bool2-v0-3/0000775000175000017500000000000013575620502013331 5ustar dwyckoff76dwyckoff76bool2-v0-3/FunctionProperties.v0000644000175000017500000025417113575574055017405 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014, Daniel Wyckoff, except for the portions so labeleled which I got from Daniel Schepler*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*The portions of this file that are marked "Schepler" were copied and pasted from Daniel Schepler's "Zorn's Lemma." Mine are marked "Wyckoff."*) Require Import Description. Require Import FunctionalExtensionality. Require Import Classical. Require Import List. Require Import Ensembles. Require Import Basics. Require Image. Require Import Omega. Require Import Setoid. Require Import Equality. Require Import EqDec. Require Import DecidableDec. (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) (*Wyckoff*) Lemma p_equal : forall {T} {P:T->Prop} {x y:T}, P x -> x = y -> P y. intros; subst; auto. Qed. (*This is helpful for FiniteMaps functionality.*) (*Wyckoff*) Definition f_pr {T U V:Type} (p:T->U->V) : (T*U) -> V := fun pr:T*U => p (fst pr) (snd pr). (*Wyckoff*) Definition f_no_pr {T U V:Type} (p:(T*U)->V) : T -> U -> V := fun (x:T) (y:U) => p (x, y). (*Wyckoff*) Lemma f_pr_f_no_pr : forall {T U V:Type} (p:(T*U)->V), p = f_pr (f_no_pr p). intros T U V p. unfold f_pr. unfold f_no_pr. apply functional_extensionality. intros pr; simpl. rewrite (surjective_pairing pr). simpl. reflexivity. Qed. (*Wyckoff*) Lemma f_no_pr_f_pr : forall {T U V:Type} (p:T->U->V), p = (f_no_pr (f_pr p)). intros T U V p. unfold f_pr. unfold f_no_pr. simpl. apply functional_extensionality. intro x. apply functional_extensionality. intro y. reflexivity. Qed. (*Wyckoff*) Definition dep_funT {T:Type} (U:T->Type) : T-> Type := fun x => U x. (*Wyckoff*) Lemma f_equal_to_existT : forall {T} {U:T->Type} {x y} (f:forall c, dep_funT U c) (pf:x = y), existT _ x (f x) = existT _ y (f y). intros; subst; auto. Qed. (*Wyckoff*) Definition dep_fun {T:Type} (U:T->Type) : Type := forall x:T, (dep_funT U x). (*Wyckoff*) Lemma f_equal_dep_app' : forall {T:Type} {x y:T} (U:T->Type) (f:dep_fun U) (pf:x = y), existT _ x (f x) = existT _ y (f y). intros; subst; auto. Qed. (*Wyckoff*) Definition prop_dep_fun {T:Type} (P:T->Prop) (U:Type): Type := forall x:T, P x -> U. (*Wyckoff*) Definition prop_dep_fun_sig {T P U} (f:prop_dep_fun P U) : {x:T|P x} -> U := fun x => f (proj1_sig x) (proj2_sig x). (*Wyckoff*) Definition prop_dep_fun2 {T:Type} (P:T->Prop) (Q:T->Prop) (U:Type) : Type := forall x:T, P x -> Q x -> U. (*Wyckoff*) Definition prop_dep_fun2_sig {T P Q U} (f:prop_dep_fun2 P Q U) : {x:T|P x /\ Q x} -> U := fun x => let (pfp, pfq) := proj2_sig x in f (proj1_sig x) pfp pfq. (*Wyckoff*) Definition prop_dep_funT {T:Type} (P:T->Prop) : Type := forall x:T, P x -> Type. (*Wyckoff*) Definition prop_dep_funT2 {T:Type} (P:T->Prop) (Q:T->Prop) : Type := forall x:T, P x -> Q x -> Type. (*Wyckoff*) Definition prop_dep_fun_dep {T} {U:T->Type} (P:forall x, U x -> Prop) V : Type := forall (x:T) (y:U x), P x y -> V. (*Wyckoff*) Definition prop_dep_fun2T {T U} (P:T->U->Prop) (V:Type) : Type := forall x y, P x y -> V. (*Wyckoff*) Definition prop_dep_fun2T' {T U} (P:T->U->Prop) (V:T->U->Type) : Type := forall x y, P x y -> V x y. (*Wyckoff*) Definition prop_dep_fun_nest2 {T:Type} {P:T->Prop} (Q:prop_dep_funT P) U : Type := forall (x:T) (px:P x), Q x px -> U. (*Wyckoff*) Lemma f_equal_dep : forall {T:Type} {P:T->Prop} {U:Type} (f:prop_dep_fun P U) (x y:T) (pfx:P x) (pfy:P y), x = y -> f x pfx = f y pfy. intros; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma f_equal_dep2T : forall {T U:Type} {P:T->U->Prop} {V:Type} (f:prop_dep_fun2T P V) (x x':T) (y y':U) (pf:P x y) (pf':P x' y'), x = x' -> y = y' -> f x y pf = f x' y' pf'. intros; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Definition prop_dep_set {T:Set} (P:T->Prop) (U:Set) : Set := forall x:T, P x -> U. (*Wyckoff*) Definition prop_dep_set_dep {T:Type} (P:T->Prop) (U:T->Set) : Type := forall x:T, P x -> U x. (*Wyckoff*) Definition forall_from_S {n U} (f:forall i (pfi:i < S n), U) : forall i (pfi:i < n), U := fun i pfi => f i (lt_S _ _ pfi). (*Wyckoff*) Definition forall_to_S {n U} (f:forall i (pfi:i < S n), U) : forall i (pfi:i < n), U := fun i pfi => f (S i) (lt_n_S _ _ pfi). (*Wyckoff*) Definition seg_fun (n:nat) (U:Type) : Type := prop_dep_fun (fun i => i < n) U. (*Wyckoff*) Definition seg_fun' (n:nat) (U:Set) : Set := prop_dep_set (fun i => i < n) U. (*Wyckoff*) Definition seg_fun0 U : seg_fun 0 U := fun _ pf => False_rect _ (lt_n_0 _ pf). (*Wyckoff*) Lemma seg_fun0_compat : forall U (f:seg_fun 0 U), seg_fun0 U = f. intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; omega. Qed. (*Wyckoff*) (*This returns a function g, such that g i _ = f i _, for i < S n -- i.e. the value at [n] is lost.*) Definition seg_fun_from_S {n:nat} {U:Type} (f:seg_fun (S n) U) : seg_fun n U := fun (m:nat) (pfm:m < n) => f m (lt_S _ _ pfm). (*Wyckoff*) Lemma seg_fun_from_S_compat : forall {U:Type} {n:nat} (f:seg_fun (S n) U) (i:nat) (pf:i < n), seg_fun_from_S f i pf = f i (lt_S _ _ pf). auto. Qed. (*Wyckoff*) Lemma seg_fun_from_S_compat' : forall {U:Type} {n:nat} (f:seg_fun (S n) U) (i:nat) (pfi:i < n) (pfi':i < S n), f i pfi' = seg_fun_from_S f i pfi. unfold seg_fun_from_S; intros; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) (* This function also derives a function of widith [n] from one of [S n], but this time, the values are translated up one shift. The returned function g obeys -- g i _ = f (S i) _. *) Definition seg_fun_to_S {n:nat} {U:Type} (f:seg_fun (S n) U) : seg_fun n U := fun i pf => f (S i) (lt_n_S _ _ pf). Lemma seg_fun_to_S_compat : forall {U:Type} {i n:nat} (f:seg_fun (S n) U) (pfi:i < n) (pfi':S i < S n), seg_fun_to_S f i pfi = f (S i) pfi'. unfold seg_fun_to_S; intros; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma seg_fun_to_S_compat' : forall {U:Type} {i n:nat} (f:seg_fun (S n) U) (pfi:S i < S n), f (S i) pfi = (seg_fun_to_S f) i (lt_S_n _ _ pfi). unfold seg_fun_to_S; intros; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma seg_fun_to_S_from_S : forall {U n} (f:seg_fun (S (S n)) U), seg_fun_to_S (seg_fun_from_S f) = seg_fun_from_S (seg_fun_to_S f). intros; apply functional_extensionality_dep. intro i. apply functional_extensionality. intro hi. unfold seg_fun_to_S, seg_fun_from_S. f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Definition seg_fun_cons {U n} (y:U) (f:seg_fun n U) : seg_fun (S n) U := fun i => match i with | 0 => fun _ => y | S i' => fun pf => let pf' := lt_S_n _ _ pf in f i' pf' end. (*Wyckoff*) Lemma seg_fun_cons0 : forall {U n} (f:seg_fun n U) (y:U) (pf:0 < S n), seg_fun_cons y f 0 pf = y. unfold seg_fun_cons; intros; auto. Qed. (*Wyckoff*) Lemma seg_fun_cons_S : forall {U n} (f:seg_fun n U) (y:U) (i:nat) (pf:S i < S n), seg_fun_cons y f (S i) pf = f i (lt_S_n _ _ pf). unfold seg_fun_cons; intros; auto. Qed. (*Wyckoff*) Lemma seg_fun_decompose_cons : forall {U n} (f:seg_fun (S n) U), f = seg_fun_cons (f 0 (lt_0_Sn _)) (seg_fun_to_S f). intros; apply functional_extensionality_dep; intro i; apply functional_extensionality; intro; simpl. unfold seg_fun_cons, seg_fun_to_S. destruct i; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma seg_fun_to_S_cons : forall {U n} y (f:seg_fun n U), seg_fun_to_S (seg_fun_cons y f) = f. intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; unfold seg_fun_to_S; rewrite seg_fun_cons_S. f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma seg_fun_from_S_cons : forall {U n} y (f:seg_fun (S n) U), seg_fun_from_S (seg_fun_cons y f) = seg_fun_cons y (seg_fun_from_S f). intros; apply functional_extensionality_dep; intro i; apply functional_extensionality; intro hi; f_equal. destruct i as [|i]; simpl. rewrite seg_fun_from_S_compat, seg_fun_cons0; auto. do 2 rewrite seg_fun_from_S_compat. rewrite seg_fun_cons_S; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Definition seg_fun_from_pred {U n} : seg_fun n U -> seg_fun (pred n) U := match n with | 0 => fun _ => seg_fun0 U | S n' => fun f => seg_fun_from_S f end. (*Wyckoff*) Definition seg_fun_to_pred {U n} : seg_fun n U -> seg_fun (pred n) U := match n with | 0 => fun _ => seg_fun0 U | S n' => fun f => seg_fun_to_S f end. (*Wyckoff*) Lemma seg_fun_to_pred_to_S : forall {U n} (f:seg_fun (S (S n)) U), seg_fun_to_S (seg_fun_to_pred f) = seg_fun_to_pred (seg_fun_to_S f). intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; unfold seg_fun_to_pred, seg_fun_to_S; auto. Qed. (*Wyckoff*) Lemma seg_fun_from_pred_from_S : forall {U n} (f:seg_fun (S (S n)) U), seg_fun_from_S (seg_fun_from_pred f) = seg_fun_from_pred (seg_fun_from_S f). intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; unfold seg_fun_from_pred, seg_fun_from_S; auto. Qed. (*Wyckoff*) Lemma seg_fun_to_pred_from_S : forall {U n} (f:seg_fun (S (S n)) U), seg_fun_to_pred (seg_fun_from_S f) = seg_fun_from_S (seg_fun_to_pred f). intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; unfold seg_fun_to_pred, seg_fun_from_S, seg_fun_to_S; auto; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma seg_fun_from_pred_to_S : forall {U n} (f:seg_fun (S (S n)) U), seg_fun_from_pred (seg_fun_to_S f) = seg_fun_to_S (seg_fun_from_pred f). intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; unfold seg_fun_from_pred, seg_fun_from_S, seg_fun_to_S; auto; f_equal; apply proof_irrelevance. Qed. (*See FunctionProperties2 for [seg_fun_[to/from|_pred] properties*) (*Wyckoff*) Definition seg_fun_dep {n} (fU:seg_fun n Type) : Type := forall i (pfi:i < n), fU i pfi. (*Wyckoff*) Definition seg_fun_dep_from_S {n} : forall {fU:seg_fun (S n) Type} (f:seg_fun_dep fU), seg_fun_dep (seg_fun_from_S fU) := fun fU f => (fun i (pfi:i < n) => f i _). (*Wyckoff*) Definition seg_fun_dep_to_S {n} : forall {fU:seg_fun (S n) Type} (f:seg_fun_dep fU), seg_fun_dep (seg_fun_to_S fU) := fun fU f => (fun i (pfi:i < n) => f (S i) (lt_n_S _ _ _)). (*Wyckoff*) Definition seg_fun_depT {S n} (U:S->Type) (f:seg_fun n S) : Type := forall i (pfi:i < n), U (f i pfi). (*Wyckoff*) Definition seg_fun_depT_to_S {T n} {U:T->Type} {f:seg_fun (S n) T} (g:seg_fun_depT U f) : seg_fun_depT U (seg_fun_to_S f) := fun i (pfi:i < n) => g (S i) (lt_n_S _ _ pfi). (*Wyckoff*) Definition seg_fun_nT n T (U:Type->nat->Type) : Type := forall i (pfi:i < n) (u:U T n), Type. (*Wyckoff*) Definition type_nat_to_S (U:Type->nat->Type) : Type -> nat -> Type:= fun T n => U T (S n). (*Wyckoff*) Definition seg_fun_nT_to_S {n T} {U:Type->nat->Type} (f:seg_fun_nT (S n) T U) : seg_fun_nT n T (type_nat_to_S U) := fun i (pfi:i < n) (u:U T (S n)) => f (S i) (lt_n_S _ _ pfi) u. (*Wyckoff*) Definition seg_fun_nT_from_S {n T} {U:Type->nat->Type} (f:seg_fun_nT (S n) T U) : seg_fun_nT n T (type_nat_to_S U) := fun i (pfi:i < n) (u:U T (S n)) => f i (lt_S _ _ pfi) u. (*Wyckoff*) Fixpoint seg_fun_app {U n} : seg_fun n U -> U -> seg_fun (S n) U := match n with | 0 => fun _ y _ _ => y | S n' => fun f y i => let f' := seg_fun_to_S f in let g := seg_fun_app f' y in match i with | 0 => fun _ => f 0 (lt_0_Sn _) | S i' => fun pfi => seg_fun_cons (f 0 (lt_0_Sn _)) g (S i') pfi end end. (*Wyckoff*) Lemma seg_fun_app_lt : forall {U n} (f:seg_fun n U) y i (pfi:i < S n) (pfi':i < n), seg_fun_app f y i pfi = f i pfi'. intros U n. induction n as [|n ih]; simpl. intros; omega. intros f y i hi hi'. destruct i as [|i]. f_equal; apply proof_irrelevance. specialize (ih (seg_fun_to_S f) y _ (lt_S_n _ _ hi) (lt_S_n _ _ hi')). rewrite ih. unfold seg_fun_to_S. f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma seg_fun_app_n : forall {U n} (f:seg_fun n U) y (pfn:n < S n), seg_fun_app f y n pfn = y. intros ? n; induction n; simpl; auto. Qed. (*Wyckoff*) Lemma seg_fun_decompose_app : forall {U n} (f:seg_fun (S n) U), f = seg_fun_app (seg_fun_from_S f) (f n (lt_n_Sn _)). intros U n f; apply functional_extensionality_dep; intro i; apply functional_extensionality; intro hi; simpl. red in hi. destruct (le_lt_eq_dec _ _ (le_S_n _ _ hi)) as [h1 | h1]. erewrite seg_fun_app_lt. unfold seg_fun_from_S; f_equal; apply proof_irrelevance. subst. erewrite seg_fun_app_n. f_equal; apply proof_irrelevance. Grab Existential Variables. assumption. Qed. (*Wyckoff*) Lemma seg_fun_from_S_app : forall {U n} (hdu:type_eq_dec U) (f:seg_fun n U) b, seg_fun_from_S (seg_fun_app f b) = f. intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; unfold seg_fun_from_S; erewrite seg_fun_app_lt. reflexivity. Qed. (*Wyckoff*) Lemma seg_fun_to_S_app : forall {U n} (hdu:type_eq_dec U) (f:seg_fun (S n) U) b, seg_fun_to_S (seg_fun_app f b) = seg_fun_app (seg_fun_to_S f) b. intros U n hdu f b; apply functional_extensionality_dep; intro; apply functional_extensionality; intro h1. unfold seg_fun_to_S. simpl. red in h1. destruct (le_lt_eq_dec _ _ (le_S_n _ _ h1)) as [h2|h2]. do 2 erewrite seg_fun_app_lt. unfold seg_fun_to_S. reflexivity. subst. do 2 erewrite seg_fun_app_n. reflexivity. Grab Existential Variables. assumption. Qed. (*Wyckoff*) Definition seg_fun_to_0S {U n} : seg_fun n U -> seg_fun n U := match n with | 0 => fun _ _ pf => False_rect _ (lt_n_0 _ pf) | S n' => fun f i pfi => seg_fun_from_S (seg_fun_cons (f 0 (lt_0_Sn _)) f) i pfi end. (*Wyckoff*) Lemma seg_fun_to_0S_0 : forall {U n} (f:seg_fun n U) (pf:0 < n), seg_fun_to_0S f 0 pf = f 0 pf. intros U n. destruct n as [|n]. intros ? h1. contradict (lt_irrefl _ h1). intros; simpl. rewrite seg_fun_from_S_compat. rewrite seg_fun_cons0. f_equal. apply proof_irrelevance. Qed. (*Wyckoff*) Lemma seg_fun_to_0S_S : forall {U n} (f:seg_fun n U) (i:nat) (pf:S i < n), seg_fun_to_0S f (S i) pf = f i (lt_trans _ _ _ (lt_n_Sn _) pf). unfold seg_fun_to_0S. destruct n as [|n]. intros ? ? h1; contradict (lt_n_0 _ h1). intros. unfold seg_fun_from_S. simpl. f_equal. apply proof_irrelevance. Qed. (*Wyckoff*) Definition forall_dep_from_S {n} {ft:seg_fun (S n) Type} (f:forall i (pfi:i < S n), ft i pfi) : forall i (pfi:i < n), seg_fun_from_S ft i pfi := fun i pfi => f i (lt_S _ _ pfi). (*Wyckoff*) Definition forall_dep_to_S {n} {ft:seg_fun (S n) Type} (f:forall i (pfi:i < S n), ft i pfi) : forall i (pfi:i < n), seg_fun_to_S ft i pfi := fun i pfi => f (S i) (lt_n_S _ _ pfi). (*Wyckoff*) Fixpoint in_seg_fun {n U} (y:U) : seg_fun n U -> Prop := match n with | 0 => fun _ => False | S n' => fun f => f 0 (lt_0_Sn _) = y \/ in_seg_fun y (seg_fun_to_S f) end. (*Decidable version of the above further below.*) (*Wyckoff*) (*Returns the first number from 0 to [(pred n)] whose value is [y]. If [y] is not in the image, then it returns [n] as a default.*) Fixpoint inv_seg_fun {U n} (pfdu:type_eq_dec U) (y:U) : seg_fun n U -> nat := match n with | 0 => fun _ => 0 | S n' => fun f => if (pfdu (f 0 (lt_0_Sn _)) y) then 0 else S (inv_seg_fun pfdu y (seg_fun_to_S f)) end. (*Wyckoff*) Lemma inv_seg_fun0: forall {U} (pfdu:type_eq_dec U) (y:U) (f:seg_fun 0 U), inv_seg_fun pfdu y f = 0. auto. Qed. (*Wyckoff*) Lemma inv_seg_fun0': forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) (pf0:0 < n), inv_seg_fun pfdu (f 0 pf0) f = 0. intros U n. destruct n as [|n]. intros; omega. intros hdu f h0; simpl. destruct (hdu (f 0 (lt_0_Sn _)) (f 0 h0)) as [|h1]; auto. contradict h1; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma inv_seg_fun_le : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) i (pfi:i < n), inv_seg_fun pfdu (f i pfi) f <= i. intros U n. induction n as [|n ih]; simpl; auto with arith. intros hdu f i hi. destruct (hdu (f 0 (lt_0_Sn n)) (f i hi)) as [| h1]; auto with arith. destruct i as [|i]. contradict h1; auto. f_equal; apply proof_irrelevance. apply le_n_S. pose proof (lt_S_n _ _ hi) as h2. specialize (ih hdu (seg_fun_to_S f) i h2). unfold seg_fun_to_S in ih. erewrite f_equal_dep. apply ih. reflexivity. Qed. (*Wyckoff*) Lemma inv_seg_fun_weak_bound : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) i (pfi:i < n) y, y = f i pfi -> inv_seg_fun pfdu y f <= i. intros U n. induction n as [|n ih]; simpl. intros; omega. intros hdu f i hi ? ?; subst. destruct (hdu (f 0 (lt_0_Sn n)) (f i hi)) as [|h1]; auto with arith. destruct i as [|i]. contradict h1; f_equal; apply proof_irrelevance. apply le_n_S. unfold seg_fun_to_S. erewrite f_equal_dep. eapply ih; auto with arith. reflexivity. Grab Existential Variables. auto with arith. Qed. (*Wyckoff*) Definition seg_pred (n:nat) (P:nat->Prop) : Type := forall i (pf:i < n), P i. (*Wyckoff*) Definition seg_pred_from_S {n:nat} {P:nat->Prop} (f:seg_pred (S n) P) : seg_pred n P := fun i (pf:i < n) => f i (lt_trans _ _ _ pf (lt_n_Sn _)). (*Wyckoff*) Definition seg_pred_type n : Type := forall i (pf:i < n), Prop. (*Wyckoff*) Definition seg_pred_type_from_S {n} (f:seg_pred_type (S n)) : seg_pred_type n := fun i (pf:i < n) => f i (lt_trans _ _ _ pf (lt_n_Sn _)). (*Wyckoff*) Definition seg_pred_dec {n:nat} (pf:seg_pred_type n) : Type := (forall i (pflt:i < n), {pf i pflt} + {~ pf i pflt}). (*Wyckoff*) Lemma seg_pred_dec_irrelevance : forall {n:nat} {pf:seg_pred_type n} (pfs pfs':seg_pred_dec pf), pfs = pfs'. intros n h1 h2 h3. apply functional_extensionality_dep. intro i. apply functional_extensionality_dep. intro hi. destruct (h2 i hi), (h3 i hi); f_equal; try apply proof_irrelevance; try contradiction. Qed. (*Wyckoff*) Definition seg_fun_eq_dec {U:Type} {n:nat} (f:seg_fun n U) : Type := forall i j (pfi:i < n) (pfj:j < n), {f i pfi = f j pfj} + {f i pfi <> f j pfj}. (*Wyckoff*) Lemma seg_fun_eq_dec_irrelevance : forall {U:Type} {n:nat} {f:seg_fun n U} (pfsf pfsf':seg_fun_eq_dec f), pfsf = pfsf'. intros U n f h1 h2. apply functional_extensionality_dep. intro i. apply functional_extensionality_dep. intro j. apply functional_extensionality_dep. intro hi. apply functional_extensionality_dep. intro hj. destruct (h1 i j hi hj), (h2 i j hi hj); f_equal; try apply proof_irrelevance; try contradiction. Qed. (*Wyckoff*) Definition seg_fun_rest {m n:nat} {U:Type} (f:seg_fun n U) (pfm:m < n) : seg_fun m U := fun (i:nat) (pfi:i < m) => f i (lt_trans _ _ _ pfi pfm). (*Wyckoff*) Lemma seg_fun_rest_compat : forall {m n:nat} {U:Type} (f:seg_fun n U) (pfm:m < n) (i:nat) (pfi:i < m), seg_fun_rest f pfm i pfi = f i (lt_trans _ _ _ pfi pfm). unfold seg_fun_rest; intros; auto. Qed. (*Wyckoff*) Definition fun_to_seg {U:Type} (f:nat->U) (n:nat) : seg_fun n U := fun (i:nat) (pfi:i < n) => f i. (*Wyckoff*) Definition fun_to_seg_compat : forall {U:Type} (f:nat->U) (n:nat) (i:nat) (pf:i < n), fun_to_seg f n i pf = f i. auto. Qed. (*Wyckoff*) Lemma seg_fun_from_S_fun_to_seg : forall {U:Type} (f:nat->U) (n:nat), seg_fun_from_S (fun_to_seg f (S n)) = fun_to_seg f n. unfold seg_fun_from_S, fun_to_seg; intros; auto. Qed. (*Wyckoff*) Definition id_seg n : seg_fun n nat := fun i _ => i. (*Wyckoff*) Definition inj_seg {U:Type} {n:nat} (f:seg_fun n U) : Prop := forall i j (pfi:i < n) (pfj:j < n), f i pfi = f j pfj -> i = j. (*Wyckoff*) Lemma inj_from_S : forall {U:Type} {n:nat} (f:seg_fun (S n) U), inj_seg f -> inj_seg (seg_fun_from_S f). intros ? ? ? h1; red in h1; red; intros; eapply h1. do 2 rewrite seg_fun_from_S_compat in H. apply H. Qed. (*Wyckoff*) Definition surj_seg {U:Type} {n:nat} (f:seg_fun n U) : Prop := forall y:U, exists i pf, f i pf = y. (*Wyckoff*) Inductive seg_fun_endo {n} : seg_fun n nat -> nat -> Prop := | seg_fun_endo_intro : forall r, forall (f:seg_fun n nat), (forall i (pf:i < n), f i pf < r) -> seg_fun_endo f r. (*Wyckoff*) Inductive inj_seg_fun_endo {n} : seg_fun n nat -> nat -> Prop := | inj_seg_fun_endo_intro : forall (f:seg_fun n nat) r, seg_fun_endo f r -> (forall i j (pfi:i < n) (pfj:j < n), f i pfi = f j pfj -> i = j) -> inj_seg_fun_endo f r. (*Wyckoff*) Inductive surj_seg_fun_endo {n} : seg_fun n nat -> nat -> Prop := | surj_seg_fun_endo_intro : forall (f:seg_fun n nat) r, seg_fun_endo f r -> (forall j, j < r -> exists s (pf:s < n), f s pf = j) -> surj_seg_fun_endo f r. (*Wyckoff*) Lemma inj_seg_fun_endo_endo : forall {n m} (f:seg_fun n nat), inj_seg_fun_endo f m -> seg_fun_endo f m. intros ? ? ? h1; destruct h1; auto. Qed. (*Wyckoff*) Lemma surj_seg_fun_endo_endo : forall {n m} (f:seg_fun n nat), surj_seg_fun_endo f m -> seg_fun_endo f m. intros ? ? ? h1; destruct h1; auto. Qed. (*Wyckoff*) Lemma seg_fun_endo_to_S : forall {n} (f:seg_fun (S n) nat), seg_fun_endo f (S n) -> seg_fun_endo (seg_fun_to_S f) (S n). intros n f h1. dependent induction h1. econstructor. intros; apply H. Qed. (*[inj] analouges are in FunctionProperties2.*) (*Wyckoff*) Definition lt_seg_fun_endo : forall {n r} (f:seg_fun n nat), seg_fun_endo f r -> forall i (pf:i < n), f i pf < r. intros ? ? ? h1; destruct h1; auto. Qed. Fixpoint in_seg_fun_dec {n U} (y:U) : seg_fun n U -> Set := match n with | 0 => fun _ => False | S n' => fun f => sumor (in_seg_fun_dec y (seg_fun_to_S f)) (f 0 (lt_0_Sn _) = y) end. Lemma in_seg_fun_dec_functional : forall {n U} (y y':U) (f g:seg_fun n U), y = y' -> f = g -> in_seg_fun_dec y f -> in_seg_fun_dec y' g. intros; subst; auto. Qed. Lemma in_seg_fun_dec0_from_S : forall {U n} (f:seg_fun (S (S n)) U), in_seg_fun_dec (f 0 (lt_0_Sn _)) (seg_fun_from_S f). intros; right; unfold seg_fun_from_S; f_equal; apply proof_irrelevance. Qed. Lemma in_seg_fun_impl_dec : forall {U n} (pfdu:type_eq_dec U) (y:U) (f:seg_fun n U), in_seg_fun y f -> in_seg_fun_dec y f. intros U n. induction n as [|n ih]; simpl. intros; contradiction. intros hdu y f h1. destruct (hdu (f 0 (lt_0_Sn n)) y) as [|h2]; subst. right; auto. left. apply ih; auto. destruct h1; subst; auto. contradict h2; auto. Qed. Lemma in_seg_fun_dec_impl : forall {U n} (y:U) (f:seg_fun n U), in_seg_fun_dec y f -> in_seg_fun y f. intros U n. induction n as [|n ih]; simpl. intros; omega. intros y f h1. destruct h1 as [h1|]; subst. apply ih in h1. right; auto. left; auto. Qed. Lemma in_seg_fun_dec_compat : forall {U n} (f:seg_fun n U) i (pfi:i < n), in_seg_fun_dec (f i pfi) f. intros U n. induction n as [|n ih]; simpl. intros; omega. intros f i h1. destruct i as [|i]. right; f_equal; apply proof_irrelevance. pose proof (lt_S_n _ _ h1) as h2. specialize (ih (seg_fun_to_S f) i h2). left; auto. unfold seg_fun_to_S in ih at 1. erewrite f_equal_dep. apply ih. reflexivity. Qed. Lemma in_seg_fun_dec_dec : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) y, (in_seg_fun_dec y f) + {in_seg_fun_dec y f -> False}. intros U n hdu. induction n as [|n ih]; simpl. intros; right; auto. intros f y. specialize (ih (seg_fun_to_S f) y). destruct ih as [h1 | h1]. left. left; auto. destruct (hdu (f 0 (lt_0_Sn _)) y) as [h2 | h2]. left; right; auto. right. intro h3. destruct h3; auto. Qed. Lemma in_seg_fun_dec' : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) y, {in_seg_fun y f} + {~in_seg_fun y f}. intros U n hdu f y. pose proof (in_seg_fun_dec_dec hdu f y) as h1. destruct h1 as [h1 | h1]. left. apply in_seg_fun_dec_impl; auto. right. intro h2. apply in_seg_fun_impl_dec in h2; auto. Qed. Lemma in_seg_fun_compat : forall {U n} (f:seg_fun n U) i (pfi:i < n), in_seg_fun (f i pfi) f. intros U n f i hi. apply in_seg_fun_dec_impl. apply in_seg_fun_dec_compat. Qed. Lemma in_seg_fun_dec_to_S : forall {U n} y (f:seg_fun (S n) U), in_seg_fun_dec y (seg_fun_to_S f) -> in_seg_fun_dec y f. intros; simpl; left; auto. Qed. Lemma in_seg_fun_to_S : forall {U n} y (f:seg_fun (S n) U), in_seg_fun y (seg_fun_to_S f) -> in_seg_fun y f. intros; simpl; right; auto. Qed. Lemma in_seg_fun_dec_to_S_neq : forall {U n} y (f:seg_fun (S n) U), in_seg_fun_dec y f -> y <> f 0 (lt_0_Sn _) -> in_seg_fun_dec y (seg_fun_to_S f). intros U n y f h1 h2; simpl; simpl in h1. destruct h1; subst; auto. contradict h2; auto. Qed. Lemma in_seg_fun_to_S_neq : forall {U n} y (f:seg_fun (S n) U), in_seg_fun y f -> y <> f 0 (lt_0_Sn _) -> in_seg_fun y (seg_fun_to_S f). intros ? ? ? ? h1 h2; simpl in h1. destruct h1; subst; auto. contradict h2; auto. Qed. (*Experiment with alternate primed version.*) Lemma in_seg_fun_from_S : forall {n U} y (f:seg_fun (S n) U), in_seg_fun_dec y (seg_fun_from_S f) -> in_seg_fun_dec y f. intros n. induction n as [|n ih]; simpl. intros; contradiction. intros U y f h1. destruct h1 as [h1|]; subst. rewrite seg_fun_to_S_from_S in h1. specialize (ih _ _ _ h1). apply in_seg_fun_dec_to_S in ih; auto. right; unfold seg_fun_from_S. f_equal; apply proof_irrelevance. Qed. Lemma in_seg_fun_dec_to_S_from_S : forall {U n} y (f:seg_fun (S n) U), in_seg_fun_dec y (seg_fun_to_S f) -> y <> f n (lt_n_Sn _) -> in_seg_fun_dec y (seg_fun_from_S f). intros U n . induction n as [|n ih]; simpl. intros; contradiction. intros y f h1 h2. destruct h1 as [h1 | h1]. apply ih in h1. left; auto. rewrite seg_fun_to_S_from_S; auto. unfold seg_fun_to_S. erewrite f_equal_dep. apply h2. reflexivity. subst. destruct (zerop n) as [|h3]; subst. unfold seg_fun_to_S in h2. contradict h2; apply f_equal_dep; reflexivity. unfold seg_fun_to_S; unfold seg_fun_to_S in ih. pose proof (in_seg_fun_dec_compat (seg_fun_from_S f) 1 (lt_n_S _ _ h3)) as h4. simpl in h4. destruct h4 as [h4 | h4]. left. rewrite seg_fun_from_S_compat in h4; auto. erewrite f_equal_dep. apply h4. reflexivity. rewrite h4. right. unfold seg_fun_from_S. erewrite f_equal_dep. reflexivity. reflexivity. Qed. Lemma in_seg_fun_from_S_neq : forall {U n} y (f:seg_fun (S n) U), in_seg_fun_dec y f -> y <> f n (lt_n_Sn _) -> in_seg_fun_dec y (seg_fun_from_S f). intros U n y f h1 h2; simpl in h1. destruct n as [|n]. destruct h1; subst; auto. contradict h2; auto. f_equal; apply proof_irrelevance. destruct h1 as [h1 | h1]. apply in_seg_fun_dec_to_S_from_S in h1; auto. subst. apply in_seg_fun_dec0_from_S. Qed. Lemma in_seg_fun_iff : forall {U n} y (f:seg_fun n U), in_seg_fun y f <-> exists i (pfi:i < n), f i pfi = y. intros U n y f; split; intro h1. revert h1. revert f y. induction n as [|n ih]; simpl. intros; contradiction. intros f y h1. destruct h1 as [|h1]; subst. exists 0, (lt_0_Sn _); auto. apply ih in h1. destruct h1 as [i [h1 ?]]; subst. exists (S i), (lt_n_S _ _ h1). unfold seg_fun_to_S; auto. destruct h1 as [i [h2 ?]]; subst. apply in_seg_fun_dec_impl. apply in_seg_fun_dec_compat. Qed. (*Wyckoff*) Lemma in_seg_fun_lt : forall {n b} (f:seg_fun n nat) y, seg_fun_endo f b -> in_seg_fun y f-> y < b. intros n b f y h2 h3. rewrite in_seg_fun_iff in h3. destruct h3 as [i [h3 ?]]; subst. destruct h2; auto. Qed. (*Wyckoff*) Lemma seg_fun_endo1_eq : forall {n} (f:seg_fun n nat), seg_fun_endo f 1 -> forall i pfi, f i pfi = 0. intros ? ? h1. inversion h1 as [? ? h2 h3]; subst. intros i hi. specialize (h2 i hi). red in h2. apply le_S_n in h2; auto with arith. Qed. (*Wyckoff*) Definition lt_surj_seg_fun_endo : forall {n r} (f:seg_fun n nat), surj_seg_fun_endo f r -> forall i (pf:i < n), f i pf < r. intros ? ? ? h1; destruct h1; intros; apply lt_seg_fun_endo; auto. Qed. (*Wyckoff*) Lemma inj_seg_fun_endo_compat : forall {n r} (f:seg_fun n nat), seg_fun_endo f r -> inj_seg f -> inj_seg_fun_endo f r. intros; constructor; auto. Qed. (*Wyckoff*) Lemma seg_fun_to_ind : forall {U n} (f:seg_fun n U) (P:forall m (f:seg_fun m U), Prop), P 0 (seg_fun0 U) -> (forall (b:U) (m:nat) (f:seg_fun m U), P m f -> P (S m) (seg_fun_cons b f)) -> P n f. intros U n. induction n as [|n ih]; simpl. intros f P h1 h2. erewrite seg_fun0_compat in h1. apply h1. intros f P h1 h2. rewrite seg_fun_decompose_cons. apply h2; auto. apply ih; auto. Qed. (*Wyckoff*) Lemma seg_fun_from_ind : forall {U n} (f:seg_fun n U) (P:forall m (f:seg_fun m U), Prop), P 0 (seg_fun0 U) -> (forall (b:U) (m:nat) (f:seg_fun m U), P m f -> P (S m) (seg_fun_app f b)) -> P n f. intros U n. induction n as [|n ih]; simpl. intros f P h2 h3. erewrite <- seg_fun0_compat; auto. intros f P h1 h2. rewrite seg_fun_decompose_app. apply h2. apply ih; auto. Qed. (*Wyckoff*) Definition seq_fun (n w:nat) (U:Type) : Type := forall m:nat, n <= m < n+w -> U. (*Wyckoff*) Definition seg_fun_to_seq {m:nat} {U:Type} (f:seg_fun m U) : seq_fun 0 m U := (fun m pf => let (_, pf') := pf in (f m pf')). (*See [ArithUtilities] for [seg_S_tl], [seq_fun_from_S], and other derived lemmas.*) (*Wyckoff*) Definition pred_impl {T} : relation (T->Prop) := fun P Q => forall x, P x -> Q x. (*Wyckoff*) Definition impl_pred {T} : relation (T->Prop) := fun P Q => pred_impl Q P. (*Wyckoff*) Definition iff_pred {T} : relation (T->Prop) := fun P Q => pred_impl P Q /\ impl_pred P Q. (*Wyckoff*) Lemma iff_pred1 : forall {T} (P Q:T->Prop), iff_pred P Q -> pred_impl P Q. intros ? ? ? h1; destruct h1; auto. Qed. (*Wyckoff*) Lemma iff_pred2 : forall {T} (P Q:T->Prop), iff_pred P Q -> impl_pred P Q. intros ? ? ? h1; destruct h1; auto. Qed. (*Wyckoff*) Lemma iff_pred_iff : forall {T} (P Q:T->Prop), iff_pred P Q <-> (forall x, P x <-> Q x). intros; red; split; intro h1; red in h1. destruct h1 as [h1 h2]. intro x. red in h1. do 2 red in h2. specialize (h1 x); specialize (h2 x); split; auto. red; split; red; try red; intros x hx; specialize (h1 x); tauto. Qed. (*Wyckoff*) Lemma refl_pred_impl : forall T, Reflexive (@pred_impl T). intro; red; intro; red; auto. Qed. (*Wyckoff*) Lemma trans_pred_impl : forall T, Transitive (@pred_impl T). intro; red; intro; red; auto. Qed. (*Wyckoff*) Lemma refl_impl_pred : forall T, Reflexive (@impl_pred T). intro; red; intro; do 2 red; auto. Qed. (*Wyckoff*) Lemma trans_impl_pred : forall T, Transitive (@impl_pred T). intro; red; intro; do 2 red; auto. Qed. (*Wyckoff*) Lemma refl_iff_pred : forall T, Reflexive (@iff_pred T). intros; red; intros; red; split; [apply refl_pred_impl | apply refl_impl_pred]. Qed. (*Wyckoff*) Lemma sym_iff_pred : forall T, Symmetric (@iff_pred T). intros; red; intro; red; intros ? h1; destruct h1 as [h1 h2]; red in h1, h2; split; [apply refl_pred_impl | apply refl_impl_pred]; auto. Qed. (*Wyckoff*) Lemma trans_iff_pred : forall T, Transitive (@iff_pred T). intro; red; intros ? ? ? h1 h2; destruct h1 as [h1a h1b], h2 as [h2a h2b]; red; split. eapply trans_pred_impl. apply h1a. auto. eapply trans_impl_pred. apply h1b. auto. Qed. (*Wyckoff*) Add Parametric Relation T : (T->Prop) (@impl_pred T) reflexivity proved by (@refl_impl_pred T) transitivity proved by (@trans_impl_pred T) as impl_pred_rel. (*Wyckoff*) Add Parametric Relation T : (T->Prop) (@pred_impl T) reflexivity proved by (@refl_pred_impl T) transitivity proved by (@trans_pred_impl T) as pred_impl_rel. (*Wyckoff*) Add Parametric Relation T : (T->Prop) (@iff_pred T) reflexivity proved by (@refl_iff_pred T) symmetry proved by (@sym_iff_pred T) transitivity proved by (@trans_iff_pred T) as iff_pred_rel. (*Wyckoff*) Definition eq_subst_l {T} {x x':T} (pfe:x = x') (f:forall y, x = y) : forall y, x' = y. intros; subst; auto. Qed. (*Wyckoff*) Definition eq_subst_r {T} {y y':T} (pfe:y = y') (f:forall x, x = y) : forall x, x = y. intros; subst; auto. Qed. (*Wyckoff*) Definition neq_subst_l {T} {x x':T} (pfe:x = x') (f:forall y, x <> y) : forall y, x' <> y. intros; subst; auto. Qed. (*Wyckoff*) Definition neq_subst_r {T} {y y':T} (pfe:y = y') (f:forall x, x <> y) : forall x, x <> y. intros; subst; auto. Qed. (*Wyckoff*) Lemma iff_pred_eq_l : forall {T} {x x':T} (pfe:x = x'), iff_pred (fun c => x = c) (fun c => x' = c). intros; subst; reflexivity. Qed. (*Wyckoff*) Lemma iff_pred_eq_r : forall {T} {y y':T} (pfe:y = y'), iff_pred (fun c => c = y) (fun c => c = y'). intros; subst; reflexivity. Qed. (*Wyckoff*) Lemma iff_pred_neq_l : forall {T} {x x':T} (pfe:x = x'), iff_pred (fun c => x <> c) (fun c => x' <> c). intros; subst; reflexivity. Qed. (*Wyckoff*) Lemma iff_pred_neq_r : forall {T} {y y':T} (pfe:y = y'), iff_pred (fun c => c <> y) (fun c => c <> y'). intros; subst; reflexivity. Qed. (*Wyckoff*) Definition proj1_sig_fun {T U:Type} {P:U->Prop} (f:T->{x:U | P x}) : T->U := fun x => proj1_sig (f x). (*Wyckoff*) Definition proj1_sig_fun_compat : forall {T U:Type} {P:U->Prop} (f:T->{x:U | P x}) (x:T), proj1_sig_fun f x = proj1_sig (f x). unfold proj1_sig_fun; intros; auto. Qed. (*Wyckoff*) Definition proj1_sig_fun1 {T U} {P:T->Prop} (Q:{x:T | P x} -> U) (pfp:forall x:T, P x) : T -> U := fun x => Q (exist P x (pfp x)). (*Wyckoff*) Definition unit_set_pred {T:Type} (U:T->Set) : Prop := forall (x:T) (a b:U x), a = b. (*Wyckoff*) Definition fexcl_mid {T:Type} (P:T->Prop) : T -> Set := (fun x => {P x} + {~P x}). (*Wyckoff*) Lemma unit_set_pred_fexcl_mid : forall {T:Type} (P:T->Prop), unit_set_pred (fexcl_mid P). intros. red. intros x a b. destruct a, b; try contradiction; auto; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Definition inj_dep {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) : Prop := forall (x y:T) (pfx:P x) (pfy:P y), f x pfx = f y pfy -> x = y. (*Wyckoff*) Definition inj_dep_contra {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) : Prop := forall (x y:T) (pfx:P x) (pfy:P y), x <> y -> f x pfx <> f y pfy. (*Wyckoff*) Lemma inj_dep_impl_contra : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U), inj_dep f -> inj_dep_contra f. intros ? ? ? ? h1; red in h1; red; intros ? ? ? ? ? h5; apply h1 in h5; auto. Qed. (*Wyckoff*) Lemma contra_impl_inj_dep : forall {T U:Type} {P:T->Prop} (pfdt:type_eq_dec T) (f:prop_dep_fun P U), inj_dep_contra f -> inj_dep f. intros ? ? ? hdt f h1; red in h1; red; intros ? ? ? ? h5; destruct (hdt x y) as [|h6]; auto; eapply h1 in h6; rewrite h5 in h6 at 1. contradict h6; reflexivity. Qed. (*Wyckoff*) Lemma inj_seg_fun0 : forall U, inj_dep (seg_fun0 U). intro; red; intros; omega. Qed. (*Wyckoff*) Definition surj_dep {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) : Prop := forall y:U, exists (x:T) (pfx:P x), f x pfx = y. (*Wyckoff*) Definition bij_dep {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) : Prop := inj_dep f /\ surj_dep f. (*Wyckoff*) Lemma bij_impl_inj_dep : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U), bij_dep f -> inj_dep f. intros ? ? ? ? h1; destruct h1; auto. Qed. (*Wyckoff*) Lemma bij_impl_surj_dep : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U), bij_dep f -> surj_dep f. intros ? ? ? ? h1; destruct h1; auto. Qed. (*Wyckoff*) (*Dependent invertibility.*) Inductive inv_dep {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) : Prop := | inv_dep_intro : forall (g:U->T) (pfp:forall u, P (g u)), (forall (x:T) (pfx:P x), g (f x pfx) = x) -> (forall (y:U), f (g y) (pfp y) = y) -> inv_dep f. (*Wyckoff*) Definition fun_in_dep_fun_compose {T U V:Type} {P:T->Prop} (f:prop_dep_fun P U) (g:U->V) : prop_dep_fun P V := fun x pf => g (f x pf). (*Wyckoff*) Definition proj1_sig_fun_dep {T U:Type} {P:T->Prop} {Q:U->Prop} (f:prop_dep_fun P {x:U | Q x}) : prop_dep_fun P U := fun_in_dep_fun_compose f (@proj1_sig _ _). (*See [FunctionProperties2] for more on [inv_dep].*) (*Wyckoff*) Lemma f_equal_prop_dep : forall {T U:Type} {P:T->Prop} {x y:T} (f:prop_dep_fun P U) (pfeq:x = y)(pfx:P x) (pfy:P y), f x pfx = f y pfy. intros ? ? ? ? ? ? ? h1 h2; subst. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. (*Wyckoff*) Lemma f_equal_prop_dep2 : forall {T U:Type} {P Q:T->Prop} {x y:T} (f:prop_dep_fun2 P Q U) (pfeq:x = y)(pfpx:P x) (pfpy:P y) (pfqx:Q x) (pfqy:Q y), f x pfpx pfqx = f y pfpy pfqy. intros ? ? ? ? ? ? ? ? h1 h2 h3 h4; subst. pose proof (proof_irrelevance _ h1 h2); subst; auto. pose proof (proof_irrelevance _ h3 h4); subst; auto. Qed. (*Wyckoff*) Lemma f_equal_seg : forall {U:Type} {n:nat} (f:seg_fun n U) (m m':nat) (pf: m < n) (pf': m' < n), m = m' -> f m pf = f m' pf'. intros; apply f_equal_prop_dep; auto. Qed. (*Wyckoff*) Lemma f_equal_seg_fun_depT : forall {S n} {U:S -> Type} {f:seg_fun n S} (g:seg_fun_depT U f) i (pfi pfi':i < n) (pfe:f i pfi = f i pfi'), existT _ (f i pfi) (g i pfi) = existT _ (f i pfi') (g i pfi'). intros S n; induction n as [|n ih]; simpl; intros U f g i h1 h2 h3; try omega. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. (*Wyckoff*) Definition id_prop_dep {T:Type} (P:T->Prop) : prop_dep_fun P T := fun x _ => x. (*Wyckoff*) Definition id_prop_dep2 {T:Type} (P:T->Prop) (Q:T->Prop) : prop_dep_fun2 P Q T := fun x _ _ => x. (*Wyckoff*) Definition id_prop_dep_sig {T:Type} (P:T->Prop) : {x:T|P x} -> T := prop_dep_fun_sig (id_prop_dep P). (*Wyckoff*) Definition id_prop_dep2_sig {T:Type} (P:T->Prop) (Q:T->Prop) : {x:T | P x /\ Q x} -> T := prop_dep_fun2_sig (id_prop_dep2 P Q). (*Wyckoff*) Definition id_prop_dep_sig' {T} (P:T->Prop) : prop_dep_fun P {x:T | P x} := fun i pf => exist P i pf. (*Wyckoff*) Lemma prop_dep_fun_ext : forall {T U:Type} (P:T->Prop) (g g':prop_dep_fun P U), (forall (x:T) (pf:P x), g x pf = g' x pf) -> g = g'. intros T U P g g' h1. apply functional_extensionality_dep. intro x. apply functional_extensionality. intro h2. apply h1. Qed. (*Wyckoff*) Lemma prop_dep_fun2_ext : forall {T U:Type} (P Q:T->Prop) (g g':prop_dep_fun2 P Q U), (forall (x:T) (pfp:P x) (pfq:Q x), g x pfp pfq = g' x pfp pfq) -> g = g'. intros T U P Q g g' h1. apply functional_extensionality_dep. intro x. apply functional_extensionality. intro h2. apply functional_extensionality. intro h3. apply h1. Qed. (*Wyckoff*) Lemma prop_dep_set_dep_ext : forall {T:Type} (P:T->Prop) (U:T->Set) (f g:prop_dep_set_dep P U), unit_set_pred U -> f = g. intros ? ? ? ? ? h1. apply functional_extensionality_dep. intro. apply functional_extensionality. intro. apply h1. Qed. (*Wyckoff*) Lemma seg_fun_ext : forall {U:Type} (n:nat) (g g':seg_fun n U), (forall (m:nat) (pf:m < n), g m pf = g' m pf) -> g = g'. intros; apply prop_dep_fun_ext; auto. Qed. (*Wyckoff*) Definition prop_dep_prop {T:Type} (P:T->Prop) : Type := forall x:T, P x -> Prop. (*Wyckoff*) Definition prop_dep_prop_u {T:Type} (P:T->Prop) (U:Type) : Type := forall x:T, P x -> U -> Prop. (* Possible name clash with something in [FunctionalExtensionality], which should be imported before this file or something. *) (*Wyckoff*) Lemma equal_f : forall {T U:Type} {f g:T->U}, f = g -> forall (x:T), f x = g x. intros; subst; auto. Qed. Lemma equal_f_dep : forall {T U} {P:T->Prop} (f g:prop_dep_fun P U), f = g -> forall x (pf:P x), f x pf = g x pf. intros; subst; auto. Qed. (*Wyckoff*) Definition not_pred {T:Type} (P:T->Prop) := fun x => ~ P x. (*Schepler*) Definition injective {X Y:Type} (f:X->Y) := forall x1 x2:X, f x1 = f x2 -> x1 = x2. (*Schepler*) Definition surjective {X Y:Type} (f:X->Y) := forall y:Y, exists x:X, f x = y. (*Schepler*) Definition bijective {X Y:Type} (f:X->Y) := injective f /\ surjective f. (*Wyckoff*) Lemma bij_impl_inj : forall {X Y:Type} (f:X->Y), bijective f -> injective f. intros ? ? ? h1; destruct h1; auto. Qed. (*Wyckoff*) Lemma bij_impl_surj : forall {X Y:Type} (f:X->Y), bijective f -> surjective f. intros ? ? ? h1; destruct h1; auto. Qed. Notation Injective := injective. (*Wyckoff*) Definition inj_contra {X Y:Type} (f:X->Y) : Prop := forall x1 x2:X, x1 <> x2 -> f x1 <> f x2. (*Wyckoff*) Lemma inj_impl_contra : forall {X Y:Type} (f:X->Y), injective f -> inj_contra f. intros; red; intros ? ? h2; contradict h2; auto. Qed. (*Wyckoff*) Lemma contra_impl_inj : forall {X Y:Type} (pfdt:type_eq_dec X) (f:X->Y), inj_contra f -> injective f. intros ? ? ? ? h1; red; intros x y h2. destruct (pfdt x y) as [|h3]; auto. apply h1 in h3; contradiction. Qed. (*Wyckoff*) Lemma inj_image_fun_prop_compat : forall {T U:Type} (f:T->U), injective f <-> Image.injective _ _ f. intros; split; auto. Qed. (*Schepler*) Inductive invertible {X Y:Type} (f:X->Y) : Prop := | intro_invertible: forall g:Y->X, (forall x:X, g (f x) = x) -> (forall y:Y, f (g y) = y) -> invertible f. (*Schepler*) Lemma unique_inverse: forall {X Y:Type} (f:X->Y), invertible f -> exists! g:Y->X, (forall x:X, g (f x) = x) /\ (forall y:Y, f (g y) = y). Proof. intros. destruct H. exists g. red; split. tauto. intros. destruct H1. extensionality y. transitivity (g (f (x' y))). rewrite H2. reflexivity. rewrite H. reflexivity. Qed. (*Schepler*) Definition function_inverse {X Y:Type} (f:X->Y) (i:invertible f) : { g:Y->X | (forall x:X, g (f x) = x) /\ (forall y:Y, f (g y) = y) } := (constructive_definite_description _ (unique_inverse f i)). (*Schepler*) Lemma bijective_impl_invertible: forall {X Y:Type} (f:X->Y), bijective f -> invertible f. Proof. intros. destruct H. assert (forall y:Y, {x:X | f x = y}). intro. apply constructive_definite_description. pose proof (H0 y). destruct H1. exists x. red; split. assumption. intros. apply H. transitivity y. auto with *. auto with *. pose (g := fun y:Y => proj1_sig (X0 y)). pose proof (fun y:Y => proj2_sig (X0 y)). simpl in H1. exists g. intro. apply H. unfold g; rewrite H1. reflexivity. intro. unfold g; apply H1. Qed. (*Schepler*) Lemma invertible_impl_bijective: forall {X Y:Type} (f:X->Y), invertible f -> bijective f. Proof. intros. destruct H. split. red; intros. congruence. red; intro. exists (g y). apply H0. Qed. (*Wyckoff*) Lemma invertible_impl_injective: forall {X Y:Type} (f:X->Y), invertible f -> injective f. intros ? ? ? h1; apply invertible_impl_bijective in h1; destruct h1; auto. Qed. (*Wyckoff*) Lemma invertible_impl_surjective: forall {X Y:Type} (f:X->Y), invertible f -> surjective f. intros ? ? ? h1; apply invertible_impl_bijective in h1; destruct h1; auto. Qed. (*Wyckoff*) Lemma invertible_impl_inv_invertible: forall {X Y:Type} (f:X->Y) (pf:invertible f), invertible (proj1_sig (function_inverse _ pf)). intros X Y f h1. destruct (function_inverse f h1) as [g h2]. simpl. destruct h2 as [h2l h2r]. apply (intro_invertible g f h2r h2l). Qed. (*Wyckoff*) Lemma distinct_functions_differ_at_point: forall {X Y:Type} (f g:X -> Y), f <> g -> (exists x:X, f x <> g x). intros X Y f g h1. apply NNPP. intro h2. pose proof (not_ex_all_not _ _ h2) as h3. simpl in h3. assert (h4:forall x:X, f x = g x). intro x. specialize (h3 x). tauto. pose proof (functional_extensionality _ _ h4). contradiction. Qed. (*Wyckoff*) Lemma bij_id : forall T, @bijective T T id. intro T. red; split; red; intros; auto. exists y; auto. Qed. (*Wyckoff*) Lemma invertible_id : forall T, @invertible T T id. intro; apply bijective_impl_invertible. apply bij_id. Qed. (*Wyckoff*) Lemma inj_id_dep : forall {T} (P:T->Prop), inj_dep (id_prop_dep P). unfold id_prop_dep; intros; red; auto. Qed. (*Wyckoff*) Lemma bij_compose : forall {T U V:Type} (f:T->U) (g:U->V), bijective f -> bijective g -> bijective (compose g f). intros T U V f g h1 h2. pose proof h1 as h1'. apply bijective_impl_invertible in h1'. destruct h1 as [h1l h1r]. destruct h2 as [h2l h2r]. red. split. red. intros x y h3. apply h2l in h3. apply h1l; auto. red. intros y. specialize (h2r y). destruct h2r as [x h2r]. subst. exists ((proj1_sig (function_inverse _ h1')) x). unfold compose. rewrite (match (proj2_sig (function_inverse f h1')) with | conj _ pf => pf end). reflexivity. Qed. (*Wyckoff*) Lemma invertible_compose : forall {T U V:Type} (f:T->U) (g:U->V), invertible f -> invertible g -> invertible (compose g f). intros; apply bijective_impl_invertible; apply bij_compose; apply invertible_impl_bijective; auto. Qed. (*Wyckoff*) Lemma inj_compose : forall {T U V:Type} (f:T->U) (g:U->V), injective f -> injective g -> injective (compose g f). intros T U V f g h1 h2. unfold compose. red. intros x y h3. apply h2 in h3. apply h1 in h3. assumption. Qed. (*Wyckoff*) Lemma inj_id : forall (T:Type), injective (id (A:=T)). intro; red; intros; simpl; auto. Qed. (*Wyckoff*) Lemma inj_dep_sig : forall {T U} {P:T->Prop} (f:prop_dep_fun P U), inj_dep f-> injective (prop_dep_fun_sig f). unfold prop_dep_fun_sig; intros T U P f h1. red in h1; red. intros x y h2; destruct x as [x hx], y as [y hy]; simpl in h2. apply h1 in h2; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma surj_dep_sig : forall {T U} {P:T->Prop} (f:prop_dep_fun P U), surj_dep f-> surjective (prop_dep_fun_sig f). unfold prop_dep_fun_sig; intros T U P f h1; red in h1; red; intro y; specialize (h1 y). destruct h1 as [x [hx ?]]; subst. exists (exist _ _ hx); simpl; auto. Qed. (*Wyckoff*) Lemma bij_dep_sig : forall {T U} {P:T->Prop} (f:prop_dep_fun P U), bij_dep f-> bijective (prop_dep_fun_sig f). intros ? ? ? ? h1; destruct h1; red; split; [apply inj_dep_sig; auto | apply surj_dep_sig; auto]. Qed. (*Wyckoff*) Definition false_fun U : False -> U := fun x => False_rect _ x. (*wyckoff*) Lemma false_fun_compat : forall {U} (f:False -> U), f = false_fun U. intros; apply functional_extensionality; intro; contradiction. Qed. (*Wyckoff*) Lemma inj_false : forall U, injective (false_fun U). intros; red; simpl; intro; contradiction. Qed. (*Wyckoff*) Lemma inj_fun_false : forall {U} (f:U->False), injective f. intros ? f; red. intros x ?. pose (f x); contradiction. Qed. (*Wyckoff*) Lemma surj_fun_false : forall {U} (f:U->False), surjective f. intros U f; red. intro; contradiction. Qed. (*Wyckoff*) Lemma bij_fun_false : forall {U} (f:U->False), bijective f. intros; red; split; [apply inj_fun_false | apply surj_fun_false]. Qed. (*Wyckoff*) Lemma inv_fun_false : forall {U} (f:U->False), invertible f. intros; apply bijective_impl_invertible; apply bij_fun_false. Qed. (*Wyckoff*) Lemma inj_false_fun_false : forall U, injective (false_fun (U->False)). intro; red; intros; contradiction. Qed. (*Wyckoff*) (*iterate unary operator*) Fixpoint iter_uo {T:Type} (uo:T->T) (x:T) (n:nat) : T := match n with | 0 => x | S n => (uo (iter_uo uo x n)) end. (*Wyckoff*) Lemma snd_iter_uo_const : forall {T U:Type} (f:T*U->T*U) (u:U), (forall pr:T*U, snd (f pr) = u) -> forall (pr:T*U) (n:nat), snd (iter_uo f pr (S n)) = u. intros T U f u h1 pr n. revert pr h1. revert u f. induction n as [|n ih]; simpl. intros u f pr h1. apply h1; auto. intros u f pr h1. specialize (ih _ _ pr h1). simpl in ih. apply h1. Qed. Section Relations. (*Wyckoff*) Lemma lt_trans' : Transitive lt. red; intros; eapply lt_trans. apply H. apply H0. Qed. (*Wyckoff*) Definition irrefl {T:Type} (R:T->T->Prop) : Prop := forall x:T, ~ R x x. (*Wyckoff*) Inductive Strict_Order {T:Type} (R:T->T->Prop) := strict_order_intro : Transitive R -> irrefl R -> Strict_Order R. (*Schepler*) Definition total_strict_order {T:Type} (R:relation T) : Prop := forall x y:T, R x y \/ x = y \/ R y x. (*Wyckoff*) Definition tso_dec {T:Type} (R:relation T) : Type := forall x y:T, {R x y} + {x = y} + {R y x}. (*Wyckoff*) Definition Rlem_dec {T:Type} (R:relation T) : Type := forall x y:T, {R x y} + {~R x y}. (*Wyckoff*) Lemma lt_irrefl_neq : forall {T:Type} (R:relation T) (a b:T), irrefl R -> R a b -> a <> b. intros; intro; subst; apply H in H0; auto. Qed. (*Wyckoff*) Lemma gt_irrefl_neq : forall {T:Type} (R:relation T) (a b:T), irrefl R -> R b a-> a <> b. intros; intro; subst; apply H in H0; auto. Qed. (*Wyckoff*) Lemma tso_dec_total_strict_order : forall {T:Type} (R:relation T), tso_dec R -> total_strict_order R. intros T R h1. red in h1; red. intros x y. specialize (h1 x y). destruct h1 as [h1 | h1]. destruct h1 as [h1 | h1]; subst. left; auto. right; left; auto. right; auto. Qed. (*Wyckoff*) Lemma irrefl_tso_dec_eq_dec : forall {T:Type} (R:relation T), irrefl R -> tso_dec R -> type_eq_dec T. intros T R h0 h1. red in h1; red. intros x y. specialize (h1 x y). destruct h1 as [[h1 | h1] | h1]. right. intro; subst. apply h0 in h1; auto. left; auto. right. intro; subst. apply h0 in h1; auto. Qed. (*Wyckoff*) Inductive sortable {T:Type} (R:relation T) : Prop := | sort_intro : total_strict_order R -> Transitive R -> irrefl R -> sortable R. (*Wyckoff*) Inductive sort_dec {T:Type} (R:relation T) : Type := | sort_dec_intro : tso_dec R -> Transitive R -> irrefl R -> sort_dec R. (*Wyckoff*) Lemma sort_dec_sortable : forall {T:Type} (R:relation T), sort_dec R -> sortable R. intros T R h1. destruct h1; constructor; auto. apply tso_dec_total_strict_order; auto. Qed. (*Wyckoff*) Lemma sort_dec_tso_dec : forall {T:Type} (R:relation T), sort_dec R -> tso_dec R. intros ? ? h1. destruct h1; auto. Defined. (*Wyckoff*) Lemma sort_dec_trans : forall {T:Type} (R:relation T), sort_dec R -> Transitive R. intros ? ? h1. destruct h1; auto. Defined. (*Wyckoff*) Lemma sort_dec_irrefl : forall {T:Type} (R:relation T), sort_dec R -> irrefl R. intros ? ? h1. destruct h1; auto. Defined. (*Wyckoff*) Lemma sort_dec_so : forall {T:Type} (R:relation T), sort_dec R -> Strict_Order R. intros ? ? h1. destruct h1; constructor; auto. Defined. (*Wyckoff*) Lemma sort_dec_eq : forall {T:Type} (R:relation T) (pf pf':sort_dec R), pf = pf'. intros T R h1 h2. destruct h1 as [ha hb hc], h2 as [ha' hb' hc']. assert (ha = ha'). apply functional_extensionality_dep. intro x. apply functional_extensionality_dep. intro y. destruct (ha x y) as [h1 | h1], (ha' x y) as [h2 | h2]; f_equal; try apply proof_irrelevance. destruct h1 as [h1 | h1], h2 as [h2 | h2]; f_equal; try apply proof_irrelevance; subst. pose proof (hc _ h1); contradiction. pose proof (hc _ h2); contradiction. destruct h1 as [h1 | h1]; subst. pose proof (hb _ _ _ h1 h2) as h3. apply hc in h3. contradiction. pose proof (hc _ h2); contradiction. destruct h2 as [h2 | h2]; subst. pose proof (hb _ _ _ h1 h2) as h3. pose proof (hc _ h3); contradiction. pose proof (hc _ h1); contradiction. subst. f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma tso_irrefl_eq_dec : forall {T:Type} (R:relation T), tso_dec R -> irrefl R -> type_eq_dec T. intros T R h1 h3. red in h1, h3; red. intros x y. specialize (h1 x y). destruct h1 as [h1 | h1]. destruct h1 as [h1 |]; subst. right. intro; subst. apply h3 in h1; auto. left; auto. right. intro; subst. apply h3 in h1; auto. Qed. (*Wyckoff*) Lemma sort_dec_eq_dec : forall {T:Type} (R:relation T), sort_dec R -> type_eq_dec T. intros T R h1. destruct h1 as [h1 ? ?]; eapply tso_irrefl_eq_dec; try apply h1; auto. Qed. (*Wyckoff*) Lemma sort_dec_tso_dec_eq : forall {T:Type} (R:relation T) (pfso:sort_dec R) (pft:tso_dec R), sort_dec_tso_dec R pfso = pft. intros T R h1 h2. unfold sort_dec_tso_dec. destruct h1 as [h1 h3 h4]. apply functional_extensionality_dep. intro x. apply functional_extensionality_dep. intro y. destruct (h1 x y) as [h5 | h5], (h2 x y) as [h6 | h6]; f_equal; try destruct h5 as [h5 | h5]; try destruct h6 as [h6 | h6]; f_equal; subst; try apply proof_irrelevance. pose proof (h4 _ h5); contradiction. pose proof (h4 _ h6); contradiction. pose proof (h3 _ _ _ h5 h6) as h7. apply h4 in h7; contradiction. pose proof (h4 _ h6); contradiction. pose proof (h3 _ _ _ h5 h6) as h7. apply h4 in h7; contradiction. pose proof (h4 _ h5); contradiction. Qed. (*Wyckoff*) (*Uses LEM*) Lemma tso_dec_compat : forall {T:Type} (R:relation T), total_strict_order R -> tso_dec R. intros T R h1 x y. red in h1. specialize (h1 x y). apply or_to_sumbool in h1. destruct h1 as [h1 | h1]. left; left; auto. apply or_to_sumbool in h1. destruct h1 as [h1 | h1]. left; right; auto. right; auto. Qed. (*Wyckoff*) (*(R Weak decidable*) Definition Rwd {T:Type} (R:relation T) : T -> T-> Type := (fun x y => {R x y} + {x = y}). (*Wyckoff*) Lemma R_impl_wd : forall {T:Type} (R:relation T) (x y:T), R x y -> Rwd R x y. intros; left; auto. Qed. (*Wyckoff*) Lemma eq_impl_wd : forall {T:Type} (R:relation T) (x y:T), x = y -> Rwd R x y. intros; right; auto. Qed. (*Wyckoff*) Lemma refl_Rwd : forall {T:Type} (R:relation T) (x:T), Rwd R x x. intros; right; auto. Qed. (*Wyckoff*) Definition Rw {T:Type} (R:relation T) : T -> T-> Prop := (fun x y => R x y \/ x = y). (*Wyckoff*) Lemma R_impl_w : forall {T:Type} (R:relation T) (x y:T), R x y -> Rw R x y. intros; left; auto. Qed. (*Wyckoff*) Lemma eq_impl_w : forall {T:Type} (R:relation T) (x y:T), x = y -> Rw R x y. intros; right; auto. Qed. (*Wyckoff*) Lemma refl_Rw : forall {T:Type} (R:relation T), Reflexive (Rw R). intros; right; auto. Qed. (*Wyckoff*) Lemma trans_Rw : forall {T:Type} (R:relation T), Transitive R -> Transitive (Rw R). intros T R ht; red; intros x y z h1 h2. red in h1, h2; red. destruct h1 as [h1|], h2 as [h2|]; subst. left; eapply ht. apply h1. assumption. left; auto. left; auto. right; auto. Qed. (*Wyckoff*) Lemma trans_R_Rw : forall {T:Type} (R:relation T), Transitive R -> forall x y z:T, R x y -> Rw R y z -> R x z. intros T R h1 x y z h2 h3. red in h3. destruct h3 as [h3|]; subst. eapply h1. apply h2. assumption. assumption. Qed. (*Wyckoff*) Lemma trans_Rw_R : forall {T:Type} (R:relation T), Transitive R -> forall x y z:T, Rw R x y -> R y z -> R x z. intros T R h1 x y z h2 h3. red in h2. destruct h2 as [h2|]; subst. eapply h1. apply h2. assumption. assumption. Qed. (*Wyckoff*) Lemma antisym_Rw : forall {T:Type} (R:relation T), Strict_Order R -> forall (x y:T), Rw R x y -> Rw R y x -> x = y. intros T R h1 x y h2 h3. red in h2, h3. destruct h2 as [h2|], h3 as [h3|]; auto. destruct h1 as [h1 h4]. specialize (h1 _ _ _ h2 h3). apply h4 in h1; contradiction. Qed. (*Wyckoff*) Lemma Rwd_impl_w : forall {T:Type} (R:relation T) (x y:T), Rwd R x y -> Rw R x y. intros T R x y h1. red in h1; red. destruct h1; [left; auto | right; auto]. Qed. (*Wyckoff*) Lemma Rw_le_iff : forall (m n:nat), Rw lt m n <-> m <= n. unfold Rw; intros; split; omega. Qed. (*Wyckoff*) Definition Rwlem_dec {T:Type} (R:relation T) : Type := forall x y:T, {Rw R x y} + {~Rw R x y}. (*Wyckoff*) Lemma R_iffn : forall {T:Type} (R:relation T) (pfso:sort_dec R) (a b:T), R a b <-> ~ Rw R b a. intros T R h1 a b; split; intro h2. intro h3. pose proof (trans_Rw_R _ (sort_dec_trans _ h1) _ _ _ h3 h2) as h4. apply (sort_dec_irrefl _ h1) in h4; auto. destruct (sort_dec_tso_dec _ h1 a b) as [h3 | h3]. destruct h3; auto; subst; auto. contradict h2; right; auto. contradict h2; left; auto. Qed. (*Wyckoff*) Lemma Rw_iffn : forall {T:Type} (R:relation T) (pfso:sort_dec R) (a b:T), Rw R a b <-> ~ R b a. intros T R h1 a b; split; intro h2. intro h3. pose proof (trans_Rw_R _ (sort_dec_trans _ h1) _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; auto. destruct (sort_dec_tso_dec _ h1 a b) as [h3 | h3]. destruct h3; auto; subst; auto. left; auto. right; auto. contradiction. Qed. (*Wyckoff*) Lemma sort_dec_Rlem_dec : forall {T:Type} (R:relation T), sort_dec R -> Rlem_dec R. intros T R h1. destruct h1 as [h1 h2 h3]. red; intros x y. destruct (h1 x y) as [h4 | h4]. destruct h4 as [h4 | h4]. left; auto. subst. right; intro h4. apply h3 in h4; auto. right. intro h5. pose proof (h2 _ _ _ h4 h5) as h6. apply h3 in h6; auto. Qed. (*Wyckoff*) Lemma sort_dec_Rwlem_dec : forall {T:Type} (R:relation T), sort_dec R -> Rwlem_dec R. intros T R h1. destruct h1 as [h1 h2 h3]. red; intros x y. destruct (h1 x y) as [h4 | h4]. destruct h4 as [h4 | h4]. left; left; auto. subst. left; right; auto. right. intro h5. pose proof (trans_Rw_R R h2 _ _ _ h5 h4) as h6. apply h3 in h6; auto. Qed. (*Wyckoff*) Definition exist_rel {T:Type} (R:relation T) (P:T->Prop) : {t:T | P t} -> {t:T | P t} -> Prop := (fun a b => R (proj1_sig a) (proj1_sig b)). (*Wyckoff*) Lemma well_founded_exist_rel : forall {T:Type} (R:relation T) (P:T->Prop), well_founded R -> well_founded (exist_rel R P). intros T R P h1. red. red in h1. intro x. destruct x as [x h2]. specialize (h1 x). induction h1 as [x h3 h4]. constructor. intros y h5. destruct y as [y h6]. red in h5. simpl in h5. apply h4; auto. Qed. (*Wyckoff*) (*This is used in well-founded-induction types, to sync up to iterators through the relation to prove statements or define values.*) Definition acc_step_exist_rel {T U:Type} (R:relation T) (P:T->Prop) (x:T) (pfp:P x) (pff:forall y:T, R y x -> U) (y:{y':T | P y'}) (pfr':(exist_rel R P) y (exist _ _ pfp)) : U := pff (proj1_sig y) pfr'. (*Wyckoff*) Definition minR {T:Type} (R:relation T) (pf:tso_dec R) : T -> T -> T := (fun (x y:T) =>if (pf x y) then x else y). (*Wyckoff*) Definition maxR {T:Type} (R:relation T) (pf:tso_dec R) : T -> T -> T := (fun (x y:T) =>if (pf x y) then y else x). (*Wyckoff*) Lemma minR_idem : forall {T:Type} (R:relation T) (pf:tso_dec R) (c:T), minR R pf c c = c. intros T R h1 c; unfold minR. destruct (h1 c c); auto. Qed. (*Wyckoff*) Lemma maxR_idem : forall {T:Type} (R:relation T) (pf:tso_dec R) (c:T), maxR R pf c c = c. intros T R h1 c; unfold maxR. destruct (h1 c c); auto. Qed. (*Wyckoff*) Lemma maxR_comm : forall {T:Type} (R:relation T) (pf:sort_dec R) (a b:T), let pftso := sort_dec_tso_dec _ pf in maxR R pftso a b = maxR R pftso b a. intros T R h1 a b htso. unfold maxR. destruct (htso a b) as [h2 | h2]. destruct h2 as [h2|]; subst. destruct (htso b a) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. reflexivity. reflexivity. destruct (htso b b) as [h2 | h2]; auto. destruct (htso b a) as [h3 | h3]; auto. pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. Qed. (*Wyckoff*) Lemma minR_comm : forall {T:Type} (R:relation T) (pf:sort_dec R) (a b:T), let pftso := sort_dec_tso_dec _ pf in minR R pftso a b = minR R pftso b a. intros T R h1 a b htso. unfold minR. destruct (htso a b) as [h2 | h2]. destruct h2 as [h2|]; subst. destruct (htso b a) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. reflexivity. reflexivity. destruct (htso b b) as [h2 | h2]; auto. destruct (htso b a) as [h3 | h3]; auto. pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. Qed. (*Wyckoff*) Lemma minR_compat_l : forall {T:Type} (R:relation T) (pf:tso_dec R) (x y:T), Rwd R (minR R pf x y) x. intros T R h1 x y. red. unfold minR. destruct (h1 x y); [right; auto | left; auto]. Qed. (*Wyckoff*) Lemma minR_compat_r : forall {T:Type} (R:relation T) (pf:tso_dec R) (x y:T), Rwd R (minR R pf x y) y. intros T R h1 x y. red. unfold minR. destruct (h1 x y); auto; try right; auto. Qed. (*Wyckoff*) Lemma maxR_compat_l : forall {T:Type} (R:relation T) (pf:tso_dec R) (x y:T), Rwd R x (maxR R pf x y). intros T R h1 x y. red. unfold maxR. destruct (h1 x y); auto; try right; auto. Qed. (*Wyckoff*) Lemma maxR_compat_r : forall {T:Type} (R:relation T) (pf:tso_dec R) (x y:T), Rwd R y (maxR R pf x y). intros T R h1 x y. red. unfold maxR. destruct (h1 x y); [right; auto | left; auto]. Qed. (*Wyckoff*) Lemma minR_functional : forall {T:Type} (R:relation T) (pf:tso_dec R) (x x' y y':T), x = x' -> y = y' -> minR R pf x y = minR R pf x' y'. intros; subst; auto. Qed. (*Wyckoff*) Lemma maxR_functional : forall {T:Type} (R:relation T) (pf:tso_dec R) (x x' y y':T), x = x' -> y = y' -> maxR R pf x y = maxR R pf x' y'. intros; subst; auto. Qed. (*Wyckoff*) Lemma lt_minR : forall {T:Type} (R:relation T) (pf:tso_dec R) (pft:Transitive R) (x y z:T), R x (minR R pf y z) -> R x y /\ R x z. intros T R h1 h2 x y z h3. unfold minR in h3. destruct (h1 y z) as [h4 | h4]. destruct h4 as [h5 | h5]. split; auto. eapply h2. apply h3. assumption. subst. split; auto. pose proof (h2 _ _ _ h3 h4) as h5. split; auto. Qed. (*Wyckoff*) Lemma minR_lt : forall {T:Type} (R:relation T) (pf:tso_dec R) (x y z:T), R (minR R pf x y) z -> R x z \/ R y z. intros T R h1 x y z h3. unfold minR in h3. destruct (h1 x y) as [h4 | h4]. left; auto. right; auto. Qed. (*Wyckoff*) Lemma lt_maxR : forall {T:Type} (R:relation T) (pf:tso_dec R) (x y z:T), R x (maxR R pf y z) -> R x y \/ R x z. intros T R h1 x y z h2. unfold maxR in h2. destruct (h1 y z) as [h3 | h3]. right; auto. left; auto. Qed. (*Wyckoff*) Lemma maxR_lt : forall {T:Type} (R:relation T) (pf:tso_dec R) (pfso:Transitive R) (x y z:T), R (maxR R pf x y) z -> R x z /\ R y z. intros T R h1 ht x y z h2. unfold maxR in h2. destruct (h1 x y) as [h3 | h3]. destruct h3 as [h3|]; subst. split; auto. eapply ht. apply h3. assumption. split; auto. split; auto. eapply ht. apply h3. assumption. Qed. (*Wyckoff*) Lemma lt_minR_eq : forall {T:Type} (R:relation T) (pfs:sort_dec R) (a b:T), R a b -> let pft := sort_dec_tso_dec _ pfs in minR R pft a b = a. intros T R h1 a b h2 h3. unfold minR. destruct (h3 a b) as [h4 | h4]; auto. pose proof (sort_dec_trans R h1 _ _ _ h2 h4) as h5. apply (sort_dec_irrefl _ h1) in h5; contradiction. Qed. (*Wyckoff*) Lemma gt_minR_eq : forall {T:Type} (R:relation T) (pfs:sort_dec R) (a b:T), R b a -> let pft := sort_dec_tso_dec _ pfs in minR R pft a b = b. intros T R h1 a b h2 h3. unfold minR. destruct (h3 a b) as [h4 | h4]; auto. destruct h4 as [h4|]; subst. pose proof (sort_dec_trans R h1 _ _ _ h4 h2) as h5. apply (sort_dec_irrefl _ h1) in h5; contradiction. apply (sort_dec_irrefl _ h1) in h2; contradiction. Qed. (*Wyckoff*) Lemma lt_maxR_eq : forall {T:Type} (R:relation T) (pfs:sort_dec R) (a b:T), R a b -> let pft := sort_dec_tso_dec _ pfs in maxR R pft a b = b. intros T R h1 a b h2 h3. unfold maxR. destruct (h3 a b) as [h4 | h4]; auto. pose proof (sort_dec_trans R h1 _ _ _ h2 h4) as h5. apply (sort_dec_irrefl _ h1) in h5; contradiction. Qed. (*Wyckoff*) Lemma gt_maxR_eq : forall {T:Type} (R:relation T) (pfs:sort_dec R) (a b:T), R b a -> let pft := sort_dec_tso_dec _ pfs in maxR R pft a b = a. intros T R h1 a b h2 h3. unfold maxR. destruct (h3 a b) as [h4 | h4]; auto. destruct h4 as [h4|]; subst. pose proof (sort_dec_trans R h1 _ _ _ h4 h2) as h5. apply (sort_dec_irrefl _ h1) in h5; contradiction. apply (sort_dec_irrefl _ h1) in h2; contradiction. Qed. (*Wyckoff*) Lemma le_minR_maxR : forall {T:Type} (R:relation T) (pfs:sort_dec R) (a b:T), let pft := sort_dec_tso_dec _ pfs in Rw R (minR R pft a b) (maxR R pft a b). intros T R h1 a b h2. unfold minR, maxR. destruct (h2 a b) as [h3 | h3]. destruct h3 as [h3 | h3]. left; auto. right; auto. left; auto. Qed. (*Wyckoff*) Lemma minR_eq_iff : forall {T:Type} (R:relation T) (pf:sort_dec R) (x y z:T), let pftso := sort_dec_tso_dec _ pf in minR R pftso x y = z <-> (x = z /\ Rw R x y) \/ (y = z /\ Rw R y x). intros T R h1 x y z ht; split; intro h2; subst; unfold minR. destruct (ht x y) as [h2 | h2]. left; auto. split; auto. destruct h2; subst; auto. left; auto. right; auto. right. split; auto. left; auto. destruct h2 as [h2 | h2]. destruct (ht x y) as [h3 | h3]. destruct h2; auto. destruct h2 as [? h2]; subst. pose proof (trans_R_Rw _ (sort_dec_trans _ h1) _ _ _ h3 h2) as h4. apply (sort_dec_irrefl R h1) in h4. contradiction. destruct h2 as [? h3]; subst. destruct (ht x z) as [h4 | h4]. destruct h4 as [h4 | h4]. pose proof (trans_R_Rw _ (sort_dec_trans _ h1) _ _ _ h4 h3) as h5. apply (sort_dec_irrefl _ h1) in h5. contradiction. assumption. reflexivity. Qed. (*Wyckoff*) Lemma maxR_eq_iff : forall {T:Type} (R:relation T) (pf:sort_dec R) (x y z:T), let pftso := sort_dec_tso_dec _ pf in maxR R pftso x y = z <-> (x = z /\ Rw R y x) \/ (y = z /\ Rw R x y). intros T R h1 x y z ht; split; intro h2; subst; unfold maxR. destruct (ht x y) as [h2 | h2]. right; auto. split; auto. destruct h2; subst; auto. left; auto. right; auto. left. split; auto. left; auto. destruct h2 as [h2 | h2]. destruct (ht x y) as [h3 | h3]. destruct h2 as [? h2]; subst. destruct h3 as [h3 | h3]. pose proof (trans_R_Rw _ (sort_dec_trans _ h1) _ _ _ h3 h2) as h4. apply (sort_dec_irrefl R h1) in h4. contradiction. auto. destruct h2 as [? h4]; subst. reflexivity. destruct h2 as [? h2]; subst. destruct (ht x z) as [h4 | h4]. reflexivity. pose proof (trans_R_Rw _ (sort_dec_trans _ h1) _ _ _ h4 h2) as h5. apply (sort_dec_irrefl _ h1) in h5. contradiction. Qed. (*Wyckoff*) Lemma tso_lt : total_strict_order lt. red. intros x y. destruct (lt_eq_lt_dec x y) as [h1 | h1]. destruct h1 as [h1 | h1]. left; auto. right; left; auto. right; right; auto. Qed. (*Wyckoff*) Lemma lt_sortable : sortable lt. apply (sort_intro _ tso_lt lt_trans lt_irrefl). Defined. (*Wyckoff*) Lemma strict_order_not_symm : forall {T:Type} (R:T->T->Prop), Strict_Order R -> forall x y:T, R x y -> ~ R y x. intros T R h1 x y h2. intro h3. destruct h1 as [h1a h1b]. red in h1a. specialize (h1a _ _ _ h2 h3). red in h1b. specialize (h1b x). contradiction. Qed. (*Wyckoff*) Lemma transitive_exist_rel : forall {T:Type} (R:relation T), Transitive R -> forall (P:T->Prop), Transitive (exist_rel R P). intros T R h1 P. red in h1. red. intros x y z h0 h0'. destruct x as [x h2], y as [y h3], z as [z h4]. red in h0, h0'. simpl in h0, h0'. specialize (h1 _ _ _ h0 h0'). red. simpl. assumption. Qed. (*Wyckoff*) Lemma irrefl_exist_rel : forall {T:Type} (R:relation T) (P:T->Prop), irrefl R -> irrefl (exist_rel R P). intros T R P h1. red in h1. red. intro x. destruct x as [x h2]. specialize (h1 x). intro h3. red in h3. simpl in h3. contradiction. Qed. (*Wyckoff*) Lemma strict_order_exist_rel : forall {T:Type} (R:relation T), Strict_Order R -> forall (P:T->Prop), Strict_Order (exist_rel R P). intros T R h1 P. destruct h1. constructor; try apply irrefl_exist_rel; try apply transitive_exist_rel; auto. Qed. (*Wyckoff*) Lemma so_lt_nat : Strict_Order lt. constructor. red; intro; apply lt_trans. red; intros; apply lt_irrefl. Qed. (*Maybe replace [tso_dec_lt_nat] with [lt_nat_tso_dec]*) (*Wyckoff*) Lemma tso_dec_lt_nat : tso_dec lt. red; apply lt_eq_lt_dec. Qed. (*Wyckoff*) Lemma sort_dec_lt_nat : sort_dec lt. constructor. apply tso_dec_lt_nat. red; apply lt_trans. red; apply lt_irrefl. Qed. (*Wyckoff*) Lemma tso_dec_lt_nat_irrelevance : forall (pf pf':forall x y:nat, {x < y} + {x = y} + {y < x}), pf = pf'. intros h1 h2. apply functional_extensionality_dep. intro m. apply functional_extensionality_dep. intro n. destruct (h1 m n) as [h3 | h3], (h2 m n) as [h4 | h4]; f_equal. destruct h3 as [h3 | h3], h4 as [h4 | h4]; f_equal; try apply proof_irrelevance; subst; try omega. destruct h3; omega. destruct h4; omega. apply proof_irrelevance. Qed. (*Wyckoff*) Lemma Rlem_dec_lt_nat : Rlem_dec lt. apply sort_dec_Rlem_dec; apply sort_dec_lt_nat. Qed. (*Wyckoff*) Definition Rcompose {T U:Type} (R:relation U) (f:T->U) : relation T := fun x y => R (f x) (f y). (*Wyckoff*) Lemma Rcompose_tso_dec : forall {T U:Type} (R:relation U) (f:T->U), injective f -> tso_dec R -> tso_dec (Rcompose R f). unfold Rcompose. intros T U R f h0 h1. red in h1; red. intros x y. specialize (h1 (f x) (f y)); auto. destruct h1 as [[h1 | h1] | h1]. left; auto. apply h0 in h1; subst. left; right; auto. right; auto. Qed. (*Wyckoff*) Lemma Rcompose_trans : forall {T U:Type} (R:relation U) (f:T->U), Transitive R -> Transitive (Rcompose R f). unfold Rcompose. intros T U R f h1. red in h1; red. intros x y z h2 h3. specialize (h1 (f x) (f y) (f z) h2 h3); auto. Qed. (*Wyckoff*) Lemma Rcompose_irrefl : forall {T U:Type} (R:relation U) (f:T->U), irrefl R -> irrefl (Rcompose R f). unfold Rcompose. intros T U R f h1. red in h1; red. intro x. apply (h1 (f x)); auto. Qed. (*Wyckoff*) Definition Rcompose_sort_dec : forall {T U:Type} (R:relation U) (f:T->U), injective f -> sort_dec R -> sort_dec (Rcompose R f). intros T U R f h0 h1. destruct h1 as [h1 h2 h3]. constructor. apply Rcompose_tso_dec; auto. apply Rcompose_trans; auto. apply Rcompose_irrefl; auto. Qed. Inductive disagrees {T U} : relation (T->U) := | disagrees_intro : forall c f g, f c <> g c -> disagrees f g. (*Do a min [seg_t] function.*) Inductive disagrees_at {T U} {f g:T->U} : disagrees f g -> T -> Prop := | disagrees_at_intro : forall c (pfn:f c <> g c), disagrees_at (disagrees_intro c f g pfn) c. (*See FunctionProperties2 for a minimal version of [disagrees_at], and decidable lemmas.*) Definition pred_eq {T} : relation (T->Prop) := fun P Q => forall x:T, P x <-> Q x. Lemma refl_pred_eq : forall T, Reflexive (@pred_eq T). unfold Reflexive, pred_eq; intros; tauto. Qed. Lemma sym_pred_eq : forall T, Symmetric (@pred_eq T). unfold Symmetric, pred_eq; intros ? ? ? h1 ?; rewrite h1; tauto. Qed. Lemma trans_pred_eq : forall T, Transitive (@pred_eq T). unfold Transitive, pred_eq; intros ? ? ? ? h1 h2 ?; rewrite h1; auto. Qed. Add Parametric Relation {T} : (T->Prop) (@pred_eq T) reflexivity proved by (refl_pred_eq T) symmetry proved by (sym_pred_eq T) transitivity proved by (trans_pred_eq T) as pred_eq_eq. Definition proj1_eq' {T} : relation {P:T->Prop & {x:T|P x}} := fun x y => proj1_sig (projT2 x) = proj1_sig (projT2 y). Lemma refl_proj1_eq' : forall T, Reflexive (@proj1_eq' T). unfold Reflexive, proj1_eq'; intros; auto. Qed. Lemma sym_proj1_eq' : forall T, Symmetric (@proj1_eq' T). unfold Symmetric, proj1_eq'; intros; auto. Qed. (*Wyckoff*) Lemma trans_proj1_eq' : forall T, Transitive (@proj1_eq' T). unfold Transitive, proj1_eq'; intros T x y z; auto. destruct x, y, z; simpl; auto. destruct s, s0, s1; simpl. intros; subst; auto. Qed. (*Wyckoff*) Add Parametric Relation {T} : {P:T->Prop & {x:T|P x}} (@proj1_eq' T) reflexivity proved by (@refl_proj1_eq' T) symmetry proved by (@sym_proj1_eq' T) transitivity proved by (@trans_proj1_eq' T) as proj1_eq_eq'. (*Wyckoff*) Definition proj1_eq {T} (P:T->Prop) : relation {x:T|P x} := fun x y => proj1_sig x = proj1_sig y. (*Wyckoff*) Lemma refl_proj1_eq : forall {T P}, Reflexive (@proj1_eq T P). unfold Reflexive, proj1_eq; intros; auto. Qed. (*Wyckoff*) Lemma sym_proj1_eq : forall {T P}, Symmetric (@proj1_eq T P). unfold Symmetric, proj1_eq; intros; auto. Qed. (*Wyckoff*) Lemma trans_proj1_eq : forall {T P}, Transitive (@proj1_eq T P). unfold Transitive, proj1_eq; intros; congruence. Qed. (*Wyckoff*) Definition incrR_fun {T U:Type} (R:relation T) (R':relation U) (f:T->U) : Prop := forall a b:T, R a b -> R' (f a) (f b). (*Wyckoff*) Lemma incrR_fun_inj : forall {T U:Type} (R:relation T) (R':relation U) (f:T->U), tso_dec R -> irrefl R'-> incrR_fun R R' f -> injective f. intros T U R R' f htso hir h1. red; intros x y h2. red in h1. destruct (htso x y) as [[h3|h3]|h3]; try apply h1 in h3; auto. rewrite h2 in h3. apply hir in h3; contradiction. rewrite h2 in h3. apply hir in h3; contradiction. Qed. (*Wyckoff*) Definition incr_fun (f:nat->nat) : Prop := forall n:nat, f n < f (S n). (*Wyckoff*) Definition incr_fun' (f:nat->nat) : Prop := incrR_fun lt lt f. (*Wyckoff*) Lemma incr_fun_iff : forall (f:nat->nat), incr_fun f <-> incr_fun' f. intro f. split; intro h1; red in h1; red. intros a b h2. revert h2. revert a h1. revert f. induction b as [|b ih]; simpl. intros; omega. intros f a h1 h2. specialize (ih f a h1). red in h2. apply le_S_n in h2. apply le_lt_eq_dec in h2. destruct h2 as [h2 | h2]. specialize (ih h2). specialize (h1 b). omega. subst. apply h1. intro; apply h1; auto with arith. Qed. (*Wyckoff*) Lemma incr_fun_S : incr_fun S. red; intros; auto with arith. Qed. (*Wyckoff*) Lemma incr_fun_inj : forall (f:nat -> nat), incr_fun f -> injective f. intros ? h1. rewrite incr_fun_iff in h1. unfold incr_fun' in h1. apply incrR_fun_inj in h1; auto. apply tso_dec_lt_nat. red; apply lt_irrefl. Qed. (*Wyckoff*) Definition weak_incrR_fun {T U:Type} (R:relation T) (R':relation U) (f:T->U) : Prop := forall a b:T, R a b -> Rw R' (f a) (f b). (*Wyckoff*) Lemma incrR_fun_impl_weak : forall {T U:Type} (R:relation T) (R':relation U) (f:T->U), incrR_fun R R' f -> weak_incrR_fun R R' f. intros T U R R' f h1. red in h1. red. intros; left; auto. Qed. (*Wyckoff*) Definition weak_incr_fun (f:nat->nat) : Prop := forall n:nat, f n <= f (S n). (*Wyckoff*) Definition weak_incr_fun' (f:nat->nat) : Prop := weak_incrR_fun lt lt f. (*Wyckoff*) Definition weak_incr_fun_iff : forall (f:nat->nat), weak_incr_fun f <-> weak_incr_fun' f. intros f; split; intro h1. red in h1; red. intros a b h2. revert h2. revert a h1. revert f. induction b as [|b ih]; simpl. intros; omega. intros f a h1 h2. specialize (ih f a h1). red in h2. apply le_S_n in h2. apply le_lt_eq_dec in h2. destruct h2 as [h2 | h2]. specialize (ih h2). specialize (h1 b). rewrite Rw_le_iff in ih. rewrite Rw_le_iff. omega. subst. rewrite Rw_le_iff. apply h1. red; red in h1; intro m; specialize (h1 m (S m)); rewrite Rw_le_iff in h1; auto. Qed. (*Wyckoff*) Lemma incr_fun_impl_weak : forall (f:nat->nat), incr_fun f -> weak_incr_fun f. intros f h1. red in h1. red. intro; auto with arith. Qed. (*Wyckoff*) Lemma weak_incr_fun_pred : weak_incr_fun pred. red; intros; simpl; auto with arith. Qed. (*Wyckoff*) Lemma weak_incr_fun_S : weak_incr_fun S. apply incr_fun_impl_weak. apply incr_fun_S. Qed. End Relations. (*Wyckoff*) Definition P_to_dep {T} (P Q:T->Prop) : {x:T | Q x} -> Prop := fun p => let (x, _) := p in P x. (*Wyckoff*) (*This is a lot of how-do-you-do about merely substituting one value of a [seg_fun] [f] at a given integer [k] with the value [y:U], and all of [f]s other points remaining the same. The simpler way to do it with [eq_nat_dec] is inappropriate for arbitrary computation (i.e. it involves equality in [Prop], rather than [Set].*) (*Wyckoff*) Fixpoint seg_fun_subst {U:Type} {n:nat} : seg_fun n U -> U -> forall k, k < n -> seg_fun n U := match n with | O => fun _ _ _ pf _ _ => False_rect _ (lt_n_0 _ pf) | S n' => fun f y k => let f' := seg_fun_to_S f in match k with | 0 => fun pfk i => match i with | 0 => fun _ => y | S i' => fun pfi => f i pfi end | S k' => fun pfk => let pfk' := lt_S_n _ _ pfk in fun i => match i with | 0 => fun pfi => f 0 pfi | S i' => fun pfi => let pfi' := lt_S_n _ _ pfi in seg_fun_subst f' y k' pfk' i' pfi' end end end. (*Wyckoff*) Lemma seg_fun_subst_compat : forall {U n} (f:seg_fun n U) y k (pfk:k < n) i (pfi:i < n), let f' := seg_fun_subst f y k pfk in (i = k -> f' i pfi = y) /\ (i <> k -> f' i pfi = f i pfi). intros U n. induction n as [|n ih]. intros; omega. intros f y k. destruct k as [|k]. intros hk i. destruct i as [|i]; simpl. intros; tauto. intros. split. intro; discriminate. intro; auto. intros hk i. destruct i as [|i]. intros; split; auto. intro; discriminate. intros hi f'. pose proof (lt_S_n _ _ hk) as hk'. pose proof (lt_S_n _ _ hi) as hi'. specialize (ih (seg_fun_to_S f) y k hk' i hi'). simpl in ih. destruct ih as [h2 h3]. split. intro h4. inversion h4; subst. specialize (h2 (eq_refl _)). rewrite <- h2. unfold f'. simpl. f_equal. apply proof_irrelevance. apply proof_irrelevance. intro h4. unfold f'. simpl. pose proof (proof_irrelevance _ hi' (lt_S_n _ _ hi)); subst. pose proof (proof_irrelevance _ hk' (lt_S_n _ _ hk)); subst. rewrite h3; auto with arith. unfold seg_fun_to_S. f_equal. apply proof_irrelevance. Qed. (*Wyckoff*) Fixpoint seg_fun_prod_type {n} : seg_fun n Type -> Type := match n with | 0 => fun _ => unit | S n' => fun f => let f' := seg_fun_to_S f in ((f 0 (lt_0_Sn _)) * (seg_fun_prod_type f'))%type end. (*Wyckoff*) Definition seg_fun_prod_type_to_S {n} {f:seg_fun (S n) Type} (g:seg_fun_prod_type f) : seg_fun_prod_type (seg_fun_to_S f) := snd g. (*Wyckoff*) Lemma seg_fun_prod_type_S_eq : forall {n} {f:seg_fun (S n) Type} (g:seg_fun_prod_type f), g = (fst g, seg_fun_prod_type_to_S g). intros; apply surjective_pairing. Qed. (*Wyckoff*) Fixpoint seg_fun_sum_type {n} : seg_fun n Type -> Type := match n with | 0 => fun _ => unit | S n' => fun f => let f' := seg_fun_to_S f in ((f 0 (lt_0_Sn _)) + (seg_fun_sum_type f'))%type end. (*Wyckoff*) Fixpoint seg_fun_prod {n} : forall {f:seg_fun n Type} (g:seg_fun_dep f), seg_fun_prod_type f := match n with | 0 => fun (f:seg_fun 0 Type) _ => tt | S n' => fun f g => let f' := seg_fun_to_S f in let g' := seg_fun_dep_to_S g in (g 0 (lt_0_Sn _), seg_fun_prod g') end. (*Wyckoff*) Definition seg_fun_prod_from_S {n} {f:seg_fun (S n) Type} (g:seg_fun_dep f) : seg_fun_prod_type (seg_fun_from_S f) := seg_fun_prod (seg_fun_dep_from_S g). (*Wyckoff*) Definition seg_fun_prod_to_S {n} {f:seg_fun (S n) Type} (g:seg_fun_dep f) : seg_fun_prod_type (seg_fun_to_S f) := seg_fun_prod (seg_fun_dep_to_S g). (*Wyckoff*) Lemma seg_fun_prod_S_eq : forall {n} {f:seg_fun (S n) Type} (g:seg_fun_dep f), seg_fun_prod g = (g 0 (lt_0_Sn _), seg_fun_prod_to_S g). auto. Qed. (*Wyckoff*) Definition seg_fun_dep_pf {U n} (f:seg_fun n U) (b:forall i (pfi:i < n), i < n) : forall i (pfi:i < n), U := fun i pfi => f i (b i pfi). (*Wyckoff*) Lemma seg_fun_dep_pf_eq : forall {U n} (f:seg_fun n U) (b:forall i (pfi:i < n), i < n), seg_fun_dep_pf f b = f. unfold seg_fun_dep_pf; intros; apply functional_extensionality_dep; intro; apply functional_extensionality_dep; intro; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma seg_fun_dep_pf_eq2 : forall {U n} (f:seg_fun n U) (b c:forall i (pfi:i < n), i < n), seg_fun_dep_pf f b = seg_fun_dep_pf f c. unfold seg_fun_dep_pf; intros; apply functional_extensionality_dep; intro; apply functional_extensionality_dep; intro; apply f_equal_dep; auto. Qed. (*Wyckoff*) Definition dep_transfer_seg_fun {U n} {f g:seg_fun n U} (pfe:f = g) (V:forall m, seg_fun m U -> Type) (Vs:V _ f) : V _ g. intros; subst; auto. Defined. (*Wyckoff*) Definition add_seg_to_identify n : seg_fun (S n) nat := match n with | 0 => fun _ pf => 0 | S n' => fun i => match i with | 0 => fun _ => 0 | S i' => fun _ => i' end end. (*Wyckoff*) Lemma lt_add_seg_to_identify : forall n i pf, 0 < n -> add_seg_to_identify n i pf < n. intro n. induction n as [|n ih]; auto; simpl. intros i ?. destruct i; simpl; auto with arith. Qed. (*Wyckoff*) Definition tl_seg_for_fictitious n : seg_fun (pred n) nat := match n with | 0 => fun _ pf => False_rect _ (lt_n_0 _ pf) | S n' => match n' with | 0 => fun _ pf => False_rect _ (lt_n_0 _ pf) | S _ => fun i _ => S i end end. (*Wyckoff*) Lemma lt_tl_seg_for_fictitious : forall n i pf, tl_seg_for_fictitious n i pf < n. intro n; destruct n; simpl; auto with arith. intros ? h1; contradict (lt_n_0 _ h1). destruct n as [|n]; auto with arith. intros ? h1; contradict (lt_n_0 _ h1). Qed. (*Wyckoff*) Lemma inj_option_map : forall {T U} (f:T->U), injective f -> injective (option_map f). intros T U f h1; red in h1; red; intros x y h2; destruct x as [x|], y as [y|]; simpl in h2; try discriminate; auto. inversion h2. apply h1 in H0; subst; auto. Qed. (*Wyckoff*) Lemma option_map_inj : forall {T U} (f:T->U), injective (option_map f) -> injective f. intros T U f h1; red in h1; red; intros x y h2. specialize (h1 (Some x) (Some y)). simpl in h1. rewrite h2 in h1. specialize (h1 eq_refl). inversion h1; auto. Qed. (*Wyckoff*) Lemma surj_option_map : forall {T U} (f:T->U), surjective f -> surjective (option_map f). intros T u f h1; red in h1; red. intro y. destruct y as [y|]. specialize (h1 y). destruct h1 as [x]; subst. exists (Some x); simpl; auto. exists None; simpl; auto. Qed. (*Wyckoff*) Lemma option_map_surj : forall {T U} (f:T->U), surjective (option_map f) -> surjective f. intros T U f h1; red in h1; red. intro y; specialize (h1 (Some y)). destruct h1 as [x h2]. destruct x as [x|]; simpl in h2. inversion h2; subst. exists x; auto. discriminate. Qed. (*Wyckoff*) Lemma bij_option_map : forall {T U} (f:T->U), bijective f -> bijective (option_map f). intros ? ? ? h1; red in h1; red; split; destruct h1; [apply inj_option_map | apply surj_option_map]; auto. Qed. (*Wyckoff*) Lemma option_map_bij : forall {T U} (f:T->U), bijective (option_map f) -> bijective f. intros ? ? ? h1; red in h1; red; split; destruct h1 as [h1 h2]; [apply option_map_inj in h1 | apply option_map_surj in h2]; auto. Qed. (*Wyckoff*) Lemma invertible_option_map : forall {T U} (f:T->U), invertible f -> invertible (option_map f). intros ? ? ? h1; apply bijective_impl_invertible; apply invertible_impl_bijective in h1; auto; apply bij_option_map; auto. Qed. (*Wyckoff*) Lemma option_map_invertible : forall {T U} (f:T->U), invertible (option_map f) -> invertible f. intros ? ? ? h1; apply invertible_impl_bijective in h1; apply bijective_impl_invertible; apply option_map_bij in h1; auto. Qed. (*In pre-image of function.*) (*Wyckoff*) Inductive in_pre_aux {X S} (v:X->S) : S -> Prop := in_pre_intro : forall x, in_pre_aux v (v x). (*Wyckoff*) Definition in_pre {X S} (s:S) (v:X->S) : Prop := in_pre_aux v s. (*Wyckoff*) Inductive inh_ind_var {X S} (v:X->S) : Prop := inh_ind_var_intro : forall s, in_pre s v -> inh_ind_var v. bool2-v0-3/ImageImplicit.v0000644000175000017500000000170113575574055016245 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-2019, Daniel Wyckoff, except for the portions so labeleled which I got from Daniel Schepler*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) (*Schepler*) (*silly citation!*) Require Export Image. Arguments Im [_] [_] _ _ _. bool2-v0-3/SetUtilities.v0000644000175000017500000153424213575574055016173 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-2016, Daniel Wyckoff, except for the portions so labeleled as (*Schepler*) which I got from Daniel Schepler. My functions will be labeled (*Wyckoff*)*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Version 8.4 pl4 Coq.*) Require Export EnsemblesImplicit. Require Export ImageImplicit. Require Import Classical. Require Import Image. Require Import LogicUtilities. Require Import Setoid. Require Import FunctionProperties. Require Import Arith. Require Import Even. Require Import Equality. Require Import Infinite_sets. Require Import ArithUtilities. Require Import Description. Require Import DecidableDec. Require Import Bool. Require Import FunctionalExtensionality. Require Import EqDec. Require Import Basics. (*This notation is to prevent the need to keep on writing [Ensembles.In] or [Ensembles.Complement] in other files that also import [List].*) Notation Inn := Ensembles.In. Notation Comp := Ensembles.Complement. (*Wyckoff*) Lemma prop_ext : forall (P Q : Prop), (P <-> Q) -> P = Q. intros P Q h1. pose proof (@Extensionality_Ensembles Prop (fun _ => P) (fun _ => Q)) as h2. hfirst h2. red; split; red; intros x h3; red in h3; red; tauto. specialize (h2 hf). pose proof (equal_f h2 True) as h3. simpl in h3; auto. Qed. (*Wyckoff*) Definition pred_ext : forall {T} (P Q:T->Prop), pred_eq P Q -> P = Q. intros; apply functional_extensionality; intro; apply prop_ext; apply H. Qed. (*This relatively "strong" assumption is used in a few places in the Boolean code explicitly since there's a vague plan to work my way around those arguments, but then some are indispensable, Overall, since this development uses LEM, this axiom doesn't really hurt, but merely indicates an inclination to want to eliminate some instances of it!*) (*Wyckoff*) Axiom ens_eq_dec : forall (T:Type), type_eq_dec (Ensemble T). (*Wyckoff*) Lemma incl_add : forall {T:Type} (A:Ensemble T) (a:T), Included A (Add A a). auto with sets. Qed. (*Wyckoff*) Lemma incl_full : forall {T} (A:Ensemble T), Included A (Full_set _). intros; red; intros; constructor. Qed. (*Wyckoff*) Lemma eq_incl : forall {T:Type} (A B:Ensemble T), A = B -> Included A B. intros; subst; auto with sets. Qed. (*Wyckoff*) Lemma in_eq_set : forall {T:Type} (A B:Ensemble T) (pf:A = B) (x:T), Inn A x -> Inn B x. intros; subst; auto. Qed. (*Wyckoff*) Lemma in_eq_set_iff : forall {T:Type} (A B:Ensemble T) (pf:A = B) (x:T), In A x <-> In B x. intros; subst; tauto. Qed. (*Wyckoff*) Lemma in_eq_elt_l : forall {T:Type} (A:Ensemble T) (x y:T) (pf:x = y), In A x -> In A y. intros; subst; auto. Qed. (*Wyckoff*) Lemma in_eq_elt_r : forall {T:Type} (A:Ensemble T) (x y:T) (pf:y = x), In A x -> In A y. intros; subst; auto. Qed. (*Wyckoff*) Lemma in_add_eq : forall {T:Type} (A:Ensemble T) (x:T), In A x -> Add A x = A. intros T A x h1. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h2. destruct h2 as [y h2l | y h2r]. assumption. destruct h2r; subst. assumption. (* >= *) red. intros y h2. left. assumption. Qed. (*Wyckoff*) Lemma eq_sing_iff : forall {T} A x, A = Singleton x <-> (Inn A x /\ forall a:T, Inn A a -> a = x). intros; split; subst; intros; subst. split. constructor. intros a h1. destruct h1; auto. destruct H as [h1 h2]. apply Extensionality_Ensembles; red; split; red; intros c h3. apply h2 in h3; subst; constructor. destruct h3; auto. Qed. (*Wyckoff*) Lemma nin_iff : forall {T:Type} (A:Ensemble T) (x:T), ~ In A x <-> Intersection A (Singleton x) = Empty_set _. intros T A x; split; intro h1. apply Extensionality_Ensembles; red; split; red; intros c h2. destruct h2 as [c h2 h3]. destruct h3; contradiction. contradiction. intro h2. pose proof (Intersection_intro _ _ _ _ h2 (In_singleton _ _ )) as h3. rewrite h1 in h3; contradiction. Qed. (*Wyckoff*) Lemma nin_neq_nin_add : forall {T:Type} (A:Ensemble T) (a:T), ~Inn A a -> forall (x:T), x <> a -> ~Inn A x -> ~Inn (Add A a) x. intros T A a h1 x h2 h3 h4. destruct h4 as [x h4 | x h5]. contradiction. destruct h5. contradict h2; auto. Qed. (*Wyckoff*) Lemma nin_add_or : forall {T:Type} (A:Ensemble T) (a x:T), ~In (Add A a) x -> a <> x /\ ~In A x. intros T A a x h1. split. intro; subst. contradict h1. right; auto. constructor. intro h2. contradict h1. left; auto. Qed. (*Wyckoff*) Lemma incl_add_nin_incl : forall {T:Type} (A B:Ensemble T) (b:T), Included A (Add B b) -> ~In A b -> Included A B. intros T A B b h1 h3. red. intros y h4. assert (h5:y <> b). intro h5. subst. contradiction. apply h1 in h4. destruct h4 as [y h4 | y h6]. assumption. destruct h6. contradict h5. reflexivity. Qed. (*Wyckoff*) (*Uses LEM.*) Lemma not_incl : forall {T:Type} (A B:Ensemble T), ~Included A B -> exists x:T, Inn A x /\ ~Inn B x. intros T A B h1. apply not_all_not_ex. intro h2. contradict h1. red; intros x h3. specialize (h2 x). tauto. Qed. (*Wyckoff*) Lemma add_preserves_inclusion : forall {T:Type} (A B:Ensemble T) (x:T), Included A B -> Included (Add A x) (Add B x). intros T A B x h1. red. intros x' h2. destruct h2 as [x' h2 | x' h3]. left. auto with sets. destruct h3. right. constructor. Qed. (*Wyckoff*) Lemma in_add_nin_dec : forall {T:Type} (A:Ensemble T) (a:T), ~Inn A a -> forall (x:T), Inn (Add A a) x -> {Inn A x} + {x = a}. intros T A a h1 x h2. destruct (classic_dec (x = a)) as [h3 | h4]. right; auto. left. apply NNPP. intro h5. pose proof (nin_neq_nin_add A a h1 x h4 h5). contradiction. Qed. (*Wyckoff*) Lemma add_not_empty : forall {T:Type} (A:Ensemble T) (x:T), Add A x <> Empty_set _. intros T A x. intro h1. pose proof (Add_intro2 _ A x) as h2. rewrite h1 in h2. contradiction. Qed. (*Wyckoff*) Lemma sub_add_same_nin : forall {T:Type} (A:Ensemble T) (x:T), ~ In A x -> (Subtract (Add A x) x) = A. intros T A x h1. apply Extensionality_Ensembles. red. split. (* => *) red. intros y h2. destruct h2 as [h2 ?]. destruct h2; try contradiction; auto. (* <= *) red. intros y h2. constructor. left; auto. intro h3. destruct h3; subst. contradiction. Qed. (*Wyckoff*) Lemma sub_add_same_in : forall {T:Type} (A:Ensemble T) (x:T), In A x -> (Subtract (Add A x) x) = Subtract A x. intros T A x h1. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h2. destruct h2 as [h3 h4]. destruct h3 as [y h3 | y h5]. constructor; auto. contradiction. (* >= *) red. intros y h2. destruct h2. constructor; try constructor; auto. Qed. (*Wyckoff*) Lemma add_sub_compat_in : forall {T:Type} (A:Ensemble T) (x:T), In A x -> (Add (Subtract A x) x) = A. intros T A x h1. apply Extensionality_Ensembles. red. split. red. intros x' h2. destruct h2 as [x' h2 | x' h3]. destruct h2; auto. inversion h3. subst. assumption. red. intros x' h2. destruct (classic (x = x')). subst. right. constructor. left. constructor. assumption. intro h3. inversion h3. subst. contradict H. reflexivity. Qed. (*Wyckoff*) Lemma add_nin_sub_compat : forall {T:Type} (A B:Ensemble T) (x:T), ~Inn A x -> Add A x = Add B x -> A = Subtract B x. intros T A B x h1 h2. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h3. pose proof (Add_intro1 _ A x y h3) as h4. rewrite h2 in h4. pose proof (Add_inv _ B x y h4) as h5. destruct h5 as [h5l | h5r]. constructor; auto. intro h6. destruct h6; subst; contradiction. subst. contradiction. (* >= *) red. intros y h3. destruct h3 as [h3 h4]. assert (h5: x <> y). intro. subst. pose proof (In_singleton _ y). contradiction. clear h4. pose proof (Add_intro1 _ B x y h3) as h4. rewrite <- h2 in h4. destruct h4 as [y h4l | y h4r]. auto. destruct h4r; subst. contradict h5. auto. Qed. (*Wyckoff*) Lemma add2 : forall {T:Type} (A:Ensemble T) (x:T), Add (Add A x) x = Add A x. intros T A x. unfold Add. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h1. destruct h1 as [y h2 | y h3]. auto. destruct h3. apply Union_intror. constructor. (* >= *) red. intros y h1. constructor. auto. Qed. (*Wyckoff*) Lemma add_comml : forall {T:Type} (A:Ensemble T) (x y:T), Included (Add (Add A x) y) (Add (Add A y) x). intros T A x y. red. intros z h1. destruct h1 as [z h2|z h3]. destruct h2 as [z h4 | z h5]. left. left. auto. destruct h5; subst. right. constructor. destruct h3; subst. left. right. constructor. Qed. (*Wyckoff*) Lemma add_comm : forall {T:Type} (A:Ensemble T) (x y:T), Add (Add A x) y = Add (Add A y) x. intros; apply Extensionality_Ensembles; red; split; apply add_comml. Qed. (*Wyckoff*) Lemma union_add_comm : forall {T:Type} (A B:Ensemble T) (x:T), Union (Add A x) B = Add (Union A B) x. intros T A B x. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h1. destruct h1 as [y h1 | y h2]. destruct h1 as [y h1 | y h3]. left. left. assumption. destruct h3. subst. right. constructor. left. right. assumption. (* >= *) red. intros y h1. destruct h1 as [y h1 | y h2]. destruct h1 as [y h1 | y h3]. left. left. assumption. right. assumption. destruct h2; subst. left. right. constructor. Qed. (*Wyckoff*) Lemma add_sing_same : forall {T:Type} (a:T), Add (Singleton a) a = Singleton a. intros T a; apply Extensionality_Ensembles; red; split; intros x h1. destruct h1 as [x h1 | x h1]; auto. destruct h1. right; auto. constructor. Qed. (*Wyckoff*) Lemma couple_add_sing : forall {T:Type} (x y:T), Couple x y = Add (Singleton x) y. intros T x y. apply Extensionality_Ensembles. red. split. red. intros z h1. destruct h1. left. constructor. right. constructor. red. intros z h1. destruct h1 as [z h1l | z h1r]. destruct h1l; subst. left. destruct h1r. right. Qed. (*Wyckoff*) Lemma add_empty_sing : forall {T:Type} (x:T), Add (Empty_set _) x = Singleton x. intros T x. apply Extensionality_Ensembles. red. split. red. intros y h1. destruct h1; try contradiction; auto. red. intros y h2. right; auto. Qed. (*Wyckoff*) Lemma couple_add_add : forall {T:Type} (x y:T), Couple x y = Add (Add (Empty_set _) x) y. intros T x y. rewrite couple_add_sing. rewrite add_empty_sing. reflexivity. Qed. (*Wyckoff*) Lemma add_nin_inj2 : forall {T:Type} (A:Ensemble T) (a a':T), ~In A a -> Add A a = Add A a' -> a = a'. intros T A a a' h0 h1. pose proof (Add_intro2 _ A a) as h2. rewrite h1 in h2. destruct h2 as [a h2 | a h2]. contradiction. destruct h2; auto. Qed. (*Wyckoff*) Lemma add_nin_inj_eq : forall {T:Type} (A A':Ensemble T) (a:T), ~In A a -> ~In A' a -> Add A a = Add A' a -> A = A'. intros T A A' a h0 h0' h1. pose proof (f_equal (fun S => Subtract S a) h1) as h2. simpl in h2. do 2 rewrite sub_add_same_nin in h2; auto. Qed. (*Wyckoff*) Lemma incl_add_nin_inj_eq : forall {T:Type} (A A':Ensemble T) (a:T), ~In A a -> ~In A' a -> Included (Add A a) (Add A' a) -> Included A A'. intros ? ? ? ? ? ? h1; red; intros. pose proof (Add_intro1 _ A a x H1) as h2. apply h1 in h2. destruct h2; subst; auto. destruct H2. contradiction. Qed. (*Wyckoff*) Lemma not_inh_empty : forall (T:Type), ~Inhabited (Empty_set T). intros ? h1; destruct h1; contradiction. Qed. Inductive well_founded_ens {T} : Ensemble T -> Prop := | well_founded_ens0 : well_founded_ens (Empty_set _) | well_founded_ensS : forall A x, ~Inn A x -> well_founded_ens A -> well_founded_ens (Add A x) | well_founded_ens_lim : forall A:Ensemble T, well_founded_ens A -> forall B:Ensemble T, well_founded_ens B -> well_founded_ens (Union A B). Inductive in_closed {T} : Ensemble T -> relation T := | in_closed_intro : forall A B x y, Inn A x -> well_founded_ens A -> Inn B x -> well_founded_ens B -> in_closed (Union A B) x y. Inductive const_on_of {T U} (f:T->U) (A:Ensemble T) : U -> Prop := const_on_of_intro : forall c, (forall x, Inn A x -> f x = c) -> const_on_of f A c. Inductive const_on {T U} (f:T->U) (A:Ensemble T) : Prop := const_on_intro : forall c, const_on_of f A c -> const_on f A. Lemma const_on_app_ex : forall {T U} (f:T->U) (A:Ensemble T), const_on f A -> Inhabited A -> exists! u, forall x:T, Inn A x -> f x = u. intros T U f A h1 h2. destruct h2 as [x h2]. destruct h1 as [y h1]. destruct h1 as [y h1]. exists y. red; split; auto. intros y' h3. pose proof (h1 x h2); subst; auto. Qed. Definition const_on_app {T U} {f:T->U} {A:Ensemble T} (pfc:const_on f A) (pfinh:Inhabited A) : U := proj1_sig (constructive_definite_description _ (const_on_app_ex f A pfc pfinh)). Definition const_on_app_compat : forall {T U} {f:T->U} {A:Ensemble T} (pfc:const_on f A) (pfinh:Inhabited A), let y := const_on_app pfc pfinh in forall x:T, Inn A x -> f x = y. unfold const_on_app; intros; destruct constructive_definite_description; auto. Qed. (*Wyckoff*) Lemma in_sub : forall {T:Type} (A:Ensemble T) (a x:T), In (Subtract A a) x -> In A x. intros ? ? ? ? h1. destruct h1; auto. Qed. (*Wyckoff*) Lemma inh_subtract_empty : forall {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (x:T), Inhabited A -> Subtract A x = Empty_set _ -> Inn A x. intros T h0 A x h1 h2. destruct h1 as [x' h1]. destruct (h0 x x') as [h3 | h3]; subst; auto. pose proof (Subtract_intro _ _ _ _ h1 h3) as h4. rewrite h2 in h4. contradiction. Qed. (*Wyckoff*) Lemma subtract_sing : forall {T:Type} (a:T), Subtract (Singleton a) a = Empty_set _. intros T a. apply Extensionality_Ensembles; red; split; auto with sets. red; intros x h1. destruct h1 as [h1 h2]. contradict h1; auto. Qed. (*Wyckoff*) Lemma subtract_sing_eq : forall {T:Type} (a b:T), a = b -> Subtract (Singleton a) b = Empty_set _. intros; subst; apply subtract_sing. Qed. (*Wyckoff*) Lemma in_sing_iff : forall {T:Type} (a b:T), In (Singleton a) b <-> a = b. intros; split; intro h1; destruct h1; subst; try constructor; auto. Qed. (*Wyckoff*) Lemma sing_not_empty : forall {T:Type} (a:T), Singleton a <> Empty_set _. intros T a h1. pose proof (In_singleton _ a) as h2. rewrite h1 in h2. contradiction. Qed. (*Wyckoff*) Lemma couple_not_empty : forall {T:Type} (a b:T), Couple a b <> Empty_set _. intros T a b h1. pose proof (Couple_l _ a b) as h2. rewrite h1 in h2. contradiction. Qed. (*Wyckoff*) Lemma card_sing : forall {T:Type} (x:T), cardinal _ (Singleton x) 1. intros T x. pose proof (add_empty_sing x) as h1. rewrite <- h1. constructor. constructor. auto with sets. Qed. (*Wyckoff*) Lemma inh_sing : forall {T:Type} (a:T), Inhabited (Singleton a). intros T a. apply Inhabited_intro with a. constructor. Qed. (*Wyckoff*) Lemma inh_add : forall {T:Type} (A:Ensemble T) (a:T), Inhabited (Add A a). intros; econstructor; right; constructor. Qed. (*Wyckoff*) Lemma inh_couple : forall {T:Type} (a b:T), Inhabited (Couple a b). intros; econstructor; left; constructor. Qed. (*psa = power set algebra.*) (*Wyckoff*) Lemma abs_sum_psa : forall {T:Type} (N M:(Ensemble T)), Union N (Intersection N M) = N. intros T N M. apply Extensionality_Ensembles. unfold Same_set. split; unfold Included. (*left*) intros x h1. inversion h1 as [x' h2 h3 | x'' h4 h5]. assumption. inversion h4 as [x''' h6 h7]. assumption. (*right*) intros x h8. apply Union_introl. assumption. Qed. (*Wyckoff*) Lemma abs_prod_psa : forall {T:Type} (N M:(Ensemble T)), Intersection N (Union N M) = N. intros T N M. apply Extensionality_Ensembles. unfold Same_set. split; unfold Included. (*left*) intros x h1. inversion h1 as [x' h2 h3 h4]. assumption. (*right*) intros x h5. assert (h6:In (Union N M) x). apply Union_introl; assumption. apply Intersection_intro; assumption. Qed. (*Wyckoff*) Lemma assoc_prod_psa : forall {T:Type} (N M P:Ensemble T), (Intersection N (Intersection M P)) = (Intersection (Intersection N M) P). intros T N M P. apply Extensionality_Ensembles. unfold Same_set. split; unfold Included. (*left*) intros x h1. inversion h1 as [x' h2 h3 h4]. inversion h3 as [x'' h5 h6 h7]. assert (h8: In (Intersection N M) x). apply Intersection_intro; assumption. apply Intersection_intro; assumption. (*right*) intros x h9. inversion h9 as [x' h10 h11 h12]. inversion h10 as [x'' h13 h14 h15]. assert (h16: In (Intersection M P) x). apply Intersection_intro; assumption. apply Intersection_intro; assumption. Qed. (*Wyckoff*) Lemma Intersection_associative : forall {T:Type} (A B C:Ensemble T), Intersection (Intersection A B) C = Intersection A (Intersection B C). intros; rewrite <- assoc_prod_psa; auto. Qed. (*Wyckoff*) Lemma sub_add_comm : forall {T:Type} (A:Ensemble T) (x y:T), (x <> y) -> (Subtract (Add A x) y) = (Add (Subtract A y) x). intros T A x y h0. apply Extensionality_Ensembles. red. split. (* <= *) red. intros z h1. destruct h1 as [h1 h2]. destruct h1 as [h1 | z h3]. left. constructor; auto. right; auto. (* >= *) red. intros z h1. destruct h1 as [z h1 | z h2]. destruct h1 as [h1 h3]. constructor. constructor; auto. auto. destruct h2; subst. constructor. apply Add_intro2. intro h1. destruct h1. contradict h0. auto. Qed. (*Wyckoff*) Lemma sub_add_eq : forall {T:Type} (A:Ensemble T) (a:T), (Subtract (Add A a) a) = Subtract A a. intros T A a. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1 h2]. rewrite in_sing_iff in h2. destruct h1 as [x h1 | x h1]. constructor; auto. rewrite in_sing_iff; auto. destruct h1. contradict h2; auto. destruct h1 as [h1 h2]. rewrite in_sing_iff in h2. constructor. left; auto. rewrite in_sing_iff. assumption. Qed. (*Wyckoff*) Lemma add_sub_eq : forall {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (a:T), Add (Subtract A a) a = Add A a. intros T h1 A a. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2 as [x h2 | x h2]. destruct h2 as [h2 h3]. left; auto. destruct h2. right. constructor; auto. destruct (h1 a x) as [h3 | h3]; subst. right. constructor; auto. left. destruct h2 as [x h2 | x h2]. constructor; auto. rewrite in_sing_iff; auto. destruct h2. contradict h3; auto. Qed. (*Wyckoff*) Lemma couple_singleton : forall {T:Type} (x:T), Couple x x = Singleton x. intros T x. apply Extensionality_Ensembles. red. split. (* <= *) red. intros ? h1. destruct h1; constructor. (* >= *) red. intros ? h1. destruct h1; constructor. Qed. (*Wyckoff*) Lemma couple_eq_sing : forall {T:Type} (a b c:T), Couple a b = Singleton c -> a = b /\ a = c. intros T a b c h1. pose proof (Couple_l _ a b) as h2. rewrite h1 in h2. destruct h2. pose proof (Couple_r _ c b) as h2. rewrite h1 in h2. destruct h2. split; auto. Qed. (*Wyckoff*) Lemma couple_inv_dec : forall {T:Type} (a b x:T), Inn (Couple a b) x -> {x = a} + {x = b}. intros T a b x h1. apply Couple_inv in h1. apply or_to_sumbool in h1. assumption. Qed. (*Wyckoff*) Lemma incl_sing_couple_l : forall {T:Type} (a b:T), Included (Singleton a) (Couple a b). intros; red; intros x h8; destruct h8; left; auto. Qed. (*Wyckoff*) Lemma incl_sing_couple_r : forall {T:Type} (a b:T), Included (Singleton b) (Couple a b). intros; red; intros x h8; destruct h8; right; auto. Qed. (*Wyckoff*) Lemma empty_union : forall {T:Type} (S:Ensemble T), Union (Empty_set T) S = S. intros T S. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h1. destruct h1; [contradiction | assumption]. (* >= *) red. intros x h1. apply Union_intror; assumption. Qed. (*Wyckoff*) Lemma empty_intersection : forall {T:Type} (S:Ensemble T), Intersection (Empty_set T) S = Empty_set T. intros T S. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h1. destruct h1; contradiction. (* >= *) auto with sets. Qed. (*Wyckoff*) Lemma Intersection_idempotent : forall {T:Type} (A:Ensemble T), Intersection A A = A. intros T A. apply Extensionality_Ensembles. red. split. red. intros ? h1. destruct h1; auto. red. intros ? h1. constructor; auto. Qed. (*Wyckoff*) Lemma singleton_inj : forall {T:Type} (a b:T), Singleton a = Singleton b -> a = b. intros T a b h1. pose proof (In_singleton _ a) as h2. rewrite h1 in h2. destruct h2; reflexivity. Qed. (*Wyckoff*) Definition sig_set {T:Type} (A:Ensemble T) := {x:T | In A x}. (*Wyckoff*) Lemma sig_set_eq : forall {T:Type} (A B:Ensemble T), A = B -> sig_set A = sig_set B. intros; subst; auto. Defined. (*Wyckoff*) Lemma f_equal_sig_set : forall {T} {A B:Ensemble T} (pfe:A = B) (x:sig_set A) (y:sig_set B), proj1_sig x = proj1_sig y -> existT id (sig_set A) x = existT id (sig_set B) y. intros ? ? ? ? ? ? h1; subst; apply proj1_sig_injective in h1; subst; auto. Qed. (*Wyckoff*) Lemma sig_set_eq_dec : forall {T:Type} (pf:type_eq_dec T) (E:Ensemble T), type_eq_dec (sig_set E). intros; apply sig_eq_dec; auto. Qed. (*Wyckoff*) Definition sig_fun_subst {T U:Type} {A B:Ensemble T} (pf:A = B) (f:sig_set A->U) : sig_set B->U := (fun c => f (exist _ (proj1_sig c) (in_eq_set _ _ (eq_sym pf) (proj1_sig c) (proj2_sig c)))). (*Wyckoff*) Lemma sig_fun_subst_compat : forall {T U:Type} {A B:Ensemble T} (pfe:A = B) (f:sig_set A->U) (x:T) (pfx:Inn B x), sig_fun_subst pfe f (exist _ _ pfx) = f (exist _ _ (in_eq_set _ _ (eq_sym pfe) x pfx)). unfold sig_fun_subst; intros; repeat f_equal. Qed. (*Wyckoff*) Definition full_sig {T:Type} (A:Ensemble T) := Full_set (sig_set A). (*Wyckoff*) Definition full_set_sig_conv {T:Type} (x:T) : sig_set (Full_set T) := exist _ _ (Full_intro _ x). (*Wyckoff*) Lemma inh_full_sig : forall {T:Type} (E:Ensemble T), Inhabited E -> Inhabited (full_sig E). intros T E h1. destruct h1 as [x h1]. apply Inhabited_intro with (exist _ _ h1). constructor. Qed. (*Wyckoff*) Lemma in_im : forall {T U} x (f:T->U), Inn (Im (Full_set _) f) (f x). intros; econstructor; constructor. Qed. (*Wyckoff*) Lemma im_singleton : forall {T U:Type} (x:T) (f:T->U), Im (Singleton x) f = Singleton (f x). intros T U x f. apply Extensionality_Ensembles. red. split. red. intros y h1. destruct h1 as [z h2]. subst. destruct h2; subst. constructor; auto. red. intros y h1. destruct h1; subst. apply Im_intro with x; auto; constructor. Qed. (*Wyckoff*) Lemma full_sig_sing : forall {T:Type} (x:T), full_sig (Singleton x) = Singleton (exist (fun c => In (Singleton x) c) _ (In_singleton _ _)). intros T x. unfold full_sig. apply Extensionality_Ensembles; red; split; red; intros a h2. clear h2. destruct a as [a h1]. destruct h1. constructor. constructor. Qed. (*Wyckoff*) Lemma im_couple : forall {T U:Type} (x y:T) (f:T->U), Im (Couple x y) f = Couple (f x) (f y). intros T U x y f. apply Extensionality_Ensembles; red; split; red; intros a h1. destruct h1 as [a h1]. subst. destruct h1; [left | right]. inversion h1; subst; [apply Im_intro with x; auto | apply Im_intro with y; auto]. left. right. Qed. (*Wyckoff*) Lemma im_preserves_inclusion : forall {T U:Type} (A B:Ensemble T) (f:T->U), Included A B -> Included (Im A f) (Im B f). intros T U A B f h1. red. intros u h2. destruct h2 as [u h2]. subst. apply h1 in h2. apply Im_intro with u; auto. Qed. (*Wyckoff*) Lemma im_im : forall {T U V:Type} (A:Ensemble T) (f:T->U) (g:U->V), Im (Im A f) g = Im A (fun x => g (f x)). intros T U V A f g. apply Extensionality_Ensembles. red. split. (* <= *) red. intros v h1. inversion h1 as [x h2 v' h3 h4]. subst. destruct h2 as [x h2 u h3]. subst. apply Im_intro with x. assumption. reflexivity. (* >= *) red. intros v h1. destruct h1 as [x h2 v' h3]. subst. apply Im_intro with (f x). apply Im_intro with x; auto. reflexivity. Qed. (*Wyckoff*) Lemma incl_im_subtract : forall {T U:Type} (A:Ensemble T) (a:T) (f:T->U), Included (Subtract (Im A f) (f a)) (Im (Subtract A a) f). intros T U A a f; red; intros x h1. destruct h1 as [h1 h2]. destruct h1 as [x h1]. subst. apply Im_intro with x; auto. constructor; auto. intro h3. destruct h3. contradict h2. constructor. Qed. (*Wyckoff*) Lemma im_subtract_inj : forall {T U:Type} (A:Ensemble T) (a:T) (f:T->U), injective f -> Im (Subtract A a) f = Subtract (Im A f) (f a). intros T U A a f h0. apply Extensionality_Ensembles; red; split; red; intros x h1; try apply incl_im_subtract; auto. destruct h1 as [x h1]. subst. destruct h1 as [h1 h2]. red in h0. constructor. apply Im_intro with x; auto. intro h3. inversion h3 as [h4]. clear h3. apply h0 in h4. subst. contradict h2. constructor. Qed. (*Wyckoff*) Lemma empty_image : forall {T U:Type} (A:Ensemble T) (f:T->U), Im A f = Empty_set _ -> A = Empty_set _. intros T U A f h1. apply NNPP. intro h2. pose proof (not_empty_Inhabited _ _ h2) as h3. destruct h3 as [x h3]. assert (h4:Inn (Im A f) (f x)). apply Im_intro with x. assumption. reflexivity. rewrite h1 in h4. contradiction. Qed. (*Wyckoff*) Lemma inh_im : forall {T U:Type} (E:Ensemble T) (f:T->U), Inhabited E -> Inhabited (Im E f). intros T U E f h1. destruct h1 as [x h1]. apply Inhabited_intro with (f x); auto. apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma im_const_fun : forall {T U:Type} (A:Ensemble T) (u:U), Inhabited A -> Im A (fun _:T => u) = Singleton u. intros T U A u h1. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2 as [x h2]; subst. constructor. destruct h2. destruct h1 as [x h1]. apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma im_full_sig_proj1_sig : forall {T:Type} (S:Ensemble T), S = Im (full_sig S) (@proj1_sig _ _). intros T S. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h1. apply Im_intro with (exist _ x h1). constructor. simpl. reflexivity. (* >= *) red. intros x h1. inversion h1 as [x' h2 ? h3 h4]. rewrite h3. apply proj2_sig. Qed. (*Wyckoff*) Lemma im_id : forall {T:Type} (S:Ensemble T), S = Im S id. intros T S. apply Extensionality_Ensembles. red. split. red. intros x h1. apply Im_intro with x; auto. red. intros x h1. destruct h1 as [x h1]. subst. unfold id. assumption. Qed. (*Wyckoff*) Lemma im_full_sig_im_eq : forall {T U V:Type} (A:Ensemble T) (f:T->U) (g:sig_set (Im A f)-> V), Im (full_sig (Im A f)) g = Im (full_sig A) (fun x:sig_set A => g (exist _ _ (Im_intro _ _ A f _ (proj2_sig x) _ (eq_refl _)))). intros T U V A f g. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. subst. clear h1. destruct x as [x h2]. destruct h2 as [x h2]. subst. apply Im_intro with (exist _ _ h2). constructor. f_equal. red. intros x h1. destruct h1 as [x h1]. subst. clear h1. destruct x as [x h1]. assert (h2:In (Im A f) (f x)). apply Im_intro with x; auto. apply Im_intro with (exist _ _ h2). constructor. f_equal. apply proj1_sig_injective. simpl. reflexivity. Qed. (*Wyckoff*) Lemma im_ext_in : forall {T U:Type} (A:Ensemble T) (f g:T->U), (forall x:T, In A x -> f x = g x) -> Im A f = Im A g. intros T U A f g h1. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. rewrite h1; auto. apply Im_intro with x; auto. red. intros x h2. destruct h2 as [x h2]. subst. rewrite <- h1; auto. apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma im_functional1 : forall {T U:Type} {A B:Ensemble T} (pf:A = B) (f:T->U), Im A f = Im B f. intros; subst; auto. Qed. (*Wyckoff*) Lemma im_functional2 : forall {T U:Type} {f g:T->U} (pf:f = g) (A:Ensemble T), Im A f = Im A g. intros; subst; auto. Qed. (*Wyckoff*) Lemma im_functional : forall {T U:Type} {A B:Ensemble T} {f g:T->U} (pfab:A = B) (pffg:f = g), Im A f = Im A g. intros; subst; auto. Qed. (*Wyckoff*) Definition im_proj2_sig {T:Type} {P:T->Prop} (A:Ensemble T) (pfx:(forall x:T, Inn A x -> P x)) : Ensemble {x | P x} := Im (full_sig A) (fun p => (exist _ _ (pfx _ (proj2_sig p)))). (*Wyckoff*) Definition im_in {T U:Type} (A:Ensemble T) (P:T->Prop) (f:{x:T|P x}->U) (pfx:(forall x:T, Inn A x -> P x)) : Ensemble U := Im (full_sig A) (fun x:{t:T|In A t}=>f (exist _ (proj1_sig x) (pfx _ (proj2_sig x)))). (*Wyckoff*) Lemma im_full_sig_subtract_in_inj : forall {T U:Type} (E:Ensemble T) (x:T) (f:T->U), injective f -> let f' := (fun a:sig_set (Subtract E x) => f (proj1_sig a)) in In E x -> Im (full_sig (Subtract E x)) f' = Subtract (Im E f) (f x). intros T U E x f h1 f' h2. apply Extensionality_Ensembles; red; split; red; intros a h3. destruct h3 as [a h3]. subst. clear h3. destruct a as [a h3]. destruct h3 as [h3 h4]. constructor. apply Im_intro with a; auto. intro h5. inversion h5 as [h6]. clear h5. red in h1. apply h1 in h6. simpl in h6. clear f'. subst. contradict h4. constructor. destruct h3 as [h3 h4]. destruct h3 as [a h3]. subst. assert (h5:In (Subtract E x) a). constructor; auto. intro h5. destruct h5. contradict h4. constructor. apply Im_intro with (exist _ _ h5). constructor. unfold f'. simpl. reflexivity. Qed. (*Wyckoff*) (*assigns image to range*) Definition sig_im_fun {T U:Type} (f:T->U) : T->sig_set (Im (Full_set T) f) := fun x => exist _ _ (Im_intro _ _ _ f _ (Full_intro _ x) _ (@eq_refl _ _)). (*Wyckoff*) Lemma surj_sig_im_fun : forall {T U:Type} (f:T->U), surjective (sig_im_fun f). intros T U f. red. intro y. destruct y as [y h1]. destruct h1 as [x]. subst. exists x. unfold sig_im_fun. apply proj1_sig_injective. simpl. reflexivity. Qed. (*Wyckoff*) Lemma inj_bij_sig_im_fun_iff : forall {T U:Type} (f:T->U), injective f <-> bijective (sig_im_fun f). intros T U f. split. intro h1. red. split. red in h1. red. intros a b h2. unfold sig_im_fun in h2. apply exist_injective in h2. apply h1; auto. apply surj_sig_im_fun. intro h1. red in h1. destruct h1 as [h1l h1r]. red in h1l. red. intros a b h2. specialize (h1l a b). apply h1l. unfold sig_im_fun. apply proj1_sig_injective. simpl. assumption. Qed. (*Wyckoff*) Lemma inj_sig_pred : forall (E:Ensemble nat), injective (fun n : sig_set (Subtract E 0) => pred (proj1_sig n)). intro E. red; intros m n h1. destruct m as [m h2], n as [n h3]. simpl in h1. destruct h2 as [h2a h2b], h3 as [h3a h3b]. destruct (zerop m) as [h4 | h4], (zerop n) as [h5 | h5]; subst. f_equal. apply proof_irrelevance. contradict h2b. constructor; auto. contradict h3b. constructor; auto. apply pred_inj0 in h1. subst. f_equal. apply proof_irrelevance. assumption. assumption. Qed. Section Disjoint. (*Wyckoff*) Definition disjoint {T:Type} (A B:Ensemble T) := Intersection A B = Empty_set _. (*Wyckoff*) Lemma sym_disjoint : forall {T:Type} (A B:Ensemble T), disjoint A B -> disjoint B A. intros T A B h1. unfold disjoint in h1. rewrite Intersection_commutative in h1. red. assumption. Qed. (*Wyckoff*) Definition disjoint_incl_comp_l : forall {T:Type} (A B:Ensemble T), disjoint A B -> Included A (Comp B). intros T A B h1. red. intros x h2. intro h3. red in h1. pose proof (Intersection_intro _ _ _ _ h2 h3) as h4. rewrite h1 in h4. contradiction. Qed. (*Wyckoff*) Definition disjoint_incl_comp_r : forall {T:Type} (A B:Ensemble T), disjoint A B -> Included B (Comp A). intros T A B h1. red. intros x h2. intro h3. red in h1. pose proof (Intersection_intro _ _ _ _ h3 h2) as h4. rewrite h1 in h4. contradiction. Qed. (*Wyckoff*) Lemma disjoint_setminus_compat : forall {T:Type} (A B:Ensemble T), disjoint A B -> Setminus A B = A. intros T A B h1. red in h1. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2; auto. constructor; auto. intro h3. pose proof (Intersection_intro _ A B _ h2 h3) as h4. rewrite h1 in h4. destruct h4. Qed. End Disjoint. Section ImProjSig. (*Wyckoff*) Definition im_proj1_sig {T:Type} {A:Ensemble T} (S:Ensemble (sig_set A)) : Ensemble T := Im S (@proj1_sig _ _). (*Wyckoff*) Definition im_proj1_sig' {T:Type} {P:T->Prop} (S:Ensemble {x:T | P x}) : Ensemble T := Im S (@proj1_sig _ _). (*Wyckoff*) Lemma im_proj1_sig_empty : forall {T:Type} (A:Ensemble T), im_proj1_sig (Empty_set (sig_set A)) = Empty_set _. unfold im_proj1_sig; intros; apply image_empty. Qed. (*Wyckoff*) Lemma im_proj1_sig_couple : forall {T:Type} {A:Ensemble T} (x y:sig_set A), im_proj1_sig (Couple x y) = Couple (proj1_sig x) (proj1_sig y). intros T A x y. unfold im_proj1_sig. apply im_couple. Qed. (*Wyckoff*) Lemma im_proj1_sig_injective : forall {T:Type} {A:Ensemble T} (R S:Ensemble (sig_set A)), im_proj1_sig R = im_proj1_sig S -> R = S. intros T A R S h1. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct x as [x h3]. assert (h4:In (im_proj1_sig R) x). apply Im_intro with (exist _ _ h3). assumption. simpl. reflexivity. rewrite h1 in h4. destruct h4 as [x h4]. subst. rewrite unfold_sig in h4. assert (h3 = proj2_sig x). apply proof_irrelevance. subst. assumption. red. intros x h2. destruct x as [x h3]. assert (h4:In (im_proj1_sig S) x). apply Im_intro with (exist _ _ h3). assumption. simpl. reflexivity. rewrite <- h1 in h4. destruct h4 as [x h4]. subst. rewrite unfold_sig in h4. assert (h3 = proj2_sig x). apply proof_irrelevance. subst. assumption. Qed. (*Wyckoff*) Lemma im_proj1_sig_undoes_im_proj2_sig : forall {T:Type} (A B:Ensemble T) (pf:Included A B), A = im_proj1_sig (im_proj2_sig A pf). intros T A B h1. unfold im_proj1_sig, im_proj2_sig. rewrite im_im. rewrite im_full_sig_proj1_sig at 1. simpl. apply im_ext_in. intros; auto. Qed. (*Wyckoff*) Lemma incl_im_proj1_sig_ens_sig_set : forall {T:Type} {A:Ensemble T} (B:Ensemble (sig_set A)), Included (im_proj1_sig B) A. intros T A B. red. intros x h1. destruct h1 as [x h1]. subst. apply proj2_sig. Qed. (*Wyckoff*) Lemma im_proj2_sig_undoes_im_proj1_sig : forall {T:Type} {B:Ensemble T} (A:Ensemble (sig_set B)), A = im_proj2_sig (im_proj1_sig A) (incl_im_proj1_sig_ens_sig_set A). intros T B A. unfold im_proj1_sig, im_proj2_sig. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct x as [x h2]. assert (h3:In (Im A (@proj1_sig _ _)) x). apply Im_intro with (exist _ _ h2). assumption. simpl. reflexivity. apply Im_intro with (exist _ _ h3). constructor. apply proj1_sig_injective. simpl. reflexivity. red. intros x h1. destruct h1 as [x h1]. subst. destruct x as [x h2]. simpl. destruct h2 as [x h2]. subst. pose proof h2 as h2'. rewrite unfold_sig in h2'. pose proof (proof_irrelevance _ (incl_im_proj1_sig_ens_sig_set A (proj1_sig x) (Im_intro (sig_set B) T A (proj1_sig (P:=fun x0 : T => In B x0)) x h2 (proj1_sig x) eq_refl)) (proj2_sig x)) as h3. rewrite h3; auto. Qed. (*Wyckoff*) Lemma im_proj2_sig_undoes_im_proj1_sig' : forall {T:Type} (B:Ensemble T) (A:Ensemble (sig_set B)) (pf:Included (im_proj1_sig A) B), A = im_proj2_sig (im_proj1_sig A) pf. intros T B A h1. pose proof (im_proj2_sig_undoes_im_proj1_sig A) as h2. pose proof (proof_irrelevance _ h1 (incl_im_proj1_sig_ens_sig_set A)) as h3. rewrite h3; auto. Qed. (*Wyckoff*) Lemma in_im_proj1_sig_iff : forall {T:Type} (B:Ensemble T) (A:Ensemble (sig_set B)) (x:sig_set B), In (im_proj1_sig A) (proj1_sig x) <-> In A x. intros T B A x. split. intro h1. inversion h1 as [i h2 ? h3 h4]. subst. apply proj1_sig_injective in h3. subst. assumption. intro h1. apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma in_im_proj1_sig : forall {T:Type} {B:Ensemble T} (A:Ensemble (sig_set B)) (x:T), In (im_proj1_sig A) x -> exists pf:In B x, In A (exist _ _ pf). intros T B A x h1. destruct h1 as [x h1]. subst. exists (proj2_sig _). rewrite <- unfold_sig. assumption. Qed. (*Wyckoff*) Lemma in_im_proj1_sig_inv_ex : forall {T:Type} (B:Ensemble T) (A:Ensemble (sig_set B)) (x:T), In (im_proj1_sig A) x -> exists! b:sig_set B, x = proj1_sig b. intros T B A x h1. destruct h1 as [x h1]. subst. exists x. red. split; auto. intros h h2. apply proj1_sig_injective in h2. assumption. Qed. (*Wyckoff*) Definition in_im_proj1_sig_inv {T:Type} (B:Ensemble T) (A:Ensemble (sig_set B)) (x:T) (pf:In (im_proj1_sig A) x) := proj1_sig (constructive_definite_description _ (in_im_proj1_sig_inv_ex B A x pf)). (*Wyckoff*) Lemma in_im_proj1_sig_inv_compat : forall {T:Type} (B:Ensemble T) (A:Ensemble (sig_set B)) (x:T) (pf:In (im_proj1_sig A) x), let b := in_im_proj1_sig_inv B A x pf in x = proj1_sig b. unfold in_im_proj1_sig_inv; intros; destruct constructive_definite_description; auto. Qed. (*Wyckoff*) Lemma in_im_proj1_sig' : forall {T:Type} (B:Ensemble T) (A:Ensemble (sig_set B)) (x:T) (pf:In (im_proj1_sig A) x), In A (in_im_proj1_sig_inv B A x pf). intros T B A x h1. pose proof (in_im_proj1_sig_inv_compat B A x h1) as h2. simpl in h2. destruct h1 as [x h1]. subst. apply proj1_sig_injective in h2. pose proof h1 as h1'. rewrite h2 in h1' at 1. assumption. Qed. (*Wyckoff*) Lemma incl_im_proj1_sig_iff : forall {T:Type} {S:Ensemble T} (A B:Ensemble (sig_set S)), Included (im_proj1_sig A) (im_proj1_sig B) <-> Included A B. intros T S A B. split. intro h1. red. intros x h2. rewrite <- in_im_proj1_sig_iff in h2. apply h1 in h2. inversion h2 as [x' h3 q h4]. subst. apply proj1_sig_injective in h4. subst. assumption. intro h1. red. intros x h2. destruct h2 as [x h2]. subst. apply h1 in h2. apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma im_proj1_sig_add_eq : forall {T:Type} {E:Ensemble T} (A:Ensemble (sig_set E)) (a:sig_set E), im_proj1_sig (Add A a) = Add (im_proj1_sig A) (proj1_sig a). intros T E A a. apply Extensionality_Ensembles; red; split. red; intros x h1. destruct h1 as [x h1]. subst. destruct h1 as [x h1 | x h2]. left. apply Im_intro with x; auto. destruct h2. right. constructor. red. intros x h1. destruct h1 as [x h1 | x h2]. destruct h1 as [x h1]. subst. apply Im_intro with x. left; auto. reflexivity. destruct h2. apply Im_intro with a. right. constructor. reflexivity. Qed. (*Wyckoff*) Lemma im_proj1_sig_subtract_eq : forall {T:Type} {E:Ensemble T} (A:Ensemble (sig_set E)) (a:sig_set E), im_proj1_sig (Subtract A a) = Subtract (im_proj1_sig A) (proj1_sig a). intros T E A a. apply Extensionality_Ensembles; red; split. red; intros x h1. destruct h1 as [x h1]. subst. destruct h1 as [h1 h2]. constructor. apply Im_intro with x; auto. intro h3. inversion h3 as [h4]. clear h3. apply proj1_sig_injective in h4. subst. contradict h2. constructor. red. intros x h1. destruct h1 as [h1 h2]. destruct h1 as [x h1]. subst. apply Im_intro with x. constructor; auto. intro h3. destruct h3. contradict h2. constructor. reflexivity. Qed. (*Wyckoff*) Lemma im_proj1_sig_sing : forall {T:Type} (E:Ensemble T) (x:sig_set E), im_proj1_sig (Singleton x) = Singleton (proj1_sig x). intros T E x. unfold im_proj1_sig. rewrite im_singleton. reflexivity. Qed. (*Wyckoff*) Lemma finite_im_proj1_sig_iff : forall {T:Type} {A:Ensemble T} (E:Ensemble (sig_set A)), Finite (im_proj1_sig E) <-> Finite E. intros T A E. split. intro h1. apply NNPP. intro h2. pose proof (Pigeonhole_bis _ _ E _ h2 h1) as h3. contradict h3. rewrite <- inj_image_fun_prop_compat. red; intros; apply proj1_sig_injective; auto. intro h1. apply finite_image; auto. Qed. (*Wyckoff*) Lemma inh_im_proj1_sig_iff : forall {T:Type} {A:Ensemble T} (E:Ensemble (sig_set A)), Inhabited (im_proj1_sig E) <-> Inhabited E. intros T A E. split. intro h1. destruct h1 as [x h1]. destruct h1 as [x h1]. subst. apply Inhabited_intro with x; auto. intro h1. destruct h1 as [x h1]. apply Inhabited_intro with (proj1_sig x). apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma im_proj1_sig_full_set_sig : forall {T:Type} (S:Ensemble T), im_proj1_sig (Full_set {x:T | In S x}) = S. intros T S. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]. subst. apply proj2_sig. apply Im_intro with (exist _ _ h1). constructor. simpl. reflexivity. Qed. (*Wyckoff*) Lemma sig_set_im_proj1_sig_in_aux : forall {T:Type} {A:Ensemble T} {E:Ensemble (sig_set A)} (x:sig_set (im_proj1_sig E)), In A (proj1_sig x). intros T A E x. destruct x as [x h1]. simpl. destruct h1 as [x h1]. subst. apply proj2_sig. Qed. (*Wyckoff*) Lemma sig_set_im_proj1_sig_in : forall {T:Type} {A:Ensemble T} {E:Ensemble (sig_set A)} (x:sig_set (im_proj1_sig E)), In E (exist _ _ (sig_set_im_proj1_sig_in_aux x)). intros T A E x. destruct x as [x h1]. simpl. destruct h1 as [x h1]. subst. destruct x as [x h2]. simpl. pose proof (proof_irrelevance _ h2 (sig_set_im_proj1_sig_in_aux (exist (fun x0 : T => In (im_proj1_sig E) x0) x (Im_intro (sig_set A) T E (proj1_sig (P:=fun x0 : T => In A x0)) (exist (fun x0 : T => In A x0) x h2) h1 x eq_refl)))) as h3. rewrite <- h3; auto. Qed. (*Wyckoff*) Lemma im_proj1_sig_eq_iff : forall {T:Type} {S:Ensemble T} (A B:Ensemble (sig_set S)), im_proj1_sig A = im_proj1_sig B <-> A = B. intros T S A B. split. intro h1. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct x as [x h3]. assert (h4:In (im_proj1_sig A) x). apply Im_intro with (exist _ _ h3); auto. rewrite h1 in h4. destruct h4 as [x h4]. subst. destruct x as [x h5]. simpl. simpl in h3. assert (h5 = h3). apply proof_irrelevance. subst. assumption. destruct x as [x h3]. assert (h4:In (im_proj1_sig A) x). rewrite h1. apply Im_intro with (exist _ _ h3); auto. destruct h4 as [x h4]. subst. destruct x as [x h5]. simpl in h2. simpl. pose proof (proof_irrelevance _ h3 h5); subst; auto. intro; subst; auto. Qed. (*Wyckoff*) Lemma intersection_im_proj1_sig_eq : forall {T:Type} {S:Ensemble T} (A B:Ensemble (sig_set S)), Intersection (im_proj1_sig A) (im_proj1_sig B) = im_proj1_sig (Intersection A B). intros T S A B. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1 h2]. destruct h1 as [x h1 ? h4], h2 as [x' h2 ? h3]. subst. apply proj1_sig_injective in h3. subst. apply Im_intro with x'; auto. constructor; auto. destruct h1 as [x h1]. subst. destruct h1 as [x h1 h2]. constructor. apply Im_intro with x; auto. apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma im_proj1_sig_disjoint_iff : forall {T:Type} {S:Ensemble T} (A B:Ensemble (sig_set S)), disjoint A B <-> disjoint (im_proj1_sig A) (im_proj1_sig B). intros T S A B. split. intro h1. red in h1. red. rewrite intersection_im_proj1_sig_eq. rewrite h1. unfold im_proj1_sig. rewrite image_empty. reflexivity. intro h1. red in h1. rewrite intersection_im_proj1_sig_eq in h1. apply im_proj1_sig_injective. unfold im_proj1_sig at 2. rewrite image_empty. assumption. Qed. Lemma in_proj1_sig_fun : forall {T U:Type} {P:U->Prop} (f:T->{x:U | P x}) (A:Ensemble T) (y:U), In (Im A (proj1_sig_fun f)) y -> P y. intros T U P f A y h1. destruct h1; subst. rewrite proj1_sig_fun_compat. apply proj2_sig. Qed. End ImProjSig. Definition fun_in_ens {T:Type} (A:Ensemble T) (U:Type) : Type := prop_dep_fun (fun x => Inn A x) U. Definition fun_in_ens' {T U} (A:Ensemble T) (B:Ensemble U) : Type := fun_in_ens A (sig_set B). Definition fun_in_set {T:Type} (A:Ensemble T) (U:Set) : Type := prop_dep_fun (fun x => Inn A x) U. (*Like the [transfer] functions in [TypeUtilities], but for dependency reasons, we'll have a minor substitute.*) Definition fun_in_ens_subst {T:Type} {A A':Ensemble T} {U:Type} (pfe:A = A') (f:fun_in_ens A U) (pfe:A = A') : fun_in_ens A' U := fun x (pfx:Inn A' x) => let pfx' := in_eq_set _ _ (eq_sym pfe) _ pfx in f x pfx'. Definition fun_in_set_subst {T:Type} {A A':Ensemble T} {U:Set} (pfe:A = A') (f:fun_in_set A U) (pfe:A = A') : fun_in_set A' U := fun x (pfx:Inn A' x) => let pfx' := in_eq_set _ _ (eq_sym pfe) _ pfx in f x pfx'. Definition fun_in_ens_subst1 {T U} {A A':Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B) (pfe:A = A') : fun_in_ens' A' B := fun x (pfx:Inn A' x) => let pfx' := in_eq_set _ _ (eq_sym pfe) _ pfx in f x pfx'. Definition fun_in_ens_subst2 {T U} {A:Ensemble T} {B B':Ensemble U} (f:fun_in_ens' A B) (pfe:B = B') : fun_in_ens' A B' := fun x (pfx:Inn A x) => exist (Inn B') _ (in_eq_set _ _ pfe _ (proj2_sig (f x pfx))). Definition fun_eq_set {T:Type} (A:Ensemble T) (U:Set) : Type := prop_dep_fun (fun B => A = B) U. Inductive im_in_ens {T U:Type} {A:Ensemble T} (f:fun_in_ens A U) : Ensemble U:= | im_in_ens_intro : forall (x:T) (pf:In A x) (y:U), y = f x pf -> In (im_in_ens f) y. Inductive im_in_set {T:Type} {U:Set} {A:Ensemble T} (f:fun_in_ens A U) : Ensemble U:= | im_in_set_intro : forall (x:T) (pf:In A x) (y:U), y = f x pf -> In (im_in_set f) y. Lemma in_im_in_ens : forall {T U:Type} {A:Ensemble T} (f:fun_in_ens A U) (x:T) (pf:Inn A x), Inn (im_in_ens f) (f x pf). intros; econstructor; reflexivity. Qed. Lemma in_im_in_set : forall {T:Type} {U:Set} {A:Ensemble T} (f:fun_in_ens A U) (x:T) (pf:Inn A x), Inn (im_in_set f) (f x pf). intros; econstructor; reflexivity. Qed. Definition inj_in_ens {T U:Type} {A:Ensemble T} (f:fun_in_ens A U) : Prop := forall (x y:T) (pfx:Inn A x) (pfy:Inn A y), f _ pfx = f _ pfy -> x = y. Definition inj_in_ens' {T U:Type} {A:Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B) : Prop := inj_in_ens f. Definition surj_in_ens {T U:Type} {A:Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B) : Prop := forall y (pfy:Inn B y), exists x (pfx:Inn A x), f x pfx = exist _ y pfy. Definition bij_in_ens {T U:Type} {A:Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B) : Prop := inj_in_ens f /\ surj_in_ens f. Lemma surj_in_ens_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B), surj_in_ens f -> im_proj1_sig (im_in_ens f) = B. intros T U A B f h1. red in h1. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2 as [x h2]; subst; apply proj2_sig. specialize (h1 x h2). destruct h1 as [y [h1 h3]]. pose proof (f_equal (@proj1_sig _ _) h3) as h4; simpl in h4; subst. apply Im_intro with (f y h1);auto. apply im_in_ens_intro with y h1; auto. Qed. (*Wyckoff*) Inductive inverses_in_ens {T U} {A:Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B) : fun_in_ens' B A -> Prop := | inverses_in_ens_intro : forall g:fun_in_ens' B A, (forall x (pfx:Inn A x), g (proj1_sig (f x pfx)) (proj2_sig (f x pfx)) = exist _ x pfx) -> (forall y (pfy:Inn B y), f (proj1_sig (g y pfy)) (proj2_sig (g y pfy)) = exist _ y pfy) -> inverses_in_ens f g. (*Wyckoff*) Inductive invertible_in_ens {T U} {A:Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B) : Prop := | invertible_in_ens_intro : forall g:fun_in_ens' B A, inverses_in_ens f g -> invertible_in_ens f. (*Wyckoff*) Definition inv_in_ens_ex : forall {T U} {A:Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B), invertible_in_ens f -> exists! g:fun_in_ens' B A, inverses_in_ens f g. intros T U A B f h1. destruct h1 as [g hinv]. exists g; auto; red; split; auto. intros h h1. destruct h1 as [k h2]. rename H into h3. destruct hinv as [q h4]. rename H into h5. apply functional_extensionality_dep. intro y. apply functional_extensionality. intro hy. specialize (h3 y hy). pose proof (f_equal (@proj1_sig _ _) h3) as h6; simpl in h6. specialize (h4 (proj1_sig (k y hy)) (proj2_sig (k y hy))). revert h4. Gen0. rewrite h6; intros h7 h8. pose proof (proof_irrelevance _ h7 hy) as h9. rewrite <- h9, h8. apply proj1_sig_injective; simpl. repeat f_equal; auto. Qed. (*Wyckoff*) Definition inv_in_ens {T U} {A:Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B) (pfinv:invertible_in_ens f) : fun_in_ens' B A := proj1_sig (constructive_definite_description _ (inv_in_ens_ex f pfinv)). (*Wyckoff*) Definition inv_in_ens_compat : forall {T U} {A:Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B) (pfinv:invertible_in_ens f), inverses_in_ens f (inv_in_ens f pfinv). unfold inv_in_ens; intros; destruct constructive_definite_description; simpl; auto. Qed. (*Wyckoff*) Lemma inv_impl_bij_in_ens : forall {T U} {A:Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B), invertible_in_ens f -> bij_in_ens f. intros T U A B f h1. destruct h1 as [g h2]. destruct h2 as [g h2]. rename H into h3. red; split. red; intros x y hx hy h1. pose proof (h2 x hx) as h2x. pose proof (h2 y hy) as h2y. pose proof (f_equal (@proj1_sig _ _) h2x) as h3x; pose proof (f_equal (@proj1_sig _ _) h2y) as h3y; simpl in h3x, h3y. rewrite h1 in h3x. rewrite <- h1 in h3y. revert h3y. subst. intro h4. rewrite h2; simpl; auto. red; intros y hy. specialize (h3 y hy). pose proof (f_equal (@proj1_sig _ _) h3) as h4; simpl in h4. exists (proj1_sig (g y hy)), (proj2_sig _); apply proj1_sig_injective; simpl; auto. Qed. (*Wyckoff*) Definition inj_in_set {T:Type} {U:Set} {A:Ensemble T} (f:fun_in_ens A U) : Prop := forall (x y:T) (pfx:Inn A x) (pfy:Inn A y), f _ pfx = f _ pfy -> x = y. (*Wyckoff*) Definition fun_in_ens_subtract {T U:Type} {A:Ensemble T} (f:fun_in_ens A U) (a:T) : fun_in_ens (Subtract A a) U. do 2 red. intros x h1. destruct h1 as [h1 h2]. do 2 red in f. refine (f _ h1). Defined. (*Wyckoff*) Definition fun_in_set_subtract {T:Type} {U:Set} {A:Ensemble T} (f:fun_in_set A U) (a:T) : fun_in_ens (Subtract A a) U. do 2 red. intros x h1. destruct h1 as [h1 h2]. do 2 red in f. refine (f _ h1). Defined. (*Wyckoff*) Definition fun_from_add {T U:Type} {A:Ensemble T} {a:T} (f:fun_in_ens (Add A a) U) : fun_in_ens A U := fun x pf => f x (Add_intro1 _ _ _ _ pf). (*Wyckoff*) Definition fun_from_add' {T:Type} {U:Set} {A:Ensemble T} {a:T} (f:fun_in_set (Add A a) U) : fun_in_set A U := fun x pf => f x (Add_intro1 _ _ _ _ pf). (*Wyckoff*) Lemma inj_fun_from_add_compat : forall {T U:Type} {A:Ensemble T} {a:T} (f:fun_in_ens (Add A a) U), inj_in_ens f -> inj_in_ens (fun_from_add f). intros T U A a f h1. red in h1; red. intros x y h2 h3 h4. specialize (h1 x y (Add_intro1 _ _ _ _ h2) (Add_intro1 _ _ _ _ h3)). unfold fun_from_add in h4. apply h1 in h4; auto. Qed. (*Wyckoff*) Lemma inj_fun_from_add_compat' : forall {T:Type} {U:Set} {A:Ensemble T} {a:T} (f:fun_in_ens (Add A a) U), inj_in_set f -> inj_in_set (fun_from_add' f). intros T U A a f h1. red in h1; red. intros x y h2 h3 h4. specialize (h1 x y (Add_intro1 _ _ _ _ h2) (Add_intro1 _ _ _ _ h3)). unfold fun_from_add in h4. apply h1 in h4; auto. Qed. (*Wyckoff*) Lemma inh_im_in_ens : forall {T U:Type} (A:Ensemble T) (f:fun_in_ens A U), Inhabited A -> Inhabited (im_in_ens f). intros T U A f h1. destruct h1 as [a h1]. econstructor. apply (im_in_ens_intro _ _ h1). reflexivity. Qed. (*Wyckoff*) Lemma inh_im_in_set : forall {T:Type} {U:Set} (A:Ensemble T) (f:fun_in_set A U), Inhabited A -> Inhabited (im_in_set f). intros T U A f h1. destruct h1 as [a h1]. econstructor. apply (im_in_set_intro _ _ h1). reflexivity. Qed. (*Wyckoff*) Lemma im_in_ens_empty : forall {T U:Type} (f:fun_in_ens (Empty_set T) U), im_in_ens f = Empty_set _. intros T U f. apply Extensionality_Ensembles; red; split; auto with sets. red; intros x h2. destruct h2 as [x h2]. contradict h2. Qed. (*Wyckoff*) Lemma fun_in_ens_subtract_compat : forall {T U:Type} {A:Ensemble T} (f:fun_in_ens A U) (a:T), let f' := fun_in_ens_subtract f a in forall (x:T) (pf:In (Subtract A a) x), f' x pf = f x (in_sub _ _ _ pf). intros T U A f. simpl. intros x h1 h2. unfold fun_in_ens_subtract. destruct h2. f_equal. apply proof_irrelevance. Qed. (*Wyckoff*) Lemma im_in_ens_add : forall {T U:Type} {A:Ensemble T} {x:T} (f:fun_in_ens (Add A x) U), ~ In A x -> im_in_ens f = Add (im_in_ens (fun_from_add f)) (f x (Add_intro2 _ A x)). intros T U A x f hn. apply Extensionality_Ensembles; red; split; red; intros c h1. destruct h1 as [c h1]; subst. destruct h1 as [c h1 | c h1]; subst. left. econstructor. unfold fun_from_add. f_equal. apply proof_irrelevance. destruct h1. right. rewrite in_sing_iff. f_equal. apply proof_irrelevance. destruct h1 as [c h1 | c h1]. destruct h1 as [c h1]; subst. econstructor. unfold fun_from_add. reflexivity. destruct h1. econstructor. reflexivity. Grab Existential Variables. assumption. Qed. (*Wyckoff*) Lemma im_in_set_add : forall {T:Type} {U:Set} {A:Ensemble T} {x:T} (f:fun_in_set (Add A x) U), ~ In A x -> im_in_set f = Add (im_in_ens (fun_from_add f)) (f x (Add_intro2 _ A x)). intros T U A x f hn. apply Extensionality_Ensembles; red; split; red; intros c h1. destruct h1 as [c h1]; subst. destruct h1 as [c h1 | c h1]; subst. left. econstructor. unfold fun_from_add. f_equal. apply proof_irrelevance. destruct h1. right. rewrite in_sing_iff. f_equal. apply proof_irrelevance. destruct h1 as [c h1 | c h1]. destruct h1 as [c h1]; subst. econstructor. unfold fun_from_add. reflexivity. destruct h1. econstructor. reflexivity. Grab Existential Variables. assumption. Qed. (*Wyckoff*) Lemma im_in_ens_add' : forall {T U:Type} {A:Ensemble T} {x:T} (f:fun_in_ens (Add A x) U), ~ In A x -> im_in_ens f = Add (im_in_ens (fun_in_ens_subtract f x)) (f x (Add_intro2 _ A x)). intros T U A x f hn. apply Extensionality_Ensembles; red; split; red; intros c h1. destruct h1 as [c h1]; subst. destruct h1 as [c h1 | c h1]; subst. left. econstructor. rewrite fun_in_ens_subtract_compat. f_equal. apply proof_irrelevance. destruct h1. right. rewrite in_sing_iff. f_equal. apply proof_irrelevance. destruct h1 as [c h1 | c h1]. destruct h1 as [c h1]; subst. destruct h1 as [h1 h2]. pose proof h2 as h2'. rewrite in_sing_iff in h2'. destruct h1 as [c h1 | c h1]. econstructor. rewrite fun_in_ens_subtract_compat. f_equal. contradiction. destruct h1. econstructor. f_equal. Grab Existential Variables. constructor. left; auto. intro h2. destruct h2. contradiction. Qed. (*Wyckoff*) Lemma im_in_set_add' : forall {T:Type} {U:Set} {A:Ensemble T} {x:T} (f:fun_in_ens (Add A x) U), ~ In A x -> im_in_ens f = Add (im_in_ens (fun_in_ens_subtract f x)) (f x (Add_intro2 _ A x)). intros T U A x f hn. apply Extensionality_Ensembles; red; split; red; intros c h1. destruct h1 as [c h1]; subst. destruct h1 as [c h1 | c h1]; subst. left. econstructor. rewrite fun_in_ens_subtract_compat. f_equal. apply proof_irrelevance. destruct h1. right. rewrite in_sing_iff. f_equal. apply proof_irrelevance. destruct h1 as [c h1 | c h1]. destruct h1 as [c h1]; subst. destruct h1 as [h1 h2]. pose proof h2 as h2'. rewrite in_sing_iff in h2'. destruct h1 as [c h1 | c h1]. econstructor. rewrite fun_in_ens_subtract_compat. f_equal. contradiction. destruct h1. econstructor. f_equal. Grab Existential Variables. constructor. left; auto. intro h2. destruct h2. contradiction. Qed. Lemma finite_im_in_ens : forall {T U:Type} {A:Ensemble T} (f:fun_in_ens A U), Finite A -> Finite (im_in_ens f). intros T U A f h1. revert f. induction h1 as [|S h2 h3 x h4]. intro. rewrite im_in_ens_empty. constructor. intro f. pose (fun_in_ens_subtract f x) as f'. pose proof (sub_add_same_nin _ _ h4) as h5. assert (h6:Included S (Subtract (Add S x) x)). auto with sets. pose (fun (x:T) (pf:In S x) => (f' x (h6 _ pf))) as f''. specialize (h3 f''). unfold f'', f' in h3. rewrite im_in_ens_add'. apply Add_preserves_Finite. eapply Finite_downward_closed. apply h3. red; intros c h7. destruct h7 as [c h7]. clear h5. subst. econstructor; auto. rewrite fun_in_ens_subtract_compat; auto. destruct h7 as [h7 h8]. destruct h7 as [c h7 | c h7]; auto. rewrite fun_in_ens_subtract_compat. f_equal. apply proof_irrelevance. contradiction. assumption. Grab Existential Variables. destruct h7 as [h7 h8]. destruct h7 as [c h7 | c h7]; auto. contradiction. Qed. Lemma im_in_ens_const_sing : forall {T:Type} {U:Set} (A:Ensemble T) (f:fun_in_set A U), Inhabited A -> (forall (x x':T) (pf:In A x) (pf':In A x'), f x pf = f x' pf') -> exists! u:U, Singleton u = im_in_set f. intros T U A f h1 h2. destruct h1 as [a h1]. exists (f _ h1). red. split. apply Extensionality_Ensembles; red; split; red; intros x h3. destruct h3. econstructor. reflexivity. destruct h3. subst. erewrite h2. constructor. intros u h3. pose proof (In_singleton _ u) as h4. rewrite h3 in h4. destruct h4 as [u h4]. subst. apply h2; auto. Qed. Lemma im_in_ens_eq : forall {T U:Type} {A:Ensemble T} (f:fun_in_ens A U), im_in_ens f = Im (full_sig A) (fun x => f (proj1_sig x) (proj2_sig x)). intros T U A f. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]; subst. apply Im_intro with (exist _ _ h1); auto; constructor. destruct h1 as [x h1]; subst. econstructor. reflexivity. Qed. Lemma im_im_in_ens : forall {T U V:Type} {E:Ensemble T} (f:fun_in_ens E U) (g:U->V), Im (im_in_ens f) g = im_in_ens (fun x (pf:Inn E x) => g (f x pf)). intros T U V E f g. apply Extensionality_Ensembles; red; split; red; intros x h1; destruct h1 as [x h1]; subst. destruct h1 as [x h1]; subst. econstructor. reflexivity. econstructor. econstructor. reflexivity. reflexivity. Qed. Definition fun_in_ens_transfer {T U:Type} {A A':Ensemble T} (f:fun_in_ens A U) (pfe:A = A') : fun_in_ens A' U := fun x (pfi:In A' x) => let pf' := in_eq_set _ _ (eq_sym pfe) _ pfi in f x pf'. Lemma fun_in_ens_transfer_compat : forall {T U:Type} {A A':Ensemble T} (f:fun_in_ens A U) (pfe:A = A') (x:T) (pfi:Inn A' x), let pf' := in_eq_set _ _ (eq_sym pfe) _ pfi in fun_in_ens_transfer f pfe x pfi = f x pf'. unfold fun_in_ens_transfer; intros; auto. Qed. Definition fun_in_set_transfer {T:Type} {U:Set} {A A':Ensemble T} (f:fun_in_ens A U) (pfe:A = A') : fun_in_set A' U := fun x (pfi:In A' x) => let pf' := in_eq_set _ _ (eq_sym pfe) _ pfi in f x pf'. Lemma fun_in_set_transfer_compat : forall {T:Type} {U:Set} {A A':Ensemble T} (f:fun_in_ens A U) (pfe:A = A') (x:T) (pfi:Inn A' x), let pf' := in_eq_set _ _ (eq_sym pfe) _ pfi in fun_in_set_transfer f pfe x pfi = f x pf'. unfold fun_in_set_transfer; intros; auto. Qed. Lemma fun_in_ens_transfer_eq_refl : forall {T U:Type} {A:Ensemble T} (f:fun_in_ens A U), fun_in_ens_transfer f eq_refl = f. intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; unfold fun_in_ens_transfer; f_equal; apply proof_irrelevance. Qed. Lemma fun_in_set_transfer_eq_refl : forall {T:Type} {U:Set} {A:Ensemble T} (f:fun_in_ens A U), fun_in_ens_transfer f eq_refl = f. intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; unfold fun_in_ens_transfer; f_equal; apply proof_irrelevance. Qed. Lemma im_in_ens_transfer : forall {T U:Type} {A B:Ensemble T} (f:fun_in_ens A U) (pfe:A = B), im_in_ens f = im_in_ens (fun_in_ens_transfer f pfe). intros; subst. rewrite fun_in_ens_transfer_eq_refl; auto. Qed. Lemma linv_in_ens_ex : forall {T U:Type} {A:Ensemble T} (f:fun_in_ens A U) (pfi:inj_in_ens f), exists! f':fun_in_ens (im_in_ens f) T, forall (x:T) (pf:In A x), f' (f x pf) (im_in_ens_intro _ _ pf _ eq_refl) = x. intros T U A f h1. assert (h2:forall y:U, Inn (im_in_ens f) y -> exists! x:T, exists pf:Inn A x, f x pf = y). intros y h2. destruct h2; subst. exists x. red. split. exists pf; auto. intros x' h2. destruct h2 as [h2 h3]. apply h1 in h3; auto. exists (fun y pf => proj1_sig (constructive_definite_description _ (h2 y pf))). red. split. intros x h3. destruct constructive_definite_description as [x' h4]. destruct h4 as [h4 h5]. pose proof (h1 _ _ _ _ h5); subst. simpl; auto. intros f' h3. apply functional_extensionality_dep. intro y. apply functional_extensionality. intro h4. destruct constructive_definite_description as [z' h5]. destruct h5 as [h5]; subst. simpl. symmetry. specialize (h3 z' h5). rewrite <- h3 at 2. f_equal. apply proof_irrelevance. Qed. Definition linv_in_ens {T U:Type} {A:Ensemble T} (f:fun_in_ens A U) (pfi:inj_in_ens f) : fun_in_ens (im_in_ens f) T := proj1_sig (constructive_definite_description _ (linv_in_ens_ex f pfi)). Lemma linv_in_ens_compat : forall {T U:Type} {A:Ensemble T} (f:fun_in_ens A U) (pfi:inj_in_ens f), let f' := linv_in_ens f pfi in forall (x:T) (pf:In A x), f' (f x pf) (im_in_ens_intro _ _ pf _ eq_refl) = x. unfold linv_in_ens; intros; destruct constructive_definite_description; auto. Qed. Lemma in_linv_in_ens : forall {T U:Type} {A:Ensemble T} (f:fun_in_ens A U) (pfi:inj_in_ens f) (y:U) (pf:Inn (im_in_ens f) y), Inn A (linv_in_ens f pfi y pf). intros T U A f h1 y h2. destruct h2; subst. rewrite linv_in_ens_compat; auto. Qed. Lemma bij_impl_inv_in_ens : forall {T U} {A:Ensemble T} {B:Ensemble U} (f:fun_in_ens' A B), bij_in_ens f -> invertible_in_ens f. intros T U A B f h1. destruct h1 as [h2 h3]. pose proof (surj_in_ens_compat _ h3) as h4. pose proof (linv_in_ens_compat f h2) as h5; simpl in h5. assert (h6:forall y (pfy:Inn B y), Inn (im_in_ens f) (exist _ y pfy)). intros y hy. pose proof hy as hy'. rewrite <- h4 in hy'. destruct hy' as [y h6]; clear h4; subst. destruct h6 as [y h6]; subst. apply im_in_ens_intro with y h6. apply proj1_sig_injective; auto. pose (fun y (pfy:Inn B y) => linv_in_ens f h2 (exist _ y pfy) (h6 y pfy)) as g. assert (h7:forall y (pfy:Inn B y), Inn A (g y pfy)). unfold g; intros; apply in_linv_in_ens. pose (fun y (pfy:Inn B y) => exist (fun c => Inn A c) (g y pfy) (h7 y pfy)) as g'. apply (invertible_in_ens_intro f g'). unfold g', g. constructor. intros x hy; apply proj1_sig_injective; simpl. specialize (h5 x hy). rewrite <- h5. apply f_equal_dep. apply proj1_sig_injective; auto. simpl. intros y hy. pose proof hy as hy'. rewrite <- h4 in hy'. destruct hy' as [y h8]; clear h4; subst. destruct h8 as [y h8]; subst. specialize (h5 y h8). apply proj1_sig_injective; simpl. f_equal. apply f_equal_dep. rewrite <- h5. apply f_equal_dep. apply proj1_sig_injective; simpl; repeat f_equal. Qed. Definition fun_from_add_im {T U V:Type} {A:Ensemble T} {a:T} {g:fun_in_ens (Add A a) U} (f:fun_in_ens (im_in_ens g) V) (pfi:inj_in_ens g) : fun_in_ens (im_in_ens (fun_from_add g)) V := fun x pf => let pfi' := inj_fun_from_add_compat _ pfi in f (g (linv_in_ens (fun_from_add g) pfi' x pf) (Add_intro1 _ _ _ _ (in_linv_in_ens (fun_from_add g) pfi' _ _))) (im_in_ens_intro g _ (Add_intro1 _ _ _ _ (in_linv_in_ens (fun_from_add g) pfi' _ _)) _ eq_refl). Lemma fun_from_add_im_compat : forall {T U V:Type} {A:Ensemble T} {a:T} {g:fun_in_ens (Add A a) U} (f:fun_in_ens (im_in_ens g) V) (pfi:inj_in_ens g) (x:T) (pf:Inn A x), fun_from_add_im f pfi (fun_from_add g x pf) (im_in_ens_intro (fun_from_add g) _ pf _ eq_refl) = f (g x (Add_intro1 _ _ _ _ pf)) (im_in_ens_intro g _ (Add_intro1 _ _ _ _ pf) _ eq_refl). unfold fun_from_add_im; intros. apply f_equal_dep. apply f_equal_dep. rewrite linv_in_ens_compat; auto. Qed. Lemma fun_from_add_im_functional : forall {T U V:Type} {A:Ensemble T} {a:T} {g:fun_in_ens (Add A a) U} (f:fun_in_ens (im_in_ens g) V) (pfi:inj_in_ens g) (y y':U) (pfeq:y = y') (pf: Inn (im_in_ens (fun_from_add g)) y) (pf':Inn (im_in_ens (fun_from_add g)) y'), fun_from_add_im f pfi y pf = fun_from_add_im f pfi y' pf'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma fun_in_set_inh_const : forall {T:Type} {U:Set} (A:Ensemble T) (f:fun_in_set A U) (pf:forall u u':U, u = u'), Inhabited A -> U. intros T U A f h0 h1. pose proof (im_in_ens_const_sing _ f h1) as h2. hfirst h2. intros x x' h3 h4. apply h0. specialize (h2 hf). refine (proj1_sig (constructive_definite_description _ h2)). Qed. (*Wyckoff*) Lemma finite_union_finite_left : forall {T:Type} (A B:Ensemble T), Finite (Union A B) -> Finite A. intros T A B h1. eapply Finite_downward_closed. apply h1. auto with sets. Qed. (*Wyckoff*) Lemma finite_union_finite_right : forall {T:Type} (A B:Ensemble T), Finite (Union A B) -> Finite B. intros T A B h1. eapply Finite_downward_closed. apply h1. auto with sets. Qed. (*Wyckoff*) Lemma finite_couple : forall {T:Type} (x y:T), Finite (Couple x y). intros T x y. destruct (eq_dec x y) as [he | hne]. subst. rewrite couple_singleton. apply Singleton_is_finite. rewrite couple_add_sing. apply Add_preserves_Finite. apply Singleton_is_finite. Qed. (*Wyckoff*) Lemma incl_subtract : forall {T:Type} (P:Ensemble T) (x:T), Included (Subtract P x) P. intros T P x. red. intros x' h1. destruct h1. assumption. Qed. (*Wyckoff*) Lemma subtract_preserves_finite : forall {T:Type} (P:Ensemble T) (x:T), Finite P -> Finite (Subtract P x). intros T P x h1. assert (h2:Included (Subtract P x) P). apply incl_subtract. apply Finite_downward_closed with P; auto. Qed. (*Wyckoff*) Lemma subtract_finite : forall {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (x:T), Finite (Subtract A x) -> Finite A. intros T h0 A x h1. pose proof (Add_preserves_Finite _ _ x h1) as h2. eapply Finite_downward_closed. apply h2. rewrite add_sub_eq; auto. auto with sets. Qed. (*Wyckoff*) Lemma couple_comm : forall {T:Type} (a b:T), Couple a b = Couple b a. intros T a b. apply Extensionality_Ensembles. red. split. (* <= *) red. intros ? h1. destruct h1; subst; constructor. (* >= *) red. intros ? h1. destruct h1; subst; constructor. Qed. (*Wyckoff*) Lemma bool_couple : Full_set bool = Couple true false. apply Extensionality_Ensembles. red. split. red. intros x ?. destruct x. left. right. red. intros; constructor. Qed. (*Wyckoff*) Lemma finite_bool : Finite (Full_set bool). rewrite bool_couple. apply finite_couple. Qed. (*Wyckoff*) Lemma finite_inh_or_empty : forall {T:Type} (A:Ensemble T), Finite A -> Inhabited A \/ A = Empty_set _. intros T A h1. induction h1 as [|A h1 h2 a h3]. right; auto. left. apply Inhabited_intro with a. right. constructor. Qed. (*Wyckoff*) (*uses LEM*) Lemma infinite_inh : forall {T:Type} (A:Ensemble T), ~ Finite A -> Inhabited A. intros T A h1. assert (h2:A <> Empty_set _). intro; subst. contradict h1. constructor. apply not_empty_Inhabited; auto. Qed. (*This is the finite anlaogue of the clasical Standard Library function, [not_empty_Inhabited] . . . but of course without LEM*) Lemma fin_not_empty_inh : forall {T:Type} (A:Ensemble T), Finite A -> A <> Empty_set _ -> Inhabited A. intros T A h1 h2. apply finite_inh_or_empty in h1. destruct h1; subst; auto. contradict h2; auto. Qed. (*Wyckoff*) Lemma fin_in_or : forall {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (x:T), Finite A -> In A x \/ ~In A x. intros T h0 A x h1. revert x. induction h1 as [|A h1 h2 a h3]. intros; right; intro; contradiction. intro x. specialize (h2 x). destruct h2 as [h2 | h2]. left. left; auto. destruct (h0 a x) as [h4 | h4]; simpl; subst. left; right; constructor; auto. right. intro h5. destruct h5 as [x h5 | x h5]; try contradiction; subst. destruct h5; contradict h4; auto. Qed. (*Wyckoff*) Lemma finite_cardinal_unq : forall {T:Type} (S:Ensemble T), Finite S -> exists! n:nat, cardinal _ S n. intros T S h1. pose proof (finite_cardinal _ _ h1) as h2. destruct h2 as [n h2]. exists n. red. split; try assumption. apply cardinal_unicity; assumption. Qed. (*Wyckoff*) Definition card_fun {T:Type} (S:Ensemble T) (pf:Finite S) : nat := proj1_sig (constructive_definite_description _ (finite_cardinal_unq S pf)). (*Wyckoff*) Lemma card_fun_compat : forall {T:Type} (S:Ensemble T) (pf:Finite S), cardinal _ S (card_fun _ pf). intros T S h1. unfold card_fun. destruct constructive_definite_description. simpl. assumption. Qed. (*Wyckoff*) Lemma card_fun_empty : forall T:Type, card_fun (Empty_set T) (Empty_is_finite T) = 0. intro T. pose proof (card_fun_compat _ (Empty_is_finite T)) as h1. pose proof (card_empty T) as h2. eapply cardinal_is_functional; auto. apply h1. assumption. Qed. (*Wyckoff*) (*Uses LEM.*) Lemma card_fun1_ex : forall {T:Type} (S:Ensemble T), exists! n:nat, (Finite S -> cardinal _ S n) /\ (~Finite S -> n = 0). intros T S. destruct (classic (Finite S)) as [h1 | h2]. exists (card_fun _ h1). red. split. split. intro h2. apply card_fun_compat. intro; contradiction. intros n h2. destruct h2 as [h2l h2r]. specialize (h2l h1). pose proof (card_fun_compat _ h1) as h3. pose proof cardinal_is_functional. pose proof (cardinal_is_functional _ _ _ h2l _ _ h3 (eq_refl _)). subst. reflexivity. exists 0. red. split. split. intros; contradiction. auto. intros n h3. destruct h3 as [h3l h3r]. symmetry. apply h3r; auto. Qed. (*Wyckoff*) Definition card_fun1 {T:Type} (S:Ensemble T) := proj1_sig (constructive_definite_description _ (card_fun1_ex S)). (*Wyckoff*) Lemma card_fun1_compat : forall {T:Type} (S:Ensemble T), let n:= card_fun1 S in (Finite S -> cardinal _ S n) /\ (~Finite S -> n = 0). intros T S. unfold card_fun1. destruct constructive_definite_description. simpl. assumption. Qed. (*Wyckoff*) Lemma card_fun1_empty : forall (T:Type), card_fun1 (Empty_set T) = 0. intro T. pose proof (card_fun1_compat (Empty_set T)) as h1. destruct h1 as [h1l h1r]. specialize (h1l (Empty_is_finite T)). pose proof (card_empty T) as h2. eapply cardinal_is_functional; auto. apply h1l. apply h2. Qed. (*Wyckoff*) (*Uses LEM.*) Lemma card_fun1_O : forall {T:Type} (A:Ensemble T), card_fun1 A = 0 -> A = Empty_set _ \/ ~Finite A. intros T A h1. pose proof (card_fun1_compat A) as h2. destruct h2 as [h2l h2r]. destruct (classic (Finite A)) as [h3 | h4]. apply h2l in h3. rewrite h1 in h3. apply cardinalO_empty in h3. left. assumption. right. assumption. Qed. (*Wyckoff*) Lemma card_fun1_sing : forall {T:Type} (x:T), card_fun1 (Singleton x) = 1. intros T x. pose proof (card_fun1_compat (Singleton x)) as h1. destruct h1 as [h1l h1r]. specialize (h1l (Singleton_is_finite _ x)). pose proof (card_sing x) as h2. eapply cardinal_is_functional; auto. apply h1l. apply h2. Qed. (*Wyckoff*) Lemma lt_card_fun1_finite : forall {T:Type} (A:Ensemble T) (n:nat), n < card_fun1 A -> Finite A. intros T A n h1. pose proof (card_fun1_compat A) as h2. destruct h2 as [h2 h3]. apply NNPP. intro h4. apply h3 in h4. rewrite h4 in h1. omega. Qed. (*Wyckoff*) Lemma card_fun1_1 : forall {T:Type} (A:Ensemble T), card_fun1 A = 1 -> exists x:T, A = Singleton x. intros T A h1. assert (h0:0 < card_fun1 A). omega. apply lt_card_fun1_finite in h0. revert h1. induction h0 as [|A h2 h3 x h4]. intro h1. rewrite card_fun1_empty in h1. omega. exists x. assert (h5:A = Empty_set _). apply NNPP. intro h5. apply not_empty_Inhabited in h5. destruct h5 as [y h5]. assert (h6:x <> y). intro h6. subst. contradiction. pose proof (card_fun1_compat (Add A x)) as h7. destruct h7 as [h7l h7r]. clear h7r. pose proof (Add_preserves_Finite _ _ x h2) as h8. apply h7l in h8. rewrite h1 in h8. pose proof card_soustr_1. pose proof (card_soustr_1 _ _ _ h8 _ (Add_intro2 _ A x)) as h9. simpl in h9. rewrite sub_add_same_nin in h9. apply cardinalO_empty in h9. subst. contradiction. assumption. subst. apply add_empty_sing; auto. Qed. (*Wyckoff*) Lemma card_fun1_ge_1 : forall {T:Type} (A:Ensemble T), Inhabited A -> Finite A -> 1 <= card_fun1 A. intros T A h1 h0. destruct (le_lt_dec 1 (card_fun1 A)) as [|h2]; auto. assert (h3:card_fun1 A = 0). omega. contradict h1. apply card_fun1_O in h3. destruct h3 as [h3|]; try contradiction. subst. intro h3. destruct h3; contradiction. Qed. (*Wyckoff*) Lemma subtract_eq_empty_card : forall {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (x:T), Subtract A x = Empty_set _ -> card_fun1 A <= 1. intros T h0 A x h1. pose proof (Empty_is_finite T) as h2. rewrite <- h1 in h2. apply subtract_finite in h2; auto. apply finite_inh_or_empty in h2. destruct h2 as [h2 | h2]. pose proof (inh_subtract_empty h0 _ _ h2 h1) as h3. pose proof (f_equal (fun S => Add S x) h1) as h4. simpl in h4. rewrite add_sub_eq in h4. rewrite in_add_eq in h4. rewrite add_empty_sing in h4. subst. rewrite card_fun1_sing; auto with arith. assumption. assumption. subst. rewrite card_fun1_empty; auto with arith. Qed. (*Wyckoff*) Lemma card_fun_card_fun1_compat : forall {T:Type} (A:Ensemble T) (pf:Finite A), card_fun A pf = card_fun1 A. intros T A h1. pose proof (card_fun_compat A h1) as h2. pose proof (card_fun1_compat A) as h3. destruct h3 as [h3l h3r]. specialize (h3l h1). eapply cardinal_is_functional; auto. apply h2. apply h3l. Qed. (*Wyckoff*) Lemma card_add_nin : forall {T:Type} (A:Ensemble T) (n:nat), cardinal _ A n -> forall x:T, ~ In A x -> cardinal _ (Add A x) (S n). intros T A n h1 x h2. constructor; auto. Qed. (*Wyckoff*) Lemma card_add_nin' : forall {T:Type} (A:Ensemble T), Finite A -> forall x:T, ~ In A x -> card_fun1 (Add A x) = S (card_fun1 A). intros T A h1 x h2. pose proof (card_fun1_compat A) as h3. destruct h3 as [h3l h3r]. specialize (h3l h1). pose proof (card_fun1_compat (Add A x)) as h4. destruct h4 as [h4l h4r]. specialize (h4l (Add_preserves_Finite _ _ _ h1)). eapply card_add_nin in h3l. eapply cardinal_is_functional. apply h4l. apply h3l. apply eq_refl. apply h2. Qed. (*Wyckoff*) Lemma incl_card_fun1 : forall {T:Type} (A B:Ensemble T), Finite B -> Included A B -> card_fun1 A <= card_fun1 B. intros T A B h4 h1. assert (h2:Finite A). eapply Finite_downward_closed; auto. apply h4; auto. assumption. pose proof (finite_cardinal _ _ h2) as h5. pose proof (finite_cardinal _ _ h4) as h6. destruct h5 as [m h5]. destruct h6 as [n h6]. pose proof (incl_card_le _ _ _ _ _ h5 h6 h1) as h7. pose proof (card_fun1_compat A) as [h8]. pose proof (card_fun1_compat B) as [h9]. specialize (h8 h2). specialize (h9 h4). pose proof (cardinal_is_functional _ _ _ h5 _ _ h8 (eq_refl _)). subst. pose proof (cardinal_is_functional _ _ _ h9 _ _ h6 (eq_refl _)). subst. assumption. Qed. (*Wyckoff*) Lemma finite_inh_or_empty_dec : forall {T:Type} (A:Ensemble T), Finite A -> {Inhabited A} + {A = Empty_set _}. intros T A h1. pose proof (finite_cardinal_unq A h1) as h2. pose proof (proj2_sig (constructive_definite_description _ h2)) as h3. simpl in h3. let P := type of h3 in match P with cardinal _ _ ?n => destruct n as [|m] end. right. apply cardinalO_empty; auto. left. inversion h3. subst. econstructor. right. constructor. Qed. (*Wyckoff*) Lemma singleton_inc : forall {T:Type} (a:T) (A:Ensemble T), Included A (Singleton a) -> A = Empty_set _ \/ A = Singleton a. intros T a A h0. pose proof (Singleton_is_finite _ a) as h6. assert (h7:cardinal _ (Singleton a) 1). assert (h7:Singleton a = Add (Empty_set _) a). unfold Add. rewrite empty_union. reflexivity. rewrite h7. constructor. constructor. intro. contradiction. pose proof (Finite_downward_closed _ _ h6 _ h0) as h8. pose proof (finite_cardinal _ _ h8) as h9. destruct h9 as [n h9]. pose proof (incl_card_le _ _ _ _ _ h9 h7 h0) as h10. pose proof (le_pred n 1 h10) as h11. simpl in h11. pose proof (le_n_0_eq _ h11) as h12. unfold pred in h12. destruct (zerop n) as [? | h13]. subst. left. apply cardinalO_empty. assumption. assert (h14: n = S O). auto with arith. subst. right. apply Extensionality_Ensembles. red. split. assumption. red. intros x h14. destruct h9. pose proof (lt_irrefl 0). contradiction. assert (h15:a = x0). unfold Included in h0. pose proof (Add_intro2 _ A x0) as h15. specialize (h0 _ h15). destruct h0; auto. subst. destruct h14; subst. apply Add_intro2. Qed. (*Wyckoff*) Lemma sing_eq_add : forall {T:Type} (A:Ensemble T) (x a:T), Singleton x = Add A a -> a = x /\ (A = Empty_set _ \/ A = Singleton x). intros T A x a h1. split. pose proof (Add_intro2 _ A a) as h2. rewrite <- h1 in h2. destruct h2; auto. assert (h2:A <> Empty_set _ -> A = Singleton x). intro h2. apply Extensionality_Ensembles; red; split; red; intros c h3. pose proof (Add_intro1 _ A a _ h3) as h4. rewrite h1. assumption. destruct h3. pose proof (Add_intro2 _ A a) as h3. rewrite <- h1 in h3. destruct h3. apply not_empty_Inhabited in h2. destruct h2 as [c h2]. pose proof (Add_intro1 _ A x _ h2) as h3. rewrite <- h1 in h3. destruct h3. assumption. tauto. Qed. (*Wyckoff*) Lemma singleton_inc_int : forall {T:Type} (a:T) (A:Ensemble T), Intersection A (Singleton a) = Empty_set _ \/ Intersection A (Singleton a) = Singleton a. intros T a A. assert (h1:Included (Intersection A (Singleton a)) (Singleton a)); auto with sets. apply singleton_inc; auto. Qed. (*Wyckoff*) Lemma card_sub_in : forall {T:Type} (A:Ensemble T), Finite A -> forall x:T, In A x -> card_fun1 A = S (card_fun1 (Subtract A x)). intros T A h1 x h2. pose proof (card_fun1_compat A) as h3. destruct h3 as [h3l h3r]. clear h3r. specialize (h3l h1). pose proof (card_fun1_compat (Subtract A x)) as h4. destruct h4 as [h4l h4r]. clear h4r. pose proof (subtract_preserves_finite _ x h1) as h5. specialize (h4l h5). rewrite <- (add_sub_compat_in A x) at 1; auto. rewrite card_add_nin'. f_equal. assumption. intro h6. destruct h6 as [h6 h7]. contradict h7. constructor. Qed. (*Wyckoff*) Lemma card_fun1_subtract : forall {T:Type} (A:Ensemble T), Finite A -> forall x:T, In A x -> card_fun1 (Subtract A x) = pred (card_fun1 A). intros T A h1 x h2. pose proof (card_sub_in A h1 x h2). omega. Qed. (*Wyckoff*) Lemma card_fun1_eq_sub : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), Finite A -> Finite B -> card_fun1 A = card_fun1 B -> forall (x:T) (y:U), In A x -> In B y -> card_fun1 (Subtract A x) = card_fun1 (Subtract B y). intros T U A B hfa hfb h1 x y h2 h3. rewrite (card_sub_in A hfa x h2) in h1. rewrite (card_sub_in B hfb y h3) in h1. apply S_inj; auto. Qed. (*Wyckoff*) Lemma cardinal_couple : forall {T:Type} (x y:T), x <> y -> cardinal _ (Couple x y) 2. intros T x y h1. rewrite couple_add_sing. constructor. apply card_sing. intro h2. destruct h2; auto. Qed. (*Wyckoff*) Lemma inclusion_intersection : forall {T:Type} (P S R : Ensemble T), (Included P S /\ Included P R) <-> Included P (Intersection S R). intros T P S R. split. (*left*) intro h1. destruct h1 as [h2 h3]. red. intros x h4. constructor; auto with sets. (*right*) intro h1. split; red. (*left*) intros x h5. red in h1. pose proof (h1 x h5) as h6. inversion h6; trivial. (*right*) intros x h7. pose proof (h1 x h7) as h6. inversion h6; trivial. Qed. (*Wyckoff*) Lemma intersection_preserves_inclusion : forall {T:Type} (P Q R:Ensemble T), Included P Q -> Included (Intersection R P) (Intersection R Q). intros T P Q R h1. red. intros x h2. destruct h2 as [x h2l h2r]. auto with sets. Qed. (*Wyckoff*) Lemma union_preserves_inclusion : forall {T:Type} (P Q R:Ensemble T), Included P Q -> Included (Union R P) (Union R Q). intros T P Q R h1. red. intros x h2. destruct h2 as [x h2 | x h3]. left; auto. apply h1 in h3. right; auto. Qed. (*Wyckoff*) Lemma inclusion_mono : forall {T:Type} (A B C D:Ensemble T), Included A B -> Included C D -> Included (Intersection A C) (Intersection B D). intros T A B C D h1 h2. red. intros x h3. destruct h3 as [x h3 h4]. apply h1 in h3. apply h2 in h4. constructor; auto. Qed. (*Wyckoff*) Lemma in_intersection_iff : forall {T:Type} (A B:Ensemble T) (x:T), In (Intersection A B) x <-> In A x /\ In B x. intros T A B x. split. intros h1. destruct h1. split; auto. intros h1. destruct h1. constructor; auto. Qed. (*Wyckoff*) Lemma in_union_iff : forall {T:Type} (A B:Ensemble T) (x:T), In (Union A B) x <-> In A x \/ In B x. intros T A B x. split. intro h1. destruct h1. left; auto. right; auto. intro h1. destruct h1. left; auto. right; auto. Qed. (*Wyckoff*) Lemma in_match_sumbool_excl_middle : forall {T U:Type} {V:U->Type} (P:forall (u:U) (v:V u), Prop) (A:forall (u:U) (v:V u) (pf:P u v), Ensemble T) (B:forall (u:U) (v:V u) (pf:~ P u v), Ensemble T) (u:U) (v:V u) (pfd:{P u v} + {~P u v}) (x:T), Inn match pfd with | left pfa => A u v pfa | right pfb => B u v pfb end x <-> (exists pfp:P u v, Inn (A u v pfp) x) \/ (exists pfnp: ~P u v, Inn (B u v pfnp) x). intros T U V P A B u v h1 x. destruct h1 as [h1 | h1]. split. intro h2. left. exists h1; auto. intro h2. destruct h2 as [h2 | h2]. destruct h2 as [h2 h3]. assert (h1 = h2). apply proof_irrelevance. subst; auto. destruct h2; contradiction. split. intro h2. right. exists h1; auto. intro h2. destruct h2 as [h2 | h2]. destruct h2; contradiction. destruct h2 as [h2 h3]. assert (h1 = h2). apply proof_irrelevance. subst; auto. Qed. (*Wyckoff*) Lemma intersection_preserves_finite : forall {T:Type} (P Q:Ensemble T), Finite P -> Finite (Intersection P Q). intros T P Q h1. assert (h2:Included (Intersection P Q) P). auto with sets. apply Finite_downward_closed with P; auto. Qed. (*Wyckoff*) Lemma add_preserves_infinite : forall {T:Type} (A:Ensemble T) (a:T), ~Finite A -> ~Finite (Add A a). intros T A a h1. intro h2. assert (h5:Included A (Add A a)). auto with sets. pose proof (Finite_downward_closed _ _ h2 _ h5). contradiction. Qed. (*Wyckoff*) Lemma in_sub_inj : forall {T:Type} (E F:Ensemble T) (x:T), In E x -> In F x -> Subtract E x = Subtract F x -> E = F. intros T E F x h1 h2 h3. apply Extensionality_Ensembles; red; split; red; intros a h5. destruct (eq_dec a x) as [h6 | h6] ;subst; auto. assert (h7:In (Subtract E x) a). constructor; auto. intro h7. destruct h7; auto. rewrite h3 in h7. destruct h7; auto. destruct (eq_dec a x) as [h6 | h6] ;subst; auto. assert (h7:In (Subtract F x) a). constructor; auto. intro h7. destruct h7; auto. rewrite <- h3 in h7. destruct h7; auto. Qed. (*Wyckoff*) Lemma inclusion_iff_intersection_eq : forall {T:Type} (P Q:Ensemble T), Included P Q <-> Intersection P Q = P. intros T P Q. split. (* -> *) intro. apply Extensionality_Ensembles. red. split; auto with sets. (* <= *) intro h1. red. intros x h2. rewrite <- h1 in h2. destruct h2; assumption. Qed. (*Wyckoff*) Lemma inclusion_iff_union : forall {T:Type} (P Q:Ensemble T), Included P Q <-> Union P Q = Q. intros T P Q. split. intro h1. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2; auto with sets. red. intros; right; auto. intro h1. red. intros. rewrite <- h1. auto with sets. Qed. (*Wyckoff*) Lemma incl_empty : forall {T:Type} (S:Ensemble T), Included (Empty_set _) S. intros T S. red. intros; contradiction. Qed. (*Wyckoff*) Lemma eq_empty : forall {T:Type} (A:Ensemble T), Included A (Empty_set _) -> A = Empty_set _. intros T A h1. pose proof (Finite_downward_closed _ _ (Empty_is_finite T) _ h1) as h2. revert h1 . induction h2; auto. intro h3. pose proof (Add_intro2 _ A x) as h4. apply h3 in h4. contradiction. Qed. (*Wyckoff*) Lemma sub_empty : forall {T:Type} (a:T), Subtract (Empty_set _) a = Empty_set _. intros T a. apply eq_empty. red; intros x h1. destruct h1; contradiction. Qed. (*Wyckoff*) Definition empty_sig_incl {T:Type} (A:Ensemble T) := exist (fun C => Included C A) (Empty_set _) (incl_empty A). (*Wyckoff*) Lemma full_inclusion : forall {T:Type} (S:Ensemble T), Included S (Full_set _). intros; red; intros; constructor. Qed. (*Wyckoff*) Lemma inclusion_reflexive : forall {T:Type} (S:Ensemble T), Included S S. auto with sets. Qed. (*Wyckoff*) Lemma inclusion_transitive : forall T:Type, Transitive (@Included T). intro T. red. intros A B C h1 h2. red. intros x h3. apply h1 in h3. apply h2 in h3. assumption. Qed. (*Wyckoff*) Lemma intersection_full_set : forall {T:Type} (S:Ensemble T), Intersection S (Full_set T) = S. intros T S. apply Extensionality_Ensembles. red. split. (* <= *) red. intros ? h. destruct h; assumption. (* >= *) red. intros x ?. constructor. assumption. constructor. Qed. (*Wyckoff*) Lemma included2_impl_union_included : forall {T:Type} (A B C:Ensemble T), ((Included A C) /\ (Included B C)) <-> Included (Union A B) C. intros T A B C. split. (* -> *) unfold Included. intro h1. destruct h1 as [h1 h2]. intros x h3. apply Union_inv in h3. destruct h3; auto with sets. (* <- *) intro h1. split. assert (Included A (Union A B)). auto with sets. auto with sets. assert (Included B (Union A B)). auto with sets. auto with sets. Qed. (*Wyckoff*) Lemma feq_im_aux : forall {T U:Type} (A B:Ensemble T) (f:T->U), A = B -> Included (Im A f) (Im B f). intros T U A B f h1. red. intros y h2. destruct h2; subst. apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma feq_im : forall {T U:Type} (A B:Ensemble T) (f:T->U), A = B -> Im A f = Im B f. intros T U A B f h1. apply Extensionality_Ensembles; red; split; apply feq_im_aux; auto. Qed. (*Wyckoff*) Lemma im_inj_inj : forall {T U:Type} (A B:Ensemble T) (f:T->U), Injective f -> Im A f = Im B f -> A = B. intros T U A B f h1 h2. apply Extensionality_Ensembles. red. split. red. intros x h3. assert (h4:In (Im A f) (f x)). apply Im_intro with x; auto. rewrite h2 in h4. inversion h4 as [y h5 ha h6]. subst. specialize (h1 _ _ h6). subst. assumption. red. intros x h3. assert (h4:In (Im B f) (f x)). apply Im_intro with x; auto. rewrite <- h2 in h4. inversion h4 as [y h5 ha h6]. subst. specialize (h1 _ _ h6). subst. assumption. Qed. (*Wyckoff*) Lemma im_union : forall {T U:Type} (A B:Ensemble T) (f:T->U), Im (Union A B) f = Union (Im A f) (Im B f). intros T U A B f. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h1. destruct h1 as [x h1 y h3]. subst. destruct h1 as [x h1a | x h1b]. left. apply Im_intro with x; auto. right. apply Im_intro with x; auto. (* >= *) red. intros x h1. destruct h1 as [x h1a | x h1b]. destruct h1a as [x h1a y h2]. subst. apply Im_intro with x. left; auto. reflexivity. destruct h1b as [x h1b y h2]. subst. apply Im_intro with x; try auto; right; auto. Qed. (*Wyckoff*) Lemma im_intersection_incl : forall {T U:Type} (A B:Ensemble T) (f:T->U), Included (Im (Intersection A B) f) (Intersection (Im A f) (Im B f)). intros T U A B f. red. intros b h1. destruct h1 as [x h1]. subst. destruct h1 as [x h1l h1r]. constructor. apply Im_intro with x; auto. apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma im_intersection_inj : forall {T U:Type} (A B:Ensemble T) (f:T->U), injective f -> Im (Intersection A B) f = Intersection (Im A f) (Im B f). intros T U A B f h0. apply Extensionality_Ensembles. red. split. apply im_intersection_incl. red. intros b h1. destruct h1 as [b h1 h2]. destruct h1 as [b h1]. subst. apply Im_intro with b. inversion h2 as [a h3 c h4]. subst. clear h2. red in h0. apply h0 in h4. subst. constructor; auto. reflexivity. Qed. (*Wyckoff*) Lemma foil_union : forall {T:Type} (A B C D:Ensemble T), Intersection (Union A B) (Union C D) = Union (Intersection A C) (Union (Intersection A D) (Union (Intersection B C) (Intersection B D))). intros T A B C D. rewrite Distributivity. rewrite Intersection_commutative at 1. rewrite (Intersection_commutative _ (Union A B) D). do 2 rewrite Distributivity. rewrite (Intersection_commutative _ C A). rewrite (Intersection_commutative _ C B). rewrite (Intersection_commutative _ D A). rewrite (Intersection_commutative _ D B). rewrite Union_associative. f_equal. rewrite <- Union_associative. rewrite <- Union_associative. f_equal. rewrite Union_commutative; auto. Qed. (*Wyckoff*) Lemma foil_intersection : forall {T:Type} (A B C D:Ensemble T), Union (Intersection A B) (Intersection C D) = Intersection (Union A C) (Intersection (Union A D) (Intersection (Union B C) (Union B D))). intros T A B C D. rewrite Distributivity'. rewrite Union_commutative at 1. rewrite (Union_commutative _ (Intersection A B) D). do 2 rewrite Distributivity'. rewrite (Union_commutative _ C A). rewrite (Union_commutative _ C B). rewrite (Union_commutative _ D A). rewrite (Union_commutative _ D B). rewrite Intersection_associative. f_equal. rewrite <- Intersection_associative. rewrite <- Intersection_associative. f_equal. rewrite Intersection_commutative; auto. Qed. (*Wyckoff*) Lemma surj_iff : forall {T U:Type} (f:T->U), surjective f <-> Im (Full_set T) f = Full_set U. intros T U f. split. (* -> *) intro h1. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x ?. constructor. (* >= *) red. intros u ?. red in h1. pose proof (h1 u) as h2. destruct h2 as [t]. apply Im_intro with t. constructor. symmetry; assumption. (* <- *) intro h1. red. intro u. pose proof (Full_intro _ u) as h2. rewrite <- h1 in h2. inversion h2 as [t]. exists t. symmetry. assumption. Qed. (*Section Abstr. -- toplevel*) (*Schepler*) Inductive characteristic_function_abstraction {X:Type} (P:X->Prop) (x:X) : Prop := | intro_characteristic_sat: P x -> In (characteristic_function_abstraction P) x. (*Schepler*) Definition characteristic_function_to_ensemble {X:Type} (P:X->Prop) : Ensemble X := characteristic_function_abstraction P. (*Schepler*) Notation "[ x : X | P ]" := (characteristic_function_to_ensemble (fun x:X => P)) (x ident). Section Abstr. (*Schepler*) Lemma characteristic_function_to_ensemble_is_identity: forall {X:Type} (P:X->Prop), [ x:X | P x ] = P. Proof. intros. apply Extensionality_Ensembles; split; red; intros. destruct H. exact H. constructor. exact H. Qed. (*Wyckoff*) Lemma sat_in : forall {T:Type} (P:T->Prop) (x:T), In [t:T | P t] x -> P x. intros ? ? ? h1; destruct h1; auto. Qed. (*Wyckoff*) Lemma sat_in' : forall {T:Type} (P:T->Prop) (x:T), In [t:T | ~ P t] x -> ~ P x. intros ? ? ? h1; destruct h1; auto. Qed. Lemma sat_in1 : forall {T:Type} (P Q:T->Prop) (x:T), In [t:T | P t /\ Q t] x -> P x. intros ? ? ? ? h1; destruct h1; tauto. Qed. Lemma sat_in1' : forall {T:Type} (P Q:T->Prop) (x:T), In [t:T | ~ P t /\ Q t] x -> ~ P x. intros ? ? ? ? h1; destruct h1; tauto. Qed. Lemma sat_in2 : forall {T:Type} (P Q:T->Prop) (x:T), In [t:T | P t /\ Q t] x -> Q x. intros ? ? ? ? h1; destruct h1; tauto. Qed. Lemma sat_in2' : forall {T:Type} (P Q:T->Prop) (x:T), In [t:T | P t /\ ~ Q t] x -> ~ Q x. intros ? ? ? ? h1; destruct h1; tauto. Qed. (*Wyckoff*) Lemma sat_nin : forall {T:Type} (P:T->Prop) (x:T), ~ In [t:T | P t] x -> ~P x. intros T P x h1. intro h2. contradict h1. constructor. assumption. Qed. (*Wyckoff*) Lemma sat_nin' : forall {T:Type} (P:T->Prop) (x:T), ~ In [t:T | ~ P t] x -> P x. intros T P x h1. apply NNPP. intro h2. contradict h1. constructor; auto. Qed. (*Wyckoff*) Lemma eq_sat : forall {T:Type} (A:Ensemble T), A = [x:T | In A x]. intros T A. rewrite characteristic_function_to_ensemble_is_identity. reflexivity. Qed. (*Wyckoff*) Lemma im_sat : forall {T U:Type} (P:T->Prop) (f:T->U), Im [x:T | P x] f = [y:U | exists x:T, P x /\ y = f x]. intros T U P f. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]. subst. destruct h1 as [h1]. constructor. exists x. split; auto. destruct h1 as [h1]. destruct h1 as [y h1]. destruct h1 as [h1 h2]. subst. apply Im_intro with y; auto. constructor; auto. Qed. (*Wyckoff*) Lemma sat_iff : forall {T:Type} (P Q:T->Prop), (forall x:T, P x <-> Q x) -> [x : T | P x] = [x : T | Q x]. intros T P Q h1. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [h2]. rewrite h1 in h2. constructor; auto. red. intros x h2. destruct h2 as [h2]. rewrite <- h1 in h2. constructor; auto. Qed. (*Wyckoff*) Lemma sig_in_sat : forall {T:Type} (P:T->Prop) (x:{t:T | P t}), In [t:T | P t] (proj1_sig x). constructor. apply proj2_sig. Qed. (*Wyckoff*) Lemma sig_in_sat2 : forall {T:Type} (P Q:T->Prop) (x:{t:T | P t}), Q (proj1_sig x) -> In [t:T | P t /\ Q t] (proj1_sig x). constructor; intros; split; [apply proj2_sig | auto]. Qed. (*Wyckoff*) Lemma sig_in_sat2' : forall {T:Type} (P Q:T->Prop) (x:{t:T | Q t}), P (proj1_sig x) -> In [t:T | P t /\ Q t] (proj1_sig x). constructor; intros; split; [auto | apply proj2_sig]. Qed. End Abstr. (*End Abstr*) (*Wyckoff*) Lemma not_inhabited_empty : forall {T:Type} (S:Ensemble T), ~ Inhabited S -> S = Empty_set _. intros T S h1. apply Extensionality_Ensembles. red. split. red. intros x h2. pose proof (Inhabited_intro _ _ _ h2). contradiction. auto with sets. Qed. (*Wyckoff*) Lemma ex_inh : forall {T:Type} (P:T->Prop), (exists x:T, P x) -> Inhabited [x:T | P x]. intros T P h1. destruct h1 as [x h1]. apply Inhabited_intro with x; constructor; auto. Qed. (*Wyckoff*) Lemma empty_alt : forall {T:Type} (A:Ensemble T), (forall x, ~In A x) -> A = Empty_set _. intros T A h1. red in h1. destruct (classic (Inhabited A)) as [h2 | h3]. destruct h2 as [x h2]. contradict h1; auto. intro h1. apply (h1 _ h2). apply not_inhabited_empty; auto. Qed. Section Complement. (*Wyckoff*) Definition excl_midP {T:Type} (P:T->Prop) : Prop := forall x:T, P x \/ ~ P x. (*Wyckoff*) Lemma excl_middle_empty : forall {T:Type} (S:Ensemble T), Intersection S (Comp S) = Empty_set _. intros T S. apply Extensionality_Ensembles. red. split. (*left*) red. intros x h1. inversion h1; contradiction. (*right*) red. intros; contradiction. Qed. (*Wyckoff*) Lemma excl_middle_full : forall {T:Type} (S:Ensemble T), Union S (Comp S) = Full_set _. intros T S. apply Extensionality_Ensembles. red. split. (*left*) red. intros. constructor. auto with sets. (*right*) red. intros x h1. case (classic (In S x)) as [h2 | h3]. apply Union_introl; trivial. apply Union_intror. auto with sets. Qed. (*Wyckoff*) Lemma complement_full : forall T:Type, Comp (Full_set T) = Empty_set T. intro T. apply Extensionality_Ensembles. red. unfold Included. split. (* left *) intros x h2. unfold Comp in h2. red in h2. elim h2. apply Full_intro. (* right *) intros x h3. contradiction. Qed. (*Wyckoff*) Lemma complement_empty : forall T:Type, Comp (Empty_set T) = Full_set T. intro T. pose proof (Complement_Complement _ (Full_set T)) as h1. rewrite complement_full in h1. assumption. Qed. (*Wyckoff*) Lemma complement_meets_non_subset : forall {T:Type} (P Q : Ensemble T), (Inhabited P) -> (~Included P Q <-> Inhabited (Intersection (Comp Q) P)). intros T P Q h1. split. (*left*) intro h2. unfold Included in h2. pose proof (not_all_ex_not _ _ h2) as h4. elim h4. intros x h5. assert (h6: (In P x) /\ (~In Q x)). tauto. destruct h6 as [h7 h8]. assert (h9: In (Intersection (Comp Q) P) x). constructor; auto with sets. apply Inhabited_intro with x; trivial. (*right*) intro h2. unfold Included. inversion h2 as [x h3]. assert (h4: exists x:T, ~(In P x -> In Q x)). exists x. assert (h5: (In P x) /\ (~ In Q x)). split. inversion h3 as [h6 h7]. assumption. inversion h3 as [y h6 h7]. compute in h7. auto. tauto. apply ex_not_not_all; trivial. Qed. (*Wyckoff*) Lemma included_empty_complement_int : forall {T:Type} (A B:Ensemble T), Included A B <-> Intersection A (Comp B) = Empty_set _. intros T A B. split. (* -> *) intros h1. apply Extensionality_Ensembles. red. split. (* <= *) red in h1. red. intros x h2. destruct h2 as [x h2l h2r]. specialize (h1 _ h2l). assert (h3:In (Intersection B (Comp B)) x). auto with sets. rewrite excl_middle_empty in h3. assumption. (* >= *) auto with sets. (* <- *) intro h1. red. intros x h2. apply NNPP. intro h3. assert (h4:In (Intersection A (Comp B)) x). constructor; assumption. rewrite h1 in h4. contradiction. Qed. (*Wyckoff*) Lemma elimination_union : forall {T:Type} (A B C:Ensemble T), A = Union B C -> C = Intersection A (Union (Comp B) C). intros T A B C h1. pose proof (f_equal (fun D => Intersection D (Union (Comp B) C)) h1) as h2. simpl in h2. rewrite foil_union in h2. rewrite excl_middle_empty in h2. rewrite Intersection_idempotent in h2. rewrite <- (Union_commutative _ C (Intersection C (Comp B))) in h2. rewrite abs_sum_psa in h2. rewrite (Union_commutative _ (Intersection B C) C) in h2. rewrite (Intersection_commutative _ B C) in h2. rewrite abs_sum_psa in h2. rewrite empty_union in h2. symmetry. assumption. Qed. (*Wyckoff*) Lemma complement_inv : forall {T:Type} (A:Ensemble T) (x:T), In (Comp A) x -> ~In A x. auto. Qed. (*Wyckoff*) Lemma complement_inclusion: forall {Y:Type} (S T:Ensemble Y), Included S T -> Included (Comp T) (Comp S). Proof. intros. red; intros. red; intro. contradiction H0. auto with sets. Qed. (*Wyckoff*) Lemma setminus_inc : forall {T:Type} (A B:Ensemble T), Included (Setminus A B) A. intros T A B. red. intros x h1. destruct h1; auto with sets. Qed. (*Wyckoff*) Lemma setminus_empty : forall {T:Type} (A:Ensemble T), Setminus (Empty_set _) A = Empty_set _. intros T A. apply Extensionality_Ensembles; red; split; auto with sets. red. intros x h1. destruct h1. contradiction. Qed. (*Wyckoff*) Lemma empty_setminus2 : forall {T:Type} (A:Ensemble T), Setminus A A = Empty_set _. intros T A. apply Extensionality_Ensembles; red; split; auto with sets. red. intros x h1. destruct h1; contradiction. Qed. (*Wyckoff*) Lemma setminus_empty' : forall {T:Type} (A:Ensemble T), Setminus A (Empty_set _) = A. intros; apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2; auto. constructor; auto. intro; contradiction. Qed. (*Wyckoff*) Lemma setminus_int_complement : forall {T:Type} (A B:Ensemble T), Setminus A B = Intersection A (Comp B). intros T A B. apply Extensionality_Ensembles. red. split. (* <= *) red. intros ? h1. destruct h1 as [h2 h3]. constructor; auto with sets. (* >= *) red. intros x h1. destruct h1 as [h2 h3]. constructor; auto with sets. Qed. (*Wyckoff*) Lemma complement_setminus_full : forall {T:Type} (A:Ensemble T), Comp A = Setminus (Full_set T) A. intros T A. apply Extensionality_Ensembles. red. split. red. intros x h1. constructor. constructor. assumption. red. intros x h1. red. red. destruct h1; auto. Qed. (*Wyckoff*) Lemma setminus_sub_sup : forall {T:Type} (A B:Ensemble T), Included A B -> Setminus A B = Empty_set _. intros T A B h1. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h2. destruct h2 as [h3 h4]. red in h1. apply h1 in h3. contradiction. (* >= *) auto with sets. Qed. (*Wyckoff*) Lemma int_setminus : forall {T:Type} (A X:Ensemble T), Intersection A (Setminus X A) = Empty_set _. intros T A X. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h1. destruct h1 as [x h2 h3]. destruct h3. contradiction. (* >= *) auto with sets. Qed. (*Wyckoff*) Lemma setminus_same_int : forall {T:Type} (X Y:Ensemble T), Setminus X (Intersection X Y) = Setminus X Y. intros T X Y. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [h1 h2]. constructor. assumption. auto with sets. red. intros x h1. constructor. destruct h1 as [h1 h2]. assumption. intro h2. destruct h2; destruct h1; contradiction. Qed. (*Wyckoff*) Lemma setminus_int_eq : forall {T:Type} (A B C:Ensemble T), Setminus (Intersection A B) C = Intersection A (Setminus B C). intros T A B C. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1 h2]. destruct h1 as [x h1a h1b]. constructor; auto. constructor; auto. destruct h1 as [x h1 h2]. destruct h2 as [h2 h3]. constructor; auto. constructor; auto. Qed. (*Wyckoff*) Lemma decompose_int_setminus : forall {T:Type} (X Y:Ensemble T), X = Union (Intersection X Y) (Setminus X Y). intros T X Y. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct (classic (In Y x)). left. constructor; auto. right. constructor; auto. red. intros x h1. destruct h1 as [x h1l | x h1r]. destruct h1l; auto. destruct h1r; auto. Qed. (*Wyckoff*) Lemma decompose_setminus_inc : forall {T:Type} (X Y:Ensemble T), Included Y X -> X = Union Y (Setminus X Y). intros T X Y h1. rewrite (decompose_int_setminus X Y) at 1. rewrite inclusion_iff_intersection_eq in h1. rewrite Intersection_commutative in h1. rewrite h1. reflexivity. Qed. (*Wyckoff*) Lemma setminus_add : forall {T:Type} (A B:Ensemble T) (a:T), Setminus (Add A a) B = Union (Setminus A B) (Setminus (Singleton a) B). intros T A B a. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [h1 h2]. destruct h1 as [x h1 | x h3]. left. constructor; auto. destruct h3. right. constructor; auto. constructor. red. intros x h1. destruct h1 as [x h1 | x h2]. destruct h1 as [h1 h2]. constructor. left; auto. assumption. destruct h2 as [h2 h3]. destruct h2. constructor. right. constructor. assumption. Qed. (*Wyckoff*) Lemma setminus_setminus : forall {T:Type} (A B C:Ensemble T), Setminus (Setminus A B) C = Setminus A (Union B C). intros T A B C. apply Extensionality_Ensembles. red; split; red; intros x h1. destruct h1 as [h1 h2]. destruct h1 as [h1 h3]. constructor; auto. intro h4. destruct h4; contradiction. destruct h1 as [h1 h2]. constructor. constructor; auto. intro h3. contradict h2. left; auto. intro h3. contradict h2. right; auto. Qed. (*Wyckoff*) (*for use in negations*) Lemma in_setminus_iff : forall {T:Type} (A B:Ensemble T) (x:T), In (Setminus A B) x <-> In A x /\ ~In B x. intros. split; intro h1. destruct h1; split; auto. destruct h1; constructor; auto. Qed. (*Wyckoff*) (*Uses LEM *) Lemma setminus_of_same : forall {T:Type} (A B:Ensemble T), Setminus A (Setminus A B) = Intersection A B. intros T A B. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1 h2]. rewrite in_setminus_iff in h2. apply not_and_or in h2. destruct h2 as [h2 | h2]; try contradiction. constructor; try apply NNPP; auto. destruct h1 as [x h1 h2]. constructor; auto. intro h3. destruct h3; contradiction. Qed. (*Wyckoff*) Lemma finite_setminus : forall {T:Type} (A B:Ensemble T), Finite A -> Finite (Setminus A B). intros T A B h1. induction h1 as [|A h1 h2 a h3]. rewrite setminus_empty. constructor. rewrite setminus_add. apply Union_preserves_Finite; auto. rewrite setminus_int_complement. eapply Finite_downward_closed. apply Singleton_is_finite. auto with sets. Qed. (*Wyckoff*) Lemma add_downward_closed : forall {T:Type} (A:Ensemble T) (x:T), Finite (Add A x) -> Finite A. intros T A x h1; eapply Finite_downward_closed. apply h1. auto with sets. Qed. (*Wyckoff*) Lemma setminus_add1 : forall {T:Type} (E:Ensemble T) (x:T), ~ In E x -> Setminus (Add E x) E = Singleton x. intros T E x h1. apply Extensionality_Ensembles. red. split. red. intros y h2. destruct h2 as [h2 h3]. destruct h2 as [y h2l | y h2r]. contradiction. assumption. red. intros y h2. destruct h2; subst. constructor; auto. right; auto. constructor. Qed. (*Wyckoff*) Lemma setminus_add2 : forall {T:Type} (E:Ensemble T) (e:T), Setminus (Add E e) (Singleton e) = Subtract E e. intros T E e. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [h1 h2]. destruct h1 as [x h1a | x h1b]. constructor; auto. contradiction. red. intros x h1. red in h1. red in h1. red in h1. destruct h1 as [h1 h2]. constructor. left. assumption. assumption. Qed. (*Wyckoff*) Lemma add_subtract_a : forall {T:Type} (A:Ensemble T) (a:T), Subtract (Add A a) a = Subtract A a. intros T A a. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h1. destruct h1 as [h1 h2]. destruct h1 as [x h1a | x h1b]. constructor; auto. contradiction. (* >= *) red. intros x h1. destruct h1 as [h1 h2]. pose proof (Add_intro1 _ A a x h1). constructor; auto. Qed. (*Wyckoff*) Lemma setminus_couple_eq : forall {T:Type} (A:Ensemble T) (a a':T), Setminus A (Couple a a') = Subtract (Subtract A a) a'. intros T A a a'. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h1. destruct h1 as [h1 h2]. constructor. constructor. assumption. intro h3. destruct h3; subst. pose proof (Couple_l _ a a'); contradiction. intro h3. destruct h3; subst. pose proof (Couple_r _ a a'); contradiction. (* >= *) red. intros x h1. destruct h1 as [h1 h2]. destruct h1 as [h1a h1b]. constructor. assumption. intro h3. destruct h3. pose proof (In_singleton _ a). contradiction. pose proof (In_singleton _ a'). contradiction. Qed. (*Wyckoff*) Lemma subtract_couple_l : forall {T:Type} (a b:T), a <> b -> Subtract (Couple a b) a = Singleton b. intros T a b h0. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1 h2]. destruct h1. contradict h2. constructor; auto. constructor; auto. destruct h1. constructor. right. intro h1. destruct h1. contradict h0; auto. Qed. (*Wyckoff*) Lemma subtract_couple_r : forall {T:Type} (a b:T), a <> b -> Subtract (Couple a b) b = Singleton a. intros T a b h1. apply neq_sym in h1. rewrite couple_comm. apply subtract_couple_l; auto. Qed. (*Wyckoff*) Lemma subtract2_comm : forall {T:Type} (S:Ensemble T) (a b:T), Subtract (Subtract S a) b = Subtract (Subtract S b) a. intros T S a b. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h1. destruct h1 as [h1a h1b]. destruct h1a as [h1aa h1ab]. constructor. constructor. assumption. assumption. assumption. (* >= *) red. intros x h1. destruct h1 as [h1a h1b]. destruct h1a as [h1aa h1ab]. constructor. constructor. assumption. assumption. assumption. Qed. (*Wyckoff*) Lemma union_singleton_eq_setminus : forall {T:Type} (A B:Ensemble T) (a b:T), Add A a = Add B b -> Setminus A (Couple a b) = Setminus B (Couple a b). intros T A B a b h1. assert (h2:Setminus (Add A a) (Couple a b) = Setminus (Add B b) (Couple a b)). f_equal; auto. repeat rewrite setminus_couple_eq in h2. rewrite add_subtract_a in h2. symmetry in h2. rewrite subtract2_comm in h2. rewrite add_subtract_a in h2. repeat rewrite <- setminus_couple_eq in h2. symmetry. rewrite couple_comm in h2. assumption. Qed. (*Wyckoff*) Lemma subtract_preserves_inclusion : forall {T:Type} (A B:Ensemble T) (x:T), Included A B -> Included (Subtract A x) (Subtract B x). intros T A B x h1. red. intros y h2. destruct h2 as [h2 h3]. constructor; auto with sets. Qed. (*Wyckoff*) Lemma subtract_nin : forall {T:Type} (A:Ensemble T) (x:T), ~In A x -> Subtract A x = A. intros T A x h1. apply Extensionality_Ensembles. red. split. red. intros y h2. destruct h2; auto. red. intros y h2. constructor; auto. intro h3. destruct h3; subst. contradiction. Qed. (*Wyckoff*) Lemma subtract_eq_sing : forall {T:Type} (pfdt:type_eq_dec T) (a b:T), Subtract (Singleton a) b = Empty_set _ -> a = b. intros T hdt a b h1. destruct (hdt a b) as [|h2]; subst; auto. rewrite subtract_nin in h1. apply sing_not_empty in h1. contradiction. rewrite in_sing_iff; auto. Qed. (*Wyckoff*) Lemma incl_add_sub_nin : forall {T:Type} (A B:Ensemble T) (b:T), ~In B b -> Included A (Add B b) -> Included (Subtract A b) B. intros T A B b h1 h2. red; intros x h3. destruct h3 as [h3 h4]. apply h2 in h3. destruct h3 as [|x h3]; auto; try contradiction. Qed. (*Wyckoff*) Lemma incl_union_setminus : forall {T:Type} (A B C:Ensemble T), Included A (Union B C) -> Included (Setminus A C) B. intros T A B C h1. red. intros x h2. destruct h2 as [h2 h3]. apply h1 in h2. destruct h2; auto with sets. contradiction. Qed. (*Wyckoff*) Lemma union_setminus : forall {T:Type} (A X:Ensemble T), Union A (Setminus X A) = Union A X. intros T A X. rewrite setminus_int_complement. rewrite Distributivity'. rewrite excl_middle_full. rewrite intersection_full_set. reflexivity. Qed. (*Wyckoff*) Lemma union_setminus_incl : forall {T:Type} (A X:Ensemble T), Included A X -> Union A (Setminus X A) = X. intros T A X h1. rewrite union_setminus. rewrite <- inclusion_iff_union. assumption. Qed. (*Wyckoff*) Lemma setminus_union : forall {T:Type} (A B:Ensemble T), Setminus (Union A B) B = Setminus A B. intros T A B. apply Extensionality_Ensembles. red; split; red. intros x h1. destruct h1 as [h1 h2]. destruct h1 as [x h1l | x h1r]. constructor; auto. contradiction. intros x h1. destruct h1 as [h1 h2]. constructor; auto. left; auto. Qed. (*Wyckoff*) Lemma union_eq_disjoint_setminus : forall {T:Type} (A B C:Ensemble T), disjoint A B -> Union A B = C -> A = Setminus C B. intros T A B C h1 h2; subst. rewrite setminus_union. symmetry. apply disjoint_setminus_compat; auto. Qed. (*Wyckoff*) Lemma im_setminus : forall {T U:Type} (A B:Ensemble T) (f:T->U), Included (Im (Setminus A B) f) (Intersection (Im A f) (Im (Comp B) f)). intros T U A B f. red. intros x h1. destruct h1 as [x h1]. subst. destruct h1 as [h1 h2]. constructor. apply Im_intro with x; auto. apply Im_intro with x. auto. reflexivity. Qed. (*Wyckoff*) Lemma incl_im_setminus : forall {T U:Type} (A B:Ensemble T), Included B A -> forall (f:T->U), Included (Setminus (Im A f) (Im B f)) (Im (Setminus A B) f). intros T U A B h1 f. assert (h2:Im A f = Union (Im B f) (Im (Setminus A B) f)). rewrite (decompose_setminus_inc _ _ h1) at 1. rewrite im_union. reflexivity. apply elimination_union in h2. rewrite Distributivity in h2. pose proof (setminus_inc A B) as h3. pose proof h3 as h3'. rewrite inclusion_iff_union in h3'. rewrite Union_commutative in h3'. apply (im_preserves_inclusion _ _ f) in h3. rewrite inclusion_iff_intersection_eq in h3. rewrite Intersection_commutative in h3. rewrite h3 in h2. symmetry in h2. rewrite <- inclusion_iff_union in h2. rewrite <- setminus_int_complement in h2. assumption. Qed. (*Wyckoff*) Lemma inj_im_setminus : forall {T U:Type} (A B:Ensemble T) (f:T->U), injective f -> Im (Setminus A B) f = Setminus (Im A f) (Im B f). intros T U A B f h1. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2 as [x h2]; subst. destruct h2 as [h2 h3]. constructor. econstructor. apply h2. reflexivity. contradict h3. inversion h3 as [? ? ? h5 h6]; subst. apply h1 in h5; subst. assumption. destruct h2 as [h2 h3]. destruct h2 as [x h2]; subst. econstructor. constructor. apply h2. contradict h3. econstructor. apply h3. reflexivity. reflexivity. Qed. (*Wyckoff*) Lemma im_complement_incl_inj : forall {T U:Type} (A:Ensemble T) (f:T->U), injective f -> Included (Im (Comp A) f) (Comp (Im A f)). intros T U A f h0. red. intros u h1. destruct h1 as [u h1]. subst. red in h1. red in h1. red. red. intro h2. inversion h2 as [x h3 h4 h5]. subst. apply h0 in h5. subst. contradiction. Qed. (*Wyckoff*) Lemma im_complement_incl_bij : forall {T U:Type} (A:Ensemble T) (f:T->U), bijective f -> Included (Im (Comp A) f) (Setminus (Im (Full_set _) f) (Im A f)). intros T U A f h0. red in h0. destruct h0 as [h0 h0']. pose proof (im_complement_incl_inj A f h0) as h1. rewrite setminus_int_complement. rewrite Intersection_commutative. rewrite surj_iff in h0'. rewrite h0'. rewrite intersection_full_set. assumption. Qed. (*Wyckoff*) Lemma im_complement_bij : forall {T U:Type} (A:Ensemble T) (f:T->U), bijective f -> Im (Comp A) f = Setminus (Im (Full_set _) f) (Im A f). intros T U A f h0. apply Extensionality_Ensembles. red. split. apply im_complement_incl_bij. assumption. red. intros u h1. red in h0. destruct h0 as [h0 h0']. rewrite surj_iff in h0'. pose proof (Full_intro _ u) as h2. rewrite <- h0' in h2. destruct h2 as [u h2]. subst. apply Im_intro with u. intro h3. inversion h1 as [h4 h5]. contradict h5. apply Im_intro with u; auto. reflexivity. Qed. (*Wyckoff*) Lemma im_set_minus_inj : forall {T U:Type} (A B:Ensemble T) (f:T->U), injective f -> Im (Setminus A B) f = Intersection (Im A f) (Im (Comp B) f). intros T U A B f h0. apply Extensionality_Ensembles. red. split. apply im_setminus. red. intros x h1. destruct h1 as [u h1 h2]. destruct h1 as [u h1]. subst. inversion h2 as [y h3 x h4]. subst. rewrite h4. apply h0 in h4. subst. apply Im_intro with y. constructor; auto. reflexivity. Qed. End Complement. (*Wyckoff*) Lemma incl_card_fun1_eq : forall {T:Type} (A B:Ensemble T), Finite B -> Included A B -> card_fun1 A = card_fun1 B -> A = B. intros T A B h1. revert A. induction h1 as [|B h2 h3 x h4]. intros; auto with sets. intros A h1 h5. pose proof (Finite_downward_closed _ _ (Add_preserves_Finite _ _ x h2) _ h1) as hf. eapply subtract_preserves_inclusion in h1. rewrite sub_add_same_nin in h1; auto. destruct (classic (In A x)) as [h6 | h7]. rewrite card_add_nin' in h5; auto. rewrite (card_sub_in _ hf x) in h5. apply S_inj in h5. specialize (h3 _ h1 h5). rewrite <- (add_sub_compat_in _ _ h6). f_equal. assumption. assumption. rewrite subtract_nin in h1; auto. apply incl_card_fun1 in h1; auto. rewrite h5 in h1; auto. rewrite card_add_nin' in h1; auto. omega. Qed. Lemma in2_card : forall (A:Ensemble nat) (c d:nat), c <> d -> Inn A c -> Inn A d -> Finite A -> 1 < card_fun1 A. intros A c d h1 h2 h3 h4. erewrite card_sub_in. erewrite card_sub_in. auto with arith. apply subtract_preserves_finite. assumption. constructor. apply h3. rewrite in_sing_iff. apply h1. assumption. assumption. Qed. (*Wyckoff*) Lemma cardinal_2 : forall {T:Type} (A:Ensemble T), cardinal _ A 2 -> exists (x y:T), x <> y /\ A = Couple x y. intros T A h1. assert (h2:Inhabited A). apply not_empty_Inhabited. intro h2. subst. inversion h1. subst. pose proof (Add_intro2 _ A x) as h3. rewrite H in h3. contradiction. destruct h2 as [x h2]. pose proof (cardinal_finite _ _ _ h1) as h4. pose proof (card_sub_in A h4 x h2) as h5. assert (ha:Finite (Subtract A x)). pose proof (incl_subtract A x) as h6. apply (Finite_downward_closed _ _ h4 _ h6). pose proof (card_fun1_compat A) as h6. simpl in h6. destruct h6 as [h6 h7]. clear h7. specialize (h6 h4). assert (h3:Inhabited (Subtract A x)). apply not_empty_Inhabited. intro h3. rewrite h3 in h5. rewrite card_fun1_empty in h5. rewrite h5 in h6. pose proof (cardinal_unicity _ _ _ h1 _ h6). omega. destruct h3 as [y h3]. pose proof (card_sub_in (Subtract A x) ha y h3) as hb. rewrite hb in h5. pose proof (card_fun1_compat (Subtract A x)) as h6'. simpl in h6'. destruct h6' as [h6' h7']. clear h7'. specialize (h6' ha). exists x, y. split. destruct h3. intro h8. subst. contradict H0. constructor. pose proof (cardinal_unicity _ _ _ h1 _ h6) as h7. rewrite h5 in h7. assert (h8:card_fun1 (Subtract (Subtract A x) y) = 0). omega. apply card_fun1_O in h8. destruct h8 as [h8l | h8r]. apply Extensionality_Ensembles. red. split. red. intros a h4'. destruct (eq_dec a x), (eq_dec a y); subst. left. left. right. assert (h5':In (Setminus A (Couple x y)) a). constructor; auto. intro h5'. destruct h5'; auto. rewrite setminus_couple_eq in h5'. rewrite h8l in h5'. contradiction. red. intros a h9. destruct h9. assumption. destruct h3; auto. pose proof (add_preserves_infinite _ y h8r) as h10. rewrite add_sub_compat_in in h10; auto. contradiction. Qed. (*Wyckoff*) Lemma card_fun1_couple : forall {T:Type} (x y:T), x <> y -> card_fun1 (Couple x y) = 2. intros T x y h1. pose proof (card_fun1_compat (Couple x y)) as h2. destruct h2 as [h2l h2r]. specialize (h2l (finite_couple x y)). pose proof (cardinal_couple _ _ h1) as h3. eapply cardinal_is_functional; auto. apply h2l. apply h3. Qed. (*Wyckoff*) Lemma incl_couple_inv : forall {T:Type} (A:Ensemble T) (b c:T), Included A (Couple b c) -> A = Empty_set _ \/ A = Singleton b \/ A = Singleton c \/ A = Couple b c. intros T A b c h1. destruct (eq_dec b c) as [he | hne]. subst. rewrite couple_singleton. rewrite couple_singleton in h1. apply singleton_inc in h1. destruct h1 as [h1l | h1r]. left. assumption. right. left. assumption. pose proof (finite_couple b c) as h2. pose proof (card_fun1_couple _ _ hne) as h3. pose proof (Finite_downward_closed _ _ h2 _ h1) as h4. pose proof (incl_card_fun1 _ _ h2 h1) as h5. rewrite h3 in h5. destruct (le_lt_eq_dec (card_fun1 A) 2 h5) as [h6 | h7]. assert (h7:card_fun1 A <= 1). auto with arith. clear h6. destruct (le_lt_eq_dec (card_fun1 A) 1 h7) as [h8 | h9]. apply lt1 in h8. clear h7. left. apply card_fun1_O in h8. destruct h8; auto. contradiction. apply card_fun1_1 in h9. destruct h9 as [x h9]. subst. red in h1. specialize (h1 _ (In_singleton _ x)). destruct h1. right. left. reflexivity. right. right. left. reflexivity. apply incl_card_fun1_eq in h1; auto. rewrite h7. symmetry. assumption. Qed. Section family1. Variable It T:Type. Definition Family := Ensemble (Ensemble T). (*Schepler*) Inductive FamilyUnion (F:Family): Ensemble T := | family_union_intro: forall (S:Ensemble T) (x:T), In F S -> In S x -> In (FamilyUnion F) x. (*Schepler*) Inductive FamilyIntersection (F:Family) : Ensemble T := | family_intersection_intro: forall x:T, (forall S:Ensemble T, In F S -> In S x) -> In (FamilyIntersection F) x. (*Schepler*) Definition IndexedFamily := It -> Ensemble T. Variable FI:IndexedFamily. (*Schepler*) Inductive IndexedUnion : Ensemble T := | indexed_union_intro: forall (i:It) (x:T), In (FI i) x -> In IndexedUnion x. (*Wyckoff*) Inductive seg_union {n} (f:seg_fun n (Ensemble T)) : Ensemble T := | seg_fun_union_intro : forall i (pfi:i < n) (x:T), Inn (f i pfi) x -> Inn (seg_union f) x. (*Schepler*) Inductive IndexedIntersection : Ensemble T := | indexed_intersection_intro: forall (x:T), (forall i:It, In (FI i) x) -> In IndexedIntersection x. (*Schepler*) Definition ImageFamily : Family:= Im (Full_set _) FI. (*Schepler*) Lemma indexed_to_family_union: IndexedUnion = FamilyUnion (ImageFamily). Proof. apply Extensionality_Ensembles. unfold Same_set. unfold Included. intuition. destruct H. apply family_union_intro with (FI i). apply Im_intro with i. constructor. reflexivity. assumption. destruct H. destruct H. apply indexed_union_intro with x0. rewrite <- H1. assumption. Qed. (*Schepler*) Lemma indexed_to_family_intersection: IndexedIntersection = FamilyIntersection ImageFamily. Proof. apply Extensionality_Ensembles. unfold Same_set. unfold Included. intuition. constructor. intros. destruct H. destruct H0. rewrite H1. apply H. constructor. intro. destruct H. apply H. apply Im_intro with i. constructor. reflexivity. Qed. End family1. (*Wyckoff*) Definition im_im_proj1_sig {T:Type} {A:Ensemble T} (F:Family (sig_set A)) : Family T := Im F (@im_proj1_sig _ _). (*Wyckoff*) Lemma im_proj1_sig_family_union : forall {T} (X:Ensemble T) (F:Family (sig_set X)), im_proj1_sig (FamilyUnion _ F) = FamilyUnion _ [S:Ensemble T | Inn (im_im_proj1_sig F) S]. intros; apply Extensionality_Ensembles; red; split; red; intros x h1. inversion h1; clear h1; subst. inversion H; clear H; subst. econstructor. constructor. econstructor. apply H0. reflexivity. econstructor. apply H1. reflexivity. inversion h1; clear h1; subst. destruct H as [h1]. inversion h1; clear h1; subst. inversion H0; clear H0; subst. econstructor. econstructor. apply H. apply H1. reflexivity. Qed. (*Wyckoff*) Lemma family_union_add : forall {T:Type} (F:Family T) (S:Ensemble T), FamilyUnion _ (Add F S) = Union S (FamilyUnion _ F). intros T F S. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h4. unfold Add in h4. inversion h4 as [A y h5 h6]. pose proof (Union_inv _ _ _ A h5) as h7. case h7 as [h8 | h9]. (*h8*) apply Union_intror. apply family_union_intro with A; assumption. (*h9*) left. pose proof (Singleton_inv _ _ _ h9) as h12. rewrite <- h12 in h6; assumption. (* >= *) red. intros x h2. pose proof (Union_inv _ _ _ x h2) as h3. case h3 as [h4 | h5]. (*h4*) pose proof (Add_intro2 _ F S) as h5. apply family_union_intro with S; assumption. (*h5*) inversion h5 as [A y h6 h7]. pose proof (Add_intro1 _ F S) as h8. pose proof (h8 A h6) as h9. apply family_union_intro with A; assumption. Qed. (*Wyckoff*) Lemma familyunion_union_eq : forall {T:Type} (F G:Family T), FamilyUnion _ (Union F G) = Union (FamilyUnion _ F) (FamilyUnion _ G). intros T F G. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [S x h1 h2]. destruct h1 as [S h1|S h1]. left. apply family_union_intro with S; auto. right. apply family_union_intro with S; auto. destruct h1 as [x h1|x h1]. destruct h1 as [S x h1 h2]. apply family_union_intro with S; auto. left; auto. destruct h1 as [S x h1 h2]. apply family_union_intro with S; auto. right; auto. Qed. (*Wyckoff*) Lemma family_union_im : forall {T U:Type} (F:Family T) (f:T->U), FamilyUnion _ (Im F (fun A:Ensemble T=>Im A f)) = Im (FamilyUnion _ F) f. intros T U F f. apply Extensionality_Ensembles. red. split. red. intros u h1. destruct h1 as [A u h2 h3]. destruct h2 as [A h2]. subst. destruct h3 as [u h3]. subst. apply Im_intro with u. apply family_union_intro with A; auto. reflexivity. red. intros u h1. destruct h1 as [u h1]. subst. destruct h1 as [A x h2 h3]. apply family_union_intro with (Im A f). apply Im_intro with A; auto. apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma family_union_im_int : forall {T:Type} (F:Family T) (A:Ensemble T), FamilyUnion _ (Im F (fun S : Ensemble T => Intersection A S)) = Intersection (FamilyUnion _ F) A. intros T F A. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [S x h1 h2]. destruct h1 as [S h1]. subst. destruct h2 as [x h2 h3]. constructor; auto. apply family_union_intro with S; auto. destruct h1 as [x h1 h2]. destruct h1 as [S x h1 h3]. eapply family_union_intro. eapply Im_intro. apply h1. apply eq_refl. constructor; auto. Qed. (*Wyckoff*) Lemma incl_family_union : forall {T:Type} (F:Family T) (E:Ensemble T), In F E -> Included E (FamilyUnion _ F). intros T F E h1 x h2. apply family_union_intro with E; auto. Qed. (*Wyckoff*) Lemma setminus_family_union : forall {T:Type} (F:Family T) (A:Ensemble T), Setminus (FamilyUnion _ (Add F A)) A = Setminus (FamilyUnion _ F) A. intros T F A. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1 h2]. destruct h1 as [S x h3 h4]. destruct h3 as [S h5|S h5]. constructor. apply family_union_intro with S; auto. assumption. destruct h5. contradiction. rewrite family_union_add. rewrite Union_commutative. rewrite setminus_union; auto. Qed. Section family2. (*Schepler*) Lemma empty_indexed_intersection: forall {T:Type} (F:IndexedFamily False T), IndexedIntersection _ _ F = (Full_set _). Proof. intros. apply Extensionality_Ensembles; red; split; red; intros; auto with sets. constructor. constructor. destruct i. Qed. Variable T:Type. (*Schepler*) Lemma empty_family_union: FamilyUnion T (Empty_set _) = (Empty_set _). Proof. apply Extensionality_Ensembles. unfold Same_set. unfold Included. intuition. destruct H. contradiction H. contradiction H. Qed. (*Wyckoff*) Lemma family_union_sing : forall (A:Ensemble T), FamilyUnion _ (Singleton A) = A. intros A. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [B x h1]. destruct h1; subst. assumption. red. intros x h1. apply family_union_intro with A; auto with sets. Qed. (*Wyckoff*) Lemma family_union_couple : forall (A B:Ensemble T), FamilyUnion _ (Couple A B) = Union A B. intros A B. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [S x h1 h2]. destruct h1; [left | right]; auto. destruct h1 as [x h1 | x h1]. apply family_union_intro with A; auto; left. apply family_union_intro with B; auto; right. Qed. (*Wyckoff*) Definition chain_set (F:Family T) : Prop := forall (S1 S2:Ensemble T), In F S1 -> In F S2 -> Included S1 S2 \/ Included S2 S1. (*Wyckoff*) Definition pairwise_disjoint (F:Family T) : Prop := forall (A B:Ensemble T), A <> B -> In F A -> In F B -> disjoint A B. (*Wyckoff*) Lemma pairwise_disjoint_empty : pairwise_disjoint (Empty_set (Ensemble T)). red. intros; contradiction. Qed. (*Wyckoff*) Lemma pairwise_disjoint_incl : forall (E F:Family T), pairwise_disjoint F -> Included E F -> pairwise_disjoint E. intros E F h1 h2. red. red in h1. red in h2. intros A B h3 h4 h5. pose proof (h2 _ h4) as h6. pose proof (h2 _ h5) as h7. apply h1; auto. Qed. (*Wyckoff*) Lemma pairwise_disjoint_add : forall (F:Family T), pairwise_disjoint F -> forall A:Ensemble T, (forall S, In F S -> disjoint S A) -> pairwise_disjoint (Add F A). intros F h1 A h2. red. intros C D h3 h4 h5. destruct h4 as [C h4 | C h4]. destruct h5 as [D h5 | D h5]. apply h1; auto. destruct h5. apply h2; auto. destruct h4. destruct h5 as [D h5 | D h5]. apply sym_disjoint. apply h2; auto. destruct h5. contradict h3. reflexivity. Qed. (*Wyckoff*) Lemma pairwise_disjoint_add_rev : forall (F:Family T) (A:Ensemble T), pairwise_disjoint (Add F A) -> pairwise_disjoint F. intros F A h1. red in h1. red. intros C D h2 h3 h4. pose proof (Add_intro1 _ _ A _ h3) as h5. pose proof (Add_intro1 _ _ A _ h4) as h6. apply h1; auto. Qed. (*Wyckoff*) Inductive inhabited_family {T:Type} (F:Family T) : Prop := inhabited_family_intro : Inhabited F -> (forall A:Ensemble T, In F A -> Inhabited A) -> inhabited_family F. (*Wyckoff*) Lemma family_intersection_relativization : forall (F:Family T) (S:Ensemble T), In F S -> FamilyIntersection _ F = FamilyIntersection _ (Im F (fun A => Intersection A S)). intros F S h1. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. constructor. intros S' h3. destruct h3 as [S' h3]. subst. constructor; apply h2; auto. red. intros x h2. destruct h2 as [x h2]. constructor. intros S' h3. specialize (h2 (Intersection S' S)). assert (h4: In (Im F (fun A : Ensemble T => Intersection A S)) (Intersection S' S)). apply Im_intro with S'. assumption. reflexivity. specialize (h2 h4). destruct h2; auto. Qed. (*Wyckoff*) Lemma family_incl_intersection_cont : forall (F F':Family T), Included F F' -> Included (FamilyIntersection _ F') (FamilyIntersection _ F). intros F F' h1. red. intros x h2. destruct h2 as [x h2]. constructor. intros S h3. apply h2; auto. Qed. (*Wyckoff*) Definition finc {T:Type} (E:Ensemble T) : Type := {F : Ensemble T | Finite F /\ Included F E}. (*Wyckoff*) Lemma family_union_im_full_set_finc_proj1_sig_eq : forall (E:Ensemble T), FamilyUnion _ (Im (Full_set (finc E)) (@proj1_sig _ _)) = E. intro E. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [S x h1 h2]. destruct h1 as [S h1]. subst. clear h1. destruct S as [S [h3 h4]]. simpl in h2. apply h4; auto. red. intros x h1. gterm1. redterm0 c. pose proof (family_union_intro _ c0 (Singleton x) x) as h2. unfold c0 in h2. apply h2. pose proof (Singleton_is_finite _ x) as h3. assert (hinc:Included (Singleton x) E). red; intros ? h4; destruct h4; auto. apply Im_intro with (exist (fun F : Ensemble T => Finite F /\ Included F E) (Singleton x) (conj h3 hinc)). constructor. constructor. constructor. Qed. (*Wyckoff*) Definition finc_union {E:Ensemble T} (A:finc E) (B:finc E) : finc E. destruct A as [A h1], B as [B h2]. destruct h1 as [h1a h1b], h2 as [h2a h2b]. refine (exist _ _ _). split. eapply Union_preserves_Finite. apply h1a. apply h2a. rewrite <- included2_impl_union_included. split; auto. Defined. (*Wyckoff*) Definition finc_intersection {E:Ensemble T} (A:finc E) (B:finc E) : finc E. destruct A as [A h1], B as [B h2]. destruct h1 as [h1a h1b], h2 as [h2a h2b]. refine (exist _ _ _). split. eapply intersection_preserves_finite. apply h1a. pose proof (Intersection_decreases_l _ A B) as h3. pose proof (inclusion_transitive T _ _ _ h3 h1b) as h4. apply h4. Defined. (*Wyckoff*) Lemma proj1_sig_finc_union : forall {E:Ensemble T} (A:finc E) (B:finc E), proj1_sig (finc_union A B) = Union (proj1_sig A) (proj1_sig B). intros E A B. destruct A as [A h1], B as [B h2]. destruct h1 as [h1 h1'], h2 as [h2 h2']. unfold finc_union. simpl. reflexivity. Qed. (*Wyckoff*) Lemma proj1_sig_finc_intersection : forall {E:Ensemble T} (A:finc E) (B:finc E), proj1_sig (finc_intersection A B) = Intersection (proj1_sig A) (proj1_sig B). intros E A B. destruct A as [A h1], B as [B h2]. destruct h1 as [h1 h1'], h2 as [h2 h2']. unfold finc_intersection. simpl. reflexivity. Qed. (*Wyckoff*) Lemma finc_union_incl_l : forall {E:Ensemble T} (A:finc E) (B:finc E), Included (proj1_sig A) (proj1_sig (finc_union A B)). intros E A B. rewrite proj1_sig_finc_union. auto with sets. Qed. (*Wyckoff*) Lemma finc_union_incl_r : forall {E:Ensemble T} (A:finc E) (B:finc E), Included (proj1_sig B) (proj1_sig (finc_union A B)). intros E A B. rewrite proj1_sig_finc_union. auto with sets. Qed. End family2. Arguments FamilyUnion [T] _ _. Arguments FamilyIntersection [T] _ _. Arguments IndexedUnion _ [T] _ _. Arguments seg_union [T] [n] _ _. Arguments IndexedIntersection _ [T] _ _. Arguments chain_set [T] _. Arguments disjoint [T] _ _. Arguments pairwise_disjoint [T] _. Section Power_set_Algebra_axioms. Variable Xt:Type. (*Wyckoff*) Lemma Union_associative : forall N M P:(Ensemble Xt), (Union N (Union M P)) = (Union (Union N M) P). intros N M P. apply Extensionality_Ensembles. unfold Same_set. split; unfold Included. (*left*) intros x H. inversion H. assert (h1 : In (Union N M) x). apply Union_introl. assumption. apply Union_introl. assumption. inversion H0. assert (h1 : In (Union N M) x). apply Union_intror. assumption. apply Union_introl. assumption. apply Union_intror. assumption. (*right*) intros x H. inversion H. inversion H0. apply (Union_introl _ _ _ _ H2). assert (h2: In (Union M P) x). apply (Union_introl _ _ _ _ H2). apply (Union_intror _ _ _ _ h2). assert (h2: In (Union M P) x). apply (Union_intror _ _ _ _ H0). apply (Union_intror _ _ _ _ h2). Qed. (*Wyckoff*) Lemma Union_commutative : forall N M:(Ensemble Xt), Union N M = Union M N. intros N M. apply Extensionality_Ensembles. unfold Same_set. split; unfold Included. (*left*) intros x h1. inversion h1 as [x' h2 h3 | x'' h4 h5]. apply Union_intror; assumption. apply Union_introl; assumption. (*right*) intros x h6. inversion h6 as [x' h7 h8 | x'' h9 h10]. apply Union_intror; assumption. apply Union_introl; assumption. Qed. (*Wyckoff*) Lemma Intersection_commutative : forall N M:(Ensemble Xt), Intersection N M = Intersection M N. intros N M. apply Extensionality_Ensembles. unfold Same_set. split; unfold Included. (*left*) intros x h1. inversion h1 as [x' h2 h3 h4]. apply Intersection_intro; assumption. (*right*) intros x h5. inversion h5 as [x' h6 h7 h8]. apply Intersection_intro; assumption. Qed. (*Wyckoff*) Lemma Distributivity : forall N M P:(Ensemble Xt), Intersection P (Union N M) = Union (Intersection P N) (Intersection P M). intros N M P. apply Extensionality_Ensembles. unfold Same_set. split; unfold Included. (*left*) intros x h1. inversion h1 as [x' h2 h3 h4]. inversion h3 as [x'' h5 h6 | x'' h7 h8]. assert (h9: In (Intersection P N) x). apply Intersection_intro; assumption. apply Union_introl; assumption. assert (h10: In (Intersection P M) x). apply Intersection_intro; assumption. apply Union_intror; assumption. (*right*) intros x h11. inversion h11 as [x' h12 h13 | x'' h14 h15]. inversion h12 as [x''' h16 h17 h18]. assert (h19 : In (Union N M) x). apply Union_introl; assumption. apply Intersection_intro; assumption. inversion h14 as [x''' h20 h21 h22]. assert (h23: In (Union N M) x). apply Union_intror; assumption. apply Intersection_intro; assumption. Qed. (*Wyckoff*) Lemma dist_prod_psa : forall N M P:(Ensemble Xt), Union P (Intersection N M) = Intersection (Union P N) (Union P M). intros N M P. apply Extensionality_Ensembles. unfold Same_set. split; unfold Included. (*left*) intros x h1. inversion h1 as [x' h2 h3 | x'' h4 h5]. assert (h6: In (Union P N) x). auto with sets. assert (h7: In (Union P M) x). auto with sets. auto with sets. inversion h4. assert (h8: In (Union P N) x). auto with sets. assert (h9: In (Union P N) x). auto with sets. auto with sets. (*right*) intros x h10. inversion h10 as [x' h11 h12 h13]. inversion h11. auto with sets. inversion h12. auto with sets. assert (h14: In (Intersection N M) x). auto with sets. auto with sets. Qed. (*Wyckoff*) Lemma comp_sum_psa: forall N:(Ensemble Xt), Union N (Comp N) = Full_set Xt. intros N. apply Extensionality_Ensembles. unfold Same_set. split; unfold Included. (*left*) intros x h1. apply Full_intro. (*right*) intros x h2. assert (h3: ~(In N x) \/ (In N x)). tauto. elim h3; auto with sets. Qed. (*Wyckoff*) Lemma comp_prod_psa: forall N:(Ensemble Xt), Intersection N (Comp N) = Empty_set Xt. intros N. apply Extensionality_Ensembles. unfold Same_set. split; unfold Included. (*left*) intros x h1. contradict h1. intro h2. inversion h2 as [x' h3 h4]. unfold In in h3. unfold In in h4. unfold Comp in h4. tauto. (*right*) intros x h5. inversion h5. Qed. End Power_set_Algebra_axioms. (*Wyckoff*) Lemma union_eq_empty : forall {T:Type} (A B:Ensemble T), Union A B = Empty_set _ -> A = Empty_set _ /\ B = Empty_set _. intros T A B h1. split; apply Extensionality_Ensembles; split; auto with sets. red. intros x h2. pose proof (Union_introl _ A B _ h2) as h3. rewrite h1 in h3. assumption. red. intros x h2. pose proof (Union_introl _ B A _ h2) as h3. rewrite Union_commutative in h3. rewrite h1 in h3. assumption. Qed. (*Wyckoff*) Lemma intersection_eq_full : forall {T:Type} (A B:Ensemble T), Intersection A B = Full_set _ -> A = Full_set _ /\ B = Full_set _. intros T A B h1. split; apply Extensionality_Ensembles; split; try (red; intros; constructor). red. intros x h2. rewrite <- h1 in h2. destruct h2; assumption. red. intros x h2. rewrite <- h1 in h2. destruct h2; assumption. Qed. (*Wyckoff*) Lemma card_disj_union : forall {T:Type} (A B:Ensemble T) (m n:nat), Intersection A B = Empty_set _ -> cardinal _ A m -> cardinal _ B n -> cardinal _ (Union A B) (m+n). intros T A B m n h1 h2 h3. revert h1. induction h2 as [|A m h2 h4 a h5]. intro h4. assert (h5:0 + n = n). auto with arith. rewrite h5. rewrite empty_union. assumption. intro h6. rewrite union_add_comm. assert (h7:S m + n = S(m+n)). auto with arith. rewrite h7. constructor. assert (h8:Included A (Add A a)). auto with sets. pose proof (intersection_preserves_inclusion _ _ B h8) as h9. rewrite (Intersection_commutative _ B (Add A a)) in h9. rewrite (Intersection_commutative _ B A) in h9. rewrite h6 in h9. assert (h10:Intersection A B = Empty_set _). auto with sets. apply h4; auto. intro h8. destruct h8 as [a h8l | a h8r]. contradiction. pose proof (Add_intro2 _ A a) as h9. assert (h10:In (Intersection (Add A a) B) a). auto with sets. rewrite h6 in h10. contradiction. Qed. (*Wyckoff*) Lemma card_disj_union' : forall {T:Type} (A B:Ensemble T) (pfa:Finite A) (pfb:Finite B), Intersection A B = Empty_set _ -> (card_fun A pfa) + (card_fun B pfb) = card_fun (Union A B) (Union_preserves_Finite _ _ _ pfa pfb). intros T A B h1 h2 h3. pose proof (card_disj_union A B (card_fun A h1) (card_fun B h2) h3) as h4. pose proof (card_fun_compat A h1) as h5. pose proof (card_fun_compat B h2) as h6. specialize (h4 h5 h6). pose proof (card_fun_compat (Union A B) (Union_preserves_Finite T A B h1 h2)) as h7. eapply cardinal_is_functional; auto. apply h4. apply h7. Qed. (*Wyckoff*) Lemma card_disj_union'' : forall {T:Type} (A B:Ensemble T), Finite A -> Finite B -> Intersection A B = Empty_set _ -> card_fun1 (Union A B) = (card_fun1 A) + (card_fun1 B). intros T A B h1 h2 h3. pose proof (card_disj_union' A B h1 h2 h3) as h4. do 3 rewrite card_fun_card_fun1_compat in h4. rewrite h4; auto. Qed. (*Wyckoff*) Lemma card_decompose_int_setminus : forall {T:Type} (X Y:Ensemble T) (m n p:nat), cardinal _ X m -> cardinal _ (Intersection X Y) n -> cardinal _ (Setminus X Y) p -> m = n + p. intros T X Y m n p h1 h2 h3. pose proof (decompose_int_setminus X Y) as h4. rewrite h4 in h1. pose proof (setminus_same_int X Y) as h5. rewrite <- h5 in h1. assert (h6:Intersection (Intersection X Y) (Setminus X (Intersection X Y)) = Empty_set _). apply int_setminus. rewrite <- h5 in h3. pose proof (card_disj_union _ _ _ _ h6 h2 h3) as h7. apply (cardinal_unicity _ _ _ h1 _ h7). Qed. (*Wyckoff*) Lemma card_decompose_int_setminus' : forall {T:Type} (X Y:Ensemble T) (m n:nat), cardinal _ X m -> cardinal _ (Intersection X Y) n -> cardinal _ (Setminus X Y) (m-n). intros T X Y m n h1 h2. pose proof (card_decompose_int_setminus X Y m n) as h3. pose proof (cardinal_finite _ _ _ h1) as h4. pose proof (setminus_inc X Y) as h5. pose proof (Finite_downward_closed). pose proof (Finite_downward_closed _ _ h4 _ h5) as h6. pose proof (finite_cardinal _ _ h6) as h7. destruct h7 as [r h7]. specialize (h3 _ h1 h2 h7). subst. assert (h8:n+r-n = r). auto with arith. rewrite h8. assumption. Qed. Section Symdiff. (*Wyckoff*) Definition Symdiff {T:Type} (A B:Ensemble T) := Union (Setminus A B) (Setminus B A). (*Wyckoff*) Lemma symdiff_ref : forall {T:Type} (A:Ensemble T), Symdiff A A = Empty_set _. intros T A. unfold Symdiff. rewrite setminus_sub_sup; auto with sets. Qed. (*Wyckoff*) Lemma symdiff_comm : forall {T:Type} (A B:Ensemble T), Symdiff A B = Symdiff B A. intros T A B. unfold Symdiff. auto with sets. Qed. (*Wyckoff*) Lemma symdiff_empty_iff_eq : forall {T:Type} (A B:Ensemble T), Symdiff A B = Empty_set _ <-> A = B. intros T A B. split. (* -> *) intro h1. unfold Symdiff in h1. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h2. apply NNPP. intro h3. pose proof (Setminus_intro _ A B x h2 h3) as h4. pose proof (Union_introl _ (Setminus A B) (Setminus B A) x h4) as h5. rewrite h1 in h5. contradiction. (* >= *) intros x h2. apply NNPP. intro h3. pose proof (Setminus_intro _ B A x h2 h3) as h4. pose proof (Union_introl _ (Setminus B A) (Setminus A B) x h4) as h5. rewrite Union_commutative in h5. rewrite h1 in h5. contradiction. (* <- *) intro h1. rewrite h1. apply symdiff_ref. Qed. (*Wyckoff*) Definition Setminus_full {T:Type} (F X Y:Ensemble T) := Intersection X (Setminus F Y). (*Wyckoff*) Lemma setminus_full_ref : forall {T:Type} (F X:Ensemble T), Setminus_full F X X = Empty_set _. intros T F X. apply Extensionality_Ensembles; split; auto with sets. red. intros x h1. destruct h1 as [x h1a h1b]. destruct h1b; contradiction. Qed. (*Wyckoff*) Lemma setminus_full_empty_inc : forall {T:Type} (F X Y:Ensemble T), Setminus_full F X Y = Empty_set _ -> Included (Intersection F X) Y. intros T F X Y h1. red. intros x h2. apply NNPP. intro h3. unfold Setminus_full in h1. destruct h2 as [x h2l h2r]. assert (h4:In (Intersection X (Setminus F Y)) x). constructor; auto; constructor; auto. rewrite h1 in h4. contradiction. Qed. (*Wyckoff*) Lemma setminus_full_empty_inc' : forall {T:Type} (F X Y:Ensemble T), Setminus_full F X Y = Empty_set _ -> Included X F -> Included X Y. intros T F X Y h1 h2. pose proof (setminus_full_empty_inc F X Y h1) as h3. rewrite inclusion_iff_intersection_eq in h2. rewrite Intersection_commutative in h2. rewrite h2 in h3. assumption. Qed. (*Wyckoff*) Definition Symdiff_full {T:Type} (F X Y:Ensemble T) := Union (Setminus_full F X Y) (Setminus_full F Y X). (*Wyckoff*) Lemma symdiff_full_ref : forall {T:Type} (F A:Ensemble T), Symdiff_full F A A = Empty_set _. intros T F A. apply Extensionality_Ensembles; split; auto with sets. red. intros x h1. destruct h1 as [x h1 |x h2]. rewrite setminus_full_ref in h1. assumption. rewrite setminus_full_ref in h2. assumption. Qed. (*Wyckoff*) Lemma symdiff_full_comm : forall {T:Type} (F A B:Ensemble T), Symdiff_full F A B = Symdiff_full F B A. intros T F A B. unfold Symdiff_full. auto with sets. Qed. (*Wyckoff*) Lemma symdiff_full_empty_iff_eq : forall {T:Type} (F A B:Ensemble T), Included A F -> Included B F -> (Symdiff_full F A B = Empty_set _ <-> A = B). intros T F A B h1 h2. split. (* -> *) intro h3. unfold Symdiff_full in h3. pose proof (union_eq_empty _ _ h3) as h4. destruct h4 as [h4l h4r]. pose proof (setminus_full_empty_inc' _ _ _ h4l h1) as h5. pose proof (setminus_full_empty_inc' _ _ _ h4r h2) as h6. apply Extensionality_Ensembles; red; split; assumption. (* <- *) intro h3. rewrite h3. apply symdiff_full_ref. Qed. End Symdiff. Section DeMorgan. Variable T:Type. (*Wyckoff*) Lemma comp_union : forall (M N : Ensemble T), Comp (Union M N) = Intersection (Comp M) (Comp N). intros M N. apply Extensionality_Ensembles. red. split. (*left*) red. intros x h1. assert (h2: Included M (Union M N)). auto with sets. assert (h3: Included N (Union M N)). auto with sets. split. pose proof (complement_inclusion _ _ h2) as h4. auto with sets. pose proof (complement_inclusion _ _ h3) as h5. auto with sets. (*right*) red. intros x h6. inversion h6 as [y h7 h8]. unfold Comp in h7. red in h7. unfold Comp in h8. red in h8. unfold Comp. red. intro h9. inversion h9. contradiction. contradiction. Qed. (*Wyckoff*) Lemma comp_int : forall (M N : Ensemble T), Comp (Intersection M N) = Union (Comp M) (Comp N). intros M N. pose proof (Complement_Complement _ M) as h1. pose proof (Complement_Complement _ N) as h2. rewrite <- h1, <- h2 at 1. pose proof (comp_union (Comp M) (Comp N)) as h3. rewrite <- h3. apply Complement_Complement. Qed. End DeMorgan. (*Wyckoff*) Lemma incl_union_inv : forall {T:Type} (A B C:Ensemble T), Included A (Union B C) -> Included A B \/ Included A C \/ (Intersection A (Symdiff B C)) <> Empty_set _. intros T A B C h1. destruct (classic (Included A (Union B (Comp C)))) as [h2 | h3]. pose proof (conj h1 h2) as h3. rewrite inclusion_intersection in h3. rewrite <- dist_prod_psa in h3. rewrite excl_middle_empty in h3. rewrite Union_commutative in h3. rewrite empty_union in h3. left; auto. destruct (classic (Included A (Union (Comp B) C))) as [h4 | h5]. rewrite Union_commutative in h1. rewrite Union_commutative in h4. pose proof (conj h1 h4) as h5. rewrite inclusion_intersection in h5. rewrite <- dist_prod_psa in h5. rewrite excl_middle_empty in h5. rewrite Union_commutative in h5. rewrite empty_union in h5. right. left. auto. rewrite included_empty_complement_int in h3. rewrite included_empty_complement_int in h5. rewrite comp_union in h5. rewrite comp_union in h3. rewrite Complement_Complement in h3. rewrite Complement_Complement in h5. pose proof (union_eq_empty (Intersection A (Intersection (Comp B) C)) (Intersection A (Intersection B (Comp C)))) as h6. pose proof (contrapos _ _ h6) as h7. clear h6. assert (h8: ~ (Intersection A (Intersection (Comp B) C) = Empty_set T /\ Intersection A (Intersection B (Comp C)) = Empty_set T) ). apply or_not_and. left. auto. specialize (h7 h8). rewrite <- Distributivity in h7. right. right. unfold Symdiff. do 2 rewrite setminus_int_complement. rewrite Union_commutative in h7. rewrite (Intersection_commutative _ (Comp B) C) in h7. assumption. Qed. (*Wyckoff*) Lemma symdiff_singleton_nin : forall {T:Type} (A:Ensemble T) (x:T), ~In A x -> Symdiff A (Singleton x) = Add A x. intros T A x h1. apply Extensionality_Ensembles. red. split. red. intros a h2. destruct h2 as [a h2 | a h3]. destruct h2 as [h2 h3]. left; auto. destruct h3 as [h3 h4]. destruct h3. right. constructor. red. intros a h2. apply Add_inv in h2. destruct h2 as [h2 | h3]. constructor. constructor; auto. intro h3. destruct h3; contradiction. subst. right. constructor. constructor. assumption. Qed. (*Wyckoff*) Lemma symdiff_sings : forall {T:Type} (b c:T), b <> c -> Symdiff (Singleton b) (Singleton c) = Couple b c. intros T b c h1. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2l | x h2r]. destruct h2l as [h2l]. destruct h2l. left. destruct h2r as [h2r]. destruct h2r. right. red. intros x h2. destruct h2. left. constructor. constructor. intro h2. destruct h2. contradict h1. reflexivity. right. constructor. constructor. intro h2. destruct h2. contradict h1. reflexivity. Qed. Section Infinite_analogues. Variable It T:Type. Variable FI : (IndexedFamily It T). (*Wyckoff*) Definition Int_Fam_Ens (P:Ensemble T) (Q:Family T) : Family T := (Im Q (fun (S:Ensemble T) => (Intersection P S))). (*Wyckoff*) Definition Union_Fam_Ens (P:Ensemble T) (Q:Family T) : Family T := (Im Q (fun (S:Ensemble T) => (Union P S))). (*Wyckoff*) Definition Subtract_Fam_Ens (F:Family T) (x:T) : Family T := (Im F (fun S => Subtract S x)). (*Wyckoff*) Definition Int_Ind_Ens (P:Ensemble T) (i:It) : (Ensemble T) := Intersection P (FI i). (*Wyckoff*) Definition Union_Ind_Ens (P:Ensemble T) (i:It) : (Ensemble T) := Intersection P (FI i). (*Wyckoff*) Lemma dist_union_infnt_1 : forall (P:Ensemble T) (Q:Family T) , Intersection P (FamilyUnion Q) = FamilyUnion (Int_Fam_Ens P Q). intros P Q. apply Extensionality_Ensembles. unfold Same_set. split. (*left*) unfold Included. intros x h1. inversion h1 as [y h2 h3 h4]. inversion h3 as [S z h5 h6 h7]. assert (h8: In (Intersection P S) x). auto with sets. apply family_union_intro with (S := (Intersection P S)). unfold Int_Fam_Ens. apply Im_intro with (x := S). assumption. reflexivity. assumption. (*right*) unfold Included. intros x h8. inversion h8 as [S y h9 h10 h11]. unfold Int_Fam_Ens in h9. apply Intersection_intro. inversion h9 as [A h12 B h13 h14]. rewrite h13 in h10. inversion h10. assumption. inversion h9 as [A h12 B h13 h14]. rewrite h13 in h10. assert (h15: In A x). inversion h10. assumption. apply family_union_intro with (S := A). assumption. assumption. Qed. (*Wyckoff*) Lemma dist_int_infnt_1 : forall (P:Ensemble T) (Q:Family T), Union P (FamilyIntersection Q) = FamilyIntersection (Union_Fam_Ens P Q). intros P Q. apply Extensionality_Ensembles. unfold Same_set. split. (*left*) unfold Included. intros x h1. inversion h1 as [y h2 | y h3]. apply family_intersection_intro. intros S h4. unfold Union_Fam_Ens in h4. inversion h4 as [S' z h5 h6 h7]. assert (h8: In (Union P S') x). auto with sets. rewrite h6. assumption. inversion h3 as [z h9 h10]. apply family_intersection_intro. intros S h11. unfold Union_Fam_Ens in h11. inversion h11 as [A h12 B h13 h14]. assert (h15: In A x). apply h9. assumption. assert (h16: In (Union P A) x). auto with sets. rewrite h13. assumption. (*right*) unfold Included. intros x h17. inversion h17 as [y h18 h19]. unfold Union_Fam_Ens in h18. assert (h20: forall (S':Ensemble T), (In Q S') -> In (Union P S') x). intros S' h50. apply h18. apply Im_intro with (x := S'). assumption. reflexivity. assert (h21 : forall S : Ensemble T, In Q S -> In P x \/ In S x). intros S h22. apply Union_inv. apply h20. assumption. assert (h22: In P x \/ (forall S : Ensemble T, (In Q S -> In S x))). apply prop_dis_forall. intro S. assert (h23: (In Q S -> In P x \/ In S x)). apply h21. tauto. case h22. auto with sets. intro h24. apply Union_intror. apply family_intersection_intro. assumption. Qed. (*Wyckoff*) Lemma dist_union_infnt_1_ind : forall (P:Ensemble T), Intersection P (IndexedUnion _ FI) = IndexedUnion _ (Int_Ind_Ens P). intro P. rewrite (indexed_to_family_union _ _ FI). rewrite indexed_to_family_union. rewrite (dist_union_infnt_1 P (ImageFamily It T FI)). apply Extensionality_Ensembles. unfold Same_set. unfold Int_Fam_Ens. unfold Int_Ind_Ens. unfold Included. split. (*left*) intros x h1. inversion h1. inversion H. unfold ImageFamily. apply family_union_intro with (S := S). unfold ImageFamily in H2. rewrite H3. inversion H2. apply Im_intro with (X := Full_set It) (x := x2). assumption. rewrite H6. reflexivity. assumption. (*right*) intros x h20. unfold ImageFamily in h20. inversion h20. apply family_union_intro with (S := S). inversion H. rewrite H3. apply Im_intro with (x := (FI x1)). unfold ImageFamily. apply Im_intro with (x := x1). assumption. reflexivity. reflexivity. assumption. Qed. (*Wyckoff*) Lemma subtract_family_union : forall (x:T) (F:Family T), Subtract (FamilyUnion F) x = FamilyUnion (Subtract_Fam_Ens F x). intros x F. apply Extensionality_Ensembles. red. split. red. intros a h1. destruct h1 as [h1 h2]. destruct h1 as [A a h1 h3]. apply family_union_intro with (Subtract A x). apply Im_intro with A; auto. constructor; auto. red. intros a h1. destruct h1 as [A a h1 h2]. destruct h1 as [A h1 ? h4]. subst. destruct h2 as [h2l h2r]. constructor. apply family_union_intro with A; auto. assumption. Qed. End Infinite_analogues. (*Wyckoff*) Lemma pairwise_disjoint_add_family_union : forall {T:Type} (F:Family T) (A:Ensemble T), ~Inn F A -> pairwise_disjoint (Add F A) -> disjoint (FamilyUnion F) A. intros T F A h0 h1. red. destruct (eq_dec F (Empty_set _)) as [hem | hnem]. rewrite hem. rewrite empty_family_union. auto with sets. pose proof (not_empty_Inhabited _ _ hnem) as hinh. destruct hinh as [B hin]. rewrite Intersection_commutative. rewrite dist_union_infnt_1. pose proof (Add_intro2 _ F A) as hadd. red in h1. assert (h2:Int_Fam_Ens T A F = Singleton (Empty_set _)). apply Extensionality_Ensembles. red. split. red. intros X h2. destruct h2 as [X h2 ? ?]. subst. pose proof (Add_intro1 _ _ A _ h2) as h3. assert (h4:A <> X). intro h4. subst. contradiction. specialize (h1 _ _ h4 hadd h3). red in h1. rewrite h1. constructor. red. intros X h2. destruct h2; subst. unfold Int_Fam_Ens. apply Im_intro with B. assumption. assert (h2:A <> B). intro h3. subst. contradiction. pose proof (Add_intro1 _ _ A _ hin) as h3. specialize (h1 _ _ h2 hadd h3). red in h1. rewrite h1. reflexivity. rewrite h2. apply family_union_sing. Qed. (*Section Characteristic. Commented to keep the notation top-level*) (*Wyckoff*) Definition power_set {T:Type} (A:Ensemble T) := [S:Ensemble T | Included S A]. (*Wyckoff*) Lemma power_set_empty : forall (T:Type), power_set (Empty_set T) = Singleton (Empty_set T). intro T. unfold power_set. apply Extensionality_Ensembles. red. split. red. intros S h1. destruct h1 as [h1]. assert (h2:Included (Empty_set _) S). auto with sets. assert (h3:S = Empty_set T). apply Extensionality_Ensembles; red; split; auto. subst. constructor. red. intros S h1. destruct h1. constructor. auto with sets. Qed. (*Wyckoff*) Lemma in_power_full : forall {T:Type} (A:Ensemble T), In (power_set (Full_set T)) A. intros T A. constructor. red; intros; constructor. Qed. (*Wyckoff*) Lemma in_power_incl : forall {T:Type} (A B X:Ensemble T), In (power_set X) B -> Included A B -> In (power_set X) A. intros T A B X h1 h2. constructor. destruct h1. auto with sets. Qed. (*Wyckoff*) Lemma in_power_intersection : forall {T:Type} (A B X:Ensemble T), In (power_set X) A -> In (power_set X) (Intersection A B). intros T A B X h1. constructor. eapply (in_power_incl _ _ _ h1). auto with sets. Qed. (*Wyckoff*) Lemma in_power_union : forall {T:Type} (A B X:Ensemble T), In (power_set X) A -> In (power_set X) B -> In (power_set X) (Union A B). intros T A B X h1 h2. destruct h1 as [h1]. destruct h2 as [h2]. constructor. auto with sets. Qed. (*Wyckoff*) Lemma in_power_comp : forall {T:Type} (A X:Ensemble T), In (power_set X) A -> In (power_set X) (Setminus X A). intros T A X h1. constructor. destruct h1 as [h1]. apply setminus_inc. Qed. Definition p_mem {T:Type} (P:T->Prop) (A:Ensemble T) : Prop := forall x:T, In A x -> P x. Lemma p_mem_downward_closed : forall {T:Type} (P:T->Prop) (A A':Ensemble T), Included A A' -> p_mem P A' -> p_mem P A. intros T P A A' h1 h2. red in h2; red; intros x h3. apply h2; auto. Qed. Lemma p_mem_setminus : forall {T:Type} (P:T->Prop) (A B:Ensemble T), p_mem P A -> p_mem P (Setminus A B). intros T P A B h1. red in h1; red. intros x h2; apply h1. destruct h2; auto. Qed. Section Finite_Families. (*Wyckoff*) Lemma Finite_Finite_Union : forall {T:Type} (F:Family T), (forall (S:(Ensemble T)), In F S -> Finite S) -> Finite F -> Finite (FamilyUnion F). intros T F h1 h2. induction h2. (*Empty*) pose proof (empty_family_union T) as h3. rewrite h3. constructor. (*Add*) rename x into S. rename A into F. rename H into h3. rename IHh2 into h4. pose proof (family_union_add F S) as h5. rewrite h5. apply Union_preserves_Finite. apply h1. auto with sets. assert (h6:forall S:Ensemble T, In F S -> Finite S). intros S0 h7. apply h1. auto with sets. apply (h4 h6). Qed. (*Wyckoff*) Definition finite_mem {T:Type} (F:Family T) : Prop := p_mem (@Finite _) F. (*Wyckoff*) Lemma finite_family_union_finite_mem : forall {T:Type} (F:Family T), Finite (FamilyUnion F) -> finite_mem F. intros T F h1. do 2 red; intros. eapply Finite_downward_closed. apply h1. apply incl_family_union; auto. Qed. (*Wyckoff*) Lemma finite_mem_empty : forall (T:Type), finite_mem (Empty_set (Ensemble T)). intro; do 2 red; intros; contradiction. Qed. (*Wyckoff*) Lemma finite_mem_sing_sing : forall {T:Type} (a:T), finite_mem (Singleton (Singleton a)). intros T a. do 2 red; intros S h2. destruct h2; subst. apply Singleton_is_finite. Qed. (*Wyckoff*) Lemma finite_mem_downward_closed : forall {T:Type} (F F':Family T), Included F F' -> finite_mem F' -> finite_mem F. intros T F F' h1 h2. do 2 red in h2. red. intros S h3. apply h1 in h3. apply h2 in h3; auto. Qed. (*Wyckoff*) Lemma finite_mem_setminus : forall {T:Type} (F F':Family T), finite_mem F -> finite_mem (Setminus F F'). intros; apply p_mem_setminus; auto. Qed. (*Wyckoff*) Lemma finite_mem_subtract : forall {T:Type} (F:Family T) (A:Ensemble T), finite_mem F -> finite_mem (Subtract F A). intros; eapply finite_mem_downward_closed. apply incl_subtract. assumption. Qed. (*Wyckoff*) (* Singleton Family decomposition of a set *) Definition SingF {T:Type} (S:Ensemble T) := Im S (@Singleton T). (*Wyckoff*) Lemma union_singF_eq : forall {T:Type} (S:Ensemble T), S = FamilyUnion (SingF S). intros T S. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h1. assert (h2:In (SingF S) (Singleton x)). unfold SingF. apply Im_intro with x; trivial. pose proof (In_singleton _ x) as h3. apply family_union_intro with (Singleton x); assumption. (* >= *) red. intros x h1. inversion h1 as [S0 ? h2 h3]. unfold SingF in h2. inversion h2 as [? h4 ? h5]. rewrite h5 in h3. inversion h3 as [h6]. rewrite <- h6; trivial. Qed. (*Wyckoff*) Lemma fin_fin_singf : forall {T:Type} (S:Ensemble T), Finite S -> Finite (SingF S). intros. apply finite_image; assumption. Qed. (*Wyckoff*) Definition ens_S (E:Ensemble nat) : Ensemble nat := Im E S. (*Wyckoff*) Definition ens_pred (E:Ensemble nat) : Ensemble nat := Im E pred. (*Wyckoff*) Lemma finite_ens_S : forall (E:Ensemble nat), Finite E -> Finite (ens_S E). unfold ens_S; intros; apply finite_image; auto. Qed. (*Wyckoff*) Lemma finite_ens_pred : forall (E:Ensemble nat), Finite E -> Finite (ens_pred E). unfold ens_pred; intros; apply finite_image; auto. Qed. (*Wyckoff*) Lemma ens_pred_undoes_ens_S : forall (E:Ensemble nat), ens_pred (ens_S E) = E. intro E. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]. subst. destruct h1 as [x h1]. subst. simpl. assumption. unfold ens_S, ens_pred. rewrite im_im. apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma incl_pos_ens_S_ens_pred_impl : forall (E:Ensemble nat) (m:nat), In (ens_S (ens_pred E)) (S m) -> m = 0 \/ In E (S m). intros E m h1. unfold ens_S, ens_pred in h1. rewrite im_im in h1. inversion h1 as [m' h2 q hq]. subst. apply S_inj in hq. subst. inversion h1 as [m h3 q hq]. subst. apply S_inj in hq. destruct (zerop (pred m')) as [h4 | h4]. left; auto. right. assert (h5:0 < m'). omega. rewrite <- (S_pred _ _ h5). apply pred_inj' in hq; auto. Qed. (*Wyckoff*) Lemma ens_S_undoes_pos_ens_pred : forall (E:Ensemble nat), ~In E 0 -> ens_S (ens_pred E) = E. intros E h1. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2 as [x h2]. subst. destruct h2 as [x h2]. subst. assert (h3:0 < x). destruct (zerop x) as [h3 | h3]. subst. contradiction. assumption. rewrite <- (S_pred _ _ h3). assumption. unfold ens_S, ens_pred. rewrite im_im. assert (h3:0 < x). destruct (zerop x) as [h3 | h3]. subst. contradiction. assumption. apply Im_intro with x. assumption. rewrite <- (S_pred _ _ h3); auto. Qed. (*Wyckoff*) Lemma ens_S_undoes_zero_ens_pred1 : forall (E:Ensemble nat), In E 0 -> ~In E 1 -> ens_S (ens_pred E) = Subtract (Add E 1) 0. intros E h1 h2. unfold ens_S, ens_pred. rewrite im_im. apply Extensionality_Ensembles; red; split; red; intros n h3. destruct h3 as [n h3]. subst. constructor. destruct n as [|n]. simpl. right. constructor. simpl. left. assumption. intro h4. inversion h4. destruct h3 as [h3 h4]. destruct h3 as [n h3 | n h3]. apply Im_intro with n. assumption. erewrite <- S_pred; auto. destruct (zerop n) as [h5 | h5]. subst. contradict h4. constructor. apply h5. destruct h3. apply Im_intro with 0; auto. Qed. (*Wyckoff*) Lemma ens_pred01_eq : forall (E:Ensemble nat), In E 0 -> In E 1 -> ens_pred E = ens_pred (Subtract E 0). intros E h0 h1. unfold ens_pred. pose proof (add_sub_compat_in _ _ h0) as h2. rewrite <- h2 at 1. rewrite Im_add. pose proof (incl_im_subtract E 0 pred) as h3. simpl. rewrite in_add_eq; auto. apply Im_intro with 1. constructor; auto. intro h4. inversion h4. simpl. reflexivity. Qed. (*Wyckoff*) Lemma ens_pred_sub_0n1 : forall (E:Ensemble nat), In E 0 -> ~In E 1 -> ens_pred (Subtract E 0) = Subtract (ens_pred E) 0. intros E h1 h2. apply Extensionality_Ensembles; red; split; red; intros m h3. destruct h3 as [m h3]. subst. destruct h3 as [h3 h4]. rewrite in_sing_iff in h4. constructor. apply Im_intro with m; auto. rewrite in_sing_iff. destruct m as [|m]. contradict h4; auto. simpl. intro h5. subst. contradiction. destruct h3 as [h3 h4]. rewrite in_sing_iff in h4. destruct h3 as [m h3]. subst. apply Im_intro with m. constructor; auto. rewrite in_sing_iff. intro h5. subst. simpl in h4. contradict h4. reflexivity. reflexivity. Qed. (*Wyckoff*) Lemma in_ens_pred_impl : forall (A:Ensemble nat) (n:nat), In (ens_pred A) n -> {n = 0} + {In A (S n)}. intros A n h1. destruct n as [|n]. left; auto. right. inversion h1 as [x h2 q hq]. subst. assert (h3:0 < x). omega. pose proof (f_equal S hq) as h4. rewrite <- (S_pred _ _ h3) in h4. subst. assumption. Qed. (*Wyckoff*) Lemma ens_S_almost_undoes_ens_pred_zero_01 : forall (E:Ensemble nat), In E 0 -> In E 1 -> ens_S (ens_pred E) = Subtract E 0. intros E h0 h1. rewrite ens_pred01_eq; auto. unfold ens_S, ens_pred. rewrite im_im. rewrite im_id. apply im_ext_in. intros x h2. destruct h2 as [h2 h3]. erewrite <- S_pred; auto. destruct (zerop x) as [h4 | h4]. subst. contradict h3. constructor. apply h4. Qed. (*Wyckoff*) Definition ens_S_inj : forall (E F:Ensemble nat), ens_S E = ens_S F -> E = F. intros E F h1. unfold ens_S in h1. apply Extensionality_Ensembles; red; split; red; intros n h2. assert (h3:In (Im E S) (S n)). apply Im_intro with n; auto. rewrite h1 in h3. inversion h3 as [m h4 h5 h6]. subst. apply S_inj in h6. subst. assumption. assert (h3:In (Im F S) (S n)). apply Im_intro with n; auto. rewrite <- h1 in h3. inversion h3 as [m h5 h6 h7]. subst. apply S_inj in h7. subst. assumption. Qed. (*Wyckoff*) Lemma intersection_ens_S : forall (E F:Ensemble nat), Intersection (ens_S E) (ens_S F) = ens_S (Intersection E F). intros E F. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1 h2]. destruct h1 as [x h1], h2 as [x' h2 ? h3]. subst. apply S_inj in h3. subst. apply Im_intro with x'. constructor; auto. reflexivity. destruct h1 as [x h1]. subst. destruct h1 as [x h1 h2]. constructor. apply Im_intro with x; auto. apply Im_intro with x; auto. Qed. (*Wyckoff*) Lemma in_ens_S : forall (E:Ensemble nat) (m:nat), In E m -> In (ens_S E) (S m). intros E m h1. apply Im_intro with m; auto. Qed. (*Wyckoff*) Lemma in_ens_pred : forall (E:Ensemble nat) (m:nat), In E (S m) -> In (ens_pred E) m. intros E m h1. apply Im_intro with (S m); auto. Qed. (*Wyckoff*) Lemma in_ens_pred_iff : forall (E:Ensemble nat) (m:nat), In (ens_pred E) m <-> (m = 0 /\ (In E 0 \/ In E 1)) \/ In E (S m). intros E m; split; intro h1. destruct h1 as [m h1]; subst. destruct m as [|m]. left; tauto. simpl; right; auto. destruct h1 as [h1 | h1]. destruct h1 as [h1 h2]; subst. destruct h2 as [h2 | h2]. econstructor. apply h2. reflexivity. econstructor. apply h2. reflexivity. econstructor. apply h1. reflexivity. Qed. (*Wyckoff*) Lemma in_ens_pred0 : forall (E:Ensemble nat), In E 0 -> In (ens_pred E) 0. intros E h1. apply Im_intro with 0; auto. Qed. (*Wyckoff*) Lemma ens_pred_pos_inj : forall (E F:Ensemble nat), ens_pred E = ens_pred F -> ~In E 0 -> ~In F 0 -> E = F. assert (h3:forall E F:Ensemble nat, ens_pred E = ens_pred F -> ~In E 0 -> ~In F 0 -> Included E F). intros E F h1 h2 h0. red; intros x h3. assert (h4:x <> 0). intro. subst. contradiction. assert (h4':0 < x). omega. rewrite (S_pred _ _ h4') in h3. pose proof (in_ens_pred _ _ h3) as h5. rewrite h1 in h5. inversion h5 as [x' h6 q hq]. subst. clear h5. apply pred_inj0 in hq; auto. subst. assumption. destruct x as [|x]. simpl in hq. contradict h4. reflexivity. destruct x' as [|x']. simpl in hq. subst. simpl in h3. contradiction. simpl in hq. subst. auto with arith. intros E F h1 h2 h0. apply Extensionality_Ensembles; red; split; red; apply h3; auto. Qed. (*Wyckoff*) Lemma nin_0_ens_S : forall (E:Ensemble nat), ~In (ens_S E) 0. intros E h1. inversion h1 as [y h2 q h3 hq]. subst. discriminate h3. Qed. (*Wyckoff*) Definition fam_S (F:Family nat) := Im F ens_S. (*Wyckoff*) Lemma finite_fam_S : forall (F:Family nat) (pf:Finite F), Finite (fam_S F). intros F h1. apply finite_image; auto. Qed. (*Wyckoff*) Lemma in_fam_S_pos : forall (F:Family nat) (A:Ensemble nat), In (fam_S F) A -> ~In A 0. intros F A h1 h2. destruct h1 as [A h1]. subst. inversion h2 as [x h3 q hq]. discriminate hq. Qed. (*Wyckoff*) Lemma impl_in_fam_S : forall (F:Family nat) (E:Ensemble nat), In F E -> In (fam_S F) (ens_S E). unfold fam_S. intros F E h1. eapply Im_intro with E; auto. Qed. (*Wyckoff*) Lemma in_fam_S_impl : forall (F:Family nat) (E:Ensemble nat), In (fam_S F) E -> In F (ens_pred E). intros F E h1. destruct h1 as [E h1]. subst. rewrite ens_pred_undoes_ens_S. assumption. Qed. (*Wyckoff*) Lemma in_fam_S_nat_pos_r : forall (F:Family nat) (E:Ensemble nat), ~In E 0 -> In F (ens_pred E) -> In (fam_S F) E. intros F E h1 h2. unfold fam_S. apply Im_intro with (ens_pred E); auto. rewrite ens_S_undoes_pos_ens_pred; auto. Qed. (*Wyckoff*) Lemma in_fam_S_nat0n1 : forall (F:Family nat) (E:Ensemble nat), In E 0 -> ~In E 1 -> In F (ens_pred E) -> In (fam_S F) (Subtract (Add E 1) 0). intros F E h1 h2 h3. unfold fam_S. apply Im_intro with (ens_pred E); auto. rewrite ens_S_undoes_zero_ens_pred1; auto. Qed. (*Wyckoff*) Lemma in_fam_ens_nat01 : forall (F:Family nat) (E:Ensemble nat), In E 0 -> In E 1 -> In F (ens_pred E) -> In (fam_S F) (Subtract E 0). intros F E h1 h2 h3. unfold fam_S. apply Im_intro with (ens_pred E); auto. rewrite ens_S_almost_undoes_ens_pred_zero_01; auto. Qed. End Finite_Families. (*Use these when the underlying type is not really thought of as the domain.*) Section FunctionsOn. (*Wyckoff*) Definition fun_dom_ran {T U:Type} (A:Ensemble T) (B:Ensemble U) (f:T->U) : Prop := forall x:T, In A x -> In B (f x). (*Wyckoff*) Definition inj_on {T U:Type} (A:Ensemble T) (f:T->U) : Prop := forall x y:T, In A x -> In A y -> f x = f y -> x = y. (*Wyckoff*) Definition surj_on {T U:Type} (A:Ensemble T) (B:Ensemble U) (f:T->U) : Prop := Included B (Im A f). (*Wyckoff*) Definition bij_on {T U:Type} (A:Ensemble T) (B:Ensemble U) (f:T->U) : Prop := inj_on A f /\ surj_on A B f. (*Wyckoff*) Definition fun_dom_ran_empty1 : forall {T U:Type} (B:Ensemble U) (f:T->U), fun_dom_ran (Empty_set _) B f. intros T U B f. red. intros; contradiction. Qed. (*Wyckoff*) Definition fun_dom_ran_empty2_iff : forall {T U:Type} (A:Ensemble T) (f:T->U), fun_dom_ran A (Empty_set _) f <-> A = Empty_set _. intros T U A f. split. intro h1. red in h1. apply Extensionality_Ensembles; red; split; red; intros x h2. specialize (h1 _ h2). contradiction. contradiction. intros; subst. apply fun_dom_ran_empty1. Qed. (*Wyckoff*) Lemma inj_on_empty : forall {T U:Type} (f:T->U), inj_on (Empty_set _) f. intros; red; intros; contradiction. Qed. (*Wyckoff*) Lemma surj_on_empty1 : forall {T U:Type} (f:T->U) (B:Ensemble U), surj_on (Empty_set _) B f -> B = Empty_set _. intros T U f B h1. red in h1. rewrite image_empty in h1. auto with sets. Qed. (*Wyckoff*) Lemma surj_on_im : forall {T U:Type} (f:T->U) (A:Ensemble T), surj_on A (Im A f) f. intros; red; red; intros; auto. Qed. (*Wyckoff*) Lemma surj_on_empty2 : forall {T U:Type} (f:T->U) (A:Ensemble T), surj_on A (Empty_set _) f. intros; red; intros; auto with sets. Qed. (*Wyckoff*) Lemma fun_dom_ran_incl1 : forall {T U:Type} (A A':Ensemble T) (B:Ensemble U) (f:T->U), fun_dom_ran A' B f -> Included A A'-> fun_dom_ran A B f. intros T U A A' B' f h1 h2. red in h1. red. intros x h3. apply (h1 _ (h2 _ h3)). Qed. (*Wyckoff*) Lemma fun_dom_ran_incl2 : forall {T U:Type} (A:Ensemble T) (B B':Ensemble U) (f:T->U), fun_dom_ran A B f -> Included B B' -> fun_dom_ran A B' f. intros T U A B B' f h1 h2. red in h1. red. intros x h3. specialize (h1 _ h3). apply h2; auto. Qed. (*Wyckoff*) Lemma fun_dom_ran_add_inj : forall {T U:Type} (A:Ensemble T) (a:T) (B:Ensemble U) (f:T->U), ~In A a -> fun_dom_ran (Add A a) B f -> inj_on (Add A a) f -> fun_dom_ran A (Subtract B (f a)) f. intros T U A a B f h1 h2 h0. red in h2. red. intros x h3. assert (h4:In (Add A a) x). auto with sets. specialize (h2 _ h4). constructor; auto. intro h5. inversion h5 as [h6]. clear h5. red in h0. pose proof (Add_intro2 _ A a) as h7. specialize (h0 _ _ h7 h4 h6). subst. contradiction. Qed. (*Wyckoff*) Lemma inj_on_add : forall {T U:Type} (A:Ensemble T) (a:T) (f:T->U), inj_on A f -> ~In (Im A f) (f a) -> inj_on (Add A a) f. intros T U A a f h1 h2. red in h1. red. intros x y h3 h4 h5. destruct h3 as [x h3|x h3]. destruct h4 as [y h4 | y h4]; subst. apply h1 in h5; auto. destruct h4. rewrite <- h5 in h2. contradict h2. apply Im_intro with x; auto. destruct h3. destruct h4 as [y h4 | y h4]. rewrite h5 in h2. contradict h2. apply Im_intro with y; auto. destruct h4. reflexivity. Qed. (*Wyckoff*) Lemma inj_on_incl : forall {T U:Type} (A A':Ensemble T) (f:T->U), Included A A' -> inj_on A' f -> inj_on A f. intros; red; intros; auto. Qed. (*Wyckoff*) Lemma surj_on_add : forall {T U:Type} (A:Ensemble T) (a:T) (B:Ensemble U) (f:T->U), ~In A a -> surj_on (Add A a) B f -> surj_on A (Subtract B (f a)) f. intros T U A a B f h1 h2. red in h2. red. red in h2. red. intros x h3. destruct h3 as [h3 h4]. specialize (h2 _ h3). destruct h2 as [x h2]. subst. destruct h2 as [x h2 | x h2]. apply Im_intro with x; auto. destruct h2. contradict h4. constructor. Qed. (*Wyckoff*) Lemma finite_inj_surj_on_impl : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (f:T->U), Finite A -> inj_on A f -> surj_on A B f -> fun_dom_ran A B f -> Im A f = B. intros T U A B f h1. revert B. induction h1 as [|A h4 h5 c h6]. intros B h1 h2 h0. apply surj_on_empty1 in h2. subst. rewrite image_empty; auto. intros B h1 h2 h0. rewrite Im_add. assert (h7:Included A (Add A c)). auto with sets. pose proof (inj_on_incl A _ f h7 h1) as h8. pose proof h2 as h2'. apply surj_on_add in h2; auto. pose proof h0 as h0'. apply fun_dom_ran_add_inj in h0'; auto. specialize (h5 _ h8 h2 h0'). rewrite h5. rewrite add_sub_compat_in; auto. red in h0. apply h0. right. constructor. Qed. End FunctionsOn. Section Extends. (*Wyckoff*) Definition extends_sig {T U:Type} {A B:Ensemble T} (f:sig_set B->U) (g:sig_set A->U) := exists (pf:Included A B), forall x:sig_set A, g x = f (exist _ _ (pf _ (proj2_sig x))). (*Wyckoff*) Lemma refl_extends_sig : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), extends_sig f f. intros T U A f. red. exists (inclusion_reflexive _ ). intro x. destruct x. simpl. f_equal. apply proj1_sig_injective. simpl. reflexivity. Qed. (*Wyckoff*) Definition extends_sig1 {T U:Type} {A:Ensemble T} (f:T->U) (g:sig_set A->U) : Prop := forall x:sig_set A, g x = f (proj1_sig x). (*Wyckoff*) (*[prim] fefers to the primitive type of the domain, as opposed to assuming a priori that the extended function's domain is a sigma type.*) Definition extends_prim {T T' U:Type} (f:T->U) (g:T'->U) : Prop := exists A:Ensemble T, T' = sig_set A /\ forall (x:T) (x':T') (pf:In A x), existT id _ x' = existT id _ (exist _ _ pf) -> f x = g x'. (*Wyckoff*) Lemma extends_sig1_impl_extends_prim : forall {T U:Type} {A:Ensemble T} (f:T->U) (g:sig_set A->U), extends_sig1 f g -> extends_prim f g. intros T U A f g h1. red in h1. red. exists A. split; auto. intros x x' h2 h3. specialize (h1 (exist _ _ h2)). simpl in h1. rewrite <- h1. f_equal. apply inj_pair2 in h3. symmetry. assumption. Qed. (*See TypeUtilities for the converse of the above*) (*Wyckoff*) Definition extends_sig_prop {T U:Type} {P P':T->Prop} (f:{x:T | P' x} ->U) (g:{x:T | P x}->U) := exists (pf:forall x:T, P x -> P' x), forall x:{t:T | P t}, g x = f (exist _ _ (pf _ (proj2_sig x))). (*Search past in_im_proj1_sig_iff to get another "extends" function, [extends_sig_sig].*) (*Wyckoff*) Definition common_extension2 {T1 T2 T U:Type} (f:T->U) (g1:T1->U) (g2:T2->U) : Prop := extends_prim f g1 /\ extends_prim f g2. (*Wyckoff*) Definition fam_fun_one_ran {U:Type} := Ensemble {T:Type & T->U}. (*Wyckoff*) Definition common_extension_fam {T U:Type} (f:T->U) (F:Ensemble {V:Type & V->U}) := forall g, In F g -> extends_prim f (projT2 g). (*Wyckoff*) Lemma common_extension2_fam_compat_iff : forall {T1 T2 T U:Type} (f:T->U) (g1:T1->U) (g2:T2->U), common_extension2 f g1 g2 <-> common_extension_fam f (Couple (existT _ _ g1) (existT _ _ g2)). intros T1 T2 T U f g1 g2. split. intro h1. red in h1. red. intros g h2. destruct h1 as [h1a h1b]. destruct h2; auto. intro h2. red in h2. red. pose proof (h2 (existT _ _ g1) (Couple_l _ _ _)) as h3. pose proof (h2 (existT _ _ g2) (Couple_r _ _ _)) as h4. simpl in h3. simpl in h4. split; auto. Qed. (*Wyckoff*) Definition directed_fun_fam {U:Type} (F:Ensemble {V:Type & V->U}) : Prop := forall f g, In F f -> In F g -> exists k, In F k /\ common_extension2 (projT2 k) (projT2 f) (projT2 g). (*Wyckoff*) Definition agree_on {T U:Type} (f g:T->U) (A:Ensemble T) : Prop := forall x:T, In A x -> f x = g x. (*Wyckoff*) Definition restriction_fun {T U:Type} (f:T->U) (A:Ensemble T) : sig_set A -> U := fun x => f (proj1_sig x). (*Wyckoff*) Definition restriction_fun2 {T U:Type} (f:T->T->U) (A:Ensemble T) : sig_set A -> sig_set A -> U := fun x y => f (proj1_sig x) (proj1_sig y). (*Wyckoff*) Definition restriction_sig {T U:Type} {A:Ensemble T} (f:sig_set A->U) (B:Ensemble T) (pf:Included B A) : sig_set B->U := fun x=>f (exist _ _ (pf _ (proj2_sig x))). (*Wyckoff*) Definition restriction_sig2 {T U:Type} {A:Ensemble T} (f:sig_set A->sig_set A->U) (B:Ensemble T) (pf:Included B A) : sig_set B->sig_set B->U := fun x y=>f (exist _ _ (pf _ (proj2_sig x))) (exist _ _ (pf _ (proj2_sig y))). (*Wyckoff*) Definition closed_fun {T:Type} {A B:Ensemble T} (f:sig_set A->sig_set A) (pf:Included B A) : Prop := forall (x:T) (pfi:In B x), In B (proj1_sig (f (exist _ _ (pf _ pfi)))). (*Wyckoff*) Definition closed_fun2 {T:Type} {A B:Ensemble T} (f:sig_set A->sig_set A->sig_set A) (pf:Included B A) : Prop := forall (x y:T) (pfx:In B x) (pfy:In B y), In B (proj1_sig (f (exist _ _ (pf _ pfx)) (exist _ _ (pf _ pfy)))). (*Wyckoff*) Definition restriction_sig' {T:Type} {A:Ensemble T} (f:sig_set A->sig_set A) (B:Ensemble T) (pf:Included B A) (pfcl:closed_fun f pf) : sig_set B->sig_set B := fun x=>exist _ _ (pfcl _ (proj2_sig x)). (*Wyckoff*) Definition restriction_sig2' {T:Type} {A:Ensemble T} (f:sig_set A->sig_set A->sig_set A) (B:Ensemble T) (pf:Included B A) (pfcl:closed_fun2 f pf) : sig_set B->sig_set B->sig_set B := fun x y=>exist _ _ (pfcl _ _ (proj2_sig x) (proj2_sig y)). Definition same_im_subsets_sig {T U:Type} {A:Ensemble T} (f:sig_set A->U) := [B:Ensemble T | exists pf:Included B A, Im (full_sig A) f = Im (full_sig B) (restriction_sig f B pf)]. Lemma extends_sig_iff : forall {T U:Type} {A B:Ensemble T} (f:sig_set B->U) (g:sig_set A->U), extends_sig f g <-> exists (pf:Included A B), restriction_sig f _ pf = g. intros T U A B f g. split. intro h1. red in h1. destruct h1 as [h1 h2]. exists h1. apply functional_extensionality. intro x. specialize (h2 x). rewrite h2. unfold restriction_sig. reflexivity. intro h1. destruct h1 as [h1 h2]. red. exists h1. intro x. subst. unfold restriction_sig. reflexivity. Qed. (*Wyckoff*) Lemma im_full_sig_restriction : forall {T U:Type} (E:Ensemble T) (p:T->U), Im (full_sig E) (restriction_fun p E) = Im E p. intros T U E p. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]; subst. apply Im_intro with (proj1_sig x). apply proj2_sig. unfold restriction_fun. reflexivity. destruct h1 as [x h1]; subst. apply Im_intro with (exist _ _ h1). constructor. unfold restriction_fun; auto. Qed. (*Wyckoff*) Lemma im_full_sig_union_eq : forall {T U:Type} {A B:Ensemble T} (f:sig_set (Union A B)->U), Im (full_sig (Union A B)) f = Union (Im (full_sig A) (restriction_sig f _ (Union_increases_l _ A B))) (Im (full_sig B) (restriction_sig f _ (Union_increases_r _ A B))). intros T U A B f. unfold restriction_sig. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]. subst. clear h1. destruct x as [x h1]. destruct h1 as [x h1 | x h1]. left. apply Im_intro with (exist _ _ h1). constructor. f_equal. apply proj1_sig_injective; auto. right. apply Im_intro with (exist _ _ h1). constructor. f_equal. apply proj1_sig_injective; auto. destruct h1 as [x h1 | x h1]. destruct h1 as [x h1]. subst. clear h1. destruct x as [x h1]. assert (h2:In (Union A B) x). left; auto with sets. apply Im_intro with (exist _ _ h2). constructor. f_equal. apply proj1_sig_injective; auto. destruct h1 as [x h1]. subst. clear h1. destruct x as [x h1]. assert (h2:In (Union A B) x). right; auto with sets. apply Im_intro with (exist _ _ h2). constructor. f_equal. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma im_full_sig_union_eq' : forall {T U:Type} {A B C:Ensemble T} (f:sig_set A->U) (pfb:Included B A) (pfc:Included C A), A = Union B C -> Im (full_sig A) f = Union (Im (full_sig B) (restriction_sig f _ pfb)) (Im (full_sig C) (restriction_sig f _ pfc)). intros T U A B C f h1 h2 h3. subst. rewrite im_full_sig_union_eq. unfold restriction_sig. f_equal. apply im_ext_in; intros; simpl; f_equal. apply proj1_sig_injective; auto. apply im_ext_in; intros; simpl; f_equal. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Definition restriction_sig_prop {T U:Type} (P P':T->Prop) (pf:forall x:T, P x -> P' x) (f:{t:T | P' t} -> U) : {t:T | P t}->U. intros x. destruct x as [x h1]. apply pf in h1. refine (f (exist _ _ h1)). Defined. (*Wyckoff*) Lemma extends_restriction_sig_prop : forall {T U:Type} (P P':T->Prop) (pf:forall x:T, P x -> P' x) (f:{t:T | P' t} -> U), extends_sig_prop f (restriction_sig_prop P P' pf f). intros T U P P' h1 f. red. exists h1. intro x. unfold restriction_sig_prop. destruct x as [x h2]. simpl. reflexivity. Qed. (*Wyckoff*) Lemma im_full_sig_add : forall {T U:Type} {A:Ensemble T} {a:T} (f:sig_set (Add A a)->U), Im (full_sig (Add A a)) f = Add (Im (full_sig A) (restriction_sig f _ (incl_add A a))) (f (exist _ _ (Add_intro2 _ A a))). intros T U A a f. apply Extensionality_Ensembles. red. split. red. intros y h1. destruct h1 as [y h1]. subst. clear h1. destruct y as [y h1]. destruct h1 as [y h1 | y h2]. left. apply Im_intro with (exist _ _ h1). constructor. unfold restriction_sig. f_equal. apply proj1_sig_injective. simpl. reflexivity. destruct h2. right. assert (h1: f (exist (In (Add A a)) a (Add_intro2 T A a)) = f (exist (fun x : T => In (Add A a) x) a (Union_intror T A (Singleton a) a (In_singleton T a)))). f_equal. apply proj1_sig_injective. simpl. reflexivity. rewrite h1. constructor. red. intros y h1. destruct h1 as [y h1 | y h2]. destruct h1 as [x h1]. subst. clear h1. destruct x as [x h1]. apply Im_intro with (exist _ _ (Add_intro1 _ A a x h1)). constructor. unfold restriction_sig. f_equal. apply proj1_sig_injective. simpl. reflexivity. destruct h2. apply Im_intro with (exist (In (Add A a)) a (Add_intro2 T A a)); auto. constructor. Qed. (*Wyckoff*) Lemma setminus_im_full_sig_incl_setminus : forall {T U:Type} (A B:Ensemble T) (pf:Included B A) (f:sig_set A->U), Included (Setminus (Im (full_sig A) f) (Im (full_sig B) (restriction_sig f _ pf))) (Im (full_sig (Setminus A B)) (restriction_sig f _ (setminus_inc A B))). intros T U A B h1 f. red. intros y h2. destruct h2 as [h2 h3]. destruct h2 as [y h2 u]. subst. clear h2. destruct y as [y h4]. assert (h5:~In B y). intro h5. contradict h3. apply Im_intro with (exist _ _ h5). constructor. unfold restriction_sig. simpl. f_equal. apply proj1_sig_injective. simpl. reflexivity. assert (h6:In (Setminus A B) y). constructor; auto. apply Im_intro with (exist _ _ h6). constructor. unfold restriction_sig. simpl. f_equal. apply proj1_sig_injective. simpl. reflexivity. Qed. End Extends. (*This Section defines the Von Neumann notion of a natural number as the set of its predecessors ([seg]) and one that counts from one ([seg_one]) (the latter so as to give more flexibility to those utilizing a back-and -forth argument, where you don't have complete control of the terms.*) Section SegSet. (*Wyckoff*) Definition seg (n:nat) := [m:nat | m < n]. (*Wyckoff*) Definition segT (n:nat) := {m:nat | m < n}. (*Wyckoff*) Inductive seg_ens (n:nat) : Ensemble nat -> Prop := | seg_ens_intro : forall (E:Ensemble nat), Included E (seg n) -> seg_ens n E. End SegSet. Notation "n @" := (segT n) (at level 30, right associativity). Section SegSet'. (*Wyckoff*) Definition seg_weak (n:nat) := [m:nat | m <= n]. (*Wyckoff*) Definition seg_one (n:nat) := [m:nat | 1 <= m /\ m <= n]. (*Wyckoff*) Definition seg_one_t (n:nat) := {m:nat | 1 <= m /\ m <= n}. (*Wyckoff*) Lemma segT_eq_dec : forall (n:nat), type_eq_dec (n @). intros n. red. intros a b. destruct a as [a ha], b as [b h]. destruct (nat_eq_dec a b) as [h1 | h1]; subst. left. f_equal. apply proof_irrelevance. right. intro h2. apply exist_injective in h2. contradiction. Qed. (*Wyckoff*) Lemma in_st : forall {n:nat} (m:n @), In (seg n) (proj1_sig m). intros n m. constructor. apply proj2_sig. Qed. (*Wyckoff*) Lemma in_seg : forall (m n:nat), m < n -> In (seg n) m. intros; constructor; auto. Qed. (*Wyckoff*) Lemma in_pred_seg : forall (n:nat), 0 < n -> Inn (seg n) (pred n). intro; constructor; auto with arith. Qed. (*Wyckoff*) Lemma seg_in : forall (m n:nat), In (seg n) m -> m < n. intros ? ? h1; destruct h1; auto. Qed. (*Wyckoff*) Lemma in_sot : forall {n:nat} (m:seg_one_t n), In (seg_one n) (proj1_sig m). intros n m. constructor. apply proj2_sig. Qed. (*Wyckoff*) Lemma seg_one_t_compat : forall (n:nat), im_proj1_sig (Full_set (seg_one_t n)) = seg_one n. intro n. apply Extensionality_Ensembles; red; split; red; intros m h1. destruct h1 as [m h1]. subst. constructor. apply proj2_sig. destruct h1 as [h1]. unfold im_proj1_sig. apply Im_intro with (exist (fun r => 1 <= r <= n) _ h1). constructor. auto. Qed. (*Wyckoff*) Lemma segT_compat : forall (n:nat), im_proj1_sig (Full_set (n @)) = seg n. intro n. apply Extensionality_Ensembles; red; split; red; intros m h1. destruct h1 as [m h1]. subst. constructor. apply proj2_sig. destruct h1 as [h1]. unfold im_proj1_sig. apply Im_intro with (exist (fun r => r < n) _ h1). constructor. auto. Qed. (*Wyckoff*) Lemma seg_one_t_lt : forall {n:nat} (m:seg_one_t n), pred (proj1_sig m) < n. intros n m; destruct m; simpl; omega. Qed. (*Wyckoff*) Lemma seg_one_t_le_le : forall {n:nat} (m:seg_one_t n), 1 <= S (pred (proj1_sig m)) <= n. intros n m; destruct m; simpl; omega. Qed. (*Wyckoff*) Lemma segT_le_le : forall {n:nat} (m:n @), 1 <= S (proj1_sig m) <= n. intros ? m; destruct m; simpl; omega. Qed. (*Wyckoff*) Definition sot_to_st {n:nat} (m:seg_one_t n) : n @. refine (exist (fun c => c < n) _ _). apply (seg_one_t_lt m). Defined. (*Wyckoff*) Lemma sot_to_st_compat : forall {n:nat} (m:seg_one_t n), proj1_sig (sot_to_st m) = pred (proj1_sig m). intros n m. destruct m as [m h1]. unfold sot_to_st. destruct h1 as [h1 h2]; simpl; auto. Qed. (*Wyckoff*) Definition st_to_sot {n:nat} (m:n @) : seg_one_t n. destruct m as [m h1]. red. assert (h2:1 <= (S m) <= n). auto with arith. refine (exist (fun a => 1 <= a <= n) _ h2). Defined. (*Wyckoff*) Lemma st_to_sot_compat : forall {n:nat} (m:n @), proj1_sig (st_to_sot m) = S (proj1_sig m). intros n m. destruct m as [m h1]; simpl; auto. Qed. (*Wyckoff*) Definition st_to_sot_eq : forall {n:nat} (m:n @), let pf := iff1 (lt_sot_iff _ _) (seg_one_t_lt (st_to_sot m)) in st_to_sot m = exist (fun c => 1 <= c <= n) _ pf. intros n m; simpl. apply proj1_sig_injective; simpl. rewrite st_to_sot_compat; simpl; auto. Qed. (*Wyckoff*) Definition sot_to_st_eq : forall {n:nat} (m:seg_one_t n), let pf := seg_one_t_lt m in sot_to_st m = exist (fun c => c < n) _ pf. intros n m; simpl. apply proj1_sig_injective; simpl; auto. Qed. (*Wyckoff*) Definition st_to_sot_to_st : forall {n:nat} (m:seg_one_t n), st_to_sot (sot_to_st m) = m. unfold st_to_sot, sot_to_st. intros n m. destruct m as [m h2]; apply proj1_sig_injective; simpl; omega. Qed. (*Wyckoff*) Definition sot_to_st_to_sot : forall {n:nat} (m:n @), sot_to_st (st_to_sot m) = m. unfold st_to_sot, sot_to_st. intros n m. destruct m as [m h2]; apply proj1_sig_injective; simpl; omega. Qed. (*Wyckoff*) Lemma full_set_sot : forall (n:nat), Full_set (seg_one_t n) = Im (Full_set (n @)) (st_to_sot). intro n. apply Extensionality_Ensembles; auto with sets. red; split; red; intros x h1. apply Im_intro with (sot_to_st x). constructor. rewrite st_to_sot_to_st; auto. constructor. Qed. (*Wyckoff*) Lemma full_set_st : forall (n:nat), Full_set (n @) = Im (Full_set (seg_one_t n)) (sot_to_st). intro n. apply Extensionality_Ensembles; auto with sets. red; split; red; intros x h1. apply Im_intro with (st_to_sot x). constructor. rewrite sot_to_st_to_sot; auto. constructor. Qed. (*Wyckoff*) Definition st_conv {n:nat} (m:n @) : sig_set (seg n) := exist _ _ (in_st m). (*Wyckoff*) Definition st_conv_rev {n:nat} (m:sig_set (seg n)) : n @ := exist _ _ (match (proj2_sig m) with | intro_characteristic_sat P => P end). (*A conversion function for functions on [seg_sset_t] (st)*) (*Wyckoff*) Definition fun_st_conv {U:Type} {n:nat} (f:n @ -> U) : sig_set (seg n) -> U. intro x. refine (f _). refine (exist _ _ (match (proj2_sig x) with | intro_characteristic_sat P => P end)). Defined. (*Wyckoff*) Definition fun_st_conv_rev {U:Type} {n:nat} (f:sig_set (seg n) -> U) : n @ -> U. intro m. refine (f _). red. refine (exist _ _ (in_st m)). Defined. (*A conversion function for functions on [seg_sset_one_t] (sot)*) (*Wyckoff*) Definition fun_sot_conv {U:Type} {n:nat} (f:seg_one_t n -> U) : sig_set (seg_one n) -> U. intro x. refine (f _). refine (exist _ _ (match (proj2_sig x) with | intro_characteristic_sat P => P end)). Defined. (*Wyckoff*) Definition fun_sot_conv_rev {U:Type} {n:nat} (f:sig_set (seg_one n) -> U) : seg_one_t n -> U. intro m. refine (f _). red. refine (exist _ _ (in_sot m)). Defined. (*Wyckoff*) Definition ens_ss_conv {n:nat} (A:Ensemble (sig_set (seg n))) : Ensemble (n @) := Im A (fun x => exist _ _ (match (proj2_sig x) with | P => match P with | intro_characteristic_sat P' => P' end end)). (*Wyckoff*) Definition ens_sso_conv {n:nat} (A:Ensemble (sig_set (seg_one n))) : Ensemble (seg_one_t n) := Im A (fun x => exist _ _ (match (proj2_sig x) with | P => match P with | intro_characteristic_sat P' => P' end end)). (*Wyckoff*) Definition ens_st_to_sot {n:nat} (A:Ensemble (n @)) : Ensemble (seg_one_t n). refine (Im A _). intro m. refine (exist _ (S (proj1_sig m)) _). apply segT_le_le. Defined. (*Wyckoff*) Definition ens_sot_to_st {n:nat} (A:Ensemble (seg_one_t n)) : Ensemble (n @). refine (Im A _). intro m. refine (exist _ (pred (proj1_sig m)) _). apply seg_one_t_lt. Defined. (*Wyckoff*) Lemma ens_st_to_sot_to_st : forall {n:nat} (A:Ensemble (seg_one_t n)), ens_st_to_sot (ens_sot_to_st A) = A. intros n A. unfold ens_st_to_sot, ens_sot_to_st. rewrite im_im. rewrite im_id. apply im_ext_in. intros m h1; simpl. apply proj1_sig_injective; simpl; auto. destruct m as [m hm]; simpl. omega. Qed. (*Wyckoff*) Lemma ens_sot_to_st_to_sot : forall {n:nat} (A:Ensemble (n @)), ens_sot_to_st (ens_st_to_sot A) = A. intros n A. unfold ens_st_to_sot, ens_sot_to_st. rewrite im_im. rewrite im_id. apply im_ext_in. intros m h1; simpl. apply proj1_sig_injective; simpl; auto. Qed. (*Wyckoff*) Lemma ens_sot_to_st_im : forall {n:nat} (A:Ensemble (seg_one_t n)), A = Im (ens_sot_to_st A) st_to_sot. intros n A. unfold ens_sot_to_st, st_to_sot. rewrite im_im. rewrite im_id at 1. apply im_ext_in. intros m h1. apply proj1_sig_injective; simpl; auto. destruct m; simpl; omega. Qed. (*Wyckoff*) Lemma ens_st_to_sot_im : forall {n:nat} (A:Ensemble (n @)), A = Im (ens_st_to_sot A) sot_to_st. intros n A. unfold ens_st_to_sot, sot_to_st. rewrite im_im. rewrite im_id at 1. apply im_ext_in. intros m h1. apply proj1_sig_injective; simpl; auto. Qed. (*Wyckoff*) Lemma im_ens_sot : forall {T:Type} {n:nat} (A:Ensemble (seg_one_t n)) (f:seg_one_t n-> T), Im A f = Im (ens_sot_to_st A) (compose f st_to_sot). intros T n A f. unfold ens_sot_to_st. rewrite im_im. apply im_ext_in. intros m h1. unfold compose, st_to_sot. f_equal. apply proj1_sig_injective; simpl. destruct m; simpl; omega. Qed. (*Wyckoff*) Lemma im_ens_st : forall {T:Type} {n:nat} (A:Ensemble (n @)) (f:n @-> T), Im A f = Im (ens_st_to_sot A) (compose f sot_to_st). intros T n A f. unfold ens_st_to_sot. rewrite im_im. apply im_ext_in. intros m h1. unfold compose, sot_to_st. f_equal. apply proj1_sig_injective; simpl; auto. Qed. (*Wyckoff*) Lemma ens_st_to_sot_full_set : forall (n:nat), ens_st_to_sot (Full_set (n @)) = Full_set (seg_one_t n). intro n. unfold ens_st_to_sot. apply Extensionality_Ensembles; auto with sets. red; split; red; intros m h1. constructor. apply Im_intro with (sot_to_st m). constructor. apply proj1_sig_injective; simpl. destruct m; simpl; omega. Qed. (*Wyckoff*) Lemma ens_sot_to_st_full_set : forall (n:nat), ens_sot_to_st (Full_set (seg_one_t n)) = Full_set (n @). intro n. unfold ens_sot_to_st. apply Extensionality_Ensembles; auto with sets. red; split; red; intros m h1. constructor. apply Im_intro with (st_to_sot m). constructor. apply proj1_sig_injective; simpl; auto. destruct m; simpl; omega. Qed. (*Wyckoff*) Definition segT_fun_to_nat_fun {T:Type} {n:nat} (f:n @->T) (def:T) := (fun m:nat => match (lt_le_dec m n) with | left P => f (exist (fun a => a < n) _ P) | _ => def end). (*Wyckoff*) Lemma segT_fun_to_nat_fun_compat : forall {T:Type} {n:nat} (f:n @->T) (def:T) (m:nat) (pf:m < n), segT_fun_to_nat_fun f def m = f (exist (fun a => a < n) _ pf). intros T n f def m h1. unfold segT_fun_to_nat_fun. destruct lt_le_dec as [h2 | h2]. f_equal. f_equal. apply proof_irrelevance. omega. Qed. (*Wyckoff*) Definition segT_fun_to_seg_one {T:Type} {n:nat} (f:n @->T) : seg_one_t n -> T. intro m. destruct m as [m h1]. destruct m as [|m]. destruct h1 as [h1]. apply le0 in h1. discriminate. rewrite <- lt_sot_iff in h1. refine (f (exist (fun c => c < n) _ h1)). Defined. (*Wyckoff*) Lemma segT_fun_to_seg_one_compat : forall {T:Type} {n:nat} (f:n @->T) (m:nat) (pf':1 <= S m <= n), let pf := iff2 (lt_sot_iff _ _) pf' in let f' := segT_fun_to_seg_one f in f' (exist (fun c => 1 <= c <= n) _ pf') = f (exist (fun c => c < n) _ pf). simpl; intros; repeat f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Definition seg_one_t_fun_to_seg {T:Type} {n:nat} (f:seg_one_t n -> T) : n @->T. intro m. destruct m as [m h1]. rewrite lt_sot_iff in h1. refine (f (exist (fun c => 1 <= c <= n) _ h1)). Defined. (*Wyckoff*) Lemma seg_one_t_fun_to_seg_compat : forall {T:Type} {n:nat} (f:seg_one_t n->T) (m:nat) (pf:m < n), let pf' := iff1 (lt_sot_iff _ _) pf in let f' := seg_one_t_fun_to_seg f in f' (exist (fun c => c < n) _ pf) = f (exist (fun c => 1 <= c <= n) _ pf'). simpl; intros; repeat f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma inj_segT_fun_to_seg_one_iff : forall {T:Type} {n:nat} (f:n @->T), injective (segT_fun_to_seg_one f) <-> injective f. intros T n f; split; intro h1; red in h1; red; intros x y h2. destruct x as [x hx], y as [y hy]. pose proof hx as h3. pose proof hy as h4. rewrite lt_sot_iff in h3, h4. specialize (h1 (exist _ _ h3) (exist _ _ h4)). hfirst h1. do 2 rewrite segT_fun_to_seg_one_compat. assert (hx = iff2 (lt_sot_iff _ _) h3). apply proof_irrelevance. assert (hy = iff2 (lt_sot_iff _ _) h4). apply proof_irrelevance. subst; auto. specialize (h1 hf). apply proj1_sig_injective; auto; simpl. apply exist_injective in h1; auto. destruct x as [x hx], y as [y hy]. destruct x as [|x], y as [|y]; try omega. do 2 rewrite segT_fun_to_seg_one_compat in h2. apply h1 in h2. apply exist_injective in h2. subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma inj_seg_one_t_fun_to_seg_iff : forall {T:Type} {n:nat} (f:seg_one_t n->T), injective (seg_one_t_fun_to_seg f) <-> injective f. intros T n f; split; intro h1; red in h1; red; intros x y h2. destruct x as [x hx], y as [y hy]. pose proof hx as h3. pose proof hy as h4. destruct x as [|x], y as [|y]; try omega. rewrite <- lt_sot_iff in h3, h4. specialize (h1 (exist (fun c => c < n) _ h3) (exist (fun c => c < n) _ h4)). do 2 rewrite seg_one_t_fun_to_seg_compat in h1. hfirst h1. assert (h5:hx = (iff1 (lt_sot_iff _ _) h3)). apply proof_irrelevance. assert (h6:hy = (iff1 (lt_sot_iff _ _) h4)). apply proof_irrelevance. subst; auto. specialize (h1 hf). apply exist_injective in h1; subst; f_equal; apply proof_irrelevance. destruct x as [x hx], y as [y hy]. pose proof hx as h3. pose proof hy as h4. rewrite lt_sot_iff in h3, h4. specialize (h1 (exist _ _ h3) (exist _ _ h4)). do 2 rewrite seg_one_t_fun_to_seg_compat in h2. assert (h5:h3 = iff1 (lt_sot_iff _ _) hx). apply proof_irrelevance. assert (h6:h4 = iff1 (lt_sot_iff _ _) hy). apply proof_irrelevance. subst. specialize (h1 h2). apply exist_injective in h1; apply S_inj in h1; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma surj_segT_fun_to_seg_one_iff : forall {T:Type} {n:nat} (f:n @->T), surjective (segT_fun_to_seg_one f) <-> surjective f. intros T n f; split; intro h1; red; red in h1; intros y; specialize (h1 y); destruct h1 as [x h2]; subst. exists (sot_to_st x). destruct x as [x h1]; subst. destruct x as [|x]; try omega. rewrite segT_fun_to_seg_one_compat. f_equal. apply proj1_sig_injective; simpl. destruct h1; simpl; auto. exists (st_to_sot x). rewrite st_to_sot_eq. rewrite segT_fun_to_seg_one_compat. f_equal; apply proj1_sig_injective; simpl; auto. rewrite st_to_sot_compat; simpl; auto. Qed. (*Wyckoff*) Lemma surj_seg_one_t_fun_to_seg_iff : forall {T:Type} {n:nat} (f:seg_one_t n->T), surjective (seg_one_t_fun_to_seg f) <-> surjective f. intros T n f; split; intro h1; red; red in h1; intros y; specialize (h1 y); destruct h1 as [x h2]; subst. exists (st_to_sot x). destruct x as [x hx]. rewrite seg_one_t_fun_to_seg_compat. rewrite st_to_sot_eq. repeat f_equal; apply proof_irrelevance. exists (sot_to_st x). rewrite sot_to_st_eq. rewrite seg_one_t_fun_to_seg_compat. f_equal; apply proj1_sig_injective; simpl; auto. pose proof (seg_one_t_lt x) as h1. destruct x as [x [h2 h3]]. simpl in h1; simpl. omega. Qed. (*Wyckoff*) Lemma bij_segT_fun_to_seg_one_iff : forall {T:Type} {n:nat} (f:n @->T), bijective (segT_fun_to_seg_one f) <-> bijective f. intros T n f; split; intro h1; red; red in h1; destruct h1 as [h1 h2]; split. rewrite inj_segT_fun_to_seg_one_iff in h1; auto. rewrite surj_segT_fun_to_seg_one_iff in h2; auto. rewrite inj_segT_fun_to_seg_one_iff; auto. rewrite surj_segT_fun_to_seg_one_iff; auto. Qed. (*Wyckoff*) Lemma bij_seg_one_t_fun_to_seg_iff : forall {T:Type} {n:nat} (f:seg_one_t n->T), bijective (seg_one_t_fun_to_seg f) <-> bijective f. intros T n f; split; intro h1; red; red in h1; destruct h1 as [h1 h2]; split. rewrite inj_seg_one_t_fun_to_seg_iff in h1; auto. rewrite surj_seg_one_t_fun_to_seg_iff in h2; auto. rewrite inj_seg_one_t_fun_to_seg_iff; auto. rewrite surj_seg_one_t_fun_to_seg_iff; auto. Qed. (*Wyckoff*) Lemma inj_st_conv_iff : forall {U:Type} {n:nat} (f:n @->U), injective f <-> injective (fun_st_conv f). intros U n f. split. intro h1. red in h1. red. intros x y h2. destruct x as [x h3], y as [y h4]. unfold fun_st_conv in h2. specialize (h1 _ _ h2). apply exist_injective in h1. simpl in h1. subst. f_equal. apply proof_irrelevance. intro h1. red in h1. red. intros x y h2. unfold fun_st_conv in h1. destruct x as [x h3], y as [y h4]. assert (h5:In (seg n) x). constructor; auto. assert (h6:In (seg n) y). constructor; auto. specialize (h1 (exist _ _ h5) (exist _ _ h6)). apply proj1_sig_injective. simpl. simpl in h1. let P := type of h1 in match P with ?h1 -> _ => assert (h7:h1) end. destruct h5 as [h5], h6 as [h6]. pose proof (proof_irrelevance _ h3 h5). pose proof (proof_irrelevance _ h4 h6). subst; auto. specialize (h1 h7). apply exist_injective in h1; auto. Qed. (*Wyckoff*) Lemma inj_sot_conv_iff : forall {U:Type} {n:nat} (f:seg_one_t n->U), injective f <-> injective (fun_sot_conv f). intros U n f. split. intro h1. red in h1. red. intros x y h2. destruct x as [x h3], y as [y h4]. unfold fun_st_conv in h2. specialize (h1 _ _ h2). apply exist_injective in h1. simpl in h1. subst. f_equal. apply proof_irrelevance. intro h1. red in h1. red. intros x y h2. unfold fun_st_conv in h1. destruct x as [x h3], y as [y h4]. assert (h5:In (seg_one n) x). constructor; auto. assert (h6:In (seg_one n) y). constructor; auto. specialize (h1 (exist _ _ h5) (exist _ _ h6)). apply proj1_sig_injective. simpl. simpl in h1. let P := type of h1 in match P with ?h1 -> _ => assert (h7:h1) end. destruct h5 as [h5], h6 as [h6]. pose proof (proof_irrelevance _ h3 h5). pose proof (proof_irrelevance _ h4 h6). subst; auto. specialize (h1 h7). apply exist_injective in h1. assumption. Qed. (*Wyckoff*) Lemma surj_st_conv_iff : forall {U:Type} {n:nat} (f:n @->U), surjective f <-> surjective (fun_st_conv f). intros U n f. split. intro h1. red in h1. red. intro y. specialize (h1 y). destruct h1 as [x h1]. subst. assert (h1:In (seg n) (proj1_sig x)). constructor; apply proj2_sig. exists (exist _ _ h1). unfold fun_st_conv. f_equal. apply proj1_sig_injective. simpl. reflexivity. intro h1. red in h1. red. intro y. specialize (h1 y). destruct h1 as [x h1]; subst. destruct x as [x h1]. destruct h1 as [h1]. exists (exist (fun a => a < n) _ h1). unfold fun_st_conv. f_equal. Qed. (*Wyckoff*) Lemma surj_sot_conv_iff : forall {U:Type} {n:nat} (f:seg_one_t n->U), surjective f <-> surjective (fun_sot_conv f). intros U n f. split. intro h1. red in h1. intro y. specialize (h1 y). destruct h1 as [x h1]. subst. assert (h1:In (seg_one n) (proj1_sig x)). constructor; apply proj2_sig. exists (exist _ _ h1). unfold fun_sot_conv. f_equal. apply proj1_sig_injective. simpl. reflexivity. intro h1. red in h1. red. intro y. specialize (h1 y). destruct h1 as [x h1]; subst. destruct x as [x h1]. destruct h1 as [h1]. exists (exist (fun a => 1 <= a <= n) _ h1). unfold fun_sot_conv. f_equal. Qed. (*Wyckoff*) Lemma bij_st_conv_iff : forall {U:Type} {n:nat} (f:n @->U), bijective f <-> bijective (fun_st_conv f). intros U n f. split. intro h1. red in h1. red. destruct h1 as [h1 h2]. rewrite inj_st_conv_iff in h1. rewrite surj_st_conv_iff in h2. split; auto. intro h1. red in h1. red. destruct h1 as [h1 h2]. rewrite <- inj_st_conv_iff in h1. rewrite <- surj_st_conv_iff in h2. split; auto. Qed. (*Wyckoff*) Lemma bij_sot_conv_iff : forall {U:Type} {n:nat} (f:seg_one_t n->U), bijective f <-> bijective (fun_sot_conv f). intros U n f. split. intro h1. red in h1. red. destruct h1 as [h1 h2]. rewrite inj_sot_conv_iff in h1. rewrite surj_sot_conv_iff in h2. split; auto. intro h1. red in h1. red. destruct h1 as [h1 h2]. rewrite <- inj_sot_conv_iff in h1. rewrite <- surj_sot_conv_iff in h2. split; auto. Qed. (*Wyckoff*) Definition st_nest {n:nat} {m:n @} (p:(proj1_sig m) @) : n @. destruct m as [m h1]. simpl in p. destruct p as [p h2]. red. refine (exist (fun a => a < n)_ (lt_trans _ _ _ h2 h1)). Defined. (*Wyckoff*) Lemma proj1_sig_st_nest_eq : forall {n:nat} {m:n @} (p:(proj1_sig m) @), proj1_sig (st_nest p) = proj1_sig p. intros n m p. destruct p as [p h1]. simpl. unfold st_nest. destruct m as [m h2]. simpl. reflexivity. Qed. (*Wyckoff*) Definition sot_nest {n:nat} {m:seg_one_t n} (p:seg_one_t (proj1_sig m)) : seg_one_t n. destruct m as [m h1]. simpl in p. destruct p as [p h2]. red. destruct h1 as [h1a h1b], h2 as [h2a h2b]. refine (exist _ _ (conj h2a (le_trans _ _ _ h2b h1b))). Defined. (*Wyckoff*) Lemma segT_0_false : forall n:0 @, False. intro n. destruct n; omega. Qed. (*Wyckoff*) Definition segT_S {n:nat} (m:n @) : (S n) @. destruct m as [m h1]. red. refine (exist (fun a => a < S n) _ (lt_n_S _ _ h1)). Defined. (*Wyckoff*) Definition segT_S_compat : forall {n:nat} (m:n @), proj1_sig (segT_S m) = S (proj1_sig m). intros ? m; destruct m; auto. Qed. (*Wyckoff*) Definition segT_S_pred {n} (m:pred n @) : n @ := exist (fun c => c < n) _ (lt_pred' _ _ (proj2_sig m)). (*Wyckoff*) Lemma segT_S_pred_compat : forall {n} (m:pred n @), proj1_sig (segT_S_pred m) = proj1_sig m. auto. Qed. (*Wyckoff*) Definition segT_S_pred_to {n} (m:S (pred n) @) (pf:0 < n) : n @ := exist (fun c => c < n) (proj1_sig m) (lt_eq_trans _ _ _ (proj2_sig m) (eq_sym (S_pred _ _ pf))). (*Wyckoff*) Lemma segT_S_pred_to_compat : forall {n} (m:S (pred n) @) (pf:0 < n), proj1_sig (segT_S_pred_to m pf) = proj1_sig m. auto. Qed. (*Wyckoff*) (*This decrements the value of m*) Definition segT_pred {n:nat} (m:n @) (pf:0 < proj1_sig m) : (pred n) @. destruct m as [m h1]. simpl in pf. pose proof (S_pred _ _ pf) as h2. rewrite h2 in h1. pose proof (lt_pred _ _ h1) as h3. refine (exist (fun c => c < pred n) _ h3). Defined. (*Wyckoff*) Definition segT_pred_compat : forall {n:nat} (m:n @) (pf:0 < proj1_sig m), proj1_sig (segT_pred m pf) = pred (proj1_sig m). intros ? m; destruct m; auto. Qed. (*Wyckoff*) Definition segT_incr {n:nat} (m:n @) : (S n) @ := exist (fun c => c < S n) (S (proj1_sig m)) (lt_n_S _ _ (proj2_sig m)). (*Wyckoff*) Definition segT_incr_compat : forall {n:nat} (m:n @), proj1_sig (segT_incr m) = S (proj1_sig m). auto. Qed. (*This only narrows the upper bound n (keeping the value of m the same)*) (*Wyckoff*) Definition segT_narrow {n:nat} (m:n @) (pfn:proj1_sig m <> pred n) : (pred n) @ := let pflt := lt_neq_impl _ _ (proj2_sig m) pfn in exist (fun c => c < pred n) (proj1_sig m) pflt. (*Wyckoff*) Lemma segT_narrow_compat : forall {n:nat} (m:n @) (pf:proj1_sig m <> pred n), proj1_sig (segT_narrow m pf) = proj1_sig m. auto. Qed. (*Wyckoff*) Definition segT_widen {n:nat} (m:n @) : S n @ := let pflt := lt_S _ _ (proj2_sig m) in exist (fun c => c < S n) (proj1_sig m) pflt. (*Wyckoff*) Lemma segT_widen_compat : forall {n:nat} (m:n @), proj1_sig (segT_widen m) = proj1_sig m. auto. Qed. Definition segT_fun n U : Type := n @ -> U. Definition segT_fun' n m : Type := n @ -> m @. Definition segged_fun T n : Type := T -> n @. Definition segged_fun_le {n m T} (f:segged_fun T n) (pfle:m >= n) : segged_fun T m := fun x => let (fx, pffx) := f x in let hfx' := lt_le_trans _ _ _ pffx pfle in exist (fun c => c < m) fx hfx'. Definition segT_fun_app {n U} (f:segT_fun (S n) U) : nat -> U := fun i => f (exist (fun c => c < S n) (nat_bound (S n) i) (lt_bound _ _)). Definition segT_fun_app' {n m} (f:segT_fun' (S n) m) : nat -> m @ := segT_fun_app f. End SegSet'. Notation "n ---> U" := (segT_fun n U) (at level 40, no associativity). Notation "n ----> m" := (segT_fun' n m) (at level 40, no associativity). Notation "T _--> n" := (segged_fun T n) (at level 40, no associativity). Notation "f |---> i" := (segT_fun_app f i) (at level 40, no associativity). Notation "f |----> i" := (segT_fun_app' f i) (at level 40, no associativity). Section SegSet''. Definition seg_fun_to_segT_fun {U n} (f:seg_fun n U) : n ---> U := fun i => f (proj1_sig i) (proj2_sig i). Lemma lt_segT_fun_app : forall {n U i} (f:segT_fun (S n) U) (pfi:i < S n), f |---> i = f (exist (fun c => c < S n) i pfi). intros n U i f hi. unfold segT_fun_app. f_equal; apply proj1_sig_injective; simpl. rewrite nat_bound_eq_same_iff; auto. Qed. Lemma ge_segT_fun_app : forall {n U i} (f:segT_fun (S n) U) (pfi:i >= S n), f |---> i = f (exist (fun c => c < S n) n (lt_n_Sn _)). intros n U i f hi. unfold segT_fun_app. f_equal; apply proj1_sig_injective; simpl. rewrite nat_bound_eq_iff. right; simpl; auto. Qed. (*Wyckoff*) Lemma segT_functional : forall (m n:nat), m = n -> m @ = n @. intros; subst; auto. Qed. (*Wyckoff*) Lemma segT_functional_eq_refl : forall (n:nat), segT_functional n n eq_refl = eq_refl. intro; apply proof_irrelevance. Qed. (*Wyckoff*) (*This preserves values.*) (*Analogous to [seg_from_S].*) Definition segged_from_Sn {T n} (f: T _--> S n) : {x:T | proj1_sig (f x) <> n} _--> n := fun x' => let (x, pfx) := x' in let pffx := lt_S_neq _ _ (proj2_sig (f x)) pfx in exist (fun c => c < n) (proj1_sig (f x)) pffx. (*Wyckoff*) (*This decrements values.*) (*Analogous to [seg_to_S]*) Definition segged_from_S0 {T n} (f: T _--> S n) : {x:T | 0 < proj1_sig (f x)} _--> n := fun x' => let (x, pfx) := x' in let pffx := lt_S_pred' _ _ pfx (proj2_sig (f x)) in exist (fun c => c < n) (pred (proj1_sig (f x))) pffx. (*Wyckoff*) Lemma segged_from_Sn_compat' : forall {T n} (f: T _--> S n) (x:{t:T | proj1_sig (f t) <> n}), proj1_sig (segged_from_Sn f x) = proj1_sig (f (proj1_sig x)). unfold segged_from_Sn; simpl; intros ? ? ? x; destruct x; auto. Qed. (*Wyckoff*) Lemma segged_from_Sn_compat : forall {T n} (f: T _--> S n) (x:{t:T | proj1_sig (f t) <> n}), let pflt := eq_lt_trans _ _ _ (eq_sym (segged_from_Sn_compat' f x)) (proj2_sig (segged_from_Sn f x)) in segged_from_Sn f x = exist (fun c => c < n) (proj1_sig (f (proj1_sig x))) pflt. unfold segged_from_Sn; intros ? ? ? x; destruct x; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Definition segged_from_S0_compat' : forall {T n} (f: T _--> S n) (x:{t:T | 0 < proj1_sig (f t)}), proj1_sig (segged_from_S0 f x) = pred (proj1_sig (f (proj1_sig x))). unfold segged_from_S0; simpl; intros ? ? ? x; destruct x; auto. Qed. (*Wyckoff*) Lemma segged_from_S0_compat : forall {T n} (f: T _--> S n) (x:{t:T | 0 < proj1_sig (f t)}), let pflt := eq_lt_trans _ _ _ (eq_sym (segged_from_S0_compat' f x)) (proj2_sig (segged_from_S0 f x)) in segged_from_S0 f x = exist (fun c => c < n) (pred (proj1_sig (f (proj1_sig x)))) pflt. unfold segged_from_S0; intros ? ? ? x; destruct x; f_equal; apply proof_irrelevance. Qed. Definition segged_cond_from_S0 {T n i} {f:T _--> S n} (pfp:forall x, proj1_sig (f x) <> S i) : forall x, (proj1_sig (segged_from_S0 f x)) <> i. intro x; destruct x as [x hx]; rewrite segged_from_S0_compat; simpl. apply S_inj_neq. rewrite <- (S_pred _ _ hx); auto. Defined. (*Wyckoff*) (*This preserves values before [i], decrements after*) (*Analogous to [remove_seg_fun] in [FunctionProperties2].*) Fixpoint segged_from_i {T n} : forall (f:T _--> n) i (pfi:i < n) (pfp:forall x, proj1_sig (f x) <> i), T _--> pred n := match n return ( forall (f:T _--> n) i (pfi:i < n) (pfp:forall x, proj1_sig (f x) <> i), T _--> pred n) with | 0 => fun _ _ pfi x => False_rect _ (lt_n_0 _ pfi) | S n' => fun f => let f' := segged_from_S0 f in let f_ := segged_from_Sn f in fun i => match i with | 0 => fun _ pfp x => f' (exist (fun c => 0 < proj1_sig (f c)) x (neq_lt0 _ (neq_sym _ _ (pfp x)))) | S i' => fun pfi pfp => let pfi' := lt_S_pred _ _ pfi in let pf0 := Olt _ _ pfi' in let pf0' := lt_neq _ _ pf0 in let pfp' := segged_cond_from_S0 pfp in let f0 := segged_from_i f' i' pfi' pfp' (n := n') in fun x => match (zerop (proj1_sig (f x))) with | left pfe => let pf0'' := eq_neq_trans _ _ _ pfe pf0' in f_ (exist (fun c => proj1_sig (f c) <> n') x pf0'') | right pflt => segT_S_pred_to (segT_S (f0 (exist (fun c => 0 < proj1_sig (f c)) x pflt))) pf0 end end end. Lemma segged_from_i_lt : forall {T n} (f:T _--> n) i j (pfi: i < n) (pfij:j < i) (pfp:forall x, proj1_sig (f x) <> i) x, proj1_sig (f x) = j -> let pfj := lt_trans _ _ _ pfij pfi in proj1_sig (segged_from_i f i pfi pfp x) = proj1_sig (f x). intros T n. revert T. induction n as [|n ih]; simpl. intros; omega. intros T f i j hi hji h1 x h2. destruct i as [|i]; simpl; subst. omega. lr_match_goal'; simpl; auto. specialize (ih _ (segged_from_S0 f) i (pred (proj1_sig (f x))) (lt_S_n _ _ hi)). hfirst ih. apply lt_S_dec in hji. destruct hji; subst; auto. omega. omega. specialize (ih hf). hfirst ih. intro y; destruct y as [y hy]. rewrite segged_from_S0_compat; simpl. specialize (h1 y). omega. rewrite segT_S_compat. specialize (ih hf0 (exist (fun c => 0 < proj1_sig (f c)) _ hr)). hfirst ih. rewrite segged_from_S0_compat; simpl; auto. specialize (ih hf1). simpl in ih. erewrite f_equal_dep. rewrite ih at 1. rewrite <- (S_pred _ _ hr); auto. apply proof_irrelevance. Qed. Lemma segged_from_i_ge : forall {T n} (f:T _--> n) i j (pfi: i < n) (pfij:i <= j) (pfp:forall x, proj1_sig (f x) <> i) x, proj1_sig (f x) = j -> proj1_sig (segged_from_i f i pfi pfp x) = pred (proj1_sig (f x)). intros T n; revert T. induction n as [|n ih]; simpl. intros; omega. intros T f i. destruct i as [|i]; simpl. intros; auto. intro j. destruct j as [|j]. intros hi h1 h2 x h4. lr_match_goal'; simpl. rewrite h4; auto. rewrite segT_S_compat. omega. intros hi hij hp x h1. lr_match_goal'; simpl. rewrite hl in h1; discriminate. rewrite segT_S_compat. specialize (ih _ (segged_from_S0 f) i j (lt_S_n _ _ hi) (ge_S_n _ _ hij) (segged_cond_from_S0 hp) (exist (fun c => 0 < proj1_sig (f c)) x hr)). hfirst ih. rewrite segged_from_S0_compat; simpl; omega. specialize (ih hf). erewrite f_equal_dep. rewrite ih at 1. rewrite segged_from_S0_compat; simpl. rewrite h1; simpl. destruct (zerop j) as [|h2]; subst. rewrite segged_from_S0_compat in e, hij; simpl in e, hij. rewrite e in hij; simpl in hij. apply le_S_n in hij. apply le0 in hij; subst. pose proof (hp x). omega. rewrite <- (S_pred _ _ h2); auto. apply proof_irrelevance. Qed. (*Wyckoff*) Lemma inj_segged_from_S0 : forall {T n} (f: T _--> S n), Injective f -> Injective (segged_from_S0 f). unfold segged_from_S0; intros T n f h1; red in h1; red; intros x y h2; destruct x as [x hx], y as [y hy]. apply exist_injective in h2. apply pred_inj0 in h2; auto. apply proj1_sig_injective in h2. apply h1 in h2; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma inj_segged_from_Sn : forall {T n} (f: T _--> S n), Injective f -> Injective (segged_from_Sn f). unfold segged_from_S0; intros T n f h1; red in h1; red; intros x y h2; destruct x as [x hx], y as [y hy]. apply exist_injective in h2. apply proj1_sig_injective in h2. apply h1 in h2; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Inductive sparse_segged_fun {T n} (f:T _--> n) : Prop := | sparse_segged_fun_intro : Injective f -> forall i, i < n -> (forall x, proj1_sig (f x) <> i) -> sparse_segged_fun f. Lemma sparse_segged_from_i : forall {T n} (f:T _--> n) i (pfi:i < n) (pfn:forall x : T, proj1_sig (f x) <> i), sparse_segged_fun (segged_from_i f i pfi pfn) -> sparse_segged_fun f. intros T n; revert T. induction n as [|n ih]. intros; omega. intros T f i hi hp h1. inversion h1 as [h2 k h3 h4]; simpl in h3. econstructor. red; intros x y h5. specialize (h2 x y). destruct (lt_le_dec (proj1_sig (f x)) i) as [h6|h6]. hfirst h2. apply proj1_sig_injective. rewrite (segged_from_i_lt _ _ _ _ h6). pose proof h6 as h6'. rewrite (segged_from_i_lt _ _ _ _ h6'). rewrite h5; auto. rewrite h5; auto. rewrite h5; auto. specialize (h2 hf); auto. hfirst h2. apply proj1_sig_injective. rewrite (segged_from_i_ge _ _ _ _ h6). pose proof h6 as h6'. rewrite (segged_from_i_ge _ _ _ _ h6'). rewrite h5; auto. rewrite h5; auto. rewrite h5; auto. specialize (h2 hf); auto. apply hi. assumption. Qed. (*Wyckoff*) Definition seg_fun_to_segged {U n} (f:seg_fun n U) : n ---> U := fun i => f (proj1_sig i) (proj2_sig _). (*Wyckoff*) Definition seg_fun_type_to_ind {n} : seg_fun n Type -> n ---> Type := fun f => seg_fun_to_segged f. (*Wyckoff*) Lemma seg_one_t_functional : forall (m n:nat), m = n -> seg_one_t m = seg_one_t n. intros; subst; auto. Defined. (*Wyckoff*) Lemma seg_one_t_functional_eq_refl : forall (n:nat), seg_one_t_functional n n eq_refl = eq_refl. intro; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma nin_seg : forall n:nat, ~In (seg n) n. intros n h1. destruct h1. omega. Qed. (*Wyckoff*) Lemma nin_seg_0 : forall n:nat, ~In (seg 0) n. intros n h1. destruct h1. omega. Qed. (*Wyckoff*) Lemma in_seg_Sn : forall n:nat, In (seg (S n)) 0. intro; constructor; auto with arith. Qed. (*Wyckoff*) Lemma in_seg_n_S : forall m n:nat, In (seg n) m -> In (seg (S n)) (S m). intros m n h1. destruct h1 as [h1]. constructor. apply lt_n_S; auto. Qed. (*Wyckoff*) Lemma seg_0_eq : seg 0 = Empty_set nat. apply Extensionality_Ensembles; red; split; auto with sets; red; intros n h1; apply nin_seg_0 in h1; contradiction. Qed. (*Wyckoff*) Lemma seg_1_eq : seg 1 = Singleton 0. apply Extensionality_Ensembles; red; split; red; intros x h1; destruct h1 as [h1]. assert (x = 0). omega. subst. constructor. constructor. omega. Qed. (*Wyckoff*) Lemma seg_eq_empty : forall n, seg n = Empty_set _ -> n = 0. intros n h1; destruct n; auto. pose proof (lt_0_Sn n) as h2. apply in_seg in h2. rewrite h1 in h2; inversion h2. Qed. (*Wyckoff*) Lemma seg_eq_singleton : forall (a n:nat), seg n = Singleton a -> n = 1 /\ a = 0. intros a n h1. pose proof (In_singleton _ a) as h3. rewrite <- h1 in h3. destruct h3 as [h3]. rewrite (S_pred _ _ h3) in h1. pose proof (lt_0_Sn (pred n)) as h4. apply in_seg in h4. rewrite h1 in h4. destruct h4; subst. pose proof (lt_n_Sn (pred n)) as h4. apply in_seg in h4. rewrite h1 in h4. inversion h4; subst. rewrite <- (S_pred _ _ h3); auto. Qed. (*Wyckoff*) Lemma seg_one_0 : seg_one 0 = Empty_set _. apply Extensionality_Ensembles; red; split; red; intros x h1; try contradiction; destruct h1; omega. Qed. (*Wyckoff*) Lemma seg_one_t_0_false : forall x:seg_one_t 0, False. intro x. destruct x; omega. Qed. (*Wyckoff*) Lemma full_set_segT0 : Full_set (0 @) = Empty_set _. apply Extensionality_Ensembles; red; split; auto with sets. red; intros m h1. pose proof (segT_0_false m). contradiction. Qed. (*Wyckoff*) Definition sot_1 : seg_one_t 1. assert (h1:1 <= 1 <= 1). auto with arith. refine (exist (fun m => 1 <= m <= 1) _ h1). Defined. (*Wyckoff*) Lemma seg_incl_seg_weak : forall (n:nat), Included (seg n) (seg_weak n). intros n. red. intros m h1. constructor. destruct h1. omega. Qed. (*Wyckoff*) Lemma seg_one_incl_seg_weak : forall n:nat, Included (seg_one n) (seg_weak n). intros; red; intros m h1; constructor; destruct h1; omega. Qed. (*Wyckoff*) Lemma seg_weak_O : seg_weak 0 = Singleton 0. apply Extensionality_Ensembles. red. split. red. intros m h1. destruct h1 as [h1]. assert (m = 0). omega. subst. constructor. red. intros m h1. destruct h1; subst. constructor. auto with arith. Qed. (*Wyckoff*) Lemma seg_weak_Sn : forall n:nat, seg_weak (S n) = Add (seg_weak n) (S n). intro n. apply Extensionality_Ensembles. red. split. red. intros m h1. destruct h1 as [h1]. destruct (le_lt_eq_dec _ _ h1) as [h2 | h3]. left. constructor. omega. subst. right. constructor. red. intros m h1. destruct h1 as [m h1l | m h1r]. destruct h1l as [h1l]. constructor. omega. destruct h1r; subst. constructor. auto with arith. Qed. (*Wyckoff*) Lemma seg_one_Sn : forall n:nat, seg_one (S n) = Add (seg_one n) (S n). intro n. apply Extensionality_Ensembles. red. split. red. intros m h1. destruct h1 as [h1]. destruct h1 as [h1 h1']. destruct (le_lt_eq_dec _ _ h1') as [h2 | h3]. left. constructor. omega. subst. right. constructor. red. intros m h1. destruct h1 as [m h1l | m h1r]. destruct h1l as [h1l]. constructor. omega. destruct h1r; subst. constructor. auto with arith. Qed. (*Wyckoff*) Lemma seg_O : seg O = Empty_set _. apply Extensionality_Ensembles. red; split; auto with sets. red. intros m h1. destruct h1; omega. Qed. (*Wyckoff*) Lemma seg_Sn : forall n:nat, seg (S n) = Add (seg n) n. intro n. apply Extensionality_Ensembles. red. split. red. intros m h1. destruct h1 as [h1]. assert (h2: m <= n). omega. destruct (le_lt_eq_dec _ _ h2) as [h3 | h4]. left. constructor. omega. subst. right. constructor. red. intros m h1. destruct h1 as [m h1l | m h1r]. destruct h1l as [h1l]. constructor. omega. destruct h1r; subst. constructor. auto with arith. Qed. (*Wyckoff*) Lemma seg_eq_couple : forall (a b n:nat), seg n = Couple a b-> n = 1 /\ a = b /\ a = 0 \/ (n = 2 /\ ((a = 0 /\ b = 1) \/ (a = 1 /\ b = 0))). intros a b n h1. destruct (nat_eq_dec a b) as [|h2]; subst. left. rewrite couple_singleton in h1. apply seg_eq_singleton in h1. destruct h1; subst; auto. right. destruct (nat_eq_dec n 2) as [h3 | h3]; subst. pose proof (in_seg _ _ (lt_0_Sn 1)) as h3. rewrite h1 in h3. inversion h3; subst. split; auto. left. pose proof (lt_n_Sn 1) as h4. apply in_seg in h4. rewrite h1 in h4. inversion h4; subst; auto. pose proof (lt_n_Sn 1) as h4. apply in_seg in h4. rewrite h1 in h4. inversion h4; subst; auto. assert (h5:n <> 1). intro; subst. rewrite seg_1_eq in h1. symmetry in h1. apply couple_eq_sing in h1. destruct h1; subst. contradict h2; auto. assert (h6: n <> 0). intro; subst. rewrite seg_0_eq in h1. symmetry in h1. apply couple_not_empty in h1; contradiction. assert (h7:0 < n). omega. pose proof h7 as h8. apply in_seg in h7. rewrite h1 in h7. inversion h7; subst. pose proof (lt_n_Sn (pred n)) as h9. rewrite <- (S_pred _ _ h8) in h9. apply in_seg in h9. rewrite h1 in h9. inversion h9. omega. subst. rewrite (S_pred _ _ h8) in h1 at 1. rewrite seg_Sn in h1. rewrite couple_add_sing in h1. apply add_nin_inj_eq in h1. apply seg_eq_singleton in h1. destruct h1 as [h1 h10]. omega. apply nin_seg. rewrite in_sing_iff; auto. pose proof (lt_n_Sn (pred n)) as h9. rewrite <- (S_pred _ _ h8) in h9. apply in_seg in h9. rewrite h1 in h9. inversion h9; subst. rewrite (S_pred _ _ h8) in h1 at 1. rewrite seg_Sn in h1. rewrite couple_comm in h1. rewrite couple_add_sing in h1. apply add_nin_inj_eq in h1. apply seg_eq_singleton in h1. destruct h1 as [h1 h10]. omega. apply nin_seg. rewrite in_sing_iff; auto. omega. Qed. (*Wyckoff*) Lemma finite_seg_weak : forall n:nat, Finite (seg_weak n). intro n. induction n as [|n h1]. rewrite seg_weak_O. apply Singleton_is_finite. rewrite seg_weak_Sn. constructor; auto. intro h2. destruct h2 as [h2]. omega. Qed. (*Wyckoff*) Lemma finite_seg : forall n:nat, Finite (seg n). intros n. pose proof (seg_incl_seg_weak n) as h1. pose proof (finite_seg_weak n) as h2. apply (Finite_downward_closed _ _ h2 _ h1). Qed. (*Wyckoff*) Lemma finite_full_set_segT : forall n:nat, Finite (Full_set (n @)). intro n. pose proof (finite_seg n) as h1. rewrite <- segT_compat in h1. rewrite finite_im_proj1_sig_iff in h1. assumption. Qed. (*Wyckoff*) Lemma finite_seg_one : forall n:nat, Finite (seg_one n). intro n. pose proof (seg_one_incl_seg_weak n) as h1. pose proof (finite_seg_weak n) as h2. apply (Finite_downward_closed _ _ h2 _ h1). Qed. (*Wyckoff*) Lemma finite_full_set_seg_one_t : forall n:nat, Finite (Full_set (seg_one_t n)). intro n. pose proof (finite_seg_one n) as h1. rewrite <- seg_one_t_compat in h1. rewrite finite_im_proj1_sig_iff in h1. assumption. Qed. (*Wyckoff*) Lemma full_set_segT_S_eq : forall m:nat, Full_set ((S m) @) = Add (Im (Full_set (m @)) (fun c => (exist (fun c => c < S m) _ (lt_trans _ _ _ (proj2_sig c) (lt_n_Sn _))))) (exist (fun c => c < S m) _ (lt_n_Sn _)). intros m. apply Extensionality_Ensembles; red; split; red; intros c h1. clear h1. destruct c as [c h1]. destruct (lt_S_dec _ _ h1) as [h2 | h2]. left. apply Im_intro with (exist (fun c => c < m) _ h2). constructor. apply proj1_sig_injective; auto. subst. right. pose proof (proof_irrelevance _ h1 (lt_n_Sn m)); subst. constructor; auto. constructor. Qed. (*Wyckoff*) Lemma inh_seg : forall (n:nat), 0 < n -> Inhabited (seg n). intros n h1. apply Inhabited_intro with 0. constructor. assumption. Qed. (*Wyckoff*) Lemma segT_pos : forall (n:nat) (m:seg_one_t n), 0 <= proj1_sig m. intros ? m. destruct m; auto with arith. Qed. (*Wyckoff*) Lemma inh_seg_weak : forall n:nat, Inhabited (seg_weak n). intro n. apply Inhabited_intro with 0. constructor; auto with arith. Qed. (*Wyckoff*) Lemma inh_seg_one : forall (n:nat), 0 < n -> Inhabited (seg_one n). intros n h1. apply Inhabited_intro with 1. constructor; auto with arith. Qed. (*Wyckoff*) Lemma seg_one_t_pos : forall (n:nat) (m:seg_one_t n), 1 <= proj1_sig m. intros ? m. destruct m as [m h1]. destruct h1; auto. Qed. (*Wyckoff*) Lemma seg_one_decompose : forall m n:nat, m < n -> seg_one n = Union (seg_one m) [r:nat | m < r <= n]. intros m n h1. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2 as [h2]. destruct h2 as [h2 h3]. destruct (lt_eq_lt_dec x m) as [h4 | h4]. destruct h4 as [h4 | h4]. left. constructor. omega. subst. left. constructor. omega. right. constructor. omega. destruct h2 as [x h2 | x h2]. destruct h2 as [h2]. constructor. omega. destruct h2 as [h2]. constructor. omega. Qed. (*Wyckoff*) Lemma incl_lt_seg : forall m n:nat, m < n -> Included (seg m) (seg n). intros m n h1. red; intros x h2. destruct h2 as [h2]. constructor. eapply lt_trans. apply h2. assumption. Qed. (*Wyckoff*) Lemma incl_le_seg : forall m n:nat, m <= n -> Included (seg m) (seg n). intros m n h1. red; intros x h2. destruct h2 as [h2]. constructor. eapply lt_le_trans. apply h2. assumption. Qed. (*Wyckoff*) Lemma incl_le_seg_one : forall m n:nat, m <= n -> Included (seg_one m) (seg_one n). intros m n h1. red; intros x h2. destruct h2 as [h2]. constructor. omega. Qed. (*Wyckoff*) Lemma incl_lt_seg_one : forall m n:nat, m < n -> Included (seg_one m) (seg_one n). intros m n h1. red; intros x h2. destruct h2 as [h2]. constructor. omega. Qed. (*Wyckoff*) Lemma sot_pred_lt_length : forall {n:nat} (m:seg_one_t n), pred (proj1_sig m) < n. intros n m. destruct m as [m h1]. destruct h1 as [h1 h2]. simpl. assert (h3:0 < m). auto with arith. apply lt_S_n. rewrite <- (S_pred _ _ h3). auto with arith. Qed. (*Wyckoff*) Definition st01 : 1 @ := exist (fun a => a < 1) _ st01_pf. (*Wyckoff*) Definition st02 : 2 @ := exist (fun a => a < 2) _ st02_pf. (*Wyckoff*) Definition st12 : 2 @ := exist (fun a => a < 2) _ st12_pf. (*Wyckoff*) Definition sot11 : seg_one_t 1 := exist (fun a => 1 <= a <= 1) _ sot11_pf. (*Wyckoff*) Definition sot12 : seg_one_t 2 := exist (fun a => 1 <= a <= 2) _ sot12_pf. (*Wyckoff*) Definition sot22 : seg_one_t 2 := exist (fun a => 1 <= a <= 2) _ sot22_pf. (*Wyckoff*) Lemma full_set_segT1 : Full_set (1 @) = Singleton st01. apply Extensionality_Ensembles; red; split; red; intros x h1; destruct x as [x h2]; subst; simpl; unfold st01. assert (x = 0). omega. subst. pose proof (proof_irrelevance _ h2 st01_pf); subst; constructor. constructor. Qed. (*Wyckoff*) Lemma full_set_seg_one_t_1 : Full_set (seg_one_t 1) = Singleton sot11. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct x as [x h2]. assert (h4:x = 1). destruct h2; auto with arith. subst. unfold sot11. pose proof (proof_irrelevance _ sot11_pf h2). subst. constructor. constructor. Qed. (*Wyckoff*) Lemma full_set_seg_one_t_2 : Full_set (seg_one_t 2) = Couple sot12 sot22. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct x as [x h2]. destruct h2 as [h2 h3]. pose proof h2 as h2'. apply le_lt_eq_dec in h2'. destruct h2' as [h4 | h4]. assert (h5:x = 2). auto with arith. subst. unfold sot22. pose proof (proof_irrelevance _ (sot22_pf) (conj h2 h3)) as h5. rewrite <- h5. right. subst. unfold sot12. pose proof (proof_irrelevance _ (sot12_pf) (conj h2 h3)) as h6. rewrite <- h6. left. constructor. Qed. (*Wyckoff*) Lemma seg_one_t_2_dec : forall (n:seg_one_t 2), {proj1_sig n = 1} + {proj1_sig n = 2}. intro n. destruct n as [n h1]. destruct h1 as [h1 h2]; simpl. apply le_lt_eq_dec in h1. destruct h1. right. omega. subst. left; auto. Qed. (*Wyckoff*) Lemma im_seg_one_t_decompose : forall {T:Type} (m n:nat) (f:seg_one_t n -> T) (pf:m < n), let f' := restriction_sig_prop _ _ (lt_prop_seg_one_t m n pf) f in let f'' := restriction_sig_prop _ _ (lt_le_prop_seg_one_t m n pf) f in Im (Full_set (seg_one_t n)) f = Union (Im (Full_set (seg_one_t m)) f') (Im (Full_set {r:nat | m < r <= n}) f''). intros T m n f h1 f' f''. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2 as [x h2]. subst. clear h2. destruct x as [x h2]. destruct h2 as [h2 h2']. destruct (le_lt_dec x m) as [h3 | h3]. left. apply Im_intro with (exist (fun s => 1 <= s <= m) _ (conj h2 h3)). constructor. simpl. f_equal. apply proj1_sig_injective; auto. right. apply Im_intro with (exist (fun s => m < s <= n) _ (conj h3 h2')). constructor. simpl. f_equal. apply proj1_sig_injective; auto. destruct h2 as [x h2 | x h2]. destruct h2 as [x h2]. subst. clear h2. destruct x as [x h2]. assert (h3: 1 <= x <= n). omega. apply Im_intro with (exist (fun s=>1 <= s <= n) _ h3). constructor. simpl. f_equal. apply proj1_sig_injective; auto. destruct h2 as [x h2]. subst. clear h2. destruct x as [x h2]. assert (h3:1 <= x <= n). omega. apply Im_intro with (exist (fun s=> 1 <= s <= n) _ h3). constructor. simpl. f_equal. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma seg_one_t_S_dec : forall {T:Type} {l:list T} (x:seg_one_t (S (length l))), {proj1_sig x = 1} + {1 <= pred (proj1_sig x) <= length l}. intros T l x. destruct x as [x h1]; simpl. destruct h1 as [h1 h2]. apply le_lt_eq_dec in h1. destruct h1 as [h1 | h1]. assert (h3:2 <= x <= S (length l)). omega. right. omega. subst. left. reflexivity. Qed. (*Wyckoff*) Lemma card_fun1_seg_eq : forall (n:nat), card_fun1 (seg n) = n. intro n. induction n as [|n h1]. rewrite seg_0_eq. apply card_fun1_empty. rewrite seg_Sn. rewrite card_add_nin'. rewrite h1; auto. apply finite_seg. intro h2. destruct h2 as [h2]. apply lt_irrefl in h2. contradiction. Qed. (*Wyckoff*) Lemma in_ens_S_seg : forall (m n:nat), m < n -> In (ens_S (seg n)) (S m). intros; apply in_ens_S; constructor; auto. Qed. (*Wyckoff*) Lemma inj_on_seg1 : forall {T:Type} (f:nat->T), inj_on (seg 1) f. intros T f. red. intros x y h1 h2 h3. destruct h1, h2; omega. Qed. (*Wyckoff*) Definition fun_to_sig_seg {U:Type} (f:nat->U) (n:nat) : seg_fun n (sig_set (Im (seg n) f)) := fun i (pf:i < n) => exist _ _ (Im_intro _ _ (seg n) f i (in_seg _ _ pf) _ eq_refl). (*Wyckoff*) Lemma fun_to_sig_seg_compat : forall {U:Type} (f:nat->U) (n:nat) (i:nat) (pf:i < n), fun_to_sig_seg f n i pf = exist _ (f i) (Im_intro _ _ (seg n) f i (in_seg _ _ pf) _ eq_refl). unfold fun_to_sig_seg; intros; auto. Qed. (*Wyckoff*) Definition seg_fun_to_ens {U:Type} {n:nat} (f:seg_fun n U) : fun_in_ens (seg n) U := fun i (pf:Inn (seg n) i) => f i (seg_in _ _ pf). (*Wyckoff*) Lemma seg_fun_to_ens_compat : forall {U:Type} {n:nat} (f:seg_fun n U) (i:nat) (pf:Inn (seg n) i), seg_fun_to_ens f i pf = f i (seg_in _ _ pf). unfold seg_fun_to_ens; intros; auto. Qed. End SegSet''. Section CartProd. (*Wyckoff*) Definition cart_prod {T U:Type} (A:Ensemble T) (B:Ensemble U) : Ensemble (T*U) := [pr:T*U | In A (fst pr) /\ In B (snd pr)]. (*Wyckoff*) Lemma in_cart_prod : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (x:T) (y:U), In A x -> In B y -> In (cart_prod A B) (x, y). intros A B x y h1 h2. constructor. simpl. split; assumption. Qed. (*Wyckoff*) Lemma in_cart_prod_comm : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (pr:T*U), In (cart_prod A B) pr -> In (cart_prod B A) (snd pr, fst pr). intros T U A B pr h1. constructor. simpl. destruct h1 as [h1]. destruct h1; split; assumption. Qed. (*Wyckoff*) Lemma cart_prod_sing : forall {T U:Type} (A:Ensemble T) (y:U), cart_prod A (Singleton y) = Im A (fun x:T => (x, y)). intros T U A y. apply Extensionality_Ensembles. red. split. (* <= *) red. intros pr h1. unfold cart_prod in h1. destruct h1 as [h1]. destruct h1 as [h1 h2]. apply Im_intro with (fst pr). assumption. pose proof (Singleton_inv _ _ _ h2) as h3. subst. apply surjective_pairing. (* >= *) red. intros pr h1. destruct h1 as [x h2 pr]. subst. unfold cart_prod. constructor. split. simpl. assumption. simpl. constructor. Qed. (*Wyckoff*) Lemma cart_prod_sing' : forall {T U:Type} (B:Ensemble U) (x:T), cart_prod (Singleton x) B = Im B (fun y:U => (x, y)). intros T U B x. apply Extensionality_Ensembles. red. split. (* <= *) red. intros pr h1. unfold cart_prod in h1. destruct h1 as [h1]. destruct h1 as [h1 h2]. apply Im_intro with (snd pr). assumption. pose proof (Singleton_inv _ _ _ h1) as h3. subst. apply surjective_pairing. (* >= *) red. intros pr h1. destruct h1 as [z h2 pr]. subst. unfold cart_prod. constructor. split. simpl. constructor. simpl. assumption. Qed. (*Wyckoff*) Lemma cart_prod_sing_rev : forall {T U:Type} (A:Ensemble T) (y:U), A = Im (cart_prod A (Singleton y)) (fun pr:T*U => fst pr). intros T U A y. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h1. apply Im_intro with (x, y). unfold cart_prod. constructor. simpl. split; auto with sets. reflexivity. (* >= *) red. intros x h1. destruct h1 as [pr h1]; subst. destruct h1 as [h1]. destruct h1; assumption. Qed. (*Wyckoff*) Lemma cart_prod_sing_rev' : forall {T U:Type} (B:Ensemble U) (x:T), B = Im (cart_prod (Singleton x) B) (fun pr:T*U => snd pr). intros T U B x. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h1. apply Im_intro with (x, y). unfold cart_prod. constructor. simpl. split; auto with sets. reflexivity. (* >= *) red. intros y h1. destruct h1 as [pr h1]; subst. destruct h1 as [h1]. destruct h1; assumption. Qed. (*Wyckoff*) Lemma card_cart_prod_sing : forall {T U:Type} (A:Ensemble T) (n:nat) (b:U), cardinal _ A n -> cardinal _ (cart_prod A (Singleton b)) n. intros T U A n b h1. rewrite cart_prod_sing. pose proof (injective_preserves_cardinal). pose proof (cardinal_finite _ _ _ h1) as h2. pose proof (finite_image _ _ _ (fun x:T=>(x, b)) h2) as h3. pose proof (finite_cardinal _ _ h3) as h4. destruct h4 as [n' h4]. assert (h5:Image.injective _ _ (fun x:T=>(x, b))). red. intros x' y' h5. injection h5. auto. pose proof (injective_preserves_cardinal _ _ _ _ _ h5 h1 _ h4) as h6. subst. assumption. Qed. (*Wyckoff*) Lemma cart_prod_sing_fin : forall {T U:Type} (A:Ensemble T) (y:U), Finite A -> Finite (cart_prod A (Singleton y)). intros T U A y h1. pose proof (cart_prod_sing A y) as h2. rewrite h2. apply finite_image. assumption. Qed. (*Wyckoff*) Lemma cart_prod_sing_fin' : forall {T U:Type} (B:Ensemble U) (x:T), Finite B -> Finite (cart_prod (Singleton x) B). intros T U B x h1. pose proof (cart_prod_sing' B x) as h2. rewrite h2. apply finite_image. assumption. Qed. (*Wyckoff*) Lemma cart_prod_sing_fin_rev : forall {T U:Type} (A:Ensemble T) (y:U), Finite (cart_prod A (Singleton y)) -> Finite A. intros T U A y h1. pose proof (cart_prod_sing_rev A y) as h2. rewrite h2. apply finite_image. assumption. Qed. (*Wyckoff*) Lemma cart_prod_sing_fin_rev' : forall {T U:Type} (B:Ensemble U) (x:T), Finite (cart_prod (Singleton x) B) -> Finite B. intros T U B x h1. pose proof (cart_prod_sing_rev' B x) as h2. rewrite h2. apply finite_image. assumption. Qed. (*Wyckoff*) Lemma card_cart_prod_sing' : forall {T U:Type} (A:Ensemble T), Finite A -> forall (b:U), card_fun1 (cart_prod A (Singleton b)) = card_fun1 A. intros T U A h1 b. pose proof (finite_cardinal _ _ h1) as h2. destruct h2 as [n h2]. pose proof (card_cart_prod_sing A n b h2) as h3. pose proof (card_fun1_compat (cart_prod A (Singleton b))) as h4. pose proof (card_fun1_compat A) as h5. destruct h4 as [h4l h4r]. destruct h5 as [h5l h5r]. specialize (h5l h1). pose proof (cart_prod_sing_fin _ b h1) as h6. specialize (h4l h6). pose proof (cardinal_is_functional _ _ _ h2 _ _ h5l (eq_refl)). subst. pose proof (cardinal_is_functional _ _ _ h3 _ _ h4l (eq_refl)). symmetry. assumption. Qed. (*Wyckoff*) Lemma cart_prod_inc_add1 : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (x:T), Included (cart_prod A B) (cart_prod (Add A x) B). intros T U A B x. red. intros pr h1. destruct h1 as [h1]. destruct h1 as [h1l h1r]. constructor. split; try assumption. left. assumption. Qed. (*Wyckoff*) Lemma cart_prod_inc_add2 : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (y:U), Included (cart_prod A B) (cart_prod A (Add B y)). intros T U A B y. red. intros pr h1. destruct h1 as [h1]. destruct h1 as [h1l h1r]. constructor. split; try assumption. left. assumption. Qed. (*Wyckoff*) Lemma cart_prod_eq : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), cart_prod A B = FamilyUnion [S:Ensemble (T*U) | exists y:U, In B y /\ S = cart_prod A (Singleton y)]. intros T U A B. apply Extensionality_Ensembles. red. split. (* <= *) red. intros pr h1. destruct h1 as [h1l h1r]. destruct h1l as [h1ll h1lr]. assert (h2:In (cart_prod A (Singleton (snd pr))) pr). rewrite cart_prod_sing. apply Im_intro with (fst pr). assumption. apply surjective_pairing. apply family_union_intro with (cart_prod A (Singleton (snd pr))). constructor. exists (snd pr). split; [assumption | reflexivity]. assumption. (* >= *) red. intros pr h1. destruct h1 as [S pr h1 h2]. destruct h1 as [h1]. destruct h1 as [y h1]. destruct h1 as [h1l h1r]. subst. rewrite cart_prod_sing in h2. destruct h2 as [x]. subst. constructor; split; assumption. Qed. (*Wyckoff*) Lemma cart_prod_eq' : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), cart_prod A B = FamilyUnion [S:Ensemble (T*U) | exists x:T, In A x /\ S = cart_prod (Singleton x) B]. intros T U A B. apply Extensionality_Ensembles. red. split. (* <= *) red. intros pr h1. destruct h1 as [h1l h1r]. destruct h1l as [h1ll h1lr]. assert (h2:In (cart_prod (Singleton (fst pr)) B) pr). rewrite cart_prod_sing'. apply Im_intro with (snd pr). assumption. apply surjective_pairing. apply family_union_intro with (cart_prod (Singleton (fst pr)) B). constructor. exists (fst pr). split; [assumption | reflexivity]. assumption. (* >= *) red. intros pr h1. destruct h1 as [S pr h1 h2]. destruct h1 as [h1]. destruct h1 as [y h1]. destruct h1 as [h1l h1r]. subst. rewrite cart_prod_sing' in h2. destruct h2 as [x]. subst. constructor; split; assumption. Qed. (*Wyckoff*) Lemma cart_prod_empty : forall {T U:Type} (B:Ensemble U), cart_prod (Empty_set T) B = Empty_set _. intros T U B. apply Extensionality_Ensembles. red. split. red. intros pr h1. unfold cart_prod in h1. destruct h1 as [h1]. destruct h1; contradiction. auto with sets. Qed. (*Wyckoff*) Lemma cart_prod_empty_empty : forall {T U:Type} (B:Ensemble U), cart_prod (Empty_set T) B = cart_prod (Empty_set _) (Empty_set _). intros T U B. rewrite cart_prod_empty. rewrite <- (cart_prod_empty (Empty_set _)). reflexivity. Qed. (*Wyckoff*) Lemma cart_prod_empty' : forall {T U:Type} (A:Ensemble T), cart_prod A (Empty_set U)= Empty_set _. intros T U A. apply Extensionality_Ensembles. red. split. red. intros pr h1. unfold cart_prod in h1. destruct h1 as [h1]. destruct h1; contradiction. auto with sets. Qed. (*Wyckoff*) Lemma cart_prod_empty_rev : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), cart_prod A B = Empty_set _ -> (A = Empty_set _ \/ B = Empty_set _). intros T U A B h1. apply NNPP. intro h2. pose proof (not_or_and _ _ h2) as h3. destruct h3 as [h3 h4]. pose proof (not_empty_Inhabited _ _ h3) as h5. pose proof (not_empty_Inhabited _ _ h4) as h6. destruct h5 as [a h5]. destruct h6 as [b h6]. assert (h7:In (cart_prod A B) (a, b)). constructor. simpl. split; assumption. rewrite h1 in h7. contradiction. Qed. (*Wyckoff*) Lemma cart_prod_empty_empty' : forall {T U:Type} (A:Ensemble U), cart_prod A (Empty_set U) = cart_prod (Empty_set _) (Empty_set _). intros T U B. rewrite cart_prod_empty'. rewrite <- (cart_prod_empty' (Empty_set _)). reflexivity. Qed. (*Wyckoff*) Lemma add_ex_family : forall {T U:Type} (A:Ensemble U) (a:U) (f:U -> Ensemble T), [S:Ensemble T | exists u:U, Inn (Add A a) u /\ S = f u] = Add [S:Ensemble T | exists u:U, Inn A u /\ S = f u] (f a). intros T U A a f. apply Extensionality_Ensembles. red. split. (* <= *) red. intros S h6. destruct h6 as [h6]. destruct h6 as [y h6]. destruct h6 as [h6l h6r]. destruct h6l as [y h6la | y h6lb]. left. constructor. exists y. split; assumption. destruct h6lb. right. rewrite h6r. constructor. (* >= *) red. intros S h6. destruct h6 as [S h6 | S h7']. destruct h6 as [h6]. destruct h6 as [y h6]. destruct h6 as [h6a h6b]. constructor. exists y. split. left. assumption. assumption. destruct h7'. constructor. exists a. split. right. constructor. reflexivity. Qed. (*Wyckoff*) Lemma cart_prod_fin : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), Finite A -> Finite B -> Finite (cart_prod A B). intros T U A B h1. induction h1. intros h2. rewrite cart_prod_empty. constructor. intro h2. rewrite cart_prod_eq. apply Finite_Finite_Union. intros S h3. destruct h3 as [h3]. destruct h3 as [y h3]. destruct h3 as [h3l h3r]. rewrite h3r. apply cart_prod_sing_fin. apply Add_preserves_Finite. assumption. induction h2 as [|B h3 h4 z h5]. gterm0. rename c into C. fold C. assert (h3:C = Empty_set _). unfold C. apply Extensionality_Ensembles. red. split. red. intros S h2. destruct h2 as [h2]. destruct h2 as [y h2]. destruct h2; contradiction. auto with sets. rewrite h3. constructor. pose proof (Add_preserves_Finite _ B z h3) as h0. specialize (IHh1 h0). rewrite cart_prod_eq in IHh1. assert (hs : Included (FamilyUnion [S:Ensemble (T*U) | exists y:U, In B y /\ S = cart_prod A (Singleton y)]) (FamilyUnion [S : Ensemble (T * U) | exists y : U, In (Add B z) y /\ S = cart_prod A (Singleton y)])). red. intros pr h6. destruct h6 as [S pr h7]. apply family_union_intro with S. destruct h7 as [h7]. destruct h7 as [y h7]. constructor. exists y. split. left. tauto. tauto. assumption. pose proof (Finite_downward_closed _ _ IHh1 _ hs) as h7. rewrite <- cart_prod_eq in h7. hfirst h4. tauto. specialize (h4 hf). pose proof (add_ex_family B z (fun u:U => cart_prod (Add A x) (Singleton u))) as h6. simpl in h6. rewrite h6. constructor. assumption. intro h9. destruct h9 as [h9]. destruct h9 as [a h9]. destruct h9 as [h9a h9b]. assert (h10:In (cart_prod (Add A x) (Singleton z)) (x, z)). constructor. simpl. split. right. constructor. constructor. rewrite h9b in h10. destruct h10 as [h10]. destruct h10 as [h10a h10b]. simpl in h10b. destruct h10b. contradiction. Qed. (*Wyckoff*) Lemma cart_prod_fin_rev1 : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), Finite (cart_prod A B) -> Inhabited B -> Finite A. intros T U A B h1 h2. dependent induction h1. destruct h2 as [y h2]. assert (h3:A = Empty_set _). apply Extensionality_Ensembles. red. split. (* <= *) red. intros z h1. assert (h3:In (cart_prod A B) (z, y)). constructor. simpl. split; assumption. rewrite <- x in h3. contradiction. (* >= *) auto with sets. rewrite h3. constructor. pose proof (Add_preserves_Finite _ _ x0 h1) as h3. rewrite x in h3. rewrite cart_prod_eq in h3. pose proof (Finite_downward_closed _ _ h3 (cart_prod A (Singleton (snd x0)))) as hf. hfirst hf. red. intros pr h4. apply family_union_intro with (cart_prod A (Singleton (snd x0))). constructor. exists (snd x0). split. pose proof (Add_intro2 _ A0 x0) as h5. rewrite x in h5. destruct h5 as [h5]. destruct h5; assumption. reflexivity. assumption. specialize (hf hf0). apply cart_prod_sing_fin_rev with (snd x0); auto. Qed. (*Wyckoff*) Lemma cart_prod_fin_rev2 : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), Finite (cart_prod A B) -> Inhabited A -> Finite B. intros T U A B h1 h2. dependent induction h1. destruct h2 as [a h2]. assert (h3:B = Empty_set _). apply Extensionality_Ensembles. red. split. (* <= *) red. intros b h1. assert (h3:In (cart_prod A B) (a, b)). constructor. simpl. split; assumption. rewrite <- x in h3. contradiction. (* >= *) auto with sets. rewrite h3. constructor. pose proof (Add_preserves_Finite _ _ x0 h1) as h3. rewrite x in h3. rewrite cart_prod_eq' in h3. assert (h4:Included (cart_prod (Singleton (fst x0)) B) (FamilyUnion [S : Ensemble (T * U) | exists x1 : T, In A x1 /\ S = cart_prod (Singleton x1) B])). red. intros pr h4. apply family_union_intro with (cart_prod (Singleton (fst x0)) B). constructor. exists (fst x0). split. pose proof (Add_intro2 _ A0 x0) as h5. rewrite x in h5. destruct h5 as [h5]. destruct h5; assumption. reflexivity. assumption. pose proof (Finite_downward_closed _ _ h3 _ h4) as h5. apply cart_prod_sing_fin_rev' with (fst x0). assumption. Qed. (*Wyckoff*) Lemma cart_prod_fin_rev : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), Finite (cart_prod A B) -> Inhabited (cart_prod A B) -> Finite A /\ Finite B. intros T U A B h1 h2. destruct h2 as [pr h2]. destruct h2 as [h2]. destruct h2 as [h2l h2r]. pose proof (Inhabited_intro _ _ _ h2l) as h3. pose proof (Inhabited_intro _ _ _ h2r) as h4. split. apply cart_prod_fin_rev1 with B; assumption. apply cart_prod_fin_rev2 with A; assumption. Qed. (*Wyckoff*) Lemma cart_prod_fin_rev_or : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), Finite (cart_prod A B) -> Finite A \/ Finite B. intros T U A B h1. destruct (classic (Inhabited A)) as [h2 | h3]. right. apply cart_prod_fin_rev2 with A; assumption. left. assert (h4:A = Empty_set _). apply Extensionality_Ensembles;split. red. intros x h4. pose proof (Inhabited_intro _ A _ h4). contradiction. auto with sets. rewrite h4. constructor. Qed. (*Wyckoff*) Lemma cart_prod_proj1_surj : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), let f := (fun pr:T*U => fst pr) in Inhabited B -> Im (cart_prod A B) f = A. intros T U A B f h1. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h2. destruct h2 as [x h2]. subst. destruct h2 as [h2]. destruct h2 as [h2l h2r]. unfold f. simpl. assumption. (* >= *) red. intros x h2. unfold f. destruct h1 as [y h1]. apply Im_intro with (x, y). split; simpl; split; assumption. simpl. reflexivity. Qed. (*Wyckoff*) Lemma cart_prod_proj2_surj : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), let f := (fun pr:T*U => snd pr) in Inhabited A -> Im (cart_prod A B) f = B. intros T U A B f h1. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h2. destruct h2 as [x h2]. subst. destruct h2 as [h2]. destruct h2 as [h2l h2r]. unfold f. simpl. assumption. (* >= *) red. intros y h2. unfold f. destruct h1 as [x h1]. apply Im_intro with (x, y). split; simpl; split; assumption. simpl. reflexivity. Qed. (*Wyckoff*) Lemma cart_prod_im : forall {T U V W:Type} (A:Ensemble T) (B:Ensemble U) (f:T->V) (g:U->W), cart_prod (Im A f) (Im B g) = Im (cart_prod A B) (fun pr => (f (fst pr), g (snd pr))). intros T U V W A B f g. apply Extensionality_Ensembles. red. split. (* <= *) red. intros pr h1. inversion h1 as [h2]. clear h1. destruct h2 as [h2l h2r]. inversion h2l as [x h3 v h4 h5]. clear h2l. inversion h2r as [y h6 w h7 h8]. clear h2r. rewrite (surjective_pairing pr). subst. rewrite h4. rewrite h7. apply Im_intro with (x, y). constructor; auto. simpl. reflexivity. (* >= *) red. intros pr h1. inversion h1 as [prx h3 pry h4 h5]. clear h1. subst. rewrite (surjective_pairing prx) in h3. destruct h3 as [h3]. simpl in h3. destruct h3 as [h3l h3r]. constructor. simpl. split. apply Im_intro with (fst prx). assumption. reflexivity. apply Im_intro with (snd prx). assumption. reflexivity. Qed. (*Wyckoff*) Lemma full_set_prod : forall {T U:Type}, Full_set (T*U) = cart_prod (Full_set T) (Full_set U). intros T U. apply Extensionality_Ensembles. red. split. red. intros pr h1. constructor. split; constructor. red. intros pr h1. constructor. Qed. End CartProd. Section Relations. (*Wyckoff*) Definition dom_rel {T U:Type} (S:Ensemble (T*U)) : Ensemble T := [x:T | exists y:U, In S (x, y)]. (*Wyckoff*) Definition ran_rel {T U:Type} (S:Ensemble (T*U)) : Ensemble U := [y:U | exists x:T, In S (x, y)]. (*Wyckoff*) Lemma dom_rel_eq : forall {T U:Type} (S:Ensemble (T*U)), dom_rel S = Im S (@fst _ _). intros T U S. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [h1]. destruct h1 as [y h1]. apply Im_intro with (x, y); auto. red. intros x h1. destruct h1 as [pr h1 x]. subst. constructor. exists (snd pr). rewrite (surjective_pairing pr) in h1. assumption. Qed. (*Wyckoff*) Lemma ran_rel_eq : forall {T U:Type} (S:Ensemble (T*U)), ran_rel S = Im S (@snd _ _). intros T U S. apply Extensionality_Ensembles. red. split. red. intros y h1. destruct h1 as [h1]. destruct h1 as [x h1]. apply Im_intro with (x, y); auto. red. intros x h1. destruct h1 as [pr h1 y]. subst. constructor. exists (fst pr). rewrite (surjective_pairing pr) in h1. assumption. Qed. (*Wyckoff*) Lemma dom_rel_empty : forall (T U:Type), dom_rel (Empty_set (T*U)) = Empty_set _. intros T U. rewrite dom_rel_eq. apply image_empty. Qed. (*Wyckoff*) Lemma ran_rel_empty : forall (T U:Type), ran_rel (Empty_set (T*U)) = Empty_set _. intros T U. rewrite ran_rel_eq. apply image_empty. Qed. (*Wyckoff*) Lemma dom_rel_singleton : forall {T U:Type} (pr:T*U), dom_rel (Singleton pr) = Singleton (fst pr). intros T U pr. apply Extensionality_Ensembles. red; split. red; intros x h1. destruct h1 as [h1]. destruct h1 as [y h1]. inversion h1. subst. clear h1. simpl. constructor. red. intros x h1. destruct h1. constructor. exists (snd pr). rewrite <- surjective_pairing. constructor. Qed. (*Wyckoff*) Lemma ran_rel_singleton : forall {T U:Type} (pr:T*U), ran_rel (Singleton pr) = Singleton (snd pr). intros T U pr. apply Extensionality_Ensembles. red; split. red; intros x h1. destruct h1 as [h1]. destruct h1 as [y h1]. inversion h1. subst. clear h1. simpl. constructor. red. intros x h1. destruct h1. constructor. exists (fst pr). rewrite <- surjective_pairing. constructor. Qed. (*Wyckoff*) Lemma dom_rel_couple : forall {T U:Type} (pr pr':T*U), dom_rel (Couple pr pr') = Couple (fst pr) (fst pr'). intros T U pr pr'. apply Extensionality_Ensembles. red; split. red; intros x h1. destruct h1 as [h1]. destruct h1 as [y h1]. inversion h1. subst. clear h1. simpl. left. subst. simpl. right. red. intros x h1. destruct h1. constructor. exists (snd pr). rewrite <- surjective_pairing. left. constructor. exists (snd pr'). rewrite <- surjective_pairing. right. Qed. (*Wyckoff*) Lemma ran_rel_couple : forall {T U:Type} (pr pr':T*U), ran_rel (Couple pr pr') = Couple (snd pr) (snd pr'). intros T U pr pr'. apply Extensionality_Ensembles. red; split. red; intros x h1. destruct h1 as [h1]. destruct h1 as [y h1]. inversion h1. subst. clear h1. simpl. left. subst. simpl. right. red. intros x h1. destruct h1. constructor. exists (fst pr). rewrite <- surjective_pairing. left. constructor. exists (fst pr'). rewrite <- surjective_pairing. right. Qed. (*Wyckoff*) Lemma dom_rel_finite : forall {T U:Type} (S:Ensemble (T*U)), Finite S -> Finite (dom_rel S). intros T U S h1. rewrite dom_rel_eq. apply finite_image; auto. Qed. (*Wyckoff*) Lemma ran_rel_finite : forall {T U:Type} (S:Ensemble (T*U)), Finite S -> Finite (ran_rel S). intros T U S h1. rewrite ran_rel_eq. apply finite_image; auto. Qed. (*Wyckoff*) Lemma dom_rel_add : forall {T U:Type} (S:Ensemble (T*U)) (pr:T*U), dom_rel (Add S pr) = Add (dom_rel S) (fst pr). intros T U S pr. unfold dom_rel. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [h1]. destruct h1 as [y h1]. inversion h1 as [pr' h1l | pr' h1r]. subst. left. constructor. exists y. assumption. inversion h1r. subst. simpl. right. constructor. red. intros x h1. destruct h1 as [x h1l | x h1r]. constructor. destruct h1l as [h1l]. destruct h1l as [y h1l]. exists y. left. assumption. inversion h1r. subst. constructor. exists (snd pr). right. rewrite surjective_pairing at 1. constructor. Qed. (*Wyckoff*) Lemma ran_rel_add : forall {T U:Type} (S:Ensemble (T*U)) (pr:T*U), ran_rel (Add S pr) = Add (ran_rel S) (snd pr). intros T U S pr. unfold dom_rel. apply Extensionality_Ensembles. red. split. red. intros y h1. destruct h1 as [h1]. destruct h1 as [x h1]. inversion h1 as [pr' h1l | pr' h1r]. subst. left. constructor. exists x. assumption. inversion h1r. subst. simpl. right. constructor. red. intros y h1. destruct h1 as [y h1l | y h1r]. constructor. destruct h1l as [h1l]. destruct h1l as [x h1l]. exists x. left. assumption. inversion h1r. subst. constructor. exists (fst pr). right. rewrite surjective_pairing at 1. constructor. Qed. (*Wyckoff*) Lemma dom_rel_incl : forall {T U:Type} (S S':Ensemble (T*U)), Included S S' -> Included (dom_rel S) (dom_rel S'). intros T U S S' h1. red. intros x h2. destruct h2 as [h2]. destruct h2 as [y h2]. constructor. exists y. auto with sets. Qed. (*Wyckoff*) Lemma ran_rel_incl : forall {T U:Type} (S S':Ensemble (T*U)), Included S S' -> Included (ran_rel S) (ran_rel S'). intros T U S S' h1. red. intros y h2. destruct h2 as [h2]. destruct h2 as [x h2]. constructor. exists x. auto with sets. Qed. (*Wyckoff*) Definition strict_included_exist_rel {T:Type} (A:Ensemble T) : {C:Ensemble T | Included C A} -> {C:Ensemble T | Included C A} -> Prop := exist_rel (@Strict_Included T) (fun C => Included C A). (*Wyckoff*) Definition decreasing_or_zero {T:Type} (R:T->T->Prop) (f:T->T) : Prop := exists z:T, f z = z /\ (forall x:T, R (f x) x \/ x = z). (*Wyckoff*) Definition decreasing_or_zero' {T U:Type} (R:T->T->Prop) (f:T->T) (z:T) : Prop := f z = z /\ (forall x:T, R (f x) x \/ x = z). (*Wyckoff*) Definition decreasing_or_zero2 {T U:Type} (R:U->U->Prop) (f:T*U->T*U) (z:U) : Prop := (forall x:T, snd (f (x, z)) = z )/\ (forall pr:T*U, R (snd (f pr)) (snd pr) \/ snd (f pr) = z). (*Wyckoff*) (*Placed here for dependency issues.*) Lemma iter_uo_decreasing_or_zero : forall {T U:Type} (R:U->U->Prop) (f:T*U->T*U) (z:U), decreasing_or_zero2 R f z -> forall (pr:T*U) (n:nat), let x' := snd (iter_uo f pr (S n)) in let x := snd (iter_uo f pr n) in R x' x \/ x' = z. intros T U R f z h1 pr n. simpl. red in h1. destruct h1 as [h1 h2]. apply h2; auto. Qed. (*Wyckoff*) Definition rel_und_set {T:Type} (R:Relation T) := [y:T | exists x:T, R x y]. (*Wyckoff*) Definition foreset {T:Type} (R:Relation T) (y:T) := [x:T | R x y]. (*Wyckoff*) Definition afterset {T:Type} (R:Relation T) (x:T) := [y:T | R x y]. (*Wyckoff*) Definition foreset' {T:Type} (R:Relation T) (S:Ensemble T) (y:T) := [x:T | R x y /\ In S y]. (*Wyckoff*) Definition afterset' {T:Type} (R:Relation T) (S:Ensemble T) (x:T) := [y:T | R x y /\ In S x]. (*Wyckoff*) Definition ex_max_set {T:Type} (R:Relation T) (S:Ensemble T) : Prop := exists x:T, In S x /\ forall w:T, In S w -> R w x \/ w = x. (*Wyckoff*) Definition ex_unq_max_set {T:Type} (R:Relation T) (S:Ensemble T) : Prop := exists! x:T, In S x /\ forall w:T, In S w -> R w x \/ w = x. (*Wyckoff*) Lemma min_set_sortable_ex : forall {T:Type} (S:Ensemble T), Finite S -> Inhabited S -> forall (R:T->T->Prop), sortable R -> exists! m:T, In S m /\ (forall x:T, In (Subtract S m) x -> R m x). intros T S h1. induction h1 as [|S h2 h3 a h4]. intro h1. destruct h1; contradiction. intros h1 R h60. inversion h60 as [h6 h6' h6'']. destruct (eq_dec S (Empty_set _)) as [h5 | h5]; subst. exists a. red. split. split. right; auto. constructor. intros x h7. rewrite add_empty_sing in h7. rewrite subtract_sing in h7. contradiction. intros x h7. destruct h7 as [h7 h8]. rewrite add_empty_sing in h7. destruct h7; auto. apply not_empty_Inhabited in h5. specialize (h3 h5 R h60). destruct h3 as [m h3]. red in h3. destruct h3 as [h3a h3b]. red in h6. specialize (h6 a m). destruct h6 as [h6 | h6]. exists a. red; split. split. right; auto. constructor. intros x h7. rewrite sub_add_same_nin in h7; auto. destruct h3a as [h3a h3c]. destruct (eq_dec x m) as [h8 | h8]. subst; auto. assert (h9:In (Subtract S m) x). constructor; auto. intro h9. destruct h9. contradict h8; auto. specialize (h3c _ h9). eapply h6'. apply h6. assumption. intros x h7. destruct h7 as [h7 h8]. destruct h7 as [x h7 | x h9]. assert (h9: In S x /\ (forall x0 : T, In (Subtract S x) x0 -> R x x0)). split; auto. intros x' h9. assert (h10:In (Subtract (Add S a) x) x'). destruct h9 as [h9 h10]. constructor. left; auto. auto. apply h8; auto. specialize (h3b _ h9). subst. specialize (h8 a). assert (h10: In (Subtract (Add S a) x) a). constructor. right. constructor. intro h10. destruct h10. apply h6'' in h6. contradiction. specialize (h8 h10). red in h6'. specialize (h6' _ _ _ h8 h6). apply h6'' in h6'. contradiction. destruct h9; auto. destruct h3a as [h3a h3c]. exists m. red; split. split. left; auto. intros x h9. destruct h9 as [h9 h10]. destruct h9 as [x h9 | x h11]. apply h3c. constructor; auto. destruct h11. destruct h6 as [h6 | h6]. subst. contradict h10; constructor; auto. assumption. intros x' h9. destruct h9 as [h9 h10]. destruct h9 as [x' h9 | x' h9]. assert (h11:In S x' /\ (forall x : T, In (Subtract S x') x -> R x' x)). split; auto. intros x h11. destruct h11 as [h11 h12]. apply h10. constructor. left; auto. assumption. apply h3b; auto. destruct h9. destruct h6 as [h6 | h11]; subst; auto. assert (h12:In (Subtract (Add S a) a) m). rewrite sub_add_same_nin; auto. specialize (h10 _ h12). specialize (h6' _ _ _ h10 h11). apply h6'' in h6'. contradiction. Qed. (*Wyckoff*) Definition min_set_sortable {T:Type} (S:Ensemble T) (pff:Finite S) (pfi: Inhabited S) (R:T->T->Prop) (pfso:sortable R) := proj1_sig (constructive_definite_description _ (min_set_sortable_ex S pff pfi R pfso)). (*Wyckoff*) Lemma min_set_sortable_compat : forall {T:Type} (S:Ensemble T) (pff:Finite S) (pfi: Inhabited S) (R:T->T->Prop) (pfso: sortable R), let m := min_set_sortable S pff pfi R pfso in In S m /\ (forall x:T, In (Subtract S m) x -> R m x). intros. unfold m, min_set_sortable. destruct constructive_definite_description; simpl; auto. Qed. (*Wyckoff*) Lemma min_set_sortable_functional : forall {T:Type} (S S':Ensemble T) (pff:Finite S) (pff':Finite S') (pfi: Inhabited S) (pfi':Inhabited S') (R:T->T->Prop) (pfso pfso':sortable R), S = S' -> min_set_sortable S pff pfi R pfso = min_set_sortable S' pff' pfi' R pfso'. intros T S S' h1 h2 h3 h4 R h5 h5' h6. subst. pose proof (proof_irrelevance _ h1 h2). pose proof (proof_irrelevance _ h3 h4). subst. apply f_equal_prop_dep; auto. Qed. Lemma min_set_sortable_sing : forall {T:Type} (a:T) (pff:Finite (Singleton a)) (pfi:Inhabited (Singleton a)) (R:T->T->Prop) (pfso:sortable R), min_set_sortable (Singleton a) pff pfi R pfso = a. intros T a h1 h2 R h3. pose proof (min_set_sortable_compat _ h1 h2 R h3) as h6. simpl in h6. destruct h6 as [h6]. inversion h6; auto. Qed. Lemma min_set_sortable_add : forall {T:Type} (A:Ensemble T) (a:T) (pff:Finite (Add A a)) (pfi:Inhabited (Add A a)) (R:T->T->Prop) (pfso:sortable R), let (pftso, _, _) := pfso in let pftso' := tso_dec_compat _ pftso in (forall (pfi':Inhabited A), let pff' := Finite_downward_closed _ _ pff _ (incl_add A a) in min_set_sortable (Add A a) pff pfi R pfso = minR R pftso' (min_set_sortable A pff' pfi' R pfso) a) /\ (A = Empty_set T -> min_set_sortable (Add A a) pff pfi R pfso = a). intros T A a h1 h2 R h3. destruct h3 as [h3 h4 h5]. split. intro h6. simpl. pose proof (min_set_sortable_compat A (Finite_downward_closed T (Add A a) h1 A (incl_add A a)) h6 R (sort_intro _ h3 h4 h5)) as h7. simpl in h7. destruct h7 as [h7 h8]. unfold min_set_sortable. destruct constructive_definite_description as [m h10]. simpl. destruct constructive_definite_description as [m' h11]. simpl. destruct h10 as [h10 h12]. destruct h11 as [h11 h13]. destruct h10 as [m h10|m h10]; subst. unfold minR. destruct (tso_dec_compat R h3 m' a) as [h14 | h14]. red in h3. pose proof (h3 m m') as h3'. destruct h3' as [h3' | h3']. apply NNPP. intro h15. specialize (h13 m). assert (h16:In (Subtract A m') m). constructor; auto. intro h16. destruct h16. red in h5. pose proof (h5 m'). contradiction. specialize (h13 h16). red in h4. pose proof (h4 m m' m h3' h13) as h4'. red in h5. pose proof (h5 m) as h5'. contradiction. destruct h3' as [|h3']; auto. specialize (h12 m'). assert (h16:In (Subtract (Add A a) m) m'). constructor. left; auto. intro h16. destruct h16. red in h5. pose proof (h5 m). contradiction. specialize (h12 h16). red in h4. pose proof (h4 _ _ _ h12 h3'). red in h5. pose proof (h5 m). contradiction. red in h3. pose proof (h3 m a) as h3'. destruct h3' as [h3' | h3']. pose proof (h13 m) as h13'. red in h4. pose proof (h4 _ _ _ h3' h14) as h15. hfirst h13'. constructor; auto. rewrite in_sing_iff. intro; subst. apply h5 in h15; auto. specialize (h13' hf). pose proof (h4 _ _ _ h13' h15) as h16. apply h5 in h16; contradiction. destruct h3' as [|h3']; auto. specialize (h12 a). hfirst h12. constructor. right; auto. constructor; auto. rewrite in_sing_iff. intro; subst. apply h5 in h3'; auto. specialize (h12 hf). pose proof (h4 _ _ _ h12 h3') as h15. apply h5 in h15; contradiction. destruct h10. unfold minR. destruct (tso_dec_compat R h3 m' a) as [h14 | h14]. destruct h14 as [h14|]; subst. specialize (h12 m'). assert (h15: In (Subtract (Add A a) a) m'). constructor. left; auto. intro h16. destruct h16. pose proof (h5 a). contradiction. specialize (h12 h15). pose proof (h4 _ _ _ h12 h14). pose proof (h5 a). contradiction. reflexivity. reflexivity. intro. subst. pose proof (add_empty_sing a) as h6. erewrite min_set_sortable_functional. apply min_set_sortable_sing. assumption. Grab Existential Variables. eapply sort_intro; auto. apply inh_sing. apply Singleton_is_finite. Qed. (*Wyckoff*) Lemma min_set_sortable_couple : forall {T:Type} (a b:T) (pff:Finite (Couple a b)) (pfi:Inhabited (Couple a b)) (R:T->T->Prop) (pfso:sortable R), R a b -> min_set_sortable (Couple a b) pff pfi R pfso = a. intros T a b hf hi R h1 h4. destruct h1 as [h1 h2 h3]. pose proof (couple_add_sing a b) as h5. assert (h6:Finite (Add (Singleton a) b)). rewrite <- h5. apply finite_couple. pose proof (min_set_sortable_functional _ _ hf h6 hi (inh_add _ _) R (sort_intro _ h1 h2 h3) (sort_intro _ h1 h2 h3) h5) as h7. rewrite h7. pose proof (min_set_sortable_add (Singleton a) b h6 (inh_add (Singleton a) b) R (sort_intro _ h1 h2 h3)) as h8. destruct h8 as [h8 h9]. specialize (h8 (inh_sing _)). simpl in h8. rewrite h8. unfold minR. destruct (tso_dec_compat R h1 (min_set_sortable (Singleton a) (Finite_downward_closed T (Add (Singleton a) b) h6 (Singleton a) (incl_add (Singleton a) b)) (inh_sing a) R (sort_intro _ h1 h2 h3))) as [h10 | h10]. rewrite min_set_sortable_sing. reflexivity. rewrite min_set_sortable_sing in h10. pose proof (h2 _ _ _ h4 h10) as h11. apply h3 in h11. contradiction. Qed. (*Wyckoff*) Lemma min_set_sortable_couple' : forall {T:Type} (a b:T) (pff:Finite (Couple a b)) (pfi:Inhabited (Couple a b)) (R:T->T->Prop) (pfso:sortable R), let (pftso, _, _) := pfso in let pftso' := tso_dec_compat _ pftso in min_set_sortable (Couple a b) pff pfi R pfso = minR R pftso' a b. intros T a b h1 h2 R h3. destruct h3 as [h3 h4 h5]. unfold minR. destruct (tso_dec_compat R h3 a b) as [h6 | h6]. destruct h6 as [h6|]; subst. apply min_set_sortable_couple; auto. pose proof (couple_singleton b) as h7. pose proof (min_set_sortable_functional _ _ h1 (Singleton_is_finite _ _) h2 (inh_sing _) R (sort_intro _ h3 h4 h5) (sort_intro _ h3 h4 h5) h7) as h8. rewrite h8. rewrite min_set_sortable_sing; auto. pose proof (couple_comm a b) as h7. pose proof (min_set_sortable_functional _ _ h1 (finite_couple _ _) h2 (inh_couple _ _) R (sort_intro _ h3 h4 h5) (sort_intro _ h3 h4 h5) h7) as h8. rewrite h8. apply min_set_sortable_couple; auto. Qed. (*Wyckoff*) Lemma max_set_sortable_ex : forall {T:Type} (S:Ensemble T), Finite S -> Inhabited S -> forall (R:T->T->Prop), sortable R -> exists! m:T, In S m /\ (forall x:T, In (Subtract S m) x -> R x m). intros T S h1. induction h1 as [|S h2 h3 a h4]. intro h1. destruct h1; contradiction. intros h1 R h60. inversion h60 as [h6 h6' h6'']. destruct (eq_dec S (Empty_set _)) as [h5 | h5]; subst. exists a. red. split. split. right; auto. constructor. intros x h7. rewrite add_empty_sing in h7. rewrite subtract_sing in h7. contradiction. intros x h7. destruct h7 as [h7 h8]. rewrite add_empty_sing in h7. destruct h7; auto. apply not_empty_Inhabited in h5. specialize (h3 h5 R h60). destruct h3 as [m h3]. red in h3. destruct h3 as [h3a h3b]. red in h6. specialize (h6 m a). destruct h6 as [h6 | h6]. exists a. red; split. split. right; auto. constructor. intros x h7. rewrite sub_add_same_nin in h7; auto. destruct h3a as [h3a h3c]. destruct (eq_dec x m) as [h8 | h8]. subst; auto. assert (h9:In (Subtract S m) x). constructor; auto. intro h9. destruct h9. contradict h8; auto. specialize (h3c _ h9). eapply h6'. apply h3c. assumption. intros x h7. destruct h7 as [h7 h8]. destruct h7 as [x h7 | x h9]. specialize (h3b x). hfirst h3b. split; auto. intros x' h9. specialize (h8 x'). hfirst h8. destruct h9 as [h9 h10]. constructor. left; auto. auto. apply h8; auto. specialize (h3b hf). subst. specialize (h8 a). hfirst h8. constructor. right. constructor. intro h10. destruct h10. apply h6'' in h6. contradiction. specialize (h8 hf0). red in h6'. specialize (h6' _ _ _ h8 h6). apply h6'' in h6'. contradiction. destruct h9; auto. destruct h3a as [h3a h3c]. exists m. red; split. split. left; auto. intros x h9. destruct h9 as [h9 h10]. destruct h9 as [x h9 | x h11]. apply h3c. constructor; auto. destruct h11. destruct h6 as [h6 | h6]. subst. contradict h10; constructor; auto. assumption. intros x' h9. destruct h9 as [h9 h10]. destruct h9 as [x' h9 | x' h9]. specialize (h3b x'). hfirst h3b. split; auto. intros x h11. destruct h11 as [h11 h12]. apply h10. constructor. left; auto. assumption. apply h3b; auto. destruct h9. destruct h6 as [h6 | h11]; subst; auto. specialize (h10 m). hfirst h10. rewrite sub_add_same_nin; auto. specialize (h10 hf). specialize (h6' _ _ _ h10 h11). apply h6'' in h6'. contradiction. Qed. (*Wyckoff*) Definition max_set_sortable {T:Type} (S:Ensemble T) (pff:Finite S) (pfi: Inhabited S) (R:T->T->Prop) (pfso:sortable R) := proj1_sig (constructive_definite_description _ (max_set_sortable_ex S pff pfi R pfso)). (*Wyckoff*) Lemma max_set_sortable_compat : forall {T:Type} (S:Ensemble T) (pff:Finite S) (pfi: Inhabited S) (R:T->T->Prop) (pfso: sortable R), let m := max_set_sortable S pff pfi R pfso in In S m /\ (forall x:T, In (Subtract S m) x -> R x m). intros. unfold m, max_set_sortable. destruct constructive_definite_description; simpl; auto. Qed. (*Wyckoff*) Lemma max_set_sortable_functional : forall {T:Type} (S S':Ensemble T) (pff:Finite S) (pff':Finite S') (pfi: Inhabited S) (pfi':Inhabited S') (R:T->T->Prop) (pfso pfso':sortable R), S = S' -> max_set_sortable S pff pfi R pfso = max_set_sortable S' pff' pfi' R pfso'. intros T S S' h1 h2 h3 h4 R h5 h5' h6. subst. pose proof (proof_irrelevance _ h1 h2). pose proof (proof_irrelevance _ h3 h4). subst. apply f_equal_prop_dep; auto. Qed. (*Wyckoff*) Lemma max_set_sortable_sing : forall {T:Type} (a:T) (pff:Finite (Singleton a)) (pfi:Inhabited (Singleton a)) (R:T->T->Prop) (pfso:sortable R), max_set_sortable (Singleton a) pff pfi R pfso = a. intros T a h1 h2 R h3. pose proof (max_set_sortable_compat _ h1 h2 R h3) as h6. simpl in h6. destruct h6 as [h6]. inversion h6; auto. Qed. (*Wyckoff*) Lemma max_set_sortable_add : forall {T:Type} (A:Ensemble T) (a:T) (pff:Finite (Add A a)) (pfi:Inhabited (Add A a)) (R:T->T->Prop) (pfso:sortable R), let (pftso, _, _) := pfso in let pftso' := tso_dec_compat _ pftso in (forall (pfi':Inhabited A), let pff' := Finite_downward_closed _ _ pff _ (incl_add A a) in max_set_sortable (Add A a) pff pfi R pfso = maxR R pftso' (max_set_sortable A pff' pfi' R pfso) a) /\ (A = Empty_set T -> max_set_sortable (Add A a) pff pfi R pfso = a). intros T A a h1 h2 R h3. simpl. destruct h3 as [h3 h4 h5]. split. intro h6. unfold max_set_sortable, maxR. destruct constructive_definite_description as [m h7]; simpl. destruct constructive_definite_description as [m' h8]; simpl. destruct h7 as [h7 h9], h8 as [h8 h10]. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. destruct hl as [hl | hl]. destruct h7 as [m h7 | m h7]; subst. pose proof (h10 m) as h10'. hfirst h10'. constructor; auto. intro h11; destruct h11. pose proof (h9 a) as h9'. hfirst h9'. constructor. right; constructor; auto. intro h11. destruct h11. apply h5 in hl; contradiction. specialize (h9' hf). pose proof (h4 _ _ _ hl h9') as h11. apply h5 in h11; contradiction. specialize (h10' hf). pose proof (h4 _ _ _ h10' hl) as h11. pose proof (h9 a) as h9'. hfirst h9'. constructor. right; constructor; auto. intro h12; destruct h12. apply h5 in h11; contradiction. specialize (h9' hf0). pose proof (h4 _ _ _ h11 h9') as h12. apply h5 in h12; contradiction. destruct h7; auto. subst. destruct h7 as [m h7 | m h7]. pose proof (h3 m a) as h3'. destruct h3' as [h3' | h3']; auto. specialize (h9 a). hfirst h9. constructor. right; constructor; auto. intro h11; destruct h11. apply h5 in h3'; contradiction. specialize (h9 hf). pose proof (h4 _ _ _ h9 h3') as h11. apply h5 in h11; contradiction. destruct h3' as [|h3']; auto. specialize (h10 m). hfirst h10. constructor; auto. intro h11; destruct h11. apply h5 in h3'; contradiction. specialize (h10 hf). pose proof (h4 _ _ _ h10 h3') as h11. apply h5 in h11; contradiction. destruct h7; auto. destruct h7 as [m h7 | m h7]; subst. pose proof (h3 m m') as h11. destruct h11 as [h11 | h11]. specialize (h9 m'). hfirst h9. constructor. left; auto. intro h12; destruct h12. apply h5 in h11; auto. specialize (h9 hf). pose proof (h4 _ _ _ h9 h11) as h12. apply h5 in h12; contradiction. destruct h11 as [|h11]; auto. specialize (h10 m). hfirst h10. constructor; auto. intro h12. destruct h12. apply h5 in h11; contradiction. specialize (h10 hf). pose proof (h4 _ _ _ h10 h11) as h12. apply h5 in h12; contradiction. destruct h7. specialize (h9 m'). hfirst h9. constructor. left; auto. intro h11; destruct h11. apply h5 in hr; contradiction. specialize (h9 hf). pose proof (h4 _ _ _ h9 hr) as h11. apply h5 in h11; contradiction. intro; subst. generalize h1, h2. rewrite add_empty_sing. intros h1' h2'. rewrite max_set_sortable_sing; auto. Qed. (*Wyckoff*) Lemma max_set_sortable_couple : forall {T:Type} (a b:T) (pff:Finite (Couple a b)) (pfi:Inhabited (Couple a b)) (R:T->T->Prop) (pfso:sortable R), R a b -> max_set_sortable (Couple a b) pff pfi R pfso = b. intros T a b hf hi R h1 h4. generalize hf hi. rewrite couple_add_sing. intros hf' hi'. pose proof (max_set_sortable_add (Singleton a) b hf' hi' R h1) as h8. simpl in h8. destruct h1 as [h1 h2 h3]. destruct h8 as [h8 h9]. specialize (h8 (inh_sing _)). simpl in h8. rewrite h8 at 1. unfold maxR. unfold maxR in h8. lr_if_goal'; auto. rewrite max_set_sortable_sing. rewrite max_set_sortable_sing in hr. pose proof (h2 _ _ _ h4 hr) as h11. apply h3 in h11; contradiction. Qed. (*Wyckoff*) Lemma max_set_sortable_couple' : forall {T:Type} (a b:T) (pff:Finite (Couple a b)) (pfi:Inhabited (Couple a b)) (R:T->T->Prop) (pfso:sortable R), let (pftso, _, _) := pfso in let pftso' := tso_dec_compat _ pftso in max_set_sortable (Couple a b) pff pfi R pfso = maxR R pftso' a b. intros T a b h1 h2 R h3. destruct h3 as [h3 h4 h5]. unfold maxR. destruct (tso_dec_compat R h3 a b) as [h6 | h6]. destruct h6 as [h6|]; subst. apply max_set_sortable_couple; auto. generalize h1 h2. rewrite couple_singleton. intros h1' h2'. rewrite max_set_sortable_sing; auto. generalize h1 h2. rewrite couple_comm. intros. rewrite max_set_sortable_couple; auto. Qed. Section FunctionallyPaired. Inductive functionally_paired {T U:Type} (A:Ensemble T) (B:Ensemble U) (S:Ensemble (T*U)) : Prop := functional_pairs_intro : (forall (x:T), In A x -> (exists! y:U, In B y /\ In S (x, y))) -> (forall (pr:T*U), In S pr -> In A (fst pr) /\ In B (snd pr)) -> functionally_paired A B S. Lemma fp_in_dom : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (S:Ensemble (T*U)), functionally_paired A B S -> forall pr:T*U, In S pr -> In A (fst pr). intros T U A B S h1 pr h2. destruct h1 as [? h1]. apply h1; assumption. Qed. Lemma fp_in_ran : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (S:Ensemble (T*U)), functionally_paired A B S -> forall pr:T*U, In S pr -> In B (snd pr). intros T U A B S h1 pr h2. destruct h1 as [? h1]. apply h1; assumption. Qed. Lemma fp_functional : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)}, functionally_paired A B S -> forall (x:T) (y1 y2:U), In S (x, y1) -> In S (x, y2) -> y1 = y2. intros T U A B S h1 x y1 y2 h3 h4. inversion h1 as [h5 h6]. pose proof (fp_in_dom _ _ _ h1 _ h3) as h7. simpl in h7. pose proof (h5 x h7) as h8. destruct h8 as [y h9]. red in h9. destruct h9 as [h9a h9b]. pose proof (fp_in_ran _ _ _ h1 _ h3) as h10. pose proof (fp_in_ran _ _ _ h1 _ h4) as h11. simpl in h10, h11. pose proof (h9b y1) as h12. pose proof (h9b y2) as h13. cut (y = y1). cut (y = y2). congruence. tauto. tauto. Qed. Definition fun_well_defined {T U:Type} (S:Ensemble (T*U)) := functionally_paired (dom_rel S) (ran_rel S) S. Lemma fun_well_defined_empty : forall (T U:Type), fun_well_defined (Empty_set (T*U)). intros T U. red. rewrite dom_rel_empty. rewrite ran_rel_empty. constructor; intros; contradiction. Qed. Lemma fun_well_defined_im: forall {T U:Type} (A:Ensemble T) (f:T->U), fun_well_defined (Im A (fun x=>(x, f x))). intros T U A f. red. constructor. intros x h1. rewrite dom_rel_eq in h1. rewrite im_im in h1. simpl in h1. destruct h1. subst. exists (f x). red. split. split. constructor. exists x. apply Im_intro with x. assumption. reflexivity. apply Im_intro with x. assumption. reflexivity. intros y h1. destruct h1 as [h1l h1r]. rewrite ran_rel_eq in h1l. rewrite im_im in h1l. simpl in h1l. destruct h1l. subst. inversion h1r. inversion H2. reflexivity. intros pr h1. split. rewrite dom_rel_eq. rewrite im_im. simpl. destruct h1. subst. simpl. apply Im_intro with x; auto. rewrite ran_rel_eq. rewrite im_im. simpl. destruct h1. subst. simpl. apply Im_intro with x; auto. Qed. Lemma fun_well_defined_incl : forall {T U:Type} (R S:Ensemble (T*U)), Included R S -> fun_well_defined S -> fun_well_defined R. intros T U R S h1 h2. constructor. intros x h3. destruct h3 as [h3]. destruct h3 as [y h3]. exists y. red. split. split. constructor. exists x. assumption. assumption. intros y' h4. inversion h2 as [h2l h2r]. assert (h5:In S (x, y)). auto with sets. specialize (h2r _ h5). simpl in h2r. destruct h2r as [h2a h2b]. pose proof (h2l _ h2a) as h2'. destruct h2' as [y'' h2']. red in h2'. destruct h2' as [h6 h7]. specialize (h7 _ (conj h2b h5)). subst. destruct h4 as [h4l h4r]. destruct h6 as [h6l h6r]. assert (h7: In S (x, y')). auto with sets. red in h2. pose proof (fp_functional h2 _ _ _ h6r h7). assumption. intros pr h3. split. constructor. exists (snd pr). rewrite (surjective_pairing pr) in h3. assumption. constructor. exists (fst pr). rewrite (surjective_pairing pr) in h3. assumption. Qed. Lemma fp_sub_add : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)} (x:T) (y y':U), y <> y' -> functionally_paired A B (Add S (x, y)) -> Subtract (Add S (x, y)) (x, y') = Add S (x, y). intros T U A B S x y y' h0 h1. apply Extensionality_Ensembles. red. split. apply incl_subtract. red. intros pr h2. constructor. assumption. intro h3. inversion h3. subst. clear h3. pose proof (Add_intro2 _ S (x, y)) as h3. pose proof (fp_functional h1 _ _ _ h2 h3). subst. contradict h0. reflexivity. Qed. Lemma fp_dom_rel : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)}, functionally_paired A B S -> dom_rel S = A. intros T U A B S h1. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [h2]. destruct h2 as [y h2]. destruct h1 as [h1l h1r]. specialize (h1r _ h2). simpl in h1r. destruct h1r; auto. red. intros x h2. constructor. destruct h1 as [h1l h1r]. specialize (h1l _ h2). destruct h1l as [y h1l]. red in h1l. destruct h1l as [[h1l ?] ?]. exists y; auto. Qed. Lemma fp_add : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)} (pf:functionally_paired A B S) (x:T) (y:U), ~ In A x -> In B y -> ~ In S (x, y) -> functionally_paired (Add A x) B (Add S (x, y)). intros T U A B S h1 x y h2 h3 h4. constructor. intros x' h5. destruct h5 as [x' h5 | x' h6]. destruct h1 as [h1l h1r]. specialize (h1l _ h5). destruct h1l as [y' h1l]. red in h1l. exists y'. red. destruct h1l as [h1a h1b]. split. destruct h1a as [h1aa h1ab]. split; auto. left. assumption. intros y'' h6. destruct h6 as [h6l h6r]. inversion h6r. subst. specialize (h1b _ (conj h6l H)). assumption. subst. inversion H. subst. contradiction. inversion h6. subst. exists y. red. split. split; try right; auto. constructor. intros y' h7. destruct h7 as [h7l h7r]. inversion h7r. subst. destruct h1 as [h1l h1r]. specialize (h1r _ H). simpl in h1r. destruct h1r; contradiction. subst. inversion H. reflexivity. intros pr h5. destruct h5 as [pr h5l |pr h5r]. destruct h1 as [h1l h1r]. specialize (h1r _ h5l). destruct h1r as [h6 h7]. split; try left; auto. inversion h5r. subst. simpl. split; try right; auto. constructor. Qed. Lemma fun_well_defined_add : forall {T U:Type} {S:Ensemble (T*U)}, fun_well_defined S -> forall (x:T) (y:U), ~In (dom_rel S) x -> fun_well_defined (Add S (x, y)). intros T U S h1 x y h2. red. constructor. intros a h3. rewrite dom_rel_add in h3. simpl in h3. destruct h3 as [a h3l | a h3r]. destruct h3l as [h3l]. destruct h3l as [b h3l]. exists b. red. split. split. rewrite ran_rel_add. simpl. left. constructor. exists a. assumption. left. assumption. intros y' h4. destruct h4 as [h4l h4r]. inversion h4r as [? h5 | ? h6]. subst. apply (fp_functional h1 _ _ _ h3l h5). subst. inversion h6. subst. pose proof (fp_in_dom _ _ _ h1 _ h3l) as h7. simpl in h7. contradiction. inversion h3r. subst. clear h3r. exists y. red. split. split. rewrite ran_rel_add. simpl. right. constructor. right. constructor. intros y' h3. destruct h3 as [h3l h3r]. inversion h3r as [? h3a | ? h3b]. subst. pose proof (fp_in_dom _ _ _ h1 _ h3a) as h4. simpl in h4. contradiction. subst. inversion h3b. reflexivity. intros pr h3. split. constructor. exists (snd pr). rewrite <- surjective_pairing. assumption. constructor. exists (fst pr). rewrite <- surjective_pairing. assumption. Qed. Lemma fun_well_defined_subtract : forall {T U:Type} {S:Ensemble (T*U)}, fun_well_defined S -> forall pr:T*U, fun_well_defined (Subtract S pr). intros T U S h1 pr. red. constructor. intros x h2. destruct h2 as [h2]. destruct h2 as [y h2]. inversion h2 as [h3 h4]. red in h1. exists y. red. split. split. constructor. exists x. assumption. assumption. intros y' h5. destruct h5 as [h5l h5r]. inversion h5r as [h6 h7]. apply (fp_functional h1 _ _ _ h3 h6). intros pr' h2. split. constructor. exists (snd pr'). rewrite <- surjective_pairing. assumption. constructor. exists (fst pr'). rewrite <- surjective_pairing. assumption. Qed. Lemma fp_cart_prod_sing : forall {T U:Type} (A:Ensemble T) (b:U), functionally_paired A (Singleton b) (cart_prod A (Singleton b)). intros T U A b. constructor. intros x h1. exists b. red. split. split. constructor. constructor. split. simpl. assumption. simpl. constructor. intros u h2. destruct h2 as [h2l h2r]. destruct h2l; subst. reflexivity. intros pr h1. destruct h1 as [h1]. assumption. Qed. Lemma fps_inc_cart_prod : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)}, functionally_paired A B S -> Included S (cart_prod A B). intros T U A B S h1. red. intros pr h2. constructor. destruct h1 as [h3 h4]. apply h4. assumption. Qed. Lemma functionally_paired_restriction : forall {T U:Type} (A A':Ensemble T) (B:Ensemble U) (S':Ensemble (T*U)), Included A A' -> let S := [pr : T * U | In S' pr /\ In A (fst pr)] : Ensemble (T * U) in functionally_paired A' B S' -> functionally_paired A B S. intros T U A A' B S' h1 S h2. destruct h2 as [h2l h2r]. constructor. intros x h3. assert (h4:In A' x). auto with sets. specialize (h2l _ h4). destruct h2l as [y h2l]. red in h2l. destruct h2l as [h2a h2b]. destruct h2a as [h2a h2a']. exists y. red. split. unfold S. split; auto. constructor. split; auto. intros u h5. destruct h5 as [h5l h5r]. specialize (h2b u). destruct h5r as [h5r]. simpl in h5r. destruct h5r as [h5ra h5rb]. specialize (h2b (conj h5l h5ra)). assumption. intros pr h3. destruct h3 as [h3]. destruct h3 as [h3l h3r]. specialize (h2r _ h3l). destruct h2r. split; auto. Qed. Lemma fp_empty1 : forall (T U:Type) (B:Ensemble U), functionally_paired (Empty_set T) B (Empty_set (T*U)). intros; constructor; intros; contradiction. Qed. Lemma fp_cart_empty11 : forall {T U V:Type} (A:Ensemble U) (B:Ensemble V), functionally_paired (cart_prod (Empty_set T) A) B (Empty_set (T*U*V)). intros T U V A B. constructor. intros ? h1. destruct h1 as [h1]. destruct h1; contradiction. intros; contradiction. Qed. Lemma fp_cart_empty21 : forall {T U V:Type} (A:Ensemble T) (B:Ensemble V), functionally_paired (cart_prod A (Empty_set U)) B (Empty_set (T*U*V)). intros T U V A B. constructor. intros ? h1. destruct h1 as [h1]. destruct h1; contradiction. intros; contradiction. Qed. Lemma fp_empty1_s : forall {T U:Type} (B:Ensemble U) (S:Ensemble (T*U)), functionally_paired (Empty_set T) B S -> S = Empty_set (T*U). intros T U B S h1. apply Extensionality_Ensembles; red; split; auto with sets. red. intros pr h2. destruct h1 as [h3 h4]. specialize (h4 _ h2). destruct h4; contradiction. Qed. Lemma no_fp_empty2 : forall {T U:Type} (A:Ensemble T) (S:Ensemble (T*U)), functionally_paired A (Empty_set U) S -> A = Empty_set _. intros T U A S h1. apply Extensionality_Ensembles. red; split; auto with sets; red. intros x h2. destruct h1 as [h3 h4]. specialize (h3 _ h2). destruct h3 as [? h3]. red in h3. destruct h3 as [h3]. destruct h3; contradiction. Qed. Lemma no_fp_empty2' : forall {T U:Type} (A:Ensemble T) (S:Ensemble (T*U)), functionally_paired A (Empty_set U) S -> S = Empty_set _. intros T U A S h1. apply Extensionality_Ensembles. red; split; auto with sets; red. intros pr h2. destruct h1 as [h3 h4]. specialize (h4 _ h2). destruct h4; contradiction. Qed. End FunctionallyPaired. Section RelExt. Definition graph {T U:Type} (f:T->U) : Ensemble (T*U) := Im (Full_set T) (fun c => (c, f c)). Definition graph_seg {T U:Type} {A:Ensemble T} (f:sig_set A->U) : Ensemble (T * U) := Im (full_sig A) (fun c => (proj1_sig c, f c)). Definition rel_ext {T U:Type} (R:Ensemble (T*U)) (f:T->U) : Prop := Included (graph f) R. Definition rel_ext_seg {T U:Type} {A:Ensemble T} (R:Ensemble (T*U)) (f:sig_set A->U) : Prop := Included (graph_seg f) R. End RelExt. Section PopEns. (*Wyckoff*) Definition pop_ens {T:Type} (S:Ensemble T) (pff:Finite S) (R:T->T->Prop) (pfso:sortable R) (def:T) : T*Ensemble T. pose proof pff as h1. apply finite_inh_or_empty in h1. apply or_to_sumbool in h1. destruct h1 as [h1 | h1]. pose (min_set_sortable S pff h1 R pfso) as m. refine (m, Subtract S m). refine (def, Empty_set _). Defined. (*Wyckoff*) Lemma incl_pop_ens : forall {T:Type} (S:Ensemble T) (pff:Finite S) (R:T->T->Prop) (pfso:sortable R) (def:T) (C:{A:Ensemble T | Included A S}), let pfc := Finite_downward_closed _ _ pff _ (proj2_sig C) in Included (snd (pop_ens (proj1_sig C) pfc R pfso def)) S. intros T S h1 R h2 def C h0. destruct h2 as [h2 h3 h4]. red. intros x h5. unfold pop_ens in h5. destruct or_to_sumbool as [h6 | h6]. simpl in h5. destruct h5 as [h5 h7]. apply (proj2_sig C); auto. simpl in h5. contradiction. Qed. (*Wyckoff*) Lemma pop_ens_functional : forall {T:Type} (S S':Ensemble T) (pff:Finite S) (pff':Finite S') (R:T->T->Prop) (pfso:sortable R) (def:T), S = S' -> pop_ens S pff R pfso def = pop_ens S' pff' R pfso def. intros T S S' h1 h2 R h3 def h6. destruct h3 as [h3 h4 h5]. subst. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. (*Wyckoff*) Lemma pop_ens_empty : forall {T:Type} (pff:Finite (Empty_set _)) (R:T->T->Prop) (pfso:sortable R) (def:T), pop_ens (Empty_set _) pff R pfso def = (def, Empty_set _). intros T h1 R h2 def. destruct h2 as [h2 h3 h4]. unfold pop_ens. destruct or_to_sumbool as [h5 | h5]. destruct h5 as [c h5]. destruct h5. reflexivity. Qed. (*Wyckoff*) Lemma pop_ens_sing : forall {T:Type} (a:T) (pff:Finite (Singleton a)) (R:T->T->Prop) (pfso:sortable R) (def:T), pop_ens (Singleton a) pff R pfso def = (a, Empty_set _). intros T a h1 R h3 def. destruct h3 as [h2 h3 h4]. unfold pop_ens. destruct or_to_sumbool as [h5 | h5]. rewrite min_set_sortable_sing. f_equal. rewrite subtract_sing; auto. pose proof (In_singleton _ a) as h6. rewrite h5 in h6. contradiction. Qed. (*Wyckoff*) Lemma pop_ens_couple : forall {T:Type} (a b:T) (pff:Finite (Couple a b)) (R:T->T->Prop) (pfso:sortable R) (def:T), let (pftso, _, _) := pfso in let pftso' := tso_dec_compat _ pftso in pop_ens (Couple a b) pff R pfso def = (minR R pftso' a b, Subtract (Couple a b) (minR R pftso' a b)). intros T a b h1 R h2 def. destruct h2 as [h2 h3 h4]. unfold pop_ens. destruct or_to_sumbool as [h5 | h5]. rewrite (min_set_sortable_couple' _ _ _ _ _ (sort_intro R h2 h3 h4)); auto. pose proof (Couple_l _ a b) as h6. rewrite h5 in h6. contradiction. Qed. (*Wyckoff*) (*This looks much more complicated than it is and basically will be used in an application of [iter_uo] in the next function to do the iteration.*) Definition pop_ens_iter_aux {T:Type} (S:Ensemble T) (pff:Finite S) (pfi:Inhabited S) (R:T->T->Prop) (pfso: sortable R) (def:T) : T * {A:Ensemble T | Included A S} -> T * {A:Ensemble T | Included A S} := let m := min_set_sortable S pff pfi R pfso in (fun C':T*{A:Ensemble T | Included A S} => let (_, C) := C' in let (x, _) := pop_ens (proj1_sig C) (Finite_downward_closed _ _ pff _ (proj2_sig C)) R pfso def in let pf := incl_pop_ens S pff R pfso def C in (x, exist (fun D => Included D S) _ pf)). (*Wyckoff*) Lemma pop_ens_iter_aux_sing : forall {T:Type} (a:T) (pff:Finite (Singleton a)) (pfi:Inhabited (Singleton a)) (R:T->T->Prop) (pfso:sortable R) (def:T) , pop_ens_iter_aux (Singleton a) pff pfi R pfso def = (fun pr:(T * {A : Ensemble T | Included A (Singleton a)}) => if (or_to_sumbool _ _ (singleton_inc _ _ (proj2_sig (snd pr)))) then (def, (exist (fun A => Included A (Singleton a)) (Empty_set _) (incl_empty _))) else (a, (exist (fun A => Included A (Singleton a)) (Empty_set _) (incl_empty _)))). intros T a h1 h2 R h3 def. destruct h3 as [h3 h4 h5]. apply functional_extensionality. intro C. destruct C as [c C]. destruct C as [C h6]. simpl. destruct or_to_sumbool as [h7 | h7]. subst. rewrite pop_ens_empty at 1. f_equal. apply proj1_sig_injective; simpl. rewrite pop_ens_empty at 1. simpl. reflexivity. subst. rewrite pop_ens_sing at 1. f_equal. apply proj1_sig_injective; simpl. rewrite pop_ens_sing. simpl. reflexivity. Qed. (*Wyckoff*) Lemma pop_ens_iter_aux_couple : forall {T:Type} (a b:T) (pff:Finite (Couple a b)) (pfi:Inhabited (Couple a b)) (R:T->T->Prop) (pfso:sortable R) (def:T), let (pftso, _, _) := pfso in let pftso' := tso_dec_compat _ pftso in pop_ens_iter_aux (Couple a b) pff pfi R pfso def = (fun pr:(T * {A : Ensemble T | Included A (Couple a b)}) => match (or_to_sumbool _ _ (incl_couple_inv _ _ _ (proj2_sig (snd pr)))) with | left _ => (def, (exist (fun A => Included A (Couple a b)) (Empty_set _) (incl_empty _))) | right P => match (or_to_sumbool _ _ P) with | left Q => (a, (exist (fun A => Included A (Couple a b)) (Empty_set _) (incl_empty _))) | right F => match (or_to_sumbool _ _ F) with | left W => (b, (exist (fun A => Included A (Couple a b)) (Empty_set _) (incl_empty _))) | right _ => ((minR R pftso' a b), (exist (fun A => Included A (Couple a b)) (Subtract (Couple a b) (minR R pftso' a b)) (incl_subtract _ _))) end end end). intros T a b h1 h2 R h3 def. destruct h3 as [h3 h4 h5]. apply functional_extensionality. intro pr. destruct pr as [x A]. destruct A as [A h6]. simpl. destruct or_to_sumbool as [h7 | h7]. subst. rewrite pop_ens_empty at 1. f_equal. apply proj1_sig_injective; simpl. rewrite pop_ens_empty. simpl. reflexivity. destruct or_to_sumbool as [h8 | h8]. subst. rewrite pop_ens_sing at 1. f_equal. apply proj1_sig_injective; simpl. rewrite pop_ens_sing. simpl; auto. destruct or_to_sumbool as [h9 | h9]. subst. rewrite pop_ens_sing at 1. f_equal. apply proj1_sig_injective; simpl. rewrite pop_ens_sing at 1. simpl; auto. subst. pose proof (pop_ens_couple a b (Finite_downward_closed T (Couple a b) h1 (Couple a b) h6) R (sort_intro R h3 h4 h5) def) as h9. simpl in h9. rewrite h9 at 1. f_equal. apply proj1_sig_injective; simpl. rewrite h9; simpl; auto. Qed. Definition pop_ens_iter {T:Type} (S:Ensemble T) (pff:Finite S) (pfi:Inhabited S) (R:T->T->Prop) (pfso:sortable R) (def:T) : nat -> T := let m := min_set_sortable S pff pfi R pfso in (fun n:nat => fst (iter_uo (pop_ens_iter_aux S pff pfi R pfso def) (m, (exist (fun C => Included C S) _ (incl_subtract S m))) n)). (*Wyckoff*) Lemma pop_ens_iter_functional : forall {T:Type} (S S':Ensemble T) (pff:Finite S) (pff':Finite S') (pfi:Inhabited S) (pfi': Inhabited S') (R:T->T->Prop) (pfso:sortable R) (def:T), S = S' -> pop_ens_iter S pff pfi R pfso def = pop_ens_iter S' pff' pfi' R pfso def. intros T S S' h1 h2 h3 h4 R h5 def h8. subst. pose proof (proof_irrelevance _ h1 h2). pose proof (proof_irrelevance _ h3 h4). subst. reflexivity. Qed. (*Wyckoff*) Lemma pop_ens_iter0 : forall {T:Type} (S:Ensemble T) (pff:Finite S) (pfi:Inhabited S) (R:T->T->Prop) (pfso:sortable R) (def:T), pop_ens_iter S pff pfi R pfso def 0 = min_set_sortable S pff pfi R pfso. intros; unfold pop_ens_iter; auto. Qed. (*Wyckoff*) Lemma pop_ens_iter_add0 : forall {T:Type} (A:Ensemble T) (a:T) (pff:Finite (Add A a)) (pfi:Inhabited (Add A a)) (R:T->T->Prop) (pfso:sortable R) (def:T), let (pftso, _, _) := pfso in let pftso' := tso_dec_compat _ pftso in (forall pf':Inhabited A, let pff' := Finite_downward_closed _ _ pff _ (incl_add A a) in pop_ens_iter (Add A a) pff pfi R pfso def 0 = minR R pftso' (min_set_sortable A pff' pf' R pfso) a) /\ (A = Empty_set _ -> pop_ens_iter (Add A a) pff pfi R pfso def 0 = a). intros T A a h1 h2 R h3 def. destruct h3 as [h3 h4 h5]. simpl. pose proof (min_set_sortable_add A a h1 h2 R (sort_intro _ h3 h4 h5)) as h7. simpl in h7. destruct h7 as [h7 h8]. rewrite pop_ens_iter0. split. intro h6. simpl. specialize (h7 h6); auto. assumption. Qed. (*Wyckoff*) Lemma pop_ens_iter_couple1 : forall {T:Type} (a b:T) (R:T->T->Prop) (pfso:sortable R) (def:T), R a b -> let (pftso, _, _) := pfso in pop_ens_iter (Couple a b) (finite_couple a b) (inh_couple a b) R pfso def 1 = b. intros T a b R h1 def h4. destruct h1 as [h1 h2 h3]. unfold pop_ens_iter. rewrite min_set_sortable_couple; auto. simpl. assert (h5:a <> b). intro. subst. pose proof (h3 b). contradiction. pose proof (subtract_couple_l _ _ h5) as h6. pose proof (pop_ens_functional _ _ (Finite_downward_closed T (Couple a b) (finite_couple a b) (Subtract (Couple a b) a) (incl_subtract (Couple a b) a)) (Singleton_is_finite _ b) R (sort_intro _ h1 h2 h3) def h6) as h7. rewrite h7 at 1. rewrite pop_ens_sing. simpl. reflexivity. Qed. End PopEns. End Relations. Lemma inhabited_family_cond : forall {T:Type} (F:Family T), inhabited_family F -> forall A:sig_set F, exists y:T, In (proj1_sig A) y. intros T F h1 A. destruct h1 as [h1 h2]. destruct A as [A h3]. simpl. apply h2 in h3. destruct h3 as [y h3]. exists y; auto. Qed. Section InverseImage. (*Schepler*) Definition inverse_image {X Y:Type} (f:X->Y) (T:Ensemble Y) : Ensemble X := [ x:X | In T (f x) ]. Hint Unfold inverse_image : sets. (*Schepler*) Lemma inverse_image_increasing: forall {X Y:Type} (f:X->Y) (T1 T2:Ensemble Y), Included T1 T2 -> Included (inverse_image f T1) (inverse_image f T2). Proof. intros. red; intros. destruct H0. constructor. auto. Qed. (*Schepler*) Lemma inverse_image_empty: forall {X Y:Type} (f:X->Y), inverse_image f (Empty_set _) = Empty_set _. Proof. intros. apply Extensionality_Ensembles; split; red; intros. destruct H as [[]]. destruct H. Qed. (*Schepler*) Lemma inverse_image_full: forall {X Y:Type} (f:X->Y), inverse_image f (Full_set _) = Full_set _. Proof. intros. apply Extensionality_Ensembles; split; red; intros; constructor; constructor. Qed. (*Schepler*) Lemma inverse_image_intersection: forall {X Y:Type} (f:X->Y) (T1 T2:Ensemble Y), inverse_image f (Intersection T1 T2) = Intersection (inverse_image f T1) (inverse_image f T2). Proof. intros. apply Extensionality_Ensembles; split; red; intros. destruct H. inversion H. constructor; constructor; trivial. destruct H as [? [] []]. constructor; constructor; trivial. Qed. (*Schepler*) Lemma inverse_image_union: forall {X Y:Type} (f:X->Y) (T1 T2:Ensemble Y), inverse_image f (Union T1 T2) = Union (inverse_image f T1) (inverse_image f T2). Proof. intros. apply Extensionality_Ensembles; split; red; intros. destruct H. inversion H. left; constructor; trivial. right; constructor; trivial. constructor. destruct H as [? []|? []]. left; trivial. right; trivial. Qed. (*Schepler*) Lemma inverse_image_complement: forall {X Y:Type} (f:X->Y) (T:Ensemble Y), inverse_image f (Comp T) = Comp (inverse_image f T). Proof. intros. apply Extensionality_Ensembles; split; red; intros. red; intro. destruct H. destruct H0. contradiction H. constructor. intro. contradiction H. constructor; trivial. Qed. (*Schepler*) Lemma inverse_image_composition: forall {X Y Z:Type} (f:Y->Z) (g:X->Y) (U:Ensemble Z), inverse_image (fun x:X => f (g x)) U = inverse_image g (inverse_image f U). Proof. intros. apply Extensionality_Ensembles; split; red; intros. constructor; constructor. destruct H. assumption. destruct H; inversion H. constructor; trivial. Qed. Hint Resolve @inverse_image_increasing : sets. Hint Rewrite @inverse_image_empty @inverse_image_full @inverse_image_intersection @inverse_image_union @inverse_image_complement @inverse_image_composition : sets. End InverseImage. Section RandomFacts. (*Wyckoff*) Definition char_fun {T:Type} (A:Ensemble T) : T->bool := fun x:T => if (classic_dec (In A x)) then true else false. (*Wyckoff*) Lemma char_fun_int : forall {T:Type} (A B:Ensemble T) (x:T), char_fun (Intersection A B) x = char_fun A x && char_fun B x. intros T A B x. pose proof (in_intersection_iff A B x) as h1. unfold char_fun. lr_if_goal'. rewrite h1 in hl. destruct hl. lr_if_goal'; simpl; auto. lr_if_goal'; auto. contradiction. contradiction. rewrite h1 in hr. apply not_and_or in hr. destruct hr as [h2 | h2]. lr_if_goal'; simpl; auto. contradiction. lr_if_goal'; simpl; auto. lr_if_goal'; simpl; auto. Qed. (*Wyckoff*) Lemma char_fun_union : forall {T:Type} (A B:Ensemble T) (x:T), char_fun (Union A B) x = char_fun A x || char_fun B x. intros T A B x. unfold char_fun. lr_if_goal'. lr_if_goal'; simpl; auto. lr_if_goal'; auto. destruct hl; contradiction. lr_if_goal'; simpl. contradict hr; left; auto. rewrite in_union_iff in hr. lr_if_goal'; auto. Qed. (*Wyckoff*) Lemma char_fun_comp : forall {T:Type} (A:Ensemble T) (x:T), char_fun (Comp A) x = negb (char_fun A x). intros T A x. unfold char_fun. lr_if_goal'; simpl; auto. lr_if_goal'; simpl; auto. destruct hl; auto. lr_if_goal'; simpl; auto. Qed. (*Wyckoff*) Lemma injective_preserves_card_fun1 : forall {T U:Type} (f:T->U), injective f -> forall (S:Ensemble T), Finite S -> card_fun1 S = card_fun1 (Im S f). intros T U f h1 S h2. pose proof (card_fun1_compat S) as h3. simpl in h3. destruct h3 as [h3 h4]. specialize (h3 h2). rewrite inj_image_fun_prop_compat in h1. symmetry. apply (injective_preserves_cardinal _ _ _ _ _ h1 h3). apply card_fun1_compat. apply finite_image; auto. Qed. (*Wyckoff*) Lemma le_im_card : forall {T U:Type} (A:Ensemble T) (f:T->U), Finite A -> card_fun1 (Im A f) <= card_fun1 A. intros T U A f h1. pose proof (card_fun1_compat A) as h3. pose proof (card_fun1_compat (Im A f)) as h4. simpl in h3, h4. destruct h3 as [h3 h5]. destruct h4 as [h4 h6]. clear h5 h6. hfirst h4. apply finite_image; auto. specialize (h4 hf). specialize (h3 h1). eapply cardinal_decreases. apply h3. apply h4. Qed. (*Wyckoff*) Lemma excl_middle_sat : forall {T:Type} (P:T->Prop), Full_set T = Union [x:T | P x] [x:T | ~ (P x)]. intros T P. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct (classic (P x)) as [h2 | h3]. left. constructor. assumption. right. constructor. assumption. red. intros; constructor. Qed. (*Wyckoff*) Lemma excl_middle_sat' : forall {T:Type} (S:Ensemble T) (P:T->Prop), S = Union [x:T | In S x /\ P x] [x:T | In S x /\ ~ (P x)]. intros T S P. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct (classic (P x)) as [h2 | h3]. left. constructor. split; auto. right. constructor. split; auto. red. intros x h1. destruct h1 as [x h1l | x h1r]. destruct h1l as [h1l]. destruct h1l; assumption. destruct h1r as [h1r]. destruct h1r; assumption. Qed. (*Wyckoff*) Lemma surj_iff' : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} (f:sig_set A -> sig_set B), surjective f <-> im_proj1_sig (Im (full_sig A) f) = B. intros T U A B f. split. intro h1. apply surj_iff in h1. rewrite h1 at 1. rewrite im_proj1_sig_full_set_sig. reflexivity. intro h1. red. intro y. destruct y as [y h2]. pose proof h2 as h2'. rewrite <- h1 in h2'. destruct h2' as [y h3]. clear h1. subst. destruct h3 as [x h3]. subst. clear h3. exists x. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Definition extends_sig_sig {T U:Type} {A B:Ensemble T} {E:Ensemble (sig_set A)} (f:sig_set B->U) (g:sig_set E->U) := exists pf:Included (im_proj1_sig E) B, forall x:sig_set (im_proj1_sig E), g (exist _ _ (sig_set_im_proj1_sig_in x)) = f (exist _ _ (pf _ (proj2_sig x))). (*Wyckoff*) Definition extends_sig_sig' {T U:Type} {A B:Ensemble T} {E:Ensemble (sig_set A)} (f:sig_set B->U) (g:sig_set E->U) := forall (pf:Included (im_proj1_sig E) B) (x:sig_set (im_proj1_sig E)), g (exist _ _ (sig_set_im_proj1_sig_in x)) = f (exist _ _ (pf _ (proj2_sig x))). (*Wyckoff*) Lemma extends_sig_sig_impl_prime : forall {T U:Type} {A B:Ensemble T} {E:Ensemble (sig_set A)} (f:sig_set B->U) (g:sig_set E->U), extends_sig_sig f g -> extends_sig_sig' f g. intros T A B E f g h1 h2. red in h2. red. intro h3. destruct h2 as [h2 h4]. pose proof (proof_irrelevance _ h2 h3); subst. assumption. Qed. (*Wyckoff*) Lemma extends_sig_sig_extends_sig1_iff : forall {T U:Type} {A:Ensemble T} (E:Ensemble (sig_set A)) (f:sig_set A->U) (g:sig_set E->U), extends_sig_sig f g <-> extends_sig1 f g. intros T U A E f g. split. intro h1. red in h1. red. destruct h1 as [h1 h2]. intro x. destruct x as [x h3]. pose proof h3 as h3'. rewrite <- in_im_proj1_sig_iff in h3'. specialize (h2 (exist _ _ h3')). simpl. simpl in h2. let P := type of h2 in match P with g ?c' = _ => pose c' as c end. let P := type of h2 in match P with _ = f ?d' => pose d' as d end. match goal with |- g ?e' = _ => pose e' as e end. assert (h4:x = d). unfold d. apply proj1_sig_injective; auto. assert (h5:e = c). unfold e, c. apply proj1_sig_injective; simpl. apply proj1_sig_injective; auto. fold e. fold c in h2. rewrite h5. rewrite h2. fold d. rewrite h4. reflexivity. intro h1. red in h1. red. assert (h2:Included (im_proj1_sig E) A). red; intros x h3. destruct h3 as [x h3]; subst. apply proj2_sig. exists h2. intro x. destruct x as [x h3]. simpl. destruct h3 as [x h3]. subst. specialize (h1 (exist _ _ h3)). simpl in h1. match goal with |- g ?c' = _ => pose c' as c end. match goal with |- _ = f ?d' => pose d' as d end. fold c d. assert (h4:c = exist _ _ h3). unfold c. apply proj1_sig_injective; simpl. apply proj1_sig_injective; auto. rewrite h4. assert (h5:x = d). unfold d. apply proj1_sig_injective; auto. rewrite <- h5; auto. Qed. (*Wyckoff*) Lemma card_fun1_im_proj1_sig : forall {T:Type} (A:Ensemble T) (S:Ensemble (sig_set A)), Finite S -> card_fun1 S = card_fun1 (im_proj1_sig S). intros; apply injective_preserves_card_fun1; auto. red; intros; apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma card_fun1_im_proj1_sig' : forall {T:Type} (P:T->Prop), let U:= {x:T | P x} in forall S:Ensemble U, Finite S -> card_fun1 S = card_fun1 (im_proj1_sig' S). intros; apply injective_preserves_card_fun1; auto. red; intros; apply proj1_sig_injective; auto. Qed. (*See [SetUtilities2], for [card_fun1_im_proj2_sig].*) Lemma im_proj1_sig_full_sig : forall {T:Type} (S:Ensemble T), S = im_proj1_sig (full_sig S). intros T S. apply Extensionality_Ensembles; red; split; red; intros x h1. apply Im_intro with (exist _ _ h1). constructor. auto. destruct h1 as [x h1]; subst. clear h1. apply proj2_sig. Qed. (*Wyckoff*) Lemma card_union : forall {T:Type} (A B:Ensemble T) (m n p:nat), cardinal _ A m -> cardinal _ B n -> cardinal _ (Intersection A B) p -> cardinal _ (Union A B) (m+n-p). intros T A B m n p h1 h2 h3. pose proof (decompose_int_setminus (Union A B) B) as h4. assert (h5:Included B (Union A B)). auto with sets. rewrite Intersection_commutative in h4. rewrite inclusion_iff_intersection_eq in h5. rewrite h5 in h4. rewrite setminus_union in h4. pose proof (int_setminus B (Union A B)) as h6. rewrite setminus_union in h6. pose proof (card_decompose_int_setminus' _ _ _ _ h1 h3) as h7. pose proof (card_disj_union _ _ _ _ h6 h2 h7) as h8. rewrite h4. assert (h9:n + (m-p) = m + n - p). rewrite (plus_comm m n). apply plus_minus_assoc. assert (h9:Included (Intersection A B) A). auto with sets. apply (incl_card_le _ _ _ _ _ h3 h1 h9). rewrite <- h9. assumption. Qed. (*Wyckoff*) Lemma card_gt_O_inh : forall {T:Type} (A:Ensemble T) (n:nat), n > O -> cardinal _ A n -> Inhabited A. intros T A n; destruct n as [|n]; simpl. intros h1 h2. omega. intros ? h1. inversion h1. apply Inhabited_intro with x. right. constructor. Qed. (*Wyckoff*) Lemma sub_empty_sing_or_empty : forall {T:Type} (A:Ensemble T) (x:T), Subtract A x = Empty_set _ -> A = Singleton x \/ A = Empty_set _. intros T A x h1. destruct (classic (In A x)) as [h2 | h3]. left. apply Extensionality_Ensembles; red; split. red. intros a h3. apply NNPP. intro h4. assert (h5:In (Subtract A x) a). constructor; auto. rewrite h1 in h5. contradiction. red. intros a h3. destruct h3; subst. assumption. right. rewrite subtract_nin in h1; auto. Qed. (*Wyckoff*) Lemma atom_sing : forall {T:Type} (E A:Ensemble T), Finite E -> Included A E -> A <> Empty_set _ -> ((forall F:Ensemble T, Included F E -> Intersection F A = A \/ Intersection F A = Empty_set _) <-> (exists (a:T), In E a /\ A = Singleton a)). intros T E A h1. revert A. induction h1 as [|E h2 h3 x h4]. intros A h1 h2. assert (h3:A = Empty_set _). auto with sets. contradiction. intros A h1 h5. split. (* -> *) intro h6. pose proof (subtract_preserves_inclusion A (Add E x) x h1) as h7. rewrite sub_add_same_nin in h7; auto. destruct (classic (In A x)) as [h8 | h9]. exists x. split. right. constructor. pose proof (Inhabited_intro _ _ _ h8) as h9. assert (h10:Included (Singleton x) (Add E x)). red. intros x' h11. destruct h11. right. constructor. specialize (h6 (Singleton x) h10). destruct h6 as [h6l | h6r]. assert (h11:Intersection (Singleton x) A = Singleton x). apply Extensionality_Ensembles. red. split. auto with sets. red. intros x' h11. destruct h11. constructor; auto. constructor. rewrite h11 in h6l. rewrite h6l. reflexivity. assert (h11:In (Intersection (Singleton x) A) x). constructor; auto. constructor. rewrite h6r in h11. contradiction. pose proof (subtract_preserves_inclusion A (Add E x) x h1) as h10. rewrite subtract_nin in h10; auto. rewrite sub_add_same_nin in h10; auto. specialize (h3 _ h10 h5). assert (h11:(forall F : Ensemble T, Included F E -> Intersection F A = A \/ Intersection F A = Empty_set T)). intros F h12. assert (h13:Included F (Add E x)). red. intros; auto with sets. apply h6; auto. rewrite h3 in h11. destruct h11 as [a h11]. destruct h11 as [h11l h11r]. exists a. split; auto. left. assumption. intros h6. destruct h6 as [a h6]. destruct h6 as [h6l h6r]. intros F h7. subst. destruct (classic (In F a)) as [h8 | h9]. left. apply Extensionality_Ensembles. red. split. auto with sets. red. intros a' h9. destruct h9; subst. constructor; auto. constructor. right. apply Extensionality_Ensembles. red. split; auto with sets. red. intros a' h10. destruct h10 as [a' h10l h10r]. destruct h10r; subst. contradiction. Qed. (*Wyckoff*) Inductive incl_sig {T:Type} (A B:Ensemble T) (pfi:Included A B) : Ensemble {x:T|In B x} := incl_sig_intro : forall (x:T) (pfa:In A x), In (incl_sig A B pfi) (exist _ _ (pfi _ pfa)). (*Wyckoff*) Lemma incl_sig_iff : forall {T:Type} (A B:Ensemble T) (pf:Included A B) x, In (incl_sig A B pf) x <-> In A (proj1_sig x). intros T A B h0 x. split. intros h1. destruct h1. simpl. assumption. intro h1. pose proof (incl_sig_intro _ _ h0 _ h1) as h2. assert (h3:x = (exist (In B) (proj1_sig x) (h0 (proj1_sig x) h1))). rewrite unfold_sig at 1. apply proj1_sig_injective. simpl. reflexivity. rewrite <- h3 in h2. assumption. Qed. (*Wyckoff*) Definition subset_sig {T:Type} (A B:Ensemble T) (pf:Included A B) : Ensemble (sig_set B) := (Im (full_sig A) (fun x => exist _ _ (pf (proj1_sig x) (proj2_sig x)))). (*Wyckoff*) Lemma subset_sig_compat : forall {T:Type} (A B:Ensemble T) (pf:Included A B) x, In (subset_sig A B pf) x <-> In A (proj1_sig x). intros T A B h0 x. split. intro h1. destruct h1 as [x h2 y h4]. subst. simpl. inversion h2. subst. apply proj2_sig. intro h1. unfold subset_sig. apply Im_intro with (exist _ _ h1). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. Qed. (*Wyckoff*) Lemma incl_sig_eq_subset_sig : forall {T:Type} (A B:Ensemble T) (pf:Included A B), incl_sig A B pf = subset_sig A B pf. intros T A B pf. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. unfold subset_sig. apply Im_intro with (exist _ _ h1). constructor. apply proj1_sig_injective. simpl. reflexivity. red. intros x h1. destruct h1 as [x h1]. subst. constructor. Qed. (*Wyckoff*) Lemma full_sig_family_union_eq : forall {T:Type} (F:Family T), full_sig (FamilyUnion F) = FamilyUnion [S:Ensemble (sig_set (FamilyUnion F)) | In F (Im S (@proj1_sig _ _))]. intros T F. apply Extensionality_Ensembles. red. split. Focus 2. red. intros x h3. constructor. red. intros x h3. destruct x as [x h4]. destruct h4 as [S x h4 h5]. assert (h7:forall s:T, In S s -> In S s). auto. apply family_union_intro with (im_in _ _ (fun x':{s:T|In S s} => (exist _ _ (family_union_intro _ _ _ (proj1_sig x') h4 (proj2_sig x')))) h7). constructor. assert (h8:S = (Im (im_in S (In S) (fun x' : {s : T | In S s} => exist (In (FamilyUnion F)) (proj1_sig x') (family_union_intro T F S (proj1_sig x') h4 (proj2_sig x'))) h7) (proj1_sig (P:=fun x0 : T => In (FamilyUnion F) x0)))). apply Extensionality_Ensembles. red. split. red. intros s h9. apply Im_intro with (exist _ _ (family_union_intro _ _ _ _ h4 h9)). unfold im_in. simpl. apply Im_intro with (exist _ _ h9). constructor. apply proj1_sig_injective. simpl. reflexivity. simpl. reflexivity. red. intros s h9. destruct h9 as [s h9]. subst. destruct h9 as [s h9]. subst. simpl. destruct s. simpl. assumption. pose proof h4 as h4'. rewrite h8 in h4'. assumption. unfold im_in. simpl. apply Im_intro with (exist _ _ h5). constructor. apply proj1_sig_injective. simpl. reflexivity. Qed. (*Wyckoff*) Lemma full_sig_decompose_setminus_incl : forall {T:Type} (A B:Ensemble T) (pf:Included B A), full_sig A = Union (Im (full_sig (Setminus A B)) (fun x => exist _ _ ((setminus_inc A B) _ (proj2_sig x)))) (Im (full_sig B) (fun x => exist _ _ (pf _ (proj2_sig x)))). intros T A B h1. apply Extensionality_Ensembles. red. split. red. intros x h2. clear h2. destruct x as [x h2]. pose proof (decompose_setminus_inc _ _ h1) as h3. pose proof h2 as h2'. rewrite h3 in h2'. destruct h2' as [x h4 | x h5]. right. apply Im_intro with (exist _ _ h4). constructor. apply proj1_sig_injective. simpl. reflexivity. left. apply Im_intro with (exist _ _ h5). constructor. apply proj1_sig_injective. simpl. reflexivity. red. intros x h2. constructor. Qed. (*Wyckoff*) Lemma full_false_empty : forall (T:Type), Full_set {_ : T | False} = Empty_set _. intro T. apply Extensionality_Ensembles. red. split; auto with sets. red. intros x. destruct x. contradiction. Qed. (*Wyckoff*) Lemma full_empty_sig_empty : forall (T:Type), full_sig (Empty_set T) = Empty_set _. intros T. apply Extensionality_Ensembles. red. split. red. intros x h1. pose proof (proj2_sig x) as h2. simpl in h2. contradiction. auto with sets. Qed. (*Wyckoff*) Lemma finite_image_rev_inj : forall {T U:Type} (A:Ensemble T) (f:T->U), injective f -> Finite (Im A f) -> Finite A. intros T U A f h1 h2. apply NNPP. intro h3. contradict h2. intro h4. pose proof (Pigeonhole_bis _ _ _ _ h3 h4). contradiction. Qed. (*Wyckoff*) Definition inv_im {T U:Type} (B:Ensemble U) (f:T->U) := [x : T | In B (f x)]. (*Wyckoff*) Definition inv_im1 {T U:Type} (f:T->U) (y:U) := [x : T | f x = y]. (*Wyckoff*) Definition inv_im_seg {U:Type} {n:nat} (f:seg_fun n U) (y:U) := [i : nat | exists pf:i < n, f i pf = y]. (*Wyckoff*) Lemma inv_im1_compat : forall {T U:Type} (f:T->U) (x:T) (y:U), Inn (inv_im1 f y) x -> f x = y. intros; destruct H; auto. Qed. (*Wyckoff*) Lemma inv_im1_compat' : forall {T U:Type} (f:T->U) (y:U), inv_im1 f y = inv_im (Singleton y) f. intros T U f y. apply Extensionality_Ensembles; red; split; red; intros x h1; destruct h1; subst; econstructor. constructor. destruct H; auto. Qed. (*Wyckoff*) Lemma inv_im_seg_compat : forall {U:Type} {n:nat} (f:seg_fun n U) (y:U) (i:nat), Inn (inv_im_seg f y) i -> exists pf:i < n, f i pf = y. intros; destruct H; auto. Qed. (*Wyckoff*) Lemma lt_inv_im_seg : forall {U:Type} {n:nat} (f:seg_fun n U) (y:U) (i:nat), Inn (inv_im_seg f y) i -> i < n. intros. destruct H. destruct H; auto. Qed. (*Wyckoff*) (*The set of all pre-images of points in [U]*) Definition inv_im1s {T U} (f:T->U) : Family T := Im (Full_set U) (inv_im1 f). (*Wyckoff*) Lemma inv_im_empty : forall {T U:Type} (f:T->U), inv_im (Empty_set U) f = Empty_set T. intros T U f. apply Extensionality_Ensembles; red; split; auto with sets. red. intros x h1. destruct h1. contradiction. Qed. (*Wyckoff*) Lemma im_inv_im : forall {T U:Type} (f:T->U) (B:Ensemble U), Im (inv_im B f) f = Intersection B (Im (Full_set T) f). intros T U f B. apply Extensionality_Ensembles. unfold inv_im. red. split. (* <= *) red. intros b h1. inversion h1 as [a h2 b' h3]. subst. constructor. destruct h2. assumption. apply Im_intro with a. constructor. reflexivity. red. intros u h1. inversion h1 as [u' h2 h3]. subst. inversion h3 as [x h4 y h5 h6]. subst. apply Im_intro with x. constructor. assumption. reflexivity. Qed. (*Wyckoff*) Lemma im_inv_im1_in : forall {T U:Type} (f:T->U) (y:U), In (Im (Full_set _) f) y -> Im (inv_im1 f y) f = Singleton y. intros T U f y h1. destruct h1; subst. apply Extensionality_Ensembles; red; split; red; intros y h1; destruct h1; subst. destruct H0. rewrite H0. constructor. econstructor. econstructor. reflexivity. reflexivity. Qed. (*Wyckoff*) Lemma im_inv_im_seg_in : forall {U:Type} {n:nat} (f:seg_fun n U) (y:U), Inn (im_in_ens f) y -> im_in_ens (fun i (pf:Inn (inv_im_seg f y) i) => f i (lt_inv_im_seg _ _ _ pf)) = Singleton y. intros U n f y h1. destruct h1; subst. red in pf. apply Extensionality_Ensembles; red; split; red; intros y h1; destruct h1; subst. destruct pf0. destruct e as [h1 h2]. rewrite in_sing_iff. rewrite <- h2. f_equal. apply proof_irrelevance. econstructor. f_equal. apply proof_irrelevance. Grab Existential Variables. econstructor. exists pf; auto. Qed. (*Wyckoff*) Lemma in_im_inv_im1 : forall {T U:Type} (f:T->U) (x:T), Inn (inv_im1 f (f x)) x. intros; econstructor; auto. Qed. (*Wyckoff*) Lemma in_inv_im_seg : forall {U:Type} {n:nat} (f:seg_fun n U) (i:nat) (pf:i < n), Inn (inv_im_seg f (f i pf)) i. intros; econstructor; auto. exists pf; auto. Qed. (*Wyckoff*) Lemma inv_im_int : forall {T U:Type} (A B:Ensemble U) (f:T->U), inv_im (Intersection A B) f = Intersection (inv_im A f) (inv_im B f). intros T U A B f. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [h1]. inversion h1 as [? h2 h3]. subst. constructor; constructor; auto. red. intros x h1. destruct h1 as [x h1 h2]. constructor. destruct h1 as [h1]. destruct h2 as [h2]. constructor; auto. Qed. (*Wyckoff*) Lemma inv_im_union : forall {T U:Type} (A B:Ensemble U) (f:T->U), inv_im (Union A B) f = Union (inv_im A f) (inv_im B f). intros T U A B f. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [h1]. inversion h1 as [? h2 |? h3]. subst. left. constructor; auto. subst. right. constructor; auto. red. intros x h1. destruct h1 as [x h1 | x h2]. destruct h1 as [h1]. constructor. left. auto. destruct h2 as [h2]. constructor. right. assumption. Qed. (*Wyckoff*) Lemma inv_im_comp : forall {T U:Type} (A:Ensemble U) (f:T->U), inv_im (Comp A) f = Comp (inv_im A f). intros T U A f. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [h1]. intro h2. destruct h2 as [h2]. auto with sets. red. intros x h1. constructor. intro h2. unfold inv_im in h1. apply complement_inv in h1. apply sat_nin in h1. contradiction. Qed. (*Wyckoff*) Lemma inh_surj : forall {T U:Type} (f:T->U) (y:U), surjective f -> Inhabited (inv_im1 f y). intros T U f y h1. red in h1. specialize (h1 y). destruct h1; subst. econstructor. econstructor. reflexivity. Qed. (*Wyckoff*) Lemma inh_surj_seg : forall {U:Type} {n:nat} (f:seg_fun n U) (y:U), surj_seg f -> Inhabited (inv_im_seg f y). intros U n f y h1. specialize (h1 y). destruct h1 as [i h2]. destruct h2; subst. econstructor. econstructor. exists x; auto. Qed. (*Wyckoff*) Lemma inh_inv_im_seg : forall {U:Type} {n:nat} (f:seg_fun n U) (y:sig_set (im_in_ens f)), Inhabited (inv_im_seg f (proj1_sig y)). intros U n f y. destruct y as [y hy]; subst. destruct hy as [hy]; subst. simpl. econstructor. econstructor. exists pf; auto. Qed. (*Wyckoff*) Lemma inh_inv_im_seg' : forall {U:Type} {n:nat} (f:nat->U) (y:sig_set (Im (seg n) f)), Inhabited (inv_im_seg (fun_to_seg f n) (proj1_sig y)). intros U n f y. destruct y as [y hy]; subst. destruct hy as [y hy]; subst. simpl. econstructor. econstructor. destruct hy as [hy]. exists hy; auto. Qed. (*Wyckoff*) Lemma inh_inv_im_sig_seg : forall {U:Type} (f:nat->U) (n:nat) (y:sig_set (Im (seg n) f)), Inhabited (inv_im_seg (fun_to_sig_seg f n) y). intros U f n y. destruct y as [y hy]; subst. destruct hy as [i hy]; subst. destruct hy as [hy]. econstructor. econstructor. exists hy. unfold fun_to_sig_seg. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma finite_inv_im_seg : forall {U:Type} {n:nat} (f:seg_fun n U) (y:U), Finite (inv_im_seg f y). intros U n f y. unfold inv_im_seg. eapply Finite_downward_closed. apply finite_seg. red. intros i h1. destruct h1 as [h1]. destruct h1; subst. constructor. apply x. Qed. (*Wyckoff*) Lemma inv_im_surj_inj : forall {T U:Type} (f:T->U), surjective f -> forall (A B:Ensemble U), inv_im A f = inv_im B f-> A = B. intros T U f h1 B C h2. red in h1. apply Extensionality_Ensembles. red. split. red. intros y h3. specialize (h1 y). destruct h1 as [x h1r]. subst. assert (h4:In (inv_im B f) x). constructor; auto. rewrite h2 in h4. destruct h4. assumption. red. intros y h3. specialize (h1 y). destruct h1 as [x h1r]. subst. assert (h4:In (inv_im C f) x). constructor; auto. rewrite <- h2 in h4. destruct h4. assumption. Qed. (*Wyckoff*) Lemma inv_im1_surj_inj : forall {T U:Type} (f:T->U) (y z:U), surjective f -> inv_im1 f y = inv_im1 f z -> y = z. intros T U f y z h1 h2. pose proof (inh_surj f y h1) as h3. destruct h3 as [x h3]. destruct h3 as [h3]; subst. pose proof (in_im_inv_im1 f x) as h3. rewrite h2 in h3. destruct h3; auto. Qed. (*Wyckoff*) Lemma inv_im_seg_surj_inj : forall {U:Type} (n:nat) (f:seg_fun n U) (y z:U), surj_seg f -> inv_im_seg f y = inv_im_seg f z -> y = z. intros U n f y z h1 h2. pose proof (inh_surj_seg _ y h1) as h3. destruct h3 as [m h4]. rewrite h2 in h4. destruct h4 as [h4]. destruct h4 as [h4 h5]; subst. pose proof (in_inv_im_seg f m h4) as h5. rewrite <- h2 in h5. destruct h5. destruct H; subst. f_equal. apply proof_irrelevance. Qed. (*Wyckoff*) Lemma inv_im_inj_sing : forall {T U:Type} (f:T->U), injective f -> forall y:U, {exists! x:T, f x = y} + {inv_im (Singleton y) f = Empty_set _}. intros T U f h1 y. destruct (classic_dec (exists! x:T, f x = y)) as [h2 | h3]. left; auto. right. apply Extensionality_Ensembles; red; split; auto with sets. red. intros x h4. destruct h4 as [h4]. inversion h4; clear h4. subst. contradict h3. exists x. red; split; auto. Qed. (*Wyckoff*) Lemma inv_im1_inj_sing : forall {T U:Type} (f:T->U), injective f -> forall y:U, {exists! x:T, f x = y} + {inv_im1 f y = Empty_set _}. intros; rewrite inv_im1_compat'; apply inv_im_inj_sing; auto. Qed. (*Wyckoff*) Lemma inv_im_inj_sing' : forall {T U:Type} (f:T->U), injective f -> forall y:U, card_fun1 (inv_im (Singleton y) f) <= 1. intros T U f h1 y. pose proof (inv_im_inj_sing f h1 y) as h2. destruct h2 as [h2a | h2b]. destruct h2a as [x h3]. red in h3. destruct h3 as [? h3]. subst. assert (h4:inv_im (Singleton (f x)) f = Singleton x). apply Extensionality_Ensembles. red. split. red. intros x' h4. destruct h4 as [h4]. inversion h4 as [h5]. clear h4. symmetry in h5. apply h3 in h5. subst. constructor. red. intros x' h4. destruct h4. constructor. constructor. rewrite h4. rewrite card_fun1_sing. auto with arith. rewrite h2b. rewrite card_fun1_empty. auto with sets. Qed. (*Wyckoff*) Lemma inv_im1_inj_sing' : forall {T U:Type} (f:T->U), injective f -> forall y:U, card_fun1 (inv_im1 f y) <= 1. intros; rewrite inv_im1_compat'; apply inv_im_inj_sing'; auto. Qed. (*Wyckoff*) Lemma left_inverse_ex : forall {T U:Type} (f:T->U), injective f -> forall (A:Ensemble T), exists! f':sig_set (Im A f)->T, forall (x:T) (pf:In A x), f' (exist _ _ (Im_intro _ _ A f _ pf _ (eq_refl _))) = x. intros T U f h1 A. assert (h2:forall y:sig_set (Im A f), exists! x:T, f x = proj1_sig y). intro y. pose proof (inv_im_inj_sing _ h1 (proj1_sig y)) as h2. destruct h2 as [h2l | h2r]; auto. destruct y as [y h3]. destruct h3 as [x h3]. subst. simpl in h2r. simpl. contradict h2r. apply Inhabited_not_empty. apply Inhabited_intro with x. constructor. constructor. exists (fun x => (proj1_sig (constructive_definite_description _ (h2 x)))). red. split. intros x h3. destruct constructive_definite_description as [x' h5]. simpl in h5. simpl. apply h1. assumption. intros f' h3. apply functional_extensionality. intro x. destruct constructive_definite_description as [x' h4]. simpl. destruct x as [x h5]. simpl in h4. destruct h5 as [x h5]. subst. apply h1 in h4. subst. symmetry. apply h3. Qed. (*Wyckoff*) Definition left_inverse {T U:Type} (f:T->U) (pf:injective f) (A:Ensemble T) : sig_set (Im A f)-> T := proj1_sig (constructive_definite_description _ (left_inverse_ex f pf A)). (*Wyckoff*) Lemma left_inverse_compat : forall {T U:Type} (f:T->U) (pf:injective f) (A:Ensemble T), let f' := left_inverse f pf A in forall (x:T) (pf:In A x), f' (exist _ _ (Im_intro _ _ A f _ pf _ (eq_refl _))) = x. intros T U f h1 S f'. unfold f'. unfold left_inverse. destruct constructive_definite_description as [f'' h2]. simpl. apply h2. Qed. (*Wyckoff*) Definition left_inverse_full {T U:Type} (f:T->U) (pf:injective f) : sig_set (Im (Full_set T) f) -> T := left_inverse f pf (Full_set T). (*Wyckoff*) Lemma left_inverse_full_compat : forall {T U:Type} (f:T->U) (pf:injective f) x, left_inverse_full f pf (exist (fun c => Inn (Im (Full_set _) f) c) _ (in_im x f)) = x. intros T U f h1 x. rewrite <- (left_inverse_compat f h1 (Full_set _) x (Full_intro _ x)). unfold left_inverse_full; f_equal; apply proj1_sig_injective; simpl. f_equal. rewrite left_inverse_compat; auto. Qed. (*Wyckoff*) Lemma left_inverse_inj : forall {T U:Type} (f:T->U) (pf:injective f) (A:Ensemble T), injective (left_inverse f pf A). intros T U f h1 A. red. intros y y' h0. destruct y as [y h2], y' as [y' h3]. pose proof (left_inverse_compat f h1 A) as h4. simpl in h4. destruct h2 as [x h2], h3 as [x' h3]. subst. do 2 rewrite h4 in h0. subst. apply proj1_sig_injective. simpl. reflexivity. Qed. (*Wyckoff*) Lemma incl_im_full_sig_im_left_inverse : forall {T U:Type} (f:T->U) (pf:injective f) (A:Ensemble T), Included (Im (full_sig (Im A f)) (left_inverse f pf A)) A. intros T U f h1 A. red. intros x h2. destruct h2 as [x h2]. subst. clear h2. destruct x as [x h2]. destruct h2 as [x h2]. subst. rewrite left_inverse_compat. assumption. Qed. (*Wyckoff*) Lemma inv_im_inj_im_compat : forall {T U:Type} (f:T->U) (pf:injective f) (B:Ensemble U), inv_im B f = Im [u:sig_set (Im (Full_set T) f) | In B (proj1_sig u)] (left_inverse_full f pf). intros T U f h1 B. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [h2]. apply Im_intro with (exist _ _ (Im_intro _ _ (Full_set T) f _ (Full_intro _ x) _ (eq_refl _))). constructor; simpl; auto. unfold left_inverse_full. pose proof (left_inverse_compat f h1 (Full_set T)) as h3. simpl in h3. rewrite h3. reflexivity. red. intros x h2. destruct h2 as [y h2]. subst. destruct y as [y h3]. destruct h3 as [x h3]. subst. inversion h2 as [h4]. clear h2. simpl in h4. constructor. pose proof (left_inverse_compat f h1 (Full_set T)) as h3'. simpl in h3'. unfold left_inverse_full. rewrite h3'. assumption. Qed. (*Wyckoff*) Lemma finite_im_full_set_f_in_b : forall {T U:Type} (B:Ensemble U), Finite B -> forall (f:T->U) (pf:injective f), Finite (Im [u : sig_set (Im (Full_set T) f) | In B (proj1_sig u)] (left_inverse f pf (Full_set T))). intros T U B h1. induction h1 as [|B h2 h3 b h4]. intros. gterm0. rename c into E. assert (h1:E = Empty_set _). apply Extensionality_Ensembles; red; split; auto with sets. red. intros x h2. destruct h2 as [x h2]. destruct h2 as [h2]. contradiction. subst. rewrite h1. constructor. intros f h1. gterm0. redterm1 c. clear c. rename c0 into E. assert (h7: E = Union [u : sig_set (Im (Full_set T) f) | In B (proj1_sig u)] [u : sig_set (Im (Full_set T) f) | proj1_sig u = b]). apply Extensionality_Ensembles. red. split. red. intros x h7. destruct h7 as [h7]. inversion h7 as [b' h8 | b' h9]; clear h7. subst. left. constructor. assumption. inversion h9. subst. right. constructor. reflexivity. red. intros x h7. destruct h7 as [x h7l | x h7r]. destruct h7l as [h7l]. constructor. left. assumption. destruct h7r as [h7r]. unfold E; clear E; subst. constructor. right. constructor. unfold E in h7. clear E. rewrite h7. rewrite im_union. apply Union_preserves_Finite. apply h3. pose proof (inv_im_inj_im_compat _ h1 (Singleton b)) as h8. assert (h9:Im [u : sig_set (Im (Full_set T) f) | In (Singleton b) (proj1_sig u)] (left_inverse_full f h1) = Im [u : sig_set (Im (Full_set T) f) | proj1_sig u = b] (left_inverse f h1 (Full_set T))). f_equal. apply Extensionality_Ensembles. red. split. red. intros x h9. destruct h9 as [h9]. inversion h9. subst. constructor. reflexivity. red. intros x h9. destruct h9 as [h9]. subst. constructor. constructor. rewrite <- h9. rewrite <- h8. pose proof (inv_im_inj_sing' _ h1 b) as h10. apply le_lt_eq_dec in h10. destruct h10 as [h10l | h10r]. assert (h11:card_fun1 (inv_im (Singleton b) f) = 0). omega. apply card_fun1_O in h11. destruct h11 as [h11l | h11r]. rewrite h11l. constructor. pose proof (approximant_can_be_any_size _ (inv_im (Singleton b) f)(inv_im (Singleton b) f) h11r 2) as h12. destruct h12 as [X h13]. destruct h13 as [h13l h13r]. destruct h13r as [h14 h15]. apply cardinal_2 in h13l. destruct h13l as [x [y [h13a h13b]]]. subst. pose proof (h15 _ (Couple_l _ x y)) as h16. pose proof (h15 _ (Couple_r _ x y)) as h17. destruct h16 as [h16], h17 as [h17]. inversion h16; subst; clear h16. inversion h17; subst; clear h17. apply h1 in H0. contradiction. apply card_fun1_1 in h10r. destruct h10r as [x' h10r]. rewrite h10r. apply Singleton_is_finite. Qed. Inductive inv_im_eq {T U V W} (f:T->U) (g:V->W) : Prop := inv_im_eq_intro : forall (h:fun_in_ens' (inv_im1s f) (inv_im1s g)), invertible_in_ens h -> (forall S (pfs:Inn (inv_im1s f) S), exists (k:fun_in_ens' S (proj1_sig (h S pfs))), invertible_in_ens k) -> inv_im_eq f g. (*Wyckoff*) Definition full_wrap_fun1 {T U:Type} (f:T -> U) : sig_set (Full_set T) -> U := (fun x => f (proj1_sig x)). (*Wyckoff*) Definition full_wrap_fun2 {T U:Type} (f:T -> U) : T -> sig_set (Full_set U) := (fun x => (exist _ _ (Full_intro _ (f x)))). (*Wyckoff*) Lemma inj_full_wrap_fun1_iff : forall {T U:Type} (f:T->U), injective f <-> injective (full_wrap_fun1 f). intros T U f. split. intro h1. red in h1. red. intros x y h4. destruct x as [x h2], y as [y h3]. unfold full_wrap_fun1 in h4. simpl in h4. specialize (h1 x y h4). subst. apply proj1_sig_injective; auto. intro h1. red in h1. red. intros x y h4. specialize (h1 (exist _ _ (Full_intro _ x)) (exist _ _ (Full_intro _ y))). unfold full_wrap_fun1 in h1. simpl in h1. specialize (h1 h4). apply exist_injective in h1. subst. reflexivity. Qed. (*Wyckoff*) Lemma surj_full_wrap_fun1_iff : forall {T U:Type} (f:T->U), surjective f <-> surjective (full_wrap_fun1 f). intros T U f. split. intro h1. red in h1. red. intro y. specialize (h1 y). destruct h1 as [x h2]. subst. exists (exist _ _ (Full_intro _ x)). unfold full_wrap_fun1. simpl. reflexivity. intro h1. red in h1. red. intro y. specialize (h1 y). destruct h1 as [x h2]. destruct x as [x h3]. unfold full_wrap_fun1 in h2. simpl in h2. subst. exists x; auto. Qed. (*Wyckoff*) Lemma bij_full_wrap_fun1_iff : forall {T U:Type} (f:T->U), bijective f <-> bijective (full_wrap_fun1 f). intros T U f. unfold bijective. rewrite inj_full_wrap_fun1_iff, surj_full_wrap_fun1_iff. tauto. Qed. (*Wyckoff*) Lemma inj_full_wrap_fun2_iff : forall {T U:Type} (f:T->U), injective f <-> injective (full_wrap_fun2 f). intros T U f. split. intro h1. red in h1. red. intros x y h4. unfold full_wrap_fun2 in h4. apply exist_injective in h4. specialize (h1 x y h4). assumption. intro h1. red in h1. red. intros x y h4. specialize (h1 x y). unfold full_wrap_fun2 in h1. simpl in h1. pose proof (f_equal (fun x => exist _ _ (Full_intro _ x)) h4) as h5. simpl in h5. specialize (h1 h5). assumption. Qed. (*Wyckoff*) Lemma surj_full_wrap_fun2_iff : forall {T U:Type} (f:T->U), surjective f <-> surjective (full_wrap_fun2 f). intros T U. split. intro h1. red in h1. red. intro y. destruct y as [y h2]. specialize (h1 y). destruct h1 as [x h3]. subst. exists x. unfold full_wrap_fun2. apply proj1_sig_injective; auto. intro h1. red in h1. red. intro y. specialize (h1 (exist _ _ (Full_intro _ y))). destruct h1 as [x h2]. unfold full_wrap_fun2 in h2. apply exist_injective in h2. subst. exists x; auto. Qed. (*Wyckoff*) Lemma bij_full_wrap_fun2_iff : forall {T U:Type} (f:T->U), bijective f <-> bijective (full_wrap_fun2 f). intros T U f. unfold bijective. rewrite inj_full_wrap_fun2_iff, surj_full_wrap_fun2_iff. tauto. Qed. (*Wyckoff*) Definition sig_fun_app {T U:Type} {A:Ensemble T} (f:sig_set A->U) (def:U) : T->U. intro x. destruct (classic_dec (In A x)) as [h1 | h2]. refine (f (exist _ _ h1)). refine def. Defined. (*Wyckoff*) (*Uses LEM*) Lemma im_sig_eq : forall {T U:Type} {P:T->Prop} (S:Ensemble {x | P x}) (f:{x | P x}-> U) (def:U), Im S f = Im [x:T | In (im_proj1_sig S) x /\ P x] (sig_fun_app f def). intros T U P S f def. unfold sig_fun_app. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]. subst. apply Im_intro with (proj1_sig x). constructor. split. rewrite in_im_proj1_sig_iff; auto. apply proj2_sig; auto. destruct classic_dec as [h2 | h2]. f_equal. apply proj1_sig_injective; auto. contradict h2. apply proj2_sig. destruct h1 as [x h1]. subst. destruct h1 as [h1]. destruct h1 as [h1 h2]. destruct h1 as [x h1]. subst. apply Im_intro with x; auto. lr_match_goal'. f_equal. apply proj1_sig_injective; auto. contradiction. Qed. (*Wyckoff*) Lemma im_full_eq : forall {T U:Type} {P:T->Prop} (f:{x | P x}-> U) (def:U), Im (Full_set {x | P x}) f = Im [x:T | P x] (sig_fun_app f def). intros T U P f def. unfold sig_fun_app. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]. subst. clear h1. apply Im_intro with (proj1_sig x). constructor. apply proj2_sig. lr_match_goal'. f_equal. apply proj1_sig_injective; auto. contradict hr. apply proj2_sig. destruct h1 as [x h1]. subst. destruct h1 as [h1]. apply Im_intro with (exist _ _ h1). constructor. lr_match_goal'. f_equal. apply proj1_sig_injective; auto. contradiction. Qed. (*Wyckoff*) Definition sig_fun_injectors {T U:Type} {A:Ensemble T} (f:sig_set A->U) : Family T := [B:Ensemble T | exists pf:Included B A, injective (restriction_sig f B pf) /\ Inn (same_im_subsets_sig f) B]. Lemma incl_sig_fun_injectors : forall {T U:Type} {A B:Ensemble T} (f:sig_set A->U), In (sig_fun_injectors f) B -> Included B A. intros T U A B f h1. destruct h1 as [h1]. destruct h1; auto. Qed. (*Wyckoff*) Definition min_sig_fun_injectors {T U:Type} {A:Ensemble T} (f:sig_set A->U) : Ensemble T := FamilyIntersection (sig_fun_injectors f). (*Wyckoff*) Lemma incl_min_sig_fun_injectors : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), Inhabited (sig_fun_injectors f) -> Included (min_sig_fun_injectors f) A. intros T U A f h0. red. intros x h1. destruct h1 as [x h1]. destruct h0 as [B h0]. specialize (h1 _ h0). destruct h0 as [h0]. destruct h0 as [h0 [h2 h3]]. apply h0; auto. Qed. (*Wyckoff*) Lemma incl_min_sig_fun_injectors' : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U) (B:Ensemble T), In (sig_fun_injectors f) B -> Included (min_sig_fun_injectors f) B. intros T U A f p h1. red. intros x h2. destruct h2 as [x h2]. apply h2; auto. Qed. (*Wyckoff*) Lemma in_sig_fun_injectors_bij : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U) (B:Ensemble T) (pf:In (sig_fun_injectors f) B), let f' := (fun x:sig_set A => (exist _ _ (Im_intro _ _ (Full_set _) f _ (Full_intro _ x) _ (eq_refl _)))) in bijective (restriction_sig f' _ (incl_sig_fun_injectors f pf )). intros T U A f B h1 f'. destruct h1 as [h1]. destruct h1 as [h1 h2]. destruct h2 as [h2 h3]. red. split. red; intros x y h4. unfold f', restriction_sig in h4. simpl in h4. apply exist_injective in h4. red in h2. unfold restriction_sig in h2. pose proof (h2 x y) as h2'. terml h4. termr h4. redterm0 x0. redterm0 c. redterm0 y0. redterm0 c1. clear c c1 x0 y0. pose proof (proof_irrelevance _ (h1 _ (proj2_sig x)) c0) as h5. pose proof (proof_irrelevance _ (h1 _ (proj2_sig y)) c2) as h6. unfold c0 in h5; unfold c2 in h6. rewrite h5, h6 in h2'. specialize (h2' h4). assumption. rewrite surj_iff. destruct h3 as [h3]. destruct h3 as [h3 h4]. apply im_proj1_sig_injective. rewrite im_proj1_sig_full_set_sig. unfold full_sig in h4. rewrite h4 at 4. unfold im_proj1_sig. rewrite im_im. apply im_ext_in. intros x ?. unfold restriction_sig, f'. simpl. f_equal. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma inj_restriction_sig : forall {T U:Type} (A B C:Ensemble T) (f:sig_set A->U) (pfbc:Included B C) (pfca:Included C A), injective (restriction_sig f C pfca) -> injective (restriction_sig f B (inclusion_transitive _ _ _ _ pfbc pfca)). intros T U A B C f h1 h2 h3. red in h3. red. intros x y h4. destruct x as [x h5], y as [y h6]. pose proof (h1 _ h5) as h5'. pose proof (h1 _ h6) as h6'. specialize (h3 (exist _ _ h5') (exist _ _ h6')). unfold restriction_sig in h4. simpl in h4. unfold restriction_sig in h3. simpl in h3. pose proof (proof_irrelevance _ (h2 x h5') (inclusion_transitive T B C A h1 h2 x h5)) as h7. pose proof (proof_irrelevance _ (h2 y h6') (inclusion_transitive T B C A h1 h2 y h6)) as h8. rewrite <- h7, <- h8 in h4. specialize (h3 h4). apply proj1_sig_injective. simpl. apply exist_injective in h3. assumption. Qed. (*Wyckoff*) Lemma im_full_sig_sub_decompose_inj : forall {T U:Type} (pfdt:type_eq_dec T) (E:Ensemble T) (x:T) (f:sig_set E ->U) (pfe:Inn E x), let f' := restriction_sig f _ (incl_subtract _ _) in Im (full_sig E) f = Add (Im (full_sig (Subtract E x)) f') (f (exist _ _ pfe)). intros T U h0 E x f h1. simpl. apply Extensionality_Ensembles; red; split; red; intros c h3. destruct h3 as [c h3]; subst. clear h3. destruct c as [c h3]. destruct (h0 x c) as [h4 | h4]; subst. right. pose proof (proof_irrelevance _ h3 h1); subst. constructor. left. rewrite <- in_sing_iff in h4. apply Im_intro with (exist (fun c => In (Subtract E x) c) c (conj h3 h4)). constructor. unfold restriction_sig. simpl. f_equal. f_equal. apply proof_irrelevance. destruct h3 as [c h3 | c h3]; subst. destruct h3 as [c h3]; subst. clear h3. destruct c as [c h3]. destruct h3 as [h3 h4]. apply Im_intro with (exist _ _ h3). constructor. unfold restriction_sig. f_equal. f_equal. apply proof_irrelevance. destruct h3. apply Im_intro with (exist _ _ h1). constructor. reflexivity. Qed. Section EquivalenceClass. (*Wyckoff*) Definition rel_classes {T:Type} (A:Ensemble T) (R:T->T->Prop) := Im A (fun x:T => [y:T | In A y /\ R x y]). (*Wyckoff*) Lemma rel_classes_empty : forall {T:Type} (R:T->T->Prop), rel_classes (Empty_set _) R = Empty_set (Ensemble T). intros T R. unfold rel_classes. rewrite image_empty. reflexivity. Qed. (*Wyckoff*) Definition rep_to_eq_class {T:Type} (A:Ensemble T) (R:T->T->Prop) (a:T) (pf:In A a) := [y:T | R a y]. (*Wyckoff*) Lemma refl_rel_classes_inhabited : forall {T:Type} (A:Ensemble T), Inhabited A -> forall (R:T->T->Prop), Reflexive R -> inhabited_family (rel_classes A R). intros T A h1 R h2. constructor. destruct h1 as [a h1]. specialize (h2 a). unfold rel_classes. apply Inhabited_intro with [y : T | In A y /\ R a y]. apply Im_intro with a; auto. intros A' h3. destruct h3 as [a' h3]. subst. apply Inhabited_intro with a'. constructor. specialize (h2 a'). split; auto. Qed. (*wyckoff*) Lemma equiv_classes_inhabited : forall {T:Type} (A:Ensemble T), Inhabited A -> forall (R:T->T->Prop), Equivalence R -> inhabited_family (rel_classes A R). intros T A h1 R h2. destruct h2 as [h2]. apply refl_rel_classes_inhabited; auto. Qed. (*Wyckoff*) Definition partition_fam {T:Type} (F:Family T) (A:Ensemble T) : Prop := pairwise_disjoint F /\ FamilyUnion F = A. (*Wyckoff*) Definition partition_fam_impl_pairwise_disjoint : forall {T:Type} {F:Family T}, pairwise_disjoint F -> partition_fam F (FamilyUnion F). unfold partition_fam; intros; auto. Qed. (*Wyckoff*) Lemma partition_fam_empty : forall {T:Type} (A:Ensemble T), partition_fam (Empty_set (Ensemble T)) A -> A = Empty_set _. intros T A h1. red in h1. destruct h1 as [h1 h2]. rewrite empty_family_union in h2; auto. Qed. (*Wyckoff*) Lemma partition_fam_add : forall {T:Type} (F:Family T) (A B:Ensemble T), ~In F A -> partition_fam (Add F A) B -> partition_fam F (Setminus B A). intros T F A B h0 h1. red in h1. red. destruct h1 as [h1 h2]. pose proof (pairwise_disjoint_add_family_union F A h0 h1) as h3. apply pairwise_disjoint_add_rev in h1. split; auto. subst. rewrite setminus_family_union. rewrite disjoint_setminus_compat; auto. Qed. (*Wyckoff*) Lemma partition_fam_eq : forall {T:Type} (F:Family T) (A B C:Ensemble T) (a:T), partition_fam F A -> In F B -> In F C -> In B a -> In C a -> B = C. intros T F A B C a h1 h2 h3 h4 h5. red in h1. destruct h1 as [h1 h6]. subst. red in h1. apply NNPP. intro h6. specialize (h1 _ _ h6 h2 h3). red in h1. contradict h1. intro h7. pose proof (Intersection_intro _ _ _ _ h4 h5) as h8. rewrite h7 in h8. contradiction. Qed. (*Wyckoff*) Lemma partition_fam_dec : forall {T:Type} (F:Family T) (A B U:Ensemble T), partition_fam F U -> In F A -> In F B -> {A = B} + {disjoint A B}. intros T F A B U h1 h2 h3. red in h1. destruct h1 as [h1 h4]. subst. red in h1. destruct (eq_dec A B) as [|h4]. left; auto. right. apply h1; auto. Qed. (*Wyckoff*) Lemma rep_to_partition_class_ex : forall {T:Type} (F:Family T) (pf:Finite F) (A:Ensemble T), partition_fam F A -> forall a:T, Inn A a -> exists! S:Ensemble T, In F S /\ In S a. intros T F h1. induction h1 as [|F h2 h3 A h4]. intros A h1 a h2. apply partition_fam_empty in h1. subst. contradiction. intros B h5 a h6. pose proof (partition_fam_add _ _ _ h4 h5) as h5'. red in h5. destruct h5 as [h5 h9]. subst. rewrite family_union_add in h6. destruct h6 as [a h6 | a h6]. exists A. red. split. split; auto. right. constructor. intros S h9. destruct h9 as [h9 h10]. red in h5. destruct h9 as [S h9 | S h9]. specialize (h5 S A). assert (h11:S <> A). intro h11. subst. contradiction. pose proof (Add_intro1 _ _ A _ h9) as h12. pose proof (Add_intro2 _ F A) as h13. specialize (h5 h11 h12 h13). red in h5. pose proof (Intersection_intro _ _ _ _ h10 h6) as h14. rewrite h5 in h14. contradiction. destruct h9; auto. rewrite setminus_family_union in h5'. red in h5'. destruct h5' as [h7 h8]. destruct h6 as [S a h6 h0]. exists S. red. split. split. left; auto. assumption. intros S' h9. destruct h9 as [h9 h10]. destruct h9 as [S' h9 | S' h9]. apply NNPP. intro h11. red in h7. specialize (h7 _ _ h11 h6 h9). red in h7. pose proof (Intersection_intro _ _ _ _ h0 h10) as h12. rewrite h7 in h12. contradiction. destruct h9. pose proof (family_union_intro _ F _ _ h6 h0) as h11. rewrite h8 in h11. destruct h11; contradiction. Qed. (*Wyckoff*) Definition rep_to_partition_class {T:Type} (F:Family T) (pff:Finite F) (A:Ensemble T) (pfp:partition_fam F A) (a:T) (pfa:In A a) : Ensemble T := proj1_sig (constructive_definite_description _ (rep_to_partition_class_ex F pff A pfp a pfa)). (*Wyckoff*) Lemma rep_to_partition_class_compat : forall {T:Type} (F:Family T) (pff:Finite F) (A:Ensemble T) (pfp:partition_fam F A) (a:T) (pfa:In A a), let S := rep_to_partition_class F pff A pfp a pfa in In F S /\ In S a. unfold rep_to_partition_class; intros; destruct constructive_definite_description; auto. Qed. (*Wyckoff*) Lemma finite_rep_to_partition_class : forall {T:Type} (F:Family T) (pff:Finite F) (A:Ensemble T) (pfp:partition_fam F A) (a:T) (pfa:In A a), Finite A -> Finite (rep_to_partition_class F pff A pfp a pfa). intros T F h1 A h2 a h4 h5. unfold rep_to_partition_class. destruct constructive_definite_description as [B h6]. simpl. red in h2. destruct h2 as [h2 h7]. subst. eapply finite_family_union_finite_mem. apply h5. destruct h6; auto. Qed. (*Wyckoff*) Lemma in_rep_to_partition_class : forall {T:Type} (F:Family T) (pff:Finite F) (A:Ensemble T) (pfp:partition_fam F A) (a b:T) (pfa:Inn A a), let S := rep_to_partition_class F pff A pfp a pfa in In S b -> In A b. intros T F h1 A h2 a b h3 S h4. red in h2. destruct h2 as [h2]. unfold S in h4. clear S. subst. pose proof (rep_to_partition_class_compat F h1 (FamilyUnion F) (conj h2 eq_refl) a h3) as h5. simpl in h5. destruct h5 as [h5 h6]. eapply family_union_intro. apply h5. assumption. Qed. (*Wyckoff*) Lemma rep_to_partition_class_functional : forall {T:Type} (F:Family T) (pff pff':Finite F) (A:Ensemble T) (pfp pfp':partition_fam F A) (a a':T) (pfa:Inn A a) (pfa':Inn A a'), a = a' -> rep_to_partition_class F pff A pfp a pfa = rep_to_partition_class F pff' A pfp' a' pfa'. intros T F h1 h2 A h3 h4 a a' h5 h6 h7. subst. pose proof (proof_irrelevance _ h1 h2). pose proof (proof_irrelevance _ h5 h6). pose proof (proof_irrelevance _ h3 h4). subst; auto. Qed. (*Wyckoff*) Lemma in_rep_to_partition_class_eq_iff : forall {T:Type} (F:Family T) (pff pff':Finite F) (A:Ensemble T) (pfp pfp':partition_fam F A) (a a':T) (pfa:In A a) (pfa':In A a'), Inn (rep_to_partition_class F pff A pfp a pfa) a' <-> rep_to_partition_class F pff A pfp a pfa = rep_to_partition_class F pff' A pfp' a' pfa'. intros T F h1 h2 A h3 h4 a a' h5 h6. unfold rep_to_partition_class. destruct constructive_definite_description as [C h9]. destruct constructive_definite_description as [D h10]. simpl. destruct h9 as [h9 h11], h10 as [h10 h12]. split. intro h8. eapply partition_fam_eq. apply h3. assumption. assumption. apply h8. apply h12. intro. subst. assumption. Qed. (*Wyckoff*) Lemma equiv_classes_partition_fam : forall {T:Type} (A:Ensemble T) (R:T->T->Prop), Equivalence R -> partition_fam (rel_classes A R) A. intros T A R h1. destruct h1 as [h1a h1b h1c]. red. split. red. intros S S' h4 h2 h3. destruct h2 as [a h2 S]. subst. destruct h3 as [a' h3 S']. subst. red. assert (h0:~R a a'). intro h5. contradict h4. apply Extensionality_Ensembles. red. split. red. intros x h6. destruct h6 as [h6]. destruct h6 as [h6a h6b]. apply h1b in h6b. specialize (h1c _ _ _ h6b h5). constructor; auto. red. intros x h6. destruct h6 as [h6]. destruct h6 as [h6a h6b]. apply h1b in h6b. apply h1b in h5. specialize (h1c _ _ _ h6b h5). constructor; auto. apply Extensionality_Ensembles; red; split; auto with sets. red. intros x h5. destruct h5 as [x h5 h6]. destruct h5 as [h5]. destruct h6 as [h6]. destruct h5 as [h5a h5b]. destruct h6 as [h6a h6b]. apply h1b in h6b. pose proof (h1c _ _ _ h5b h6b) as h7. contradiction. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [S x h2 h3]. destruct h2 as [y h2 S]. subst. destruct h3 as [h3]. destruct h3; auto. red. intros x h2. apply family_union_intro with [y:T |In A y /\ R x y]. unfold rel_classes. apply Im_intro with x; auto. specialize (h1a x). constructor; auto. Qed. (*Wyckoff*) Lemma partition_fam_S : forall (F:Family nat) (A:Ensemble nat), Finite F -> partition_fam F A-> partition_fam (fam_S F) (ens_S A). intros F A h1 h2. red in h2. red. destruct h2 as [h2 h3]. subst. red in h2. split. red. intros A B h3 h4 h5. destruct h4 as [A h4]. destruct h5 as [B h5]. subst. assert (h6:A <> B). intro h6. subst. contradict h3; auto. specialize (h2 _ _ h6 h4 h5). red in h2. red. rewrite intersection_ens_S. rewrite h2. apply image_empty. unfold ens_S, fam_S. rewrite <- family_union_im. unfold ens_S. reflexivity. Qed. (*Wyckoff*) Lemma rep_to_partition_class_fam_S_eq : forall (F:Family nat) (pff:Finite F) (A:Ensemble nat) (pfp:partition_fam F A) (n:nat) (pfn:Inn A n), rep_to_partition_class (fam_S F) (finite_fam_S _ pff) (ens_S A) (partition_fam_S _ _ pff pfp) (S n) (in_ens_S _ _ pfn) = ens_S (rep_to_partition_class F pff A pfp n pfn). intros F h1 A h2 n h3. unfold rep_to_partition_class, ens_S. destruct constructive_definite_description as [B h4]. destruct constructive_definite_description as [C h5]. simpl. destruct h4 as [h4 h6], h5 as [h5 h7]. destruct h4 as [B h4]. subst. inversion h6 as [m h11 p h12]. subst. apply S_inj in h12. subst. clear h6. pose proof (rep_to_partition_class_ex F h1 A h2 m h3) as h12. destruct h12 as [S h12]. red in h12. destruct h12 as [h12 h13]. pose proof (h13 _ (conj h5 h7)). subst. pose proof (h13 _ (conj h4 h11)). subst. reflexivity. Qed. (*Wyckoff*) Lemma rep_to_partition_class_fam_S_eq' : forall (F:Family nat) (pff:Finite F) (A:Ensemble nat) (pfp:partition_fam F A) (n:nat) (pfn:Inn A n) (pfSn:Inn (ens_S A) (S n)), rep_to_partition_class (fam_S F) (finite_fam_S _ pff) (ens_S A) (partition_fam_S _ _ pff pfp) (S n) pfSn = ens_S (rep_to_partition_class F pff A pfp n pfn). intros F h1 A h2 n h3 h4. pose proof (rep_to_partition_class_fam_S_eq F h1 A h2 n h3) as h5. pose proof (proof_irrelevance _ h4 (in_ens_S _ _ h3)); subst. assumption. Qed. (*Wyckoff*) Lemma in_rep_to_partition_class_fam_S : forall (F:Family nat) (pff:Finite F) (A:Ensemble nat) (pfp:partition_fam F A) (m n:nat) (pfn:Inn A n), In (rep_to_partition_class F pff A pfp n pfn) m -> In (rep_to_partition_class (fam_S F) (finite_fam_S _ pff) (ens_S A) (partition_fam_S _ _ pff pfp) (S n) (in_ens_S _ _ pfn)) (S m). intros F h1 A h2 m n h3 h4. pose proof (rep_to_partition_class_compat (fam_S F) (finite_fam_S F h1) (ens_S A) (partition_fam_S F A h1 h2) (S n) (in_ens_S A n h3)) as h5. simpl in h5. destruct h5 as [h5 h6]. inversion h5 as [B h8 ? h9]. subst. rewrite h9. rewrite rep_to_partition_class_fam_S_eq in h9. apply ens_S_inj in h9. rewrite h9 in h4. econstructor. apply h4. reflexivity. Qed. (*Converse*) (*Wyckoff*) Lemma in_rep_to_partition_class_fam_S_impl : forall (F:Family nat) (pff:Finite F) (A:Ensemble nat) (pfp:partition_fam F A) (m n:nat) (pfn:Inn A n), In (rep_to_partition_class (fam_S F) (finite_fam_S _ pff) (ens_S A) (partition_fam_S _ _ pff pfp) (S n) (in_ens_S _ _ pfn)) (S m) -> In (rep_to_partition_class F pff A pfp n pfn) m. intros F h1 A h2 m n h3 h4. pose proof (rep_to_partition_class_compat (fam_S F) (finite_fam_S F h1) (ens_S A) (partition_fam_S F A h1 h2) (S n) (in_ens_S A n h3)) as h5. simpl in h5. destruct h5 as [h5 h6]. inversion h5 as [B h8 ? h9]. subst. rewrite h9 in h4. inversion h4 as [m' h10 n' h11 h12]. apply S_inj in h11. subst. clear h4. rewrite rep_to_partition_class_fam_S_eq in h9. apply ens_S_inj in h9. subst. assumption. Qed. (*Wyckoff*) Definition im_rel {T U:Type} (f:T->U) := (fun x y:T=>f x = f y). (*Wyckoff*) Lemma equiv_im_rel : forall {T U:Type} (f:T->U), Equivalence (im_rel f). intros T U f. constructor. red. intro x. red. reflexivity. red. intros x y h1. red in h1. red. rewrite h1. reflexivity. red. intros x y z h1 h2. red in h1, h2. red. rewrite h1, h2. reflexivity. Qed. (*Wyckoff*) Definition rel_classes_im_rel {T U:Type} (f:T->U) : Ensemble (Ensemble T) := rel_classes (Full_set T) (im_rel f). (*Wyckoff*) Definition rel_classes_im_rel_set {T U:Type} (A:Ensemble T) (f:T->U) : Ensemble (Ensemble T) := rel_classes A (im_rel f). (*Wyckoff*) Lemma partition_fam_rel_classes_im_rel_set : forall {T U:Type} (A:Ensemble T) (f:T->U), partition_fam (rel_classes_im_rel_set A f) A. intros T U A f. apply equiv_classes_partition_fam. apply equiv_im_rel. Qed. (*Wyckoff*) Definition rel_classes_im_rel_sig_fun {T U:Type} {A:Ensemble T} (f:sig_set A->U) : Ensemble (Ensemble T) := Im (rel_classes (full_sig A) (im_rel f)) (fun A=>Im A (@proj1_sig _ _)). (*Wyckoff*) Lemma rel_classes_im_rel_eq_inv_im_singletons : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), rel_classes_im_rel_sig_fun f = [S:Ensemble T | exists u:U, In (Im (full_sig A) f) u /\ S = (Im (inv_im (Singleton u) f) (@proj1_sig _ _))]. intros T U A f. apply Extensionality_Ensembles. red. split. red. intros B h1. destruct h1 as [B h1]. subst. destruct h1 as [x h1]. subst. clear h1. constructor. exists (f x). split. apply Im_intro with x. constructor. reflexivity. f_equal. unfold inv_im. unfold im_rel. apply Extensionality_Ensembles. red. split. red. intros x' h1. destruct h1 as [h1]. destruct h1 as [h1l h1r]. clear h1l. constructor. rewrite h1r. constructor. red. intros x' h1. destruct h1 as [h1']. inversion h1'. constructor. rewrite H0. split; auto. constructor. red. intros B h1. destruct h1 as [h1]. destruct h1 as [u h1]. destruct h1 as [h1l h1r]. subst. destruct h1l as [x h1]. subst. clear h1. unfold rel_classes_im_rel_sig_fun. unfold rel_classes. rewrite im_im. apply Im_intro with x. constructor. f_equal. unfold inv_im. apply Extensionality_Ensembles. red. split. red. intros x' h1. destruct h1 as [h1]. inversion h1 as [h2]. constructor. split; try constructor. red. assumption. red. intros x' h1. destruct h1 as [h1]. destruct h1 as [h0 h1]. clear h0. constructor. red in h1. rewrite h1. constructor. Qed. (*Wyckoff*) Lemma rel_classes_im_rel_eq_inv_im_singletons' : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), rel_classes_im_rel_sig_fun f = Im (Im (full_sig A) f) (fun u:U => Im (inv_im (Singleton u) f) (@proj1_sig _ _)). intros T U A f. rewrite rel_classes_im_rel_eq_inv_im_singletons. rewrite im_im. apply Extensionality_Ensembles. red. split. red. intros B h1. destruct h1 as [h1]. destruct h1 as [u h1]. destruct h1 as [h1l h1r]. subst. destruct h1l as [u h1]. subst. apply Im_intro with u. assumption. reflexivity. red. intros B h1. destruct h1 as [x h1]. subst. clear h1. constructor. exists (f x). split. apply Im_intro with x. constructor. reflexivity. reflexivity. Qed. (*Wyckoff*) Lemma empty_rel_classes_im_rel_eq_inv_im_singletons : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), rel_classes_im_rel_sig_fun f = Empty_set _ -> A = Empty_set _. intros T U A f h1. apply Extensionality_Ensembles; red; split; auto with sets. red. intros x h2. unfold rel_classes_im_rel_sig_fun in h1. unfold rel_classes in h1. rewrite im_im in h1. assert (h3:In (Im (full_sig A) (fun x : sig_set A => Im [y : sig_set A | In (full_sig A) y /\ im_rel f x y] (proj1_sig (P:=fun x0 : T => In A x0)))) ((fun x : sig_set A => Im [y : sig_set A | In (full_sig A) y /\ im_rel f x y] (proj1_sig (P:=fun x0 : T => In A x0))) (exist _ _ h2))). apply Im_intro with (exist _ _ h2). constructor. reflexivity. rewrite h1 in h3. contradiction. Qed. (*Wyckoff*) Lemma inh_rel_classes_im_rel_sig_fun : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), forall (S:Ensemble T), Inn (rel_classes_im_rel_sig_fun f) S -> Inhabited S. intros T U A f S h2. destruct h2 as [S h2]. subst. destruct h2 as [x h2]. subst. clear h2. apply Inhabited_intro with (proj1_sig x). apply Im_intro with x. constructor. split. constructor. red. reflexivity. reflexivity. Qed. (*Wyckoff*) Lemma inh_rel_classes_im_rel_set : forall {T U:Type} {A:Ensemble T} (f:T->U), forall (S:Ensemble T), Inn (rel_classes_im_rel_set A f) S -> Inhabited S. intros T U A f S h1. destruct h1 as [a h1]. subst. apply Inhabited_intro with a. constructor. split; auto. red. auto. Qed. (*Wyckoff*) Lemma rel_classes_im_rel_sig_fun_consistent_membership : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U) (B:Ensemble T), In (rel_classes_im_rel_sig_fun f) B -> forall (x y:sig_set A), (f x = f y) -> In B (proj1_sig x) -> In B (proj1_sig y). intros T U A f B h1 x y h2 h3. destruct h1 as [B h1]. subst. destruct h1 as [x' h1]. subst. apply Im_intro with y. constructor. split. constructor. inversion h3 as [y' h5]. subst. destruct h5 as [h5]. destruct h5 as [h5l h5r]. clear h5l. red in h5r. red. rewrite <- h2. rewrite h5r. f_equal. destruct y', x. simpl in H. subst. apply proj1_sig_injective. simpl. reflexivity. reflexivity. Qed. (*Wyckoff*) Lemma partition_fam_rel_classes_im_rel_sig_fun : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), partition_fam (rel_classes_im_rel_sig_fun f) A. intros T U A f. pose proof (equiv_classes_partition_fam (full_sig A) _ (equiv_im_rel f)) as h1. red in h1. destruct h1 as [h1l h1r]. red. split. intros S S' h4 h2 h3. destruct h2 as [S h2]. destruct h3 as [S' h3]. subst. red in h1l. assert (h5:S <> S'). intro h5. subst. contradict h4. reflexivity. specialize (h1l _ _ h5 h2 h3). red in h1l. red. apply NNPP. intro h6. apply not_empty_Inhabited in h6. destruct h6 as [x h6]. destruct h6 as [x h6 h7]. destruct h6 as [x h6]. destruct x as [x h8]. subst. simpl in h7. destruct h7 as [x h7]. destruct x as [x h9]. subst. simpl in h6, h8. pose proof (proof_irrelevance _ h8 h9); subst. assert (h10:In (Intersection S S') (exist _ _ h9)). auto with sets. rewrite h1l in h10. contradiction. unfold rel_classes_im_rel_sig_fun. rewrite family_union_im. rewrite h1r. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. apply proj2_sig. red. intros x h2. apply Im_intro with (exist _ _ h2); auto. constructor. Qed. (*Wyckoff*) Lemma im_elt_of_rel_classes_im_rel_sig_fun : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U) (B:Ensemble T) (pfinc:Included B A), In (rel_classes_im_rel_sig_fun f) B -> forall (b:T) (pfb:In B b), Im (full_sig B) (fun x => f (exist _ _ (pfinc _ (proj2_sig x)))) = Singleton (f (exist _ _ (pfinc _ pfb))). intros T U A f B h1 h2 b h3. apply Extensionality_Ensembles. red. split. red. intros u h4. destruct h4 as [u h4]. subst. clear h4. destruct u as [u h4]. destruct h2 as [B h2]. subst. simpl. destruct h2 as [x h2]. subst. clear h2. destruct h3 as [b h3]. subst. destruct h3 as [h3]. destruct h3 as [h3a h3b]. destruct h4 as [u h4]. subst. destruct h4 as [h4]. destruct h4 as [h4a h4b]. red in h3b. red in h4b. destruct b as [b h5]. destruct u as [u h6]. simpl. pose proof h4b as h4b'. rewrite h3b in h4b'. gterm0. redterm0 c. gterm1. redterm0 c1. redterm0 c2. assert (h7: c3 = exist _ _ h5). apply proj1_sig_injective; simpl; auto. fold c3. assert (h8:c0 = exist _ _ h6). unfold c0. apply proj1_sig_injective; simpl; auto. fold c0. rewrite h7, h8 at 1. rewrite h4b' at 1. constructor. red. intros x h4. destruct h4. apply Im_intro with (exist _ _ h3). constructor. simpl. reflexivity. Qed. (*Wyckoff*) Lemma rel_classes_im_rel_sig_fun_setminus : forall {T U:Type} (A B:Ensemble T) (f:sig_set A -> U), In (rel_classes_im_rel_sig_fun f) B -> rel_classes_im_rel_sig_fun (restriction_sig f (Setminus A B) (setminus_inc A B)) = Subtract (rel_classes_im_rel_sig_fun f) B. intros T U A B f h1. unfold restriction_sig. do 2 rewrite rel_classes_im_rel_eq_inv_im_singletons. apply Extensionality_Ensembles. red. split. red. intros S h2. destruct h2 as [h2]. destruct h2 as [u h2]. destruct h2 as [h2 h3]. subst. destruct h2 as [u h2]. subst. clear h2. destruct u as [u h2]. destruct h2 as [h2 h3]. constructor. constructor. exists (f (exist _ _ h2)). split. apply Im_intro with (exist _ _ h2). constructor. reflexivity. unfold restriction_sig. simpl. pose proof (proof_irrelevance _ h2 (setminus_inc A B u (conj h2 h3))) as h4. rewrite <- h4. apply Extensionality_Ensembles. red. split. red. intros x h5. destruct h5 as [x h5]. clear h4. subst. destruct x as [x h6]. destruct h6 as [h6 h7]. simpl. apply Im_intro with (exist _ _ h6). inversion h5 as [h8]. clear h5. simpl in h8. constructor. inversion h8 as [h9]. clear h8. pose proof (proof_irrelevance _ h6 (setminus_inc A B x (conj h6 h7))) as h10. rewrite h9. rewrite <- h10. constructor. simpl. reflexivity. red. intros x h5. destruct h5 as [x h5]. clear h4. subst. destruct h5 as [h5]. inversion h5 as [h6]. clear h5. assert (h7:~In B (proj1_sig x)). intro h8. symmetry in h6. pose proof (rel_classes_im_rel_sig_fun_consistent_membership f B h1 _ _ h6 h8) as h9. simpl in h9. contradiction. destruct x as [x h8]. simpl in h7. simpl. assert (h9:In (Setminus A B) x). auto with sets. apply Im_intro with (exist _ _ h9). constructor. simpl. rewrite h6. pose proof (proof_irrelevance _ h8 (setminus_inc A B x h9)) as h10. subst. constructor. simpl. reflexivity. intro h4. inversion h4 as [h5]. clear h4. simpl in h5. termr h5. assert (h6:Inn y u). unfold y. assert (h6:In (Setminus A B) u). auto with sets. apply Im_intro with (exist _ _ h6). constructor. simpl. pose proof (proof_irrelevance _ h6 (conj h2 h3)) as h7. rewrite h7. constructor. simpl. reflexivity. unfold y in h6. rewrite <- h5 in h6. contradiction. red. intros S h2. destruct h2 as [h2 h3]. assert (h4:S <> B). intro h4. subst. contradict h3. constructor. clear h3. destruct h2 as [h2]. destruct h2 as [u h2]. destruct h2 as [h2 h3]. subst. destruct h2 as [x h2]. subst. clear h2. destruct x as [x h2]. constructor. exists (f (exist _ _ h2)). split. assert (h3:~In B x). nterml h4. assert (h5:Inn x0 x). unfold x0. apply Im_intro with (exist _ _ h2). constructor. constructor. simpl. reflexivity. intro h6. assert (h7:In (rel_classes_im_rel_sig_fun f) x0). unfold x0. rewrite rel_classes_im_rel_eq_inv_im_singletons. constructor. exists (f (exist _ _ h2)). split. apply Im_intro with (exist _ _ h2). constructor. reflexivity. f_equal. pose proof (partition_fam_rel_classes_im_rel_sig_fun f) as h8. red in h8. destruct h8 as [h8 h9]. red in h8. specialize (h8 _ _ h4 h7 h1). red in h8. terml h8. assert (h10:Inn x1 x). unfold x1. auto with sets. unfold x1 in h10. rewrite h8 in h10. contradiction. assert (h5:In (Setminus A B) x). auto with sets. apply Im_intro with (exist _ _ h5). constructor. simpl. f_equal. apply proj1_sig_injective. reflexivity. apply Extensionality_Ensembles. red. split. red. intros x' h5. destruct h5 as [x' h5]. subst. destruct x' as [x' h6]. inversion h5 as [h7]. clear h5. inversion h7 as [h8]. clear h7. simpl. nterml h4. assert (h7:In (rel_classes_im_rel_sig_fun f) x0). unfold x0. rewrite rel_classes_im_rel_eq_inv_im_singletons. constructor. exists (f (exist _ _ h2)). split. apply Im_intro with (exist _ _ h2). constructor. reflexivity. f_equal. unfold x0 in h7. pose proof (partition_fam_rel_classes_im_rel_sig_fun f) as h9. red in h9. destruct h9 as [h9 h10]. specialize (h9 _ _ h4 h7 h1). red in h9. assert (h11:Inn x0 x). unfold x0. apply Im_intro with (exist _ _ h2). constructor. constructor. reflexivity. pose proof (rel_classes_im_rel_sig_fun_consistent_membership _ _ h7 _ _ h8 h11) as h12. simpl in h12. assert (h13:~In B x'). intro h14. terml h9. assert (h15:Inn x1 x'). unfold x1. auto with sets. unfold x1 in h15. rewrite h9 in h15. contradiction. assert (h14:In (Setminus A B) x'). auto with sets. apply Im_intro with (exist _ _ h14). constructor. simpl. rewrite h8 at 1. pose proof (proof_irrelevance _ h6 (setminus_inc A B x' h14)) as h15. rewrite h15. constructor. reflexivity. red. intros x' h5. destruct h5 as [x' h5]. subst. destruct h5 as [h5]. destruct x' as [x' h6]. simpl in h5. inversion h5 as [h7]. clear h5. simpl. destruct h6 as [h6 h8]. apply Im_intro with (exist _ _ h6). constructor. rewrite h7 at 1. pose proof (proof_irrelevance _ h6 (setminus_inc A B x' (conj h6 h8))) as h9. rewrite h9 at 2. constructor. reflexivity. Qed. (*Wyckoff*) Lemma in_rep_to_partition_class_fam_S_seg : forall (F:Family nat) (pff:Finite F) (m n p:nat) (pfp:partition_fam F (seg n)) (pflt:m < n), In (rep_to_partition_class F pff (seg n) pfp m (in_seg _ _ pflt)) p -> In (rep_to_partition_class (fam_S F) (finite_fam_S _ pff) (ens_S (seg n)) (partition_fam_S _ _ pff pfp) (S m) (in_ens_S_seg _ _ pflt)) (S p). intros F h1 m n p h2 h3 h4. apply in_rep_to_partition_class_fam_S in h4. pose proof (proof_irrelevance _ (in_ens_S (seg n) m (in_seg m n h3)) (in_ens_S_seg m n h3)) as h5. rewrite h5 in h4; auto. Qed. (*Wyckoff*) Lemma in_rep_to_partition_class_fam_S_seg_impl : forall (F:Family nat) (pff:Finite F) (m n p:nat) (pfp:partition_fam F (seg n)) (pflt:m < n), In (rep_to_partition_class (fam_S F) (finite_fam_S _ pff) (ens_S (seg n)) (partition_fam_S _ _ pff pfp) (S m) (in_ens_S_seg _ _ pflt)) (S p) -> In (rep_to_partition_class F pff (seg n) pfp m (in_seg _ _ pflt)) p. intros F h1 m n p h2 h3 h4. pose proof (proof_irrelevance _ (in_ens_S_seg _ _ h3) (in_ens_S _ _ (in_seg _ _ h3))) as h5. rewrite h5 in h4. apply in_rep_to_partition_class_fam_S_impl in h4; auto. Qed. End EquivalenceClass. Section EvenOddPos. (*Wyckoff*) Definition pos_evens := [n:nat | 0 < n /\ even n]. (*Wyckoff*) Definition pos_odds := [n:nat | odd n]. (*Wyckoff*) Definition pos_evens_t := {n:nat | 0 < n /\ even n}. (*Wyckoff*) Definition pos_odds_t := {n:nat | odd n}. (*Wyckoff*) Definition pet_lt := exist_rel lt (fun n:nat => 0 < n /\ even n). (*Wyckoff*) Definition pot_lt := exist_rel lt (fun n:nat => odd n). (*Wyckoff*) Definition pet_le := exist_rel le (fun n:nat => 0 < n /\ even n). (*Wyckoff*) Definition pot_le := exist_rel le (fun n:nat => odd n). (*Wyckoff*) Lemma pet_le_refl : forall x, pet_le x x. intro; red; red; auto with arith. Qed. (*Wyckoff*) Lemma pot_le_refl : forall x, pot_le x x. intro; red; red; auto with arith. Qed. (*Wyckoff*) Lemma pet_ge_2 : forall m:pos_evens_t, 2 <= (proj1_sig m). intro m. destruct m as [m h1]. simpl. destruct h1 as [h1 h2]. apply even_decompose in h2. destruct h2 as [k h2]. omega. Qed. (*Wyckoff*) Lemma pet_ge_1 : forall m:pos_evens_t, 1 <= (proj1_sig m). intro m. pose proof (pet_ge_2 m). auto with arith. Qed. (*Wyckoff*) Definition o1 : pos_odds_t := exist _ _ odd_1. (*Wyckoff*) Definition e2 : pos_evens_t := exist _ _ (conj lt_0_2 even_2). (*Wyckoff*) Lemma pos_evens_eq : pos_evens = Im (Full_set nat) (fun n => 2*n + 2). apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1]. destruct h1 as [h1 h2]. apply even_decompose in h2. destruct h2 as [k h2]. assert (h3:k > 0). omega. apply S_pred in h3. apply Im_intro with (pred k). constructor. omega. destruct h1 as [x h1]. subst. constructor. split; try omega. apply even_even_plus. apply even_mult_l. repeat constructor. repeat constructor. Qed. (*Wyckoff*) Lemma pos_odds_eq : pos_odds = Im (Full_set nat) (fun n => 2*n + 1). apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1]. apply odd_decompose in h1. destruct h1 as [k h1]. subst. apply Im_intro with k; constructor. destruct h1 as [x h1]. subst. constructor. apply odd_plus_r. apply even_mult_l. repeat constructor. repeat constructor. Qed. (*Wyckoff*) Definition pet_pred (n:pos_evens_t) : pos_odds_t. pose proof (pet_ge_2 n) as h0. destruct n as [n h1]. destruct h1 as [h1 h2]. simpl in h0. rewrite <- (odd_pred_iff n h1) in h2. refine (exist _ _ h2). Defined. (*Wyckoff*) (*Note: This returns 2 if n = 1, kind of like pred 0 = 0*) Definition pot_pred (n:pos_odds_t) : pos_evens_t. destruct (le_lt_dec 3 (proj1_sig n)) as [h1 | h2]. destruct n as [n h3]. simpl in h1. assert (h4:0 < n). destruct (zerop n) as [h4 | h4]. subst. inversion h1. assumption. rewrite <- (even_pred_iff n h4) in h3. assert (h5:0 < pred n). red. apply le_S_n. simpl. rewrite <- (S_pred _ _ h4). auto with arith. refine (exist _ _ (conj h5 h3)). assert (h6:0 < 2). red. auto with arith. refine (exist _ _ (conj h6 even_2)). Defined. (*Wyckoff*) Lemma pot_pred_one : forall (pf:odd 1), pot_pred (exist _ _ pf) = exist _ _ (conj lt_0_2 even_2). intro h1. unfold pot_pred. simpl. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma le_proj1_sig_pot_pred : forall (m:nat) (n:pos_odds_t), even m -> m <= proj1_sig n -> m <= proj1_sig (pot_pred n). intros m n h1 h2. unfold pot_pred. destruct le_lt_dec as [h3 | h3]. destruct n as [n h4]. simpl. simpl in h2, h3. apply le_even_odd; auto. simpl. omega. Qed. (*Wyckoff*) Lemma le_proj1_sig_pet_pred : forall (m:nat) (n:pos_evens_t), odd m -> m <= proj1_sig n -> m <= proj1_sig (pet_pred n). intros m n h1 h2. unfold pet_pred. destruct n as [n h4]. destruct h4 as [h4 h5]. simpl. simpl in h2. apply le_odd_even; auto. Qed. (*Wyckoff*) Lemma pet_le_pot_pred : forall (m:pos_evens_t) (n:pos_odds_t), 1 < proj1_sig n -> pet_le m (pot_pred n) -> proj1_sig m <= proj1_sig n. intros m n h0 h1. red in h1. unfold exist_rel in h1. unfold pot_pred in h1. destruct le_lt_dec as [h2 | h2]. simpl in h1. destruct n as [n h3]. simpl in h1. simpl. omega. simpl in h1. omega. Qed. (*Wyckoff*) Lemma le_pot_le_pet_pred : forall (n:pos_evens_t) (o:pos_odds_t), proj1_sig o <= proj1_sig n -> pot_le o (pet_pred n). intros n o h1. red. unfold exist_rel. unfold pet_pred. destruct n as [n h2], o as [o h3]. simpl. destruct h2 as [h2 h4]. simpl. simpl in h1. destruct n as [|n]. assert (h5:o = 0). auto with arith. subst. inversion h3. simpl. destruct (le_lt_eq_dec _ _ h1); try omega. subst. pose proof (not_even_and_odd _ h4 h3). contradiction. Qed. (*Wyckoff*) Definition pos_evens_fun_to_sig_fun {T:Type} (f:pos_evens_t -> T) : sig_set (pos_evens)->T. intro n. destruct n as [n h1]. destruct h1 as [h1]. refine (f (exist _ _ h1)). Defined. (*Wyckoff*) Definition pos_odds_fun_to_sig_fun {T:Type} (f:pos_odds_t -> T) : sig_set (pos_odds)->T. intro n. destruct n as [n h1]. destruct h1 as [h1]. refine (f (exist _ _ h1)). Defined. (*Wyckoff*) Lemma inj_pos_evens_fun_to_sig_fun_iff : forall {T:Type} (f:pos_evens_t -> T), injective f <-> injective (pos_evens_fun_to_sig_fun f). intros T f. split. intro h1. red in h1. red. intros x y h2. unfold pos_evens_fun_to_sig_fun in h2. destruct x as [x h3], y as [y h4]. destruct h3 as [h3], h4 as [h4]. apply h1 in h2. apply proj1_sig_injective; simpl. apply exist_injective in h2. assumption. intro h1. red in h1. red. intros x y h2. destruct x as [x h3], y as [y h4]. assert (h5:In pos_evens x). constructor; auto. assert (h6:In pos_evens y). constructor; auto. specialize (h1 (exist _ _ h5) (exist _ _ h6)). unfold pos_evens_fun_to_sig_fun in h1. destruct h5 as [h5], h6 as [h6]. pose proof (proof_irrelevance _ h3 h5); subst. pose proof (proof_irrelevance _ h4 h6); subst. specialize (h1 h2). apply exist_injective in h1. subst. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma inj_pos_odds_fun_to_sig_fun_iff : forall {T:Type} (f:pos_odds_t -> T), injective f <-> injective (pos_odds_fun_to_sig_fun f). intros T f. split. intro h1. red in h1. red. intros x y h2. unfold pos_odds_fun_to_sig_fun in h2. destruct x as [x h3], y as [y h4]. destruct h3 as [h3], h4 as [h4]. apply h1 in h2. apply proj1_sig_injective; simpl. apply exist_injective in h2. assumption. intro h1. red in h1. red. intros x y h2. destruct x as [x h3], y as [y h4]. assert (h5:In pos_odds x). constructor; auto. assert (h6:In pos_odds y). constructor; auto. specialize (h1 (exist _ _ h5) (exist _ _ h6)). unfold pos_odds_fun_to_sig_fun in h1. destruct h5 as [h5], h6 as [h6]. pose proof (proof_irrelevance _ h3 h5). pose proof (proof_irrelevance _ h4 h6); subst. specialize (h1 h2). apply exist_injective in h1. subst. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma surj_pos_evens_fun_to_sig_fun_iff : forall {T:Type} (f:pos_evens_t -> T), surjective f <-> surjective (pos_evens_fun_to_sig_fun f). intros T f. split. intro h1. red in h1. red. intro y. specialize (h1 y). destruct h1 as [x h1]. subst. destruct x as [x h1]. assert (h2:In pos_evens x). constructor; auto. exists (exist _ _ h2). unfold pos_evens_fun_to_sig_fun. destruct h2. f_equal. apply proj1_sig_injective; auto. intro h1. red in h1. red. intro y. specialize (h1 y). destruct h1 as [x h1]. destruct x as [x h2]. destruct h2 as [h2]. exists (exist (fun s=>0 < s /\ even s) _ h2). subst. unfold pos_evens_fun_to_sig_fun. reflexivity. Qed. (*Wyckoff*) Lemma surj_pos_odds_fun_to_sig_fun_iff : forall {T:Type} (f:pos_odds_t -> T), surjective f <-> surjective (pos_odds_fun_to_sig_fun f). intros T f. split. intro h1. red in h1. red. intro y. specialize (h1 y). destruct h1 as [x h1]. subst. destruct x as [x h1]. assert (h2:In pos_odds x). constructor; auto. exists (exist _ _ h2). unfold pos_odds_fun_to_sig_fun. destruct h2. f_equal. apply proj1_sig_injective; auto. intro h1. red in h1. red. intro y. specialize (h1 y). destruct h1 as [x h1]. destruct x as [x h2]. destruct h2 as [h2]. exists (exist (fun s=>odd s) _ h2). subst. unfold pos_odds_fun_to_sig_fun. reflexivity. Qed. (*Wyckoff*) Lemma bij_pos_evens_fun_to_sig_fun_iff : forall {T:Type} (f:pos_evens_t -> T), bijective f <-> bijective (pos_evens_fun_to_sig_fun f). intros T f. unfold bijective. rewrite inj_pos_evens_fun_to_sig_fun_iff. rewrite surj_pos_evens_fun_to_sig_fun_iff. tauto. Qed. (*Wyckoff*) Lemma bij_pos_odds_fun_to_sig_fun_iff : forall {T:Type} (f:pos_odds_t -> T), bijective f <-> bijective (pos_odds_fun_to_sig_fun f). intros T f. unfold bijective. rewrite inj_pos_odds_fun_to_sig_fun_iff. rewrite surj_pos_odds_fun_to_sig_fun_iff. tauto. Qed. End EvenOddPos. Section Unite. (*These "unite functions" take an infinite function on all [odds/evens] and an [[even/odd]/[odd/even]] number and a finite function on some [evens/odds] and unites it all into a single finite function on 1 <= m <= n (seg_one_t n).*) (*Wyckoff*) Definition pos_odds_fun_unite_finite_evens_aux {T:Type} (p:pos_odds_t->T) (n:pos_evens_t) (p':{m:pos_evens_t | pet_le m n}->T) : seg_one_t (proj1_sig n) -> T. intro m. destruct m as [m h1]. destruct (even_odd_dec m) as [h2 | h2]. assert (h3:0 < m). destruct h1. auto with arith. pose (exist (fun a => 0 < a /\ even a) _ (conj h3 h2)) as m'. assert (h4:pet_le m' n). red. red. unfold m'. simpl. destruct h1; auto with arith. refine (p' (exist (fun a => pet_le a n) _ h4)). refine (p (exist _ _ h2)). Defined. (*Wyckoff*) Lemma pos_odds_fun_unite_finite_evens_ex : forall {T:Type} (p:pos_odds_t->T) (n:pos_evens_t) (p':{m:pos_evens_t | pet_le m n}->T), exists !f : seg_one_t (proj1_sig n) -> T, (forall (o:pos_odds_t) (pf:proj1_sig o <= proj1_sig n), p o = f (exist (fun m => 1 <= m <= (proj1_sig n)) _ (conj (odd_ge_1 _ (proj2_sig o)) pf))) /\ (forall (e:pos_evens_t) (pf:pet_le e n), p' (exist (fun m => pet_le m n) _ pf) = f (exist (fun m => 1 <= m <= proj1_sig n) _ (conj (pet_ge_1 e) pf))). intros T p n p'. exists (pos_odds_fun_unite_finite_evens_aux p n p'). red. split. split. intros o h1. unfold pos_odds_fun_unite_finite_evens_aux. destruct even_odd_dec as [h2 | h2]. destruct o as [o h3]. simpl in h2. pose proof (not_even_and_odd _ h2 h3). contradiction. destruct o as [o h3]. f_equal. apply proj1_sig_injective; auto. intros e h4. destruct e as [e h3]. unfold pos_odds_fun_unite_finite_evens_aux. destruct even_odd_dec as [h5 | h5]. f_equal. apply proj1_sig_injective; simpl. apply proj1_sig_injective; auto. simpl in h5. destruct n as [n h6]. destruct h3 as [h3 h7]. pose proof (not_even_and_odd _ h7 h5). contradiction. intros f h1. apply functional_extensionality. intros x. destruct x as [m h2]. destruct h1 as [h1a h1b]. unfold pos_odds_fun_unite_finite_evens_aux. destruct even_odd_dec as [h3 | h3]. assert (h4: 0 < m). omega. specialize (h1b (exist _ _ (conj h4 h3)) (match h2 with | conj _ H0 => H0 end)). terml h1b. gterml. fold x0. fold x in h1b. assert (h6:x = x0). unfold x, x0. f_equal. apply proj1_sig_injective; auto. simpl. apply proj1_sig_injective. simpl. reflexivity. rewrite <- h6 at 1. rewrite h1b. f_equal. apply proj1_sig_injective; auto. destruct h2 as [h2 h4]. specialize (h1a (exist _ _ h3) h4). simpl in h1a. rewrite h1a. f_equal. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Definition pos_odds_fun_unite_finite_evens {T:Type} (p:pos_odds_t->T) (n:pos_evens_t) (p':{m:pos_evens_t | pet_le m n}->T) : seg_one_t (proj1_sig n) -> T := proj1_sig (constructive_definite_description _ (pos_odds_fun_unite_finite_evens_ex p n p')). (*Wyckoff*) Lemma pos_odds_fun_unite_finite_evens_compat : forall {T:Type} (p:pos_odds_t->T) (n:pos_evens_t) (p':{m:pos_evens_t | pet_le m n}->T), let f := pos_odds_fun_unite_finite_evens p n p' in (forall (o:pos_odds_t) (pf:proj1_sig o <= proj1_sig n), p o = f (exist (fun m => 1 <= m <= (proj1_sig n)) _ (conj (odd_ge_1 _ (proj2_sig o)) pf))) /\ (forall (e:pos_evens_t) (pf:pet_le e n), p' (exist (fun m => pet_le m n) _ pf) = f (exist (fun m => 1 <= m <= proj1_sig n) _ (conj (pet_ge_1 e) pf))). unfold pos_odds_fun_unite_finite_evens; intros; destruct constructive_definite_description; auto. Qed. (*Wyckoff*) Definition pos_odds_fun_unite_finite_evens_aux' {T:Type} (p:pos_odds_t->T) (n:pos_odds_t) (p':{m:pos_evens_t | pet_le m (pot_pred n)}->T) : seg_one_t (proj1_sig n) -> T. intro m. destruct m as [m h1]. destruct (even_odd_dec m) as [h2 | h2]. assert (h3:0 < m). destruct h1 as [h1 h3]. red. assumption. pose (exist (fun a => 0 < a /\ even a) _ (conj h3 h2)) as m'. assert (h4:pet_le m' (pot_pred n)). red. red. apply le_proj1_sig_pot_pred; auto. unfold m'. simpl. destruct h1; auto. refine (p' (exist (fun m:pos_evens_t => pet_le m (pot_pred n)) _ h4)). refine (p (exist _ _ h2)). Defined. (*Wyckoff*) Lemma pos_odds_fun_unite_finite_evens_ex' : forall {T:Type} (p:pos_odds_t->T) (n:pos_odds_t) (p':{m:pos_evens_t | pet_le m (pot_pred n)}->T), exists! f:seg_one_t (proj1_sig n) -> T, (forall (o:pos_odds_t) (pf:proj1_sig o <= proj1_sig n), p o = f (exist (fun m => 1 <= m <= (proj1_sig n)) _ (conj (odd_ge_1 _ (proj2_sig o)) pf))) /\ (forall (e:pos_evens_t) (pf:pet_le e (pot_pred n)) (pflt: 1 < proj1_sig n), p' (exist (fun m => pet_le m (pot_pred n)) _ pf) = f (exist (fun m => 1 <= m <= proj1_sig n) _ (conj (pet_ge_1 e) (pet_le_pot_pred _ _ pflt pf)))). intros T p n p'. exists (pos_odds_fun_unite_finite_evens_aux' p n p'). split. split. intros o h1. unfold pos_odds_fun_unite_finite_evens_aux'. destruct even_odd_dec as [h2 | h2]. destruct o as [o h3]. simpl in h2. pose proof (not_even_and_odd _ h2 h3). contradiction. destruct o as [o h3]. f_equal. apply proj1_sig_injective; auto. intros e h4 hlt. destruct e as [e h3]. unfold pos_odds_fun_unite_finite_evens_aux'. destruct even_odd_dec as [h5 | h5]. f_equal. apply proj1_sig_injective; simpl. apply proj1_sig_injective; auto. simpl in h5. destruct n as [n h6]. destruct h3 as [h3 h7]. pose proof (not_even_and_odd _ h7 h5). contradiction. intros f h1. destruct h1 as [h1 h2]. unfold pos_odds_fun_unite_finite_evens_aux'. apply functional_extensionality. intro m. simpl. destruct m as [m h3]. destruct even_odd_dec as [h4 | h4]. assert (h5:0 < m). omega. specialize (h2 (exist _ _ (conj h5 h4))). hfirst h2. red. red. simpl. unfold pot_pred. destruct (le_lt_dec 3 (proj1_sig n)). simpl. destruct n as [n h6]. simpl. simpl in h3, l. destruct n as [|n]. omega. pose proof h6 as h6'. rewrite odd_Sn_iff in h6'. simpl. apply even_decompose in h6'. destruct h6' as [k h6']. pose proof h4 as h4'. apply even_decompose in h4'. destruct h4' as [k' h7]. omega. simpl. omega. specialize (h2 hf). red in hf. red in hf. simpl in hf. unfold pot_pred in hf. destruct n as [n h8]. simpl in h2, hf. simpl in h3. destruct (lt_eq_lt_dec 1 n) as [h10 | h10]. destruct h10 as [h10 | h10]. specialize (h2 h10). terml h2. gterml. fold x0. fold x in h2. assert (he:x = x0). unfold x, x0. f_equal. apply proj1_sig_injective; simpl. apply proj1_sig_injective; simpl. reflexivity. rewrite <- he at 1. rewrite h2. f_equal. apply proj1_sig_injective; simpl. reflexivity. subst. assert (h10: m = 1). omega. subst. inversion h4. inversion H0. omega. clear h2. destruct h3 as [h3 h5]. specialize (h1 (exist _ _ h4)). simpl in h1. specialize (h1 h5). rewrite h1. f_equal. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Definition pos_odds_fun_unite_finite_evens' {T:Type} (p:pos_odds_t->T) (n:pos_odds_t) (p':{m:pos_evens_t | pet_le m (pot_pred n)}->T) : seg_one_t (proj1_sig n) -> T := proj1_sig (constructive_definite_description _ (pos_odds_fun_unite_finite_evens_ex' p n p')). (*Wyckoff*) Lemma pos_odds_fun_unite_finite_evens_compat' : forall {T:Type} (p:pos_odds_t->T) (n:pos_odds_t) (p':{m:pos_evens_t | pet_le m (pot_pred n)}->T), let f := pos_odds_fun_unite_finite_evens' p n p' in (forall (o:pos_odds_t) (pf:proj1_sig o <= proj1_sig n), p o = f (exist (fun m => 1 <= m <= (proj1_sig n)) _ (conj (odd_ge_1 _ (proj2_sig o)) pf))) /\ (forall (e:pos_evens_t) (pf:pet_le e (pot_pred n)) (pflt: 1 < proj1_sig n), p' (exist (fun m => pet_le m (pot_pred n)) _ pf) = f (exist (fun m => 1 <= m <= proj1_sig n) _ (conj (pet_ge_1 e) (pet_le_pot_pred _ _ pflt pf)))). unfold pos_odds_fun_unite_finite_evens'; intros; destruct constructive_definite_description; auto. Qed. (*Wyckoff*) Definition pos_evens_fun_unite_finite_odds_aux {T:Type} (p:pos_evens_t->T) (n:pos_odds_t) (p':{m:pos_odds_t | pot_le m n}->T) : seg_one_t (proj1_sig n) -> T. intro m. destruct m as [m h1]. destruct (even_odd_dec m) as [h2 | h2]. assert (h3:0 < m). destruct h1 as [h1 h3]. red. assumption. pose (exist (fun a => 0 < a /\ even a) _ (conj h3 h2)) as m'. refine (p m'). pose (exist _ _ h2) as m'. assert (h4:pot_le m' n). red. red. unfold m'. simpl. destruct h1; auto. refine (p' (exist (fun a => pot_le a n) _ h4)). Defined. (*Wyckoff*) Lemma pos_evens_fun_unite_finite_odds_ex : forall {T:Type} (p:pos_evens_t->T) (n:pos_odds_t) (p':{m:pos_odds_t | pot_le m n}->T), exists! f:seg_one_t (proj1_sig n) -> T, (forall (o:pos_odds_t) (pf:proj1_sig o <= proj1_sig n), p' (exist (fun m => pot_le m n) _ pf) = f (exist (fun m => 1 <= m <= proj1_sig n) _ (conj (odd_ge_1 _ (proj2_sig o)) pf))) /\ (forall (e:pos_evens_t) (pf:proj1_sig e <= proj1_sig n), p e = f (exist (fun m => 1 <= m <= (proj1_sig n)) _ (conj (pet_ge_1 e) pf))). intros T p n p'. exists (pos_evens_fun_unite_finite_odds_aux p n p'). red. split. split. intros o h1. unfold pos_evens_fun_unite_finite_odds_aux. simpl. destruct o as [o h3]. destruct even_odd_dec as [h2 | h2]. pose proof (not_even_and_odd _ h2 h3). contradiction. f_equal. apply proj1_sig_injective; simpl. apply proj1_sig_injective; auto. intros e h1. destruct e as [e h3]. unfold pos_evens_fun_unite_finite_odds_aux. simpl. destruct even_odd_dec as [h2 | h2]. f_equal. apply proj1_sig_injective; simpl. reflexivity. destruct h3 as [h3 h3']. pose proof (not_even_and_odd _ h3' h2). contradiction. intros f h1. destruct h1 as [h1 h2]. unfold pos_evens_fun_unite_finite_odds_aux. apply functional_extensionality. intro m. destruct m as [m h3]. destruct h3 as [h3 h4]. destruct even_odd_dec as [h5 | h5]. assert (h6:0 < m). red. assumption. specialize (h2 (exist (fun a => 0 < a /\ even a) _ (conj h6 h5))). simpl in h2. specialize (h2 h4). assert (h7:p (exist _ _ (conj h6 h5)) = p (exist _ _ (conj h3 h5))). f_equal. apply proj1_sig_injective; auto. rewrite h7 in h2. rewrite h2 at 1. f_equal. apply proj1_sig_injective; auto. specialize (h1 (exist _ _ h5)). simpl in h1. specialize (h1 h4). terml h1. gterml. fold x in h1; fold x0. assert (h6:x = x0). unfold x, x0. f_equal. rewrite h6 in h1. rewrite h1. f_equal. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Definition pos_evens_fun_unite_finite_odds {T:Type} (p:pos_evens_t->T) (n:pos_odds_t) (p':{m:pos_odds_t | pot_le m n}->T) : seg_one_t (proj1_sig n) -> T := proj1_sig (constructive_definite_description _ (pos_evens_fun_unite_finite_odds_ex p n p')). (*Wyckoff*) Lemma pos_evens_fun_unite_finite_odds_compat : forall {T:Type} (p:pos_evens_t->T) (n:pos_odds_t) (p':{m:pos_odds_t | pot_le m n}->T), let f := pos_evens_fun_unite_finite_odds p n p' in (forall (o:pos_odds_t) (pf:proj1_sig o <= proj1_sig n), p' (exist (fun m => pot_le m n) _ pf) = f (exist (fun m => 1 <= m <= proj1_sig n) _ (conj (odd_ge_1 _ (proj2_sig o)) pf))) /\ (forall (e:pos_evens_t) (pf:proj1_sig e <= proj1_sig n), p e = f (exist (fun m => 1 <= m <= (proj1_sig n)) _ (conj (pet_ge_1 e) pf))). unfold pos_evens_fun_unite_finite_odds; intros; destruct constructive_definite_description; auto. Qed. (*Wyckoff*) Definition pos_evens_fun_unite_finite_odds_aux' {T:Type} (p:pos_evens_t->T) (n:pos_evens_t) (p':{m:pos_odds_t | pot_le m (pet_pred n)}->T) : seg_one_t (proj1_sig n) -> T. intro m. destruct m as [m h1]. destruct (even_odd_dec m) as [h2 | h2]. assert (h3:0 < m). red. destruct h1; auto. pose (exist (fun a => 0 < a /\ even a) _ (conj h3 h2)) as m'. refine (p m'). pose (exist _ _ h2) as m'. assert (h3:pot_le m' (pet_pred n)). red. red. unfold m'. simpl. apply le_proj1_sig_pet_pred; auto. destruct h1; auto. refine (p' (exist (fun a => pot_le a (pet_pred n)) _ h3)). Defined. (*Wyckoff*) Lemma pos_evens_fun_unite_finite_odds_ex' : forall {T:Type} (p:pos_evens_t->T) (n:pos_evens_t) (p':{m:pos_odds_t | pot_le m (pet_pred n)}->T), exists! f:seg_one_t (proj1_sig n) -> T, (forall (o:pos_odds_t) (pf:proj1_sig o <= proj1_sig n), p' (exist (fun m => pot_le m (pet_pred n)) _ (le_pot_le_pet_pred _ _ pf)) = f (exist (fun m => 1 <= m <= proj1_sig n) _ (conj (odd_ge_1 _ (proj2_sig o)) pf))) /\ (forall (e:pos_evens_t) (pf:proj1_sig e <= proj1_sig n), p e = f (exist (fun m => 1 <= m <= (proj1_sig n)) _ (conj (pet_ge_1 e) pf))). intros T p n p'. exists (pos_evens_fun_unite_finite_odds_aux' p n p'). red. split. split. intros o h1. unfold pos_evens_fun_unite_finite_odds_aux'. destruct o as [o h3]. destruct even_odd_dec as [h2 | h2]. simpl in h2. pose proof (not_even_and_odd _ h2 h3). contradiction. f_equal. apply proj1_sig_injective; simpl. apply proj1_sig_injective; auto. intros e h1. destruct e as [e h4]. unfold pos_evens_fun_unite_finite_odds_aux'. destruct even_odd_dec as [h2 | h2]. simpl. f_equal. apply proj1_sig_injective; simpl. reflexivity. simpl in h2. destruct h4 as [h4 h5]. pose proof (not_even_and_odd _ h5 h2). contradiction. intros f h1. destruct h1 as [h1 h2]. apply functional_extensionality. intro m. destruct m as [m h3]. unfold pos_evens_fun_unite_finite_odds_aux'. destruct even_odd_dec as [h4 | h4]. destruct h3 as [h3 h5]. assert (h6:0 < m). red. assumption. specialize (h2 (exist _ _ (conj h6 h4))). simpl in h2. specialize (h2 h5). terml h2. gterml. fold x in h2. fold x0. assert (h7:x = x0). unfold x, x0. f_equal. apply proj1_sig_injective; auto. rewrite h7 in h2. rewrite h2 at 1. f_equal. apply proj1_sig_injective; auto. specialize (h1 (exist _ _ h4)). simpl in h1. destruct h3 as [h3 h5]. specialize (h1 h5). gterml. terml h1. assert (h6:x = x0). unfold x, x0. f_equal; apply proj1_sig_injective; simpl; auto. unfold x, x0 in h6. rewrite h6, h1; f_equal; apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Definition pos_evens_fun_unite_finite_odds' {T:Type} (p:pos_evens_t->T) (n:pos_evens_t) (p':{m:pos_odds_t | pot_le m (pet_pred n)}->T) := proj1_sig (constructive_definite_description _ (pos_evens_fun_unite_finite_odds_ex' p n p')). (*Wyckoff*) Lemma pos_evens_fun_unite_finite_odds_compat' : forall {T:Type} (p:pos_evens_t->T) (n:pos_evens_t) (p':{m:pos_odds_t | pot_le m (pet_pred n)}->T), let f:= pos_evens_fun_unite_finite_odds' p n p' in (forall (o:pos_odds_t) (pf:proj1_sig o <= proj1_sig n), p' (exist (fun m => pot_le m (pet_pred n)) _ (le_pot_le_pet_pred _ _ pf)) = f (exist (fun m => 1 <= m <= proj1_sig n) _ (conj (odd_ge_1 _ (proj2_sig o)) pf))) /\ (forall (e:pos_evens_t) (pf:proj1_sig e <= proj1_sig n), p e = f (exist (fun m => 1 <= m <= (proj1_sig n)) _ (conj (pet_ge_1 e) pf))). unfold pos_evens_fun_unite_finite_odds'; intros; destruct constructive_definite_description; auto. Qed. End Unite. End RandomFacts. Notation "pr ||-> x" := (sig_fun_app (fst pr) (snd pr) x) (at level 20). Section RandomFacts'. (*Wyckoff*) Lemma incl_im_im_left_inverse_compose : forall {T U V:Type} (f:T->U) (pff:injective f) (A:Ensemble T) (deft:T) (g:T->V), Included (Im (Im A f) (fun u : U => g ((left_inverse f pff A, deft) ||-> u))) (Im A g). intros T U V f h1 A deft g. red. intros y h2. rewrite im_im in h2. destruct h2 as [x h2]. subst. apply Im_intro with x; auto. simpl. unfold sig_fun_app. lr_match_goal'. f_equal. erewrite f_equal_dep. rewrite left_inverse_compat. reflexivity. reflexivity. contradict hr. apply Im_intro with x; auto. Grab Existential Variables. assumption. Qed. (*Wyckoff*) Lemma in_eq_rect : forall {T U:Type} (pf:T=U) (A:Ensemble T) (x:U), In (eq_rect _ (fun V=>Ensemble V) A _ pf) x <-> In A (eq_rect_r id x pf). intros; subst; simpl; tauto. Qed. (*Wyckoff*) Lemma in_eq_rect_r : forall {T U:Type} (pf:T=U) (A:Ensemble U) (x:T), In (eq_rect_r (fun V=>Ensemble V) A pf) x <-> In A (eq_rect _ id x _ pf). intros; subst; simpl; tauto. Qed. (*Wyckoff*) Definition max_set_nat (E:Ensemble nat) (pf:Finite E) (pfi:Inhabited E) := max_set_sortable E pf pfi lt lt_sortable. (*Wyckoff*) Lemma max_set_nat_compat : forall (A:Ensemble nat) (pff:Finite A) (pfi:Inhabited A), let n := max_set_nat A pff pfi in In A n /\ (forall m:nat, In (Subtract A n) m -> m < n). intros A h1 h2 n. unfold n, max_set_nat, max_set_sortable. destruct constructive_definite_description as [n' h3]. simpl. assumption. Qed. (*Wyckoff*) Lemma max_set_nat_functional : forall (A A':Ensemble nat) (pff:Finite A) (pff':Finite A') (pfi:Inhabited A) (pfi':Inhabited A'), A = A' -> max_set_nat A pff pfi = max_set_nat A' pff' pfi'. intros A A' h1 h1' h2 h2' ?. subst. pose proof (proof_irrelevance _ h1 h1'). pose proof (proof_irrelevance _ h2 h2'). subst; auto. Qed. (*Wyckoff*) Lemma max_set_nat_sing : forall (a:nat), max_set_nat (Singleton a) (Singleton_is_finite _ _) (inh_sing _) = a. intro a. pose proof (max_set_nat_compat (Singleton a) (Singleton_is_finite _ _) (inh_sing _)) as h1. simpl in h1. destruct h1 as [h1 h2]. rewrite in_sing_iff in h1; auto. Qed. (*Wyckoff*) Lemma lt_max_set_nat_add : forall (A:Ensemble nat) (pf:Finite A) (pfi:Inhabited A) (m:nat), let pf' := Add_preserves_Finite _ A m pf in let pfi' := inh_add A m in m = max_set_nat (Add A m) pf' pfi' -> ~Inn A m -> forall (n:nat), Inn A n -> n < m. intros A h1 h2 m h3 h4 h5 h6 n h7. pose proof (max_set_nat_compat (Add A m) h3 h4) as h8. destruct h8 as [h8 h9]. specialize (h9 n). inversion h8 as [a hm | a hm]. rewrite h5 in h6. contradiction. rewrite in_sing_iff in hm. rewrite hm. apply h9. constructor. left; auto. intro h10; destruct h10. rewrite h5 in h6. contradiction. Qed. (*Wyckoff*) Definition max_set_nat_sig_ex : forall {P:nat->Prop} (A:Ensemble {m:nat | P m}), Finite A -> Inhabited A -> exists! n:{m:nat | P m}, In A n /\ (forall r:{m:nat | P m}, In A r -> (proj1_sig r) <= (proj1_sig n)). intros P A h1 h2. rewrite <- finite_im_proj1_sig_iff in h1. rewrite <- inh_im_proj1_sig_iff in h2. pose proof (max_set_nat_compat _ h1 h2) as h3. simpl in h3. destruct h3 as [h3 h4]. inversion h3 as [n h6 q h7]. subst. exists n. red. split. split; auto. intros r h8. destruct (nat_eq_dec (proj1_sig r) (proj1_sig n)) as [h9 | h9]. rewrite h9. auto with arith. specialize (h4 (proj1_sig r)). hfirst h4. constructor. apply Im_intro with r; auto. rewrite in_sing_iff. rewrite <- h7 in h9. apply neq_sym; auto. specialize (h4 hf). rewrite h7 in h4. auto with arith. intros m' h8. destruct h8 as [h8 h9]. specialize (h4 (proj1_sig m')). specialize (h9 n h6). apply le_lt_eq_dec in h9. destruct h9 as [h9 | h9]. hfirst h4. constructor. apply Im_intro with m'; auto. rewrite in_sing_iff. intro h10. rewrite <- h7 in h9 at 1. omega. specialize (h4 hf). rewrite h7 in h4. pose proof (lt_trans _ _ _ h4 h9) as h10. apply lt_irrefl in h10; contradiction. apply proj1_sig_injective in h9; auto. Qed. (*Wyckoff*) Definition max_set_nat_sig {P:nat->Prop} (A:Ensemble {m:nat | P m}) (pff:Finite A) (pfi:Inhabited A) := proj1_sig (constructive_definite_description _ (max_set_nat_sig_ex A pff pfi)). (*Wyckoff*) Definition max_set_nat_sig_compat : forall {P:nat->Prop} (A:Ensemble {m:nat | P m}) (pff:Finite A) (pfi:Inhabited A), let n := max_set_nat_sig A pff pfi in In A n /\ (forall r:{m:nat | P m}, In A r -> (proj1_sig r) <= (proj1_sig n)). intros; unfold n, max_set_nat_sig; intros; destruct constructive_definite_description; auto. Qed. (*Wyckoff*) Definition min_set_nat (E:Ensemble nat) (pf:Finite E) (pfi:Inhabited E) := min_set_sortable E pf pfi lt lt_sortable. (*Wyckoff*) Lemma min_set_nat_functional : forall (E E':Ensemble nat) (pf:Finite E) (pf':Finite E') (pfi:Inhabited E) (pfi':Inhabited E'), E = E' -> min_set_nat E pf pfi = min_set_nat E' pf' pfi'. intros; subst. f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma min_set_nat_compat : forall (E:Ensemble nat) (pf:Finite E) (pfi:Inhabited E), let m := min_set_nat E pf pfi in In E m /\ (forall n:nat, In (Subtract E m) n -> m < n). intros E h1 h2. simpl. unfold min_set_nat. pose proof (min_set_sortable_compat E h1 h2 lt lt_sortable) as h3. simpl in h3. destruct h3 as [h3 h4]. split; auto. Qed. (*Wyckoff*) Lemma min_set_nat_add : forall (E:Ensemble nat) (pff:Finite E) (pfi:Inhabited E) (m:nat), min_set_nat (Add E m) (Add_preserves_Finite _ _ _ pff) (inh_add E m) = min (min_set_nat E pff pfi) m. intros E h1 h2 m. unfold min_set_nat. pose proof (min_set_sortable_add E m (Add_preserves_Finite _ _ _ h1) (inh_add _ _) lt (sort_intro _ tso_lt lt_trans lt_irrefl)) as h3. destruct h3 as [h3 h4]. specialize (h3 h2). simpl in h3. unfold lt_sortable. rewrite h3. unfold minR. lr_if_goal; fold hlr; destruct hlr as [[ha | ha] | ha]. rewrite min_l. f_equal. apply proof_irrelevance. erewrite min_set_sortable_functional; auto. apply lt_le_weak in ha. apply ha; auto. rewrite min_r. rewrite <- ha. apply min_set_sortable_functional; auto. erewrite min_set_sortable_functional in ha; auto. rewrite <- ha. apply le_refl. rewrite min_r; auto with arith. erewrite min_set_sortable_functional in ha; auto. apply lt_le_weak in ha. apply ha. Qed. (*Wyckoff*) Lemma min_set_nat_0 : forall (E:Ensemble nat) (pff:Finite E) (pfi:Inhabited E), Inn E 0 -> min_set_nat E pff pfi = 0. intros E h1 h2 h3. destruct (zerop (min_set_nat E h1 h2)) as [h4 | h4]; auto. pose proof (min_set_nat_compat E h1 h2) as h5. simpl in h5. destruct h5 as [h5 h6]. specialize (h6 0). hfirst h6. constructor; auto. rewrite in_sing_iff. omega. specialize (h6 hf). omega. Qed. (*Wyckoff*) Lemma in_im_seg_sig : forall {T:Type} {D:Ensemble T} (f:nat->sig_set D) (n:nat) (x:sig_set (Im (seg n) (fun i => proj1_sig (f i)))), Inn D (proj1_sig x). intros T D f n x. destruct x as [x hx]; simpl. destruct hx as [x]; subst. apply proj2_sig. Qed. (*Wyckoff*) Definition inv_im_seg_min {U:Type} (f:nat->U) (n:nat) (y:sig_set (Im (seg n) f)) := min_set_nat (inv_im_seg (fun_to_sig_seg f n) y) (finite_inv_im_seg (fun_to_sig_seg f n) y) (inh_inv_im_sig_seg f n y). (*Wyckoff*) Lemma inv_im_seg_min_compat : forall {U:Type} (f:nat->U) (n:nat) (y:sig_set (Im (seg n) f)), f (inv_im_seg_min f n y) = proj1_sig y. intros U f n y. unfold inv_im_seg_min. pose proof (min_set_nat_compat (inv_im_seg (fun_to_sig_seg f n) y) (finite_inv_im_seg (fun_to_sig_seg f n) y) (inh_inv_im_sig_seg f n y)) as h1. simpl in h1. destruct h1 as [h1 h2]. destruct h1 as [h1]. destruct h1 as [h1 h3]. rewrite fun_to_sig_seg_compat in h3. rewrite <- h3 at 4. simpl; auto. Qed. (*Wyckoff*) Lemma inv_im_seg_min_compat' : forall {U:Type} (f:nat->U) (n:nat) (y:sig_set (Im (seg n) f)), Inn (inv_im1 f (proj1_sig y)) (inv_im_seg_min f n y). intros U f n y. unfold inv_im1. constructor. apply inv_im_seg_min_compat. Qed. (*Wyckoff*) Lemma inv_im_seg_min_sig_compat : forall {U:Type} {E:Ensemble U} (f:nat->sig_set E) (n:nat) (y:sig_set (Im (seg n) (proj1_sig_fun f))), f (inv_im_seg_min (proj1_sig_fun f) n y) = exist _ _ (in_proj1_sig_fun f _ _ (proj2_sig y)). intros U E f n y. pose proof (inv_im_seg_min_compat (proj1_sig_fun f) n y) as h1. rewrite proj1_sig_fun_compat in h1. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma lt_inv_im_seg_min : forall {U:Type} (f:nat->U) (n:nat) (y:sig_set (Im (seg n) f)), inv_im_seg_min f n y < n. intros U f n y. unfold inv_im_seg_min. pose proof (min_set_nat_compat (inv_im_seg (fun_to_sig_seg f n) y) (finite_inv_im_seg (fun_to_sig_seg f n) y) (inh_inv_im_sig_seg f n y)) as h1. simpl in h1. destruct h1 as [h1 h2]. destruct h1 as [h1]. destruct h1; auto. Qed. (*Wyckoff*) Lemma in_inv_im_seg_min : forall {U:Type} (f:nat->U) (pf:surjective f) (n:nat) (y:sig_set (Im (seg n) f)), Inn (inv_im_seg (fun_to_sig_seg f n) y) (inv_im_seg_min f n y). intros; constructor. exists (lt_inv_im_seg_min _ _ _). destruct y as [y hy]. destruct hy as [i hy]; subst. rewrite fun_to_sig_seg_compat. apply proj1_sig_injective; simpl. rewrite inv_im_seg_min_compat; simpl; auto. Qed. (*Wyckoff*) Lemma le_inv_im_seg_min : forall {U:Type} (f:nat->U) (pf:surjective f) (i n:nat) (pf:Inn (seg n) i), inv_im_seg_min f n (exist _ _ (Im_intro _ _ (seg n) f i pf _ eq_refl)) <= i. intros U f h1 i n h2. destruct h2 as [h2]. unfold inv_im_seg_min. glel. redterm0 xle. redterm1 xle. pose proof (min_set_nat_compat _ c0 c) as h3. unfold c, c0 in h3; simpl in h3. destruct h3 as [h3 h4]. destruct (le_lt_dec xle i) as [h5 | h5]. assumption. specialize (h4 i). hfirst h4. constructor. constructor. exists h2. rewrite fun_to_sig_seg_compat. apply proj1_sig_injective; auto. rewrite in_sing_iff. intro h6. unfold xle in h5. rewrite h6 in h5. apply lt_irrefl in h5; auto. specialize (h4 hf). omega. Qed. (*Wyckoff*) Definition nat_sig_to_fun {U V:Type} {D:Ensemble U} {E:Ensemble V} (f:nat->sig_set D) (g:nat->sig_set E) (n:nat) : sig_set (Im (seg n) (proj1_sig_fun f)) -> V := fun x => proj1_sig (g (inv_im_seg_min (proj1_sig_fun f) n x)). (*Wyckoff*) Lemma bij_seg : forall {T:Type} (A:Ensemble T), Finite A -> exists f:sig_set (seg (card_fun1 A)) -> sig_set A, bijective f. intros T A h1. induction h1 as [|A h1 h2 x h3]. rewrite card_fun1_empty. exists (fun x:sig_set (seg 0) => False_rect (sig_set (Empty_set T)) (nin_seg_0 _ (proj2_sig x))). red. split. red. intros x y h1. destruct x as [x h2]. pose proof (nin_seg_0 _ h2). contradiction. red. intro y. destruct y as [y h1]. contradict h1. intro h1. contradiction. rewrite card_add_nin'; auto. rewrite seg_Sn. destruct h2 as [f h2]. red in h2. destruct h2 as [h2a h2b]. exists (fun c:sig_set (Add (seg (card_fun1 A)) (card_fun1 A)) => match (in_add_nin_dec _ _ (nin_seg _) _ (proj2_sig c)) with | left P => exist (fun a => In (Add A x) a) _ (Add_intro1 _ _ x _ (proj2_sig (f (exist _ _ P)))) | right _ => exist _ _ (Add_intro2 _ A x) end). red. split. red. intros m n h0. lr_match' h0. rename hl into h4. lr_match' h0. rename hl into h6. apply exist_injective in h0. apply proj1_sig_injective in h0. apply h2a in h0. apply exist_injective in h0. apply proj1_sig_injective in h0. assumption. apply exist_injective in h0. subst. contradict h3. apply proj2_sig. rename hr into h5. lr_match' h0. rename hl into h6. apply exist_injective in h0. subst. contradict h3. apply proj2_sig. rename hr into h7. rewrite <- h7 in h5. apply proj1_sig_injective in h5. assumption. red. intro y. destruct y as [y h4]. destruct h4 as [y h4 | y h5]. red in h2b. specialize (h2b (exist _ _ h4)). destruct h2b as [a h2b]. destruct a as [a h6]. exists (exist _ _ (Add_intro1 _ _ (card_fun1 A) _ h6)). simpl. lr_match_goal'. rename hl into h7. apply proj1_sig_injective; simpl. pose proof (f_equal (@proj1_sig _ _) h2b) as h8. simpl in h8. subst. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; simpl. subst. pose proof (nin_seg _ h6). contradiction. destruct h5. exists (exist _ _ (Add_intro2 _ (seg (card_fun1 A)) (card_fun1 A))). simpl. lr_match_goal'. rename hl into h4. apply proj1_sig_injective; simpl. pose proof (nin_seg _ h4). contradiction. apply proj1_sig_injective; simpl. reflexivity. Qed. (*Use is made of Full_set rather than T to exploit the previous theorem easier. See [SetUtilities2], for the version of this with an arbitrary set instead of Full_set. *) (*Wyckoff*) Lemma bij_seg_inh_full_iff : forall {T:Type} (f:sig_set (seg (card_fun1 (Full_set T))) -> sig_set (Full_set _)), bijective f -> forall (P:T->Prop), (Inhabited [x:T | P x] <-> Inhabited [n : nat | exists pf : In (seg (card_fun1 (Full_set T))) n, P (proj1_sig (f (exist _ _ pf)))]). intros T f h1 P. apply bijective_impl_invertible in h1. pose (proj1_sig (function_inverse _ h1)) as f'. split. intro h2. destruct h2 as [x h2]. destruct h2 as [h2]. apply Inhabited_intro with (proj1_sig (f' (exist _ _ (Full_intro _ x)))). constructor. exists (proj2_sig _). pose proof (proj2_sig (function_inverse f h1)) as h3. simpl in h3. destruct h3 as [h3 h4]. gterm0. assert (h5:c = x). unfold c. rewrite <- unfold_sig. unfold f'. rewrite h4. simpl. reflexivity. unfold c in h5. rewrite h5 at 1. assumption. intro h2. destruct h2 as [n h2]. destruct h2 as [h2]. destruct h2 as [h2 h3]. apply Inhabited_intro with (proj1_sig (f (exist _ _ h2))). constructor. assumption. Qed. Definition Im_iter {T:Type} (f:T->T) (x0:T) (n:nat): Ensemble T := Im (seg n) (fun m:nat => iter_uo f x0 m). Lemma im_iter_zero : forall {T:Type} (f:T->T) (x0:T), Im_iter f x0 0 = Empty_set _. intros T f x0. unfold Im_iter. rewrite seg_O. rewrite image_empty. reflexivity. Qed. Lemma im_iter_Sn : forall {T:Type} (f:T->T) (x0:T) (n:nat), Im_iter f x0 (S n) = Add (Im_iter f x0 n) (iter_uo f x0 n). intros T f x0 n. unfold Im_iter. rewrite seg_Sn. rewrite Im_add. reflexivity. Qed. Lemma finite_im_iter : forall {T:Type} (f:T->T) (x0:T) (n:nat), Finite (Im_iter f x0 n). intros T f x0 n. unfold Im_iter. apply finite_image. apply finite_seg. Qed. Lemma im_iter_strict_order_Sn : forall {T:Type} (f:T->T) (x0:T) (n:nat) (R:T->T->Prop), Strict_Order R -> (forall x:T, R (f x) x) -> R (iter_uo f x0 (S n)) (iter_uo f x0 n). intros T f x0 n R h1 h2. apply h2. Qed. Lemma im_iter_strict_order_preservers_lt : forall {T:Type} (f:T->T) (x0:T) (m n:nat), m < n -> forall (R:T->T->Prop), Strict_Order R -> (forall x:T, R (f x) x) -> R (iter_uo f x0 n) (iter_uo f x0 m). intros T f x0 m n. revert m. induction n as [|n h1]. intros; omega. intros m h2 R h3 h4. red in h2. apply le_S_n in h2. apply le_lt_eq_dec in h2. destruct h2 as [h5 | h6]. specialize (h1 m h5 R h3 h4). simpl. pose proof (h4 (iter_uo f x0 n)) as h6. destruct h3 as [h3a h3b]. eapply h3a. apply h6. apply h1. subst. apply h4. Qed. Lemma card_im_iter_strict_order : forall {T:Type} (f:T->T) (x0:T) (n:nat) (R:T->T->Prop), Strict_Order R -> (forall x:T, R (f x) x) -> card_fun1 (Im_iter f x0 n) = n. intros T f x0 n R h1 h2. revert R h1 h2. induction n as [|n h1]. intros R h1 h2. rewrite im_iter_zero. rewrite card_fun1_empty. reflexivity. intros R h2 h3. specialize (h1 R h2 h3). rewrite im_iter_Sn. rewrite card_add_nin'. rewrite h1. reflexivity. apply finite_im_iter. intro h4. inversion h4 as [m h5 x h6]. clear h1. subst. destruct h5 as [h5]. pose proof (im_iter_strict_order_preservers_lt f x0 m n h5 R h2 h3) as h8. rewrite h6 in h8. destruct h2 as [h2a h2b]. apply h2b in h8; auto. Qed. Lemma strict_order_decreasing_uo_infinite : forall {T:Type}, Inhabited (Full_set T) -> forall (f:T->T) (R:T->T->Prop), Strict_Order R -> (forall x:T, R (f x) x) -> ~Finite (Full_set T). intros T h1 f R h2 h3 h4. pose (card_fun1 (Full_set T)) as n. destruct h1 as [x h1]. pose proof (card_im_iter_strict_order f x (S n) R h2 h3) as h5. pose proof (incl_card_fun1 (Im_iter f x (S n)) _ h4) as h7. hfirst h7. red; intros; constructor. specialize (h7 hf). rewrite h5 in h7. unfold n in h7. omega. Qed. (*Wyckoff*) Lemma le_prop_impl : forall {T:Type} {R:T->T->Prop}, Transitive R -> forall (t t':T), R t t' -> forall (x:T), (x = t \/ R x t) -> (x = t' \/ R x t'). intros T R h1 t t' h2 x h3. destruct h3 as [h3 | h4]; subst. right; auto. red in h1. specialize (h1 _ _ _ h4 h2). right; auto. Qed. (*Wyckoff*) Definition le_sig_conv_ex : forall {T:Type} {R:T->T->Prop} {t t':T}, Transitive R -> forall (pf:R t t') (x':{x:T | t = x \/ R x t}), exists ! x'':{x:T | t' = x \/ R x t'}, proj1_sig x' = proj1_sig x''. intros T R t t' h0 h1 x'. destruct x' as [x' h2]. destruct h2 as [h2 | h3]. subst. exists (exist (fun x:T => t' = x \/ R x t') _ (or_intror h1)). red; simpl; split. reflexivity. intros x'' h2. destruct x'' as [x'' h3]. subst. simpl in h1. simpl. apply proj1_sig_injective; auto. red in h0. specialize (h0 _ _ _ h3 h1). exists (exist (fun x:T => t' = x \/ R x t') _ (or_intror h0)). red; split; simpl; auto. intros x'' h4. destruct x'' as [x'' h5]. simpl in h4. subst. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Definition le_sig_conv {T:Type} {R:T->T->Prop} {t t':T} (pft:Transitive R) (pfr:R t t') (x':{x:T | t = x \/ R x t}) : {x:T | t' = x \/ R x t'} := proj1_sig (constructive_definite_description _ (le_sig_conv_ex pft pfr x')). (*Wyckoff*) Lemma le_sig_conv_cond : forall {T:Type} {R:T->T->Prop} {t t':T} (pft:Transitive R) (pfr:R t t') (x':{x:T | t = x \/ R x t}), let x'' := le_sig_conv pft pfr x' in proj1_sig x' = proj1_sig x''. intros. unfold x'', le_sig_conv. destruct constructive_definite_description as [z h2]. simpl. assumption. Qed. (*Wyckoff*) Lemma overlapping_functions_constructive_definite_description : forall {T U:Type} (R:T->T->Prop) (pft:Transitive R) (P:{x:T & ({y : T | y = x \/ R y x} -> U)}->Prop) (pf: forall x : T, exists ! f : {y : T | y = x \/ R y x} -> U, P (existT _ x f)), let psi := (fun x:T => proj1_sig (constructive_definite_description _ (pf x))) in (forall (x y:T) (pfr:R x y) (f:{y0 : T | y0 = y \/ R y0 y}->U), let f' := restriction_sig_prop (fun t:T => t = x \/ R t x) (fun t:T => t = y \/ R t y) (le_prop_impl pft _ _ pfr) f in P (existT _ y f) -> P (existT _ x f')) -> forall (x y:T) (pfr:R x y), psi x (exist _ _ (or_introl eq_refl)) = psi y (exist _ _ (or_intror pfr)). intros T U R h0 P h1 psi hp x y h2. unfold psi. destruct constructive_definite_description as [f h3]. destruct constructive_definite_description as [g h4]. simpl. specialize (hp x y h2 g h4). pose proof (h1 x) as h5. destruct h5 as [f' h5]. red in h5. destruct h5 as [h5 h6]. apply h6 in hp. subst. apply h6 in h3. subst. unfold restriction_sig_prop. f_equal. apply proj1_sig_injective; auto. Qed. End RandomFacts'. Arguments FamilyUnion [T] F _. Arguments FamilyIntersection [T] F _. Arguments IndexedUnion [It] [T] FI _. Arguments IndexedIntersection [It] [T] FI _. Arguments Union_associative [Xt] N M P. Arguments assoc_prod_psa [T] N M P. Arguments Union_commutative [Xt] N M. Arguments abs_sum_psa [T] N M. Arguments abs_prod_psa [T] N M. Arguments comp_sum_psa [Xt] N. Arguments comp_prod_psa [Xt] N. Section ArithSets. (*Wyckoff*) Definition lb_nat (A:Ensemble nat) (b:nat) := forall n:nat, In A n -> b <= n. (*Wyckoff*) Lemma lb_nat_lt : forall (A:Ensemble nat) (b:nat), lb_nat A b -> forall (b':nat), b' < b -> lb_nat A b'. intros A b h1 b' h2. red in h1. red. intros n h3. specialize (h1 n h3). omega. Qed. (*Wyckoff*) Definition glb_nat (A:Ensemble nat) (b:nat) := lb_nat A b /\ forall b':nat, lb_nat A b' -> b' <= b. (*Wyckoff*) Definition nat_translate (n:nat) := [m:nat | m >= n]. (*Wyckoff*) Lemma nat_translate_inj : forall (m n:nat), nat_translate m = nat_translate n -> m = n. intros m n h1. destruct (lt_eq_lt_dec m n) as [h2 | h3]. destruct h2 as [h2 | h3]. assert (h3:m <= pred n). destruct (zerop n) as [h3 | h4]. subst. simpl. omega. omega. assert (h4:In (nat_translate m) (pred n)). constructor. omega. rewrite h1 in h4. destruct h4 as [h4]. assert (h5:n = 0). omega. subst. omega. assumption. assert (h3':n <= pred m). destruct (zerop m) as [h3' | h4']. subst. simpl. omega. omega. assert (h4:In (nat_translate n) (pred m)). constructor. omega. rewrite <- h1 in h4. destruct h4 as [h4]. assert (h5:n = 0). omega. subst. omega. Qed. (*Wyckoff*) Definition nat_translatable (S:Ensemble nat) := exists n:nat, S = nat_translate n. (*Wyckoff*) Definition nat_translatable' (S:Ensemble nat) := exists! n:nat, S = nat_translate n. (*Wyckoff*) Definition nat_translatable_endpoint (S:Ensemble nat) (pf:nat_translatable' S) := proj1_sig (constructive_definite_description _ pf). (*Wyckoff*) Lemma nat_translatable_endpoint_compat : forall (S:Ensemble nat) (pf:nat_translatable' S), S = nat_translate (nat_translatable_endpoint S pf). intros S h1. unfold nat_translatable_endpoint. destruct constructive_definite_description as [n h2]. subst. simpl. reflexivity. Qed. (*Wyckoff*) Lemma nat_translatable_prime_compat : forall S:Ensemble nat, nat_translatable S <-> nat_translatable' S. intro S. split. intro h1. destruct h1 as [n h1]. subst. red. exists n. red. split; auto. intros n' h1. apply nat_translate_inj in h1. assumption. intro h1. destruct h1 as [n h1]. red in h1. destruct h1 as [h1 h2]. subst. red. exists n. reflexivity. Qed. (*Wyckoff*) Lemma nat_translate_0 : nat_translate 0 = Full_set nat. apply Extensionality_Ensembles. red; split; red; intros n h1; try constructor; omega. Qed. (*Wyckoff*) Lemma nat_translatable_iff : forall A:Ensemble nat, Inhabited A -> (nat_translatable A <-> (forall m n:nat, In A m -> m <= n -> In A n)). intros A h1. split. intros h2 m n h3 h4. red in h2. destruct h2 as [r h2]. subst. destruct h3 as [h3]. constructor. omega. intro h2. destruct h1 as [n h1]. red. revert h2 h1. revert A. induction n as [|n h1]. intros A h1 h2. exists 0. rewrite nat_translate_0. apply Extensionality_Ensembles; red; split; intros m h3; try constructor. clear h3. apply (h1 0 m h2). omega. intros A h2 h3. specialize (h1 A h2). destruct (classic (In A n)) as [h4 | h5]. apply h1; auto. exists (S n). apply Extensionality_Ensembles. red. split. red. intros m h6. constructor. destruct (lt_eq_lt_dec m (S n)) as [h7 | h8]. destruct h7 as [h7 | h9]. assert (h8:m <= n). omega. specialize (h2 m n h6 h8). contradiction. subst. omega. omega. red. intros m h6. destruct h6 as [h6]. apply (h2 _ _ h3). omega. Qed. (*Wyckoff*) Definition nat_translatable_glb_compat : forall (S:Ensemble nat) (pf:nat_translatable' S), glb_nat S (nat_translatable_endpoint S pf). intros S h1. unfold nat_translatable_endpoint. destruct constructive_definite_description as [b h2]. subst. simpl. red. split. red. intros n h2. destruct h2. auto with arith. intros b' h2. red in h2. apply h2. constructor. auto with arith. Qed. (*Wyckoff*) Lemma step_le_compat : forall A:Ensemble nat, (forall m n:nat, In A m -> m <= n -> In A n) <-> (forall n:nat, In A n -> In A (S n)). intro A. split. intro h1. intros n h2. apply (h1 n (S n) h2). omega. intro h1. intros m n. revert m h1. revert A. induction n as [|n ih]. intros. assert (m=0). omega. subst. assumption. intros A m h1 h2 h3. specialize (ih A). pose proof (le_lt_eq_dec _ _ h3) as h4. destruct h4 as [h4 | h5]. assert (h5:m <= n). omega. clear h4. specialize (ih _ h1 h2 h5). apply h1; auto. subst. assumption. Qed. (*Wyckoff*) Lemma nat_translatable_iff' : forall A:Ensemble nat, Inhabited A -> (nat_translatable A <-> (forall n:nat, In A n -> In A (S n))). intros A h0. rewrite nat_translatable_iff; auto. rewrite step_le_compat. reflexivity. Qed. (*Wyckoff*) Lemma nat_translatable_boundary' : forall (P:nat->Prop), P 0 -> Inhabited [n:nat | ~ P n] -> nat_translatable [n:nat | ~ P n] -> exists ! n:nat, P n /\ ~ P (S n). intros P h0 h1 h2. pose proof h2 as h2'. rewrite nat_translatable_prime_compat in h2. pose proof (nat_translatable_glb_compat _ h2) as h3. rewrite nat_translatable_iff' in h2'. destruct (zerop (nat_translatable_endpoint [n:nat | ~ P n] h2)) as [h4 | h5]. pose proof (nat_translatable_endpoint_compat [n:nat | ~ P n] h2) as h5. rewrite h4 in h5. rewrite nat_translate_0 in h5. pose proof (Full_intro _ 0) as h6. rewrite <- h5 in h6. destruct h6. contradiction. exists (pred (nat_translatable_endpoint [n : nat | ~ P n] h2)). red. split. split. red in h3. destruct h3 as [h3 hl]. red in h3. apply NNPP. intro h6. assert (h7:In [n0:nat | ~ P n0] (pred (nat_translatable_endpoint [n : nat | ~ P n] h2))). constructor. assumption. specialize (h3 _ h7). omega. rewrite <- (S_pred _ _ h5). pose proof (nat_translatable_endpoint_compat [n:nat | ~ P n] h2) as he. assert (h6:In (nat_translate (nat_translatable_endpoint [n : nat | ~ P n] h2)) (nat_translatable_endpoint [n : nat | ~ P n] h2)). constructor. auto with arith. rewrite <- he in h6. destruct h6 as [h6]. assumption. rename h5 into hlt. intros n h5. destruct h5 as [h5 h6]. apply S_inj. rewrite <- (S_pred _ _ hlt). red in h3. destruct h3 as [h3a h3b]. destruct (lt_eq_lt_dec (nat_translatable_endpoint [n0 : nat | ~ P n0] h2) (S n)) as [h7 | h8]. destruct h7 as [h7 | h8]. assert (h8:lb_nat [n0 : nat | ~ P n0] n). red. intros m h8. destruct h8 as [h8]. destruct (lt_eq_lt_dec n m) as [h9 | h10]. destruct h9 as [h9 | h11]. omega. subst. contradiction. assert (h11:In [n : nat | ~ P n] m). constructor; auto. rewrite <- step_le_compat in h2'. assert (h13:m <= n). omega. specialize (h2' m n h11 h13). destruct h2'. contradiction. specialize (h3b n h8). assert (h9: n = nat_translatable_endpoint [n0 : nat | ~ P n0] h2). omega. subst. contradict h5. pose proof (nat_translatable_endpoint_compat [n:nat | ~ P n] h2) as h5. assert (h6':In (nat_translate (nat_translatable_endpoint [n : nat | ~ P n] h2)) (nat_translatable_endpoint [n : nat | ~ P n] h2)). constructor. auto with arith. rewrite <- h5 in h6'. destruct h6'. assumption. assumption. assert (h9:In [n:nat | ~ P n] (S n)). constructor; auto. pose proof (nat_translatable_endpoint_compat [n:nat | ~ P n] h2) as h5'. rewrite h5' in h9. destruct h9 as [h9]. omega. assumption. Qed. End ArithSets. (*Wyckoff*) Definition odd_nat1t_fun_unite_finite_even {T:Type} (p:odd_nat1t->T) (n:even_nat1t) (p':{m:even_nat1t | en1t_le m n}->T) : seg_one_t (nat_of_en1t n) -> T. intro m. destruct m as [m h1]. destruct (even_odd_dec m) as [h2 | h2]. assert (h3:0 < m). destruct h1. auto with sets. pose (gt0_to_1t h3) as m'. assert (h4: even1t m'). red. unfold m'. rewrite match_gt0_to_1t_even_iff. rewrite gt0_to_1t_compat. assumption. pose (even_nat1t_intro _ h4) as m''. pose proof (nat_of_en1t_le_compat m'' n) as h5. assert (h6: nat_of_en1t m'' <= nat_of_en1t n). simpl. unfold m'. rewrite gt0_to_1t_compat. destruct h1; auto. specialize (h5 h6). refine (p' (exist (fun m => en1t_le m n) _ h5)). assert (h3:0 < m). destruct h1; auto with arith. assert (h4:odd1t (gt0_to_1t h3)). assert (h5:~even m). intro h6. apply (not_even_and_odd m h6 h2); auto. rewrite (even_gt0_to_1t_iff m h3) in h5. destruct (even_odd_dec_1t (gt0_to_1t h3)) as [|h6]; try contradiction. assumption. refine (p (odd_nat1t_intro _ h4)). Defined. (*Wyckoff*) Lemma odd_nat1t_fun_unite_finite_even_compat : forall {T:Type} (p:odd_nat1t->T) (n:even_nat1t) (p':{m:even_nat1t | en1t_le m n}->T), let f := odd_nat1t_fun_unite_finite_even p n p' in (forall (o:odd_nat1t) (pf:nat_of_on1t o <= nat_of_en1t n), p o = f (exist (fun m => (1 <= m <= nat_of_en1t n)) _ (conj (le_1_odd1t o) pf))) /\ (forall (e:even_nat1t) (pf:nat_of_en1t e <= nat_of_en1t n), p' (exist (fun m => en1t_le m n) _ (nat_of_en1t_le_compat _ _ pf)) = f (exist (fun m => (1 <= m <= nat_of_en1t n)) _ (conj (en1t_ge_1 e) pf))). intros T p n p' f. split. intros o h1. unfold f, odd_nat1t_fun_unite_finite_even. destruct even_odd_dec as [h2 | h2]. destruct o as [o' h3]. simpl in h2. red in h3. destruct o' as [o'' o']. simpl in h2. red in h3. pose proof (not_even_and_odd _ h2 h3). contradiction. f_equal. destruct even_odd_dec_1t as [h3 | h3]. red in h3. pose proof (iff1 (match_gt0_to_1t_even_iff _ (le_1_odd1t o)) h3) as h3'. rewrite gt0_to_1t_compat in h3'. destruct o as [o' o]. destruct o' as [o'' o']. simpl in h3', h2. pose proof (not_even_and_odd _ h3' h2). contradiction. destruct o as [o' o]. destruct o' as [o'' o']. apply odd_nat1t_functional. apply nat1t_functional; simpl. unfold gt0_to_1t. destruct constructive_definite_description as [a h5]; simpl. destruct a; simpl in h5; auto. intros e h1. unfold f, odd_nat1t_fun_unite_finite_even. destruct even_odd_dec as [h2 | h2]. f_equal. apply proj1_sig_injective; simpl. destruct e as [e' e]. destruct e' as [e'' e']. apply even_nat1t_functional. apply nat1t_functional. unfold gt0_to_1t. destruct Description.constructive_definite_description as [n' h3]; simpl. destruct n'. simpl in h3; auto. destruct e as [e' h3]. destruct e' as [e'' e']. simpl. simpl in h2. red in h3. red in h3. pose proof (not_even_and_odd _ h3 h2). contradiction. Qed. (*Wyckoff*) Definition odd_nat1t_fun_unite_finite_even' {T:Type} (p:odd_nat1t->T) (n:odd_nat1t) (p':{m:even_nat1t | en1t_le m (on1t_pred' n)}->T) : seg_one_t (nat_of_on1t n) -> T. intro m. destruct m as [m h1]. destruct (even_odd_dec m) as [h2 | h2]. assert (h3:0 < m). destruct h1. auto with sets. pose (gt0_to_1t h3) as m'. assert (h4: even1t m'). red. unfold m'. rewrite match_gt0_to_1t_even_iff. rewrite gt0_to_1t_compat. assumption. pose (even_nat1t_intro _ h4) as m''. pose proof (nat_of_en1t_le_pred_compat m'' n) as h5. assert (h6: nat_of_en1t m'' <= nat_of_on1t n ). unfold m''. simpl. unfold m'. rewrite gt0_to_1t_compat. destruct h1; auto. specialize (h5 h6). refine (p' (exist (fun m => en1t_le m (on1t_pred' n)) _ h5)). refine (p _). refine (odd_to_on1t _). apply h2. Defined. (*Wyckoff*) Lemma odd_nat1t_fun_unite_finite_even_compat' : forall {T:Type} (p:odd_nat1t->T) (n:odd_nat1t) (p':{m:even_nat1t | en1t_le m (on1t_pred' n)}->T), let f := odd_nat1t_fun_unite_finite_even' p n p' in (forall (o:odd_nat1t) (pf:nat_of_on1t o <= nat_of_on1t n), p o = f (exist (fun m => (1 <= m <= nat_of_on1t n)) _ (conj (le_1_odd1t o) pf))) /\ (forall (e:even_nat1t) (pf:nat_of_en1t e <= nat_of_on1t n), p' (exist (fun m => en1t_le m (on1t_pred' n)) _ (nat_of_en1t_le_pred_compat _ _ pf)) = f (exist (fun m => (1 <= m <= nat_of_on1t n)) _ (conj (en1t_ge_1 e) pf))). intros T p n p' f. split. intros o h1. unfold f, odd_nat1t_fun_unite_finite_even'. destruct even_odd_dec as [h2 | h2]. destruct o as [o' h3]. simpl in h2. red in h3. destruct o' as [o'' o']. simpl in h2. red in h3. pose proof (not_even_and_odd _ h2 h3). contradiction. f_equal. destruct o as [o' o]. destruct o' as [o'' o']. apply odd_nat1t_functional. unfold odd_to_on1t, odd_to_1t. simpl. apply nat_of_1t_inj. simpl. rewrite gt0_to_1t_compat. reflexivity. intros e h1. unfold f, odd_nat1t_fun_unite_finite_even'. destruct even_odd_dec as [h2 | h2]. f_equal. apply proj1_sig_injective; simpl. destruct e as [e' e]. destruct e' as [e'' e']. apply even_nat1t_functional. apply nat1t_functional. unfold gt0_to_1t. destruct constructive_definite_description as [n' h3]. simpl. destruct n'; simpl in h3; auto. destruct e as [e' h3]. destruct e' as [e'' e']. simpl; simpl in h2. do 2 red in h3. pose proof (not_even_and_odd _ h3 h2). contradiction. Qed. (*Wyckoff*) Definition even_nat1t_fun_unite_finite_odd {T:Type} (p:even_nat1t->T) (n:odd_nat1t) (p':{m:odd_nat1t | on1t_le m n}->T) : seg_one_t (nat_of_on1t n) -> T. intro m. destruct m as [m h1]. destruct (even_odd_dec m) as [h2 | h2]. assert (h3:0 < m). destruct h1. auto with sets. assert (h4:even1t (gt0_to_1t h3)). red. rewrite match_gt0_to_1t_even_iff. rewrite gt0_to_1t_compat. assumption. refine (p (even_nat1t_intro _ h4)). pose (odd_to_1t h2) as m'. pose proof (odd_to_1t_compat h2) as h4. pose (odd_nat1t_intro _ h4) as m''. assert (h6:nat_of_on1t m'' <= nat_of_on1t n). unfold m''. simpl. unfold odd_to_1t. rewrite gt0_to_1t_compat. destruct h1; auto. pose proof (nat_of_on1t_le_compat m'' n) as h5. specialize (h5 h6). refine (p' (exist (fun m => on1t_le m n) _ h5)). Defined. (*Wyckoff*) Lemma even_nat1t_fun_unite_finite_odd_compat : forall {T:Type} (p:even_nat1t->T) (n:odd_nat1t) (p':{m:odd_nat1t | on1t_le m n}->T), let f := even_nat1t_fun_unite_finite_odd p n p' in (forall (e:even_nat1t) (pf:nat_of_en1t e <= nat_of_on1t n), p e = f (exist (fun m => (1 <= m <= nat_of_on1t n)) _ (conj (en1t_ge_1 e) pf))) /\ (forall (o:odd_nat1t) (pf:nat_of_on1t o <= nat_of_on1t n), p' (exist (fun m => on1t_le m n) _ (nat_of_on1t_le_compat _ _ pf)) = f (exist (fun m => (1 <= m <= nat_of_on1t n)) _ (conj (le_1_odd1t o) pf))). intros T p n p' f. split. intros e h1. unfold f, even_nat1t_fun_unite_finite_odd. destruct even_odd_dec as [h2 | h2]. f_equal. destruct e as [e' e]. destruct e' as [e'' e']. apply even_nat1t_functional. pose proof (nat1t_functional (nat1t_intro e') (gt0_to_1t (en1t_ge_1 (even_nat1t_intro (nat1t_intro e') e)))) as h4. simpl in h4. hfirst h4. unfold gt0_to_1t. destruct Description.constructive_definite_description as [a h5]. simpl. destruct a. simpl in h5. subst. reflexivity. specialize (h4 hf). assumption. destruct e as [e' h3]. simpl in h2. red in h3. destruct e' as [e'' e']. simpl in h2. red in h3. pose proof (not_even_and_odd _ h3 h2). contradiction. intros o h1. unfold f, even_nat1t_fun_unite_finite_odd. destruct even_odd_dec as [h2 | h2]. destruct o as [o' h3]. destruct o' as [o'' o']. simpl. simpl in h2. red in h3. red in h3. pose proof (not_even_and_odd _ h2 h3). contradiction. f_equal. apply proj1_sig_injective; simpl. destruct o as [o' o]. destruct o' as [o'' o']. apply odd_nat1t_functional. pose proof (nat1t_functional (nat1t_intro o') (odd_to_1t h2)) as h3. apply h3. clear h3. unfold odd_to_1t, gt0_to_1t. destruct constructive_definite_description as [n' h3]. simpl. destruct n'. simpl in h3. subst. reflexivity. Qed. (*Wyckoff*) Definition even_nat1t_fun_unite_finite_odd' {T:Type} (p:even_nat1t->T) (n:even_nat1t) (p':{m:odd_nat1t | on1t_le m (en1t_pred n)}->T) : seg_one_t (nat_of_en1t n) -> T. intro m. destruct m as [m h1]. destruct (even_odd_dec m) as [h2 | h2]. assert (h3:0 < m). destruct h1. auto with sets. assert (h4:even1t (gt0_to_1t h3)). red. rewrite match_gt0_to_1t_even_iff. rewrite gt0_to_1t_compat. assumption. refine (p (even_nat1t_intro _ h4)). pose (odd_to_1t h2) as m'. pose proof (odd_to_1t_compat h2) as h4. pose (odd_nat1t_intro _ h4) as m''. pose proof (nat_of_on1t_le_pred_compat m'' n) as h5. assert (h6:nat_of_on1t m'' <= nat_of_en1t n). unfold m''. simpl. unfold odd_to_1t. rewrite gt0_to_1t_compat. destruct h1; auto. specialize (h5 h6). refine (p' (exist (fun m => on1t_le m (en1t_pred n)) _ h5)). Defined. (*Wyckoff*) Lemma even_nat1t_fun_unite_finite_odd_compat' : forall {T:Type} (p:even_nat1t->T) (n:even_nat1t) (p':{m:odd_nat1t | on1t_le m (en1t_pred n)}->T), let f := even_nat1t_fun_unite_finite_odd' p n p' in (forall (e:even_nat1t) (pf:nat_of_en1t e <= nat_of_en1t n), p e = f (exist (fun m => (1 <= m <= nat_of_en1t n)) _ (conj (en1t_ge_1 e) pf))) /\ (forall (o:odd_nat1t) (pf:nat_of_on1t o <= nat_of_en1t n), p' (exist (fun m => on1t_le m (en1t_pred n)) _ (nat_of_on1t_le_pred_compat _ _ pf)) = f (exist (fun m => (1 <= m <= nat_of_en1t n)) _ (conj (le_1_odd1t o) pf))). intros T p n p' f. split. intros e h1. unfold f, even_nat1t_fun_unite_finite_odd'. destruct even_odd_dec as [h2 | h2]. f_equal. destruct e as [e' e]. destruct e' as [e'' e']. apply even_nat1t_functional. apply nat1t_functional. pose proof (nat1t_functional (nat1t_intro e') (gt0_to_1t (en1t_ge_1 (even_nat1t_intro (nat1t_intro e') e)))) as h4. hfirst h4. unfold gt0_to_1t. destruct Description.constructive_definite_description as [a h5]. simpl. destruct a. simpl in h5. subst. reflexivity. specialize (h4 hf). assumption. destruct e as [e' h3]. simpl in h2. red in h3. destruct e' as [e'' e']. simpl in h2. red in h3. pose proof (not_even_and_odd _ h3 h2). contradiction. intros o h1. unfold f, even_nat1t_fun_unite_finite_odd'. destruct even_odd_dec as [h2 | h2]. destruct o as [o' h3]. destruct o' as [o'' o']. simpl. simpl in h2. red in h3. red in h3. pose proof (not_even_and_odd _ h2 h3). contradiction. f_equal. apply proj1_sig_injective; simpl. destruct o as [o' o]. destruct o' as [o'' o']. apply odd_nat1t_functional. pose proof (nat1t_functional (nat1t_intro o') (odd_to_1t h2)) as h3. apply h3. clear h3. unfold odd_to_1t, gt0_to_1t. destruct constructive_definite_description as [n' h3]. simpl. destruct n'. simpl in h3. subst. reflexivity. Qed. bool2-v0-3/Subalgebras.v0000644000175000017500000343653313575574055016004 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-17, Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import SetUtilities. Require Import LogicUtilities. Require Export BoolAlgBasics. Require Import FieldsOfSets. Require Import InfiniteOperations. Require Import TypeUtilities. Require Import Arith. Require Import FunctionProperties. Require Import Equality. Require Import FunctionalExtensionality. Require Import Description. Require Import FiniteMaps. Require Import DecidableDec. Require Import Equality. Require Import FiniteOperations. Require Import FiniteMaps2. Require Import NPeano. Require Import SetUtilities2. Require Import SetUtilities3. Require Import ArithUtilities. Require Import ClassicalChoice. Require Import ListUtilities2. Require Import EqDec. Require Import SetUtilities2_5. Require Import SetUtilities1_5. Open Scope ba_scope. (*See Section ParametricAnalogues for Bool_Alg_p subalgebras (i.e. with an underlying type, which makes a posteriori extensions possible, and makes arbitrarily infinite families of subalgebras possible.*) Section AlgClosed. Variable B:Bool_Alg. Let Bt := bt B. Variable A:(Ensemble Bt). Definition SubBtype := {a:Bt |In A a}. Definition Bplus_sub (x y:SubBtype) := (proj1_sig x) + (proj1_sig y). Definition Btimes_sub (x y:SubBtype) := (proj1_sig x) * (proj1_sig y). Definition Bcomp_sub (x:SubBtype) := -(proj1_sig x). Notation "x +|' y" := (Bplus_sub x y) (at level 50, left associativity). Notation "x *|' y" := (Btimes_sub x y) (at level 40, left associativity). Notation "-|' x" := (Bcomp_sub x) (at level 30). Definition Plus_closed_sub : Prop := (forall (x y:SubBtype), In A (x +|' y)). Definition Times_closed_sub : Prop := (forall (x y:SubBtype), In A (x *|' y)). Definition One_closed_sub : Prop := In A 1. Definition Zero_closed_sub : Prop := In A 0. Definition Comp_closed_sub : Prop := (forall x:SubBtype, In A (-|' x)). Lemma Plus_closed'' : forall (x y:Bt), Plus_closed_sub -> In A x -> In A y -> In A (x+y). intros x y h1 h2 h3. unfold Plus_closed_sub in h1. apply h1 with (x := exist (In A) x h2) (y := exist (In A) y h3). Qed. Lemma Times_closed'' : forall (x y:Bt), Times_closed_sub -> In A x -> In A y -> In A (x*y). intros x y h1 h2 h3. unfold Times_closed_sub in h1. apply h1 with (x := exist (In A) x h2) (y := exist (In A) y h3). Qed. Lemma Comp_closed'' : forall (b:Bt), Comp_closed_sub -> (In A b) -> In A (-b). intros b h1 h2. unfold Comp_closed_sub in h1. apply h1 with (x := exist (In A) b h2). Qed. Record alg_closed : Prop := { P_c : Plus_closed_sub; T_c : Times_closed_sub; O_c : One_closed_sub; Z_c : Zero_closed_sub; C_c : Comp_closed_sub}. Variable Ac : alg_closed. Definition Bc' := Build_Bconst SubBtype (Full_set SubBtype) (fun (x y:SubBtype) => exist (In A) (x +|' y) (P_c Ac x y)) (fun (x y:SubBtype) => exist (In A) (x *|' y) (T_c Ac x y)) (exist (In A ) 1 (O_c Ac)) (exist (In A ) 0 (Z_c Ac)) (fun (x:SubBtype) => (exist (In A) (-|' x) (C_c Ac x))). Section Bc''. Infix "+|" := (Bplus Bc') (at level 50, left associativity). Infix "*|" := (Btimes Bc') (at level 40, left associativity). Notation "0|" := (Bzero Bc'). Notation "1|" := (Bone Bc'). Notation "-| x" := ((Bcomp Bc') x) (at level 30). (*These are just all the basic axioms applied to the subset type SubBtype, derived from the main boolean algebra B *) Lemma assoc_sum' : forall n m p:SubBtype, n +| (m +| p) = n +| m +| p. intros n m p. red in n. red in m. red in p. inversion n as [x]. inversion m as [y]. inversion p as [z]. simpl. unfold Bplus_sub. simpl. apply proj1_sig_injective. simpl. apply assoc_sum. Qed. Lemma assoc_prod' : forall n m p:SubBtype, n *| (m *| p) = n *| m *| p. intros n m p. red in n. red in m. red in p. inversion n as [x]. inversion m as [y]. inversion p as [z]. simpl. unfold Btimes_sub. simpl. apply proj1_sig_injective. simpl. apply assoc_prod. Qed. Lemma comm_sum' : forall n m:SubBtype, n +| m = m +| n. intros n m. red in n. red in m. inversion n. inversion m. simpl. unfold Bplus_sub. simpl. apply proj1_sig_injective. simpl. apply comm_sum. Qed. Lemma comm_prod' : forall n m:SubBtype, n *| m = m *| n. intros n m. red in n. red in m. inversion n. inversion m. simpl. unfold Btimes_sub. simpl. apply proj1_sig_injective. simpl. apply comm_prod. Qed. Lemma abs_sum' : forall n m:SubBtype, n +| (n *| m) = n. intros n m. red in n. red in m. inversion n. inversion m. simpl. unfold Btimes_sub. unfold Bplus_sub. simpl. apply proj1_sig_injective. simpl. apply abs_sum. Qed. Lemma abs_prod': forall n m:SubBtype, n *| (n +| m) = n. intros n m. red in n. red in m. inversion n. inversion m. simpl. unfold Btimes_sub. unfold Bplus_sub. simpl. apply proj1_sig_injective. simpl. apply abs_prod. Qed. Lemma dist_sum' : forall n m p:SubBtype, p *| (n +| m) = p *| n + p *| m. intros n m p. red in n. red in m. red in p. inversion n. inversion m. inversion p. simpl. unfold Btimes_sub. unfold Bplus_sub. simpl. apply proj1_sig_injective. simpl. apply dist_sum. Qed. Lemma dist_prod': forall n m p:SubBtype, p +| (n *| m) = (p +| n) *| (p +| m). intros n m p. red in n. red in m. red in p. inversion n. inversion m. inversion p. simpl. unfold Btimes_sub. unfold Bplus_sub. simpl. apply proj1_sig_injective. simpl. apply dist_prod. Qed. Lemma comp_sum': forall n:SubBtype, n +| (-| n) = 1. intros n. red in n. inversion n. simpl. unfold Bplus_sub. unfold Bcomp_sub. simpl. apply proj1_sig_injective. simpl. apply comp_sum. Qed. Lemma comp_prod': forall n:SubBtype, n *| (-| n) = 0. intros n. red in n. inversion n. simpl. unfold Bplus_sub. unfold Bcomp_sub. simpl. apply proj1_sig_injective. simpl. apply comp_prod. Qed. Lemma und_set' : (BS Bc') = Full_set (Btype Bc'). unfold Bc'. simpl. reflexivity. Qed. Definition Subalg := Build_Bool_Alg Bc' und_set' assoc_sum' assoc_prod' comm_sum' comm_prod' abs_sum' abs_prod' dist_sum' dist_prod' comp_sum' comp_prod'. (*Being closed under 2 operations is sufficient to ensure that a subset is a subalgebra *) Lemma two_ops_imp_times : Comp_closed_sub -> Plus_closed_sub -> Times_closed_sub. intros h1 h2. unfold Times_closed_sub. intros x y. unfold SubBtype in x. unfold SubBtype in y. assert (h3: proj1_sig x * proj1_sig y = -(-(proj1_sig x) + -(proj1_sig y))). apply times_eq_plus. unfold Btimes_sub. rewrite h3. apply Comp_closed'' with (b:= (-proj1_sig x + -proj1_sig y)). apply h1. apply Plus_closed'' with (x := -proj1_sig x) (y := -proj1_sig y). apply h2. apply Comp_closed''. apply h1. apply (proj2_sig x). apply Comp_closed''. apply h1. apply (proj2_sig y). Qed. Lemma two_ops_imp_plus : Comp_closed_sub -> Times_closed_sub -> Plus_closed_sub. intros h1 h2. unfold Plus_closed_sub. intros x y. unfold SubBtype in x. unfold SubBtype in y. assert (h3: proj1_sig x + proj1_sig y = -(-(proj1_sig x) * -(proj1_sig y))). apply plus_eq_times. unfold Bplus_sub. rewrite h3. apply Comp_closed'' with (b:= (-proj1_sig x * -proj1_sig y)). apply h1. apply Times_closed'' with (x := -proj1_sig x) (y := -proj1_sig y). apply h2. apply Comp_closed''. apply h1. apply (proj2_sig x). apply Comp_closed''. apply h1. apply (proj2_sig y). Qed. End Bc''. End AlgClosed. Arguments alg_closed [B] _. Arguments Subalg [B] _ _. Arguments Plus_closed_sub [B] _. Arguments Times_closed_sub [B] _. Arguments Comp_closed_sub [B] _. Arguments Zero_closed_sub [B] _. Arguments One_closed_sub [B] _. (*More [alg_closed] theorems (without section variables).*) Lemma alg_closed_ba_ens : forall B:Bool_Alg, alg_closed (ba_ens B). intro B. constructor; red; intros; constructor. Qed. Lemma alg_closed_subst_iff : forall {B:Bool_Alg} (E F:Ensemble (bt B)), E = F -> (alg_closed E <-> alg_closed F). intros B E F h1. subst. tauto. Qed. Lemma alg_closed_improper_subalg : forall {B:Bool_Alg} (S:Ensemble (bt B)) (S':Ensemble (bt (Subalg (ba_ens B) (alg_closed_ba_ens B)))), S = im_proj1_sig S' -> (alg_closed S <-> alg_closed S'). intros B S S' h1. subst. split. intro h1. destruct h1 as [h1 h2 h3 h4 h5]. red in h1, h2, h3, h4, h5. constructor; red. intros x y. unfold SubBtype in h1. destruct x as [x h6], y as [y h7]. destruct x as [x h8], y as [y h9]. match goal with |- @In ?T ?S ?x => pose x as a end. assert (h10:In (im_proj1_sig S') x). apply Im_intro with (exist (fun a : Btype (Bc B) => In (ba_ens B) a) x h8). assumption. simpl. reflexivity. assert (h11:In (im_proj1_sig S') y). apply Im_intro with (exist (fun a : Btype (Bc B) => In (ba_ens B) a) y h9). assumption. simpl. reflexivity. specialize (h1 (exist _ _ h10) (exist _ _ h11)). inversion h1 as [z h12 ? h13]. subst. clear h1. assert (h14:z = a). unfold a. destruct z as [z h14]. apply proj1_sig_injective. simpl. simpl in h13. subst. f_equal. unfold Bplus_sub. simpl. reflexivity. fold a. rewrite <- h14 at 1; auto. intros x y. unfold SubBtype in h1. destruct x as [x h6], y as [y h7]. destruct x as [x h8], y as [y h9]. assert (h10:In (im_proj1_sig S') x). apply Im_intro with (exist (fun a : Btype (Bc B) => In (ba_ens B) a) x h8). assumption. simpl. reflexivity. assert (h11:In (im_proj1_sig S') y). apply Im_intro with (exist (fun a : Btype (Bc B) => In (ba_ens B) a) y h9). assumption. simpl. reflexivity. specialize (h2 (exist _ _ h10) (exist _ _ h11)). inversion h2 as [z h12 ? h13]. subst. clear h2. match goal with |-In ?S ?x => pose x as a end. assert (h14:z = a). unfold a. destruct z as [z h14]. apply proj1_sig_injective. simpl. simpl in h13. subst. f_equal. unfold Bplus_sub. simpl. reflexivity. fold a. rewrite <- h14 at 1; auto. inversion h3 as [a h6 ? h7]. subst. assert (h8:a = (Bone (Bc (@Subalg B (ba_ens B) (alg_closed_ba_ens B)))) ). apply proj1_sig_injective. rewrite h7 at 1. reflexivity. subst. assumption. inversion h4 as [a h6 ? h7]. subst. assert (h8:a = (Bzero (Bc (@Subalg B (ba_ens B) (alg_closed_ba_ens B)))) ). apply proj1_sig_injective. rewrite h7 at 1. reflexivity. subst. assumption. intro x. unfold SubBtype in h1. destruct x as [x h6]. destruct x as [x h8]. assert (h10:In (im_proj1_sig S') x). apply Im_intro with (exist (fun a : Btype (Bc B) => In (ba_ens B) a) x h8). assumption. simpl. reflexivity. specialize (h5 (exist _ _ h10)). inversion h5 as [z h12 ? h13]. subst. clear h5. match goal with |- In ?S ?x => pose x as a end. assert (h14:z = a). unfold a. destruct z as [z h14]. apply proj1_sig_injective. simpl. simpl in h13. subst. f_equal. unfold Bplus_sub. simpl. reflexivity. fold a. rewrite <- h14 at 1; auto. intro h1. destruct h1 as [h1 h2 h3 h4 h5]. red in h1, h2, h3, h4, h5. constructor; red. intros x y. destruct x as [x h6], y as [y h7]. destruct h6 as [x h6], h7 as [y h7]. subst. destruct x as [x h8], y as [y h9]. subst. simpl. unfold SubBtype in h1. simpl in h1. unfold SubBtype in h1. simpl in h1. specialize (h1 (exist _ _ h6) (exist _ _ h7)). let P := type of h1 in match P with In ?S ?x => pose x as a end. apply Im_intro with a; auto. intros x y. destruct x as [x h6], y as [y h7]. destruct h6 as [x h6], h7 as [y h7]. subst. destruct x as [x h8], y as [y h9]. subst. simpl. unfold SubBtype in h1. simpl in h1. unfold SubBtype in h1. simpl in h1. specialize (h2 (exist _ _ h6) (exist _ _ h7)). let P := type of h2 in match P with In ?S ?x => pose x as a end. apply Im_intro with a; auto. eapply Im_intro. apply h3. simpl. reflexivity. eapply Im_intro. apply h4. simpl. reflexivity. intros x. destruct x as [x h6]. destruct h6 as [x h6]. subst. destruct x as [x h8]. subst. simpl. unfold SubBtype in h5. simpl in h5. unfold SubBtype in h5. simpl in h5. specialize (h5 (exist _ _ h6)). let P := type of h5 in match P with In ?S ?x => pose x as a end. apply Im_intro with a; auto. Qed. Lemma two_ops_plus_closed : forall {B:Bool_Alg} (A:Ensemble (bt B)), Inhabited A -> Comp_closed_sub A -> Plus_closed_sub A -> alg_closed A. intros B A h0 h1 h2. assert (h3:One_closed_sub A). destruct h0 as [x h0]. red in h1. specialize (h1 (exist _ _ h0)). red in h2. specialize (h2 (exist _ _ h0) (exist _ _ h1)). unfold Bplus_sub in h2. simpl in h2. unfold Bcomp_sub in h2. simpl in h2. rewrite comp_sum in h2. red. assumption. constructor; auto. apply two_ops_imp_times; auto. red in h3. red in h1. specialize (h1 (exist _ _ h3)). unfold Bcomp_sub in h1. simpl in h1. rewrite one_comp in h1. red. assumption. Qed. Lemma two_ops_times_closed : forall {B:Bool_Alg} (A:Ensemble (bt B)), Inhabited A -> Comp_closed_sub A -> Times_closed_sub A -> alg_closed A. intros B A h0 h1 h2. assert (h3:Zero_closed_sub A). destruct h0 as [x h0]. red in h1. specialize (h1 (exist _ _ h0)). red in h2. specialize (h2 (exist _ _ h0) (exist _ _ h1)). unfold Btimes_sub in h2. simpl in h2. unfold Bcomp_sub in h2. simpl in h2. rewrite comp_prod in h2. red. assumption. constructor; auto. apply two_ops_imp_plus; auto. red in h3. red in h1. specialize (h1 (exist _ _ h3)). unfold Bcomp_sub in h1. simpl in h1. rewrite zero_comp in h1. red. assumption. Qed. Lemma subalg_functional : forall {B:Bool_Alg} (A A':Ensemble (Btype (Bc B))), A = A' -> forall (pfa:alg_closed A) (pfa':alg_closed A'), Subalg A pfa = Subalg A' pfa'. intros B A A' h1. rewrite h1. intros h2 h3. pose proof (proof_irrelevance _ h2 h3); subst; auto. Qed. Section Closed. Variable B:Bool_Alg. Let Bt := Btype (Bc B). Lemma plus_closed : forall (S:Ensemble Bt), alg_closed S -> forall (x y:Bt), In S x -> In S y -> In S (x + y). intros S h1 x y h2 h3. pose proof (P_c _ _ h1) as h4. red in h4. specialize (h4 (exist _ _ h2) (exist _ _ h3)). unfold Bplus_sub in h4. simpl in h4. assumption. Qed. Lemma times_closed : forall (S:Ensemble Bt), alg_closed S -> forall (x y:Bt), In S x -> In S y -> In S (x * y). intros S h1 x y h2 h3. pose proof (T_c _ _ h1) as h4. red in h4. specialize (h4 (exist _ _ h2) (exist _ _ h3)). unfold Btimes_sub in h4. simpl in h4. assumption. Qed. Lemma comp_closed : forall (S:Ensemble Bt), alg_closed S -> forall x:Bt, In S x -> In S (- x). intros S h1 x h2. pose proof (C_c _ _ h1) as h4. red in h4. specialize (h4 (exist _ _ h2)). unfold Bcomp_sub in h4. simpl in h4. assumption. Qed. Lemma zero_closed : forall (S:Ensemble Bt), alg_closed S -> In S 0. intros S h1. pose proof (Z_c _ _ h1) as h4. red in h4. assumption. Qed. Lemma one_closed : forall (S:Ensemble Bt), alg_closed S -> In S 1. intros S h1. pose proof (O_c _ _ h1) as h4. red in h4. assumption. Qed. Lemma closed_intersection : forall (E F:Ensemble Bt), alg_closed E -> alg_closed F -> alg_closed (Intersection E F). intros E F h1 h2. destruct h1 as [h1a h1b h1c h1d h1e]. destruct h2 as [h2a h2b h2c h2d h2e]. red in h1a, h1b, h1c, h1d, h1e, h2a, h2b, h2c, h2d, h2e. constructor. red. intros x y. destruct x as [x h3], y as [y h4]. destruct h3 as [x h3a h3b], h4 as [y h4a h4b]. specialize (h1a (exist _ _ h3a) (exist _ _ h4a)). specialize (h2a (exist _ _ h3b) (exist _ _ h4b)). unfold Bplus_sub. simpl. unfold Bplus_sub in h1a, h2a. simpl in h1a, h2a. constructor; auto. red. intros x y. destruct x as [x h3], y as [y h4]. destruct h3 as [x h3a h3b], h4 as [y h4a h4b]. specialize (h1b (exist _ _ h3a) (exist _ _ h4a)). specialize (h2b (exist _ _ h3b) (exist _ _ h4b)). unfold Btimes_sub. simpl. unfold Btimes_sub in h1a, h2a. simpl in h1a, h2a. constructor; auto. red. constructor; auto. red. constructor; auto. red. intro x. destruct x as [x h3]. destruct h3 as [x h3 h4]. specialize (h1e (exist _ _ h3)). specialize (h2e (exist _ _ h4)). unfold Bcomp_sub. simpl. unfold Bcomp_sub in h1e, h2e. simpl in h1e, h2e. constructor; auto. Qed. End Closed. Lemma subalg_preserves_atom : forall {B:Bool_Alg} (A:Ensemble (bt B)) (pfc:alg_closed A) (x:bt B) (pfi:In A x), atom x -> atom (exist _ _ pfi) (B:=Subalg _ pfc). intros B A h1 x h2. intro h3. red in h3. red. destruct h3 as [h3 h4]. split. intro h5. subst. contradict h3. simpl in h5. pose proof (f_equal (@proj1_sig _ _) h5) as h5'. clear h5. simpl in h5'. assumption. intros b h5. destruct b as [b h6]. assert (h7:le b x). red in h5. red. simpl in h5. apply exist_injective in h5. unfold Bplus_sub in h5. simpl in h5. assumption. specialize (h4 _ h7). destruct h4 as [h4a | h4b]; subst. left. apply proj1_sig_injective. simpl. reflexivity. right. apply proj1_sig_injective. simpl. reflexivity. Qed. Section Families_Ind_Families_Subalgebras. Variable B:Bool_Alg. Let Bt := Btype (Bc B). Variable It : Type. Variable A:Ensemble Bt. (*Subalgebra -- Indexed representation*) Record SubalgI := { sai_A:It->(Ensemble Bt); sai_closed:forall (i:It), alg_closed (sai_A i); sai_B := fun (i:It) => Subalg (sai_A i) (sai_closed i) }. (*Subalgebra -- Family representation*) Record SubalgF := { saf_A:Family Bt; saf_closed : forall (S:Ensemble Bt), In saf_A S -> alg_closed S; saf_B := fun (S:{S':Ensemble Bt | In saf_A S'}) => Subalg (proj1_sig S) (saf_closed (proj1_sig S) (proj2_sig S))}. Variable SI:SubalgI. Variable SF:SubalgF. Definition int_sai_A := IndexedIntersection (sai_A SI). Definition int_saf_A := FamilyIntersection (saf_A SF). Lemma Plus_closed_i : Plus_closed_sub int_sai_A. red. intros x y. assert (h1: In int_sai_A (proj1_sig x)). apply proj2_sig. assert (h2: In int_sai_A (proj1_sig y)). apply proj2_sig. inversion h1. inversion h2. assert (h3: forall (i:It), Plus_closed_sub (sai_A SI i)). intro i. apply (P_c _ _ (sai_closed SI i)). apply indexed_intersection_intro. intro i. apply Plus_closed''. apply (h3 i). apply (H i). apply (H1 i). Qed. Lemma Plus_closed_F : Plus_closed_sub int_saf_A. red. intros x y. constructor. intros S h1. apply Plus_closed''. apply (P_c _ S (saf_closed SF S h1)). destruct x as [? h2]. simpl. inversion h2 as [? h3]. apply (h3 S h1). destruct y as [? h4]. simpl. inversion h4 as [? h5]. apply (h5 S h1). Qed. Lemma Times_closed_i : Times_closed_sub int_sai_A. red. intros x y. assert (h1: In int_sai_A (proj1_sig x)). apply proj2_sig. assert (h2: In int_sai_A (proj1_sig y)). apply proj2_sig. inversion h1. inversion h2. assert (h3: forall (i:It), Times_closed_sub (sai_A SI i)). intro i. apply (T_c _ _ (sai_closed SI i)). apply indexed_intersection_intro. intro i. apply Times_closed''. apply (h3 i). apply (H i). apply (H1 i). Qed. Lemma Times_closed_F : Times_closed_sub int_saf_A. red. intros x y. constructor. intros S h1. apply Times_closed''. apply (T_c _ S (saf_closed SF S h1)). destruct x as [? h2]. simpl. inversion h2 as [? h3]. apply (h3 S h1). destruct y as [? h4]. simpl. inversion h4 as [? h5]. apply (h5 S h1). Qed. Lemma Comp_closed_i : Comp_closed_sub int_sai_A. red. intro x. assert (h1: In int_sai_A (proj1_sig x)). apply proj2_sig. assert (h2: forall (i:It), Comp_closed_sub (sai_A SI i)). intro i. apply (C_c _ _ (sai_closed SI i)). apply indexed_intersection_intro. intro i. apply Comp_closed''. apply (h2 i). inversion h1. apply (H i). Qed. Lemma Comp_closed_F : Comp_closed_sub int_saf_A. red. intros x. constructor. intros S h1. apply Comp_closed''. apply (C_c _ S (saf_closed SF S h1)). destruct x as [? h2]. simpl. inversion h2 as [? h3]. apply (h3 S h1). Qed. Lemma One_closed_i : One_closed_sub int_sai_A. red. apply indexed_intersection_intro. intro i. pose proof (O_c _ _ (sai_closed SI i)) as h1. unfold One_closed_sub in h1. assumption. Qed. Lemma One_closed_F : One_closed_sub int_saf_A. red. constructor. intros S h1. apply (O_c _ S (saf_closed SF S h1)). Qed. Lemma Zero_closed_i : Zero_closed_sub int_sai_A. red. apply indexed_intersection_intro. intro i. pose proof (Z_c _ _ (sai_closed SI i)) as h1. unfold Zero_closed_sub in h1. assumption. Qed. Lemma Zero_closed_F : Zero_closed_sub int_saf_A. red. constructor. intros S h1. apply (Z_c _ S (saf_closed SF S h1)). Qed. Lemma int_sai_alg_closed : alg_closed int_sai_A. constructor. apply Plus_closed_i. apply Times_closed_i. apply One_closed_i. apply Zero_closed_i. apply Comp_closed_i. Qed. Lemma int_saf_alg_closed : alg_closed int_saf_A. constructor. apply Plus_closed_F. apply Times_closed_F. apply One_closed_F. apply Zero_closed_F. apply Comp_closed_F. Qed. Definition int_sub_algi := Subalg _ int_sai_alg_closed. Definition int_sub_algF := Subalg _ int_saf_alg_closed. (*This defines the family of all subalgebra sets that contain A, as a set*) Definition S_clo_cont_A := [S:Ensemble Bt | Included A S /\ alg_closed S]. (*Ditto as a subalgebra family *) Definition SA_cont_A : SubalgF. pose S_clo_cont_A as F. red in F. assert (h1: forall (S:Ensemble Bt), In F S -> alg_closed S). intros S h2. inversion h2 as [h3]. destruct h3 as [h3l h3r]. assumption. refine (Build_SubalgF F h1). Defined. Lemma saf_A_eq : saf_A SA_cont_A = S_clo_cont_A. apply Extensionality_Ensembles. red. split. (* <= *) red. intros S h1. inversion h1. constructor; assumption. (* >= *) red. intros S h1. unfold S_clo_cont_A in h1. inversion h1. simpl. constructor; assumption. Qed. End Families_Ind_Families_Subalgebras. Arguments SA_cont_A [B] _. Arguments int_saf_A [B] _ _. Arguments int_saf_alg_closed [B] _. Arguments S_clo_cont_A [B] _ _. Arguments saf_A [B] _ _. Arguments saf_A_eq [B] _. Arguments saf_closed [B] _ _ _. Section GenAlg. Variable B:Bool_Alg. Let Bt := bt B. Definition Gen_Ens (A:Ensemble Bt) : Ensemble Bt := FamilyIntersection (S_clo_cont_A A). Lemma closed_gen_ens : forall (E:Ensemble Bt), alg_closed (Gen_Ens E). intro E. unfold Gen_Ens. constructor. (* + *) red. intros x y. destruct x as [x h1]. destruct y as [y h2]. destruct h1 as [x h1]. destruct h2 as [y h2]. constructor. intros S h3. pose proof (h1 _ h3) as h4. pose proof (h2 _ h3) as h5. unfold Bplus_sub. simpl. destruct h3 as [h3]. destruct h3 as [h3l h3r]. apply plus_closed; auto. (* * *) red. intros x y. destruct x as [x h1]. destruct y as [y h2]. destruct h1 as [x h1]. destruct h2 as [y h2]. constructor. intros S h3. pose proof (h1 _ h3) as h4. pose proof (h2 _ h3) as h5. unfold Bplus_sub. simpl. destruct h3 as [h3]. destruct h3 as [h3l h3r]. apply times_closed; auto. (* 1 *) red. constructor. intros S h3. destruct h3 as [h3]. destruct h3 as [h3l h3r]. apply one_closed; auto. (* 0 *) red. constructor. intros S h3. destruct h3 as [h3]. destruct h3 as [h3l h3r]. apply zero_closed; auto. (* - *) red. intro x. destruct x as [x h1]. destruct h1 as [x h1]. constructor. intros S h3. pose proof (h1 _ h3) as h4. unfold Bcomp_sub. simpl. destruct h3 as [h3]. destruct h3 as [h3l h3r]. apply comp_closed; auto. Qed. Definition Gen (E:Ensemble Bt) := Subalg _ (closed_gen_ens E). Lemma gen_ens_minimal : forall (A S:Ensemble Bt), let F := (S_clo_cont_A A) in In F S -> Included (Gen_Ens A) S. intros A S F h0. unfold Gen_Ens. red. intros x h2. inversion h2 as [? h3]. subst. apply h3. assumption. Qed. Lemma gen_ens_includes : forall (S:Ensemble Bt), Included S (Gen_Ens S). intro S. red. intros x h1. constructor. intros S0 h2. inversion h2 as [h3]. destruct h3. auto with sets. Qed. Lemma gen_minimal'': forall (A S:Ensemble Bt), alg_closed S -> Included A S -> Included (int_saf_A (SA_cont_A A)) S. intros A S h1 h2. unfold int_saf_A. pose proof (saf_A_eq A) as h3. rewrite h3. red. intros x h4. inversion h4 as [? h5]. apply h5. constructor. split; assumption. Qed. Lemma incl_gen_ens_intersection : forall (E F:Ensemble Bt), Included (Gen_Ens (Intersection E F)) (Intersection (Gen_Ens E) (Gen_Ens F)). intros E F. red; intros x h1. destruct h1 as [x h1]. constructor. apply h1. constructor. split. pose proof (gen_ens_includes E) as h2. assert (h3:Included (Intersection E F) E). auto with sets. auto with sets. apply closed_gen_ens. apply h1. constructor. split. pose proof (gen_ens_includes F) as h2. assert (h3:Included (Intersection E F) F). auto with sets. auto with sets. apply closed_gen_ens. Qed. Lemma closed_intersection_gen_ens : forall (E F:Ensemble Bt), alg_closed (Intersection (Gen_Ens E) (Gen_Ens F)). intros E F. pose proof (closed_gen_ens E) as h2. pose proof (closed_gen_ens F) as h3. apply closed_intersection; auto. Qed. Lemma gen_ens_union : forall (E F:Ensemble Bt), Gen_Ens (Union E F) = Gen_Ens (Union (Gen_Ens E) (Gen_Ens F)). intros E F. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. apply h1. constructor. split. red. intros a h2. destruct h2 as [a h2 | a h3]. apply gen_ens_includes. left. apply gen_ens_includes; auto. apply gen_ens_includes. right. apply gen_ens_includes; auto. apply closed_gen_ens. red. intros x h1. destruct h1 as [x h1]. apply h1. constructor. split. red. intros a h2. constructor. intros S h3. destruct h3 as [h3]. destruct h2 as [a h2 | a h4]. destruct h2 as [a h2]. apply h2. destruct h3 as [h3a h3b]. constructor. split; auto with sets. destruct h3 as [h3a h3b]. destruct h4 as [a h4]. apply h4. constructor. split; auto with sets. apply closed_gen_ens. Qed. End GenAlg. Arguments alg_closed [B] _. Arguments Gen_Ens [B] _ _. Arguments Gen [B] _. Section Gen_examples. Variable B:Bool_Alg. Let Bt := Btype (Bc B). Example alg_closed_two : alg_closed (@Couple Bt 0 1). constructor. (* + *) red. intros x y. destruct x as [x h2]. destruct y as [y h3]. unfold Bplus_sub. simpl. pose proof (Couple_inv _ _ _ _ h2) as h4. pose proof (Couple_inv _ _ _ _ h3) as h5. destruct h4 as [h4l | h4r]. (*h4l*) destruct h5 as [h5l | h5r]. (*h5l*) rewrite h4l. rewrite h5l. pose proof (@zero_sum B 0) as h6. rewrite h6. apply Couple_l. (*h5r*) rewrite h4l. rewrite h5r. pose proof (@one_sum B 0) as h6. rewrite h6. apply Couple_r. (*h4r*) destruct h5 as [h5l | h5r]. (*h5l*) rewrite h4r. rewrite h5l. pose proof (@zero_sum B 1) as h6. rewrite h6. apply Couple_r. (*h5r*) rewrite h4r. rewrite h5r. pose proof (@one_sum B 1) as h6. rewrite h6. apply Couple_r. (* * *) red. intros x y. destruct x as [x h1]. destruct y as [y h2]. unfold Btimes_sub. simpl. pose proof (Couple_inv _ _ _ _ h1) as h3. pose proof (Couple_inv _ _ _ _ h2) as h4. destruct h3 as [h3l | h3r]. (*h3l*) destruct h4 as [h4l | h4r]. (*h4l*) rewrite h3l. rewrite h4l. pose proof (@zero_prod B 0) as h5. rewrite h5. apply Couple_l. (*h4r*) rewrite h3l. rewrite h4r. pose proof (@one_prod B 0) as h5. rewrite h5. apply Couple_l. (*h3r*) destruct h4 as [h4l | h4r]. (*h4l*) rewrite h3r. rewrite h4l. pose proof (@zero_prod B 1) as h5. rewrite h5. apply Couple_l. (*h4r*) rewrite h3r. rewrite h4r. pose proof (@one_prod B 1) as h5. rewrite h5. apply Couple_r. (* 1 *) red. apply Couple_r. (* 0 *) red. apply Couple_l. (* - *) red. intro x. destruct x as [x h1]. unfold Bcomp_sub. simpl. pose proof (Couple_inv _ _ _ _ h1) as h2. destruct h2 as [h2l | h2r]. (*h2l*) rewrite h2l. pose proof (@zero_comp B) as h3. rewrite h3. apply Couple_r. (*h2r*) rewrite h2r. pose proof (@one_comp B) as h3. rewrite h3. apply Couple_l. Qed. Definition two : Bool_Alg. pose (@Couple Bt 0 1) as A. pose proof (alg_closed_two) as h1. refine (Subalg A h1). Defined. Example two_smallest_subalgebra : forall (A:Ensemble Bt), alg_closed A -> let C := (Couple 0 1) in Included C A. intros A h1 C. red. intros x h2. pose proof (Couple_inv _ _ _ _ h2) as h3. destruct h3 as [h3l | h3r]. rewrite h3l. apply (Z_c _ _ h1). rewrite h3r. apply (O_c _ _ h1). Qed. Lemma gen_ens_closed_eq : forall (S:Ensemble Bt), alg_closed S -> Gen_Ens S = S. intros S h1. apply Extensionality_Ensembles. red. split. (* <= *) apply gen_ens_minimal. constructor. split. auto with sets. assumption. (* >= *) apply gen_ens_includes. Qed. (*"If E is empty and A is not degenerate, then the subalgebra generated by E is the smallest possible subalgebra of A, namely 2."*) Example gen_ens_empty : Gen_Ens (Empty_set Bt) = (Couple 0 1). unfold Gen_Ens. apply Extensionality_Ensembles. red. split. (* <= *) apply gen_ens_minimal. constructor. split; auto with sets. apply alg_closed_two. (* >= *) red. intros x h1. constructor. intros S h2. inversion h2 as [h3]. destruct h3 as [? h4]. pose proof (Couple_inv _ _ _ _ h1) as [h5 | h6]. rewrite h5. apply (Z_c _ _ h4). rewrite h6. apply (O_c _ _ h4). Qed. Example gen_empty : Gen (Empty_set Bt) = two. unfold Gen. unfold two. generalize (closed_gen_ens B (Empty_set Bt)). rewrite gen_ens_empty. intros h1. pose proof (proof_irrelevance _ h1 (alg_closed_two)); subst; auto. Qed. End Gen_examples. Section More_Gen_Ens. (*"A simple but useful remark for subsets E and F of a Boolean algebra A is that if F is included in E, then the subalgebra generated by F is included (as a subalgebra) in the subalgebra generated by E."*) Lemma gen_ens_preserves_inclusion : forall {B:Bool_Alg} (S U: Ensemble (Btype (Bc B))), Included S U -> Included (Gen_Ens S) (Gen_Ens U). intros B S U h1. unfold Gen_Ens. red. intros x h2. inversion h2 as [? h3]; subst. constructor. intros S0 h5. apply h3. constructor. destruct h5 as [h5]. destruct h5; auto with sets. Qed. Lemma two_incl_gen_ens : forall {B:Bool_Alg} (E:Ensemble (bt B)), Included (Couple 0 1) (Gen_Ens E). intros B E. pose proof (gen_ens_preserves_inclusion (Empty_set _) E) as h1. rewrite gen_ens_empty in h1; apply h1; auto with sets. Qed. Lemma not_incl_gen_ens_neq_0 : forall {B:Bool_Alg} (E:Ensemble (bt B)), forall x:bt B, ~Inn (Gen_Ens E) x -> x <> 0. intros B E x h1 h2. subst. pose proof (two_incl_gen_ens E) as h2. auto with sets. Qed. Lemma not_incl_gen_ens_neq_1 : forall {B:Bool_Alg} (E:Ensemble (bt B)), forall x:bt B, ~Inn (Gen_Ens E) x -> x <> 1. intros B E x h1 h2. subst. pose proof (two_incl_gen_ens E) as h2. auto with sets. Qed. Lemma inhabited_gen_ens : forall {B:Bool_Alg} (E:Ensemble (bt B)), Inhabited (Gen_Ens E). intros B E. pose proof (two_incl_gen_ens E) as h1. pose proof (Couple_l _ (Bzero (Bc B)) 1) as h2. apply h1 in h2; econstructor; apply h2. Qed. (*"The relation of one Boolean algebra being a subalgebra of another is a partial order on the class of all Boolean algebras." Technically I don't define [subalg_of] exactly like that, since I presuppose an underlying algebra B, but my definition is equivalent to the informal definition given in the quotes.*) Definition subalg_of {B:Bool_Alg} (A1 A2:{A:Ensemble (Btype (Bc B)) | alg_closed A}) : Prop := (Included (proj1_sig A1) (proj1_sig A2)). Lemma order_subalg_of : forall B:Bool_Alg, Order _ (@subalg_of B). constructor. (*refl*) red. intro. red. auto with sets. (*trans*) red. unfold subalg_of. intros. auto with sets. (*antisymm*) red. unfold subalg_of. intros A1 A2 h1 h2. apply proj1_sig_injective; auto with sets. Qed. (*"The family is said to be directed if any two members B_i and B_j of the family are always subalgebras of some third member Bk."*) Definition directed_sa {A:Bool_Alg} (F:(SubalgF A)) : Prop := forall (S1 S2:Ensemble (Btype (Bc A))), In (saf_A F) S1 -> In (saf_A F) S2 -> (exists (S3:Ensemble (Btype (Bc A))), In (saf_A F) S3 /\ Included S1 S3 /\ Included S2 S3). (*"Lemma 1. The union of a non-empty, directed family of subalgebras is again a subalgebra."*) Lemma union_directed_Subalgf_closed : forall {B:Bool_Alg} (F:(SubalgF B)), directed_sa F -> Inhabited (saf_A F) -> alg_closed (FamilyUnion (saf_A F)). intros B F h1 h2. constructor. (* + *) red. intros p q. destruct p as [p h3]. destruct q as [q h4]. unfold Bplus_sub. simpl. inversion h3 as [B_i ? h5 h6]. inversion h4 as [B_j ? h7 h8]. unfold directed_sa in h1. pose proof (h1 _ _ h5 h7) as h9. elim h9. intros B_k h10. destruct h10 as [h10 [h11l h11r]]. assert (h12:In B_k p). auto with sets. assert (h13:In B_k q). auto with sets. pose proof (saf_closed F B_k h10) as h14. pose proof (P_c _ _ h14) as h15. red in h15. pose proof (h15 (exist _ p h12) (exist _ q h13)) as h16. unfold Bplus_sub in h16. simpl in h16. apply family_union_intro with B_k; assumption. (* * *) red. intros p q. destruct p as [p h3]. destruct q as [q h4]. unfold Btimes_sub. simpl. inversion h3 as [B_i ? h5 h6]. inversion h4 as [B_j ? h7 h8]. unfold directed_sa in h1. pose proof (h1 _ _ h5 h7) as h9. elim h9. intros B_k h10. destruct h10 as [h10 [h11l h11r]]. assert (h12:In B_k p). auto with sets. assert (h13:In B_k q). auto with sets. pose proof (saf_closed F B_k h10) as h14. pose proof (T_c _ _ h14) as h15. red in h15. pose proof (h15 (exist _ p h12) (exist _ q h13)) as h16. unfold Btimes_sub in h16. simpl in h16. apply family_union_intro with B_k; assumption. (* 1 *) red. inversion h2 as [S h3]. pose proof (saf_closed F S h3) as h4. pose proof (O_c _ _ h4) as h5. red in h5. apply family_union_intro with S; assumption. (* 0 *) red. inversion h2 as [S h3]. pose proof (saf_closed F S h3) as h4. pose proof (Z_c _ _ h4) as h5. red in h5. apply family_union_intro with S; assumption. (* - *) red. intro p. destruct p as [p h3]. unfold Bcomp_sub. simpl. inversion h3 as [B_i ? h4 h5]. pose proof (saf_closed F B_i h4) as h6. pose proof (C_c _ _ h6) as h7. red in h7. pose proof (h7 (exist _ p h5)) as h8. red in h8. unfold Bcomp_sub in h8. simpl in h8. apply family_union_intro with B_i; assumption. Qed. (* "More precisely, a family {B_i} of subalgebras is a chain if for any two members Bi and Bj of the family, either Bi is a subalgebra of Bj , or vice versa." *) Definition chain_subalg {B:Bool_Alg} (F:SubalgF B) : Prop := chain_set (saf_A F). Lemma chain_impl_directed : forall {B:Bool_Alg} (F:SubalgF B), chain_subalg F -> directed_sa F. intros B F h1. red in h1. red. intros S1 S2 h2 h3. red in h1. pose proof (h1 _ _ h2 h3) as h4. case h4 as [h4l | h4r]. (* left *) exists S2. split. assumption. split. assumption. auto with sets. (* right *) exists S1. split. assumption. split. auto with sets. assumption. Qed. Lemma union_chain_Subalgf_closed : forall {B:Bool_Alg} (F:SubalgF B), chain_subalg F -> Inhabited (saf_A F) -> alg_closed (FamilyUnion (saf_A F)). intros B F h1 h2. apply union_directed_Subalgf_closed. apply chain_impl_directed; assumption. assumption. Qed. (*"Corollary 1. Let A be a Boolean algebra generated by a set E, and for each
finite subset F of E, let B_F be the subalgebra of A generated by F. The family {B_F:F <= E and F is finite}. is directed, and its union is A." [or in my notation]: Assume B is generated by a set E, and for each finite subset S of E, let A be the subalgebra of B generated by S. The family {A : A is generated by S and S <= E and S is finite} is directed, and its union is B.*) Lemma family_included_subs_directed_eq_B : forall {B:Bool_Alg}, let Bt := (Btype (Bc B)) in forall (E:Ensemble Bt) (F:SubalgF B), (Full_set Bt) = Gen_Ens E -> (saf_A F) = [A:Ensemble Bt | exists (S:Ensemble Bt), A = Gen_Ens S /\ Included S E /\ Finite S] -> directed_sa F /\ FamilyUnion (saf_A F) = (Full_set Bt). intros B Bt E F h1 h2. assert (h3:directed_sa F). red. intros S1' S2' h3 h4. rewrite h2. rewrite h2 in h3. rewrite h2 in h4. inversion h3 as [h5]. inversion h4 as [h6]. elim h5. intros S1 h7. destruct h7 as [h7a [h7b h7c]]. elim h6. intros S2 h8. destruct h8 as [h8a [h8b h8c]]. assert (h9 :Included S1 (Union S1 S2)). auto with sets. assert (h10:Included S2 (Union S1 S2)). auto with sets. pose proof (gen_ens_preserves_inclusion S1 (Union S1 S2) h9) as h11. pose proof (gen_ens_preserves_inclusion S2 (Union S1 S2) h10) as h12. exists (Gen_Ens (Union S1 S2)). constructor. constructor. exists (Union S1 S2). split. reflexivity. split. auto with sets. apply Union_preserves_Finite;assumption. rewrite h7a. rewrite h8a. split; assumption. split. assumption. (* F's union is Full_set *) pose proof (@union_directed_Subalgf_closed B) as h4. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x ?. constructor. (* >= *) assert (h5:Included E (FamilyUnion (saf_A F))). red. intros x h6. pose proof (In_singleton _ x) as h7. apply family_union_intro with (Gen_Ens (Singleton x)). rewrite h2. constructor. exists (Singleton x). split. reflexivity. split. red. intros x' h8. pose proof (Singleton_inv _ _ _ h8) as h9. rewrite <- h9. assumption. apply Singleton_is_finite. constructor. intros S h8. inversion h8 as [h9]. destruct h9 as [h9l h9r]. auto with sets. rewrite h1. apply gen_ens_minimal. constructor. split. assumption. pose proof (Inhabited_intro _ (saf_A F) (Couple 0 1)) as h6. assert (h7:In (saf_A F) (Couple 0 1)). rewrite h2. constructor. exists (Empty_set Bt). split. (*left*) rewrite <- gen_ens_empty. reflexivity. (*right*) split. auto with sets. constructor. pose proof (h6 h7) as h8. pose proof (h4 _ h3 h8). assumption. Qed. Lemma family_included_subs_closed : forall {A:Bool_Alg}, let At:=Btype (Bc A) in forall (E S:Ensemble At), In [A:Ensemble At | exists (S:Ensemble At), A = Gen_Ens S /\ Included S E /\ Finite S] S -> alg_closed S. intros A At E S h1. destruct h1 as [h1]. destruct h1 as [F h1]. destruct h1 as [h1a [h1b h1c]]. subst. apply closed_gen_ens. Qed. Lemma family_included_subs_directed_eq_B' : forall {B:Bool_Alg}, let Bt := (Btype (Bc B)) in forall (E:Ensemble Bt), let saf :=[A:Ensemble Bt | exists S:Ensemble Bt, A = Gen_Ens S /\ Included S E /\ Finite S] in let F := Build_SubalgF B saf (family_included_subs_closed E) in directed_sa F /\ FamilyUnion saf = Gen_Ens E. intros B Bt E saf F. assert (h1:directed_sa F). intros S1 S2 h1 h2. simpl in h1. simpl in h2. destruct h1 as [h1]. destruct h1 as [S h1]. destruct h1 as [h1a [h1b h1c]]. destruct h2 as [h2]. destruct h2 as [S' h2]. destruct h2 as [h2a [h2b h2c]]. subst. assert (h9 :Included S (Union S S')). auto with sets. assert (h10:Included S' (Union S S')). auto with sets. pose proof (gen_ens_preserves_inclusion S (Union S S') h9) as h11. pose proof (gen_ens_preserves_inclusion S' (Union S S') h10) as h12. exists (Gen_Ens (Union S S')). split; try tauto. simpl. constructor. exists (Union S S'). split; auto. split; try apply Union_preserves_Finite; auto with sets. pose proof (@union_directed_Subalgf_closed B) as h4. split. apply h1. apply Extensionality_Ensembles. red. split. red. intros x h5. destruct h5 as [S x h5 h6]. destruct h5 as [h5]. destruct h5 as [S' h5]. destruct h5 as [h5a [h5b h5c]]. subst. pose proof (gen_ens_preserves_inclusion _ _ h5b x h6) as h7. assumption. red. intros S' h5. assert (h6:Included E (FamilyUnion saf)). red. intros x h6. pose proof (In_singleton _ x) as h7. apply family_union_intro with (Gen_Ens (Singleton x)). constructor. exists (Singleton x). split. reflexivity. split. red. intros x' h8. pose proof (Singleton_inv _ _ _ h8) as h9. rewrite <- h9. assumption. apply Singleton_is_finite. constructor. intros S h8. inversion h8 as [h9]. destruct h9 as [h9l h9r]. auto with sets. assert (h7:In (S_clo_cont_A E) (FamilyUnion saf)). constructor. split; auto. specialize (h4 _ h1). apply h4. apply (Inhabited_intro _ (saf_A F) (Couple 0 1)). constructor. exists (Empty_set _). split. rewrite gen_ens_empty. reflexivity. split; auto with sets. pose proof (gen_ens_minimal _ E (FamilyUnion saf) h7) as h8. auto with sets. Qed. (*For any Boolean algebra B, apply the preceding corollary to the set E = B to conclude that every Boolean algebra is the (directed) union of its finitely generated subalgebras.*) Corollary family_included_subs_directed_eq_B_full_set : forall {B:Bool_Alg}, let Bt:=(Btype (Bc B)) in forall (F:SubalgF B), (saf_A F) = [A:Ensemble Bt | exists (S:Ensemble Bt), A = Gen_Ens S /\ Finite S] -> directed_sa F /\ FamilyUnion (saf_A F) = Full_set Bt. Proof. intros B Bt F h1. pose proof (family_included_subs_directed_eq_B (Full_set Bt) F) as h2. assert (h3:Full_set Bt = Gen_Ens (Full_set Bt)). apply Extensionality_Ensembles. red. split. (* <= *) red. intros x ?. unfold Gen_Ens. constructor. intros S h4. inversion h4 as [h5]. destruct h5; auto with sets. (* >= *) red. intros. constructor. apply h2; auto. rewrite h1. apply Extensionality_Ensembles. red. split. (* <= *) red. intros S h4. inversion h4 as [h5]. destruct h5 as [E h6]. destruct h6. constructor. exists E. repeat split; auto with sets. (* >= *) red. intros S h4. inversion h4 as [h5]. destruct h5 as [E h6]. destruct h6 as [? []]. constructor. exists E. split; assumption. Qed. Lemma gen_ba_eq_ba : forall {B:Bool_Alg} (E:Ensemble (bt B)), ba_ens B = Gen_Ens E -> Subalg _ (alg_closed_ba_ens B) = Gen E. intros B E h1. apply bc_inj. unfold Subalg, Gen. unfold Subalg. simpl. pose proof (subsetT_eq_compat _ _ _ _ (alg_closed_ba_ens B) (closed_gen_ens B E) h1) as h2. dependent rewrite -> h2. reflexivity. Qed. Lemma subalg_gen_ens : forall {B:Bool_Alg} (E:Ensemble (bt B)), Subalg (Gen_Ens E) (closed_gen_ens _ E) = Gen E. intros B E. apply bc_inj. simpl. reflexivity. Qed. Lemma gen_ba_eq_type : forall {B:Bool_Alg} (E:Ensemble (bt B)), ba_ens B = Gen_Ens E -> bt (Gen E) = sig_set (ba_ens B). intros B E h1. destruct B. unfold bt, Gen. simpl. unfold SubBtype. simpl. destruct Bc. simpl. simpl in h1. rewrite <- h1. unfold ba_ens. simpl. reflexivity. Qed. Lemma bt_gen : forall {B:Bool_Alg} (E:Ensemble (bt B)), bt (Gen E) = sig_set (Gen_Ens E). intros B E. unfold bt, Gen. simpl. unfold SubBtype, sig_set. reflexivity. Qed. End More_Gen_Ens. Require Import ListUtilities. Section Closed_Finite. Variable B:Bool_Alg. Let Bt:= bt B. Lemma plus_list_closed : forall (S:Ensemble Bt) (l:list Bt), alg_closed S -> Included (list_to_set l) S -> Inn S (plus_list l). intros S l. induction l as [|a l h1]. intros h1 h2. simpl. apply zero_closed. assumption. intros h2 h3. simpl. simpl in h3. assert (h4:Included (list_to_set l) (Add (list_to_set l) a)). auto with sets. assert (h5:Included (list_to_set l) S). auto with sets. specialize (h1 h2 h5). apply plus_closed; auto. pose proof (Add_intro2 _ (list_to_set l) a). auto with sets. Qed. Lemma times_list_closed : forall (S:Ensemble Bt) (l:list Bt), alg_closed S -> Included (list_to_set l) S -> Inn S (times_list l). intros S l. induction l as [|a l h1]. intros h1 h2. simpl. apply one_closed. assumption. intros h2 h3. simpl. simpl in h3. assert (h4:Included (list_to_set l) (Add (list_to_set l) a)). auto with sets. assert (h5:Included (list_to_set l) S). auto with sets. specialize (h1 h2 h5). apply times_closed; auto. pose proof (Add_intro2 _ (list_to_set l) a). auto with sets. Qed. Lemma plus_set_closed : forall (R S:Ensemble Bt) (pf:Finite R), alg_closed S -> Included R S -> Inn S (plus_set R pf). intros R S h1 h2 h3. pose proof (finite_set_list _ h1) as h4. destruct h4 as [l h4]. subst. pose proof (plus_set_compat' _ _ h1 (eq_refl _)) as h4. rewrite h4. apply plus_list_closed; auto. Qed. Lemma times_set_closed : forall (R S:Ensemble Bt) (pf:Finite R), alg_closed S -> Included R S -> Inn S (times_set R pf). intros R S h1 h2 h3. pose proof (finite_set_list _ h1) as h4. destruct h4 as [l h4]. subst. pose proof (times_set_compat' _ _ h1 (eq_refl _)) as h4. rewrite h4. apply times_list_closed; auto. Qed. End Closed_Finite. Section SubCompat. Variable B:Bool_Alg. Let Bt:= Btype (Bc B). Lemma plus_list_sub_compat : forall (A:Ensemble Bt) (pf:alg_closed A) (l:list (bt (Subalg _ pf))), plus_list (map (@proj1_sig _ _) l) = proj1_sig (plus_list l). intros S h1 l. induction l as [|a l h2]. simpl. reflexivity. simpl. rewrite h2. unfold Bplus_sub. reflexivity. Qed. Lemma times_list_sub_compat : forall (A:Ensemble Bt) (pf:alg_closed A) (l:list (bt (Subalg _ pf))), times_list (map (@proj1_sig _ _) l) = proj1_sig (times_list l). intros S h1 l. induction l as [|a l h2]. simpl. reflexivity. simpl. rewrite h2. unfold Btimes_sub. reflexivity. Qed. Lemma plus_set_sub_compat : forall (A:Ensemble Bt) (pf:alg_closed A) (E:Ensemble (bt (Subalg _ pf))) (pfe:Finite E) (pfi:Finite (Im E (@proj1_sig _ _))), plus_set (Im E (@proj1_sig _ _)) pfi = proj1_sig (plus_set E pfe). intros S h1 E h2 h3. pose proof (finite_set_list _ h2) as h4. pose proof (finite_set_list _ h3) as h5. destruct h4 as [l h4]. destruct h5 as [l' h5]. subst. generalize dependent h3. rewrite <- map_im_compat. intro h3. pose proof plus_set_compat'. rewrite (plus_set_compat' _ _ h3 (eq_refl _)). rewrite (plus_set_compat' _ _ h2 (eq_refl _)). unfold Bt. rewrite (plus_list_sub_compat _ h1). reflexivity. Qed. Lemma times_set_sub_compat : forall (A:Ensemble Bt) (pf:alg_closed A) (E:Ensemble (bt (Subalg _ pf))) (pfe:Finite E) (pfi:Finite (Im E (@proj1_sig _ _))), times_set (Im E (@proj1_sig _ _)) pfi = proj1_sig (times_set E pfe). intros S h1 E h2 h3. pose proof (finite_set_list _ h2) as h4. pose proof (finite_set_list _ h3) as h5. destruct h4 as [l h4]. destruct h5 as [l' h5]. subst. generalize dependent h3. rewrite <- map_im_compat. intro h3. pose proof plus_set_compat'. rewrite (times_set_compat' _ _ h3 (eq_refl _)). rewrite (times_set_compat' _ _ h2 (eq_refl _)). unfold Bt. rewrite (times_list_sub_compat _ h1). reflexivity. Qed. (*This section is for the unfortunate syntactic difference between the subalgebra of a full algebra (improper subalgebra), and the full algebra itself. These are formally different, and if that bothers you, then you can always use [Bool_Alg_p]s, in "Section ParametricAnalogues". With those, you can easily express the reflexivity of the subalgebra relation, at the price of an extra primitive type variable. For families of Boolean algebras, that section is essential.*) Section SubalgBaEnsCompat. (*This definition is used to assert that a predicate on elements of an algebra behave the same, whether or not those elements are considered to be elements of a the improper subalgebra or elements of the algebra itself.*) Definition full_sig_equivalent (P:{C:Bool_Alg & (bt C)}->Prop) := forall (x:Bt), P (existT _ B x) <-> P (existT _ (Subalg _ (alg_closed_ba_ens B)) (exist _ _ (Full_intro _ x))). Lemma full_sig_equivalent_atom : full_sig_equivalent (fun x:{C:Bool_Alg & (bt C)} => atom (projT2 x)). red. simpl. intro x. split. intro h1. red in h1. red. destruct h1 as [h1l h1r]. split. intro h2. pose proof (f_equal (@proj1_sig _ _) h2) as h3. simpl in h3. contradiction. intros b h2. destruct b as [b]. red in h2. simpl in h2. unfold Bplus_sub in h2. simpl in h2. pose proof (f_equal (@proj1_sig _ _) h2) as h4. simpl in h4. assert (h5:le b x). red. assumption. specialize (h1r _ h5). destruct h1r as [h1a | h1b]. subst. left. apply proj1_sig_injective. simpl. reflexivity. clear h4. subst. right. apply proj1_sig_injective. simpl. reflexivity. intro h1. red in h1. red. destruct h1 as [h1l h1r]. split. intro h2. subst. contradict h1l. apply proj1_sig_injective. simpl. reflexivity. intros b h2. specialize (h1r (exist _ _ (Full_intro _ b))). let P := type of h1r in match P with ?P -> ?Q => (pose P as h3) end. assert (h4:h3). red. simpl. unfold Bplus_sub. apply proj1_sig_injective. simpl. red in h2. assumption. specialize (h1r h4). simpl in h1r. destruct h1r as [h1a | h1b]. pose proof (f_equal (@proj1_sig _ _) h1a) as h1a'. clear h1a. simpl in h1a'. left. assumption. pose proof (f_equal (@proj1_sig _ _) h1b) as h1b'. clear h1b. simpl in h1b'. right. assumption. Qed. End SubalgBaEnsCompat. End SubCompat. Inductive sign:Set := pls|mns. Definition signl:list sign := pls::mns::nil. Definition signe:Ensemble sign := Couple pls mns. Lemma sign_eq_dec : type_eq_dec sign. red. intros x y. destruct x, y. left; auto. right; intro; discriminate. right; intro; discriminate. left; auto. Qed. Lemma in_sign : forall s:sign, Inn signe s. intro s; destruct s. left. right. Qed. Lemma in_signl : forall s:sign, In s signl. intro s; destruct s. left; auto. right; auto. left; auto. Qed. Definition sign_flip (s:sign) := match s with | pls => mns | mns => pls end. Notation "-- x" := (sign_flip x) (at level 30). Lemma in_sign_flip : forall s:sign, Inn signe (sign_flip s). intro s. destruct s; constructor. Qed. (*Switches the 0 index in sign of a finite two-valued Boolean map*) Definition head_flip_fin_map_seg {n:nat} (a:Fin_map (seg (S n)) signe mns) : Fin_map (seg (S n)) signe mns := fin_map_switch nat_eq_dec a 0%nat (-- (a |->0%nat)) (in_seg_Sn _) (in_sign_flip _). Lemma head_flip_fin_map_seg_0_compat : forall {n:nat} (a:Fin_map (seg (S n)) signe mns), head_flip_fin_map_seg a |-> 0%nat = -- a |->0%nat. intros n a. unfold head_flip_fin_map_seg. match goal with |- fin_map_switch _ ?a ?b ?c ?pfa ?pfb |-> _ = _ => pose proof (fin_map_switch_compat nat_eq_dec a b c pfa pfb) as h1 end. simpl in h1. destruct h1; auto. Qed. Lemma head_flip_fin_map_seg_S_compat : forall {n:nat} (a:Fin_map (seg (S n)) signe mns) (m:nat), m < n -> head_flip_fin_map_seg a |-> S m = a |-> S m. intros n a m h0. unfold head_flip_fin_map_seg. match goal with |- fin_map_switch _ ?a ?b ?c ?pfa ?pfb |-> _ = _ => pose proof (fin_map_switch_compat nat_eq_dec a b c pfa pfb) as h1 end. simpl in h1. destruct h1 as [h1 h2]. apply h2; try omega. constructor. apply lt_n_S; auto. Qed. Lemma sign_dec : forall s s':sign, {s = s'} + {s <> s'}. intros s s'. destruct s, s'; auto; try right. intro h1. discriminate h1. intro h1. discriminate h1. Qed. Lemma sign_opp : forall s s':sign, s <> s' -> s = --s'. unfold sign_flip; intros s s'; destruct s, s'; auto. intro h1; contradict h1; auto. intro h1; contradict h1; auto. Qed. Lemma doub_opp : forall s:sign, --(--s) = s. intro s; destruct s; simpl; auto. Qed. (*Whereas Givant/Halmos uses "p", I use epsilon notation of the Handbook of Boolean Algebras (vol.1 Koppelberg/Monk); but with arguments in the same order as [Givant/Halmos]'s "p", and with the latter's semantics.*) Definition eps {B:Bool_Alg} (b:(bt B)) (sgn:sign) := if sgn then b else -b. Definition eps' {B:Bool_Alg} (pr:(bt B)*sign) := eps (fst pr) (snd pr). Lemma card_signe : card_fun1 signe = 2%nat. apply card_fun1_couple. intro; discriminate. Qed. Lemma list_to_set_sign : list_to_set signl = signe. unfold signl. simpl. rewrite <- couple_add_add. unfold signe. rewrite couple_comm. reflexivity. Qed. Lemma signe_finite : Finite signe. rewrite <- list_to_set_sign. apply list_to_set_finite. Qed. Lemma eps_eps'_compat : forall {B:Bool_Alg} (x:(bt B)), eps x = (fun y:sign => eps' (x, y)). intro x. unfold eps. unfold eps'. simpl. unfold eps. reflexivity. Qed. Lemma eps_inj : forall {B:Bool_Alg} (x y:bt B) (s s':sign), eps x s = eps y s' -> (x = y /\ s = s') \/ (x = -y /\ s = --s'). intros B x y s s' h1. unfold eps in h1. destruct s, s'. left; auto. right; auto. rewrite <- h1. rewrite <- doub_neg. right; auto. apply comp_unq in h1. left; auto. Qed. Lemma eps_opp : forall {B:Bool_Alg} (x:bt B) (s:sign), eps x (--s) = - (eps x s). unfold sign_flip; intros B x s; destruct s; simpl; auto. rewrite <- doub_neg; auto. Qed. Lemma eps_covers : forall {B:Bool_Alg} (E:Ensemble (bt B)) (pf:Finite E) (x:bt B), plus_set (Im signe (eps x)) (finite_image _ _ signe (eps x) (signe_finite)) = 1. intros B E h1 x. assert (h2:list_to_set ((eps x pls) :: (- (eps x pls)) :: nil) = (Im signe (eps x))). apply Extensionality_Ensembles. red. split. red. intros y h3. rewrite <- list_to_set_in_iff in h3. simpl in h3. destruct h3 as [h3a | [h3b | h3c]]. subst. apply Im_intro with pls. constructor. simpl. reflexivity. apply Im_intro with mns. constructor. simpl. subst. reflexivity. contradiction. red. intros y h3. destruct h3 as [y h3 z]. subst. unfold signe in h3. destruct h3. simpl. right. constructor. simpl. left. right. constructor. rewrite (plus_set_compat' _ (eps x pls :: - eps x pls :: nil)). simpl. rewrite zero_sum. rewrite comp_sum. reflexivity. unfold bt. unfold bt in h2. rewrite h2. reflexivity. Qed. Section HasComps. (*This is used in conjunction with [el_prod_l_firstn]*) Inductive has_comps_elt {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:bt B) (a:Fin_map (seg m) signe mns) : Prop := | has_comps_elt_intro : Finite P -> forall (c d:nat) (pfc:c < m) (pfd:d < m), c <> d -> Inn P c -> Inn P d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = x -> nth_lt l d (lt_le_trans _ _ _ pfd pfm) = x -> a |-> c <> a |-> d -> has_comps_elt _ _ pfm P x a. (*This is used in conjunction with [el_prod_l]*) Definition has_comps_elt' {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (x:bt B) (a:Fin_map (seg (length l)) signe mns) : Prop := has_comps_elt l (length l) (le_refl _) P x a. Inductive has_sames_elt {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:bt B) (a:Fin_map (seg m) signe mns) : Prop := | has_sames_elt_intro : Finite P -> forall (c d:nat) (pfc:c < m) (pfd:d < m), c <> d -> Inn P c -> Inn P d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = x -> nth_lt l d (lt_le_trans _ _ _ pfd pfm) = x -> a |-> c = a |-> d -> has_sames_elt _ _ pfm P x a. Definition has_sames_elt' {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (x:bt B) (a:Fin_map (seg (length l)) signe mns) : Prop := has_sames_elt l (length l) (le_refl _) P x a. Lemma has_comps_elt_finite : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:bt B) (a:Fin_map (seg m) signe mns), has_comps_elt l m pfm P x a -> Finite P. intros ? ? ? ? ? ? ? h1. destruct h1; auto. Qed. Lemma has_comps_elt_finite' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (x:bt B) (a:Fin_map (seg (length l)) signe mns), has_comps_elt' l P x a -> Finite P. intros ? ? ? ? ? h1; red in h1. apply has_comps_elt_finite in h1; auto. Qed. Lemma has_sames_elt_finite : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:bt B) (a:Fin_map (seg m) signe mns), has_sames_elt l m pfm P x a -> Finite P. intros ? ? ? ? ? ? ? h1. destruct h1; auto. Qed. Lemma has_sames_elt_finite' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (x:bt B) (a:Fin_map (seg (length l)) signe mns), has_sames_elt' l P x a -> Finite P. intros ? ? ? ? ? h1; red in h1. apply has_sames_elt_finite in h1; auto. Qed. Lemma has_comps_elt_functional : forall {B:Bool_Alg} (l l':list (bt B)) (m:nat) (x x':bt B) (pfm:m <= length l) (pfm':m <= length l') (P:Ensemble nat) (a:Fin_map (seg m) signe mns), l = l' -> x = x' -> has_comps_elt _ _ pfm P x a -> has_comps_elt _ _ pfm' P x' a. simpl; intros; subst. pose proof (proof_irrelevance _ pfm pfm'); subst; auto. Qed. Lemma has_comps_elt_functional' : forall {B:Bool_Alg} (l l':list (bt B)) (m:nat) (x x':bt B) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), l = l' -> x = x' -> has_comps_elt' _ P x a -> has_comps_elt' _ P x' a. intros; subst; auto. Qed. Lemma has_sames_elt_functional : forall {B:Bool_Alg} (l l':list (bt B)) (m:nat) (x x':bt B) (pfm:m <= length l) (pfm':m <= length l') (P:Ensemble nat) (a:Fin_map (seg m) signe mns), l = l' -> x = x' -> has_sames_elt _ _ pfm P x a -> has_sames_elt _ _ pfm' P x a. simpl; intros; subst. pose proof (proof_irrelevance _ pfm pfm'); subst; auto. Qed. Lemma has_sames_elt_functional' : forall {B:Bool_Alg} (l l':list (bt B)) (x x':bt B) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), l = l' -> x = x' -> has_sames_elt' _ P x a -> has_sames_elt' _ P x' a. intros; subst; auto. Qed. Lemma has_comps_elt_dec : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (C:Ensemble nat) (x:bt B) (a:Fin_map (seg m) signe mns), Finite C -> {has_comps_elt l m pfm C x a} + {~has_comps_elt l m pfm C x a}. intros B l m h1 C x F hc. destruct (nil_dec' l) as [|hn]; subst. simpl in h1. pose proof (le0 _ h1); subst. right. intro h2. inversion h2. omega. destruct (in_dec ba_eq_dec x l) as [h2 | h2]. pose (fun c => exists d (pfc:c < m) (pfd:d < m), Inn C c /\ Inn C d /\ c <> d /\ nth_lt l c (lt_le_trans _ _ _ pfc h1) = x /\ nth_lt l d (lt_le_trans _ _ _ pfd h1) = x /\ F |-> c <> F |-> d) as P. assert (h3:pred_dec P). red. intro n. unfold P. pose (fun d pfd => exists pfc : n < m, n <> d /\ Inn C n /\ Inn C d /\ nth_lt l n (lt_le_trans n m (length l) pfc h1) = x /\ nth_lt l d (lt_le_trans d m (length l) pfd h1) = x /\ F |-> n <> F |-> d) as R. assert (h4:pred_dep_nat_dec R). red. intros r h3. unfold R. destruct (lt_le_dec n m) as [h4 | h4]. destruct (nat_eq_dec n r) as [h5 | h5]. right. intro h6. destruct h6; omega. destruct (ba_eq_dec (nth_lt l n (lt_le_trans n m (length l) h4 h1)) x) as [h8 | h8]. destruct (ba_eq_dec (nth_lt l r (lt_le_trans r m (length l) h3 h1)) x) as [h9 | h9]. destruct (sign_eq_dec (F |-> n) (F |-> r)) as [h10 | h10]. right. intro h11. destruct h11 as [? [? [? [? [? [? ?]]]]]]. contradiction. destruct (in_fin_ens_dec nat_eq_dec C hc r) as [hir | hir]. destruct (in_fin_ens_dec nat_eq_dec C hc n) as [hin | hin]. left. exists h4. repeat split; auto. right. intro h11. destruct h11 as [? [? [? [? [? [? ?]]]]]]. contradiction. right. intro h11. destruct h11 as [? [? [? [? [? [? ?]]]]]]. contradiction. right. intro h11. destruct h11 as [? [? [? [? [? [? ?]]]]]]. contradiction. right; intro h11; destruct h11 as [? [? [? [? [? [? ?]]]]]]. rewrite <- H2 in h8. contradict h8; apply nth_lt_functional3. right. intro h5. destruct h5; omega. pose proof (lt_ex_dep_nat_dec _ h4) as h5. unfold R in h5. destruct h5 as [h5 | h5]. left. destruct h5 as [r [h5 [h6 [h7 [h8 [h9 [h10 [h11 h12]]]]]]]]. exists r, h6, h5. repeat split; auto. right. contradict h5. destruct h5 as [r [h5 [h6 [h7 [h8 [h9 [h10 [h11 h12]]]]]]]]. exists r, h6, h5. repeat split; auto. pose proof (lt_ex_dec _ m h3) as h4. unfold P in h4. destruct h4 as [h4 | h4]. left. destruct h4 as [r [h4 [s [h6 [h7 [h8 [h9 [h10 [h11 [h12 h13]]]]]]]]]]. econstructor. assumption. apply h10. assumption. assumption. apply h11. apply h12. apply h13. right. contradict h4. destruct h4. exists c. split; auto. exists d, pfc, pfd. split; auto. right. intro h3. destruct h3; subst. contradict h2. apply in_nth_lt. Qed. Lemma has_comps_elt_dec' : forall {B:Bool_Alg} (l:list (bt B)) (C:Ensemble nat) (x:bt B) (a:Fin_map (seg (length l)) signe mns), Finite C -> {has_comps_elt' l C x a} + {~has_comps_elt' l C x a}. intros; unfold has_comps_elt'; apply has_comps_elt_dec; auto. Qed. Lemma has_sames_elt_dec : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (C:Ensemble nat) (x:bt B) (a:Fin_map (seg m) signe mns), Finite C -> {has_sames_elt l m pfm C x a} + {~has_sames_elt l m pfm C x a}. intros B l m h1 C x F hc. destruct (nil_dec' l) as [|hn]; subst. simpl in h1. pose proof (le0 _ h1); subst. right. intro h2. inversion h2. omega. destruct (in_dec ba_eq_dec x l) as [h2 | h2]. pose (fun c => exists d (pfc:c < m) (pfd:d < m), Inn C c /\ Inn C d /\ c <> d /\ nth_lt l c (lt_le_trans _ _ _ pfc h1) = x /\ nth_lt l d (lt_le_trans _ _ _ pfd h1) = x /\ F |-> c = F |-> d) as P. assert (h3:pred_dec P). red. intro n. unfold P. pose (fun d pfd => exists pfc : n < m, n <> d /\ Inn C n /\ Inn C d /\ nth_lt l n (lt_le_trans n m (length l) pfc h1) = x /\ nth_lt l d (lt_le_trans d m (length l) pfd h1) = x /\ F |-> n = F |-> d) as R. assert (h4:pred_dep_nat_dec R). red. intros r h3. unfold R. destruct (lt_le_dec n m) as [h4 | h4]. destruct (nat_eq_dec n r) as [h5 | h5]. right. intro h6. destruct h6; omega. destruct (ba_eq_dec (nth_lt l n (lt_le_trans n m (length l) h4 h1)) x) as [h8 | h8]. destruct (ba_eq_dec (nth_lt l r (lt_le_trans r m (length l) h3 h1)) x) as [h9 | h9]. destruct (sign_eq_dec (F |-> n) (F |-> r)) as [h10 | h10]. destruct (in_fin_ens_dec nat_eq_dec C hc r) as [hir | hir]. destruct (in_fin_ens_dec nat_eq_dec C hc n) as [hin | hin]. left. exists h4. repeat split; auto. right. intro h11. destruct h11 as [? [? [? [? [? [? ?]]]]]]. contradiction. right. intro h11. destruct h11 as [? [? [? [? [? [? ?]]]]]]. contradiction. right. intro h11. destruct h11 as [? [? [? [? [? [? ?]]]]]]. contradiction. right. intro h11. destruct h11 as [? [? [? [? [? [? ?]]]]]]. contradiction. right; intro h11; destruct h11 as [? [? [? [? [? [? ?]]]]]]. rewrite <- H2 in h8. contradict h8; apply nth_lt_functional3. right. intro h5. destruct h5; omega. pose proof (lt_ex_dep_nat_dec _ h4) as h5. unfold R in h5. destruct h5 as [h5 | h5]. left. destruct h5 as [r [h5 [h6 [h7 [h8 [h9 [h10 [h11 h12]]]]]]]]. exists r, h6, h5. repeat split; auto. right. contradict h5. destruct h5 as [r [h5 [h6 [h7 [h8 [h9 [h10 [h11 h12]]]]]]]]. exists r, h6, h5. repeat split; auto. pose proof (lt_ex_dec _ m h3) as h4. unfold P in h4. destruct h4 as [h4 | h4]. left. destruct h4 as [r [h4 [s [h6 [h7 [h8 [h9 [h10 [h11 [h12 h13]]]]]]]]]]. econstructor. assumption. apply h10. assumption. assumption. apply h11. apply h12. apply h13. right. contradict h4. destruct h4. exists c. split; auto. exists d, pfc, pfd. split; auto. right. intro h3. destruct h3; subst. contradict h2. apply in_nth_lt. Qed. Lemma has_sames_elt_dec' : forall {B:Bool_Alg} (l:list (bt B)) (C:Ensemble nat) (x:bt B) (a:Fin_map (seg (length l)) signe mns), Finite C -> {has_sames_elt' l C x a} + {~has_sames_elt' l C x a}. intros; unfold has_sames_elt'; apply has_sames_elt_dec; auto. Qed. Lemma has_comps_elt_S : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfsm:S m <= length l) (C:Ensemble nat) (x:bt B) (a:Fin_map (seg (S m)) signe mns), let pfm := le_trans _ _ _ (le_n_Sn m) pfsm in has_comps_elt l m pfm C x (fin_map_seg_rest a m (lt_n_Sn _)) -> has_comps_elt l (S m) pfsm C x a. intros B l m h1 C x a h2 h3. inversion h3; subst. do 2 rewrite fin_map_seg_rest_compat in H5; auto. econstructor. assumption. apply H0. assumption. assumption. apply nth_lt_functional3. rewrite <- H4 at 1. apply nth_lt_functional3. assumption. Grab Existential Variables. omega. omega. Qed. (*In these lemmas [has_comps] and [has_sames], it's oftentimes inconvenient to use [econstructor] rather than explicit constructor application.*) Inductive has_comps {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) : Prop := | has_comps_intro : forall (c d:nat) (pfc:c < m) (pfd:d < m), Finite P -> Inn P c -> Inn P d -> c <> d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = nth_lt l d (lt_le_trans _ _ _ pfd pfm) -> a |-> c <> a |-> d -> has_comps _ _ pfm P a. Definition has_comps' {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) : Prop := has_comps l (length l) (le_refl _) P a. Inductive has_sames {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) : Prop := | has_sames_intro : forall (c d:nat) (pfc:c < m) (pfd:d < m), Finite P -> Inn P c -> Inn P d -> c <> d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = nth_lt l d (lt_le_trans _ _ _ pfd pfm) -> a |-> c = a |-> d -> has_sames _ _ pfm P a. Definition has_sames' {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) : Prop := has_sames l (length l) (le_refl _) P a. Lemma has_comps_finite : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps l m pfm P a -> Finite P. intros ? ? ? ? ? ? h1. destruct h1; auto. Qed. Lemma has_comps_finite' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps' l P a -> Finite P. unfold has_comps'; intros; eapply has_comps_finite; apply H. Qed. Lemma has_sames_finite : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames l m pfm P a -> Finite P. intros ? ? ? ? ? ? h1. destruct h1; auto. Qed. Lemma has_sames_finite' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames' l P a -> Finite P. unfold has_sames'; intros; eapply has_sames_finite; apply H. Qed. Lemma has_comps_functional : forall {B:Bool_Alg} (l l':list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) (pfe:l = l'), let pfm' := le_length_subst _ _ _ pfe pfm in has_comps _ _ pfm P a -> has_comps _ _ pfm' P a. simpl; intros; subst. pose proof (proof_irrelevance _ pfm (le_length_subst l' l' m eq_refl pfm)) as h1. rewrite <- h1; auto. Qed. Lemma has_comps_functional' : forall {B:Bool_Alg} (l l':list (bt B)) (m:nat) (pfm:m <= length l) (pfm':m <= length l') (P:Ensemble nat) (a:Fin_map (seg m) signe mns), l = l' -> has_comps _ _ pfm P a -> has_comps _ _ pfm' P a. simpl; intros; subst. pose proof (proof_irrelevance _ pfm pfm'); subst; auto. Qed. Lemma has_sames_functional : forall {B:Bool_Alg} (l l':list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) (pfe:l = l'), let pfm' := le_length_subst _ _ _ pfe pfm in has_sames _ _ pfm P a -> has_sames _ _ pfm' P a. simpl; intros; subst. pose proof (proof_irrelevance _ pfm (le_length_subst l' l' m eq_refl pfm)) as h1. rewrite <- h1; auto. Qed. Lemma has_sames_functional' : forall {B:Bool_Alg} (l l':list (bt B)) (m:nat) (pfm:m <= length l) (pfm':m <= length l') (P:Ensemble nat) (a:Fin_map (seg m) signe mns), l = l' -> has_sames _ _ pfm P a -> has_sames _ _ pfm' P a. simpl; intros; subst. pose proof (proof_irrelevance _ pfm pfm'); subst; auto. Qed. Lemma has_comps_transfer : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (m':nat) (pfm:m <= length l) (pfm':m' <= length l) (pfe:m = m') (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps l m pfm P a -> let a' := fin_map_seg_transfer a pfe in has_comps l m' pfm' P a'. intros B l m m' h1 h2 h3 P a h4; simpl. subst. eapply has_comps_functional'. reflexivity. rewrite fin_map_seg_transfer_eq_refl. apply h4. Qed. Lemma has_sames_transfer : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (m':nat) (pfm:m <= length l) (pfm':m' <= length l) (pfe:m = m') (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames l m pfm P a -> let a' := fin_map_seg_transfer a pfe in has_sames l m' pfm' P a'. intros B l m m' h1 h2 h3 P a h4; simpl. subst. eapply has_sames_functional'. reflexivity. rewrite fin_map_seg_transfer_eq_refl. apply h4. Qed. Lemma has_comps_functional_rev : forall {B:Bool_Alg} (l:list (bt B)) (m n:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) (pfem:m = n), let pfn := eq_le_trans _ _ _ (eq_sym pfem) pfm in let a' := fin_map_seg_transfer a pfem in has_comps l n pfn P a' -> has_comps l m pfm P a. simpl; intros Bp l m n h1 P a h2 h3; subst. rewrite fin_map_seg_transfer_eq_refl in h3. simpl. pose proof (proof_irrelevance _ (eq_le_trans _ _ _ eq_refl h1) h1) as h4. rewrite <- h4; auto. Qed. Lemma has_comps_functional_rev' : forall {B:Bool_Alg} (l l':list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) (pfel:l = l'), let pfell := length_functional _ _ pfel in let a' := fin_map_seg_transfer a pfell in has_comps' l' P a' -> has_comps' l P a. simpl; intros; subst; simpl. pose proof (proof_irrelevance _ (eq_refl _) (length_functional l' l' eq_refl)) as h1. rewrite <- h1 in H. rewrite <- fin_map_seg_transfer_eq_refl; auto. Qed. Lemma has_sames_functional_rev : forall {B:Bool_Alg} (l:list (bt B)) (m n:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) (pfem:m = n), let pfn := eq_le_trans _ _ _ (eq_sym pfem) pfm in let a' := fin_map_seg_transfer a pfem in has_sames l n pfn P a' -> has_sames l m pfm P a. simpl; intros Bp l m n h1 P a h2 h3; subst. rewrite fin_map_seg_transfer_eq_refl in h3. simpl in h3. pose proof (proof_irrelevance _ h1 (eq_le_trans _ _ _ eq_refl h1)) as h4. rewrite h4; auto. Qed. Lemma has_sames_functional_rev' : forall {B:Bool_Alg} (l l':list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) (pfel:l = l'), let pfell := length_functional _ _ pfel in let a' := fin_map_seg_transfer a pfell in has_sames' l' P a' -> has_sames' l P a. simpl; intros; subst; simpl. rename H into h. pose proof (proof_irrelevance _ (eq_refl _) (length_functional l' l' eq_refl)) as h1. rewrite <- h1 in h. rewrite fin_map_seg_transfer_eq_refl in h; auto. Qed. Lemma has_comps_transfer_rev : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (m':nat) (pfm:m <= length l) (pfm':m' <= length l) (pfe:m = m') (P:Ensemble nat) (a:Fin_map (seg m) signe mns), let a' := fin_map_seg_transfer a pfe in has_comps l m' pfm' P a' -> has_comps l m pfm P a. simpl; intros B l m m' h1 h2 h3 P a h4; subst. rewrite fin_map_seg_transfer_eq_refl in h4. eapply has_comps_functional'. reflexivity. apply h4. Qed. Lemma has_sames_transfer_rev : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (m':nat) (pfm:m <= length l) (pfm':m' <= length l) (pfe:m = m') (P:Ensemble nat) (a:Fin_map (seg m) signe mns), let a' := fin_map_seg_transfer a pfe in has_sames l m' pfm' P a' -> has_sames l m pfm P a. simpl; intros B l m m' h1 h2 h3 P a h4; subst. rewrite fin_map_seg_transfer_eq_refl in h4. eapply has_sames_functional'. reflexivity. apply h4. Qed. Lemma has_comps_iff : forall {B:Bool_Alg} {l:list (bt B)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps _ _ pfm P a <-> Finite P /\ (exists (x:bt B) (r:nat) (pfr:r < m), let pfr := lt_le_trans _ _ _ pfr pfm in Inn P r /\ nth_lt l r pfr = x /\ has_comps_elt _ _ pfm P x a). intros B l m h1 P a; split; intro h2. destruct h2. split; auto. exists (nth_lt _ _ (lt_le_trans _ _ _ pfc h1)). exists c, pfc. simpl. repeat split; auto. econstructor. assumption. apply H2. assumption. assumption. reflexivity. rewrite H3. reflexivity. assumption. destruct h2 as [h2 [x [r [h3 [h4 [h5 h6]]]]]]; subst. inversion h6. econstructor. assumption. apply H1. apply H2. assumption. rewrite H3 at 1. rewrite <- H4. reflexivity. assumption. Qed. Lemma has_comps_iff' : forall {B:Bool_Alg} {l:list (bt B)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps' _ P a <-> Finite P /\ (exists (x:bt B) (r:nat) (pfr:r < length l), Inn P r /\ nth_lt l r pfr = x /\ has_comps_elt' _ P x a). unfold has_comps'; intros. rewrite has_comps_iff; unfold has_comps_elt'; auto. intuition. destruct H1 as [x [r [h2 [h3 [h4 h5]]]]]. exists x, _, h2. repeat split; auto. rewrite <- h4. apply nth_lt_functional3. destruct H1 as [x [r [h2 [h3 [h4 h5]]]]]. exists x, _, h2. repeat split; auto. rewrite <- h4. apply nth_lt_functional3. Qed. Lemma has_sames_iff : forall {B:Bool_Alg} {l:list (bt B)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames _ _ pfm P a <-> Finite P /\ (exists (x:bt B) (r:nat) (pfr:r < m), let pfr := lt_le_trans _ _ _ pfr pfm in Inn P r /\ nth_lt l r pfr = x /\ has_sames_elt _ _ pfm P x a). intros B l m h1 P a; split; intro h2. destruct h2. split; auto. exists (nth_lt _ _ (lt_le_trans _ _ _ pfc h1)). exists c, pfc. simpl. repeat split; auto. econstructor. assumption. apply H2. assumption. assumption. reflexivity. rewrite H3. reflexivity. assumption. destruct h2 as [h2 [x [r [h3 [h4 [h5 h6]]]]]]; subst. inversion h6. econstructor. assumption. apply H1. apply H2. assumption. rewrite H3 at 1. rewrite <- H4. reflexivity. assumption. Qed. Lemma has_sames_iff' : forall {B:Bool_Alg} {l:list (bt B)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames' _ P a <-> Finite P /\ (exists (x:bt B) (r:nat) (pfr:r < length l), Inn P r /\ nth_lt l r pfr = x /\ has_sames_elt' _ P x a). unfold has_sames'; intros. rewrite has_sames_iff; unfold has_sames_elt'; auto. intuition. destruct H1 as [x [r [h2 [h3 [h4 h5]]]]]. exists x, _, h2. repeat split; auto. rewrite <- h4. apply nth_lt_functional3. destruct H1 as [x [r [h2 [h3 [h4 h5]]]]]. exists x, _, h2. repeat split; auto. rewrite <- h4. apply nth_lt_functional3. Qed. Lemma has_comps_gt_0 : forall {B:Bool_Alg} {l:list (bt B)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps _ _ pfm P a -> 0 < m. intros B l m hm P a h1. destruct m as [|m]. destruct h1. simpl in pfc, pfd. omega. auto with arith. Qed. Lemma has_comps_gt_0' : forall {B:Bool_Alg} {l:list (bt B)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps' _ P a -> l <> nil. intros B l P a h1. destruct l as [|c l]. inversion h1. simpl in pfc; omega. intro; subst; discriminate. Qed. Lemma has_sames_gt_0 : forall {B:Bool_Alg} {l:list (bt B)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames _ _ pfm P a -> 0 < m. intros B l m hm P a h1. destruct m as [|m]. destruct h1. simpl in pfc, pfd. omega. auto with arith. Qed. Lemma has_sames_gt_0' : forall {B:Bool_Alg} {l:list (bt B)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames' _ P a -> l <> nil. intros B l P a h1. destruct l as [|c l]. inversion h1. simpl in pfc; omega. intro; subst; discriminate. Qed. Lemma has_comps_inh : forall {B:Bool_Alg} {l:list (bt B)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps _ _ pfm P a -> Inhabited P. intros B l m hm P a h1. destruct h1. econstructor. apply H0. Qed. Lemma has_comps_inh' : forall {B:Bool_Alg} {l:list (bt B)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps' _ P a -> Inhabited P. intros ? ? ? ? h1; red in h1; apply has_comps_inh in h1; auto. Qed. Lemma has_sames_inh : forall {B:Bool_Alg} {l:list (bt B)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames _ _ pfm P a -> Inhabited P. intros B l m hm P a h1. destruct h1. econstructor. apply H0. Qed. Lemma has_sames_inh' : forall {B:Bool_Alg} {l:list (bt B)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames' _ P a -> Inhabited P. intros ? ? ? ? h1; red in h1; apply has_sames_inh in h1; auto. Qed. Lemma has_comps_dup : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps _ _ pfm P a -> ~NoDup (firstn m l). intros B l a h0 P F h1. destruct h1 as [m n h1 h2 h3]. eapply nth_lt_dup'. apply H1. erewrite nth_lt_firstn_eq. symmetry. erewrite nth_lt_firstn_eq. symmetry. erewrite nth_lt_functional3 in H2. rewrite H2 at 1. apply nth_lt_functional3. Grab Existential Variables. rewrite length_firstn; auto. rewrite length_firstn; auto. Qed. Lemma has_comps_dup' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps' _ P a -> ~NoDup l. unfold has_comps'; intros ? ? ? ? h1; apply has_comps_dup in h1. rewrite firstn_length in h1; auto. Qed. Lemma has_sames_dup : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames _ _ pfm P a -> ~NoDup (firstn m l). intros B l a h0 P F h1. destruct h1 as [m n h1 h2 h3]. eapply nth_lt_dup'. apply H1. erewrite nth_lt_firstn_eq. symmetry. erewrite nth_lt_firstn_eq. symmetry. erewrite nth_lt_functional3 in H2. rewrite H2 at 1. apply nth_lt_functional3. Grab Existential Variables. rewrite length_firstn; auto. rewrite length_firstn; auto. Qed. Lemma has_sames_dup' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames' _ P a -> ~NoDup l. unfold has_sames'; intros. apply has_sames_dup in H; auto; rewrite firstn_length in H; auto. Qed. Lemma has_sames_cons : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (b:bt B) (a:Fin_map (seg (S m)) signe mns) (pfm:m <= length l) (pfsm:S m <= length (b::l)) (P:Ensemble nat), has_sames _ _ pfm P (fin_map_seg_pred a) -> has_sames _ _ pfsm (ens_S P) a. intros B l m b a hm hsm P h1. destruct h1 as [r s h2 h3 h4]. do 2 rewrite fin_map_seg_pred_compat in H3. econstructor. apply finite_ens_S; auto. econstructor. apply H. reflexivity. econstructor. apply H0. reflexivity. contradict H1; auto with arith. do 2 erewrite nth_lt_cons'; simpl. erewrite nth_lt_functional3. rewrite H2 at 1. apply nth_lt_functional3. assumption. assumption. assumption. assumption. Grab Existential Variables. auto with arith. auto with arith. auto with arith. auto with arith. Qed. Lemma has_sames_cons_in : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (b:bt B) (a:Fin_map (seg (S m)) signe mns) (pfm:m <= length l) (pfsm:S m <= length (b::l)) (P:Ensemble nat), In b l -> ~Inn P 0%nat -> has_sames _ _ pfsm P a -> has_sames _ _ pfm (ens_pred P) (fin_map_seg_pred a). intros B l m b a hm hsm P hz hi h1. destruct h1 as [r s h2 h3 h4]. econstructor. apply finite_ens_pred; auto. econstructor. apply H. reflexivity. econstructor. apply H0. reflexivity. destruct (zerop r) as [|hzr]; subst. contradiction. destruct (zerop s) as [|hzs]; subst. contradiction. omega. destruct r as [|r], s as [|s]. contradict H1; auto. contradiction. contradiction. rewrite nth_lt_cons in H2. simpl. erewrite nth_lt_functional3. rewrite H2 at 1. rewrite nth_lt_cons. apply nth_lt_functional3. destruct r as [|r], s as [|s]. contradict H2; auto. contradiction. contradiction. simpl. rewrite fin_map_seg_pred_compat; auto with arith. rewrite fin_map_seg_pred_compat; auto with arith. Grab Existential Variables. destruct s as [|s]. contradiction. auto with arith. omega. Qed. Lemma has_comps_cons : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (b:bt B) (a:Fin_map (seg (S m)) signe mns) (pfm:m <= length l) (pfsm:S m <= length (b::l)) (P:Ensemble nat), has_comps _ _ pfm P (fin_map_seg_pred a) -> has_comps _ _ pfsm (ens_S P) a. intros B l m b a hm hsm P h1. destruct h1 as [r s h2 h3 h4]. do 2 rewrite fin_map_seg_pred_compat in H3. econstructor. apply finite_ens_S; auto. econstructor. apply H. reflexivity. econstructor. apply H0. reflexivity. omega. do 2 erewrite nth_lt_cons'; simpl. erewrite nth_lt_functional3. rewrite H2 at 1. apply nth_lt_functional3. intro h5. contradiction. assumption. assumption. assumption. Grab Existential Variables. auto with arith. auto with arith. auto with arith. auto with arith. Qed. Lemma has_comps_incl : forall {B:Bool_Alg} {l:list (bt B)} {m:nat} (pfm:m <= length l) (C C':Ensemble nat) (F:Fin_map (seg m) signe mns), Included C C' -> Finite C' -> has_comps _ _ pfm C F -> has_comps _ _ pfm C' F. intros B l m hm C C' F h1 hf h2. destruct h2 as [r s hr hs h0 h4 h5 h6 h7]. apply h1 in h4. apply h1 in h5. econstructor. assumption. apply h4. apply h5. assumption. apply h7. assumption. Qed. Lemma has_comps_incl' : forall {B:Bool_Alg} {l:list (bt B)} (C C':Ensemble nat) (F:Fin_map (seg (length l)) signe mns), Included C C' -> Finite C' -> has_comps' _ C F -> has_comps' _ C' F. unfold has_comps'; intros; eapply has_comps_incl. apply H. apply H0. assumption. Qed. Lemma has_comps_cons_ens_pred : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (b:bt B) (a:Fin_map (seg (S m)) signe mns) (pfm:m <= length l) (pfsm:S m <= length (b::l)) (P:Ensemble nat), Finite P -> has_comps _ _ pfm (ens_pred P) (fin_map_seg_pred a) -> has_comps _ _ pfsm P a \/ Inn P 0%nat /\ exists (n:nat) (pfnm:n < m), let pfnm' := lt_le_trans _ _ _ pfnm pfm in hd b l = nth_lt l n pfnm' /\ Inn P (S n) /\ a |-> 1%nat <> a |-> S n. intros B l m b a hm hsm P h0 h1. destruct (in_fin_ens_dec nat_eq_dec P h0 0%nat) as [hi | hi]. destruct (in_fin_ens_dec nat_eq_dec P h0 1%nat) as [hi' | hi']. rewrite ens_pred01_eq in h1; auto. inversion h1. assert (hclt:c < length (b::l)). simpl. apply (lt_trans _ _ _ pfc hsm). assert (hdlt:d < length (b::l)). simpl. apply (lt_trans _ _ _ pfd hsm). destruct H0 as [r h5]; subst. destruct H1 as [s h6]; subst. destruct h5 as [h5 h7], h6 as [h6 h8]. rewrite in_sing_iff in h7, h8. do 2 rewrite fin_map_seg_pred_compat in H4. destruct (zerop r) as [|hzr]; subst. contradict h7; auto. destruct (zerop s) as [|hzs]; subst. contradict h8; auto. rewrite <- (S_pred _ _ hzr) in H4. rewrite <- (S_pred _ _ hzs) in H4. left. pose proof (has_comps_intro _ _ hsm P a r s) as hh. hfirst hh. destruct r; auto with arith. specialize (hh hf). hfirst hh. destruct s; auto with arith. specialize (hh hf0 h0 h5 h6). hfirst hh. intro; subst. contradict H2; auto. specialize (hh hf1). pose proof (nth_lt_compat (b::l) r (lt_le_trans r (S m) (length (b :: l)) hf hsm)) as hb. simpl in hh, hb. rewrite hb in hh. hfirst hh. destruct r, s; auto. contradict h7; auto. contradict h8; auto. simpl in H3. rewrite nth_lt_compat in H3. erewrite nth_indep. rewrite H3 at 1. rewrite nth_lt_compat. apply nth_indep. simpl in hclt, hdlt. omega. omega. apply (hh hf2 H4). assumption. assumption. assumption. inversion h1. do 2 rewrite fin_map_seg_pred_compat in H4; auto. destruct (zerop c) as [|h6]; subst. right. split; auto. inversion H0; subst. assert (h7:x = 0%nat). symmetry in H6. apply pred0_dec in H6. destruct H6; subst; auto. contradiction. subst. destruct H1 as [d]; subst. rewrite nth_lt_0_hd' in H3. assert (hz:0%nat < d). destruct (zerop d); subst; auto with arith. exists _, pfd. simpl. rewrite <- (hd_hd_compat' _ _ b) in H3. rewrite H3. do 2 split; auto. rewrite <- (S_pred _ _ hz); auto. destruct (zerop d) as [|h6']; subst. right. split; auto. inversion H1; subst. assert (h7:x = 0%nat). symmetry in H6. apply pred0_dec in H6. destruct H6; subst; auto. contradiction. subst. destruct H0 as [c]; subst. rewrite nth_lt_0_hd' in H3. assert (hz:0%nat < c). destruct (zerop c); subst; auto with arith. exists _, pfc. simpl. rewrite <- (hd_hd_compat' _ _ b) in H3. rewrite H3. do 2 split; auto. rewrite <- (S_pred _ _ hz); auto. destruct H0; subst. destruct H1; subst. assert (h5:0 < x). destruct (zerop x) as [|h5]; subst. simpl in h6. apply Lt.lt_irrefl in h6; contradiction. assumption. destruct (zerop x0) as [|h5']; subst. simpl in h6'. apply Lt.lt_irrefl in h6'; contradiction. rewrite <- (S_pred _ _ h5) in H4. rewrite <- (S_pred _ _ h5') in H4. left. pose proof (has_comps_intro _ _ hsm P a x x0) as hh. hfirst hh. omega. specialize (hh hf). hfirst hh. omega. specialize (hh hf0 h0 H0 H1). hfirst hh. intro; subst. contradict H4; auto. specialize (hh hf1). simpl in hh. hfirst hh. destruct x, x0; auto. apply Lt.lt_irrefl in h5; contradiction. apply Lt.lt_irrefl in h5'; contradiction. simpl in H3. rewrite nth_lt_compat in H3. erewrite nth_indep. rewrite H3 at 1. rewrite nth_lt_compat. apply nth_indep. omega. omega. specialize (hh hf2). apply hh; auto. left. inversion h1. do 2 rewrite fin_map_seg_pred_compat in H4; auto. destruct H0; subst. destruct H1; subst. econstructor. assumption. apply H0. apply H1. intro; subst. contradict H2; auto. destruct x as [|x], x0 as [|x0]; subst; simpl; auto; try contradiction. simpl in H3. pose proof (lt_le_trans _ _ _ pfc hm) as h5. simpl in h5. rewrite nth_lt_compat in H3. destruct x as [|x]. pose proof (nth_indep l b (nth_lt l 0 (lt_le_trans 0 m (length l) pfc hm)) (lt_le_trans 0 m (length l) pfc hm)) as he. rewrite he. rewrite H3. rewrite nth_lt_compat. apply nth_indep. destruct x0 as [|x0]; auto with arith. apply (lt_le_trans (S x0) m (length l) pfd hm). pose proof (nth_indep l b (nth_lt l (S x) (lt_le_trans (S x) m (length l) pfc hm)) (lt_le_trans (S x) m (length l) pfc hm)) as he'. rewrite he'. rewrite H3 at 1. rewrite nth_lt_compat. apply nth_indep. destruct x0; auto. simpl in hsm. apply (lt_le_trans _ _ _ pfd hm). apply (lt_le_trans (S x0) m (length l) pfd hm). destruct (zerop x0); subst. simpl in H2. destruct (zerop x) as [|hz]; subst. contradict H2; auto. rewrite (S_pred _ _ hz) in H4; auto. rewrite <- (S_pred _ _ l0) in H4. destruct (zerop x) as [|hz]; subst. contradiction. rewrite <- (S_pred _ _ hz) in H4. assumption. Grab Existential Variables. destruct x0; auto with arith. destruct x; auto with arith. Qed. Lemma has_comps_cons_in : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (b:bt B) (a:Fin_map (seg (S m)) signe mns) (pfm:m <= length l) (pfsm:S m <= length (b::l)) (P:Ensemble nat), In b l -> ~Inn P 0%nat -> has_comps _ _ pfsm P a -> has_comps _ _ pfm (ens_pred P) (fin_map_seg_pred a). intros B l m b a hm hsm P hz hi h1. destruct h1 as [r s h2 h3 h4]. destruct (zerop r) as [|hzr]; subst. contradiction. destruct (zerop s) as [|hzs]; subst. contradiction. econstructor. apply finite_ens_pred; auto. econstructor. apply H. reflexivity. econstructor. apply H0. reflexivity. contradict H1. apply pred_inj0 in H1; auto. rewrite (nth_lt_cons' _ _ _ _ hzr) in H2. rewrite (nth_lt_cons' _ _ _ _ hzs) in H2. erewrite nth_lt_functional3. rewrite H2 at 1. apply nth_lt_functional3. rewrite fin_map_seg_pred_compat. rewrite fin_map_seg_pred_compat. contradict H2; auto. simpl. rewrite <- (S_pred _ _ hzr), <- (S_pred _ _ hzs) in H2. contradiction. omega. omega. Grab Existential Variables. omega. omega. Qed. Lemma has_comps_rest_impl : forall {B:Bool_Alg} {l:list (bt B)} (m:nat) (pfm : m <= length l) (pfsm : S m <= length l) (C:Ensemble nat) (F:Fin_map (seg (S m)) signe mns), Finite C -> has_comps _ _ pfm C (fin_map_seg_rest F m (lt_n_Sn m)) -> has_comps _ _ pfsm C F. intros B l m h1 h2 C F hc h3. destruct h3 as [r s hr hs hf h4 h5 h6 h7 h8]. simpl in hr, hs. do 2 rewrite fin_map_seg_rest_compat in h8; auto. econstructor; simpl. assumption. apply h4. apply h5. assumption. simpl in h7. erewrite nth_lt_functional3. rewrite h7 at 1. apply nth_lt_functional3. assumption. Grab Existential Variables. simpl. apply lt_S; auto. simpl. apply lt_S; auto. Qed. Lemma has_comps_rest_impl' : forall {B:Bool_Alg} {l:list (bt B)} (m:nat) (pfm : m <= length l) (pfsm : S m <= length l) (C:Ensemble nat) (F:Fin_map (seg (S m)) signe mns), has_comps _ _ pfm C (fin_map_seg_rest F m (lt_n_Sn m)) -> has_comps _ _ pfsm (Add C m) F. intros B l m h1 h2 C F h3. inversion h3 as [r s hr hs hf h4 h5 h6 h7 h8]. clear h3. do 2 rewrite fin_map_seg_rest_compat in h8; auto. simpl in hr, hs. econstructor; simpl. apply Add_preserves_Finite; auto. left. apply h4. left. apply h5. assumption. simpl in h7. erewrite nth_lt_functional3. rewrite h7 at 1. apply nth_lt_functional3. assumption. Grab Existential Variables. simpl; apply lt_S; auto. simpl; apply lt_S; auto. Qed. Lemma has_comps_S_rest_iff : forall {B:Bool_Alg} {l:list (bt B)} (m:nat) (pfm : m <= length l) (pfsm : S m <= length l) (C:Ensemble nat) (F:Fin_map (seg (S m)) signe mns), Finite C -> (has_comps _ _ pfsm C F <-> has_comps _ _ pfm C (fin_map_seg_rest F m (lt_n_Sn m)) \/ (Inn C m /\ exists (r:nat) (pfr:r < m), let pft := lt_le_trans _ _ _ pfr pfm in Inn C r /\ nth_lt l r pft = nth_lt l m pfsm /\ F |-> r <> F |-> m)). intros B l m h1 h2 C F h3; split; intro h4. inversion h4. red in pfc, pfd. pose proof (le_lt_eq_dec _ _ pfc) as h8. pose proof (le_lt_eq_dec _ _ pfd) as h9. destruct h8 as [h8 | h8], h9 as [h9 | h9]. apply lt_S_n in h8. apply lt_S_n in h9. pose proof (has_comps_intro _ _ h1 C (fin_map_seg_rest F m (lt_n_Sn _)) c d h8 h9 H H0 H1 H2) as hi. hfirst hi. erewrite nth_lt_functional3. rewrite H3 at 1. apply nth_lt_functional3. specialize (hi hf). do 2 rewrite fin_map_seg_rest_compat in hi. specialize (hi H4). left; auto. assumption. assumption. assumption. apply S_inj in h9; subst. right. split; auto. exists c. ex_goal. omega. exists hex. simpl; split; auto. erewrite nth_lt_functional3. rewrite H3 at 1. split; auto. apply nth_lt_functional3. apply S_inj in h8; subst. right. split; auto. exists d. ex_goal. omega. exists hex. simpl. apply neq_sym in H4. split; auto. split; auto. symmetry. erewrite nth_lt_functional3. rewrite H3 at 1. apply nth_lt_functional3. apply S_inj in h8; apply S_inj in h9; subst. contradict H4; auto. destruct h4 as [h4 | h4]. inversion h4. pose proof (has_comps_intro l (S m) h2 C F _ _ (lt_trans _ _ _ pfc (lt_n_Sn _)) (lt_trans _ _ _ pfd (lt_n_Sn _)) h3 H0 H1 H2). apply H5. erewrite nth_lt_functional3. rewrite H3 at 1. apply nth_lt_functional3. do 2 rewrite fin_map_seg_rest_compat in H4; auto. destruct h4 as [hi [r [h4 [h5 h6]]]]. apply (has_comps_intro l (S m) h2 C F _ _ (lt_trans _ _ _ h4 (lt_n_Sn _)) (lt_n_Sn _) h3 h5 hi). intro; subst. pose proof (Lt.lt_irrefl _ h4); contradiction. erewrite nth_lt_functional3. destruct h6 as [h6 h7]. rewrite h6 at 1. apply nth_lt_functional3. destruct h6; auto. Qed. Lemma has_comps_dec : forall {B:Bool_Alg} {l:list (bt B)} {m:nat} (pfm : m <= length l) (F:Fin_map (seg m) signe mns) (P:Ensemble nat), Finite P -> {has_comps _ _ pfm P F} + {~has_comps _ _ pfm P F}. intros B l m h1 F P h2. pose (fun x => exists (c d : nat) (pfc : c < m) (pfd : d < m), c <> d /\ Inn P c /\ Inn P d /\ nth_lt l c (lt_le_trans c m (length l) pfc h1) = x /\ nth_lt l d (lt_le_trans d m (length l) pfd h1) = x /\ F |-> c <> F |-> d ) as R. pose proof (l_ex_dec ba_eq_dec R l) as h3. hfirst h3. red. intro x. unfold R. pose (fun c pfc => exists (d:nat) (pfd : d < m), Inn P c /\ Inn P d /\ nth_lt l c (lt_le_trans c m (length l) pfc h1) = x /\ nth_lt l d (lt_le_trans d m (length l) pfd h1) = x /\ F |-> c <> F |-> d) as Q. pose proof (lt_ex_dep_nat_dec Q) as h4. hfirst h4. red. unfold Q. intros c h5. pose (fun d pfd => Inn P c /\ Inn P d /\ nth_lt l c (lt_le_trans c m (length l) h5 h1) = x /\ nth_lt l d (lt_le_trans d m (length l) pfd h1) = x /\ F |-> c <> F |-> d) as G. pose proof (lt_ex_dep_nat_dec G) as h6. hfirst h6. unfold G; red. intros d h7. destruct (in_fin_ens_dec nat_eq_dec P h2 c) as [h8 | h8]. destruct (in_fin_ens_dec nat_eq_dec P h2 d) as [h9 | h9]. destruct (ba_eq_dec (nth_lt l c (lt_le_trans c m (length l) h5 h1)) x) as [h10 | h10]. destruct (ba_eq_dec (nth_lt l d (lt_le_trans d m (length l) h7 h1)) x) as [h11 | h11]. destruct (sign_eq_dec (F |-> c) (F |-> d)) as [h12 | h12]. right. intro h13. destruct h13 as [? [? [? [? ?]]]]. contradiction. left. repeat split; auto. right. intro h13. destruct h13 as [? [? [? [? ?]]]]. contradiction. right. intro h13. destruct h13 as [? [? [? [? ?]]]]. contradiction. right. intro h13. destruct h13 as [? [? [? [? ?]]]]. contradiction. right. intro h13. destruct h13 as [? [? [? [? ?]]]]. contradiction. specialize (h6 hf). unfold G in h6. destruct h6 as [h6 | h6]. left. destruct h6 as [d [h7 [h8 [h9 [h10 [h11 h12]]]]]]. exists d, h7. repeat split; auto. right. contradict h6. destruct h6 as [d [h7 [h8 [h9 [h10 [h11 h12]]]]]]. exists d, h7. repeat split; auto. specialize (h4 hf). unfold Q in h4. destruct h4 as [h4 | h4]. left. destruct h4 as [c [h7 [d [h9 [h10 [h11 [h12 [h13 h14]]]]]]]]. exists c, d, h7, h9. repeat split; auto. right. contradict h4. destruct h4 as [c [h7 [d [h9 [h10 [h11 [h12 [h13 [h14 h15]]]]]]]]]. exists c, d, h7, h9. repeat split; auto. specialize (h3 hf). destruct h3 as [h3 | h3]. left. rewrite has_comps_iff. split; auto. unfold R in h3. destruct h3 as [x [h3 [r [s [h6 [h7 [h8 [h9 [h10 [h11 [h12 h13]]]]]]]]]]]. exists x, r, h6. simpl. repeat split; auto. econstructor. assumption. apply h8. assumption. assumption. apply h11. apply h12. assumption. right. contradict h3. destruct h3. exists (nth_lt _ _ (lt_le_trans _ _ _ pfc h1)). split. apply in_nth_lt. unfold R. exists c, d, pfc, pfd. repeat split; auto. Qed. Lemma has_comps_dec' : forall {B:Bool_Alg} {l:list (bt B)} (F:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> {has_comps' _ P F} + {~has_comps' _ P F}. unfold has_comps'; intros; apply has_comps_dec; auto. Qed. Lemma has_sames_dec : forall {B:Bool_Alg} {l:list (bt B)} {m:nat} (pfm : m <= length l) (F:Fin_map (seg m) signe mns) (P:Ensemble nat), Finite P -> {has_sames _ _ pfm P F} + {~has_sames _ _ pfm P F}. intros B l m h1 F P h2. pose (fun x => exists (c d : nat) (pfc : c < m) (pfd : d < m), c <> d /\ Inn P c /\ Inn P d /\ nth_lt l c (lt_le_trans c m (length l) pfc h1) = x /\ nth_lt l d (lt_le_trans d m (length l) pfd h1) = x /\ F |-> c = F |-> d ) as R. pose proof (l_ex_dec ba_eq_dec R l) as h3. hfirst h3. red. intro x. unfold R. pose (fun c pfc => exists (d:nat) (pfd : d < m), c <> d /\ Inn P c /\ Inn P d /\ nth_lt l c (lt_le_trans c m (length l) pfc h1) = x /\ nth_lt l d (lt_le_trans d m (length l) pfd h1) = x /\ F |-> c = F |-> d) as Q. pose proof (lt_ex_dep_nat_dec Q) as h4. hfirst h4. red. unfold Q. intros c h5. pose (fun d pfd => c <> d /\ Inn P c /\ Inn P d /\ nth_lt l c (lt_le_trans c m (length l) h5 h1) = x /\ nth_lt l d (lt_le_trans d m (length l) pfd h1) = x /\ F |-> c = F |-> d) as G. pose proof (lt_ex_dep_nat_dec G) as h6. hfirst h6. unfold G; red. intros d h7. destruct (in_fin_ens_dec nat_eq_dec P h2 c) as [h8 | h8]. destruct (in_fin_ens_dec nat_eq_dec P h2 d) as [h9 | h9]. destruct (ba_eq_dec (nth_lt l c (lt_le_trans c m (length l) h5 h1)) x) as [h10 | h10]. destruct (ba_eq_dec (nth_lt l d (lt_le_trans d m (length l) h7 h1)) x) as [h11 | h11]. destruct (sign_eq_dec (F |-> c) (F |-> d)) as [h12 | h12]. destruct (nat_eq_dec c d) as [h13 | h13]. right. intro h14. destruct h14; contradiction. left. repeat split; auto. right. intro h13. destruct h13 as [? [? [? [? [? ?]]]]]. contradiction. right. intro h13. destruct h13 as [? [? [? [? [? ?]]]]]. contradiction. right. intro h13. destruct h13 as [? [? [? [? [? ?]]]]]. unfold G in h6; unfold Q in h4. clear G Q. contradiction. right. intro h13. destruct h13 as [? [? [? [? ?]]]]. contradiction. right. intro h13. destruct h13 as [? [? [? [? ?]]]]. contradiction. specialize (h6 hf). unfold G in h6. destruct h6 as [h6 | h6]. left. destruct h6 as [d [h7 [h8 [h9 [h10 [h11 [h12 h13]]]]]]]. exists d, h7. repeat split; auto. right. contradict h6. destruct h6 as [d [h7 [h8 [h9 [h10 [h11 [h12 h13]]]]]]]. exists d, h7. repeat split; auto. specialize (h4 hf). unfold Q in h4. destruct h4 as [h4 | h4]. left. destruct h4 as [c [h7 [d [h9 [h10 [h11 [h12 [h13 [h14 h15]]]]]]]]]. exists c, d, h7, h9. repeat split; auto. right. contradict h4. destruct h4 as [c [h7 [d [h9 [h10 [h11 [h12 [h13 [h14 h15]]]]]]]]]. exists c, d, h7, h9. repeat split; auto. specialize (h3 hf). destruct h3 as [h3 | h3]. left. rewrite has_sames_iff. split; auto. unfold R in h3. destruct h3 as [x [h3 [r [s [h6 [h7 [h8 [h9 [h10 [h11 [h12 h13]]]]]]]]]]]. exists x, r, h6. simpl. repeat split; auto. econstructor. assumption. apply h8. assumption. assumption. apply h11. apply h12. assumption. right. contradict h3. destruct h3. exists (nth_lt _ _ (lt_le_trans _ _ _ pfc h1)). split. apply in_nth_lt. unfold R. exists c, d, pfc, pfd. repeat split; auto. Qed. Lemma has_sames_dec' : forall {B:Bool_Alg} {l:list (bt B)} (F:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> {has_sames' _ P F} + {~has_sames' _ P F}. unfold has_sames'; intros; apply has_sames_dec; auto. Qed. Lemma not_has_comps_or_sames_nin_iff : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (a:Fin_map (seg m) signe mns) (P:Ensemble nat), Finite P -> (~has_comps _ _ pfm P a /\ ~has_sames _ _ pfm P a <-> (forall (r s:nat) (pfr:r < m) (pfs:s < m), r <> s -> let pfr' := lt_le_trans _ _ _ pfr pfm in let pfs' := lt_le_trans _ _ _ pfs pfm in nth_lt l r pfr' = nth_lt l s pfs' -> ~Inn P r \/ ~Inn P s)). intros B l m h1 a P h2; split; intro h3. destruct h3 as [h3 h4]. simpl. destruct (zerop m) as [|h5]; subst. simpl. intros; omega. intros r s h7 h8 h11 h12. destruct (in_fin_ens_dec nat_eq_dec P h2 r) as [h13 | h13]. destruct (in_fin_ens_dec nat_eq_dec P h2 s) as [h14 | h14]. destruct (sign_eq_dec (a |-> r) (a |-> s)) as [hs | hs]. contradict h4. apply (has_sames_intro l m h1 P a _ _ h7 h8 h2 h13 h14 h11 h12 hs). contradict h3. apply (has_comps_intro l m h1 P a _ _ h7 h8 h2 h13 h14 h11 h12); auto. right; auto. left; auto. split. intro h4. inversion h4; subst. specialize (h3 _ _ pfc pfd H2 H3). destruct h3; contradiction. intro h4. destruct h4. specialize (h3 _ _ pfc pfd H2 H3). destruct h3; contradiction. Qed. Lemma not_has_comps_or_sames_nin_iff' : forall {B:Bool_Alg} (l:list (bt B)) (a:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> (~has_comps' _ P a /\ ~has_sames' _ P a <-> (forall (r s:nat) (pfr:r < length l) (pfs:s < length l), r <> s -> nth_lt l r pfr = nth_lt l s pfs -> ~Inn P r \/ ~Inn P s)). unfold has_comps', has_sames'; intros; rewrite not_has_comps_or_sames_nin_iff. simpl. intuition. eapply H0; auto. erewrite nth_lt_functional3. rewrite H2 at 1. apply nth_lt_functional3. eapply H0; auto. apply H2. assumption. Grab Existential Variables. assumption. assumption. Qed. Lemma not_has_comps_or_sames_elt_nin_iff : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (x:bt B) (a:Fin_map (seg m) signe mns) (P:Ensemble nat), Finite P -> (~has_comps_elt _ _ pfm P x a /\ ~has_sames_elt _ _ pfm P x a <-> (forall (r s: nat) (pfr: r < m) (pfs: s < m), let pfr' := lt_le_trans r m (length l) pfr pfm in let pfs' := lt_le_trans s m (length l) pfs pfm in r <> s -> nth_lt l r pfr' = x -> nth_lt l s pfs' = x -> ~ Inn P r \/ ~Inn P s)). intros B l m h1 x a P h2; split; intro h3. intros r s h5 h6 h7 h8 h9 h10 h11; subst. destruct h3 as [h3 h4]. destruct (in_fin_ens_dec nat_eq_dec P h2 r) as [h14 | h14]. destruct (in_fin_ens_dec nat_eq_dec P h2 s) as [h12 | h12]. destruct (sign_eq_dec (a |-> r) (a |-> s)) as [h13 | h13]. contradict h4. econstructor. assumption. apply h9. assumption. assumption. apply nth_lt_functional3. rewrite <- h11. apply nth_lt_functional3. assumption. contradict h3. econstructor. assumption. apply h9. assumption. assumption. apply nth_lt_functional3. rewrite <- h11. apply nth_lt_functional3. assumption. right; auto. left; auto. destruct (has_comps_elt_dec l m h1 P x a h2) as [h4 | h4]. inversion h4; subst. specialize (h3 c d pfc pfd H0 eq_refl H4). tauto. split; auto. destruct (has_sames_elt_dec l m h1 P x a h2) as [h5 | h5]. inversion h5; subst. specialize (h3 c d pfc pfd H0 eq_refl H4). tauto. auto. Grab Existential Variables. assumption. assumption. assumption. assumption. Qed. Lemma not_has_comps_or_sames_elt_nin_iff' : forall {B:Bool_Alg} (l:list (bt B)) (x:bt B) (a:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> (~has_comps_elt' _ P x a /\ ~has_sames_elt' _ P x a <-> (forall (r s: nat) (pfr: r < length l) (pfs: s < length l), r <> s -> nth_lt l r pfr = x -> nth_lt l s pfs = x -> ~Inn P r \/ ~Inn P s)). unfold has_comps_elt', has_sames_elt'; intros B l x a P h1. erewrite not_has_comps_or_sames_elt_nin_iff; auto. simpl. intuition; subst. eapply H; auto. apply nth_lt_functional3. rewrite <- H2. apply nth_lt_functional3. eapply H; auto. rewrite <- H2. apply nth_lt_functional3. Grab Existential Variables. assumption. assumption. assumption. Qed. Lemma has_dec : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (a:Fin_map (seg m) signe mns) (P:Ensemble nat), Finite P -> {has_comps _ _ pfm P a} + {has_sames _ _ pfm P a} + {forall (r s:nat) (pfr:r < m) (pfs:s < m), r <> s -> let pfr' := lt_le_trans _ _ _ pfr pfm in let pfs' := lt_le_trans _ _ _ pfs pfm in nth_lt l r pfr' = nth_lt l s pfs' -> ~Inn P r \/ ~Inn P s}. intros B l m h1 a P h2. destruct (has_comps_dec h1 a P h2) as [h3 | h3]. left; auto. destruct (has_sames_dec h1 a P h2) as [h4 | h4]. left; right. assumption. pose proof (not_has_comps_or_sames_nin_iff l m h1 a P h2) as h5. hl h5. split; auto. rewrite h5 in hl. right. assumption. Qed. Lemma has_dec' : forall {B:Bool_Alg} (l:list (bt B)) (a:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> {has_comps' _ P a} + {has_sames' _ P a} + {forall (r s:nat) (pfr:r < length l) (pfs:s < length l), r <> s -> nth_lt l r pfr = nth_lt l s pfs -> ~Inn P r \/ ~Inn P s}. unfold has_comps', has_sames'; intros B l a P h1. pose proof (has_dec _ (length l) (le_refl _) a P h1) as h2. destruct h2 as [[h2 | h2] | h2]. left; auto. left; right; auto. right. intros r s h3 h4 h5 h6. specialize (h2 r s h3 h4 h5). simpl in h2. hfirst h2. erewrite nth_lt_functional3. rewrite h6 at 1. apply nth_lt_functional3. apply h2; auto. Qed. Lemma has_elt_dec : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (x:bt B) (a:Fin_map (seg m) signe mns) (P:Ensemble nat), Finite P -> {has_comps_elt _ _ pfm P x a} + {has_sames_elt _ _ pfm P x a} + {(forall (r s: nat) (pfr: r < m) (pfs: s < m), let pfr' := lt_le_trans r m (length l) pfr pfm in let pfs' := lt_le_trans s m (length l) pfs pfm in r <> s -> nth_lt l r pfr' = x -> nth_lt l s pfs' = x -> ~ Inn P r \/ ~Inn P s)}. intros B l m h1 x a P h2. destruct (has_comps_elt_dec _ _ h1 P x a h2) as [h3 | h3]. left; auto. destruct (has_sames_elt_dec _ _ h1 P x a h2) as [h4 | h4]. left; right; auto. pose proof (not_has_comps_or_sames_elt_nin_iff l m h1 x a P h2) as h5. hl h5. split; auto. rewrite h5 in hl. right; auto. Qed. Lemma has_elt_dec' : forall {B:Bool_Alg} (l:list (bt B)) (x:bt B) (a:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> {has_comps_elt' _ P x a} + {has_sames_elt' _ P x a} + {(forall (r s: nat) (pfr: r < length l) (pfs: s < length l), r <> s -> nth_lt l r pfr = x -> nth_lt l s pfs = x -> ~ Inn P r \/ ~Inn P s)}. unfold has_comps_elt', has_sames_elt'. intros B l x a P h1. pose proof (has_elt_dec _ (length l) (le_refl _) x a P h1) as h2. destruct h2 as [[h2 | h2] | h2]. left; auto. left; right; auto. right. intros r s h3 h4 h5 h6 h7; subst. simpl in h2. specialize (h2 _ _ h3 h4 h5). apply h2. apply nth_lt_functional3. rewrite <- h7; apply nth_lt_functional3. Qed. Lemma not_has_comps_elt_const : forall {B:Bool_Alg} {l:list (bt B)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (x:bt B) (F:Fin_map (seg m) signe mns), Finite P -> ~has_comps_elt _ _ pfm P x F -> forall (c d:nat) (pfc:c < m) (pfd: d < m), let pfc' := lt_le_trans _ _ _ pfc pfm in let pfd' := lt_le_trans _ _ _ pfd pfm in c <> d -> Inn P c -> Inn P d -> nth_lt l c pfc' = x -> nth_lt l d pfd' = x -> F |-> c = F |-> d. intros B l m h1 P x F hf h2 c d hc hd h3 h4 h5 h6 h7 h8 h9; subst. destruct (sign_dec (F |-> c) (F |-> d)) as [| h10]; auto. contradict h2. econstructor. assumption. apply h5. assumption. assumption. apply nth_lt_functional3. rewrite <- h9. apply nth_lt_functional3. assumption. Grab Existential Variables. assumption. assumption. Qed. Lemma not_has_comps_elt_const' : forall {B:Bool_Alg} {l:list (bt B)} (P:Ensemble nat) (x:bt B) (F:Fin_map (seg (length l)) signe mns), Finite P -> ~has_comps_elt' _ P x F -> forall (c d:nat) (pfc:c < length l) (pfd: d < length l), c <> d -> Inn P c -> Inn P d -> nth_lt l c pfc = x -> nth_lt l d pfd = x -> F |-> c = F |-> d. intros; eapply not_has_comps_elt_const. apply H. apply H0. assumption. assumption. assumption. subst. apply nth_lt_functional3. rewrite <- H5. apply nth_lt_functional3. Grab Existential Variables. assumption. assumption. Qed. Lemma has_comps_S_impl : forall {A:Bool_Alg} {l:list (bt A)} (m:nat) (pfm : m <= length l) (pfsm : S m <= length l) (C:Ensemble nat) (F:Fin_map (seg (S m)) signe mns), Finite C -> has_comps _ _ pfsm C F -> (has_comps _ _ pfm C (fin_map_seg_rest F m (lt_n_Sn m))) \/ (Inn C m /\ (exists (r:nat) (pfr:r < m), Inn C r /\ nth_lt l r (lt_le_trans _ _ _ pfr pfm) = nth_lt l m (lt_le_trans _ _ _ (lt_n_Sn m) pfsm) /\ F |-> r <> F |-> m)). intros A l m h1 h2 C F hf h3. destruct h3 as [r s h6 h7 h0 h8 h9 h10 h11 h12]. simpl in h6, h7, h11. destruct (nat_eq_dec r m) as [h13 | h13]; subst. right. split; auto. exists s. ex_goal. omega. exists hex. split. assumption. split. erewrite nth_lt_functional3. unfold bt in h11. unfold bt. rewrite <- h11 at 1. apply nth_lt_functional3. apply neq_sym; auto. destruct (nat_eq_dec s m) as [h14 | h14]; subst. right. split; auto. exists r. ex_goal. omega. exists hex. split. assumption. split. erewrite nth_lt_functional3. unfold bt in h11. unfold bt. rewrite h11 at 1. apply nth_lt_functional3. assumption. left. econstructor. assumption. apply h8. apply h9. assumption. simpl. erewrite nth_lt_functional3. unfold bt in h11. unfold bt. rewrite h11 at 1. apply nth_lt_functional3. rewrite fin_map_seg_rest_compat. rewrite fin_map_seg_rest_compat. assumption. omega. omega. Grab Existential Variables. simpl. omega. simpl. omega. Qed. Lemma has_comps_rest_impl_final : forall {A:Bool_Alg} {l:list (bt A)} (m:nat) (pfm : m <= length l) (pfsm : S m <= length l) (C:Ensemble nat) (F:Fin_map (seg (S m)) signe mns), Finite C -> has_comps _ _ pfm C (fin_map_seg_rest F m (lt_n_Sn m)) -> has_comps _ _ pfsm C F. intros A l m h1 h2 C F hc h3. destruct h3 as [r s hr hs hf h4 h5 h6 h7 h8]. simpl in hr, hs. do 2 rewrite fin_map_seg_rest_compat in h8; auto. econstructor; simpl. assumption. apply h4. apply h5. assumption. simpl in h7. erewrite nth_lt_functional3. rewrite h7 at 1. apply nth_lt_functional3. assumption. Grab Existential Variables. simpl. apply lt_S; auto. simpl. apply lt_S; auto. Qed. Lemma has_comps_linds_occ_impl : forall {A:Bool_Alg} (la:list (bt A)) (m:nat) (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (x:bt A), has_comps la m pfm (list_to_set (linds_occ ba_eq_dec (firstn m la) x)) F -> has_comps la m pfm (seg m) F. intros A la m h1 F x h2. inversion h2. econstructor. apply finite_seg. constructor. apply pfc. constructor. apply pfd. assumption. apply H3. assumption. Qed. Lemma has_comps_linds_occ_impl' : forall {A:Bool_Alg} (la:list (bt A)) (F:Fin_map (seg (length la)) signe mns) (x:bt A), has_comps' la (list_to_set (linds_occ ba_eq_dec la x)) F -> has_comps' la (seg (length la)) F. unfold has_comps'; intros; eapply has_comps_linds_occ_impl. rewrite firstn_length. apply H. Qed. Lemma nin_has_comps_impl_removelast : forall {A:Bool_Alg} {l:list (bt A)} {m:nat} (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns), ~In (nth_lt l m pfsm) (firstn m l) -> forall x:(bt A), In x (firstn m l) -> has_comps l (S m) pfsm (list_to_set (linds_occ ba_eq_dec (firstn (S m) l) x)) F -> let pfm := le_length_removelast l m pfsm in has_comps (removelast l) m pfm (list_to_set (linds_occ ba_eq_dec (firstn m (removelast l)) x)) (fin_map_seg_rest F m (lt_n_Sn _)). intros A l m h1 F h2 x h3 h4 h5. inversion h4; subst. red in pfc, pfd. pose proof pfc as hc. pose proof pfd as hd. apply le_S_n in hc. apply le_S_n in hd. apply le_lt_eq_dec in hc. apply le_lt_eq_dec in hd. destruct hc as [hc |], hd as [hd|]; subst. econstructor. apply list_to_set_finite. rewrite firstn_removelast; auto with arith. rewrite <- list_to_set_in_iff in H0. rewrite <- list_to_set_in_iff. eapply in_nin_linds_occ_firstn_S. apply hc. assumption. apply h2. assumption. rewrite firstn_removelast; auto with arith. rewrite <- list_to_set_in_iff in H1. rewrite <- list_to_set_in_iff. eapply in_nin_linds_occ_firstn_S. apply hd. assumption. apply h2. assumption. assumption. erewrite nth_lt_removelast. erewrite nth_lt_removelast. apply H3. red; apply ba_eq_dec. red; apply ba_eq_dec. rewrite fin_map_seg_rest_compat; auto with arith. rewrite fin_map_seg_rest_compat; auto with arith. rewrite <- list_to_set_in_iff in H0. rewrite in_linds_occ_iff in H0. destruct H0 as [o [h6 [h7 h8]]]; subst. revert H3. Gen0. rewrite lind_occ_firstn. intros h8 h9. erewrite nth_lt_functional3 in h9. rewrite lind_occ_compat in h9. subst. contradict h2. erewrite nth_lt_functional3. apply h3. assumption. rewrite <- list_to_set_in_iff in H1. rewrite in_linds_occ_iff in H1. destruct H1 as [o [h6 [h7 h8]]]; subst. symmetry in H3. revert H3. Gen0. rewrite lind_occ_firstn. intros h8 h9. erewrite nth_lt_functional3 in h9. rewrite lind_occ_compat in h9. subst. contradict h2. erewrite nth_lt_functional3. apply h3. assumption. contradict H2; auto. Grab Existential Variables. assumption. assumption. Qed. Lemma has_comps_removelast_impl : forall {A:Bool_Alg} {l:list (bt A)} (m:nat) (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns) (x:bt A), let pfm := le_length_removelast l m pfsm in In x (firstn m l) -> ~ In (nth_lt l m pfsm) (firstn m l) -> has_comps (removelast l) m pfm (list_to_set (linds_occ ba_eq_dec (firstn m l) x)) (fin_map_seg_rest F m (lt_n_Sn m)) -> has_comps l (S m) pfsm (list_to_set (linds_occ ba_eq_dec (firstn (S m) l) x)) F. intros A l m h1 F x h2 hi hn h3. inversion h3; subst. econstructor. apply list_to_set_finite. rewrite <- list_to_set_in_iff in H0. rewrite <- list_to_set_in_iff. eapply in_linds_occ_firstn_le. apply h1. auto with arith. apply pfc. apply H0. rewrite <- list_to_set_in_iff in H1. rewrite <- list_to_set_in_iff. eapply in_linds_occ_firstn_le. apply h1. auto with arith. apply pfd. apply H1. assumption. do 2 erewrite nth_lt_removelast in H3. apply H3. red; apply ba_eq_dec. red; apply ba_eq_dec. red; apply ba_eq_dec. rewrite fin_map_seg_rest_compat in H4; auto. rewrite fin_map_seg_rest_compat in H4; auto. Grab Existential Variables. omega. auto with arith. auto with arith. Qed. Lemma has_comps_removelast_impl' : forall {A:Bool_Alg} (l:list (bt A)) (m:nat) (pfm:m <= length (removelast l)) (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns), has_comps (removelast l) m pfm (list_to_set (linds_occ ba_eq_dec (firstn m (removelast l)) (nth_lt l m pfsm))) (fin_map_seg_rest F m (lt_n_Sn m)) -> has_comps l (S m) pfsm (list_to_set (linds_occ ba_eq_dec (firstn (S m) l) (nth_lt l m pfsm))) F. intros A l m h1 h2 F h3. inversion h3. econstructor. apply list_to_set_finite. rewrite <- list_to_set_in_iff in H0. rewrite <- list_to_set_in_iff. eapply in_linds_occ_firstn_le. apply h2. auto with arith. apply pfc. rewrite firstn_removelast in H0; auto. rewrite <- list_to_set_in_iff in H1. rewrite <- list_to_set_in_iff. eapply in_linds_occ_firstn_le. apply h2. auto with arith. apply pfd. rewrite firstn_removelast in H1; auto. assumption. do 2 erewrite nth_lt_removelast in H3. apply H3. red; apply ba_eq_dec. red; apply ba_eq_dec. red; apply ba_eq_dec. do 2 rewrite fin_map_seg_rest_compat in H4; auto. Grab Existential Variables. omega. auto with arith. auto with arith. Qed. Lemma has_comps_removelast_impl'' : forall {A:Bool_Alg} (l:list (bt A)) (m:nat) (pfm:m <= length (removelast l)) (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns) (x:bt A), has_comps (removelast l) m pfm (list_to_set (linds_occ ba_eq_dec (firstn m (removelast l)) x)) (fin_map_seg_rest F m (lt_n_Sn m)) -> has_comps l (S m) pfsm (list_to_set (linds_occ ba_eq_dec (firstn (S m) l) x)) F. intros A l m h1 h2 F x h3. inversion h3. econstructor. apply list_to_set_finite. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in H0. eapply in_linds_occ_firstn_le. apply h2. auto with arith. apply pfc. rewrite firstn_removelast in H0; auto with arith. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in H1. eapply in_linds_occ_firstn_le. apply h2. auto with arith. apply pfd. rewrite firstn_removelast in H1; auto with arith. assumption. do 2 erewrite nth_lt_removelast in H3. apply H3. red; apply ba_eq_dec. red; apply ba_eq_dec. red; apply ba_eq_dec. do 2 rewrite fin_map_seg_rest_compat in H4; auto with arith. Grab Existential Variables. eapply lt_trans. apply pfc. apply h2. auto with arith. auto with arith. Qed. Lemma has_comps_removelast_iff : forall {A:Bool_Alg} (l:list (bt A)) (m:nat) (pfm: m <= length l) (pfm':m <= length (removelast l)) (x:bt A) (F:Fin_map (seg m) signe mns), has_comps (removelast l) m pfm' (list_to_set (linds_occ ba_eq_dec (firstn m l) x)) F <-> has_comps l m pfm (list_to_set (linds_occ ba_eq_dec (firstn m l) x)) F. intros A l m h1 h2 x F; split; intro h3. inversion h3. econstructor. apply list_to_set_finite. apply H0. apply H1. assumption. do 2 erewrite nth_lt_removelast in H3. apply H3. red; apply ba_eq_dec. red; apply ba_eq_dec. red; apply ba_eq_dec. assumption. inversion h3. econstructor. apply list_to_set_finite. apply H0. apply H1. assumption. erewrite nth_lt_removelast. erewrite nth_lt_removelast. apply H3. red; apply ba_eq_dec. red; apply ba_eq_dec. assumption. Grab Existential Variables. omega. assumption. omega. assumption. assumption. Qed. Lemma impl_has_comps_removelast_neq : forall {A:Bool_Alg} (l:list (bt A)) (m:nat) (pfm:m <= length (removelast l)) (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns) (x:bt A), x <> nth_lt l m pfsm -> has_comps l (S m) pfsm (list_to_set (linds_occ ba_eq_dec (firstn (S m) l) x)) F -> has_comps (removelast l) m pfm (list_to_set (linds_occ ba_eq_dec (firstn m (removelast l)) x)) (fin_map_seg_rest F m (lt_n_Sn m)). intros A l m h1 h2 F x h3 h4. inversion h4. econstructor. apply list_to_set_finite. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in H0. eapply in_nin_linds_occ_firstn_S' in H0. rewrite firstn_removelast. apply H0. assumption. assumption. apply h3. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in H1. eapply in_nin_linds_occ_firstn_S' in H1. rewrite firstn_removelast. apply H1. assumption. assumption. apply h3. assumption. erewrite nth_lt_removelast. erewrite nth_lt_removelast. apply H3. red; apply ba_eq_dec. red; apply ba_eq_dec. pose proof pfd as hd. apply le_S_n in hd. apply le_lt_eq_dec in hd. destruct hd as [|hd]; auto. pose proof pfc as hc. apply le_S_n in hc. apply le_lt_eq_dec in hc. destruct hc as [|hc]; auto. rewrite fin_map_seg_rest_compat. rewrite fin_map_seg_rest_compat; auto. auto. subst. rewrite <- list_to_set_in_iff in H1. rewrite in_linds_occ_iff in H1. destruct H1 as [o [h6 [h7 h8]]]; subst. symmetry in H3. revert H3. Gen0. rewrite lind_occ_firstn. intros h8 h9. pose proof (lind_occ_compat ba_eq_dec l o x (in_firstn l (S m) x h6)) as h10. unfold bt in h9, h10. erewrite nth_lt_functional3 in h10. rewrite h10 in h9 at 1. rewrite h9 in h3. contradict h3. apply nth_lt_functional3. assumption. subst. rewrite <- list_to_set_in_iff in H1. rewrite in_linds_occ_iff in H1. destruct H1 as [o [h6 [h7 h8]]]. rewrite lind_occ_firstn in h8; auto. revert h8. Gen0. intros h9 h10. subst. erewrite nth_lt_functional3 in h3. erewrite lind_occ_compat in h3. contradict h3; auto. Grab Existential Variables. pose proof pfd as hd. apply le_S_n in hd. apply le_lt_eq_dec in hd. destruct hd as [|hd]; auto. rewrite <- list_to_set_in_iff in H1. rewrite in_linds_occ_iff in H1. destruct H1 as [o [h6 [h7 h8]]]. rewrite lind_occ_firstn in h8; auto. revert h8. Gen0. intros h9 h10. subst. subst. contradict h3. erewrite nth_lt_functional3. erewrite lind_occ_compat at 1; auto. pose proof pfc as hc. apply le_S_n in hc. apply le_lt_eq_dec in hc. destruct hc as [|hc]; auto. subst. rewrite <- list_to_set_in_iff in H1. rewrite in_linds_occ_iff in H1. destruct H1 as [o [h6 [h7 h8]]]. rewrite lind_occ_firstn in h8; auto. revert h8. Gen0. intros h9 h10. subst. subst. symmetry in H3. erewrite nth_lt_functional3 in H3. unfold bt in H3. erewrite lind_occ_compat in H3 at 1; auto. rewrite H3 in h3. contradict h3. apply nth_lt_functional3. Qed. Lemma has_comps_S_linds_occ_in : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns), has_comps l (S m) pfsm (list_to_set (linds_occ ba_eq_dec (firstn (S m) l) (nth_lt l m pfsm))) F -> In (nth_lt l m pfsm) (firstn m l). intros B l m h1 F h2. inversion h2. rewrite <- list_to_set_in_iff in H0. rewrite <- list_to_set_in_iff in H1. rewrite in_linds_occ_iff in H0, H1. destruct H0 as [o [h5 [h6 h7]]], H1 as [o' [h5' [h6' h7']]]; subst. pose proof (lt_lind_occ ba_eq_dec (firstn (S m) l) o (nth_lt l m h1) h5) as h8. pose proof (lt_lind_occ ba_eq_dec (firstn (S m) l) o' (nth_lt l m h1) h5) as h9. rewrite length_firstn in h8, h9; auto. red in h8, h9. apply le_S_n in h8. apply le_S_n in h9. apply le_lt_eq_dec in h8. apply le_lt_eq_dec in h9. destruct h8 as [h8 | h8]. rewrite <- (length_firstn l) in h8. pose proof (in_nth_lt _ _ h8) as h10. erewrite nth_lt_functional3 in h10. unfold bt in h10. pose proof (lind_occ_compat ba_eq_dec (firstn (S m) l) o (nth_lt l m h1) h5) as h11. rewrite nth_lt_firstn_eq in h10, h11. unfold bt in h10, h11. rewrite <- h11 at 1. unfold bt. erewrite nth_lt_functional3. apply h10. auto with arith. destruct h9 as [h9 | h9]. rewrite <- (length_firstn l) in h9. pose proof (in_nth_lt _ _ h9) as h10. unfold bt in h10. erewrite nth_lt_functional3 in h10. unfold bt in h10. pose proof (lind_occ_compat ba_eq_dec (firstn (S m) l) o' (nth_lt l m h1) h5) as h11. rewrite nth_lt_firstn_eq in h10, h11. unfold bt in h10, h11. rewrite <- h11 at 1. unfold bt. unfold bt in h10. erewrite nth_lt_functional3. apply h10. auto with arith. pose proof (proof_irrelevance _ h5 h5'). contradict H2. unfold bt in h8, h9. unfold bt. congruence. Grab Existential Variables. assumption. assumption. Qed. Inductive only_has_comps {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) : Prop := | only_has_comps_intro : Finite P -> (forall (c d:nat) (pfc:c < m) (pfd:d < m), Inn P c -> Inn P d -> c <> d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = nth_lt l d (lt_le_trans _ _ _ pfd pfm) -> a |-> c <> a |-> d) -> only_has_comps l m pfm P a. Definition only_has_comps' {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) : Prop := only_has_comps l (length l) (le_refl _) P a. Inductive only_has_sames {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) : Prop := | only_has_sames_intro : Finite P -> (forall (c d:nat) (pfc:c < m) (pfd:d < m), Inn P c -> Inn P d -> c <> d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = nth_lt l d (lt_le_trans _ _ _ pfd pfm) -> a |-> c = a |-> d) -> only_has_sames l m pfm P a. Definition only_has_sames' {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) : Prop := only_has_sames l (length l) (le_refl _) P a. Lemma only_has_comps_compat : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), Finite P -> (only_has_comps l m pfm P a <-> ~has_sames l m pfm P a). intros B l m h1 P a hf; split; intro h2. intro h3. destruct h2, h3. eapply H0 in H6; auto. apply H5. econstructor. assumption. intros c d h3 h4 h5 h6 h7 h8 h9. contradict h2. econstructor. assumption. apply h5. apply h6. assumption. apply h8. assumption. Qed. Lemma only_has_comps_compat' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), Finite P -> (only_has_comps' l P a <-> ~has_sames' l P a). unfold only_has_comps', has_sames'; intros; apply only_has_comps_compat; auto. Qed. Lemma only_has_sames_compat : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), Finite P -> (only_has_sames l m pfm P a <-> ~has_comps l m pfm P a). intros B l m h1 P a hf; split; intro h2. intro h3. destruct h2, h3. contradict H6. eapply H0. assumption. assumption. assumption. apply H5. econstructor. assumption. intros c d h3 h4 h5 h6 h7 h8. destruct (sign_dec (a |-> c) (a |-> d)) as [|h10]; auto. contradict h2. econstructor. assumption. apply h5. apply h6. assumption. apply h8. assumption. Qed. Lemma only_has_sames_compat' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), Finite P -> (only_has_sames' l P a <-> ~has_comps' l P a). unfold only_has_comps', has_sames'; intros; apply only_has_sames_compat; auto. Qed. Lemma only_has_comps_finite : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), only_has_comps l m pfm P a -> Finite P. intros; destruct H; auto. Qed. Lemma only_has_comps_finite' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), only_has_comps' l P a -> Finite P. unfold only_has_comps'; intros; eapply only_has_comps_finite; apply H. Qed. Lemma only_has_sames_finite : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), only_has_sames l m pfm P a -> Finite P. intros; destruct H; auto. Qed. Lemma only_has_sames_finite' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), only_has_sames' l P a -> Finite P. unfold only_has_sames'; intros; eapply only_has_sames_finite; apply H. Qed. Lemma only_has_sames_dec : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), Finite P -> {only_has_sames l m pfm P a} + {has_comps l m pfm P a}. intros B l m h1 P a hf. destruct (has_comps_dec h1 a P hf) as [h2 | h2]. right; auto. left. rewrite only_has_sames_compat; auto. Qed. Lemma only_has_sames_dec' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), Finite P -> {only_has_sames' l P a} + {has_comps' l P a}. unfold only_has_sames'; intros; apply only_has_sames_dec; auto. Qed. Lemma only_has_comps_dec : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), Finite P -> {only_has_comps l m pfm P a} + {has_sames l m pfm P a}. intros B l m h1 P a hf. destruct (has_sames_dec h1 a P hf) as [h2 | h2]. right; auto. left. rewrite only_has_comps_compat; auto. Qed. Lemma only_has_comps_dec' : forall {B:Bool_Alg} (l:list (bt B)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), Finite P -> {only_has_comps' l P a} + {has_sames' l P a}. unfold only_has_comps'; intros; apply only_has_comps_dec; auto. Qed. Lemma linds_eq_determines_has_comps : forall {A B:Bool_Alg} (l:list (bt B)) (l':list (bt A)) (P:Ensemble nat) (m:nat) (pfm:m <= length l) (pfe:length l = length l') (a:Fin_map (seg m) signe mns), Finite P -> let pfm' := le_eq_trans _ _ _ pfm pfe in linds ba_eq_dec (firstn m l) = linds ba_eq_dec (firstn m l') -> (has_comps l m pfm P a <-> has_comps l' m pfm' P a). intros A B l l' P m h1 h2 a hf h3 h4; split; intro h5. destruct h5 as [c d h5 h6 h7 h8 h9 h10 h11 h12]. econstructor. assumption. apply h8. apply h9. assumption. pose proof (nth_lt_firstn_eq l c m) as hnc. hfirst hnc. rewrite length_firstn; auto. specialize (hnc hf0). pose proof (nth_lt_firstn_eq l d m) as hnd. hfirst hnd. rewrite length_firstn; auto. specialize (hnd hf1). erewrite nth_lt_functional3 in h11. rewrite <- hnc in h11 at 1. symmetry in h11. erewrite nth_lt_functional3 in h11. rewrite <- hnd in h11 at 1. symmetry in h11. erewrite linds_nth_lt_compat_iff in h11. erewrite nth_lt_functional2. erewrite <- (nth_lt_firstn_eq l' c m). symmetry. erewrite nth_lt_functional2. erewrite <- (nth_lt_firstn_eq l' d m). symmetry. apply h11. reflexivity. reflexivity. assumption. destruct h5 as [c d h5 h6 h7 h8 h9 h10 h11 h12]. econstructor. assumption. apply h8. apply h9. assumption. pose proof (nth_lt_firstn_eq l' c m) as hnc. hfirst hnc. rewrite length_firstn; auto. specialize (hnc hf0). pose proof (nth_lt_firstn_eq l' d m) as hnd. hfirst hnd. rewrite length_firstn; auto. specialize (hnd hf1). erewrite nth_lt_functional3 in h11. rewrite <- hnc in h11 at 1. symmetry in h11. erewrite nth_lt_functional3 in h11. rewrite <- hnd in h11 at 1. symmetry in h11. erewrite linds_nth_lt_compat_iff in h11. erewrite nth_lt_functional2. erewrite <- (nth_lt_firstn_eq l c m). symmetry. erewrite nth_lt_functional2. erewrite <- (nth_lt_firstn_eq l d m). symmetry. apply h11. reflexivity. reflexivity. assumption. Grab Existential Variables. symmetry. apply h4. assumption. assumption. apply h4. assumption. assumption. Qed. Lemma linds_eq_determines_has_sames : forall {A B:Bool_Alg} (l:list (bt A)) (l':list (bt B)) (P:Ensemble nat) (m:nat) (pfm:m <= length l) (pfe:length l = length l') (a:Fin_map (seg m) signe mns), Finite P -> let pfm' := le_eq_trans _ _ _ pfm pfe in linds ba_eq_dec (firstn m l) = linds ba_eq_dec (firstn m l') -> (has_sames l m pfm P a <-> has_sames l' m pfm' P a). intros A B l l' P m h1 h2 a hf h3 h4;split; intro h5. destruct h5 as [c d h5 h6 h7 h8 h9 h10 h11 h12]. econstructor. assumption. apply h8. apply h9. assumption. pose proof (nth_lt_firstn_eq l c m) as hnc. hfirst hnc. rewrite length_firstn; auto. specialize (hnc hf0). pose proof (nth_lt_firstn_eq l d m) as hnd. hfirst hnd. rewrite length_firstn; auto. specialize (hnd hf1). erewrite nth_lt_functional3 in h11. rewrite <- hnc in h11 at 1. symmetry in h11. erewrite nth_lt_functional3 in h11. rewrite <- hnd in h11 at 1. symmetry in h11. erewrite linds_nth_lt_compat_iff in h11. erewrite nth_lt_functional2. erewrite <- (nth_lt_firstn_eq l' c m). symmetry. erewrite nth_lt_functional2. erewrite <- (nth_lt_firstn_eq l' d m). symmetry. apply h11. reflexivity. reflexivity. assumption. destruct h5 as [c d h5 h6 h7 h8 h9 h10 h11 h12]. econstructor. assumption. apply h8. apply h9. assumption. pose proof (nth_lt_firstn_eq l' c m) as hnc. hfirst hnc. rewrite length_firstn; auto. specialize (hnc hf0). pose proof (nth_lt_firstn_eq l' d m) as hnd. hfirst hnd. rewrite length_firstn; auto. specialize (hnd hf1). erewrite nth_lt_functional3 in h11. rewrite <- hnc in h11 at 1. symmetry in h11. erewrite nth_lt_functional3 in h11. rewrite <- hnd in h11 at 1. symmetry in h11. erewrite linds_nth_lt_compat_iff in h11. erewrite nth_lt_functional2. erewrite <- (nth_lt_firstn_eq l c m). symmetry. erewrite nth_lt_functional2. erewrite <- (nth_lt_firstn_eq l d m). symmetry. apply h11. reflexivity. reflexivity. assumption. Grab Existential Variables. symmetry. apply h4. assumption. assumption. apply h4. assumption. assumption. Qed. Lemma linds_eq_determines_has_comps' : forall {A B:Bool_Alg} (l:list (bt B)) (l':list (bt A)) (P:Ensemble nat) (pfe:length l = length l') (a:Fin_map (seg (length l)) signe mns), Finite P -> let a' := fin_map_seg_transfer a pfe in linds ba_eq_dec l = linds ba_eq_dec l' -> (has_comps' l P a <-> has_comps' l' P a'). unfold has_comps'; intros A B l l' P h1 a hf h2. pose proof (linds_eq_determines_has_comps l l' P (length l) (le_refl _) h1 a hf) as h3. simpl in h3. hfirst h3. rewrite firstn_length, h1, firstn_length; auto. specialize (h3 hf0). rewrite h3. split; intro h4. eapply has_comps_transfer in h4. apply h4. eapply has_comps_transfer_rev. apply h4. Qed. Lemma linds_eq_determines_has_sames' : forall {A B:Bool_Alg} (l:list (bt B)) (l':list (bt A)) (P:Ensemble nat) (pfe:length l = length l') (a:Fin_map (seg (length l)) signe mns), Finite P -> let a' := fin_map_seg_transfer a pfe in linds ba_eq_dec l = linds ba_eq_dec l' -> (has_sames' l P a <-> has_sames' l' P a'). unfold has_sames'; intros A B l l' P h1 a hf h2. pose proof (linds_eq_determines_has_sames l l' P (length l) (le_refl _) h1 a hf) as h3. simpl in h3. hfirst h3. rewrite firstn_length, h1, firstn_length; auto. specialize (h3 hf0). rewrite h3. split; intro h4. eapply has_sames_transfer in h4. apply h4. eapply has_sames_transfer_rev. apply h4. Qed. Lemma only_has_sames_map_seg_to_fin_map_ex : forall {B:Bool_Alg} {n:nat} (f:nat->bt B) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns), only_has_sames' _ (seg n) F -> exists! F':Fin_map (Im (seg n) f) signe mns, forall i:nat, i < n -> F' |-> (f i) = F |-> i. intros B n f F h1. red in h1. destruct h1 as [h1 h2]. pose (fun x:sig_set (Im (seg n) f) => let pfi := proj2_sig x in let pfi0 := in_eq_set _ _ (eq_sym (list_to_set_map_seg_ f n)) _ pfi in let pfi1 := iff2 (list_to_set_in_iff _ _) pfi0 in F |-> lind ba_eq_dec _ _ pfi1) as fF. pose (sig_fun_to_fin_map_ran ba_eq_dec fF (finite_image _ _ _ _ (finite_seg _)) mns signe signe_finite) as F'. hfirst F'. red; intros x h4; destruct h4 as [x h4]; destruct (fF x); subst. left. right. exists (F' hf). red. split. intros i h3. unfold F'. erewrite sig_fun_to_fin_map_ran_compat. unfold fF. simpl. match goal with |- F |-> lind _ _ _ ?pf = _ => generalize pf end. intro h4. destruct (nat_eq_dec i (lind ba_eq_dec _ _ h4)) as [h5|h5]. rewrite <- h5; auto. apply neq_sym in h5. eapply h2. constructor. simpl. eapply lt_eq_trans. apply lt_lind. apply length_map_seg. constructor; auto. simpl. assumption. do 2 rewrite nth_lt_map_seg'. do 2 rewrite fun_to_seg_compat. apply nat_fun_lind_map_seg'; auto. intros F0 h3. apply fin_map_ext. intros y h4. destruct h4 as [m h4]; subst. destruct h4 as [h4]. specialize (h3 _ h4). rewrite h3. unfold F'. erewrite sig_fun_to_fin_map_ran_compat. unfold fF. simpl. match goal with |- F |-> ?x' = _ => destruct (nat_eq_dec m x') as [h5 | h5] end. rewrite <- h5; auto. symmetry. eapply h2. constructor; auto. constructor. eapply lt_eq_trans. apply lt_lind. apply length_map_seg. assumption. do 2 rewrite nth_lt_map_seg'. do 2 rewrite fun_to_seg_compat. symmetry. apply nat_fun_lind_map_seg'; auto. Grab Existential Variables. apply lt_lind. rewrite length_map_seg; auto. apply Im_intro with m. constructor; auto. reflexivity. rewrite length_map_seg; auto. apply lt_lind. apply Im_intro with i. constructor; auto. reflexivity. Qed. Definition only_has_sames_map_seg_to_fin_map {B:Bool_Alg} {n:nat} (f:nat->bt B) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pf:only_has_sames' _ (seg n) F) : Fin_map (Im (seg n) f) signe mns := proj1_sig (constructive_definite_description _ (only_has_sames_map_seg_to_fin_map_ex f F pf)). Lemma only_has_sames_map_seg_to_fin_map_compat : forall {B:Bool_Alg} {n:nat} (f:nat->bt B) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pf:only_has_sames' _ (seg n) F), let F' := only_has_sames_map_seg_to_fin_map f F pf in forall i:nat, i < n -> F' |-> (f i) = F |-> i. simpl; unfold only_has_sames_map_seg_to_fin_map; intros; destruct constructive_definite_description; auto. Qed. Lemma fin_map_im_seg_to_seg_length_ex : forall {B:Bool_Alg} {n:nat} (f:nat->bt B) (F:Fin_map (Im (seg n) f) signe mns), exists! F':Fin_map (seg (length (map_seg (fun i (_:i < n) => f i)))) signe mns, forall i (pf:i < n), F' |-> i = F |-> f i. intros B n f F. pose (fun y:sig_set (seg (length (map_seg (fun (i : nat) (_ : i < n) => f i)))) => F |-> f (proj1_sig y)) as Ff. pose (sig_fun_to_fin_map_ran nat_eq_dec Ff (finite_seg (length (map_seg (fun i (_:i < n) => f i)))) mns _ signe_finite) as F'. hfirst F'. red; intros x ?; destruct x. left. right. exists (F' hf). red. split. intros i h1. unfold F'. erewrite sig_fun_to_fin_map_ran_compat. unfold Ff; simpl; auto. intros G h1. unfold F'. apply fin_map_ext. intros i h2. destruct h2 as [h2]. rewrite length_map_seg in h2. erewrite sig_fun_to_fin_map_ran_compat. unfold Ff. simpl. symmetry. apply h1; auto. Grab Existential Variables. constructor. rewrite length_map_seg; auto. constructor. rewrite length_map_seg; auto. Qed. Definition fin_map_im_seg_to_seg_length {B:Bool_Alg} {n:nat} (f:nat->bt B) (F:Fin_map (Im (seg n) f) signe mns) : Fin_map (seg (length (map_seg (fun i (_:i < n) => f i)))) signe mns := proj1_sig (constructive_definite_description _ (fin_map_im_seg_to_seg_length_ex f F)). Lemma fin_map_im_seg_to_seg_length_compat : forall {B:Bool_Alg} {n:nat} (f:nat->bt B) (F:Fin_map (Im (seg n) f) signe mns), let F' := fin_map_im_seg_to_seg_length f F in forall i (pf:i < n), F' |-> i = F |-> f i. simpl; unfold fin_map_im_seg_to_seg_length; intros; destruct constructive_definite_description; auto. Qed. End HasComps. Section ElProdL. Variable A:Bool_Alg. Let At := Btype (Bc A). (*This is more "natural" than the eventual versions. and will be used in the list version of the _countable_ monomorphism (resp. homomorphism) extension criterion.*) Definition el_prod_l {l:list At} (F:Fin_map (seg (length l)) signe mns) := times_list (map_seg (fun i (pf:i < length l) => eps (nth_lt l i pf) (F |-> i))). Definition el_prod_l_compose {B:Bool_Alg} {l:list (bt B)} (F:Fin_map (seg (length l)) signe mns) (g:bt B->At):= times_list (map_seg (fun i (pf:i < length l) => eps (g (nth_lt l i pf)) (F |-> i))). (*This is more "specialized" and will be used in the list version of the _finite_ monomorphism (resp. homomorphism) extension criterion.*) Definition el_prod_l_firstn {l:list At} {m:nat} (pf:m <= length l) (F:Fin_map (seg m) signe mns) := let F' := fin_map_seg_transfer F (eq_sym (length_firstn l m pf)) in el_prod_l F'. Lemma el_prod_l_to_firstn : forall {l:list At} (F:Fin_map (seg (length l)) signe mns), el_prod_l F = el_prod_l_firstn (le_refl _) F. unfold el_prod_l_firstn, el_prod_l; intros. f_equal. pose proof (length_firstn l (length l) (le_refl _)) as h1. symmetry in h1. pose proof (map_seg_functional h1 (fun (i : nat) (pf : i < length l) => eps (nth_lt l i pf) (F |-> i))) as h2. rewrite h2. apply map_seg_ext. intros i h3. rewrite transfer_seg_fun_compat. repeat f_equal. rewrite nth_lt_firstn_eq. apply nth_lt_functional3. rewrite fin_map_seg_transfer_compat'; auto. Qed. Lemma el_prod_l0 : forall (F:Fin_map (seg (length nil)) signe mns), el_prod_l F = bone A. intro; unfold el_prod_l; auto. Qed. Lemma el_prod_l_compose0 : forall {B:Bool_Alg} (F:Fin_map (seg (length nil)) signe mns) (g:bt B->At), el_prod_l_compose F g = bone A. intro; unfold el_prod_l; auto. Qed. Lemma el_prod_l_firstn0 : forall (l:list At) (F:Fin_map (seg 0) signe mns), el_prod_l_firstn (le_0_n (length l)) F = bone _. intros; unfold el_prod_l_firstn, el_prod_l; auto. Qed. Lemma has_comps_el_prod_l_firstn : forall (l:list At) (m:nat) (pfle:m <= length l) (F:Fin_map (seg m) signe mns), has_comps l m pfle (seg m) F -> el_prod_l_firstn pfle F = 0. intros l m h1 F h2. destruct h2. unfold el_prod_l_firstn, el_prod_l. apply (times_list_opps _ _ (eps (nth_lt l c (lt_le_trans c m (length l) pfc h1)) (F |-> c)) (eps (nth_lt l d (lt_le_trans d m (length l) pfd h1)) (F |-> d))). rewrite in_map_seg_iff. exists c. ex_goal. rewrite length_firstn; auto. exists hex. rewrite nth_lt_firstn_eq. f_equal. apply nth_lt_functional3. rewrite fin_map_seg_transfer_compat'; auto. rewrite in_map_seg_iff. exists d. ex_goal. rewrite length_firstn; auto. exists hex. rewrite nth_lt_firstn_eq. f_equal. apply nth_lt_functional3. rewrite fin_map_seg_transfer_compat'; auto. destruct (F |-> c), (F |-> d); simpl. contradict H4; auto. rewrite <- doub_neg; auto. f_equal; auto. contradict H4; auto. Qed. Lemma el_prod_l_firstn_0 : forall {l:list At} {m:nat} (pfm: m <= length l) (F:Fin_map (seg m) signe mns) (x:At), In x (firstn m l) -> let P := list_to_set (linds_occ ba_eq_dec (firstn m l) x) in has_comps _ _ pfm P F -> el_prod_l_firstn pfm F = 0. intros l m h1 F x h2 P h3. destruct h3. unfold el_prod_l_firstn, el_prod_l. apply (times_list_opps _ _ (eps (nth_lt l c (lt_le_trans c m (length l) pfc h1)) (F |-> c)) (eps (nth_lt l d (lt_le_trans d m (length l) pfd h1)) (F |-> d))). rewrite in_map_seg_iff. exists c. ex_goal. rewrite length_firstn; auto. exists hex. rewrite nth_lt_firstn_eq. f_equal. apply nth_lt_functional3. rewrite fin_map_seg_transfer_compat'; auto. rewrite in_map_seg_iff. exists d. ex_goal. rewrite length_firstn; auto. exists hex. rewrite nth_lt_firstn_eq. f_equal. apply nth_lt_functional3. rewrite fin_map_seg_transfer_compat'; auto. unfold eps. destruct (F |-> c), (F |-> d). contradict H4; auto. rewrite <- doub_neg; auto. f_equal; auto. contradict H4; auto. Qed. Lemma el_prod_l_0 : forall {l:list At} (F:Fin_map (seg (length l)) signe mns) (x:At), In x l -> let P := list_to_set (linds_occ ba_eq_dec l x) in has_comps' _ P F -> el_prod_l F = 0. intros; rewrite el_prod_l_to_firstn; eapply el_prod_l_firstn_0. rewrite firstn_length. apply H. unfold has_comps' in H0. unfold P in H0. rewrite firstn_length; auto. Qed. Lemma has_comps_el_prod_l0 : forall (l:list At) (F:Fin_map (seg (length l)) signe mns), has_comps' l (seg (length l)) F -> el_prod_l F = 0. intros l F h1. inversion h1. unfold el_prod_l. pose proof (in_nth_lt _ _ pfc) as h2. match goal with |- times_list (map_seg ?f') = 0 => pose f' as f end. pose proof (in_map_seg_iff f (eps (nth_lt l c pfc) (F |-> c))) as h6. pose proof (in_map_seg_iff f (eps (nth_lt l d pfd) (F |-> d))) as h7. hr h6. exists c, pfc. unfold f; auto. hr h7. exists d, pfd. unfold f; auto. rewrite <- h6 in hr. rewrite <- h7 in hr0. eapply times_list_opps. apply hr. apply hr0. unfold eps. unfold At, bt. unfold bt in H4. let P := type of H4 in match P with ?fc' <> _ => destruct fc' end. let P := type of H4 in match P with _ <> ?fd' => destruct fd' end. contradict H4; auto. rewrite <- doub_neg. erewrite nth_lt_functional3. rewrite H3 at 1. apply nth_lt_functional3. let P := type of H4 in match P with _ <> ?fd' => destruct fd' end. f_equal. erewrite nth_lt_functional3. rewrite H3 at 1. apply nth_lt_functional3. contradict H4; auto. Qed. Lemma el_prod_l_firstn_S : forall (l:list At) (m:nat) (pf:S m <= length l) (F:Fin_map (seg (S m)) signe mns), let pf' := lt_le_trans _ _ _ (lt_n_Sn m) pf in let pf'' := lt_le_weak _ _ pf' in el_prod_l_firstn pf F = (eps (nth_lt l m pf') (F |-> m)) * (el_prod_l_firstn pf'' (fin_map_seg_rest F _ (lt_n_Sn m))). intros l m h1 F pf'. unfold el_prod_l_firstn, el_prod_l. pose proof (length_firstn_S_eq _ _ pf') as h2. rewrite (map_seg_functional h2); simpl. rewrite transfer_seg_fun_compat. rewrite times_list_app. rewrite comm_prod. simpl. f_equal. rewrite (nth_lt_firstn_eq l _ _ (transfer_dep_r h2 (lt_n_Sn (length (firstn m l))))) at 1. rewrite one_prod. f_equal. apply nth_lt_functional2. rewrite length_firstn; auto with arith. rewrite length_firstn; auto with arith. pose proof (fin_map_seg_transfer_compat F ( (eq_sym (length_firstn l (S m) h1)))) as h3. assert (h4:m < length (firstn (S m) l)). rewrite length_firstn; auto. specialize (h3 (exist (fun c=> c < length (firstn (S m) l)) _ h4)). simpl in h3. rewrite h3. f_equal. rewrite segT_transfer_r_compat at 1. simpl; auto. f_equal. f_equal. apply functional_extensionality_dep. intro c. apply functional_extensionality. intro h3. unfold seg_fun_from_S. rewrite transfer_seg_fun_compat. f_equal. rewrite nth_lt_firstn_eq. rewrite (nth_lt_firstn_eq l c (S m)). f_equal. apply proof_irrelevance. rewrite fin_map_seg_transfer_compat'; auto. rewrite fin_map_seg_transfer_compat'; auto. rewrite fin_map_seg_rest_compat; auto. rewrite length_firstn in h3; auto. pose proof (length_firstn l (S m) h1) as h4. unfold At in h4. unfold At. auto with arith. rewrite length_firstn in h3; auto with arith. destruct l. omega. simpl. rewrite length_firstn; auto with arith. Qed. Lemma el_prod_l_firstn_removelast_compat : forall {l:list At} {m:nat} (pfm: m <= length l) (pfm':m <= length (removelast l)) (F:Fin_map (seg (S m)) signe mns), el_prod_l_firstn pfm (fin_map_seg_rest F m (lt_n_Sn _)) = el_prod_l_firstn pfm' (fin_map_seg_rest F m (lt_n_Sn _)). intros l m h1 h2 F. unfold el_prod_l_firstn, el_prod_l. f_equal. rewrite map_seg_eq_iff. ex_goal. rewrite length_firstn, length_map_seg, length_firstn; auto with arith. exists hex. intros c h3 h4. symmetry. erewrite nth_lt_functional3. rewrite nth_lt_map_seg at 1. f_equal. gen0. erewrite firstn_removelast. intros; apply nth_lt_functional3. assumption. rewrite fin_map_seg_transfer_compat'. rewrite fin_map_seg_transfer_compat'. reflexivity. assumption. pose proof h3 as h3'. rewrite length_firstn in h3'; auto with arith. rewrite length_firstn; auto with arith. Grab Existential Variables. pose proof h3 as h3'. rewrite length_firstn in h3'; auto with arith. rewrite length_firstn; auto with arith. Qed. Lemma in_last_el_prod_l_firstn_dec : forall {l:list At} {m:nat} (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns), let pfm := le_trans _ _ _ (le_n_Sn _) pfsm in let x := nth_lt l m pfsm in let P := list_to_set (linds_occ ba_eq_dec (firstn (S m) l) x) in In x (firstn m l) -> Finite P -> {has_comps_elt _ _ pfsm P (nth_lt l m pfsm) F /\ el_prod_l_firstn pfsm F = 0} + {~has_comps_elt _ _ pfsm P (nth_lt l m pfsm) F /\ el_prod_l_firstn pfsm F = el_prod_l_firstn pfm (fin_map_seg_rest F _ (lt_n_Sn _))}. intros l m h1 F h2 x P h3 hf. destruct (has_comps_elt_dec l (S m) h1 P (nth_lt l m h1) F hf) as [h4 | h4]. left. split; auto. inversion h4. unfold el_prod_l_firstn, el_prod_l. apply (times_list_opps _ (map_seg (fun (i : nat) (pf : i < length (firstn (S m) l)) => eps (nth_lt (firstn (S m) l) i pf) (fin_map_seg_transfer F (eq_sym (length_firstn l (S m) h1)) |-> i))) (eps (nth_lt l c (lt_le_trans c (S m) (length l) pfc h1)) (F |-> c)) (eps (nth_lt l d (lt_le_trans d (S m) (length l) pfd h1)) (F |-> d))). rewrite in_map_seg_iff. exists c. ex_goal. rewrite length_firstn; auto. exists hex. f_equal. rewrite nth_lt_firstn_eq. apply nth_lt_functional3. rewrite fin_map_seg_transfer_compat'; auto. rewrite in_map_seg_iff. exists d. ex_goal. rewrite length_firstn; auto. exists hex. f_equal. rewrite nth_lt_firstn_eq. apply nth_lt_functional3. rewrite fin_map_seg_transfer_compat'; auto. unfold eps. destruct (F |-> c), (F |-> d). contradict H5; auto. rewrite <- doub_neg. unfold At in H3, H4. unfold At. erewrite nth_lt_functional3. rewrite H3 at 1. rewrite <- H4; auto. f_equal. unfold At in H3, H4. unfold At. erewrite nth_lt_functional3. rewrite H3 at 1. rewrite <- H4; auto. contradict H5; auto. right. split; auto. unfold x in h3, P. clear x. pose proof (not_has_comps_elt_const h1 P _ F hf h4 (lind ba_eq_dec _ _ h3) m) as h6. hfirst h6. eapply lt_trans. apply lt_lind. rewrite length_firstn; auto with arith. specialize (h6 hf0 (lt_n_Sn _)). simpl in h6. hfirst h6. intro h7. pose proof (lt_lind ba_eq_dec (firstn m l) _ h3) as h8. rewrite length_firstn in h8; auto with arith. omega. specialize (h6 hf1). unfold P in h6. hfirst h6. rewrite <- list_to_set_in_iff. rewrite in_linds_occ_iff. exists 0%nat. ex_goal. pose proof h3 as h3'. rewrite in_firstn_iff in h3'. destruct h3' as [c [h8 h9]]. subst. unfold At. pose proof (in_firstn_iff l (nth_lt l m h1) (S m) h1) as hi. unfold At in hi. rewrite hi. exists m. exists (lt_n_Sn _). apply nth_lt_functional3. exists hex. split. pose proof (count_occ_In ba_eq_dec (firstn (S m) l) (nth_lt l m h1)) as h7. rewrite h7 in hex. auto with arith. unfold lind. rewrite lind_occ_firstn. rewrite lind_occ_firstn. apply lind_occ_functional; auto. pose proof (count_occ_In ba_eq_dec (firstn m l) (nth_lt l m h1)) as h7. unfold At in h7, h3. pose proof h3 as h3'. rewrite h7 in h3'. auto with arith. pose proof (count_occ_In ba_eq_dec (firstn (S m) l) (nth_lt l m h1)) as h7. rewrite h7 in hex. auto with arith. specialize (h6 hf2). hfirst h6. rewrite <- list_to_set_in_iff. rewrite in_linds_occ_iff. pose proof h3 as h3'. unfold At in h3'. rewrite in_firstn_iff in h3'. exists (pred (count_occ ba_eq_dec (firstn (S m) l) (nth_lt l m h1))). ex_goal. destruct h3' as [c [h8 h9]]. subst. unfold At. pose proof (in_firstn_iff l (nth_lt l m h1) (S m) h1) as hi. unfold At in hi. rewrite hi. exists m. exists (lt_n_Sn _). apply nth_lt_functional3. exists hex. split. rewrite (count_occ_In ba_eq_dec) in hex. omega. apply lind_occ_firstn_pred_count_occ_nth_lt. specialize (h6 hf3). pose proof (lind_firstn ba_eq_dec l (nth_lt l m h1) m h3) as h7. simpl in h7. unfold At, bt in h7. hfirst h6. gen0. unfold At, bt. rewrite h7. erewrite lind_functional. intro h8. erewrite nth_lt_functional3. erewrite lind_compat at 1. apply nth_lt_functional3. apply nth_lt_functional3. specialize (h6 hf4). hfirst h6. apply nth_lt_functional3. specialize (h6 hf5). rewrite el_prod_l_firstn_S. unfold el_prod_l_firstn, el_prod_l. rewrite times_list_times_in. f_equal. apply map_seg_ext. intros d h8. f_equal. rewrite fin_map_seg_transfer_compat'. rewrite fin_map_seg_transfer_compat'. reflexivity. assumption. assumption. rewrite in_map_seg_iff. exists (lind ba_eq_dec (firstn m l) (nth_lt l m h1) h3). ex_goal. apply lt_lind. exists hex. f_equal. symmetry. erewrite <- nth_lt_functional3. rewrite <- hf5 at 1. rewrite nth_lt_firstn_eq. symmetry. erewrite nth_lt_functional3. unfold At, bt in hf4. unfold At, bt. rewrite hf4 at 1. apply nth_lt_functional3. rewrite fin_map_seg_transfer_compat'. rewrite fin_map_seg_rest_compat. assumption. rewrite length_firstn in hex; auto with arith. assumption. Grab Existential Variables. assumption. apply in_nth_lt. auto with arith. assumption. Qed. Definition fun_for_el_prod_l_firstn {la:list At} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (x:At) (pf:In x (firstn m la)) : At := if (has_comps_dec pfm F (list_to_set (linds_occ ba_eq_dec (firstn m la) x)) (list_to_set_finite _)) then bzero _ else eps x (F |-> lind ba_eq_dec _ _ pf). Definition fun_for_el_prod_l {la:list At} (F:Fin_map (seg (length la)) signe mns) (x:At) (pf:In x la) : At := if (has_comps_dec' F (list_to_set (linds_occ ba_eq_dec la x)) (list_to_set_finite _)) then bzero _ else eps x (F |-> lind ba_eq_dec _ _ pf). Lemma fun_for_el_prod_l_to_firstn : forall {la:list At} (F:Fin_map (seg (length la)) signe mns) (x:At) (pf:In x la), fun_for_el_prod_l F x pf = fun_for_el_prod_l_firstn (le_refl _) F x (in_subst_l _ _ _ (eq_sym (firstn_length _)) pf). intros la F x h1. unfold fun_for_el_prod_l, fun_for_el_prod_l_firstn. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. reflexivity. unfold has_comps' in hl. rewrite firstn_length in hr'. contradiction. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. rewrite firstn_length in hl'. unfold has_comps' in hr. contradiction. repeat f_equal. symmetry. gen0. rewrite firstn_length. intros; repeat f_equal. apply proof_irrelevance. Qed. Lemma fun_for_el_prod_l_firstn_functional : forall {la la':list At} {m:nat} (pfm:m <= length la) (pfm':m <= length la') (F:Fin_map (seg m) signe mns) (x x':At) (pf:In x (firstn m la)) (pf':In x' (firstn m la')), x = x' -> la = la' -> fun_for_el_prod_l_firstn pfm F x pf = fun_for_el_prod_l_firstn pfm' F x' pf'. intros; subst; repeat f_equal; try apply proof_irrelevance. Qed. Lemma fun_for_el_prod_l_functional : forall {la:list At} (F:Fin_map (seg (length la)) signe mns) (x:At) (pf:In x la) (x':At) (pf':In x' la), x = x' -> fun_for_el_prod_l F x pf = fun_for_el_prod_l F x' pf'. intros; subst; repeat f_equal; apply proof_irrelevance. Qed. Lemma fun_for_el_prod_l_functional' : forall {la la':list At} (F:Fin_map (seg (length la)) signe mns) (x:At) (pf:In x la) (x':At) (pf':In x' la') (pfl:la = la'), x = x' -> let F' := fin_map_seg_transfer F (f_equal (@length At) pfl) in fun_for_el_prod_l F x pf = fun_for_el_prod_l F' x' pf'. simpl; intros; subst; repeat f_equal. apply fin_map_ext. intros x h1. rewrite fin_map_seg_transfer_compat'; auto. destruct h1; auto. apply proof_irrelevance. Qed. (*This domain is [linds _(firstn m la)].*) Definition fun_for_el_prod_l_linds_firstn {la:list At} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (ln:list nat) (pf:In ln (linds ba_eq_dec (firstn m la))) : At := if (has_comps_dec pfm F (list_to_set ln) (list_to_set_finite _)) then Bzero (Bc A) else let (x, pf) := constructive_definite_description _ (iff1 (in_linds_iff _ _ _) pf) in let (pf', _) := pf in eps x (F |-> lind ba_eq_dec _ _ pf'). (*This domain is [linds _ la].*) Definition fun_for_el_prod_l_linds {la:list At} (F:Fin_map (seg (length la)) signe mns) (ln:list nat) (pf:In ln (linds ba_eq_dec la)) : At := if (has_comps_dec' F (list_to_set ln) (list_to_set_finite _)) then Bzero (Bc A) else let (x, pf) := constructive_definite_description _ (iff1 (in_linds_iff _ _ _) pf) in let (pf', _) := pf in eps x (F |-> lind ba_eq_dec _ _ pf'). Lemma fun_for_el_prod_l_linds_firstn_functional : forall {la la':list At} {m:nat} (pfm:m <= length la) (pfm':m <= length la') (F:Fin_map (seg m) signe mns) (ln ln':list nat) (pf:In ln (linds ba_eq_dec (firstn m la))) (pf':In ln' (linds ba_eq_dec (firstn m la'))), la = la' -> ln = ln' -> fun_for_el_prod_l_linds_firstn pfm F ln pf = fun_for_el_prod_l_linds_firstn pfm' F ln' pf'. intros; subst; repeat f_equal; try apply proof_irrelevance. Qed. Lemma fun_for_el_prod_l_linds_functional : forall {la:list At} (F:Fin_map (seg (length la)) signe mns) (ln ln':list nat) (pf:In ln (linds ba_eq_dec la)) (pf':In ln' (linds ba_eq_dec la)), ln = ln' -> fun_for_el_prod_l_linds F ln pf = fun_for_el_prod_l_linds F ln' pf'. intros; subst; repeat f_equal; try apply proof_irrelevance. Qed. Lemma fun_for_el_prod_l_linds_functional' : forall {la la':list At} (F:Fin_map (seg (length la)) signe mns) (ln ln':list nat) (pf:In ln (linds ba_eq_dec la)) (pf':In ln' (linds ba_eq_dec la')) (pfe:la = la'), ln = ln' -> let F' := fin_map_seg_transfer F (f_equal (@length At) pfe) in fun_for_el_prod_l_linds F ln pf = fun_for_el_prod_l_linds F' ln' pf'. simpl; intros; subst; repeat f_equal; try apply proof_irrelevance; simpl; rewrite fin_map_seg_transfer_eq_refl; auto. Qed. Lemma fun_for_el_prod_l_linds_to_firstn : forall {la:list At} (F:Fin_map (seg (length la)) signe mns) (ln:list nat) (pf:In ln (linds ba_eq_dec la)), let pf' := in_linds_subst ba_eq_dec pf (eq_sym (firstn_length _)) in fun_for_el_prod_l_linds F ln pf = fun_for_el_prod_l_linds_firstn (le_refl _) F ln pf'. intros la F ln h1 h2. unfold fun_for_el_prod_l_linds, fun_for_el_prod_l_linds_firstn. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']; auto. unfold has_comps' in hl. contradiction. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']; auto. destruct constructive_definite_description as [x h8]; simpl. destruct h8 as [h8]; subst. unfold has_comps' in hr. contradiction. destruct constructive_definite_description as [x h8]; simpl. destruct h8 as [h8 h9]. unfold h2. clear h2. subst. destruct constructive_definite_description as [x' h8']; simpl. destruct h8' as [h8' h9]. rewrite firstn_length in h9. apply inj_linds_occ in h9; subst. repeat f_equal. symmetry. gen0. rewrite firstn_length. intro; f_equal. apply proof_irrelevance. assumption. rewrite firstn_length in h8'; auto. Qed. Lemma fun_for_el_prod_l_firstn_has : forall {la:list At} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (x:At) (pf:In x (firstn m la)), has_comps _ _ pfm (list_to_set (linds_occ ba_eq_dec (firstn m la) x)) F -> fun_for_el_prod_l_firstn pfm F x pf = 0. unfold fun_for_el_prod_l_firstn; intros; lr_if_goal; fold hlr; destruct hlr; try contradiction; auto. Qed. Lemma fun_for_el_prod_l_has : forall {la:list At} (F:Fin_map (seg (length la)) signe mns) (x:At) (pf:In x la), has_comps' _ (list_to_set (linds_occ ba_eq_dec la x)) F -> fun_for_el_prod_l F x pf = 0. intros; rewrite fun_for_el_prod_l_to_firstn. apply fun_for_el_prod_l_firstn_has; auto. unfold has_comps' in H. rewrite firstn_length; auto. Qed. Lemma fun_for_el_prod_l_firstn_not_has : forall {la:list At} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (x:At) (pf:In x (firstn m la)), ~ has_comps _ _ pfm (list_to_set (linds_occ ba_eq_dec (firstn m la) x)) F -> fun_for_el_prod_l_firstn pfm F x pf = eps x (F |-> lind ba_eq_dec _ _ pf). unfold fun_for_el_prod_l_firstn; intros; lr_if_goal; fold hlr; destruct hlr; try contradiction; auto. Qed. Lemma fun_for_el_prod_l_not_has : forall {la:list At} (F:Fin_map (seg (length la)) signe mns) (x:At) (pf:In x la), ~ has_comps' _ (list_to_set (linds_occ ba_eq_dec la x)) F -> fun_for_el_prod_l F x pf = eps x (F |-> lind ba_eq_dec _ _ pf). intros; rewrite fun_for_el_prod_l_to_firstn. rewrite fun_for_el_prod_l_firstn_not_has. repeat f_equal. gen0. rewrite firstn_length. intro; f_equal; apply proof_irrelevance. unfold has_comps' in H. rewrite firstn_length; auto. Qed. Lemma fun_for_el_prod_l_firstn_S_neq : forall {la:list At} {n:nat} (pfn:S n <= length la) (F:Fin_map (seg (S n)) signe mns) (x:At) (pfins:In x (firstn (S n) la)) (pfin: In x (firstn n la)), x <> nth_lt la n pfn -> let pfm' := lt_le_weak _ _ (lt_le_trans _ _ _ (lt_n_Sn _) pfn) in fun_for_el_prod_l_firstn pfn F x pfins = fun_for_el_prod_l_firstn pfm' (fin_map_seg_rest F n (lt_n_Sn _)) x pfin. intros la n h1 F x h2 h3 hn hle. unfold fun_for_el_prod_l_firstn. lr_if_goal; fold hlr; destruct hlr as [h4 | h4]. lr_if_goal; fold hlr; destruct hlr as [h5 | h5]. reflexivity. pose proof (has_comps_S_impl _ hle h1 _ _ (list_to_set_finite _) h4) as h6. destruct h6 as [h6 | h6]. rewrite (firstn_S_eq _ _ h1) in h4, h6. rewrite linds_occ_app_nin_r in h6. contradiction. assumption. contradict hn. simpl in hn. destruct hn; auto. contradiction. destruct h6 as [h6 [m [h8 [h9 [h10 h11]]]]]. rewrite <- list_to_set_in_iff in h9. rewrite in_linds_occ_iff in h9. destruct h9 as [m' [h9 [h12 h15]]]. rewrite <- list_to_set_in_iff in h6. rewrite in_linds_occ_iff in h6. destruct h6 as [n' [h6 [h13 h14]]]. pose proof (lind_occ_compat ba_eq_dec (firstn (S n) la) n' x h6) as h16. revert h16. Gen0. rewrite h14. intros h17 h18. rewrite nth_lt_firstn_eq in h18. rewrite <- h18 in hn. contradict hn. apply nth_lt_functional3. lr_if_goal; fold hlr; destruct hlr as [h5 | h5]. pose proof (has_comps_rest_impl_final _ hle h1 (list_to_set (linds_occ ba_eq_dec (firstn n la) x)) F (list_to_set_finite _) h5) as h6. rewrite (firstn_S_eq _ _ h1) in h4. rewrite linds_occ_app_nin_r in h4. contradiction. assumption. intro h7. simpl in h7. destruct h7; subst. contradict hn; auto. contradiction. rewrite fin_map_seg_rest_compat. f_equal. f_equal. do 2 rewrite lind_firstn. f_equal. apply proof_irrelevance. pose proof (lt_lind ba_eq_dec (firstn n la) x h3) as h6. rewrite length_firstn in h6; auto. Qed. Lemma fun_for_el_prod_l_firstn_S_eq_has_comps : forall {la:list At} {n:nat} (pfn:S n <= length la) (F:Fin_map (seg (S n)) signe mns) (pfins:In (nth_lt la n pfn) (firstn (S n) la)) (pfin: In (nth_lt la n pfn) (firstn n la)), let pfn' := lt_le_weak _ _ (lt_le_trans _ _ _ (lt_n_Sn _) pfn) in has_comps _ _ pfn' (list_to_set (linds_occ ba_eq_dec (firstn n la) (nth_lt la n pfn))) (fin_map_seg_rest F n (lt_n_Sn _)) -> fun_for_el_prod_l_firstn pfn F _ pfins = 0. intros la n h1 F h2 h3 h4 hc. unfold fun_for_el_prod_l_firstn. lr_if_goal; fold hlr; destruct hlr as [h5 | h5]; auto. contradict h5. rewrite (firstn_S_eq _ _ h1). rewrite linds_occ_app_in; auto. rewrite list_to_set_app, linds_occ_cons_eq; simpl. rewrite add_empty_sing; simpl. rewrite <- plus_n_O. rewrite length_firstn; auto with arith. eapply has_comps_rest_impl'. apply hc. left; auto. Qed. Lemma fun_for_el_prod_l_firstn_S_eq_not_has_comps : forall {la:list At} {n:nat} (pfn:S n <= length la) (F:Fin_map (seg (S n)) signe mns) (pfins:In (nth_lt la n pfn) (firstn (S n) la)) (pfin: In (nth_lt la n pfn) (firstn n la)), let pfn' := lt_le_weak _ _ (lt_le_trans _ _ _ (lt_n_Sn _) pfn) in ~has_comps _ _ pfn' (list_to_set (linds_occ ba_eq_dec (firstn n la) (nth_lt la n pfn))) (fin_map_seg_rest F n (lt_n_Sn _)) -> ~has_comps _ _ pfn (list_to_set (linds_occ ba_eq_dec (firstn (S n) la) (nth_lt la n pfn))) F -> fun_for_el_prod_l_firstn pfn F _ pfins = fun_for_el_prod_l_firstn pfn' (fin_map_seg_rest F n (lt_n_Sn _)) _ pfin. intros la n h1 F h2 h3 h4 h5 h8. rewrite fun_for_el_prod_l_firstn_not_has; auto. rewrite fun_for_el_prod_l_firstn_not_has; auto. rewrite fin_map_seg_rest_compat. do 2 rewrite lind_firstn. repeat f_equal. apply proof_irrelevance. eapply lt_eq_trans. apply lt_lind. rewrite length_firstn; auto. Qed. (*This iterates over [m] and [list_singularize la].*) Definition list_for_el_prod_l_firstn {la:list At} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) : list At := map_in_dep (fun_in_dep_incl _ _ (fun_for_el_prod_l_firstn pfm F) (incl_list_singularize ba_eq_dec _)). (*This iterates over [linds (firstn m la)].*) Definition list_for_el_prod_l_linds_firstn {la:list At} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) : list At := map_in_dep (fun_for_el_prod_l_linds_firstn pfm F). (*This iterates over [list_singularize la].*) Definition list_for_el_prod_l {la:list At} (F:Fin_map (seg (length la)) signe mns) : list At := map_in_dep (fun_in_dep_incl _ _ (fun_for_el_prod_l F) (incl_list_singularize ba_eq_dec _)). (*This iterates over [linds la].*) Definition list_for_el_prod_l_linds {la:list At} (F:Fin_map (seg (length la)) signe mns) : list At := map_in_dep (fun_for_el_prod_l_linds F). Lemma list_for_el_prod_l_to_firstn : forall {la:list At} (F:Fin_map (seg (length la)) signe mns), list_for_el_prod_l F = list_for_el_prod_l_firstn (le_refl _) F. intros la F. unfold list_for_el_prod_l_firstn, list_for_el_prod_l, fun_in_dep_incl. pose proof (firstn_length la) as h1. pose proof (f_equal (fun c => list_singularize ba_eq_dec c) h1) as h2. simpl in h2. rewrite (map_in_dep_functional _ h2). apply map_in_dep_ext. intros x h3. apply functional_extensionality. intro h4. rewrite transfer_fun_in_dep_compat. rewrite fun_for_el_prod_l_to_firstn. repeat f_equal. apply proof_irrelevance. Qed. Lemma list_for_el_prod_l_linds_to_firstn : forall {la:list At} (F:Fin_map (seg (length la)) signe mns), list_for_el_prod_l_linds F = list_for_el_prod_l_linds_firstn (le_refl _) F. intros la F. unfold list_for_el_prod_l_linds, list_for_el_prod_l_linds_firstn. rewrite map_in_dep_eq_iff. ex_goal. rewrite length_map_in_dep. repeat f_equal. rewrite firstn_length; auto. exists hex. intros n h1. erewrite nth_lt_map_in_dep. erewrite fun_for_el_prod_l_linds_to_firstn. pose proof (firstn_length la) as h2. pose proof (f_equal (fun c => linds ba_eq_dec c) (eq_sym h2)) as h3. simpl in h3. pose proof (nth_lt_functional1 _ _ n h3 h1) as h4. gen0. erewrite h4. intro. repeat f_equal. apply proof_irrelevance. Grab Existential Variables. rewrite firstn_length; auto. Qed. Lemma in_list_for_el_prod_l_firstn_iff : forall {la:list At} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (y:At), In y (list_for_el_prod_l_firstn pfm F) <-> exists (x:At) (pfx:In x (firstn m la)), y = fun_for_el_prod_l_firstn pfm F x pfx. intros la m h1 F y; split; intro h2. unfold list_for_el_prod_l_firstn in h2. unfold At, bt in h2. rewrite in_map_in_dep_iff in h2. destruct h2 as [x [h2 h3]]. unfold fun_in_dep_incl in h3. subst. pose proof h2 as h2'. unfold At, bt in h2'. rewrite <- list_singularize_in_iff in h2'. exists x, h2'. f_equal. apply proof_irrelevance. destruct h2 as [x [h2 ?]]; subst. unfold list_for_el_prod_l_firstn. unfold At, bt. rewrite in_map_in_dep_iff. exists x. ex_goal. rewrite <- list_singularize_in_iff; auto. exists hex. unfold fun_in_dep_incl. f_equal. apply proof_irrelevance. Qed. Lemma in_list_for_el_prod_l_iff : forall {la:list At} (F:Fin_map (seg (length la)) signe mns) (y:At), In y (list_for_el_prod_l F) <-> exists (x:At) (pfx:In x la), y = fun_for_el_prod_l F x pfx. intros la F y. pose proof (in_list_for_el_prod_l_firstn_iff (le_refl _) F y) as h1. rewrite list_for_el_prod_l_to_firstn. rewrite h1. intuition. destruct H0 as [x [h3 h4]]; subst. pose proof h3 as h3'. rewrite firstn_length in h3'. exists x, h3'. rewrite fun_for_el_prod_l_to_firstn. f_equal. apply proof_irrelevance. destruct H1 as [x [h2 h3]]; subst. exists x. ex_goal. rewrite firstn_length; auto. exists hex. rewrite fun_for_el_prod_l_to_firstn; f_equal; apply proof_irrelevance. Qed. Lemma list_for_el_prod_l_firstn_eq_list_for_el_prod_l_linds_firstn : forall {la:list At} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns), list_for_el_prod_l_firstn pfm F = list_for_el_prod_l_linds_firstn pfm F. intros la m F. unfold list_for_el_prod_l_firstn, list_for_el_prod_l_linds_firstn. symmetry. rewrite map_in_dep_eq_iff. ex_goal. rewrite length_linds. rewrite length_map_in_dep; auto. exists hex. intros n h1. erewrite nth_lt_map_in_dep. unfold fun_in_dep_incl. unfold fun_for_el_prod_l_linds_firstn, fun_for_el_prod_l_firstn. lr_if_goal; fold hlr; destruct hlr as [h2 | h2]; auto. lr_if_goal; fold hlr; destruct hlr as [h3 | h3]; auto. unfold linds in h2. erewrite nth_lt_map in h2. pose proof (h3 h2); contradiction. lr_if_goal; fold hlr; destruct hlr as [h4 | h4]; auto. unfold linds in h2. erewrite nth_lt_map in h2. pose proof (h2 h4); contradiction. destruct constructive_definite_description as [x h5]. destruct h5 as [h5 h6]. unfold linds in h6. erewrite nth_lt_map in h6. erewrite (linds_occ_compat' _ _ _ h5) in h6. pose proof h1 as h1'. rewrite length_linds in h1'. pose proof (in_nth_lt (list_singularize ba_eq_dec (firstn m la)) _ h1') as h7. rewrite <- list_singularize_in_iff in h7. erewrite nth_lt_functional3 in h6. erewrite (linds_occ_compat' ba_eq_dec (firstn m la) _ h7) in h6 at 1. apply inj_dep_linds_occ' in h6. subst. f_equal. f_equal. apply proof_irrelevance. f_equal. apply lind_functional. apply nth_lt_functional3. Grab Existential Variables. rewrite <- length_linds; auto. rewrite <- length_linds; auto. Qed. Lemma list_for_el_prod_l_eq_list_for_el_prod_l_linds : forall {la:list At} (F:Fin_map (seg (length la)) signe mns), list_for_el_prod_l F = list_for_el_prod_l_linds F. intros la F. rewrite list_for_el_prod_l_to_firstn, list_for_el_prod_l_linds_to_firstn, list_for_el_prod_l_firstn_eq_list_for_el_prod_l_linds_firstn; auto. Qed. Lemma has_comps_elt_seg_impl : forall (la:list At) (m r:nat) (pfm:m <= length la) (pfr:r < m) (F:Fin_map (seg m) signe mns), has_comps_elt la m pfm (seg m) (nth_lt la r (lt_le_trans r m (length la) pfr pfm)) F -> has_comps la m pfm (list_to_set (linds_occ ba_eq_dec (firstn m la) (nth_lt la r (lt_le_trans _ _ _ pfr pfm)))) F. intros la m r h1 h2 F h3. inversion h3. pose proof (length_firstn _ m h1) as h6. symmetry in h6. pose proof (nth_lt_firstn_eq _ _ m (lt_eq_trans _ _ _ pfc h6)) as h7. pose proof (nth_lt_firstn_eq _ _ m (lt_eq_trans _ _ _ pfd h6)) as h8. erewrite nth_lt_functional3 in H3. erewrite nth_lt_functional3 in H4. unfold At, bt in h7, H3. rewrite <- h7 in H3 at 1. unfold At, bt in h8, H4. rewrite <- h8 in H4 at 1. pose proof (ex_lind_occ_nth_lt ba_eq_dec _ _ (lt_eq_trans c m (length (firstn m la)) pfc h6)) as h9. pose proof (ex_lind_occ_nth_lt ba_eq_dec _ _ (lt_eq_trans d m (length (firstn m la)) pfd h6)) as h10. destruct h9 as [o1 [h9 h11]], h10 as [o2 [h12 h13]]. econstructor. apply list_to_set_finite. rewrite <- list_to_set_in_iff. rewrite in_linds_occ_iff. exists o1. ex_goal. erewrite nth_lt_functional3. rewrite <- nth_lt_firstn_eq. apply in_nth_lt. exists hex. split. unfold At, bt in H3. unfold At, bt. rewrite <- H3; auto. erewrite lind_occ_functional. apply h11. reflexivity. reflexivity. erewrite nth_lt_functional3. unfold At, bt. rewrite <- H3 at 1. reflexivity. rewrite <- list_to_set_in_iff. rewrite in_linds_occ_iff. exists o2. ex_goal. erewrite nth_lt_functional3. rewrite <- nth_lt_firstn_eq. apply in_nth_lt. exists hex. split. unfold At, bt in H4. unfold At, bt. rewrite <- H4; auto. erewrite lind_occ_functional. apply h13. reflexivity. reflexivity. erewrite nth_lt_functional3. unfold At, bt. rewrite <- H4 at 1. reflexivity. assumption. erewrite nth_lt_functional3. rewrite <- h7 at 1. symmetry. erewrite nth_lt_functional3. rewrite <- h8 at 1. rewrite H3, H4. reflexivity. assumption. Grab Existential Variables. rewrite length_firstn; auto. rewrite length_firstn; auto. destruct H1; auto. destruct H1; auto. Qed. Lemma has_comps_elt_seg_impl' : forall (la:list At) (r:nat) (pfr:r < length la) (F:Fin_map (seg (length la)) signe mns), has_comps_elt' la (seg (length la)) (nth_lt la r pfr) F -> has_comps' la (list_to_set (linds_occ ba_eq_dec la (nth_lt la r pfr))) F. unfold has_comps', has_comps_elt'. intros la r h1 F h2. pose proof (has_comps_elt_seg_impl la (length la) _ (le_refl _) h1 F) as h3. hfirst h3. erewrite nth_lt_functional3. apply h2. specialize (h3 hf). pose proof (proof_irrelevance _ h1 (lt_le_trans r (length la) (length la) h1 (le_refl (length la)))) as h4. rewrite <- h4 in h3. pose proof (linds_occ_functional ba_eq_dec la _ _ _ (eq_refl (nth_lt la r h1)) (eq_sym (firstn_length _))) as h5. unfold At, bt in h3, h5. simpl in h3, h5. rewrite <- h5 in h3; auto. Qed. Lemma has_comps_times_list_list_for_el_prod : forall {la:list At} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns), has_comps _ _ pfm (seg m) F -> times_list (list_for_el_prod_l_firstn pfm F) = 0. intros la m h1 F h2. pose proof h2 as h2'. rewrite has_comps_iff in h2. destruct h2 as [h2 [x [r [h3 h4]]]]. simpl in h4. destruct h4 as [h4 [h5 h6]]; subst. unfold list_for_el_prod_l_firstn. inversion h6. unfold list_for_el_prod_l_firstn. pose proof pfc as hc. rewrite <- (length_firstn _ _ h1) in hc. pose proof (in_nth_lt (firstn m la) _ hc) as h9. pose proof (list_singularize_in_iff ba_eq_dec (firstn m la) (nth_lt (firstn m la) c hc)) as h9'. unfold At in h9, h9'. rewrite h9' in h9. rewrite nth_lt_firstn_eq in h9. pose proof (in_map_in_dep (fun_in_dep_incl (firstn m la) (list_singularize ba_eq_dec (firstn m la)) (fun_for_el_prod_l_firstn h1 F) (incl_list_singularize ba_eq_dec (firstn m la))) _ h9) as h11. rewrite <- (times_list_times_in _ _ _ h11). unfold fun_in_dep_incl, fun_for_el_prod_l_firstn. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. rewrite zero_prod_l; auto. apply has_comps_elt_seg_impl in h6. contradict hr. erewrite nth_lt_functional3. rewrite H3 at 1. assumption. Qed. Lemma has_comps_times_list_list_for_el_prod' : forall {la:list At} (F:Fin_map (seg (length la)) signe mns), has_comps' _ (seg (length la)) F -> times_list (list_for_el_prod_l F) = 0. intros la F h1. red in h1. apply has_comps_times_list_list_for_el_prod in h1. rewrite list_for_el_prod_l_to_firstn; auto. Qed. Lemma incl_list_to_set_list_for_el_prod_l_firstn_S : forall {la:list At} {m:nat} (pfsm:S m <= length la) (F:Fin_map (seg (S m)) signe mns), let pfm := le_trans _ _ _ (le_n_Sn _) pfsm in Included (list_to_set (list_for_el_prod_l_firstn pfsm F)) (list_to_set (list_for_el_prod_l_firstn pfm (fin_map_seg_rest F m (lt_n_Sn m)) ++ (fun_for_el_prod_l_firstn pfsm F _ (in_nth_lt_firstn la _ _ pfsm (lt_n_Sn _))) :: nil)). intros la m h1 F h2. red; intros y h3. rewrite <- list_to_set_in_iff in h3. rewrite list_to_set_app. rewrite in_list_for_el_prod_l_firstn_iff in h3. destruct h3 as [x [h3 h4]]; subst. generalize (in_nth_lt_firstn la m (S m) h1 (lt_n_Sn m)). erewrite nth_lt_firstn_eq. intro hi. simpl. rewrite add_empty_sing. pose proof h3 as h3'. rewrite (firstn_S_eq _ _ h1) in h3'. apply in_app_or in h3'. destruct h3' as [h4 | h4]. destruct (ba_eq_dec x (nth_lt la _ h1)) as [|h5]; subst. simpl. right. rewrite in_sing_iff. unfold At. match goal with |- fun_for_el_prod_l_firstn _ _ (nth_lt _ _ ?pf) _ = _ => pose proof (proof_irrelevance _ pf h1) as h5 end. terml h5. pose proof (nth_lt_functional3 la m h1 x) as h6. unfold x in h6. generalize hi. simpl. unfold At, bt in h6. unfold At, bt. rewrite <- h6. intro h7. f_equal. apply proof_irrelevance. left. rewrite <- list_to_set_in_iff. rewrite in_list_for_el_prod_l_firstn_iff. exists x, h4. erewrite fun_for_el_prod_l_firstn_S_neq. f_equal. apply proof_irrelevance. assumption. simpl in h4. destruct h4; subst; try contradiction. right. rewrite in_sing_iff. match goal with |- fun_for_el_prod_l_firstn _ _ (nth_lt _ _ ?pf) _ = _ => assert (h5:pf = h1) end. apply proof_irrelevance. terml h5. pose proof (nth_lt_functional3 la m h1 x) as h6. unfold x in h6. generalize hi. simpl. unfold At, bt in h6. unfold At, bt. rewrite <- h6. intro h7. f_equal. apply proof_irrelevance. Qed. Lemma list_to_set_list_for_el_prod_l_firstn_S_eq_has_has : forall {la:list At} {m:nat} (pfsm:S m <= length la) (F:Fin_map (seg (S m)) signe mns), let pfm := le_trans _ _ _ (le_n_Sn _) pfsm in has_comps _ _ pfm (list_to_set (linds_occ ba_eq_dec (firstn m la) (nth_lt la m pfsm))) (fin_map_seg_rest F m (lt_n_Sn _)) -> has_comps _ _ pfsm (list_to_set (linds_occ ba_eq_dec (firstn (S m) la) (nth_lt la m pfsm))) F -> list_to_set (list_for_el_prod_l_firstn pfsm F) = list_to_set (list_for_el_prod_l_firstn pfm (fin_map_seg_rest F m (lt_n_Sn m))). intros la m h1 F hle h2 h3. apply Extensionality_Ensembles; red; split; intros x h4. apply incl_list_to_set_list_for_el_prod_l_firstn_S in h4. unfold fun_for_el_prod_l_firstn in h4. lr_if h4; fold hlr in h4; destruct hlr as [h5 | h5]; auto. rewrite list_to_set_app in h4. simpl in h4. rewrite add_empty_sing in h4. change (Inn (Add (list_to_set (list_for_el_prod_l_firstn (le_trans m (S m) (length la) (le_n_Sn m) h1) (fin_map_seg_rest F m (lt_n_Sn m)))) 0) x) in h4. rewrite in_add_eq in h4. assumption. rewrite <- list_to_set_in_iff. rewrite in_list_for_el_prod_l_firstn_iff. unfold fun_for_el_prod_l_firstn. exists (nth_lt la m h1). ex_goal. inversion h2. rewrite <- list_to_set_in_iff in H0. rewrite in_linds_occ_iff in H0. destruct H0 as [? [? ?]]; auto. exists hex. lr_if_goal; fold hlr; destruct hlr as [h6 | h6]; auto. contradiction. contradict h5. rewrite nth_lt_firstn_eq. erewrite nth_lt_functional3. unfold At, bt in h3. unfold At, bt. apply h3. rewrite <- list_to_set_in_iff in h4. rewrite <- list_to_set_in_iff. rewrite in_list_for_el_prod_l_firstn_iff in h4. rewrite in_list_for_el_prod_l_firstn_iff. destruct h4 as [y [h4 h5]]; subst. destruct (ba_eq_dec y (nth_lt la m h1)) as [h5 | h5]; subst. exists (nth_lt la m h1). ex_goal. erewrite firstn_S_eq. apply in_app_l; auto. exists hex. rewrite fun_for_el_prod_l_firstn_has. rewrite fun_for_el_prod_l_firstn_has; auto. assumption. exists y. ex_goal. erewrite firstn_S_eq. apply in_app_l; auto. exists hex. erewrite fun_for_el_prod_l_firstn_S_neq. f_equal. apply proof_irrelevance. assumption. Grab Existential Variables. assumption. assumption. Qed. Lemma has_comps_elt_in0_list_for_el_prod_l_firstn : forall {l:list At} {m:nat} (pfm: m <= length l) (F:Fin_map (seg m) signe mns) (x:At), has_comps_elt l m pfm (list_to_set (linds_occ ba_eq_dec (firstn m l) x)) x F -> In 0 (list_for_el_prod_l_firstn pfm F). intros l m h1 F x h2. unfold list_for_el_prod_l_firstn. unfold At, bt. rewrite in_map_in_dep_iff. inversion h2; subst. pose proof H1 as h10; pose proof H2 as h20. rewrite <- list_to_set_in_iff in H1, H2. rewrite in_linds_occ_iff in H1, H2. destruct H1 as [o [h3 [h4 h5]]], H2 as [o' [h3' [h4' h5']]]. unfold fun_in_dep_incl, fun_for_el_prod_l_firstn. pose proof h3 as h6. rewrite list_singularize_in_iff in h6. exists _, h6. lr_if_goal; fold hlr; destruct hlr as [hl | hr]; auto. contradict hr. econstructor. apply list_to_set_finite. apply h10. apply h20. assumption. rewrite H4 at 1. apply nth_lt_functional3. assumption. Grab Existential Variables. assumption. Qed. Lemma has_comps_elt_in0_list_for_el_prod_l_firstn' : forall {l:list At} (F:Fin_map (seg (length l)) signe mns) (x:At), has_comps_elt' l (list_to_set (linds_occ ba_eq_dec l x)) x F -> In 0 (list_for_el_prod_l F). intros l F x h1. red in h1. rewrite <- firstn_length in h1 at 4. apply has_comps_elt_in0_list_for_el_prod_l_firstn in h1. rewrite list_for_el_prod_l_to_firstn; auto. Qed. Lemma nin_list_for_el_prod_l_firstn_S : forall {l:list At} {m:nat} (pfsm: S m <= length l) (F:Fin_map (seg (S m)) signe mns), let pfm := le_length_removelast _ _ pfsm in ~ In (nth_lt l m pfsm) (firstn m l) -> list_for_el_prod_l_firstn pfsm F = list_for_el_prod_l_firstn pfm (fin_map_seg_rest F _ (lt_n_Sn _)) ++ eps (nth_lt l m pfsm) (F |-> m) :: nil. intros l m h1 F h2 h3. unfold list_for_el_prod_l_firstn, At, bt, fun_in_dep_incl. pose proof (in_nth_lt _ _ h1) as h4. pose proof (in_firstn_iff l (nth_lt l m h1) _ h1) as h5. hr h5. exists m. ex_goal. omega. exists hex. apply nth_lt_functional3. rewrite <- h5 in hr. assert (h6: list_singularize ba_eq_dec (firstn (S m) l) = list_singularize ba_eq_dec (firstn m (removelast l)) ++ (nth_lt l m h1) :: nil). rewrite (firstn_S_eq _ _ h1). rewrite firstn_removelast; auto with arith. rewrite list_singularize_rev_cons. f_equal. rewrite nin_remove_eq; auto. rewrite <- list_singularize_in_iff; auto. rewrite (map_in_dep_functional _ h6). rewrite map_in_dep_app. simpl. unfold fun_from_app2. rewrite transfer_fun_in_dep_compat. f_equal. apply map_in_dep_ext. intros x h7. unfold fun_from_app1; simpl. apply functional_extensionality. intros h8. rewrite transfer_fun_in_dep_compat. unfold fun_for_el_prod_l_firstn. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. lr_if_goal; fold hlr; destruct hlr as [hl'' | hr'']; auto. contradict hr''. pose proof (firstn_removelast _ _ h2) as h9. unfold At, bt in h9. unfold At, bt. rewrite h9 in h7, h8. rewrite <- list_singularize_in_iff in h7, h8. rewrite h9. unfold h2. pose proof (nin_has_comps_impl_removelast h1 F h3 _ h7 hl') as h10. simpl in h10. rewrite firstn_removelast in h10; auto. lr_if_goal; fold hlr; destruct hlr as [hl0 | hr0]. contradict hr'. rewrite <- list_singularize_in_iff in h7, h8. rewrite firstn_removelast in h7, h8, hl0. apply has_comps_removelast_impl in hl0; auto. apply h2. apply h2. apply h2. rewrite fin_map_seg_rest_compat. do 2 rewrite lind_firstn. repeat f_equal. rewrite lind_removelast_in. apply lind_functional; auto. eapply lt_eq_trans. apply lt_lind. rewrite firstn_removelast; auto with arith. rewrite length_firstn; auto. auto with arith. unfold fun_for_el_prod_l_firstn. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. apply has_comps_S_linds_occ_in in hl'. contradiction. rewrite nin_lind_nth_lt_firstn_S; auto. Qed. Lemma in_not_has_comps_elt_list_for_el_prod_l_firstn_S : forall {l:list At} {m:nat} (pfsm: S m <= length l) (F:Fin_map (seg (S m)) signe mns), let pfm := le_length_removelast _ _ pfsm in In (nth_lt l m pfsm) (firstn m l) -> ~has_comps_elt l (S m) pfsm (list_to_set (linds_occ ba_eq_dec (firstn (S m) l) (nth_lt l m pfsm))) (nth_lt l m pfsm) F -> list_to_set (list_for_el_prod_l_firstn pfsm F) = list_to_set (list_for_el_prod_l_firstn pfm (fin_map_seg_rest F _ (lt_n_Sn _))). intros l m h1 F h2 h3 h4. unfold list_for_el_prod_l_firstn, At, bt. pose proof (incl_firstn_le l _ _ (le_n_Sn m)) as h5. apply (incl_list_singularize_compat ba_eq_dec) in h5. match goal with |- list_to_set (map_in_dep ?f') = _ => pose f' as f end. match goal with |- _ = list_to_set (map_in_dep ?f') => pose f' as g end. pose proof (list_to_set_map_in_dep_decompose ba_eq_dec _ _ f h5) as h6. simpl in h6. fold f g. unfold At, bt. unfold At, bt in h6. rewrite h6 at 1. match goal with |- Union (list_to_set (map_in_dep ?f')) _ = _ => pose f' as h end. match goal with |- Union _ (list_to_set (map_in_dep ?f')) = _ => pose f' as k end. redterm2 k. redterm3 k. pose (list_minus ba_eq_dec c c0) as l0. assert (h7:l0 = nil). unfold l0, c, c0. apply impl_list_minus_nil. red. intros x h8. unfold At, bt. pose proof (list_singularize_in_iff ba_eq_dec (firstn m l) x) as hl. unfold At, bt in hl. rewrite <- hl. unfold At, bt in h8. pose proof (firstn_S_eq _ _ h1) as hf. simpl in hf. rewrite <- list_singularize_in_iff in h8. unfold At, bt in hf. rewrite hf in h8. apply in_app_or in h8. destruct h8 as [|h8]; auto. destruct h8 as [|h8]; subst. assumption. simpl in h8. contradiction. fold l0. rewrite (map_in_dep_functional _ h7). simpl. rewrite Union_commutative. rewrite empty_union. f_equal. pose proof (firstn_removelast l m) as h9. hfirst h9. rewrite length_removelast. omega. specialize (h9 hf). pose proof (f_equal (fun l':list At => list_singularize ba_eq_dec l') h9) as h10. simpl in h10. rewrite (map_in_dep_functional _ h10). apply map_in_dep_ext. intros x h11. apply functional_extensionality. intro h12. unfold f, g. rewrite transfer_fun_in_dep_compat. unfold fun_in_dep_incl. destruct (ba_eq_dec x (nth_lt l m h1)) as [h13 | h13]; subst. unfold fun_for_el_prod_l_firstn, fun_in_dep_decompose1. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. pose proof (has_comps_iff h1 (list_to_set (linds_occ ba_eq_dec (firstn (S m) l) (nth_lt l m h1))) F) as h14. rewrite h14 in hl. destruct hl as [h15 [x [r [h17 h18]]]]. simpl in h18. destruct h18 as [h18 [h19 h20]]. subst. rewrite <- list_to_set_in_iff in h18. rewrite in_linds_occ_iff in h18. destruct h18 as [s [h19 [h21 h22]]]; subst. revert h20. change (lind_occ ba_eq_dec (firstn (S m) l) s (nth_lt l m h1) h19 < S m) in h17. intro h20. change (has_comps_elt l (S m) h1 (list_to_set (linds_occ ba_eq_dec (firstn (S m) l) (nth_lt l m h1))) (nth_lt l (lind_occ ba_eq_dec (firstn (S m) l) s (nth_lt l m h1) h19) (lt_le_trans _ _ _ h17 h1)) F) in h20. revert h20. match goal with |- has_comps_elt _ _ _ _ (nth_lt _ _ ?pf) _ -> _ => generalize pf end. pose proof (lind_occ_firstn ba_eq_dec _ (nth_lt _ _ h1) (S m) h19 _ h21) as h22. unfold At, bt in h22. unfold At, bt. rewrite h22. intros h23 h25. pose proof (lind_occ_compat ba_eq_dec l s (nth_lt l m h1) (in_firstn l (S m) (nth_lt l m h1) h19)) as h24. pose proof (proof_irrelevance _ h23 (lt_lind_occ ba_eq_dec l s (nth_lt l m h1) (in_firstn l (S m) (nth_lt l m h1) h19))); subst. unfold At, bt in h24, h25. rewrite h24 in h25. contradiction. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. apply has_comps_removelast_impl' in hl'. contradiction. rewrite fin_map_seg_rest_compat. repeat f_equal. do 2 rewrite lind_firstn. rewrite lind_removelast_in. apply lind_functional. reflexivity. eapply lt_eq_trans. apply lt_lind. rewrite length_firstn; auto. unfold fun_for_el_prod_l_firstn. unfold fun_in_dep_decompose1. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']; auto. pose proof (impl_has_comps_removelast_neq _ _ h2 _ _ _ h13 hl). contradiction. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']; auto. pose proof (has_comps_removelast_impl'' _ _ h2 h1 _ _ hl'). contradiction. f_equal. rewrite fin_map_seg_rest_compat. f_equal. do 2 rewrite lind_firstn. rewrite lind_removelast_in. apply lind_functional; auto. eapply lt_eq_trans. apply lt_lind. rewrite length_firstn; auto. Qed. (*This is the goal of all this preamble, proving that we can access the elemetary product corresponding to each 2-valued map on list indices, with a product of a _real_ Boolean list, and not some composite of list and map. Yay!*) Lemma el_prod_l_firstn_eq : forall {l:list At} {m:nat} (pfm:m <= length l) (F:Fin_map (seg m) signe mns), el_prod_l_firstn pfm F = times_list (list_for_el_prod_l_firstn pfm F). intros l m. revert l. induction m as [|m ih]; simpl. intros l h1 F. unfold el_prod_l_firstn. unfold el_prod_l. simpl; auto. intros l h1 F. destruct (in_dec ba_eq_dec (nth_lt l m h1) (firstn m l)) as [h2 | h2]. pose proof (in_last_el_prod_l_firstn_dec h1 F h2 (list_to_set_finite _)) as h3. destruct h3 as [[h3 h4] | [h3 h4]]. rewrite h4. symmetry. rewrite <- h4. symmetry. specialize (ih (removelast l)). hfirst ih. rewrite length_removelast. omega. specialize (ih hf (fin_map_seg_rest F _ (lt_n_Sn _))). rewrite h4. rewrite el_prod_l_firstn_S in h4. pose proof (el_prod_l_firstn_removelast_compat (lt_le_weak m (length l) (lt_le_trans m (S m) (length l) (lt_n_Sn m) h1)) hf F) as h5. rewrite h5 in h4 at 1. rewrite ih in h4. symmetry. apply times_list_0. eapply has_comps_elt_in0_list_for_el_prod_l_firstn. apply h3. rewrite h4, ih. do 2 rewrite times_list_compat. apply in_not_has_comps_elt_list_for_el_prod_l_firstn_S in h3; auto. symmetry. gen0. unfold At, bt in h3. unfold At, bt. rewrite h3. intro h5. apply times_set_functional. unfold list_for_el_prod_l_firstn. pose proof (le_length_removelast _ _ h1) as h6. pose proof (firstn_removelast l m h6) as h7. pose proof (f_equal (fun l' => list_singularize ba_eq_dec l') h7) as h8. simpl in h8. rewrite (map_in_dep_functional _ h8). f_equal. apply map_in_dep_ext. intros x h9. unfold fun_in_dep_incl. gterml. gtermr. apply functional_extensionality. intro h10. rewrite transfer_fun_in_dep_compat. unfold fun_for_el_prod_l_firstn. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']; auto. rewrite firstn_removelast in hl; auto. pose proof (has_comps_removelast_iff _ _ (le_trans m (S m) (length l) (le_n_Sn m) h1) (le_length_removelast l m h1) x (fin_map_seg_rest F m (lt_n_Sn _))) as h11. unfold At, bt in h11, hl. rewrite h11 in hl at 1. contradiction. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']; auto. pose proof (has_comps_removelast_iff _ _ (le_trans m (S m) (length l) (le_n_Sn m) h1) (le_length_removelast l m h1) x (fin_map_seg_rest F m (lt_n_Sn _))) as h11. unfold At, bt in h11, hl', hr. rewrite firstn_removelast in hr. rewrite <- h11 in hl' at 1. contradiction. assumption. repeat f_equal. gen0. rewrite firstn_removelast. intro. apply lind_functional; auto. assumption. specialize (ih (removelast l) (le_length_removelast _ _ h1) (fin_map_seg_rest F m (lt_n_Sn _))). rewrite (nin_list_for_el_prod_l_firstn_S _ _ h2). unfold At in ih. unfold At. rewrite times_list_app. rewrite <- ih at 1. rewrite el_prod_l_firstn_S. rewrite comm_prod. f_equal. erewrite el_prod_l_firstn_removelast_compat. reflexivity. simpl. rewrite one_prod. repeat f_equal. apply proof_irrelevance. Qed. Lemma el_prod_l_eq : forall {l:list At} (F:Fin_map (seg (length l)) signe mns), el_prod_l F = times_list (list_for_el_prod_l F). intros l F. rewrite list_for_el_prod_l_to_firstn. rewrite el_prod_l_to_firstn. rewrite el_prod_l_firstn_eq; auto. Qed. End ElProdL. Section NormalForm. Variable A:Bool_Alg. Let At := Btype (Bc A). Definition eps_map {E:Ensemble At} (f:Fin_map E signe mns) : At->At := (fun i:At => eps i (f |-> i)). Definition eps_map' {E:Ensemble At} {R:Ensemble sign} (f:Fin_map E R mns) : At->At := (fun i:At => eps i (f |-> i)). Lemma eps_maps_dec : forall {E:Ensemble At} (f g:Fin_map E signe mns) (x:At), Inn E x -> {eps_map f x = eps_map g x} + {eps_map f x = -(eps_map g x)}. intros E f g x h1. unfold eps_map. destruct (f |-> x); destruct (g |-> x); simpl; auto. right. rewrite <- doub_neg. reflexivity. Qed. Definition eps_map_compose {B:Bool_Alg} {E:Ensemble (bt B)} (f:Fin_map E signe mns) (g:(bt B)->At) : (bt B) -> At := (fun i:(bt B) => eps (g i) (f |-> i)). Definition eps_map_compose' {B:Bool_Alg} {E:Ensemble (bt B)} {R:Ensemble sign} (f:Fin_map E R mns) (g:(bt B)->At) : (bt B) -> At := (fun i:(bt B) => eps (g i) (f |-> i)). Lemma eps_maps_compose_dec : forall {B:Bool_Alg} {E:Ensemble (bt B)} (f f':Fin_map E signe mns) (g:(bt B)->At) (x:(bt B)), Inn E x -> {eps_map_compose f g x = eps_map_compose f' g x} + {eps_map_compose f g x = -(eps_map_compose f' g x)}. intros B E f f' g x h1. unfold eps_map_compose. destruct (f |-> x); destruct (f' |-> x); simpl; auto. right. rewrite <- doub_neg. reflexivity. Qed. Definition im_eps {E:Ensemble At} (f:Fin_map E signe mns) : Ensemble At := Im E (eps_map f). Definition im_eps' {E:Ensemble At} {R:Ensemble sign} (f:Fin_map E R mns) : Ensemble At := Im E (eps_map' f). Definition im_eps_compose {B:Bool_Alg} {E:Ensemble (bt B)} (f:Fin_map E signe mns) (g:(bt B)->At) : Ensemble At := Im E (eps_map_compose f g). Definition im_eps_compose' {B:Bool_Alg} {E:Ensemble (bt B)} {R:Ensemble sign} (f:Fin_map E R mns) (g:(bt B)->At) : Ensemble At := Im E (eps_map_compose' f g). Lemma finite_im_eps : forall {E:Ensemble At} (f:Fin_map E signe mns), Finite (im_eps f). intros E f. apply finite_image. apply (fin_map_fin_dom f). Qed. Definition finite_im_eps_compose : forall {B:Bool_Alg} {E:Ensemble (bt B)} (f:Fin_map E signe mns) (g:(bt B)->At), Finite (im_eps_compose f g). intros B E f g. apply finite_image. apply (fin_map_fin_dom f). Qed. Lemma im_eps_im_eps_compose_compat : forall {B:Bool_Alg} {E:Ensemble (bt B)} (g:sig_set E->At) (f : Fin_map (Im (full_sig E) g) signe mns) (pf:Finite E), im_eps f = im_eps_compose (fin_map_im_full_sig_eq ba_eq_dec f pf 0) (sig_fun_app g 0). intros B E g f h1. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. unfold im_eps_compose. destruct h2 as [x h2]. subst. clear h2. destruct x as [x h2]. apply Im_intro with x; auto. unfold eps_map, eps_map_compose. f_equal. unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h3 | h4]. f_equal. apply proj1_sig_injective. simpl. reflexivity. contradiction. pose proof (fin_map_im_full_sig_eq_compat ba_eq_dec f h1 0) as h3. simpl in h3. rewrite h3. f_equal. unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h3' | h4']. f_equal. apply proj1_sig_injective. simpl. reflexivity. contradiction. assumption. red. intros x h2. destruct h2 as [x h2]. subst. unfold im_eps. rewrite im_im. apply Im_intro with (exist _ _ h2). constructor. unfold eps_map_compose, eps_map. f_equal. unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h3' | h4']. f_equal. apply proj1_sig_injective. simpl. reflexivity. contradiction. pose proof (fin_map_im_full_sig_eq_compat ba_eq_dec f h1 0) as h3. simpl in h3. rewrite h3. f_equal. unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h3' | h4']. f_equal. apply proj1_sig_injective. simpl. reflexivity. contradiction. assumption. Qed. (*Keep in mind that I'm using the name "elementary product" from [Koppelberg/Monk], although there they don't use finite maps, unlike in [Givant/Halmos] and my code.*) Definition el_prod {E:Ensemble At} (f:Fin_map E signe mns) : At. pose (fun i:At => eps i (f |-> i)) as p. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ E p h1) as h2. refine (times_set _ h2). Defined. Definition el_sum {E:Ensemble At} (f:Fin_map E signe mns) : At. pose (fun i:At => eps i (f |-> i)) as p. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ E p h1) as h2. refine (plus_set _ h2). Defined. Definition el_prod_compose {B:Bool_Alg} {E:Ensemble (bt B)} (g:(bt B)->At) (f:Fin_map E signe mns) : At. pose (fun i:(bt B) => eps (g i) (f |-> i)) as p. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ E p h1) as h2. refine (times_set _ h2). Defined. (*This is to interface with some specific cases where we need "injectors" of g, i.e. injective restrictions of g, whose image agrees with g on E *) Definition el_prod_compose' {B:Bool_Alg} {E:Ensemble (bt B)} (g:(bt B)->At) (f:Fin_map E signe mns) : At. pose (fun i:sig_set E => g (proj1_sig i)) as g'. pose (fun i:sig_set E => eps (g' i) (f |-> (proj1_sig i))) as p. pose proof (fin_map_fin_dom f) as h1. rewrite finite_full_sig_iff in h1. pose proof (finite_image _ _ (full_sig E) p h1) as h2. refine (times_set _ h2). Defined. Definition fin_map_seg_to_firstn {m:nat} {l:list At} (pf:m <= length l) (f:Fin_map (Full_set (seg_one_t m)) signe mns) : Fin_map (Full_set (seg_one_t (length (firstn m l)))) signe mns := fin_map_seg_one_transfer f (eq_sym (length_firstn l m pf)). Lemma el_prod_compose_el_prod_compose'_compat : forall {B:Bool_Alg} {E:Ensemble (bt B)} (g:(bt B)->At), el_prod_compose g = el_prod_compose' (E:=E) g. intros B E g. unfold el_prod_compose, el_prod_compose'. apply functional_extensionality. intro f. apply times_set_functional. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. subst. apply Im_intro with (exist _ _ h1). constructor. simpl. reflexivity. red. intros x h1. destruct h1 as [x h1]. subst. destruct x as [x h2]. simpl. apply Im_intro with x; auto. Qed. Definition el_sum_compose {B:Bool_Alg} {E:Ensemble (bt B)} (g:(bt B)->At) (f:Fin_map E signe mns) : At. pose (fun i:(bt B) => eps (g i) (f |-> i)) as p. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ E p h1) as h2. refine (plus_set _ h2). Defined. Lemma el_prod_el_prod_compose_compat : forall {B:Bool_Alg} {E:Ensemble (bt B)} {g:sig_set E->At} (f:Fin_map (Im (full_sig E) g) signe mns) (pf:Finite E), el_prod f = el_prod_compose (sig_fun_app g 0) (fin_map_im_full_sig_eq ba_eq_dec f pf 0). intros B E g f h1. unfold el_prod, el_prod_compose. apply times_set_functional. rewrite im_im. pose proof (fin_map_im_full_sig_eq_compat ba_eq_dec f h1 0) as h5. simpl in h5. apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [x h3]. subst. destruct x as [x h4]. apply Im_intro with x. assumption. specialize (h5 _ h4). rewrite h5. assert (h6:sig_fun_app g 0 x = g (exist _ _ h4)). unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h7 | h8]. pose proof (proof_irrelevance _ h7 h4); subst; auto; contradiction. contradiction. rewrite h6. f_equal. red. intros x h6. destruct h6 as [x h6]. subst. apply Im_intro with (exist _ _ h6). constructor. specialize (h5 _ h6). assert (h0:sig_fun_app g 0 x = g (exist _ _ h6)). unfold sig_fun_app. lr_match_goal'; auto. repeat f_equal; apply proof_irrelevance. contradiction. rewrite h5, h0; auto. Qed. Lemma el_sum_el_sum_compose_compat : forall {B:Bool_Alg} {E:Ensemble (bt B)} {g:sig_set E->At} (f:Fin_map (Im (full_sig E) g) signe mns) (pf:Finite E), el_sum f = el_sum_compose (sig_fun_app g 0) (fin_map_im_full_sig_eq ba_eq_dec f pf 0). intros B E g f h1. unfold el_sum, el_sum_compose. pose proof (fin_map_im_full_sig_eq_compat ba_eq_dec f h1 0) as h5. simpl in h5. apply plus_set_functional. rewrite im_im. apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [x h3]. subst. destruct x as [x h4]. apply Im_intro with x. assumption. specialize (h5 _ h4). rewrite h5. assert (h6:sig_fun_app g 0 x = g (exist _ _ h4)). unfold sig_fun_app. lr_match_goal'; auto. repeat f_equal; apply proof_irrelevance. contradiction. rewrite h6. f_equal. red. intros x h6. destruct h6 as [x h6]. subst. apply Im_intro with (exist _ _ h6). constructor. specialize (h5 _ h6). assert (h0:sig_fun_app g 0 x = g (exist _ _ h6)). unfold sig_fun_app. lr_match_goal'; auto. repeat f_equal; apply proof_irrelevance. contradiction. rewrite h5, h0; auto. Qed. Lemma el_prod_el_prod_compose_compat' : forall {B:Bool_Alg} {E:Ensemble (bt B)} (f : Fin_map E signe mns) (g:sig_set E->At) (f': Fin_map (Im (full_sig E) g) signe mns), (forall b:(bt B), Inn E b -> f' |-> ((g, (Bzero (Bc A))) ||-> b) = f |-> b) -> el_prod_compose (sig_fun_app g 0) f = el_prod f'. intros B E f g f' h1. pose proof (fin_map_fin_dom f) as h0. unfold el_prod_compose, el_prod. apply times_set_functional. rewrite im_im. apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [x h3]. subst. specialize (h1 _ h3). rewrite <- h1. apply Im_intro with (exist _ _ h3). constructor. f_equal. unfold sig_fun_app. lr_match_goal'; auto. repeat f_equal; apply proof_irrelevance. contradiction. f_equal. simpl. unfold sig_fun_app. lr_match_goal'; auto. repeat f_equal; apply proof_irrelevance. contradiction. f_equal. simpl. red. intros x h3. destruct h3 as [x h3]. subst. destruct x as [x h4]. clear h3. apply Im_intro with x. assumption. unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h4' | h5']. assert (h4 = h4'). apply proof_irrelevance. subst. f_equal. specialize (h1 _ h4'). rewrite <- h1. f_equal. unfold sig_fun_app. lr_match_goal'; simpl; auto. repeat f_equal; apply proof_irrelevance. contradiction. simpl. f_equal. contradiction. contradiction. Qed. Lemma times_set_def_im_rel_class_im_rel_set_fun_value : forall {B:Bool_Alg} (E S:Ensemble (bt B)), Finite E -> forall (g:bt B->At), Inn (rel_classes_im_rel_set E g) S -> forall x:(bt B), Inn S x -> times_set_def _ 0 (Im S g) = g x. intros B E S h0 g h1 x ha. unfold times_set_def. destruct classic_dec as [h2 | h3]. rewrite <- times_set_sing. apply times_set_functional. apply Extensionality_Ensembles. red. split. red. intros x'' h4. destruct h4 as [x'' h4]. subst. destruct h1 as [x' h1]. subst. destruct h4 as [h4]. destruct h4 as [h4 h5]. destruct ha as [ha]. destruct ha as [ha hb]. red in hb, h5. rewrite <- hb, h5. constructor. red. intros x'' h4. destruct h4. apply Im_intro with x; auto. contradict h3. apply finite_image. eapply Finite_downward_closed. apply h0. destruct h1 as [S h1]. subst. red. intros x' h2. destruct h2 as [h2]. destruct h2; auto. Qed. Lemma im_rel_classes_im_rel_set : forall {B:Bool_Alg} {E:Ensemble (bt B)} (pf:Finite E) (g:bt B->At), Im (rel_classes_im_rel_set E g) (fun S => (times_set_def _ 0 (Im S g))) = Im E g. intros B E h1 g. assert (h4: forall x, Inn E x -> Im [y : bt B | Inn E y /\ im_rel g x y] g = Singleton (g x)). intros x' hin. apply Extensionality_Ensembles. red. split. red. intros x'' h4. destruct h4 as [x'' h4]. subst. destruct h4 as [h4]. destruct h4 as [h4 h5]. red in h5. rewrite h5. constructor. red. intros x'' h4. destruct h4. apply Im_intro with x'. constructor. split; auto. red. auto. auto. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. destruct h2 as [x h2 q h3]. subst. apply Im_intro with x; auto. rewrite <- times_set_sing. unfold times_set_def. destruct classic_dec as [h3 | h4']. apply times_set_functional. rewrite h4. reflexivity. assumption. contradict h4'. apply finite_image. eapply Finite_downward_closed. apply h1. red. intros x' h3. destruct h3 as [h3]. destruct h3; auto. red. intros y h2. destruct h2 as [y h2]. subst. apply Im_intro with [x:(bt B) | Inn E x /\ im_rel g y x]. unfold rel_classes_im_rel_set. unfold rel_classes. apply Im_intro with y. assumption. reflexivity. unfold times_set_def. destruct classic_dec as [h3 | h4']. rewrite <- (times_set_sing (g y)). apply times_set_functional. rewrite h4; auto. contradict h4'. apply finite_image. eapply Finite_downward_closed. apply h1. red. intros x' h3. destruct h3 as [h3]. destruct h3; auto. Qed. Lemma el_prod_le_ai : forall {E:Ensemble At} (a:Fin_map E signe mns) (i:At), Inn E i -> le (el_prod a) (eps i (a |-> i)). intros E a i h1. unfold el_prod. apply le_times_set. apply Im_intro with i; auto. Qed. Lemma ai_le_el_sum : forall {E:Ensemble At} (a:Fin_map E signe mns) (i:At), Inn E i -> le (eps i (a |-> i)) (el_sum a). intros E a i h1. unfold el_sum. apply le_plus_set. apply Im_intro with i; auto. Qed. Lemma el_prod_extends_le : forall {E F:Ensemble At} (a:Fin_map E signe mns) (b:Fin_map F signe mns), extends b a -> le (el_prod b) (el_prod a). intros E F a b h1. unfold el_prod. apply times_set_inc_le. red. intros x h2. red in h1. destruct h1 as [h1 h3]. destruct h2 as [x h2]. subst. specialize (h3 _ h2). rewrite h3. apply Im_intro with x. apply h1; auto. reflexivity. Qed. Lemma non_zero_el_prod_inj : forall {E:Ensemble At} (f g:Fin_map E signe mns), el_prod f <> 0 -> el_prod f = el_prod g -> im_eps f = im_eps g. intros E f g h1 h2. apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [x h3]. subst. pose proof (el_prod_le_ai f _ h3) as h4. destruct (eps_maps_dec f g x h3) as [h5 | h6]. rewrite h5. apply Im_intro with x; auto. pose proof (el_prod_le_ai g _ h3) as h7. rewrite <- h2 in h7. unfold eps_map in h6. rewrite h6 in h4. pose proof (mono_prod _ _ _ _ h7 h4) as h8. rewrite idem_prod in h8. rewrite comp_prod in h8. apply le_x_0 in h8. contradiction. red. intros x h3. destruct h3 as [x h3]. subst. pose proof (el_prod_le_ai g _ h3) as h4. destruct (eps_maps_dec f g x h3) as [h5 | h6]. rewrite <- h5. apply Im_intro with x; auto. pose proof (el_prod_le_ai f _ h3) as h7. rewrite h2 in h7. unfold eps_map in h6. rewrite h6 in h7. pose proof (mono_prod _ _ _ _ h4 h7) as h8. rewrite idem_prod in h8. rewrite comp_prod in h8. apply le_x_0 in h8. rewrite h2 in h1. contradiction. Qed. Lemma el_prod_disjoint : forall (E:Ensemble At) (a b:Fin_map E signe mns), a <> b -> (el_prod a) * (el_prod b) = 0. intros E a b h3. pose proof (distinct_fin_maps_differ_at_point sign_eq_dec a b h3) as h4. destruct h4 as [i h4]. destruct h4 as [h4l h4r]. pose proof (el_prod_le_ai a i h4l) as h5. pose proof (el_prod_le_ai b i h4l) as h6. destruct (a |->i) ; destruct (b |-> i). contradict h4r. reflexivity. simpl in h5. simpl in h6. pose proof (mono_prod _ _ _ _ h5 h6) as h7. rewrite comp_prod in h7. red in h7. rewrite zero_sum in h7. assumption. simpl in h5. simpl in h6. pose proof (mono_prod _ _ _ _ h5 h6) as h7. rewrite (comm_prod _ (- i) i) in h7. rewrite comp_prod in h7. red in h7. rewrite zero_sum in h7. assumption. contradict h4r. reflexivity. Qed. Lemma el_prod_compose_le_ai : forall {B:Bool_Alg} {E:Ensemble (bt B)} (a:Fin_map E signe mns) (g:(bt B)->At) (i:(bt B)), Inn E i -> le (el_prod_compose g a) (eps (g i) (a |-> i)). intros B E a g i h1. unfold el_prod_compose. apply le_times_set. apply Im_intro with i; auto. Qed. Lemma ai_le_el_sum_compose : forall {B:Bool_Alg} {E:Ensemble (bt B)} (a:Fin_map E signe mns) (g:(bt B)->At) (i:(bt B)), Inn E i -> le (eps (g i) (a |-> i)) (el_sum_compose g a). intros B E a g i h1. unfold el_sum_compose. apply le_plus_set. apply Im_intro with i; auto. Qed. Lemma el_prod_compose_extends_le : forall {B:Bool_Alg} {E F:Ensemble (bt B)} (a:Fin_map E signe mns) (b:Fin_map F signe mns) (g:(bt B)->At), extends b a -> le (el_prod_compose g b) (el_prod_compose g a). intros B E F a b g h1. unfold el_prod_compose. apply times_set_inc_le. red. intros x h2. red in h1. destruct h1 as [h1 h3]. destruct h2 as [x h2]. subst. specialize (h3 _ h2). rewrite h3. apply Im_intro with x. apply h1; auto. reflexivity. Qed. Lemma el_prod_compose_disjoint : forall {B:Bool_Alg} (E:Ensemble (bt B)) (g:(bt B)->At) (a b:Fin_map E signe mns), a <> b -> (el_prod_compose g a) * (el_prod_compose g b) = 0. intros B E g a b h3. pose proof (distinct_fin_maps_differ_at_point sign_eq_dec _ _ h3) as h4. destruct h4 as [i h4]. destruct h4 as [h4l h4r]. pose proof (el_prod_compose_le_ai a g i h4l) as h5. pose proof (el_prod_compose_le_ai b g i h4l) as h6. destruct (a |->i) ; destruct (b |-> i). contradict h4r. reflexivity. simpl in h5. simpl in h6. pose proof (mono_prod _ _ _ _ h5 h6) as h7. rewrite comp_prod in h7. red in h7. rewrite zero_sum in h7. assumption. simpl in h5. simpl in h6. pose proof (mono_prod _ _ _ _ h5 h6) as h7. rewrite (comm_prod _ (- (g i)) (g i)) in h7. rewrite comp_prod in h7. red in h7. rewrite zero_sum in h7. assumption. contradict h4r. reflexivity. Qed. Lemma non_zero_el_prod_compose_inj : forall {B:Bool_Alg} {E:Ensemble (bt B)} (f f':Fin_map E signe mns) (g:(bt B)->At), el_prod_compose g f <> 0 -> el_prod_compose g f = el_prod_compose g f' -> im_eps_compose f g = im_eps_compose f' g. intros B E f f' g h1 h2. apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [x h3]. subst. pose proof (el_prod_compose_le_ai f g _ h3) as h4. destruct (eps_maps_compose_dec f f' g x h3) as [h5 | h6]. rewrite h5. apply Im_intro with x; auto. pose proof (el_prod_compose_le_ai f' g _ h3) as h7. rewrite <- h2 in h7. unfold eps_map_compose in h6. rewrite h6 in h4. pose proof (mono_prod _ _ _ _ h7 h4) as h8. rewrite idem_prod in h8. rewrite comp_prod in h8. apply le_x_0 in h8. contradiction. red. intros x h3. destruct h3 as [x h3]. subst. pose proof (el_prod_compose_le_ai f' g _ h3) as h4. destruct (eps_maps_compose_dec f f' g x h3) as [h5 | h6]. rewrite <- h5. apply Im_intro with x; auto. pose proof (el_prod_compose_le_ai f g _ h3) as h7. rewrite h2 in h7. unfold eps_map_compose in h6. rewrite h6 in h7. pose proof (mono_prod _ _ _ _ h4 h7) as h8. rewrite idem_prod in h8. rewrite comp_prod in h8. apply le_x_0 in h8. rewrite h2 in h1. contradiction. Qed. Lemma non_zero_el_prod_compose_constant : forall {B:Bool_Alg} {E:Ensemble (bt B)} (g:bt B->At) (f:Fin_map E signe mns), el_prod_compose g f <> 0 -> forall (x x':bt B), Inn E x -> Inn E x' -> g x = g x' -> f |-> x = f |-> x'. intros B E g f h1 x x' hx hx' h2. pose proof (sign_dec (f |-> x) (f |-> x')) as h3. destruct h3 as [|h3]; auto. unfold el_prod_compose in h1. assert (h4: (eps (g x) (f |-> x)) * (eps (g x') (f |-> x')) = 0). unfold eps. simpl. destruct (f |-> x); destruct (f |-> x'). contradict h3. reflexivity. rewrite h2. apply comp_prod. rewrite h2. rewrite comm_prod. apply comp_prod. contradict h3. reflexivity. assert (h5:Included (Couple x x') E). red; intros e h5. destruct h5; auto. pose proof (decompose_setminus_inc _ _ h5) as h6. pose proof (fin_map_fin_dom f) as h0. pose proof (f_equal (fun S => Im S (fun i : bt B => eps (g i) (f |-> i))) h6) as h7. simpl in h7. let P := type of h7 in match P with ?F = ?G => pose G as S end. assert (h8:Finite S). unfold S. rewrite <- h7. apply finite_image; auto. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (bt B) (bt A) E (fun i : bt B => eps (g i) (f |-> i)) (fin_map_fin_dom f)) h8 h7) as h9. dependent rewrite -> h9 in h1. pose proof (im_union (Couple x x') (Setminus E (Couple x x')) (fun i : bt B => eps (g i) (f |-> i))) as h10. assert (h11:Finite (Union (Im (Couple x x') (fun i : bt B => eps (g i) (f |-> i))) (Im (Setminus E (Couple x x')) (fun i : bt B => eps (g i) (f |-> i))))). rewrite <- h10. auto. pose proof (subsetT_eq_compat _ _ _ _ h8 h11 h10) as h12. dependent rewrite -> h12 in h1. pose proof (finite_couple x x') as h13. assert (h14:Finite (Setminus E (Couple x x'))). eapply Finite_downward_closed. apply h0. red. intros e h15. destruct h15; auto. pose proof (finite_image _ _ _ (fun i:bt B => eps (g i) (f |-> i)) h13) as h16. pose proof (finite_image _ _ _ (fun i:bt B => eps (g i) (f |-> i)) h14) as h17. pose proof (times_set_union' _ _ h16 h17 h11) as h18. unfold bt, At in h1, h18. rewrite h18 in h1. assert (h19: times_set (Im (Couple x x') (fun i : Btype (Bc B) => eps (g i) (f |-> i))) h16 = 0). assert (h20: Im (Couple x x') (fun i : Btype (Bc B) => eps (g i) (f |-> i)) = Couple (eps (g x) (f |-> x)) (eps (g x') (f |-> x'))). apply Extensionality_Ensembles. red. split. red. intros a h19. destruct h19 as [a h19 q h20]. rewrite h20. clear h20 q. destruct h19. left. right. red. intros a h19. destruct h19. apply Im_intro with x. left. reflexivity. apply Im_intro with x'. right. reflexivity. assert (h21:Finite (Couple (eps (g x) (f |-> x)) (eps (g x') (f |-> x')))). rewrite <- h20. assumption. pose proof (subsetT_eq_compat _ _ _ _ h16 h21 h20) as h22. dependent rewrite -> h22. rewrite times_set_couple'. assumption. rewrite h19 in h1 at 1. rewrite comm_prod in h1. rewrite zero_prod in h1. contradict h1. reflexivity. Qed. Lemma el_prod_disjoint' : forall (E:Ensemble At) (X Y:Ensemble (Fin_map E signe mns)), Inhabited [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] -> (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod (fst x) * el_prod (snd x))) = Singleton 0. intros E X Y h1. apply Extensionality_Ensembles. red. split. red. intros x h2. inversion h2 as [x' h3 y h4 h5]. subst. clear h2. destruct h3 as [h3]. destruct h3 as [h3l h3r]. destruct h3l as [h3l]. destruct h3l as [h3a h3b]. pose proof (el_prod_disjoint _ (fst x') (snd x') h3r) as h4. rewrite h4. constructor. destruct h1 as [pr h1]. destruct h1 as [h1]. destruct h1 as [h1l h1r]. pose proof (el_prod_disjoint _ (fst pr) (snd pr) h1r) as h2. red. intros x h1. destruct h1; subst. rewrite <- h2. apply Im_intro with pr. constructor; auto. reflexivity. Qed. Lemma el_prod_compose_disjoint' : forall {B:Bool_Alg} (E:Ensemble (bt B)) (g:(bt B)->At) (X Y:Ensemble (Fin_map E signe mns)), Inhabited [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] -> (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_compose g (fst x) * el_prod_compose g (snd x))) = Singleton 0. intros B E g X Y h1. apply Extensionality_Ensembles. red. split. red. intros x h2. inversion h2 as [x' h3 y h4 h5]. subst. clear h2. destruct h3 as [h3]. destruct h3 as [h3l h3r]. destruct h3l as [h3l]. destruct h3l as [h3a h3b]. pose proof (el_prod_compose_disjoint _ g (fst x') (snd x') h3r) as h4. rewrite h4. constructor. destruct h1 as [pr h1]. destruct h1 as [h1]. destruct h1 as [h1l h1r]. pose proof (el_prod_compose_disjoint _ g (fst pr) (snd pr) h1r) as h2. red. intros x h1. destruct h1; subst. rewrite <- h2. apply Im_intro with pr. constructor; auto. reflexivity. Qed. Lemma plus_set_el_prod_disjoint : forall (E:Ensemble At) (X Y:Ensemble (Fin_map E signe mns)) (pf:Finite [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x]), plus_set (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod (fst x) * el_prod (snd x))) (finite_image _ _ _ _ pf) = 0. intros E X Y h1. match goal with |- plus_set (Im ?C _) _ = _ => pose C as S end. destruct (classic (Inhabited S)) as [h2 | h3]. pose proof (el_prod_disjoint' _ X Y h2) as h3. match goal with |- plus_set _ ?pf = _ => generalize pf end. unfold bt in h3. unfold bt. rewrite h3. intro. rewrite plus_set_sing' at 1. reflexivity. pose proof (not_inhabited_empty _ h3) as h4. subst. pose proof (image_empty _ _ (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod (fst x) * el_prod (snd x))) as h5. rewrite <- h4 in h5. rewrite <- plus_set_empty. apply plus_set_functional; auto. Qed. Lemma plus_set_el_prod_disjoint' : forall (E:Ensemble At) (X Y:Ensemble (Fin_map E signe mns)) (pf:Finite [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x]) (pfi:Finite (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod (fst x) * el_prod (snd x)))), plus_set (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod (fst x) * el_prod (snd x))) pfi = 0. intros E X Y h1 h2. pose proof (plus_set_el_prod_disjoint _ X Y h1) as h3. rewrite <- h3. apply plus_set_functional. reflexivity. Qed. Lemma plus_set_el_prod_compose_disjoint : forall {B:Bool_Alg} (E:Ensemble (bt B)) (g:(bt B)->At) (X Y:Ensemble (Fin_map E signe mns)) (pf:Finite [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x]), plus_set (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_compose g (fst x) * el_prod_compose g (snd x))) (finite_image _ _ _ _ pf) = 0. intros B E g X Y h1. match goal with |- plus_set (Im ?C _) _ = _ => pose C as S end. destruct (classic (Inhabited S)) as [h2 | h3]. pose proof (el_prod_compose_disjoint' _ g X Y h2) as h3. match goal with |- plus_set _ ?pf = _ => generalize pf end. unfold bt in h3. unfold bt. rewrite h3. intro. apply plus_set_sing'. pose proof (not_inhabited_empty _ h3) as h4. match goal with |- plus_set (Im ?S ?h) _ = _ => pose h as f end. pose proof (image_empty _ _ f) as h5. fold f. rewrite <- h4 in h5. rewrite <- plus_set_empty. apply plus_set_functional. fold f; auto. Qed. Lemma plus_set_el_prod_compose_disjoint' : forall {B:Bool_Alg} (E:Ensemble (bt B)) (g:(bt B)->At) (X Y:Ensemble (Fin_map E signe mns)) (pf:Finite [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x]) (pfi:Finite (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_compose g (fst x) * el_prod_compose g (snd x)))), plus_set (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_compose g (fst x) * el_prod_compose g (snd x))) pfi = 0. intros B E g X Y h1 h2. pose proof (plus_set_el_prod_compose_disjoint _ g X Y h1) as h3. rewrite <- h3. apply plus_set_functional. reflexivity. Qed. Lemma el_prod_covers : forall (E:Ensemble At) (pf:Finite E), plus_set (Im (Full_set (Fin_map E signe mns)) el_prod) (finite_image _ _ _ _ (finite_fin_maps _ _ mns pf signe_finite)) = 1. intros E h1. assert (h2:times_plus_fin_pair_map1 ba_eq_dec h1 (fun_to_fin_map (impl_type_eq_dec_prod ba_eq_dec sign_eq_dec) _ 0 (cart_prod_fin _ _ h1 signe_finite) eps') = 1). unfold times_plus_fin_pair_map1. unfold fin_map_times. match goal with |- times_set _ ?pf = _ => generalize pf end. unfold plus_fin_pair_map1. rewrite im_fin_map_app_undoes_fun_to_fin_map. intros h2. match goal with |- times_set ?R _ = _ => pose R as S end. fold S. assert (h3:S = Im E (fun x:At => plus_set (Im signe (fun y => (eps' (x, y)))) (finite_image _ _ _ _ signe_finite))). apply im_ext_in. intros x h. apply plus_set_functional. apply im1_fun_to_fin_map; auto. revert h2. fold S. unfold At, bt in h3. unfold At, bt. rewrite h3. intro ha. unfold At, bt in h1. unfold At, bt. rewrite <- times_set_sing. destruct (finite_inh_or_empty _ h1) as [hinh | hninh]. apply times_set_functional. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2 q hq]. rewrite hq. clear hq q. match goal with |- Inn _ ?y => pose y as a end. pose proof (eps_eps'_compat x) as h0. unfold bt, At in h0. pose proof (eps_covers _ h1 x) as h8. unfold At, bt in h8. rewrite h8 at 1. constructor. destruct hinh as [e hinh]. pose proof (eps_covers _ h1 e) as hc. red. intros x h2. destruct h2. apply Im_intro with e. assumption. rewrite <- hc. apply plus_set_functional. apply im_ext_in. intros s hin. rewrite eps_eps'_compat. reflexivity. rewrite times_set_sing. assert (h5:S = Empty_set _). unfold S in h3. unfold S. subst. rewrite image_empty. reflexivity. unfold S in h3. subst. rewrite h3 in h5. rewrite <- times_set_empty. apply times_set_functional; auto. rewrite <- h2. pose proof (complete_dist_times_plus1 ba_eq_dec sign_eq_dec _ _ mns h1 signe_finite eps') as hc. simpl in hc. unfold At in hc. unfold At. rewrite hc at 1. unfold plus_times_fun_all_maps1. apply plus_set_functional. f_equal. apply functional_extensionality. intro F. unfold el_prod. unfold times_fun_fin_map1. unfold fin_map_times. apply times_set_functional. apply Extensionality_Ensembles. red. split. intro y. intros h3. destruct h3 as [e h4 y h5]. apply Im_intro with e. assumption. rewrite fun_to_fin_map_compat. rewrite fun_to_fin_map_compat. unfold eps'. simpl. assumption. split. simpl. split; auto. apply fin_map_app_in. assumption. assumption. red. intros y h3. destruct h3 as [e h4 y h5]. rewrite fun_to_fin_map_compat in h5. rewrite fun_to_fin_map_compat in h5. apply Im_intro with e. assumption. unfold eps' in h5. simpl in h5. assumption. split. simpl. split; auto. apply fin_map_app_in. assumption. assumption. Qed. Lemma el_prod_covers' : forall (E:Ensemble At) (pf:Finite E) (pf:Finite (Im (Full_set (Fin_map E signe mns)) el_prod)), plus_set (Im (Full_set (Fin_map E signe mns)) el_prod) pf = 1. intros E h1 h2. pose proof (el_prod_covers _ h1) as h3. rewrite <- h3. apply plus_set_functional. reflexivity. Qed. Lemma el_prod_ex_non_zero_el_prod_compose : forall {B:Bool_Alg} {E:Ensemble (bt B)} (g:bt B->At) (a:Fin_map E signe mns), el_prod_compose g a <> 0 -> exists a':Fin_map (Im E g) signe mns, el_prod a' = el_prod_compose g a. intros B E g a h0. pose proof (non_zero_el_prod_compose_constant _ _ h0) as h0'. pose proof (partition_fam_rel_classes_im_rel_set E g) as h1. red in h1. destruct h1 as [h1 h2]. symmetry in h2. pose proof (im_rel_classes_im_rel_set (fin_map_fin_dom a) g) as h3. assert (h4:forall y, Inn (Im E g) y -> exists! s:sign, forall x:(bt B), Inn E x -> y = g x -> s = a |-> x). intros y h4. destruct h4 as [x h4]. clear h2. subst. exists (a |-> x). red. split. intros y h5. apply h0'; auto. intros s h5. specialize (h5 _ h4 eq_refl). rewrite h5. reflexivity. pose (fun i:sig_set (Im E g) => (proj1_sig (constructive_definite_description _ (h4 _ (proj2_sig i))))) as f. pose (sig_fun_to_fin_map ba_eq_dec f (finite_image _ _ _ g (fin_map_fin_dom a)) mns) as a'. assert (h5:Included (Im (full_sig (Im E g)) f) signe). red. intros x h5. destruct x; constructor. pose (fin_map_new_ran a' signe_finite h5) as a''. exists a''. unfold el_prod, el_prod_compose. apply times_set_functional. rewrite im_im. apply im_ext_in. intros x h6. f_equal. unfold a''. rewrite <- fin_map_new_ran_compat. unfold a'. assert (h7:Inn (Im E g) (g x)). apply Im_intro with x. assumption. reflexivity. pose proof (sig_fun_to_fin_map_compat ba_eq_dec f (finite_image (bt B) At E g (fin_map_fin_dom a)) mns _ h7) as he. unfold At in he. unfold At. rewrite he at 1. unfold f. simpl. destruct constructive_definite_description as [s h8]. simpl. rewrite (h8 _ h6) at 1. reflexivity. reflexivity. apply Im_intro with x; auto. Qed. Lemma non_zero_el_prod_compose_ex_el_prod : forall {B:Bool_Alg} {E:Ensemble (bt B)}, Finite E -> forall (g:bt B->At) (a:Fin_map (Im E g) signe mns), el_prod a <> 0 -> exists a':Fin_map E signe mns, el_prod_compose g a' = el_prod a. intros B E h0 g a h1. pose (fun e:bt B => a |-> (g e)) as f. pose (fun_to_fin_map ba_eq_dec E mns h0 f) as a'. assert (h2:Included (Im E f) signe). red. intros x h2. destruct h2 as [x h2]. subst. destruct (f x); constructor. pose (fin_map_new_ran a' signe_finite h2) as a''. exists a''. unfold el_prod_compose, el_prod. apply times_set_functional. rewrite im_im. apply im_ext_in. intros x h3. f_equal. unfold a''. rewrite <- fin_map_new_ran_compat. unfold a'. rewrite fun_to_fin_map_compat; auto. assumption. Qed. Lemma el_prod_compose_covers : forall {B:Bool_Alg} (E:Ensemble (bt B)) (pf:Finite E) (g:(bt B)->At), plus_set (Im (Full_set (Fin_map E signe mns)) (el_prod_compose g)) (finite_image _ _ _ _ (finite_fin_maps _ _ mns pf signe_finite)) = 1. intros B E h1 g. pose proof (excl_middle_sat (fun f:(Fin_map E signe mns) => el_prod_compose g f <> 0)) as h2. let P := type of h2 in match P with _ = Union ?S _ => pose S as C end. let P := type of h2 in match P with _ = Union _ ?S => pose S as D end. fold D C in h2. pose proof (f_equal (fun S => Im S (el_prod_compose g)) h2) as h3. simpl in h3. let P := type of h3 in match P with | _ = ?S => assert (h4:Finite S) end. rewrite <- h2. apply finite_image. apply finite_fin_maps. assumption. apply signe_finite. match goal with |- plus_set _ ?pf = _ => pose pf as hyp end. pose proof (subsetT_eq_compat _ _ _ _ hyp h4 h3) as h5. fold hyp. unfold At, bt in h5. unfold At, bt. simpl in h5. simpl. dependent rewrite -> h5. pose proof (im_union C D (el_prod_compose g)) as h6. let P := type of h6 in match P with _ = ?S => assert (h7:Finite S) end. rewrite <- h6. assumption. pose proof (subsetT_eq_compat _ _ _ _ h4 h7 h6) as h8. unfold At, bt in h8. unfold At, bt. simpl in h8. simpl. dependent rewrite -> h8. assert (h9:Finite (Im C (el_prod_compose g))). eapply Finite_downward_closed. apply h7. auto with sets. assert (h10:Finite (Im D (el_prod_compose g))). eapply Finite_downward_closed. apply h7. auto with sets. pose proof (plus_set_union' _ _ h9 h10 h7) as h11. unfold At, bt in h11. unfold At, bt. rewrite h11 at 1. assert (h12:plus_set (Im D (el_prod_compose g)) h10 = 0). destruct (classic (Inhabited D)) as [h13 | h14]. destruct h13 as [x h13]. destruct h13 as [h13]. rewrite <- (plus_set_sing 0). apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros y h14. destruct h14 as [y h14]. subst. destruct h14 as [h14]. apply NNPP in h14. rewrite h14. constructor. red. intros y h14. destruct h14. apply Im_intro with x. constructor. assumption. apply NNPP in h13. rewrite h13. reflexivity. apply not_inhabited_empty in h14. pose proof (image_empty (Fin_map E signe mns) _ (el_prod_compose g)) as h15. rewrite <- h14 in h15. pose proof (Empty_is_finite At) as h16. pose proof (subsetT_eq_compat _ _ _ _ h10 h16 h15) as h17. dependent rewrite -> h17. apply plus_set_empty'. unfold At, bt in h12. unfold At, bt. rewrite h12 at 1. rewrite zero_sum. pose proof (finite_image _ _ _ g h1) as h0. assert (h2':times_plus_fin_pair_map1 ba_eq_dec h0 (fun_to_fin_map (impl_type_eq_dec_prod ba_eq_dec sign_eq_dec) _ 0 (cart_prod_fin (Im E g) _ h0 signe_finite) eps') = 1). unfold times_plus_fin_pair_map1, fin_map_times. match goal with |- times_set _ ?pf = _ => generalize pf end. unfold plus_fin_pair_map1. rewrite im_fin_map_app_undoes_fun_to_fin_map. intros h2'. match goal with |- times_set ?S _ = _ => pose S as E' end. fold E'. assert (h3': E' = Im (Im E g) (fun x:At => plus_set (Im signe (fun y => (eps' (x, y)))) (finite_image _ _ _ _ signe_finite))). unfold E'. do 2 rewrite im_im. apply Extensionality_Ensembles. red. split. red. intros b h3'. destruct h3' as [a h4' b h5']. apply Im_intro with a. assumption. rewrite h5'. apply plus_set_functional. apply im1_fun_to_fin_map; auto. apply Im_intro with a; auto. red. intros b h3'. destruct h3' as [a h4' b h5']. apply Im_intro with a. assumption. rewrite h5'. apply plus_set_functional. symmetry. apply im1_fun_to_fin_map; auto. apply Im_intro with a; auto. unfold At, bt in h3'. unfold At, bt. fold E' in h2'. revert h2'. rewrite h3'. clear h3'. clear h12 h11 h10 h9 h8 h7 h6 h5 h4 h3 h2. pose proof (finite_inh_or_empty _ h1) as h1'. destruct h1' as [hinh | hninh]. destruct hinh as [e hin]. intro h2'. rewrite <- times_set_sing. apply times_set_functional. rewrite im_im. apply Extensionality_Ensembles. red. split. red. intros a h6'. destruct h6' as [x h6' a h7']. assert (heps: (fun y : sign => eps' (g x, y)) = eps (g x)). apply functional_extensionality. intro s. rewrite eps_eps'_compat. reflexivity. pose proof (eps_covers _ h0 (g x)) as h8'. rewrite <- h8'. rewrite <- h7' at 1. constructor. red; intros x ho. destruct ho. apply Im_intro with e; auto. pose proof (eps_covers _ h0 (g e)) as h8'. rewrite <- h8'. f_equal. subst. intros h3. rewrite <- times_set_empty. apply times_set_functional. rewrite im_im. rewrite image_empty; auto. unfold bt in h2'. unfold bt. rewrite <- h2'. pose proof (complete_dist_times_plus1 ba_eq_dec sign_eq_dec _ _ mns h0 signe_finite eps') as he. simpl in he. unfold At in he. unfold At. rewrite he at 1. unfold plus_times_fun_all_maps1. pose proof (excl_middle_sat (fun f:(Fin_map (Im E g) signe mns) => el_prod f <> 0)) as h2''. let P := type of h2'' in match P with _ = Union ?S _ => pose S as C'' end. let P := type of h2'' in match P with _ = Union _ ?S => pose S as D'' end. fold D'' C'' in h2''. pose proof (f_equal (fun S => Im S (times_fun_fin_map1 sign_eq_dec eps')) h2'') as h3''. simpl in h3''. let P := type of h3'' in match P with _ = ?S => assert (h4'': Finite S) end. rewrite <- h2''. apply finite_image. apply finite_fin_maps. assumption. apply signe_finite. match goal with |- _ = plus_set _ ?pf => pose pf as hf end. fold hf. pose proof (subsetT_eq_compat _ _ _ _ hf h4'' h3'') as h5''. unfold At, bt in h5''. unfold At, bt. simpl in h5''. simpl. pose proof (proof_irrelevance _ h0 (finite_image _ _ E g h1)) as ho. dependent rewrite -> h5''. pose proof (im_union C'' D'' (times_fun_fin_map1 sign_eq_dec eps')) as h6''. unfold At, bt in h6''. unfold At, bt. let P := type of h6'' in match P with _ = ?S => assert (h7'':Finite S) end. unfold At, bt. rewrite <- h6''. assumption. pose proof (subsetT_eq_compat _ _ _ _ h4'' h7'' h6'') as h8''. dependent rewrite -> h8''. assert (h9'':Finite (Im C'' (times_fun_fin_map1 sign_eq_dec eps'))). eapply Finite_downward_closed. apply h7''. auto with sets. assert (h10'':Finite (Im D'' (times_fun_fin_map1 sign_eq_dec eps'))). eapply Finite_downward_closed. apply h7''. auto with sets. pose proof (plus_set_union' _ _ h9'' h10'' h7'') as h11''. unfold At, bt in h11''. unfold At, bt. rewrite h11'' at 1. assert (h12'': plus_set (Im D'' (times_fun_fin_map1 sign_eq_dec eps')) h10'' = 0). destruct (classic (Inhabited [x : Fin_map (Im E g) signe mns | ~ el_prod x <> 0])) as [h13 | h14]. destruct h13 as [x h13]. destruct h13 as [h13]. rewrite <- (plus_set_sing 0). apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros y h14. destruct h14 as [y h14 q hq]. rewrite hq. clear hq q. destruct h14 as [h14]. apply NNPP in h14. assert (h15:times_fun_fin_map1 sign_eq_dec eps' y = 0). unfold times_fun_fin_map1. unfold fin_map_times. unfold el_prod in h14. rewrite <- h14. apply times_set_functional. apply im_ext_in. intros d h16. rewrite fun_to_fin_map_compat. rewrite fun_to_fin_map_compat. rewrite eps_eps'_compat. reflexivity. constructor. simpl. split; auto. apply fin_map_app_in. simpl in h16. simpl. apply h16. destruct h16 as [c h16 q hq]. rewrite hq. clear hq q. apply Im_intro with c; auto. rewrite h15. constructor. red. intros d h14. destruct h14. apply Im_intro with x. constructor. assumption. unfold times_fun_fin_map1. unfold fin_map_times. unfold el_prod in h13. apply NNPP in h13. rewrite <- h13. apply times_set_functional. apply im_ext_in. intros d h16. rewrite fun_to_fin_map_compat. rewrite fun_to_fin_map_compat. rewrite eps_eps'_compat. reflexivity. constructor. simpl. split; auto. apply fin_map_app_in. simpl in h16. simpl. apply h16. destruct h16 as [c h16 q hq]. rewrite hq. clear q hq. apply Im_intro with c; auto. apply not_inhabited_empty in h14. pose proof (Empty_is_finite (Fin_map (Im E g) signe mns)) as h15. rewrite <- h14 in h15. apply finite_image with (f:=times_fun_fin_map1 sign_eq_dec eps') in h15. pose proof (feq_im _ _ (times_fun_fin_map1 sign_eq_dec eps') h14) as h16. rewrite image_empty in h16. pose proof (subsetT_eq_compat _ _ _ _ h10'' (Empty_is_finite _) h16) as h17. fold D'' in h17. unfold bt, At in h17. unfold bt, At. simpl in h17. simpl. dependent rewrite -> h17. rewrite plus_set_empty'. reflexivity. unfold At, bt in h12''. unfold At, bt. rewrite h12'' at 1. rewrite zero_sum. apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros x h13. destruct h13 as [f h13 q hq]. rewrite hq. clear hq q. destruct h13 as [h13]. pose proof (el_prod_ex_non_zero_el_prod_compose g f h13) as h14. destruct h14 as [f' h14]. apply Im_intro with f'. constructor. unfold bt, At in h14. unfold bt, At. simpl in h14. simpl. simpl. rewrite h14 at 1. assumption. rewrite <- h14. unfold el_prod, times_fun_fin_map1. unfold fin_map_times. apply times_set_functional. apply im_ext_in. intros x h15. rewrite fun_to_fin_map_compat. rewrite fun_to_fin_map_compat. rewrite eps_eps'_compat. reflexivity. constructor. simpl. split; auto. apply fin_map_app_in. destruct h15 as [x h15 q hq]. rewrite hq. clear hq q. apply Im_intro with x. assumption. reflexivity. assumption. red. intros x h13. destruct h13 as [f h13 q hq]. rewrite hq. clear hq q. destruct h13 as [h13]. pose proof (non_zero_el_prod_compose_ex_el_prod h1 g f h13) as h14. destruct h14 as [f' h14]. apply Im_intro with f'. constructor. rewrite h14. assumption. rewrite h14. unfold times_fun_fin_map1, el_prod. unfold fin_map_times. apply times_set_functional. apply im_ext_in. intros x h15. rewrite fun_to_fin_map_compat. rewrite fun_to_fin_map_compat. rewrite eps_eps'_compat. reflexivity. constructor. simpl. split; auto. apply fin_map_app_in. destruct h15 as [x h15 q hq]. rewrite hq. clear hq q. apply Im_intro with x. assumption. reflexivity. assumption. Qed. Definition non_zero_el_prod_maps (E:Ensemble At) := [a:Fin_map E signe mns | el_prod a <> 0]. Lemma non_zero_el_prod_maps_fin : forall (E:Ensemble At), Finite E -> Finite (non_zero_el_prod_maps E). intros E h1. eapply Finite_downward_closed. apply finite_fin_maps; auto. apply signe_finite. red; intros; constructor. Qed. Definition non_zero_el_prod_compose_maps {B:Bool_Alg} (g:(bt B)->At) (E:Ensemble (bt B)) := [a:Fin_map E signe mns | el_prod_compose g a <> 0]. Lemma non_zero_el_prod_compose_maps_fin : forall {B:Bool_Alg} {E:Ensemble (bt B)} (g:bt B->At), Finite E -> Finite (non_zero_el_prod_compose_maps g E). intros B E g h1. eapply Finite_downward_closed. apply finite_fin_maps; auto. apply signe_finite. red; intros; constructor. Qed. Definition plus_subset_non_zero_el_prod_maps (E:Ensemble At) (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_maps E)) : At := plus_set (Im S el_prod) (finite_image _ _ _ el_prod (Finite_downward_closed _ _ (non_zero_el_prod_maps_fin _ pfe) _ pfi)). Definition plus_subset_non_zero_el_prod_compose_maps {B:Bool_Alg} (E:Ensemble (bt B)) (pfe:Finite E) (g:(bt B)->At) (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_compose_maps g E)) : At := plus_set (Im S (el_prod_compose g)) (finite_image _ _ _ (el_prod_compose g) (finite_fin_map_ens S pfe signe_finite)). Lemma inclusion_plus_subset_non_zero_el_prod_maps_preserves_int : forall (E:Ensemble At) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_maps E) -> Included Y (non_zero_el_prod_maps E) -> Included (Intersection X Y) (non_zero_el_prod_maps E). intros E h1 X Y h2 h3. red. intros a h4. destruct h4 as [a h4l h4r]. constructor. red in h3. specialize (h3 _ h4r). destruct h3. assumption. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_maps_preserves_union : forall (E:Ensemble At) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_maps E) -> Included Y (non_zero_el_prod_maps E) -> Included (Union X Y) (non_zero_el_prod_maps E). intros E h1 X Y h2 h3. red. intros a h4. destruct h4 as [a h4l | a h4r]. red in h2. specialize (h2 _ h4l). assumption. red in h3. specialize (h3 _ h4r). assumption. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp : forall (E:Ensemble At) (pfe:Finite E) (X:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_maps E) -> Included (Setminus (non_zero_el_prod_maps E) X) (non_zero_el_prod_maps E). intros E h1 X h2. red. intros a h4. destruct h4 as [h4l]. assumption. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_maps_preserves_setminus : forall (E:Ensemble At) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_maps E) -> Included (Setminus X Y) (non_zero_el_prod_maps E). intros E h1 X Y h2. red. intros a h3. destruct h3 as [h3a h3b]. auto with sets. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_maps_preserves_symdiff_full : forall (E:Ensemble At) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_maps E) -> Included Y (non_zero_el_prod_maps E) -> Included (Symdiff_full (non_zero_el_prod_maps E) X Y) (non_zero_el_prod_maps E). intros E h1 X Y h2 h3. red. intros a h4. destruct h4 as [a h4l | a h4r]. destruct h4l as [a h4a h4b]. auto with sets. destruct h4r as [a h4a h4b]. auto with sets. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_union : forall {B:Bool_Alg} {E:Ensemble (bt B)} (pfe:Finite E) (g:(bt B)->At) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_compose_maps g E) -> Included Y (non_zero_el_prod_compose_maps g E) -> Included (Union X Y) (non_zero_el_prod_compose_maps g E). intros B E h1 g X Y h2 h3. red. intros a h4. destruct h4 as [a h4l | a h4r]. red in h2. specialize (h2 _ h4l). assumption. red in h3. specialize (h3 _ h4r). assumption. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_int : forall {B:Bool_Alg} {E:Ensemble (bt B)} (pfe:Finite E) (g:(bt B)->At) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_compose_maps g E) -> Included Y (non_zero_el_prod_compose_maps g E) -> Included (Intersection X Y) (non_zero_el_prod_compose_maps g E). intros B E h1 g X Y h2 h3. red. intros a h4. destruct h4 as [a h4l h4r]. constructor. red in h3. specialize (h3 _ h4r). destruct h3. assumption. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_comp : forall {B:Bool_Alg} {E:Ensemble (bt B)} (pfe:Finite E) (g:(bt B)->At) (X:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_compose_maps g E) -> Included (Setminus (non_zero_el_prod_compose_maps g E) X) (non_zero_el_prod_compose_maps g E). intros B E h1 g X h2. red. intros a h4. destruct h4 as [h4l]. assumption. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_symdiff_full : forall {B:Bool_Alg} {E:Ensemble (bt B)} (pfe:Finite E) (g:(bt B)->At) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_compose_maps g E) -> Included Y (non_zero_el_prod_compose_maps g E) -> Included (Symdiff_full (non_zero_el_prod_compose_maps g E) X Y) (non_zero_el_prod_compose_maps g E). intros B E h1 g X Y h2 h3. red. intros a h4. destruct h4 as [a h4l | a h4r]. destruct h4l as [a h4a h4b]. auto with sets. destruct h4r as [a h4a h4b]. auto with sets. Qed. Lemma plus_subset_non_zero_el_prod_maps_empty : forall (E:Ensemble At) (pfe:Finite E), plus_subset_non_zero_el_prod_maps E pfe _ (incl_empty _) = 0. intros. unfold plus_subset_non_zero_el_prod_maps. match goal with |- plus_set _ ?pf = _ => generalize pf end. rewrite image_empty. intros; apply plus_set_empty'. Qed. Lemma le_plus_subset_non_zero_el_prod_maps : forall (E:Ensemble At) (pfe:Finite E) (X:Ensemble (Fin_map E signe mns)) (pf:Included X (non_zero_el_prod_maps E)) (a:Fin_map E signe mns), Inn X a -> le (el_prod a) (plus_subset_non_zero_el_prod_maps E pfe _ pf). intros E h1 X h2 a h3. unfold plus_subset_non_zero_el_prod_maps. apply le_plus_set. apply Im_intro with a; auto. Qed. Lemma le_plus_subset_non_zero_el_prod_compose_maps : forall {B:Bool_Alg} (g:bt B->At) (E:Ensemble (bt B)) (pfe:Finite E) (X:Ensemble (Fin_map E signe mns)) (pf:Included X (non_zero_el_prod_compose_maps g E)) (a:Fin_map E signe mns), Inn X a -> le (el_prod_compose g a) (plus_subset_non_zero_el_prod_compose_maps E pfe _ X pf). intros B g E h1 X h2 a h3. unfold plus_subset_non_zero_el_prod_compose_maps. apply le_plus_set. apply Im_intro with a; auto. Qed. Lemma plus_subset_non_zero_el_prod_maps_empty_rev : forall (E:Ensemble At) (pfe:Finite E) (X:Ensemble (Fin_map E signe mns)) (pf:Included X (non_zero_el_prod_maps E)), plus_subset_non_zero_el_prod_maps E pfe _ pf = 0 -> X = Empty_set _. intros E h1 X h2 h3. apply NNPP. intros h4. pose proof (not_empty_Inhabited _ _ h4) as h5. destruct h5 as [a h5]. unfold non_zero_el_prod_maps in h2. assert (h6:Inn [a:Fin_map E signe mns | el_prod a <> 0] a). auto with sets. destruct h6 as [h6]. pose proof (le_plus_subset_non_zero_el_prod_maps _ h1 _ h2 _ h5) as h7. contradict h6. rewrite h3 in h7. apply le_x_0. assumption. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_empty_rev : forall {B:Bool_Alg} (g:bt B->At) (E:Ensemble (bt B)) (pfe:Finite E) (X:Ensemble (Fin_map E signe mns)) (pf:Included X (non_zero_el_prod_compose_maps g E)), plus_subset_non_zero_el_prod_compose_maps E pfe g _ pf = 0 -> X = Empty_set _. intros B g E h1 X h2 h3. apply NNPP. intros h4. pose proof (not_empty_Inhabited _ _ h4) as h5. destruct h5 as [a h5]. unfold non_zero_el_prod_compose_maps in h2. assert (h6:Inn [a:Fin_map E signe mns | el_prod_compose g a <> 0] a). auto with sets. destruct h6 as [h6]. pose proof (le_plus_subset_non_zero_el_prod_compose_maps g _ h1 _ h2 _ h5) as h7. contradict h6. rewrite h3 in h7. apply le_x_0. assumption. Qed. Lemma plus_set_el_prod_zero : forall (E:Ensemble At) (pf:Finite _), plus_set (Im [x : Fin_map E signe mns | el_prod x = 0] el_prod) pf = 0. intros E h1. destruct (classic (Im [x : Fin_map E signe mns | el_prod x = 0] el_prod = Empty_set _)) as [h2 | h3]. generalize h1. rewrite h2. intro. apply plus_set_empty'. pose proof (not_empty_Inhabited _ _ h3) as h4. rewrite <- plus_set_sing. apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros x h15. inversion h15 as [a h16 y]. subst. inversion h15 as [b h17 z]. subst. destruct h17 as [h17]. rewrite <- H in h17. rewrite h17. constructor. red. intros x h15. destruct h15. subst. destruct h4 as [a h4]. inversion h4. subst. inversion h4. subst. inversion H0. rewrite H2 in H1. rewrite <- H1. apply Im_intro with x. constructor. reflexivity. reflexivity. Qed. Lemma plus_set_el_prod_compose_zero : forall {B:Bool_Alg} (g:bt B->At) (E:Ensemble (bt B)) (pf:Finite _), plus_set (Im [x : Fin_map E signe mns | el_prod_compose g x = 0] (el_prod_compose g)) pf = 0. intros B g E h1. destruct (classic (Im [x : Fin_map E signe mns | el_prod_compose g x = 0] (el_prod_compose g) = Empty_set _)) as [h2 | h3]. generalize h1. rewrite h2. intro. apply plus_set_empty'. pose proof (not_empty_Inhabited _ _ h3) as h4. rewrite <- plus_set_sing. apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros x h15. inversion h15 as [a h16 y]. subst. inversion h15 as [b h17 z]. subst. destruct h17 as [h17]. rewrite <- H in h17. rewrite h17. constructor. red. intros x h15. destruct h15. subst. destruct h4 as [a h4]. inversion h4. subst. inversion h4. subst. inversion H0. rewrite H2 in H1. rewrite <- H1. apply Im_intro with x. constructor. reflexivity. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_maps_full : forall (E:Ensemble At) (pfe:Finite E), plus_subset_non_zero_el_prod_maps E pfe _ (inclusion_reflexive _) = 1. intros E h1. unfold plus_subset_non_zero_el_prod_maps. unfold non_zero_el_prod_maps. pose proof (excl_middle_sat (fun a:(Fin_map E signe mns) => el_prod a = 0)) as h2. pose proof f_equal. pose proof (feq_im _ _ el_prod h2) as h3. simpl in h3. rewrite im_union in h3. pose proof (finite_fin_maps _ _ mns h1 signe_finite) as h4. pose proof (finite_image _ _ _ el_prod h4) as h5. pose proof h5 as h6. rewrite h3 in h6. pose proof f_equal. pose proof (plus_set_functional _ _ h3 h5 h6) as h7. pose proof (el_prod_covers' _ h1 h5) as h8. rewrite h8 in h7. symmetry in h7. pose proof (plus_set_union). assert (h9:Included (Im [x : Fin_map E signe mns | el_prod x = 0] el_prod) (Im (Full_set (Fin_map E signe mns)) el_prod)). rewrite h3. auto with sets. assert (h10:Included (Im [x : Fin_map E signe mns | el_prod x <> 0] el_prod) (Im (Full_set (Fin_map E signe mns)) el_prod)). rewrite h3. auto with sets. pose proof (Finite_downward_closed _ _ h5 _ h9) as h11. pose proof (Finite_downward_closed _ _ h5 _ h10) as h12. pose proof (plus_set_union _ _ h11 h12) as h13. pose proof (proof_irrelevance _ h6 (Union_preserves_Finite (Btype (Bc A)) (Im [x : Fin_map E signe mns | el_prod x = 0] el_prod) (Im [x : Fin_map E signe mns | el_prod x <> 0] el_prod) h11 h12)); subst. unfold At, bt in h7. unfold At, bt in h13. rewrite h7 in h13. pose proof plus_subset_non_zero_el_prod_maps_empty. pose proof (plus_set_el_prod_zero _ h11) as h15. unfold At in h15. unfold At in h13. rewrite h15 in h13. rewrite comm_sum, zero_sum in h13. rewrite h13. apply plus_set_functional; auto. Qed. Lemma plus_subset_non_zero_el_prod_maps_int : forall (E:Ensemble At) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_maps E)) (pfy:Included Y (non_zero_el_prod_maps E)), (plus_subset_non_zero_el_prod_maps E pfe _ pfx) * (plus_subset_non_zero_el_prod_maps E pfe _ pfy) = plus_subset_non_zero_el_prod_maps E pfe (Intersection X Y) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_int _ pfe X Y pfx pfy). intros E h1 X Y h2 h3. unfold plus_subset_non_zero_el_prod_maps. rewrite dist_set_plus2. pose proof (cart_prod_im X Y el_prod el_prod) as h4. pose proof (excl_middle_sat' (cart_prod X Y) (fun pr => fst pr = snd pr)) as h5. let P := type of h5 in match P with _ = Union ?S _ => pose S as C end. let P := type of h5 in match P with _ = Union _ ?S => pose S as D end. simpl in C, D. fold C D in h5. rewrite h5 in h4. clear h5. rewrite im_union in h4. pose (fun p : Btype (Bc A) * Btype (Bc A) => fst p * snd p) as f. pose (fun pr : Fin_map E signe mns * Fin_map E signe mns => (el_prod (fst pr), el_prod (snd pr))) as psi. pose (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod (fst x) * el_prod (snd x)) as psi'. pose proof (feq_im _ _ f h4) as h5. fold psi in h5, h4. rewrite im_union in h5. do 2 rewrite im_im in h5. simpl in h5. pose proof (finite_fin_map_ens X h1 signe_finite) as h6. pose proof (finite_fin_map_ens Y h1 signe_finite) as h7. pose proof (finite_image _ _ _ el_prod h6) as h8. pose proof (finite_image _ _ _ el_prod h7) as h9. pose proof (cart_prod_fin _ _ h8 h9) as h10. pose proof (finite_image _ _ _ f h10) as h11. rewrite h5 in h11. fold C D in h11. match goal with |- plus_set _ ?pf = _ => pose pf as hf end. fold hf. pose proof (plus_set_functional _ _ h5 hf h11) as h12. assert (h13:Included (Im C psi) (Union (Im C psi) (Im D psi))). auto with sets. assert (h14:Included (Im D psi) (Union (Im C psi) (Im D psi))). auto with sets. assert (h13':Included (Im C psi') (Union (Im C psi') (Im D psi'))). auto with sets. assert (h14':Included (Im D psi') (Union (Im C psi') (Im D psi'))). auto with sets. pose proof (Finite_downward_closed _ _ h11 _ h13') as h15. clear h13'. pose proof (Finite_downward_closed _ _ h11 _ h14') as h16. clear h14'. simpl in h12. unfold f, psi in h12. simpl in h12. fold C D psi' in h12. pose proof (finite_fin_map_squared_ens D h1 signe_finite) as h13'. pose proof (plus_set_el_prod_disjoint' _ X Y h13' h16) as hp. fold D psi' in hp. rewrite (plus_set_union' _ _ h15 h16 h11) in h12 at 1. rewrite hp in h12. rewrite zero_sum in h12. unfold At in h12. unfold At. simpl in At. simpl. rewrite h12 at 1. apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros x h17. inversion h17 as [a h18 b h19 h20]. subst. destruct h18 as [h18]. destruct h18 as [h18l h18r]. destruct h18l as [h18l]. destruct h18l as [h18a h18b]. unfold psi'. rewrite h18r. rewrite idem_prod. rewrite h18r in h18a. apply Im_intro with (snd a). split; auto. reflexivity. red. intros x h17. inversion h17 as [a h18 b h19 h20]. subst. destruct h18 as [a h18l h18r]. apply Im_intro with (a, a). constructor. simpl. split; auto. split; auto. simpl. unfold psi'. rewrite idem_prod. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_maps_union : forall (E:Ensemble At) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_maps E)) (pfy:Included Y (non_zero_el_prod_maps E)), (plus_subset_non_zero_el_prod_maps E pfe _ pfx) + (plus_subset_non_zero_el_prod_maps E pfe _ pfy) = plus_subset_non_zero_el_prod_maps E pfe (Union X Y) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_union _ pfe X Y pfx pfy). intros E h1 X Y h2 h3. unfold plus_subset_non_zero_el_prod_maps. match goal with |- _ = plus_set _ ?pf => generalize pf end. rewrite <- plus_set_union. intro h4. apply plus_set_functional. rewrite im_union. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_maps_comp : forall (E:Ensemble At) (pfe:Finite E) (X:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_maps E)), - (plus_subset_non_zero_el_prod_maps E pfe _ pfx) = plus_subset_non_zero_el_prod_maps E pfe _ (inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp _ pfe X pfx). intros E h1 X h2. symmetry. apply comp_char. rewrite plus_subset_non_zero_el_prod_maps_int. rewrite <- (plus_subset_non_zero_el_prod_maps_empty E h1). unfold plus_subset_non_zero_el_prod_maps. apply plus_set_functional. apply feq_im. apply int_setminus. rewrite plus_subset_non_zero_el_prod_maps_union. rewrite <- (plus_subset_non_zero_el_prod_maps_full E h1). unfold plus_subset_non_zero_el_prod_maps. apply plus_set_functional. apply feq_im. apply union_setminus_incl; auto. Qed. Lemma plus_subset_non_zero_el_prod_maps_setminus : forall (E:Ensemble At) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_maps E)) (pfy:Included Y (non_zero_el_prod_maps E)), (plus_subset_non_zero_el_prod_maps E pfe _ pfx) * - (plus_subset_non_zero_el_prod_maps E pfe _ pfy) = plus_subset_non_zero_el_prod_maps E pfe (Intersection X (Setminus (non_zero_el_prod_maps E) Y)) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_int _ pfe X _ pfx (inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp _ pfe Y pfy)). intros E h1 X Y h2 h3. rewrite <- plus_subset_non_zero_el_prod_maps_int. rewrite <- plus_subset_non_zero_el_prod_maps_comp. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_maps_symdiff_full : forall (E:Ensemble At) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_maps E)) (pfy:Included Y (non_zero_el_prod_maps E)), (plus_subset_non_zero_el_prod_maps E pfe _ pfx) /_\ (plus_subset_non_zero_el_prod_maps E pfe _ pfy) = plus_subset_non_zero_el_prod_maps E pfe (Symdiff_full (non_zero_el_prod_maps E) X Y) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_symdiff_full _ pfe X Y pfx pfy). intros E h1 X Y h2 h3. unfold sym_diff. do 2 rewrite plus_subset_non_zero_el_prod_maps_comp. do 2 rewrite plus_subset_non_zero_el_prod_maps_int. rewrite plus_subset_non_zero_el_prod_maps_union. unfold Symdiff_full. unfold Setminus_full. unfold plus_subset_non_zero_el_prod_maps. apply plus_set_functional. reflexivity. Qed. Definition non_zero_el_prod_maps_of_set {E:Ensemble At} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)) := Intersection S (non_zero_el_prod_maps E). Lemma non_zero_el_prod_maps_of_set_eq : forall {E:Ensemble At} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)), non_zero_el_prod_maps_of_set pfe S = [f:Fin_map E signe mns | Inn S f /\ el_prod f <> 0]. intros E h1 S. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [f h2 h3]. destruct h3 as [h3]. constructor. split; auto. red. intros x h2. destruct h2 as [h2]. destruct h2 as [h2a h2b]. constructor; auto. constructor. assumption. Qed. Lemma non_zero_el_prod_maps_of_set_decompose : forall {E:Ensemble At} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)), S = Union (non_zero_el_prod_maps_of_set pfe S) [f:Fin_map E signe mns | Inn S f /\ el_prod f = 0]. intros E h1 S. apply Extensionality_Ensembles. red. split. red. intros f h2. destruct (classic (el_prod f = 0)) as [h3 | h4]. right. constructor. split; auto. left. constructor; auto. constructor; auto. red. intros f h2. destruct h2 as [f h2 | f h3]. destruct h2; auto. destruct h3 as [h3]. destruct h3; auto. Qed. Lemma incl_non_zero_el_prod_maps_of_set : forall {E:Ensemble At} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)), Included (non_zero_el_prod_maps_of_set pfe S) (non_zero_el_prod_maps E). intros E h1 S. unfold non_zero_el_prod_maps_of_set. auto with sets. Qed. Lemma finite_non_zero_el_prod_maps_of_set : forall {E:Ensemble At} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)), Finite (non_zero_el_prod_maps_of_set pfe S). intros E h1 S. eapply Finite_downward_closed. apply (non_zero_el_prod_maps_fin _ h1). apply incl_non_zero_el_prod_maps_of_set. Qed. Definition non_zero_el_prod_compose_maps_of_set {B:Bool_Alg} {E:Ensemble (bt B)} (pfe:Finite E) (g:bt B->At) (S:Ensemble (Fin_map E signe mns)) := Intersection S (non_zero_el_prod_compose_maps g E). Lemma non_zero_el_prod_compose_maps_of_set_eq : forall {B:Bool_Alg} {E:Ensemble (bt B)} (pfe:Finite E) (g:bt B->At) (S:Ensemble (Fin_map E signe mns)), non_zero_el_prod_compose_maps_of_set pfe g S = [f:Fin_map E signe mns | Inn S f /\ el_prod_compose g f <> 0]. intros B E h1 g S. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [f h2 h3]. destruct h3 as [h3]. constructor. split; auto. red. intros x h2. destruct h2 as [h2]. destruct h2 as [h2a h2b]. constructor; auto. constructor. assumption. Qed. Lemma non_zero_el_prod_compose_maps_of_set_decompose : forall {B:Bool_Alg} {E:Ensemble (bt B)} (pfe:Finite E) (g:bt B->At) (S:Ensemble (Fin_map E signe mns)), S = Union (non_zero_el_prod_compose_maps_of_set pfe g S) [f:Fin_map E signe mns | Inn S f /\ el_prod_compose g f = 0]. intros B E h1 g S. apply Extensionality_Ensembles. red. split. red. intros f h2. destruct (classic (el_prod_compose g f = 0)) as [h3 | h4]. right. constructor. split; auto. left. constructor; auto. constructor; auto. red. intros f h2. destruct h2 as [f h2 | f h3]. destruct h2; auto. destruct h3 as [h3]. destruct h3; auto. Qed. Lemma incl_non_zero_el_prod_compose_maps_of_set : forall {B:Bool_Alg} {E:Ensemble (bt B)} (pfe:Finite E) (g:bt B->At) (S:Ensemble (Fin_map E signe mns)), Included (non_zero_el_prod_compose_maps_of_set pfe g S) (non_zero_el_prod_compose_maps g E). intros B E h1 g S. unfold non_zero_el_prod_compose_maps_of_set. auto with sets. Qed. Lemma finite_non_zero_el_prod_compose_maps_of_set : forall {B:Bool_Alg} {E:Ensemble (bt B)} (pfe:Finite E) (g:bt B->At) (S:Ensemble (Fin_map E signe mns)), Finite (non_zero_el_prod_compose_maps_of_set pfe g S). intros B E h1 g S. eapply Finite_downward_closed. apply (non_zero_el_prod_compose_maps_fin g h1). apply incl_non_zero_el_prod_compose_maps_of_set. Qed. Lemma plus_set_zero_el_prod_maps_of_set : forall {E:Ensemble At}, Finite E -> forall (S:Ensemble (Fin_map E signe mns)) (pf:Finite (Im [f : Fin_map E signe mns | Inn S f /\ el_prod f = 0] el_prod)), plus_set _ pf = 0. intros E h1 S h2. destruct (classic (Inhabited (Im [f : Fin_map E signe mns | Inn S f /\ el_prod f = 0] el_prod))) as [h3 | h4]. destruct h3 as [f h3]. destruct h3 as [f h3]. subst. destruct h3 as [h3]. rewrite <- (plus_set_sing 0). apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros g h4. destruct h4 as [g h4]. subst. destruct h4 as [h4]. destruct h4 as [h4 h5]. rewrite h5. constructor. red. intros x h4. destruct h4. apply Im_intro with f. constructor. assumption. destruct h3 as [h3 h4]. rewrite h4. reflexivity. apply not_inhabited_empty in h4. pose proof (Empty_is_finite At) as h5. pose proof (subsetT_eq_compat _ _ _ _ h2 h5 h4) as h6. dependent rewrite -> h6. apply plus_set_empty'. Qed. Lemma plus_subset_el_prod_maps_eq_same_non_zero : forall {E:Ensemble At} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)), plus_set (Im S el_prod) (finite_image _ _ _ el_prod (finite_fin_map_ens S pfe signe_finite)) = plus_set (Im (non_zero_el_prod_maps_of_set pfe S) el_prod) (finite_image _ _ _ el_prod (finite_non_zero_el_prod_maps_of_set pfe S)). intros E h1 S. pose proof (non_zero_el_prod_maps_of_set_decompose h1 S) as h2. pose proof (f_equal (fun X=>Im X el_prod) h2) as h3. simpl in h3. assert (h4:Finite (Im (Union (non_zero_el_prod_maps_of_set h1 S) [f : Fin_map E signe mns | Inn S f /\ el_prod f = 0]) el_prod)). rewrite <- h3. apply finite_image. apply finite_fin_map_ens. assumption. apply signe_finite. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Fin_map E signe mns) At S el_prod (finite_fin_map_ens S h1 signe_finite)) h4 h3) as h5. dependent rewrite -> h5. pose proof (im_union (non_zero_el_prod_maps_of_set h1 S) [f : Fin_map E signe mns | Inn S f /\ el_prod f = 0] el_prod) as h6. assert (h7:Finite ( Union (Im (non_zero_el_prod_maps_of_set h1 S) el_prod) (Im [f : Fin_map E signe mns | Inn S f /\ el_prod f = 0] el_prod))). rewrite <- h6. assumption. pose proof (subsetT_eq_compat _ _ _ _ h4 h7 h6) as h8. dependent rewrite -> h8. assert (h9:Finite (Im (non_zero_el_prod_maps_of_set h1 S) el_prod)). eapply Finite_downward_closed. apply h7. auto with sets. assert (h10:Finite (Im [f : Fin_map E signe mns | Inn S f /\ el_prod f = 0] el_prod)). eapply Finite_downward_closed. apply h7. auto with sets. pose proof (plus_set_union' _ _ h9 h10 h7) as h11. unfold At, bt in h11. unfold At, bt. rewrite h11 at 1. rewrite plus_set_zero_el_prod_maps_of_set. rewrite zero_sum. apply plus_set_functional. reflexivity. assumption. Qed. Lemma plus_set_zero_el_prod_compose_maps_of_set : forall {B:Bool_Alg} {E:Ensemble (bt B)}, Finite E -> forall (g:bt B->At) (S:Ensemble (Fin_map E signe mns)) (pf:Finite (Im [f : Fin_map E signe mns | Inn S f /\ el_prod_compose g f = 0] (el_prod_compose g))), plus_set _ pf = 0. intros B E h1 g S h2. destruct (classic (Inhabited (Im [f : Fin_map E signe mns | Inn S f /\ el_prod_compose g f = 0] (el_prod_compose g)))) as [h3 | h4]. destruct h3 as [f h3]. destruct h3 as [f h3]. subst. destruct h3 as [h3]. rewrite <- (plus_set_sing 0). apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros k h4. destruct h4 as [k h4]. subst. destruct h4 as [h4]. destruct h4 as [h4 h5]. rewrite h5. constructor. red. intros x h4. destruct h4. apply Im_intro with f. constructor. assumption. destruct h3 as [h3 h4]. rewrite h4. reflexivity. apply not_inhabited_empty in h4. pose proof (Empty_is_finite At) as h5. pose proof (subsetT_eq_compat _ _ _ _ h2 h5 h4) as h6. dependent rewrite -> h6. apply plus_set_empty'. Qed. Lemma plus_subset_el_prod_compose_maps_eq_same_non_zero : forall {B:Bool_Alg} {E:Ensemble (bt B)} (pfe:Finite E) (g:bt B->At) (S:Ensemble (Fin_map E signe mns)), plus_set (Im S (el_prod_compose g)) (finite_image _ _ _ (el_prod_compose g) (finite_fin_map_ens S pfe signe_finite)) = plus_set (Im (non_zero_el_prod_compose_maps_of_set pfe g S) (el_prod_compose g)) (finite_image _ _ _ (el_prod_compose g) (finite_non_zero_el_prod_compose_maps_of_set pfe g S)). intros B E h1 g S. pose proof (non_zero_el_prod_compose_maps_of_set_decompose h1 g S) as h2. pose proof (f_equal (fun X=>Im X (el_prod_compose g)) h2) as h3. simpl in h3. assert (h4:Finite (Im (Union (non_zero_el_prod_compose_maps_of_set h1 g S) [f : Fin_map E signe mns | Inn S f /\ el_prod_compose g f = 0]) (el_prod_compose g))). rewrite <- h3. apply finite_image. apply finite_fin_map_ens. assumption. apply signe_finite. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Fin_map E signe mns) At S (el_prod_compose g) (finite_fin_map_ens S h1 signe_finite)) h4 h3) as h5. dependent rewrite -> h5. pose proof (im_union (non_zero_el_prod_compose_maps_of_set h1 g S) [f : Fin_map E signe mns | Inn S f /\ el_prod_compose g f = 0] (el_prod_compose g)) as h6. assert (h7:Finite ( Union (Im (non_zero_el_prod_compose_maps_of_set h1 g S) (el_prod_compose g)) (Im [f : Fin_map E signe mns | Inn S f /\ el_prod_compose g f = 0] (el_prod_compose g)))). rewrite <- h6. assumption. assert (h9:Finite (Im (non_zero_el_prod_compose_maps_of_set h1 g S) (el_prod_compose g))). eapply Finite_downward_closed. apply h7. auto with sets. assert (h10:Finite (Im [f : Fin_map E signe mns | Inn S f /\ el_prod_compose g f = 0] (el_prod_compose g))). eapply Finite_downward_closed. apply h7. auto with sets. pose proof (subsetT_eq_compat _ _ _ _ h4 h7 h6) as h8. dependent rewrite -> h8. pose proof (plus_set_union' _ _ h9 h10 h7) as h11. unfold At, bt in h11. unfold At, bt. rewrite h11 at 1. pose proof (plus_set_zero_el_prod_compose_maps_of_set h1 g _ h10) as h12. unfold bt, At in h12. unfold bt, At. rewrite h12 at 1. rewrite zero_sum. apply plus_set_functional. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_empty : forall {B:Bool_Alg} (E:Ensemble (bt B)) (pfe:Finite E) (g:(bt B)->At), plus_subset_non_zero_el_prod_compose_maps E pfe g _ (incl_empty _) = 0. intros. unfold plus_subset_non_zero_el_prod_compose_maps. pose proof (image_empty (Fin_map E signe mns) _ (el_prod_compose g)) as h1. match goal with |- plus_set _ ?pf = _ => generalize pf end. rewrite h1. intro. apply plus_set_empty'. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_full : forall {B:Bool_Alg} (E:Ensemble (bt B)) (pfe:Finite E) (g:(bt B)->At), plus_subset_non_zero_el_prod_compose_maps E pfe g _ (inclusion_reflexive _) = 1. intros B E h1 g. unfold plus_subset_non_zero_el_prod_compose_maps. unfold non_zero_el_prod_compose_maps. pose proof (excl_middle_sat (fun a:(Fin_map E signe mns) => el_prod_compose g a = 0)) as h2. pose proof (feq_im _ _ (el_prod_compose g) h2) as h3. simpl in h3. rewrite im_union in h3. pose proof (finite_fin_maps _ _ mns h1 signe_finite) as h4. pose proof (finite_image _ _ _ (el_prod_compose g) h4) as h5. pose proof h5 as h6. rewrite h3 in h6. pose proof f_equal. pose proof (plus_set_functional _ _ h3 h5 h6) as h7. pose proof (el_prod_compose_covers _ h1 g) as h8. terml h8. redterm0 x. pose proof (proof_irrelevance _ c h5) as h9. unfold c in h9; subst. rewrite h8 in h7 at 1. symmetry in h7. erewrite plus_set_union' in h7. rewrite plus_set_el_prod_compose_zero, zero_sum_l in h7. apply h7. Grab Existential Variables. apply (Finite_downward_closed _ _ c). red; intros. destruct H0 as [F]; subst. apply Im_intro with F; auto. constructor. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_int : forall {B:Bool_Alg} (E:Ensemble (bt B)) (pfe:Finite E) (g:(bt B)->At) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_compose_maps g E)) (pfy:Included Y (non_zero_el_prod_compose_maps g E)), plus_subset_non_zero_el_prod_compose_maps E pfe g (Intersection X Y) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_int pfe g X Y pfx pfy) = plus_subset_non_zero_el_prod_compose_maps E pfe g X pfx * plus_subset_non_zero_el_prod_compose_maps E pfe g Y pfy. intros B E h1 g X Y h2 h3. unfold plus_subset_non_zero_el_prod_compose_maps. rewrite dist_set_plus2. pose proof (cart_prod_im X Y (el_prod_compose g) (el_prod_compose g)) as h4. pose proof (excl_middle_sat' (cart_prod X Y) (fun pr => fst pr = snd pr)) as h5. rewrite h5 in h4. clear h5. rewrite im_union in h4. let P := type of h4 in match P with _ = Union (Im ?C' _) _ => pose C' as C end. let P := type of h4 in match P with _ = Union _ (Im ?D' _) => pose D' as D end. let P := type of h4 in match P with _ = Union (Im _ ?f') _ => pose f' as f end. fold C D f in h4. match goal with |- _ = plus_set _ ?pf => pose pf as hyp1 end. fold hyp1. pose proof (feq_im _ _ (fun p : Btype (Bc A) * Btype (Bc A) => fst p * snd p) h4) as h5. rewrite im_union in h5. *do 2 rewrite im_im in h5. simpl in h5. unfold At in C, D. let P := type of h5 in match P with _ = Union (Im _ ?psi') _ => pose psi' as psi end. fold psi in h5. pose proof (finite_fin_map_ens X h1 signe_finite) as h6. pose proof (finite_fin_map_ens Y h1 signe_finite) as h7. pose proof (finite_image _ _ _ (el_prod_compose g) h6) as h8. pose proof (finite_image _ _ _ (el_prod_compose g) h7) as h9. pose proof (cart_prod_fin _ _ h8 h9) as h10. pose proof (finite_image _ _ _ (fun p : Btype (Bc A) * Btype (Bc A) => fst p * snd p) h10) as h11. rewrite h5 in h11. pose proof (plus_set_functional _ _ h5 hyp1 h11) as h12. assert (h13:Included (Im C psi) (Union (Im C psi) (Im D psi))). auto with sets. assert (h14:Included (Im D psi) (Union (Im C psi) (Im D psi))). auto with sets. pose proof (Finite_downward_closed _ _ h11 _ h13) as h15. clear h13. pose proof (Finite_downward_closed _ _ h11 _ h14) as h16. clear h14. rewrite (plus_set_union' _ _ h15 h16 h11) in h12 at 1. pose proof (finite_fin_map_squared_ens D h1 signe_finite) as h13. pose proof (plus_set_el_prod_compose_disjoint' _ g X Y h13 h16) as hp. fold D psi in hp. rewrite hp in h12. rewrite zero_sum in h12. unfold At. unfold At in h12. rewrite h12 at 1. apply plus_set_functional. symmetry. apply Extensionality_Ensembles. red. split. red. intros x h17. inversion h17 as [a h18 b h19 h20]. subst. destruct h18 as [h18]. destruct h18 as [h18l h18r]. destruct h18l as [h18l]. destruct h18l as [h18a h18b]. unfold psi. rewrite h18r. rewrite idem_prod. rewrite h18r in h18a. apply Im_intro with (snd a). split; auto. reflexivity. red. intros x h17. inversion h17 as [a h18 b h19 h20]. subst. destruct h18 as [a h18l h18r]. apply Im_intro with (a, a). constructor. simpl. split; auto. split; auto. simpl. unfold psi. simpl. rewrite idem_prod. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_union : forall {B:Bool_Alg} (E:Ensemble (bt B)) (pfe:Finite E) (g:(bt B)->At) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_compose_maps g E)) (pfy:Included Y (non_zero_el_prod_compose_maps g E)), plus_subset_non_zero_el_prod_compose_maps E pfe g (Union X Y) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_union pfe g X Y pfx pfy) = plus_subset_non_zero_el_prod_compose_maps E pfe g X pfx + plus_subset_non_zero_el_prod_compose_maps E pfe g Y pfy. intros B E h1 g X Y hx hy. unfold plus_subset_non_zero_el_prod_compose_maps. pose proof (im_union X Y (el_prod_compose g)) as h2. assert (h3:Finite (Im X (el_prod_compose g))). apply finite_image. apply finite_fin_map_ens; auto. apply signe_finite. assert (h4:Finite (Im Y (el_prod_compose g))). apply finite_image. apply finite_fin_map_ens; auto. apply signe_finite. match goal with |- plus_set _ ?pf = _ => pose pf as hyp end. fold hyp. pose proof (subsetT_eq_compat _ _ _ _ hyp (Union_preserves_Finite _ _ _ h3 h4) h2) as h5. dependent rewrite -> h5. rewrite plus_set_union. f_equal. apply plus_set_functional; auto. apply plus_set_functional; auto. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_comp : forall {B:Bool_Alg} (E:Ensemble (bt B)) (pfe:Finite E) (g:(bt B)->At) (X:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_compose_maps g E)), plus_subset_non_zero_el_prod_compose_maps E pfe g (Setminus (non_zero_el_prod_compose_maps g E) X) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_comp pfe g X pfx) = - plus_subset_non_zero_el_prod_compose_maps E pfe g X pfx. intros B E h1 g X h2. assert (h3:(plus_subset_non_zero_el_prod_compose_maps E h1 g X h2) * plus_subset_non_zero_el_prod_compose_maps E h1 g (Setminus (non_zero_el_prod_compose_maps g E) X) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_comp h1 g X h2) = 0). rewrite <- plus_subset_non_zero_el_prod_compose_maps_int. rewrite <- (plus_subset_non_zero_el_prod_compose_maps_empty E h1 g). unfold plus_subset_non_zero_el_prod_compose_maps. apply plus_set_functional. apply feq_im. apply int_setminus. assert (h4:(plus_subset_non_zero_el_prod_compose_maps E h1 g X h2) + plus_subset_non_zero_el_prod_compose_maps E h1 g (Setminus (non_zero_el_prod_compose_maps g E) X) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_comp h1 g X h2) = 1). rewrite <- plus_subset_non_zero_el_prod_compose_maps_union. rewrite <- (plus_subset_non_zero_el_prod_compose_maps_full E h1 g). unfold plus_subset_non_zero_el_prod_maps. apply plus_set_functional. apply feq_im. apply union_setminus_incl. assumption. pose proof (comp_char _ _ h3 h4) as h5. assumption. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_symdiff_full : forall {B:Bool_Alg} (E:Ensemble (bt B)) (pfe:Finite E) (g:bt B->At) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_compose_maps g E)) (pfy:Included Y (non_zero_el_prod_compose_maps g E)), (plus_subset_non_zero_el_prod_compose_maps E pfe g _ pfx) /_\ (plus_subset_non_zero_el_prod_compose_maps E pfe g _ pfy) = plus_subset_non_zero_el_prod_compose_maps E pfe g (Symdiff_full (non_zero_el_prod_compose_maps g E) X Y) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_symdiff_full pfe g X Y pfx pfy). intros B E h1 g X Y h2 h3. unfold sym_diff. do 2 rewrite <- plus_subset_non_zero_el_prod_compose_maps_comp. do 2 rewrite <- plus_subset_non_zero_el_prod_compose_maps_int. rewrite <- plus_subset_non_zero_el_prod_compose_maps_union. unfold Symdiff_full. unfold Setminus_full. unfold plus_subset_non_zero_el_prod_compose_maps. apply plus_set_functional. reflexivity. Qed. Lemma closed_all_plus_subsets_non_zero_el_prod_maps : forall (E:Ensemble At) (pfe:Finite E), alg_closed [a:At | exists (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_maps E)), a = plus_subset_non_zero_el_prod_maps E pfe S pfi]. intros E h1. constructor. (* + *) red. intros x y. constructor. destruct x as [x h2]. destruct y as [y h3]. destruct h2 as [h2]. destruct h3 as [h3]. simpl. destruct h2 as [X h2]. destruct h3 as [Y h3]. destruct h2 as [h2a h2b]. destruct h3 as [h3a h3b]. exists (Union X Y). exists (inclusion_plus_subset_non_zero_el_prod_maps_preserves_union _ h1 _ _ h2a h3a). unfold Bplus_sub. simpl. rewrite h2b. rewrite h3b. apply plus_subset_non_zero_el_prod_maps_union. (* * *) red. intros x y. constructor. destruct x as [x h2]. destruct y as [y h3]. destruct h2 as [h2]. destruct h3 as [h3]. simpl. destruct h2 as [X h2]. destruct h3 as [Y h3]. destruct h2 as [h2a h2b]. destruct h3 as [h3a h3b]. exists (Intersection X Y). exists (inclusion_plus_subset_non_zero_el_prod_maps_preserves_int _ h1 _ _ h2a h3a). unfold Btimes_sub. simpl. rewrite h2b. rewrite h3b. apply plus_subset_non_zero_el_prod_maps_int. (* 1 *) red. constructor. exists (non_zero_el_prod_maps E). exists (inclusion_reflexive _). symmetry. apply plus_subset_non_zero_el_prod_maps_full. (* 0 *) red. constructor. exists (Empty_set _). exists (incl_empty _). symmetry. apply plus_subset_non_zero_el_prod_maps_empty. (* - *) red. intro x. constructor. destruct x as [x h2]. destruct h2 as [h2]. simpl. destruct h2 as [X h2]. destruct h2 as [h2a h2b]. exists (Setminus (non_zero_el_prod_maps E) X). exists (inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp _ h1 X h2a). unfold Bcomp_sub. simpl. rewrite h2b. apply plus_subset_non_zero_el_prod_maps_comp. Qed. Lemma closed_all_plus_subsets_non_zero_el_prod_compose_maps : forall {B:Bool_Alg} (E:Ensemble (bt B)) (pfe:Finite E) (g:bt B->At), alg_closed [a:At | exists (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_compose_maps g E)), a = plus_subset_non_zero_el_prod_compose_maps E pfe g S pfi]. intros B E h1 g. constructor. (* + *) red. intros x y. constructor. destruct x as [x h2]. destruct y as [y h3]. destruct h2 as [h2]. destruct h3 as [h3]. simpl. destruct h2 as [X h2]. destruct h3 as [Y h3]. destruct h2 as [h2a h2b]. destruct h3 as [h3a h3b]. exists (Union X Y). exists (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_union h1 g _ _ h2a h3a). unfold Bplus_sub. simpl. rewrite h2b. rewrite h3b. rewrite <- (plus_subset_non_zero_el_prod_compose_maps_union E h1 g _ _ h2a h3a). reflexivity. (* * *) red. intros x y. constructor. destruct x as [x h2]. destruct y as [y h3]. destruct h2 as [h2]. destruct h3 as [h3]. simpl. destruct h2 as [X h2]. destruct h3 as [Y h3]. destruct h2 as [h2a h2b]. destruct h3 as [h3a h3b]. exists (Intersection X Y). exists (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_int h1 g _ _ h2a h3a). unfold Btimes_sub. simpl. rewrite h2b. rewrite h3b. rewrite <- (plus_subset_non_zero_el_prod_compose_maps_int E h1 g _ _ h2a h3a). reflexivity. (* 1 *) red. constructor. exists (non_zero_el_prod_compose_maps g E). exists (inclusion_reflexive _). symmetry. rewrite <- (plus_subset_non_zero_el_prod_compose_maps_full E h1 g). reflexivity. (* 0 *) red. constructor. exists (Empty_set _). exists (incl_empty _). symmetry. rewrite <- (plus_subset_non_zero_el_prod_compose_maps_empty E h1 g). reflexivity. (* - *) red. intro x. constructor. destruct x as [x h2]. destruct h2 as [h2]. destruct h2 as [X h2]. destruct h2 as [h2a h2b]. exists (Setminus (non_zero_el_prod_compose_maps g E) X). exists (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_comp h1 g _ h2a). unfold Bcomp_sub. simpl. rewrite h2b. rewrite <- (plus_subset_non_zero_el_prod_compose_maps_comp E h1 g _ h2a). reflexivity. Qed. Lemma el_prod_sing_inc : forall (E:Ensemble At) (pf:Finite E) (a:Fin_map E signe mns), el_prod a <> 0 -> Included (Singleton a) (non_zero_el_prod_maps E). intros E h1 a h2. red. intros a' h3. destruct h3; subst. unfold non_zero_el_prod_maps. constructor. assumption. Qed. Lemma el_prod_sing : forall (E:Ensemble At) (pf:Finite E) (a:Fin_map E signe mns) (pfel:el_prod a <> 0), el_prod a = plus_subset_non_zero_el_prod_maps _ pf (Singleton a) (el_prod_sing_inc _ pf _ pfel). intros E h1 a h2. unfold plus_subset_non_zero_el_prod_maps. assert (h3:Im (Singleton a) el_prod = (Singleton (el_prod a))). apply Extensionality_Ensembles. red. split. red. intros x' h13. inversion h13. subst. destruct H. subst. constructor. red. intros x' h13. destruct h13. apply Im_intro with a. constructor. reflexivity. generalize dependent ( (finite_image (Fin_map E signe mns) At (Singleton a) el_prod (Finite_downward_closed (Fin_map E signe mns) (non_zero_el_prod_maps E) (non_zero_el_prod_maps_fin E h1) (Singleton a) (el_prod_sing_inc E h1 a h2)))). rewrite h3. intro h4. rewrite plus_set_sing'. reflexivity. Qed. Lemma atom_non_zero_el_prod : forall (E:Ensemble At) (pf:Finite E), let C := (Subalg _ (closed_all_plus_subsets_non_zero_el_prod_maps E pf)) in forall (atm:Btype (Bc C)), atom atm <-> exists a:Fin_map E signe mns, Inn (non_zero_el_prod_maps E) a /\ proj1_sig atm = el_prod a. intros E h1 C atm. split. intros h2. pose proof h2 as h3. rewrite atom_iff in h3. pose proof (proj2_sig atm) as h4. destruct h4 as [h4]. destruct h4 as [X h4]. destruct h4 as [h4 h5]. assert (h6: forall Y:Ensemble (Fin_map E signe mns), Included Y (non_zero_el_prod_maps E) -> Intersection Y X = X \/ Intersection Y X = Empty_set _). intros Y h7. assert (h8:Inn [a : At | exists (S : Ensemble (Fin_map E signe mns)) (pfi : Included S (non_zero_el_prod_maps E)), a = plus_subset_non_zero_el_prod_maps E h1 S pfi] (plus_subset_non_zero_el_prod_maps E h1 _ h7)). constructor. exists Y, h7. reflexivity. pose (exist _ _ h8) as pY. red in h2. destruct h2 as [h2l h2r]. specialize (h3 pY). destruct h3 as [h3l h3r]. destruct h3l as [h3a | h3b]. (* h3a *) red in h3a. rewrite eq_ord in h3a. simpl in h3a. unfold Btimes_sub in h3a. rewrite unfold_sig in h3a. pose proof (exist_injective _ _ _ _ _ h3a) as h6. clear h3a. unfold pY in h6. rewrite h5 in h6. simpl in h6. rewrite plus_subset_non_zero_el_prod_maps_int in h6. left. assert (h9:Included (Intersection X Y) X). auto with sets. rewrite Intersection_commutative. apply NNPP. intro h10. assert (h11:Strict_Included (Intersection X Y) X). auto with sets. pose proof (Strict_Included_inv _ _ _ h11) as h12. destruct h12 as [h12l h12r]. clear h12l. rewrite setminus_same_int in h12r. pose proof (Inhabited_not_empty _ _ h12r) as h13. assert (h14:Included (Setminus X Y) (non_zero_el_prod_maps E)). red. intros x h13'. destruct h13'. auto with sets. pose proof (plus_subset_non_zero_el_prod_maps_empty_rev _ h1 _ h14) as h15. assert (h16:plus_subset_non_zero_el_prod_maps E h1 (Setminus X Y) h14 <> 0). tauto. assert (h17:Setminus X Y = Intersection X (Setminus (non_zero_el_prod_maps E) Y)). apply Extensionality_Ensembles. red. split. red. intros b h17. destruct h17 as [h17 h18]. constructor; auto with sets. red. intros b h17. destruct h17 as [b h17 h18]. constructor; auto with sets. destruct h18. assumption. pose proof (inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp _ h1 _ h7) as h18. assert (h19:plus_subset_non_zero_el_prod_maps E h1 (Setminus X Y) h14 = plus_subset_non_zero_el_prod_maps E h1 _ (inclusion_plus_subset_non_zero_el_prod_maps_preserves_int _ h1 _ _ h4 (inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp _ h1 _ h7))). unfold plus_subset_non_zero_el_prod_maps. apply plus_set_functional. apply feq_im. assumption. rewrite h19 in h16. rewrite <- plus_subset_non_zero_el_prod_maps_int in h16. rewrite <- plus_subset_non_zero_el_prod_maps_comp in h16. unfold At in h16. rewrite <- le_iff in h16. unfold le in h16. rewrite eq_ord in h16. rewrite <- plus_subset_non_zero_el_prod_maps_int in h6. contradiction. (* h3b *) right. do 2 rewrite unfold_sig in h3b. unfold pY in h3b. simpl in h3b. pose proof (exist_injective _ _ _ _ _ h3b) as h6. unfold Btimes_sub in h6. simpl in h6. rewrite h5 in h6. rewrite plus_subset_non_zero_el_prod_maps_int in h6. rewrite Intersection_commutative. apply (plus_subset_non_zero_el_prod_maps_empty_rev _ h1 _ (inclusion_plus_subset_non_zero_el_prod_maps_preserves_int _ h1 _ _ h4 h7)). assumption. assert (h9:X <> Empty_set _). intro h10. generalize dependent h4. rewrite h10. intro h4. pose proof (proof_irrelevance _ h4 (incl_empty (non_zero_el_prod_maps E))); subst. rewrite plus_subset_non_zero_el_prod_maps_empty. intro h12. red in h2. destruct h2 as [h2l h2r]. contradict h2l. rewrite (unfold_sig _ atm). rewrite (unfold_sig _ (Bzero (Bc C))). apply proj1_sig_injective. simpl. assumption. pose proof (non_zero_el_prod_maps_fin _ h1) as h10. pose proof (atom_sing _ _ h10 h4 h9) as h11. rewrite h11 in h6. destruct h6 as [a h6]. exists a. destruct h6 as [h6l h6r]. generalize dependent h4. rewrite h6r. intros h4 h5. unfold plus_subset_non_zero_el_prod_maps in h5. assert (h12:Im (Singleton a) (el_prod) = Singleton (el_prod a)). apply Extensionality_Ensembles. red. split. red. intros x' h13. inversion h13. subst. destruct H. subst. constructor. red. intros x' h13. destruct h13. apply Im_intro with a. constructor. reflexivity. generalize dependent (finite_image (Fin_map E signe mns) At (Singleton a) el_prod (Finite_downward_closed (Fin_map E signe mns) (non_zero_el_prod_maps E) (non_zero_el_prod_maps_fin E h1) (Singleton a) h4)). rewrite h12. intros h13 h14. rewrite plus_set_sing' in h14. split; auto. intro h2. destruct h2 as [a h2]. pose proof h2 as h2'. rewrite atom_iff. intro b. split. pose proof (proj2_sig b) as h3. simpl in h3. destruct h3 as [h3]. destruct h3 as [S h3]. destruct h3 as [h3 h4]. pose proof (proj2_sig atm) as h5. simpl in h5. destruct h5 as [h5]. destruct h5 as [A' h5]. destruct h5 as [h5 h6]. destruct (classic (A' = Empty_set _)) as [h7 | h8]. generalize dependent h5. rewrite h7. intros h5 h6. pose proof (proof_irrelevance _ h5 (incl_empty (non_zero_el_prod_maps E))); subst. rewrite plus_subset_non_zero_el_prod_maps_empty in h6. right. rewrite (unfold_sig _ atm). rewrite (unfold_sig _ b). rewrite (unfold_sig _ (Bzero (Bc C))). simpl. apply proj1_sig_injective. simpl. unfold Btimes_sub. simpl. rewrite h6. rewrite comm_prod. apply zero_prod. assert (h9:plus_subset_non_zero_el_prod_maps E h1 A' h5 <> 0). intro h10. pose proof (plus_subset_non_zero_el_prod_maps_empty_rev _ h1 _ h5 h10) as h11. contradiction. rewrite <- h6 in h9. destruct h2 as [h2'' h2]. rewrite h2 in h9. pose proof el_prod_sing. pose proof (el_prod_sing _ h1 a h9) as h10. rewrite h10 in h2. pose proof (@plus_subset_non_zero_el_prod_maps_int). pose proof (plus_subset_non_zero_el_prod_maps_int _ h1 _ _ (el_prod_sing_inc E h1 a h9) h3) as h11. rewrite <- h2 in h11. rewrite <- h4 in h11. pose proof (singleton_inc_int a S) as h12. destruct h12 as [h12l | h12r]. generalize dependent (inclusion_plus_subset_non_zero_el_prod_maps_preserves_int E h1 (Singleton a) S (el_prod_sing_inc E h1 a h9) h3). rewrite Intersection_commutative in h12l. rewrite h12l. intros h13 h14. pose proof (proof_irrelevance _ h13 (incl_empty (non_zero_el_prod_maps E))); subst. rewrite plus_subset_non_zero_el_prod_maps_empty in h14. right. rewrite (unfold_sig _ atm). rewrite (unfold_sig _ b). rewrite (unfold_sig _ (Bzero (Bc C))). simpl. apply proj1_sig_injective. simpl. unfold Btimes_sub. simpl. assumption. left. red. rewrite eq_ord. generalize dependent (inclusion_plus_subset_non_zero_el_prod_maps_preserves_int E h1 (Singleton a) S (el_prod_sing_inc E h1 a h9) h3). rewrite Intersection_commutative in h12r. rewrite h12r. intros h13 h14. pose proof (proof_irrelevance _ h13 (el_prod_sing_inc E h1 a h9)); subst. rewrite <- h2 in h14. rewrite (unfold_sig _ atm). rewrite (unfold_sig _ b). simpl. unfold Btimes_sub. simpl. apply proj1_sig_injective. simpl; auto. unfold le. rewrite eq_ord. intro h3. pose proof (proj2_sig b) as h4. destruct h3 as [h3l h3r]. pose proof h3l as h3l'. rewrite h3r in h3l. destruct h2 as [h2l h2r]. rewrite (unfold_sig _ atm) in h3l. rewrite (unfold_sig _ (Bzero (Bc C))) in h3l. pose proof (exist_injective _ _ _ _ _ h3l) as h8. simpl in h8. rewrite <- h8 in h2r. unfold non_zero_el_prod_maps in h2l. destruct h2l as [h2l]. symmetry in h2r. contradiction. Qed. Lemma plus_subset_non_zero_el_prod_maps_inj : forall (E:Ensemble At) (pfe:Finite E) (S T:Ensemble (Fin_map E signe mns)) (pfs:Included S (non_zero_el_prod_maps E)) (pft:Included T (non_zero_el_prod_maps E)), plus_subset_non_zero_el_prod_maps E pfe S pfs = plus_subset_non_zero_el_prod_maps E pfe T pft -> S = T. intros E h1 S T h2 h3 h4. pose proof (plus_subset_non_zero_el_prod_maps_empty _ h1) as h5. pose proof (symdiff_full_ref (non_zero_el_prod_maps E) S) as h6. generalize dependent (incl_empty (non_zero_el_prod_maps E)). rewrite <- h6. pose proof (plus_subset_non_zero_el_prod_maps_symdiff_full _ h1 _ _ h2 h2) as h7. rewrite h4 in h7 at 2. generalize dependent (inclusion_plus_subset_non_zero_el_prod_maps_preserves_symdiff_full E h1 S S h2 h2). intros h7 h8 h9 h10. pose proof (proof_irrelevance _ h7 h9); subst. rewrite h10 in h8. rewrite plus_subset_non_zero_el_prod_maps_symdiff_full in h8. pose proof (inclusion_plus_subset_non_zero_el_prod_maps_preserves_symdiff_full E h1 S T h2 h3) as h11. pose proof (plus_subset_non_zero_el_prod_maps_empty_rev _ h1 (Symdiff_full (non_zero_el_prod_maps E) S T) h11) as h12. generalize dependent (inclusion_plus_subset_non_zero_el_prod_maps_preserves_symdiff_full E h1 S T h2 h3). intro h13. pose proof (proof_irrelevance _ h11 h13); subst. intro h14. specialize (h12 h14). rewrite symdiff_full_empty_iff_eq in h12; auto. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_inj : forall {B:Bool_Alg} (g:bt B->At) (E:Ensemble (bt B)) (pfe:Finite E) (S T:Ensemble (Fin_map E signe mns)) (pfs:Included S (non_zero_el_prod_compose_maps g E)) (pft:Included T (non_zero_el_prod_compose_maps g E)), plus_subset_non_zero_el_prod_compose_maps E pfe g S pfs = plus_subset_non_zero_el_prod_compose_maps E pfe g T pft -> S = T. intros B g E h1 S T h2 h3 h4. pose proof (plus_subset_non_zero_el_prod_compose_maps_empty _ h1 g) as h5. pose proof (symdiff_full_ref (non_zero_el_prod_compose_maps g E) S) as h6. generalize dependent (incl_empty (non_zero_el_prod_compose_maps g E)). rewrite <- h6. pose proof (plus_subset_non_zero_el_prod_compose_maps_symdiff_full _ h1 g _ _ h2 h2) as h7. rewrite h4 in h7 at 2. generalize dependent (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_symdiff_full h1 g S S h2 h2). intros h7 h8 h9 h10. pose proof (proof_irrelevance _ h7 h9); subst. rewrite h10 in h8. rewrite plus_subset_non_zero_el_prod_compose_maps_symdiff_full in h8. pose proof (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_symdiff_full h1 g S T h2 h3) as h11. pose proof (plus_subset_non_zero_el_prod_compose_maps_empty_rev g _ h1 (Symdiff_full (non_zero_el_prod_compose_maps g E) S T) h11) as h12. generalize dependent (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_symdiff_full h1 g S T h2 h3). intro h13. pose proof (proof_irrelevance _ h11 h13); subst. intro h14. specialize (h12 h14). rewrite symdiff_full_empty_iff_eq in h12; auto. Qed. Lemma el_prod_eq_plus_subset_non_zero_el_prod_maps_extends : forall (E F:Ensemble At) (pfe:Finite E), Included F E -> forall (b:Fin_map F signe mns), let L := [a : Fin_map E signe mns | Inn (non_zero_el_prod_maps E) a /\ extends a b] in exists pf:Included L (non_zero_el_prod_maps E), el_prod b = plus_subset_non_zero_el_prod_maps E pfe L pf. intros E F h1 h2 b L. ex_goal. red. intros x h3. destruct h3 as [h3]. destruct h3; auto. exists hex. assert (h4:forall a:Fin_map E signe mns, Inn L a -> Included (Im F (fun i : At => eps i (b |-> i))) (Im E (fun i : At => eps i (a |-> i)))). intros a h5. red. intros p h6. destruct h5 as [h5]. destruct h5 as [h5l h5r]. red in h5r. destruct h5r as [h5a h5b]. destruct h6 as [f h6 p]. subst. rewrite h5b; auto. apply Im_intro with f. auto with sets. reflexivity. assert (h5:forall a:Fin_map E signe mns, Inn L a -> le (el_prod a) (el_prod b)). intros a h6. unfold el_prod. apply times_set_inc_le. apply h4; auto. assert (h6:forall a:Fin_map E signe mns, Inn (Setminus (non_zero_el_prod_maps E) L) a -> (el_prod a) * (el_prod b) = 0). intros a h7. destruct h7 as [h7 h8]. unfold L in h8. assert (h9:~extends a b). intro h10. assert (h11:Inn [a0 : Fin_map E signe mns | Inn (non_zero_el_prod_maps E) a0 /\ extends a0 b] a). constructor. split; assumption. contradiction. pose proof (not_extends_differs_at_point sign_eq_dec _ _ h2 h9) as h10. destruct h10 as [i h10]. destruct h10 as [h10l h10r]. pose proof (el_prod_le_ai b _ h10l) as h11. assert (h12:Inn E i). auto with sets. pose proof (el_prod_le_ai a _ h12) as h13. pose proof (mono_prod _ _ _ _ h11 h13) as h14. destruct (b |-> i); destruct (a |-> i). contradict h10r. reflexivity. simpl in h14. rewrite comp_prod in h14. apply le_x_0. rewrite comm_prod. assumption. simpl in h14. rewrite (comm_prod _ (-i) i) in h14. rewrite comp_prod in h14. apply le_x_0. rewrite comm_prod. assumption. contradict h10r. reflexivity. pose proof (one_prod (el_prod b)) as h7. rewrite <- (plus_subset_non_zero_el_prod_maps_full _ h1) in h7. pose proof (decompose_setminus_inc _ _ hex) as h8. assert (h9: plus_subset_non_zero_el_prod_maps E h1 (non_zero_el_prod_maps E) (inclusion_reflexive (non_zero_el_prod_maps E)) = plus_subset_non_zero_el_prod_maps E h1 (Union L (Setminus (non_zero_el_prod_maps E) L)) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_union _ h1 _ _ hex (inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp _ h1 _ hex))). unfold plus_subset_non_zero_el_prod_maps. apply plus_set_functional. apply feq_im. assumption. rewrite h9 in h7. unfold plus_subset_non_zero_el_prod_maps in h7. rewrite dist_set_plus1 in h7. revert h7. Gen0. rewrite im_im, im_union. intros h10 h7. assert (h11:Included (Im L (fun x:Fin_map E signe mns => el_prod b * el_prod x)) (Union (Im L (fun x:Fin_map E signe mns => el_prod b * el_prod x)) (Im (Setminus (non_zero_el_prod_maps E) L) (fun x:Fin_map E signe mns => el_prod b * el_prod x)))). auto with sets. assert (h12:Included (Im (Setminus (non_zero_el_prod_maps E) L) (fun x:Fin_map E signe mns => el_prod b * el_prod x)) (Union (Im L (fun x:Fin_map E signe mns => el_prod b * el_prod x)) (Im (Setminus (non_zero_el_prod_maps E) L) (fun x:Fin_map E signe mns => el_prod b * el_prod x)))). auto with sets. assert (h13:Finite (Im L (fun x:Fin_map E signe mns => el_prod b * el_prod x))). eapply Finite_downward_closed. apply h10. apply h11. assert (h14:Finite (Im (Setminus (non_zero_el_prod_maps E) L) (fun x:Fin_map E signe mns => el_prod b * el_prod x))). eapply Finite_downward_closed. apply h10. apply h12. assert (h16:plus_set (Im (Setminus (non_zero_el_prod_maps E) L) (fun x:Fin_map E signe mns => el_prod b * el_prod x)) h14 = 0). destruct (classic (Inhabited (Setminus (non_zero_el_prod_maps E) L))) as [h17 | h18]. rewrite <- (plus_set_sing 0). apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros x h18. destruct h18 as [x h18]. subst. rewrite comm_prod. rewrite h6; auto. constructor. red. intros z h18. destruct h18; subst. destruct h17 as [a h17]. apply Im_intro with a; auto. rewrite comm_prod. rewrite h6; auto. pose proof (not_inhabited_empty _ h18) as h19. pose proof (image_empty _ _ (fun x : Fin_map E signe mns => el_prod b * el_prod x)) as h20. rewrite <- h19 in h20. rewrite <- plus_set_empty. apply plus_set_functional. assumption. unfold plus_subset_non_zero_el_prod_maps. rewrite <- h7. pose proof (im_im (Union L (Setminus (non_zero_el_prod_maps E) L)) el_prod (fun y:bt A => el_prod b * y)) as h17. simpl in h17. pose proof (im_union L (Setminus (non_zero_el_prod_maps E) L) (fun x : Fin_map E signe mns => el_prod b * el_prod x)) as hiu. rewrite hiu in h17. rewrite <- decompose_setminus_inc in h17. let P := type of h10 in match P with Finite ?S => pose S as C end. fold C in h10. match goal with |- ?c = _ => assert (heq:c = plus_set C h10) end. apply plus_set_functional. unfold C. rewrite <- h17; auto. unfold At, bt in heq. unfold At, bt. simpl in heq. simpl. rewrite heq at 1. pose proof (plus_set_union (Im L (fun x : Fin_map E signe mns => el_prod b * el_prod x)) (Im (Setminus (non_zero_el_prod_maps E) L) (fun x : Fin_map E signe mns => el_prod b * el_prod x)) h13 h14) as hpsu. terml hpsu. redterm0 x. pose proof (proof_irrelevance _ c h10) as hpi. unfold c in hpi; subst. unfold bt, At, C in hpsu. unfold bt, At, C. simpl. simpl in hpsu. unfold At. simpl. rewrite hpsu. unfold bt, At in h16. rewrite h16 at 1. rewrite zero_sum. apply plus_set_functional. apply im_ext_in. clear x. intros x h17'. specialize (h5 x h17'). unfold le in h5. rewrite eq_ord in h5. rewrite <- h5 at 2. rewrite comm_prod. reflexivity. assumption. Qed. Lemma el_prod_compose_eq_plus_subset_non_zero_el_prod_compose_maps_extends : forall {B:Bool_Alg} (E F:Ensemble (bt B)) (pfe:Finite E), Included F E -> forall (g:bt B->At) (b:Fin_map F signe mns), let L := [a : Fin_map E signe mns | Inn (non_zero_el_prod_compose_maps g E) a /\ extends a b] in exists pf:Included L (non_zero_el_prod_compose_maps g E), el_prod_compose g b = plus_subset_non_zero_el_prod_compose_maps E pfe g L pf. intros B E F h1 h2 g b L. ex_goal. red. intros x h3. destruct h3 as [h3]. destruct h3; auto. exists hex. unfold plus_subset_non_zero_el_prod_compose_maps. pose proof (one_prod (el_prod_compose g b)) as h7. rewrite <- (plus_subset_non_zero_el_prod_compose_maps_full _ h1 g) in h7. pose proof (decompose_setminus_inc _ _ hex) as h8. assert (h9: plus_subset_non_zero_el_prod_compose_maps E h1 g (non_zero_el_prod_compose_maps g E) (inclusion_reflexive (non_zero_el_prod_compose_maps g E)) = plus_subset_non_zero_el_prod_compose_maps E h1 g (Union L (Setminus (non_zero_el_prod_compose_maps g E) L)) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_union h1 g _ _ hex (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_comp h1 g _ hex))). unfold plus_subset_non_zero_el_prod_compose_maps. apply plus_set_functional. apply feq_im. assumption. rewrite h9 in h7. unfold plus_subset_non_zero_el_prod_compose_maps in h7. rewrite dist_set_plus1 in h7. revert h7. Gen0. rewrite im_im, im_union. intros h10 h7. rewrite <- h7 at 1. assert (h14:Finite (Im (Setminus (non_zero_el_prod_compose_maps g E) L) (fun x:Fin_map E signe mns => el_prod_compose g b * el_prod_compose g x))). eapply Finite_downward_closed. apply h10. assert (h11:Included (Im L (fun x:Fin_map E signe mns => el_prod_compose g b * el_prod_compose g x)) (Union (Im L (fun x:Fin_map E signe mns => el_prod_compose g b * el_prod_compose g x)) (Im (Setminus (non_zero_el_prod_compose_maps g E) L) (fun x:Fin_map E signe mns => el_prod_compose g b * el_prod_compose g x)))). auto with sets. auto with sets. assert (h6:forall a:Fin_map E signe mns, Inn (Setminus (non_zero_el_prod_compose_maps g E) L) a -> (el_prod_compose g a) * (el_prod_compose g b) = 0). intros a h12. destruct h12 as [h12 h13]. unfold L in h13. assert (h19:~extends a b). intro h15. assert (h11:Inn [a0 : Fin_map E signe mns | Inn (non_zero_el_prod_compose_maps g E) a0 /\ extends a0 b] a). constructor. split; assumption. contradiction. pose proof (not_extends_differs_at_point sign_eq_dec _ _ h2 h19) as h15. destruct h15 as [i h15]. destruct h15 as [h15l h15r]. pose proof (el_prod_compose_le_ai b g _ h15l) as h11. assert (h17:Inn E i). auto with sets. pose proof (el_prod_compose_le_ai a g _ h17) as h18. pose proof (mono_prod _ _ _ _ h11 h18) as h24. destruct (b |-> i); destruct (a |-> i). contradict h15r. reflexivity. simpl in h24. rewrite comp_prod in h24. apply le_x_0. rewrite comm_prod. assumption. simpl in h24. rewrite (comm_prod _ (-g i) (g i)) in h24. rewrite comp_prod in h24. apply le_x_0. rewrite comm_prod. assumption. contradict h15r. reflexivity. assert (h16:plus_set (Im (Setminus (non_zero_el_prod_compose_maps g E) L) (fun x:Fin_map E signe mns => el_prod_compose g b * el_prod_compose g x)) h14 = 0). destruct (classic (Inhabited (Setminus (non_zero_el_prod_compose_maps g E) L))) as [h17 | h18]. rewrite <- (plus_set_sing 0). apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros x h18. destruct h18 as [x h18]. subst. rewrite comm_prod. rewrite h6; auto. constructor. red. intros z h18. destruct h18; subst. destruct h17 as [a h17]. apply Im_intro with a; auto. rewrite comm_prod. rewrite h6; auto. pose proof (not_inhabited_empty _ h18) as h19. pose proof (image_empty _ _ (fun x : Fin_map E signe mns => el_prod_compose g b * el_prod_compose g x)) as h20. rewrite <- h19 in h20. rewrite <- plus_set_empty. apply plus_set_functional. assumption. erewrite plus_set_union'. rewrite comm_sum. erewrite plus_set_functional. rewrite h16 at 1. rewrite zero_sum_l. apply plus_set_functional. apply im_ext_in; intros x hx. unfold el_prod_compose. gtermr. gterml. redterm1 x0. fold c. assert (h4:forall a:Fin_map E signe mns, Inn L a -> Included (Im F (fun i : bt B => eps (g i) (b |-> i))) (Im E (fun i : bt B => eps (g i) (a |-> i)))). intros a h5. red. intros p h11. destruct h5 as [h5]. destruct h5 as [h5l h5r]. red in h5r. destruct h5r as [h5a h5b]. destruct h11 as [f h11 p]. subst. rewrite h5b; auto. apply Im_intro with f. auto with sets. reflexivity. specialize (h4 _ hx). redterm0 c. redterm0 y. pose proof (times_set_inc_le _ _ c0 c1 h4) as h17. red in h17. unfold c. rewrite eq_ord in h17. unfold c0, c1 in h17. rewrite comm_prod, h17; auto. reflexivity. Grab Existential Variables. assumption. eapply Finite_downward_closed. apply h10. auto with sets. Qed. Lemma el_prod_inc_c : forall (E F:Ensemble At) (pfe:Finite E) (b:Fin_map F signe mns), Included F E -> let C := [a:At | exists (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_maps E)), a = plus_subset_non_zero_el_prod_maps E pfe S pfi] in Inn C (el_prod b). intros E F h1 b h2 C. pose proof (el_prod_eq_plus_subset_non_zero_el_prod_maps_extends E F h1 h2 b) as h3. simpl in h3. destruct h3 as [h3 h4]. rewrite h4. constructor. exists [a : Fin_map E signe mns | Inn (non_zero_el_prod_maps E) a /\ extends a b]. exists h3. reflexivity. Qed. Lemma el_prod_compose_inc_c : forall {B:Bool_Alg} (g:bt B->At) (E F:Ensemble (bt B)) (pf:Finite E) (b:Fin_map F signe mns), Included F E -> let C := [a:At | exists (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_compose_maps g E)), a = plus_subset_non_zero_el_prod_compose_maps E pf g S pfi] in Inn C (el_prod_compose g b). intros B g E F h1 b h2 C. pose proof (el_prod_compose_eq_plus_subset_non_zero_el_prod_compose_maps_extends E F h1 h2 g b) as h3. simpl in h3. destruct h3 as [h3 h4]. rewrite h4. constructor. exists [a : Fin_map E signe mns | Inn (non_zero_el_prod_compose_maps g E) a /\ extends a b]. exists h3. reflexivity. Qed. Lemma elt_eq_plus_subset_non_zero_el_prod_maps_at_one : forall (E:Ensemble At) (pfe:Finite E) (r:At), Inn E r -> let L := [a : Fin_map E signe mns | Inn (non_zero_el_prod_maps E) a /\ a |-> r = pls] in exists pf:Included L (non_zero_el_prod_maps E), r = plus_subset_non_zero_el_prod_maps E pfe L pf. intros E h1 x h2. pose proof (In_singleton _ x) as h3. assert (h4:Included (Singleton x) E). red. intros x' h4. destruct h4; subst. assumption. assert (h6:functionally_paired (Singleton x) signe (Singleton (x, pls))). constructor. intros x' h5. exists pls. red. split. split. unfold signe. left. destruct h5; subst. constructor. intros s h6. destruct h6 as [h6l h6r]. inversion h6r. reflexivity. intros pr h5. inversion h5 as [h6]. rewrite (surjective_pairing pr) in h6. inversion h6. subst. simpl. split. constructor. left. pose proof (Singleton_is_finite _ x) as h7. pose (fin_map_intro _ _ mns ba_eq_dec h7 signe_finite _ h6) as F. assert (h8: x = el_prod F). unfold F. unfold el_prod. symmetry. gen0. rewrite im_singleton. intro. rewrite times_set_sing'. simpl. pose proof (fin_fps_to_f_s_compat ba_eq_dec h7 h6 mns _ h3) as h9. inversion h9 as [he]. unfold At in he. unfold At. rewrite <- he at 1. reflexivity. simpl. ex_goal. red. intros f h9. destruct h9 as [h9]. destruct h9; auto. exists hex. pose proof (el_prod_eq_plus_subset_non_zero_el_prod_maps_extends _ _ h1 h4 F) as h10. simpl in h10. destruct h10 as [h10 h11]. rewrite <- h8 in h11 at 1. rewrite h11 at 1. assert (h12:[a : Fin_map E signe mns | Inn (non_zero_el_prod_maps E) a /\ a |-> x = pls] = [a : Fin_map E signe mns | Inn (non_zero_el_prod_maps E) a /\ extends a F]). apply Extensionality_Ensembles. red. split. red. intros f h12. destruct h12 as [h12]. destruct h12 as [h12 h13]. constructor. split; auto. red. split; auto. intros x' h14. destruct h14. simpl. rewrite h13. pose proof (fin_fps_to_f_s_compat ba_eq_dec h7 h6 mns x h3) as h14. inversion h14 as [he]. unfold At in he. unfold At. rewrite he; auto. red. intros f h12. destruct h12 as [h12]. destruct h12 as [h12 h13]. constructor. split; auto. red in h13. destruct h13 as [h13 h14]. specialize (h14 _ h3). rewrite <- h14. simpl. pose proof (fin_fps_to_f_s_compat ba_eq_dec h7 h6 mns x h3) as h15. inversion h15 as [he]. unfold At in he. unfold At. rewrite he; auto. pose proof (subsetT_eq_compat _ (fun S => Included S (non_zero_el_prod_maps E)) _ _ hex h10 h12) as h13. dependent rewrite -> h13. reflexivity. Qed. Lemma elt_eq_plus_subset_non_zero_el_prod_compose_maps_at_one : forall {B:Bool_Alg} (E:Ensemble (bt B)) (pfe:Finite E) (g:bt B->At) (r:bt B), Inn E r -> let L := [a : Fin_map E signe mns | Inn (non_zero_el_prod_compose_maps g E) a /\ a |-> r = pls] in exists pf:Included L (non_zero_el_prod_compose_maps g E), g r = plus_subset_non_zero_el_prod_compose_maps E pfe g L pf. intros B E h1 g x h2. pose proof (In_singleton _ x) as h3. assert (h4:Included (Singleton x) E). red. intros x' h4. destruct h4. assumption. assert (h6:functionally_paired (Singleton x) signe (Singleton (x, pls))). constructor. intros x' h5. exists pls. red. split. split. unfold signe. left. destruct h5; subst. constructor. intros s h6. destruct h6 as [h6l h6r]. inversion h6r. reflexivity. intros pr h5. inversion h5 as [h6]. rewrite (surjective_pairing pr) in h6. inversion h6. subst. simpl. split. constructor. left. pose proof (Singleton_is_finite _ x) as h7. pose (fin_map_intro _ _ mns ba_eq_dec h7 signe_finite _ h6) as F. assert (h8: g x = el_prod_compose g F). unfold F. unfold el_prod_compose. symmetry. gen0. rewrite im_singleton. intro h8. rewrite times_set_sing'. simpl. pose proof (fin_fps_to_f_s_compat ba_eq_dec h7 h6 mns _ h3) as h9. inversion h9. simpl. reflexivity. simpl. assert (h9: Included [a : Fin_map E signe mns | Inn (non_zero_el_prod_compose_maps g E) a /\ a |-> x = pls] (non_zero_el_prod_compose_maps g E)). red. intros f h9. destruct h9 as [h9]. destruct h9; auto. exists h9. pose proof (el_prod_compose_eq_plus_subset_non_zero_el_prod_compose_maps_extends _ _ h1 h4 g F) as h10. simpl in h10. destruct h10 as [h10 h11]. rewrite <- h8 in h11 at 1. rewrite h11 at 1. assert (h12:[a : Fin_map E signe mns | Inn (non_zero_el_prod_compose_maps g E) a /\ a |-> x = pls] = [a : Fin_map E signe mns | Inn (non_zero_el_prod_compose_maps g E) a /\ extends a F]). apply Extensionality_Ensembles. red. split. red. intros f h12. destruct h12 as [h12]. destruct h12 as [h12 h13]. constructor. split; auto. red. split; auto. intros x' h14. destruct h14. simpl. rewrite h13. pose proof (fin_fps_to_f_s_compat ba_eq_dec h7 h6 mns x h3) as h14. inversion h14. reflexivity. red. intros f h12. destruct h12 as [h12]. destruct h12 as [h12 h13]. constructor. split; auto. red in h13. destruct h13 as [h13 h14]. specialize (h14 _ h3). rewrite <- h14. simpl. pose proof (fin_fps_to_f_s_compat ba_eq_dec h7 h6 mns x h3) as h15. inversion h15. reflexivity. pose proof (subsetT_eq_compat _ (fun S => Included S (non_zero_el_prod_compose_maps g E)) _ _ h9 h10 h12) as h13. dependent rewrite -> h13. reflexivity. Qed. Lemma el_prod_empty : forall (F:Fin_map (Empty_set At) signe mns), el_prod F = 1. intros F. unfold el_prod. generalize dependent (finite_image At (bt A) (Empty_set At) (fun i : At => eps i (F |-> i)) (fin_map_fin_dom F)). rewrite image_empty. intro h1. apply times_set_empty'. Qed. Lemma el_prod_compose_empty : forall {B:Bool_Alg} (g:bt B->At) (F:Fin_map (Empty_set (bt B)) signe mns), el_prod_compose g F = 1. intros B g F. unfold el_prod_compose. generalize dependent (finite_image (bt B) (bt A) (Empty_set (bt B)) (fun i : bt B => eps (g i) (F |-> i)) (fin_map_fin_dom F)). rewrite image_empty. intro h1. apply times_set_empty'. Qed. Lemma el_prod_decompose_restriction : forall (E F:Ensemble At) (pf:Included F E) (a:Fin_map E signe mns), el_prod a = el_prod (restriction a pf) * el_prod (restriction a (setminus_inc E F)). intros E F h1 a. unfold el_prod. pose proof (decompose_setminus_inc _ _ h1) as h2. generalize (finite_image At (bt A) E (fun i : At => eps i (a |-> i)) (fin_map_fin_dom a)). pose proof (feq_im _ _ (fun i : At => eps i (a |-> i)) h2) as h3. rewrite im_union in h3. rewrite h3. intro h4. rewrite <- times_set_union. apply times_set_functional. apply Extensionality_Ensembles. red. split. red. intros y h5. destruct h5 as [y h5l | y h5r]. left. destruct h5l as [x h5l y h6]. rewrite h6. apply Im_intro with x. assumption. f_equal. rewrite restriction_compat; auto. right. destruct h5r as [x h5l y h6]. rewrite h6. apply Im_intro with x. assumption. f_equal. rewrite restriction_compat; auto. red. intros y h5. destruct h5 as [y h5l | y h5r]. left. destruct h5l as [x h5l y h6]. rewrite h6. apply Im_intro with x; auto. f_equal. rewrite restriction_compat; auto. destruct h5r as [x h5r y h6]. rewrite h6. right. apply Im_intro with x; auto. f_equal. rewrite restriction_compat; auto. Qed. Lemma el_prod_compose_decompose_restriction : forall {B:Bool_Alg} (g:bt B->At) (E F:Ensemble (bt B)) (pf:Included F E) (a:Fin_map E signe mns), el_prod_compose g a = el_prod_compose g (restriction a pf) * el_prod_compose g (restriction a (setminus_inc E F)). intros B g E F h1 a. unfold el_prod_compose. pose proof (decompose_setminus_inc _ _ h1) as h2. generalize (finite_image (bt B) (bt A) E (fun i : bt B => eps (g i) (a |-> i)) (fin_map_fin_dom a)). pose proof (feq_im _ _ (fun i : bt B => eps (g i) (a |-> i)) h2) as h3. rewrite im_union in h3. rewrite h3. intro h4. rewrite <- times_set_union. apply times_set_functional. apply Extensionality_Ensembles. red. split. red. intros y h5. destruct h5 as [y h5l | y h5r]. left. destruct h5l as [x h5l y h6]. rewrite h6. apply Im_intro with x. assumption. f_equal. rewrite restriction_compat; auto. right. destruct h5r as [x h5l y h6]. rewrite h6. apply Im_intro with x. assumption. f_equal. rewrite restriction_compat; auto. red. intros y h5. destruct h5 as [y h5l | y h5r]. left. destruct h5l as [x h5l y h6]. rewrite h6. apply Im_intro with x; auto. f_equal. rewrite restriction_compat; auto. destruct h5r as [x h5r y h6]. rewrite h6. right. apply Im_intro with x; auto. f_equal. rewrite restriction_compat; auto. Qed. Lemma eps_in_closed : forall (S:Ensemble At), alg_closed S -> forall x:At, Inn S x -> forall s:sign, Inn S (eps x s). intros S h1 x h2 s. destruct s. simpl. assumption. simpl. pose proof (C_c _ _ h1) as h3. red in h3. specialize (h3 (exist _ _ h2)). unfold Bcomp_sub in h3. simpl in h3. assumption. Qed. (*This might seem too lazy, but the only other option to define this function is to end the section here and reset the section variables and start a new section, then go back to these section variables.*) Lemma eps_in_closed' : forall {B:Bool_Alg} (S:Ensemble (bt B)), alg_closed S -> forall x:bt B, Inn S x -> forall s:sign, Inn S (eps x s). intros B S h1 x h2 s. destruct s. simpl. assumption. simpl. pose proof (C_c _ _ h1) as h3. red in h3. specialize (h3 (exist _ _ h2)). unfold Bcomp_sub in h3. simpl in h3. assumption. Qed. Lemma in_gen_ens_im_g_plus_subset_el_prod_compose_maps : forall {B:Bool_Alg} (E:Ensemble (bt B)) (pfe:Finite E) (g:(bt B)->At) (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_compose_maps g E)), Inn (Gen_Ens (Im E g)) (plus_subset_non_zero_el_prod_compose_maps _ pfe g S pfi). intros B E h1 g S h0. pose proof (closed_gen_ens _ (Im E g)) as h2. unfold plus_subset_non_zero_el_prod_compose_maps. apply plus_set_closed; auto. red. intros x h3. destruct h3 as [f h3]. subst. unfold el_prod_compose. apply times_set_closed; auto. red. intros x h4. destruct h4 as [x h4]. subst. apply eps_in_closed; auto. apply gen_ens_includes. apply Im_intro with x; auto. Qed. Lemma el_prod_in_gen : forall (E:Ensemble At) (F:Fin_map E signe mns), Inn (Gen_Ens E) (el_prod F). intros E F. pose proof (fin_map_fin_dom F) as h1. revert F. induction h1 as [|E h1 h2 x h3]. intros F. rewrite gen_ens_empty. rewrite el_prod_empty. right. intro F. assert (h4:Included E (Add E x)). auto with sets. specialize (h2 (restriction F h4)). rewrite (el_prod_decompose_restriction _ _ h4). assert (h5: el_prod (restriction F (setminus_inc (Add E x) E)) = eps x (F |-> x)). unfold el_prod. generalize (finite_image At (bt A) (Setminus (Add E x) E) (fun i : At => eps i (restriction F (setminus_inc (Add E x) E) |-> i)) (fin_map_fin_dom (restriction F (setminus_inc (Add E x) E)))). pose proof (setminus_add1 _ _ h3) as h5. pose proof (feq_im _ _ (fun i : At => eps i (restriction F (setminus_inc (Add E x) E) |-> i)) h5) as h6. rewrite h6. rewrite im_singleton. intro h7. rewrite times_set_sing'. f_equal. rewrite restriction_compat. reflexivity. constructor. right. constructor. assumption. rewrite h5. pose proof (gen_ens_preserves_inclusion _ _ h4) as h6. assert (h7:Inn (Gen_Ens (Add E x)) (el_prod (restriction F h4))). auto with sets. pose proof (closed_gen_ens _ (Add E x)) as h8. assert (h9:Inn (Gen_Ens (Add E x)) x). constructor. intros S h9. destruct h9 as [h9]. destruct h9 as [h9l h9r]. pose proof (Add_intro2 _ E x). auto with sets. pose proof (eps_in_closed _ h8 _ h9 (F |-> x)) as h10. apply times_closed; auto. Qed. Lemma plus_subset_non_zero_el_prod_maps_in_gen : forall (E:Ensemble At) (pf:Finite E) (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_maps E)), Inn (Gen_Ens E) (plus_subset_non_zero_el_prod_maps E pf S pfi). intros E h1 S. pose proof (finite_fin_map_ens S h1 signe_finite) as h2. induction h2 as [|S h3 h4 a h5]. intro h2. unfold plus_subset_non_zero_el_prod_maps. gterm0. redterm0 c. fold c0. generalize c0. rewrite image_empty. intro h3. rewrite plus_set_empty'. pose proof (closed_gen_ens _ E) as h4. apply zero_closed; auto. intro h6. assert (h7:Included S (Add S a)). auto with sets. assert (h8:Included S (non_zero_el_prod_maps E)). auto with sets. specialize (h4 h8). unfold plus_subset_non_zero_el_prod_maps. unfold Add. gterm0. redterm0 c. fold c0. generalize c0. rewrite im_union. intro h9. unfold plus_subset_non_zero_el_prod_maps in h4. pose proof (finite_image _ _ _ el_prod h3) as h10. pose proof (Singleton_is_finite _ a) as h11. pose proof (finite_image _ _ _ el_prod h11) as h12. pose proof (plus_set_union' _ _ h10 h12 h9) as h13. unfold At in h13. unfold At. rewrite h13 at 1. term0 h4. redterm0 c1. pose proof (proof_irrelevance _ c2 h10) as he. unfold c2 in he; subst. pose proof (closed_gen_ens _ E) as h15. apply plus_closed; auto. assert (h16: (plus_set (Im (Singleton a) el_prod) h12) = el_prod a). rewrite <- plus_set_sing. apply plus_set_functional. rewrite im_singleton. reflexivity. unfold At in h16. rewrite h16. apply el_prod_in_gen. Qed. Lemma gen_e_eq_c : forall (E:Ensemble At) (pf:Finite E), let C := [a:At | exists (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_maps E)), a = plus_subset_non_zero_el_prod_maps E pf S pfi] in Gen_Ens E = C. intros E h1 C. apply Extensionality_Ensembles. red. split. assert (h2:Included E C). red. intros x h2. pose proof (In_singleton _ x) as h3. assert (h4:Included (Singleton x) E). red. intros x' h4. destruct h4; subst. assumption. assert (h6:functionally_paired (Singleton x) signe (Singleton (x, pls))). constructor. intros x' h5. exists pls. red. split. split. unfold signe. left. destruct h5; subst. constructor. intros s h6. destruct h6 as [h6l h6r]. inversion h6r. reflexivity. intros pr h5. inversion h5 as [h6]. rewrite (surjective_pairing pr) in h6. inversion h6. subst. simpl. split. constructor. left. pose proof (Singleton_is_finite _ x) as h7. pose (fin_map_intro _ _ mns ba_eq_dec h7 signe_finite _ h6) as F. assert (h8: x = el_prod F). unfold F. unfold el_prod. symmetry. gen0. rewrite im_singleton. intro. rewrite times_set_sing'. simpl. pose proof (fin_fps_to_f_s_compat ba_eq_dec h7 h6 mns _ h3) as h9. inversion h9 as [he]. unfold At in he. unfold At. rewrite <- he at 1. simpl. reflexivity. rewrite h8. apply el_prod_inc_c. assumption. pose proof (gen_ens_preserves_inclusion _ _ h2) as h3. rewrite <- gen_ens_closed_eq. assumption. apply closed_all_plus_subsets_non_zero_el_prod_maps. red. intros x h2. destruct h2 as [h2]. destruct h2 as [S h2]. destruct h2 as [h2 h3]. subst. apply plus_subset_non_zero_el_prod_maps_in_gen. Qed. Lemma el_prod_compose_in_gen : forall {B:Bool_Alg} (E:Ensemble (bt B)) (g:bt B->At) (F:Fin_map E signe mns), Inn (Gen_Ens (Im E g)) (el_prod_compose g F). intros B E g F. pose proof (fin_map_fin_dom F) as h1. revert F. induction h1 as [|E h1 h2 x h3]. intros F. rewrite image_empty. rewrite gen_ens_empty. rewrite el_prod_compose_empty. right. intro F. assert (h4:Included E (Add E x)). auto with sets. specialize (h2 (restriction F h4)). rewrite (el_prod_compose_decompose_restriction g _ _ h4). assert (h5: el_prod_compose g (restriction F (setminus_inc (Add E x) E)) = eps (g x) (F |-> x)). unfold el_prod_compose. gterml. redterm0 x0. fold c. generalize c. pose proof (setminus_add1 _ _ h3) as h5. pose proof (feq_im _ _ (fun i : bt B => eps (g i) (restriction F (setminus_inc (Add E x) E) |-> i)) h5) as h6. rewrite h6, im_singleton. intro h7. rewrite times_set_sing'. f_equal. rewrite restriction_compat. reflexivity. constructor. right. constructor. assumption. rewrite h5. assert (h4':Included (Im E g) (Im (Add E x) g)). red. intros a h5'. destruct h5' as [a h6']. subst. apply Im_intro with a. left. assumption. reflexivity. pose proof (gen_ens_preserves_inclusion _ _ h4') as h6. assert (h7:Inn (Gen_Ens (Im (Add E x) g)) (el_prod_compose g (restriction F h4))). auto with sets. pose proof (closed_gen_ens _ (Im (Add E x) g)) as h8. assert (h9:Inn (Gen_Ens (Im (Add E x) g)) (g x)). constructor. intros S h9. destruct h9 as [h9]. destruct h9 as [h9l h9r]. pose proof (Add_intro2 _ E x). auto with sets. pose proof (eps_in_closed _ h8 _ h9 (F |-> x)) as h10. apply times_closed; auto. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_in_gen : forall {B:Bool_Alg} (g:bt B->At) (E:Ensemble (bt B)) (pf:Finite E) (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_compose_maps g E)), Inn (Gen_Ens (Im E g)) (plus_subset_non_zero_el_prod_compose_maps E pf g S pfi). intros B g E h1 S. pose proof (finite_fin_map_ens S h1 signe_finite) as h2. induction h2 as [|S h3 h4 a h5]. intro h2. unfold plus_subset_non_zero_el_prod_compose_maps. gterm0. redterm0 c. fold c0. generalize c0. rewrite image_empty. intro h3. rewrite plus_set_empty'. pose proof (closed_gen_ens _ (Im E g)) as h4. apply zero_closed; auto. intro h6. assert (h7:Included S (Add S a)). auto with sets. assert (h8:Included S (non_zero_el_prod_compose_maps g E)). auto with sets. specialize (h4 h8). unfold plus_subset_non_zero_el_prod_compose_maps. unfold Add. gterm0. redterm0 c. fold c0. generalize c0. rewrite im_union. intro h9. unfold plus_subset_non_zero_el_prod_compose_maps in h4. pose proof (finite_image _ _ _ (el_prod_compose g) h3) as h10. pose proof (Singleton_is_finite _ a) as h11. pose proof (finite_image _ _ _ (el_prod_compose g) h11) as h12. pose proof (plus_set_union' _ _ h10 h12 h9) as h13. unfold At in h13. unfold At. rewrite h13 at 1. term0 h4. redterm0 c1. pose proof (proof_irrelevance _ h10 c2) as he. unfold c2 in he. subst. pose proof (closed_gen_ens _ (Im E g)) as h15. apply plus_closed; auto. assert (h16: (plus_set (Im (Singleton a) (el_prod_compose g)) h12) = el_prod_compose g a). rewrite <- plus_set_sing. apply plus_set_functional. rewrite im_singleton. reflexivity. unfold At in h16. rewrite h16. apply el_prod_compose_in_gen. Qed. Lemma gen_e_eq_im_c : forall {B:Bool_Alg} (g:bt B->At) (E:Ensemble (bt B)) (pf:Finite E) , let C := [a:At | exists (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_compose_maps g E)), a = plus_subset_non_zero_el_prod_compose_maps E pf g S pfi] in Gen_Ens (Im E g) = C. intros B g E h1 C. apply Extensionality_Ensembles. red. split. assert (h2:Included (Im E g) C). red. intros x h2. destruct h2 as [x h2]. subst. assert (h3:Included (Singleton x) E). red. intros y h4. destruct h4. assumption. assert (h6:functionally_paired (Singleton x) signe (Singleton (x, pls))). constructor. intros x' h5. exists pls. red. split. split. unfold signe. left. destruct h5; subst. constructor. intros s h6. destruct h6 as [h6l h6r]. inversion h6r. reflexivity. intros pr h5. inversion h5 as [h6]. rewrite (surjective_pairing pr) in h6. inversion h6. subst. simpl. split. constructor. left. pose proof (Singleton_is_finite _ x) as h7. pose (fin_map_intro _ _ mns ba_eq_dec h7 signe_finite _ h6) as F. assert (h8: g x = el_prod_compose g F). unfold F. unfold el_prod_compose. symmetry. gen0. rewrite im_singleton. intro. rewrite times_set_sing'. simpl. pose proof (fin_fps_to_f_s_compat ba_eq_dec h7 h6 mns _ (In_singleton _ x)) as h9. inversion h9. simpl. reflexivity. rewrite h8. apply el_prod_compose_inc_c. assumption. pose proof (gen_ens_preserves_inclusion _ _ h2) as h3. rewrite <- gen_ens_closed_eq. assumption. apply closed_all_plus_subsets_non_zero_el_prod_compose_maps. red. intros x h2. destruct h2 as [h2]. destruct h2 as [S h2]. destruct h2 as [h2 h3]. subst. apply plus_subset_non_zero_el_prod_compose_maps_in_gen. Qed. (*The main theorem of this section, along with unprimed eponym.*) Theorem normal_form' : forall (E:Ensemble At) (pfe:Finite E), let Bt := (Btype (Bc (Gen E))) in (forall p:Bt, atom p <-> ((p <> 0) /\ exists f:Fin_map E signe mns, proj1_sig p = el_prod f)) /\ (forall b:Bt, exists !(S:Ensemble (Fin_map E signe mns)), exists (pf:Included S (non_zero_el_prod_maps E)), plus_subset_non_zero_el_prod_maps E pfe S pf = proj1_sig b). intros E h1 Bt. pose proof (gen_e_eq_c _ h1) as h2. pose proof (closed_gen_ens _ E) as h3. pose proof (closed_all_plus_subsets_non_zero_el_prod_maps _ h1) as h4. unfold Bt. unfold Gen. pose (fun (D:Ensemble At) (pf:alg_closed D) => (forall p : Btype (Bc (Subalg D pf)), atom p <-> p <> 0 /\ (exists f : Fin_map E signe mns, proj1_sig p = el_prod f)) /\ (forall b : Btype (Bc (Subalg D pf)), exists ! S : Ensemble (Fin_map E signe mns), exists pf : Included S (non_zero_el_prod_maps E), plus_subset_non_zero_el_prod_maps E h1 S pf = proj1_sig b)) as f. assert (h5:f _ (closed_gen_ens _ E) <-> f _ h4). generalize (closed_gen_ens _ E). rewrite h2. intro h6. pose proof (proof_irrelevance _ h4 h6); subst. tauto. unfold f in h5. simpl in h5. simpl. rewrite h5. split. intro p. split. intro h8. pose proof (atom_non_zero_el_prod _ h1 p) as h9. unfold atom in h8. unfold atom in h9. pose proof (proof_irrelevance _ h4 (closed_all_plus_subsets_non_zero_el_prod_maps E h1)); subst. pose proof h8 as h8'. destruct h8' as [h8l h8r]. simpl in h8l. split. unfold At in h8l. unfold At. assumption. rewrite h9 in h8. destruct h8 as [a h8]. exists a. destruct h8; assumption. intro h6. pose proof (atom_non_zero_el_prod _ h1 p) as h9. destruct h6 as [h6l h6r]. assert (h10: (exists a : Fin_map E signe mns, Inn (non_zero_el_prod_maps E) a /\ proj1_sig p = el_prod a)). destruct h6r as [a h6r]. exists a. split. constructor. rewrite <- h6r. intro h10. rewrite (unfold_sig _ p) in h6l. contradict h6l. apply proj1_sig_injective. simpl. assumption. assumption. rewrite <- h9 in h10. unfold atom. unfold atom in h10. pose proof (proof_irrelevance _ h4 (closed_all_plus_subsets_non_zero_el_prod_maps E h1)); subst. assumption. intro b. pose proof (proj2_sig b) as h7. simpl in h7. destruct h7 as [h7]. destruct h7 as [S [h7 h8]]. exists S. red. split. exists h7. symmetry. assumption. intros S' h10. destruct h10 as [h10l h10r]. rewrite h8 in h10r. symmetry in h10r. apply (plus_subset_non_zero_el_prod_maps_inj _ h1 _ _ h7 h10l h10r). Qed. Lemma gen_e_eq_c_unq : forall (E:Ensemble At) (pf:Finite E), let C := [a:At | exists! (S:Ensemble (Fin_map E signe mns)), exists (pfi:Included S (non_zero_el_prod_maps E)), plus_subset_non_zero_el_prod_maps E pf S pfi = a] in Gen_Ens E = C. intros E h1 C. unfold C. pose proof (normal_form' _ h1) as h2. simpl in h2. destruct h2 as [h2l h2r]. pose proof (gen_e_eq_c _ h1) as h3. simpl in h3. apply Extensionality_Ensembles. red. split. red. intros x h4. specialize (h2r (exist _ _ h4)). constructor. simpl in h2r. assumption. red. intros x h4. destruct h4 as [h4]. rewrite h3. constructor. destruct h4 as [S h4]. red in h4. destruct h4 as [h4a h4b]. destruct h4a as [h4a h4c]. exists S. exists h4a. rewrite h4c. reflexivity. Qed. Lemma plus_set_incl_sig : forall (E F:Ensemble At) (pfi:Included E F) (pfe:Finite E) (pfss:Finite (incl_sig E F pfi)) (pff:alg_closed F), proj1_sig (@plus_set (Subalg _ pff) (incl_sig E _ pfi) pfss) = plus_set E pfe. intros E F h1 h2 h3 h4. assert (h5:E = Im (incl_sig E F h1) (@proj1_sig _ _)). apply Extensionality_Ensembles. red. split. red. intros x h5. apply Im_intro with (exist _ _ (h1 _ h5)). constructor. simpl. reflexivity. red. intros x h5. destruct h5 as [x h5]. subst. rewrite <- (incl_sig_iff _ _ h1). assumption. pose proof h2 as h2'. rewrite h5 in h2'. assert (h6:plus_set E h2 = plus_set _ h2'). apply plus_set_functional. assumption. rewrite h6. unfold At. pose proof h2 as h2''. rewrite <- (finite_incl_sig_iff _ _ h1) in h2''. pose proof (plus_set_sub_compat _ _ h4 _ h2'' h2') as h9. pose proof (proof_irrelevance _ h2'' h3) as h10. rewrite h10 in h9. unfold At. unfold At in h9. rewrite <- h9. apply plus_set_functional. reflexivity. Qed. Lemma times_set_incl_sig : forall (E F:Ensemble At) (pfi:Included E F) (pfe:Finite E) (pfss:Finite (incl_sig E F pfi)) (pff:alg_closed F), proj1_sig (@times_set (Subalg _ pff) (incl_sig E _ pfi) pfss) = times_set E pfe. intros E F h1 h2 h3 h4. assert (h5:E = Im (incl_sig E F h1) (@proj1_sig _ _)). apply Extensionality_Ensembles. red. split. red. intros x h5. apply Im_intro with (exist _ _ (h1 _ h5)). constructor. simpl. reflexivity. red. intros x h5. destruct h5 as [x h5]. subst. rewrite <- (incl_sig_iff _ _ h1). assumption. pose proof h2 as h2'. rewrite h5 in h2'. assert (h6:times_set E h2 = times_set _ h2'). apply times_set_functional. assumption. rewrite h6. unfold At. pose proof h2 as h2''. rewrite <- (finite_incl_sig_iff _ _ h1) in h2''. pose proof (times_set_sub_compat _ _ h4 _ h2'' h2') as h9. pose proof (proof_irrelevance _ h2'' h3) as h7. rewrite <- h7. unfold At. unfold At in h9. rewrite <- h9. apply times_set_functional. reflexivity. Qed. (* More faithful to Givant/Halmos' form of the normal form. *) Theorem normal_form : forall (E:Ensemble At), let Bt := (Btype (Bc (Gen E))) in Finite E -> (forall p:Bt, atom p <-> ((p <> 0) /\ exists f:Fin_map E signe mns, proj1_sig p = el_prod f)) /\ (forall b:Bt, exists !(S:Ensemble Bt), exists pf:Finite S, (forall x:Bt, Inn S x -> atom x) /\ plus_set S pf = b). intros E Bt h1. pose proof (normal_form' E h1) as h2. destruct h2 as [h2l h2r]. split. assumption. intros b. specialize (h2r b). destruct h2r as [S h2r]. pose (Im S el_prod) as imS. assert (h3':Included (Im S el_prod) (Gen_Ens E)). red. intros y h4. destruct h4 as [y h4 x]. subst. apply el_prod_in_gen. exists (incl_sig _ _ h3'). red in h2r. red. destruct h2r as [h2a h2b]. destruct h2a as [h6 h7]. pose proof (finite_fin_map_ens S h1 signe_finite) as h3. pose proof (finite_image _ _ _ el_prod h3) as h4. rewrite <- (finite_incl_sig_iff _ _ h3') in h4. split. exists h4. split. intros x h5. inversion h5 as [c h8]. subst. destruct h8 as [a h8 y]. subst. assert (h9: (exist (Inn (Gen_Ens E)) (el_prod a) (h3' (el_prod a) (Im_intro (Fin_map E signe mns) At S el_prod a h8 (el_prod a) eq_refl))) <> (Bzero (Bc (Gen E))) /\ (exists f:Fin_map E signe mns, proj1_sig (exist (Inn (Gen_Ens E)) (el_prod a) (h3' (el_prod a) (Im_intro (Fin_map E signe mns) At S el_prod a h8 (el_prod a) eq_refl))) = el_prod f)). split. assert (h10:Inn (non_zero_el_prod_maps E) a). auto with sets. destruct h10 as [h10]. intro h11. pose proof (exist_injective _ _ _ _ _ h11). contradiction. exists a. simpl. reflexivity. rewrite <- h2l in h9. assumption. apply proj1_sig_injective. rewrite <- h7. unfold plus_subset_non_zero_el_prod_maps. pose proof (finite_image _ _ _ el_prod h3) as h8. pose proof (closed_gen_ens _ E) as h9. pose proof (plus_set_incl_sig _ _ h3' h8 h4 h9) as h10. symmetry. erewrite f_equal_dep. rewrite <- h10 at 1. pose proof (proof_irrelevance _ h9 (closed_gen_ens A E)); subst; auto. auto. intros D h6'. destruct h6' as [h6l h6r]. destruct h6r as [h6a h6b]. specialize (h2b [a:(Fin_map E signe mns) | Inn D (exist _ (el_prod a) (el_prod_in_gen E a)) /\ el_prod a <> 0]). assert (h7': Included [a : Fin_map E signe mns | Inn D (exist _ (el_prod a) (el_prod_in_gen E a)) /\ el_prod a <> 0] (non_zero_el_prod_maps E)). red. intros a h7'. destruct h7' as [h7']. destruct h7'. constructor. assumption. assert (h8:plus_subset_non_zero_el_prod_maps E h1 [a : Fin_map E signe mns | Inn D (exist _ (el_prod a) (el_prod_in_gen E a)) /\ el_prod a <> 0] h7' = proj1_sig b). unfold plus_subset_non_zero_el_prod_maps. assert (h8:Finite [a:Fin_map E signe mns | Inn D (exist _ (el_prod a) (el_prod_in_gen E a)) /\ el_prod a = 0]). apply finite_fin_map_ens. assumption. apply signe_finite. pose proof (finite_image _ _ _ el_prod h8) as h9. assert (h10:plus_set (Im [a:Fin_map E signe mns | Inn D (exist _ (el_prod a) (el_prod_in_gen E a)) /\ el_prod a = 0] el_prod) h9 = 0). destruct (classic (Inhabited [a:Fin_map E signe mns | Inn D (exist _ (el_prod a) (el_prod_in_gen E a)) /\ el_prod a = 0])) as [h10 | h11]. rewrite <- plus_set_sing. apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros x h11. destruct h11 as [x h11 x' h12]. subst. destruct h11 as [h11]. destruct h11 as [h11a h11b]. rewrite h11b. constructor. red. intros x h11. destruct h11; subst. destruct h10 as [x h10]. destruct h10 as [h10]. apply Im_intro with x. constructor. assumption. symmetry. destruct h10; auto. pose proof (not_inhabited_empty _ h11) as h12. rewrite <- plus_set_empty. apply plus_set_functional. pose proof (feq_im _ _ el_prod h12) as h13. rewrite image_empty in h13. assumption. subst. gterml. pose proof (f_equal (Bplus _ x) h10) as h11. unfold x in h11. rewrite zero_sum in h11. unfold At in h11. unfold At. rewrite <- h11 at 1. rewrite <- plus_set_union. pose proof (finite_image _ _ _ (@proj1_sig _ _) h6l) as h12. pose proof (plus_set_sub_compat _ _ (closed_gen_ens _ E) _ h6l h12) as h13. unfold At in h13. unfold At. unfold Gen. rewrite <- h13. unfold Gen in h11. apply plus_set_functional. rename x into oldx. assert (h12':(Union [a : Fin_map E signe mns | Inn D (exist _ (el_prod a) (el_prod_in_gen E a)) /\ el_prod a <> 0] [a : Fin_map E signe mns | Inn D (exist _ (el_prod a) (el_prod_in_gen E a)) /\ el_prod a = 0]) = [a : Fin_map E signe mns | Inn D (exist _ (el_prod a) (el_prod_in_gen E a))]). apply Extensionality_Ensembles. red. split. red. intros x h12'. destruct h12' as [x h12l | x h12r]. destruct h12l as [h12l]. destruct h12l; constructor; auto. destruct h12r as [h12r]. destruct h12r; constructor; auto. red. intros x h12'. destruct h12' as [h12']. destruct (classic_dec (el_prod x = 0)) as [h13' | h14']. right. constructor; auto. left. constructor; auto. rewrite <- im_union. unfold At in h12'. unfold At. rewrite h12' at 1. apply Extensionality_Ensembles. red. split. red. intros x h13'. destruct h13' as [x h13']. subst. destruct h13' as [h13']. apply Im_intro with (exist (Inn (Gen_Ens E)) (el_prod x) (el_prod_in_gen E x)). assumption. simpl. reflexivity. red. intros x h14. destruct h14 as [x h14]. subst. pose proof h6a as h15. clear h6a. specialize (h15 _ h14). rewrite h2l in h15. destruct h15 as [h15l h15r]. destruct h15r as [a h15r]. apply Im_intro with a. constructor. assert (h16:(exist (Inn (Gen_Ens E)) (el_prod a) (el_prod_in_gen E a)) = x). apply proj1_sig_injective. simpl. rewrite <- h15r. reflexivity. rewrite h16 at 1. assumption. assumption. hfirst h2b. exists h7'. assumption. specialize (h2b hf). subst. apply Extensionality_Ensembles. red. split. red. intros x h10. destruct h10 as [y h10]. destruct h10 as [x h10 y h11]. subst. destruct h10 as [h10]. destruct h10 as [h10l h10r]. term0 h10l. gterm0. assert (h11:c = c0). unfold c, c0; apply proj1_sig_injective; auto. unfold c, c0 in h11. simpl. unfold At in h11. unfold At. rewrite <- h11 at 1. assumption. red. intros x h11. specialize (h6a _ h11). pose proof h6a as h12. clear h6a. rewrite h2l in h12. destruct h12 as [h12l h12r]. destruct h12r as [a h13]. gterm1. redterm2 c. assert (h14:Inn c0 (proj1_sig x)). unfold c0. rewrite h13. apply Im_intro with a. constructor. split. rewrite (proj1_sig_eq_iff x) in h13. destruct h13 as [h13]; subst. erewrite f_equal_dep. apply h11. reflexivity. intro h14. rewrite h14 in h13. contradict h12l. apply proj1_sig_injective. simpl. assumption. reflexivity. unfold c0 in h14. pose proof (incl_sig_intro _ _ h3' _ h14) as h16. erewrite (unfold_sig _ x). pose proof (proof_irrelevance _ (h3' (proj1_sig x) h14) (proj2_sig _)) as h15. rewrite <- h15; auto. Qed. Corollary normal_form'' : forall (E:Ensemble At), let Bt := (Btype (Bc (Gen E))) in Finite E -> (forall p:Bt, atom p <-> ((p <> 0) /\ (exists! (S:Ensemble At), exists (pf:Finite S), proj1_sig p = times_set _ pf /\ exists f:Fin_map E signe mns, S = im_eps f))) /\ (forall b:Bt, exists !(S:Ensemble Bt), exists pf:Finite S, (forall x:Bt, Inn S x -> atom x) /\ plus_set S pf = b). intros E Bt h1. pose proof (normal_form _ h1) as h2. destruct h2 as [h2l h2r]. split; auto. clear h2r. intro p. split. intro h3. rewrite h2l in h3. destruct h3 as [h3l h3r]. split; auto. destruct h3r as [f h4]. exists (im_eps f). red. split. exists (finite_im_eps f). split. rewrite h4. unfold el_prod. apply times_set_functional. unfold im_eps. reflexivity. exists f. reflexivity. intros S h5. destruct h5 as [h5 [h6 h7]]. destruct h7 as [f' h8]. subst. rewrite h4 in h6. assert (h7:el_prod f <> 0). intro h7. contradict h3l. rewrite <- h4 in h7. apply proj1_sig_injective. simpl. assumption. assert (h8:el_prod f = el_prod f'). rewrite h6. unfold el_prod. apply times_set_functional. reflexivity. apply (non_zero_el_prod_inj f f' h7 h8). intro h3. destruct h3 as [h3l h3r]. destruct h3r as [S [h4 h5]]. destruct h4 as [h4a [h4b h4c]]. destruct h4c as [f h4c]. assert (h6:exists f : Fin_map E signe mns, proj1_sig p = el_prod f). exists f. rewrite h4b. unfold el_prod. apply times_set_functional. rewrite h4c. reflexivity. specialize (h2l p). pose proof (conj h3l h6) as h7. simpl in h7. simpl in h2l. pose proof (iff2 h2l h7) as h8. assumption. Qed. Corollary normal_form_gen_ba : forall (E:Ensemble (bt A)), Finite E -> ba_ens A = Gen_Ens E -> (forall p : At, atom p <-> p <> 0 /\ (exists f : Fin_map E signe mns, p = el_prod f)) /\ (forall b : At, exists ! S : Ensemble At, exists pf : Finite S, (forall x : At, Inn S x -> atom x) /\ plus_set S pf = b). intros E h1 h2. pose proof (normal_form _ h1) as h3. pose proof (gen_ba_eq_ba _ h2) as h4. unfold Gen in h3. pose (fun (C:Ensemble At) (pfc:alg_closed C) => (forall p : Btype (Bc (Subalg C pfc)), atom p <-> p <> 0 /\ (exists f : Fin_map E signe mns, proj1_sig p = el_prod f)) /\ (forall b : Btype (Bc (Subalg C pfc)), exists ! S : Ensemble (Btype (Bc (Subalg C pfc))), exists pf : Finite S, (forall x : Btype (Bc (Subalg C pfc)), Inn S x -> atom x) /\ plus_set S pf = b)) as P. assert (h5:P _ (closed_gen_ens A E) <-> P _ (alg_closed_ba_ens A)). pose proof (subsetT_eq_compat _ _ _ _ (alg_closed_ba_ens A) (closed_gen_ens A E) h2) as h6. dependent rewrite <- h6. tauto. pose proof (full_sig_equivalent_atom A) as h9. red in h9. simpl in h9. unfold P in h5. simpl in h5, h3. pose proof (iff1 h5 h3) as h6. clear h5 h3. destruct h6 as [h6 h7]. split. intro p. split. intro h8. specialize (h6 (exist _ _ (Full_intro _ p))). rewrite h9 in h8. unfold At in h6. simpl in h6, h8. rewrite h6 in h8. destruct h8 as [h8l h8r]. split. intro h10. subst. contradict h8l. apply proj1_sig_injective. simpl. reflexivity. assumption. intro h8. destruct h8 as [h8l h8r]. assert (h10: (exist _ _ (Full_intro _ p)) <> exist (Inn (Full_set (Btype (Bc A)))) 0 (Z_c A (Full_set (Btype (Bc A))) (alg_closed_ba_ens A))). intro h10. apply exist_injective in h10. contradiction. assert (h11:p = proj1_sig (exist (Inn (Full_set At)) p (Full_intro At p))). reflexivity. pose proof (conj h10 h8r) as h12. pose proof (h6 (exist (Inn (Full_set At)) p (Full_intro At p))) as h13. unfold ba_ens in h13. simpl in h13, h12. pose proof (iff2 h13 h12) as h14. clear h13 h12. rewrite h9. assumption. intros b. specialize (h7 (exist _ _ (Full_intro _ b))). destruct h7 as [S h7]. destruct h7 as [h7l h7r]. destruct h7l as [h7a [h7b h7c]]. exists (Im S (@proj1_sig _ _)). red. split. exists (finite_image _ _ _ _ h7a). split. intros x h8. destruct h8 as [x h8]. subst. specialize (h7b _ h8). destruct x as [x h10]. pose proof (proof_irrelevance _ h10 (Full_intro (Btype (Bc A)) x)); subst. pose proof (iff2 (h9 x) h7b) as h11. simpl. assumption. pose proof (f_equal (@proj1_sig _ _) h7c) as h10. simpl in h10. rewrite <- h10. rewrite (plus_set_sub_compat _ _ (alg_closed_ba_ens A) S h7a). reflexivity. intros S' h8. specialize (h7r (Im S' (fun x => (exist _ _ (Full_intro _ x))))). destruct h8 as [h8a [h8b h8c]]. simpl in h7r. assert (h10:exists pf : Finite (Im S' (fun x : At => exist (Inn (Full_set At)) x (Full_intro At x))), (forall x : SubBtype A (ba_ens A), Inn (Im S' (fun x0 : At => exist (Inn (Full_set At)) x0 (Full_intro At x0))) x -> atom (B:=Subalg (ba_ens A) (alg_closed_ba_ens A)) x) /\ plus_set (B:=Subalg (ba_ens A) (alg_closed_ba_ens A)) (Im S' (fun x : At => exist (Inn (Full_set At)) x (Full_intro At x))) pf = exist (Inn (Full_set At)) b (Full_intro At b)). exists (finite_image _ _ _ _ h8a). split. intros x h10. destruct h10 as [x h10]. subst. specialize (h8b _ h10). rewrite h9 in h8b. assumption. apply proj1_sig_injective. simpl. rewrite <- h8c. assert (h11:S' = Im (Im S' (fun x : At => exist (Inn (Full_set At)) x (Full_intro At x))) (@proj1_sig _ _)). rewrite im_im. simpl. rewrite im_id at 1. f_equal. assert (h12:Finite (Im (Im S' (fun x : At => exist (Inn (Full_set At)) x (Full_intro At x))) (proj1_sig (P:=Inn (Full_set At))))). apply finite_image; apply finite_image; auto. pose proof (subsetT_eq_compat _ _ _ _ h8a h12 h11) as h13. assert (h14:plus_set S' h8a = plus_set _ h12). dependent rewrite <- h13. reflexivity. rewrite h14. symmetry. pose proof (plus_set_sub_compat _ _ (alg_closed_ba_ens A) _ (finite_image At {x | Inn (Full_set At) x} S' (fun x : At => exist (Inn (Full_set At)) x (Full_intro At x)) h8a) h12) as h15. unfold ba_ens, At in h15. unfold At. simpl in h15. simpl. rewrite h15 at 1. reflexivity. specialize (h7r h10). clear h10. rewrite h7r. rewrite im_im. simpl. rewrite im_id. f_equal. Qed. Corollary normal_form_gen_ba' : forall (E:Ensemble (bt A)), Finite E -> ba_ens A = Gen_Ens E -> (forall p:At, atom p <-> ((p <> 0) /\ (exists! (S:Ensemble At), exists (pf:Finite S), p = times_set _ pf /\ exists f:Fin_map E signe mns, S = im_eps f))) /\ (forall b:At, exists !(S:Ensemble At), exists pf:Finite S, (forall x:At, Inn S x -> atom x) /\ plus_set S pf = b). intros E h1 h2. pose proof (normal_form'' _ h1) as h3. pose proof (gen_ba_eq_ba _ h2) as h4. unfold Gen in h3. pose (fun (C:Ensemble At) (pfc:alg_closed C) => (forall p : Btype (Bc (Subalg C pfc)), atom p <-> p <> 0 /\ (exists ! S : Ensemble At, exists pf : Finite S, proj1_sig p = times_set S pf /\ (exists f : Fin_map E signe mns, S = im_eps f))) /\ (forall b : Btype (Bc (Subalg C pfc)), exists ! S : Ensemble (Btype (Bc (Subalg C pfc))), exists pf : Finite S, (forall x : Btype (Bc (Subalg C pfc)), Inn S x -> atom x) /\ plus_set S pf = b)) as P. assert (h5:P _ (closed_gen_ens A E) <-> P _ (alg_closed_ba_ens A)). pose proof (subsetT_eq_compat _ _ _ _ (alg_closed_ba_ens A) (closed_gen_ens A E) h2) as h6. dependent rewrite <- h6. tauto. pose proof (full_sig_equivalent_atom A) as h9. red in h9. simpl in h9. unfold P in h5. simpl in h5, h3. pose proof (iff1 h5 h3) as h6. clear h5 h3. destruct h6 as [h6 h7]. split. intro p. split. intro h8. specialize (h6 (exist _ _ (Full_intro _ p))). rewrite h9 in h8. unfold At in h6. simpl in h6, h8. rewrite h6 in h8. destruct h8 as [h8l h8r]. split. intro h10. subst. contradict h8l. apply proj1_sig_injective. simpl. reflexivity. assumption. intro h8. destruct h8 as [h8l h8r]. assert (h10: (exist _ _ (Full_intro _ p)) <> exist (Inn (Full_set (Btype (Bc A)))) 0 (Z_c A (Full_set (Btype (Bc A))) (alg_closed_ba_ens A))). intro h10. apply exist_injective in h10. contradiction. assert (h11:p = proj1_sig (exist (Inn (Full_set At)) p (Full_intro At p))). reflexivity. pose proof (conj h10 h8r) as h12. pose proof (h6 (exist (Inn (Full_set At)) p (Full_intro At p))) as h13. unfold ba_ens in h13. simpl in h13, h12. pose proof (iff2 h13 h12) as h14. clear h13 h12. rewrite h9. assumption. intros b. specialize (h7 (exist _ _ (Full_intro _ b))). destruct h7 as [S h7]. destruct h7 as [h7l h7r]. destruct h7l as [h7a [h7b h7c]]. exists (Im S (@proj1_sig _ _)). red. split. exists (finite_image _ _ _ _ h7a). split. intros x h8. destruct h8 as [x h8]. subst. specialize (h7b _ h8). destruct x as [x h10]. pose proof (proof_irrelevance _ h10 (Full_intro (Btype (Bc A)) x)); subst. pose proof (iff2 (h9 x) h7b) as h11. simpl. assumption. pose proof (f_equal (@proj1_sig _ _) h7c) as h10. simpl in h10. rewrite <- h10. rewrite (plus_set_sub_compat _ _ (alg_closed_ba_ens A) S h7a). reflexivity. intros S' h8. specialize (h7r (Im S' (fun x => (exist _ _ (Full_intro _ x))))). destruct h8 as [h8a [h8b h8c]]. simpl in h7r. assert (h10:exists pf : Finite (Im S' (fun x : At => exist (Inn (Full_set At)) x (Full_intro At x))), (forall x : SubBtype A (ba_ens A), Inn (Im S' (fun x0 : At => exist (Inn (Full_set At)) x0 (Full_intro At x0))) x -> atom (B:=Subalg (ba_ens A) (alg_closed_ba_ens A)) x) /\ plus_set (B:=Subalg (ba_ens A) (alg_closed_ba_ens A)) (Im S' (fun x : At => exist (Inn (Full_set At)) x (Full_intro At x))) pf = exist (Inn (Full_set At)) b (Full_intro At b)). exists (finite_image _ _ _ _ h8a). split. intros x h10. destruct h10 as [x h10]. subst. specialize (h8b _ h10). rewrite h9 in h8b. assumption. apply proj1_sig_injective. simpl. rewrite <- h8c. assert (h11:S' = Im (Im S' (fun x : At => exist (Inn (Full_set At)) x (Full_intro At x))) (@proj1_sig _ _)). rewrite im_im. simpl. rewrite im_id at 1. f_equal. assert (h12:Finite (Im (Im S' (fun x : At => exist (Inn (Full_set At)) x (Full_intro At x))) (proj1_sig (P:=Inn (Full_set At))))). apply finite_image; apply finite_image; auto. pose proof (subsetT_eq_compat _ _ _ _ h8a h12 h11) as h13. assert (h14:plus_set S' h8a = plus_set _ h12). dependent rewrite <- h13. reflexivity. rewrite h14. symmetry. pose proof (plus_set_sub_compat _ _ (alg_closed_ba_ens A) _ (finite_image At {x | Inn (Full_set At) x} S' (fun x : At => exist (Inn (Full_set At)) x (Full_intro At x)) h8a) h12) as h15. unfold ba_ens, At in h15. unfold At. simpl in h15. simpl. rewrite h15 at 1. reflexivity. specialize (h7r h10). clear h10. rewrite h7r. rewrite im_im. simpl. rewrite im_id. f_equal. Qed. Corollary normal_form_gen_im : forall {B:Bool_Alg} (E:Ensemble (bt B)), Finite E -> forall (g:sig_set E->At), let C := Gen (Im (full_sig E) g) in (forall p : (bt C), atom p <-> p <> 0 /\ (exists f : Fin_map E signe mns, proj1_sig p = el_prod_compose (sig_fun_app g 0) f)) /\ (forall b : (bt C), exists ! S : Ensemble (bt C), exists pf : Finite S, (forall x : (bt C), Inn S x -> atom x) /\ plus_set S pf = b). intros B E h1 g C. pose proof h1 as h1'. rewrite finite_full_sig_iff in h1. pose proof (finite_image _ _ (full_sig E) g h1) as h5. pose proof (normal_form _ h5) as h6. destruct h6 as [h6a h6b]. split; auto. intro p. specialize (h6a p). unfold C. rewrite h6a. split. intro h7. destruct h7 as [h7l h7r]. split; auto. destruct h7r as [f h7r]. rewrite h7r. exists (fin_map_im_full_sig_eq ba_eq_dec f h1' 0). rewrite <- (el_prod_el_prod_compose_compat _ h1'). reflexivity. intro h7. destruct h7 as [h7l h7r]. split; auto. destruct h7r as [f h7r]. rewrite h7r. pose proof (ex_same_im_subset_sig_inj ba_eq_dec h1' g) as h8. destruct h8 as [D [h9 [h10 h11]]]. destruct h10 as [h10]. destruct h10 as [h10 h12]. rewrite h12. pose proof (proof_irrelevance _ h9 h10); subst. pose proof (fin_map_to_fin_map_im_full_sig_comp ba_eq_dec (restriction f h10) (restriction_sig g D h10) 0 h11) as h8. destruct h8 as [F' h8]. exists F'. unfold el_prod_compose, el_prod. apply times_set_functional. rewrite im_im. apply Extensionality_Ensembles. red. split. red. intros x h13. destruct h13 as [x h13]. subst. assert (h14:Inn (Im (full_sig E) g) (g (exist _ _ h13))). apply Im_intro with (exist _ _ h13). constructor. reflexivity. rewrite h12 in h14. inversion h14 as [d h15 a h16]. subst. apply Im_intro with d. constructor. f_equal. unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h17 | h18]. pose proof (proof_irrelevance _ h17 h13); subst; auto. contradiction. specialize (h8 _ (proj2_sig d)). simpl in h8. unfold sig_fun_app in h8. destruct (classic_dec (Inn D (proj1_sig d))) as [h17 | h18]. destruct d as [d h19]. simpl in h8. simpl in h17. pose proof (proof_irrelevance _ h19 h17); subst. unfold At in h8. unfold At. rewrite h8 at 1. rewrite restriction_compat. pose proof (el_prod_compose_le_ai f (sig_fun_app g 0) x h13) as h18. pose proof (el_prod_compose_le_ai f (sig_fun_app g 0) d (h10 _ h17)) as h19. rewrite <- h7r in h18, h19. unfold eps in h18, h19. unfold sig_fun_app in h18, h19. destruct (classic_dec (Inn E x)) as [h20 | h21]. destruct (classic_dec (Inn E d)) as [h22 | h23]. pose proof (proof_irrelevance _ h13 h20); subst. rewrite h16 in h18. unfold restriction_sig in h18. simpl in h18. pose proof (proof_irrelevance _ (h10 d h17) h22); subst. destruct (f |-> x), (f |-> d); auto. pose proof (mono_prod _ _ _ _ h18 h19) as h22. rewrite comp_prod in h22. rewrite idem_prod in h22. apply le_x_0 in h22. contradict h7l. apply proj1_sig_injective; auto. pose proof (mono_prod _ _ _ _ h19 h18) as h22. rewrite comp_prod in h22 at 1. rewrite idem_prod in h22. apply le_x_0 in h22. contradict h7l. apply proj1_sig_injective. assumption. contradict h23. apply h10. assumption. contradiction. assumption. contradict h18. apply proj2_sig. red. intros x h13. inversion h13 as [d h14 q h15]. subst. clear h13 h14. pose proof (h8 _ (proj2_sig d)) as h13. simpl in h13. unfold sig_fun_app in h13. destruct (classic_dec (Inn D (proj1_sig d))) as [h14 | h15]. destruct d as [d h15]. simpl in h14, h13. pose proof (proof_irrelevance _ h14 h15); subst. unfold At in h13. unfold At. rewrite h13 at 1. unfold restriction_sig. simpl. apply Im_intro with d. apply h10; auto. f_equal. unfold sig_fun_app. destruct (classic_dec (Inn E d)) as [h17 | h18]. f_equal. apply proj1_sig_injective. reflexivity. contradict h18. apply h10; auto. rewrite restriction_compat. reflexivity. assumption. contradict h15. apply proj2_sig. Qed. Lemma finite_atom_set_fin_gen : forall E:Ensemble At, Finite E -> Finite (atom_set (Gen E)). intros E h1. pose proof (normal_form E h1) as h2. unfold atom_set, bt. destruct h2 as [h2l h2r]. rewrite (sat_iff _ _ h2l). assert (h3:Finite [x:Btype (Bc (Gen E)) | (exists f : Fin_map E signe mns, proj1_sig x = el_prod f)]). gterm0. assert (h4:c = Im (Full_set (Fin_map E signe mns)) (fun f => (exist _ _ (el_prod_in_gen E f)))). unfold c. apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [h3]. destruct h3 as [F h3]. apply Im_intro with F. constructor. apply proj1_sig_injective. simpl. apply h3. red. intros x h3. destruct h3 as [F h3]. subst. constructor. simpl. exists F. reflexivity. unfold c in h4. rewrite h4. apply finite_image. apply finite_fin_maps. assumption. apply signe_finite. eapply Finite_downward_closed. apply h3. red. intros x h4. destruct h4 as [h4]. destruct h4 as [h4 h5]. constructor. assumption. Qed. Corollary card_atom_set_fin_gen : forall E:Ensemble At, Finite E -> (card_fun1 (atom_set (Gen E)) <= 2^(card_fun1 E)). intros E h1. pose proof (normal_form _ h1) as h2. unfold atom_set, bt. destruct h2 as [h2l h2r]. assert (h0: Finite (Im (Full_set (Fin_map E signe mns)) (fun f : Fin_map E signe mns => exist (Inn (Gen_Ens E)) (el_prod f) (el_prod_in_gen E f)))). apply finite_image. apply finite_fin_maps. assumption. apply signe_finite. rewrite (sat_iff _ _ h2l). assert (h4: [x : Btype (Bc (Gen E)) | exists f : Fin_map E signe mns, proj1_sig x = el_prod f] = Im (Full_set (Fin_map E signe mns)) (fun f => (exist _ _ (el_prod_in_gen E f)))). apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [h3]. destruct h3 as [F h3]. apply Im_intro with F. constructor. apply proj1_sig_injective. simpl. apply h3. red. intros x h3. destruct h3 as [F h3]. subst. constructor. simpl. exists F. reflexivity. terml h4. glel. fold xle. redterm0 xle. assert (h3:xle <= card_fun1 x). unfold x, xle. apply incl_card_fun1. rewrite h4. apply h0. red. clear x. intros x h3. destruct h3 as [h3]. destruct h3 as [? h3]. constructor. assumption. unfold x, xle in h3. rewrite h4 in h3. pose proof (finite_cardinal _ _ h0) as h5. destruct h5 as [n h5]. pose proof signe_finite as h6. pose proof (card_fin_maps ba_eq_dec _ _ mns h1 h6) as h7. rewrite card_signe in h7. rewrite FiniteT_card_fun1_compat in h7. pose proof (card_fun1_compat (Full_set (Fin_map E signe mns))) as h8. destruct h8 as [h8l h8r]. specialize (h8l (finite_fin_maps _ _ mns h1 h6)). unfold At in h7. unfold At in h8l. rewrite h7 in h8l at 1. pose proof (cardinal_decreases _ _ _ _ _ h8l _ h5) as h9. ler h3. redterm0 xle0. pose proof (card_fun1_compat c0) as h10. unfold c0 in h10. destruct h10 as [h10l h10r]. specialize (h10l h0). pose proof (cardinal_is_functional _ _ _ h10l _ _ h5 (eq_refl _)). subst. clear h10r h5 h8r. clear h0 h4 h6 h7. clear h2l h2r. clear h8l h10l. eapply le_trans. apply h3. apply h9. Qed. (*Uses LEM ([ens_eq_dec]).*) Corollary card_fin_gen : forall E:Ensemble At, Finite E -> (card_fun1 (Gen_Ens E) = 2^(card_fun1 (atom_set (Gen E)))). intros E h1. pose proof (normal_form _ h1) as h2. destruct h2 as [h2l h2r]. assert (h3: forall S:Ensemble (Btype (Bc (Gen E))), Inn (power_set (atom_set (Gen E))) S -> Finite S). intros S h3. destruct h3 as [h3]. eapply Finite_downward_closed; auto. apply finite_atom_set_fin_gen. apply h1. apply h3. assert (h4:forall (S:Ensemble (Btype (Bc (Gen E)))), Inn (power_set (atom_set (Gen E))) S -> exists ! b, Inn (Full_set (Btype (Bc (Gen E)))) b /\ exists pf:Finite S, (forall x : Btype (Bc (Gen E)), Inn S x -> atom x) /\ (plus_set S pf = b)). intros S h4. exists (plus_set S (h3 S h4)). red. split. split. constructor. exists (h3 S h4). split. destruct h4 as [h4]. red in h4. apply h4. reflexivity. intros b h5. destruct h5 as [h5 h6]. subst. destruct h6 as [h6 h7]. destruct h7; subst. apply plus_set_functional. reflexivity. assert (h5:forall b:Btype (Bc (Gen E)), Inn (Full_set (Btype (Bc (Gen E)))) b -> exists ! S:Ensemble (Btype (Bc (Gen E))), Inn (power_set (atom_set (Gen E))) S /\ exists pf : Finite S, (forall x : Btype (Bc (Gen E)), Inn S x -> atom x) /\ plus_set S pf = b). intros b ?. specialize (h2r b). destruct h2r as [S h2r]. red in h2r. destruct h2r as [h2a h2b]. destruct h2a as [h2c h2d]. exists S. red. split. split. constructor. red. destruct h2d as [h2d ?]. intros x h5. constructor. apply h2d; auto. exists h2c. subst. assumption. intros x h5. destruct h5 as [? h5]. apply h2b; auto. clear h2r. pose proof (finite_atom_set_fin_gen _ h1) as h6. pose proof (power_set_finite _ h6) as h7. pose proof (bij_ex_impl_eq_card' (ens_eq_dec (bt (Gen E))) _ _ h7 _ h4 h5) as h8. rewrite card_power_set in h8. rewrite h8. simpl. destruct (classic_dec (Finite (Gen_Ens E))) as [h9 | h10]. pose proof h9 as h9'. apply Finite_ens_type in h9. unfold SubBtype. unfold sig_set in h9. rewrite <- FiniteT_card_fun1_compat with h9. rewrite FiniteT_card_fun1_compat. pose proof (card_fun_full_sig_eq _ h9') as h10. do 2 rewrite card_fun_card_fun1_compat in h10. assumption. pose proof (card_fun1_compat (Gen_Ens E)) as h11. destruct h11 as [h11l h11r]. specialize (h11r h10). rewrite h11r. pose proof (card_fun1_compat (Full_set (SubBtype A (Gen_Ens E)))) as h12. destruct h12 as [h12l h12r]. assert (h13: ~ Finite (Full_set (SubBtype A (Gen_Ens E))) ). rewrite Finite_FiniteT_iff. intro h13. apply FiniteT_sig_Finite in h13. contradiction. rewrite h12r; auto. red; intros; apply ba_eq_dec. assumption. Qed. Corollary finite_fin_gen : forall E:Ensemble At, Finite E -> Finite (Gen_Ens E). intros E h1. pose proof (card_fin_gen _ h1) as h2. apply NNPP. intro h3. pose proof (card_fun1_compat (Gen_Ens E)) as h4. destruct h4 as [h4l h4r]. specialize (h4r h3). rewrite h4r in h2. assert (h5:0 < 2). omega. pose proof (O_lt_pow _ (card_fun1 (atom_set (Gen E))) h5). omega. Qed. Corollary card_fin_gen_le : forall E:Ensemble At, Finite E -> card_fun1 (Gen_Ens E) <= 2 ^ (2 ^ (card_fun1 E)). intros E h1. pose proof (card_atom_set_fin_gen _ h1) as h2. rewrite card_fin_gen; auto. apply pow_mono. auto with arith. assumption. Qed. Lemma incl_comp_set_closed : forall {B:Bool_Alg} (D E:Ensemble (Btype (Bc B))), alg_closed E -> Included D E -> Included (comp_set D) E. intros B D E h1 hi. red. intros x h2. destruct h2 as [x h2]. subst. apply comp_closed; auto. Qed. Definition comp_add_set {B:Bool_Alg} (A:Ensemble (Btype (Bc B))) := Union A (Im A (Bcomp (Bc B))). Lemma incl_comp_add_set : forall {B:Bool_Alg} (D:Ensemble (Btype (Bc B))), Included D (comp_add_set D). intros B D. unfold comp_add_set. auto with sets. Qed. Lemma incl_comp_add_set_closed : forall {B:Bool_Alg} (D E:Ensemble (Btype (Bc B))), alg_closed E -> Included D (comp_add_set E) -> Included D E. intros B D E h1 h2. red. red in h2. intros x h3. specialize (h2 _ h3). destruct h2 as [x h2l | x h2r]. assumption. destruct h2r as [x h4]. subst. apply comp_closed; auto. Qed. Lemma comp_add_set_add : forall {B:Bool_Alg} (D:Ensemble (Btype (Bc B))) (a:Btype (Bc B)), comp_add_set (Add D a) = Union (comp_add_set D) (Couple a (-a)). intros B D a. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1 | x h2]. destruct h1 as [x h1 | x h3]. left. left. assumption. destruct h3; subst. right. left. destruct h2 as [x h2]. subst. destruct h2 as [x h2l | x h2r]. left. right. econstructor. apply h2l. reflexivity. destruct h2r; subst. right. right. red. intros x h1. destruct h1 as [x h1 | x h2]. destruct h1 as [x h1 | x h3]. constructor. left. assumption. destruct h3 as [x h3]. subst. right. econstructor. left. apply h3. reflexivity. destruct h2. left. right. constructor. right. econstructor. right. constructor. reflexivity. Qed. Lemma comp_add_set_preserves_inclusion : forall {B:Bool_Alg} (D E:Ensemble (Btype (Bc B))), Included D E -> Included (comp_add_set D) (comp_add_set E). intros B D E h1. red. intros x h2. destruct h2 as [x h2l |x h2r]. left. apply (h1 _ h2l). right. destruct h2r as [x h2]. subst. apply Im_intro with x; auto. Qed. Lemma incl_comp_add_set_comp_add_set_incl : forall {B:Bool_Alg} (D E:Ensemble (Btype (Bc B))), Included D (comp_add_set E) -> Included (comp_add_set D) (comp_add_set E). intros B D E h1. red. intros x h2. destruct h2 as [x h2 | x h3]. apply h1; auto. destruct h3 as [x h3]. subst. apply h1 in h3. destruct h3 as [x h3 | x h4]. right. apply Im_intro with x; auto. destruct h4 as [x h4]. subst. rewrite <- doub_neg. left. assumption. Qed. Lemma im_eps_fin_map_incl_comp_add_set : forall {D:Ensemble At} (f:Fin_map D signe mns), Included (Im D (fun i => eps i (f |-> i))) (comp_add_set D). intros D f. red. intros x h1. destruct h1 as [x h1]. subst. destruct (f |-> x). simpl. constructor. assumption. simpl. unfold comp_add_set. right. apply Im_intro with x; auto. Qed. Lemma comp_add_set_union : forall (D E:Ensemble At), comp_add_set (Union D E) = Union (comp_add_set D) (comp_add_set E). intros D E. apply Extensionality_Ensembles. red; split; red; intros x h1. destruct h1 as [x h1 | x h2]. destruct h1 as [x h1 | x h2]. left. left. assumption. right. left. assumption. rewrite im_union in h2. destruct h2 as [x h2 | x h3]. left. right. assumption. right. right. assumption. destruct h1 as [x h1 | x h2]. destruct h1 as [x h1 | x h2]. left. left. assumption. right. rewrite im_union. left. assumption. destruct h2 as [x h2 | x h3]. left. right. assumption. right. rewrite im_union. right. assumption. Qed. Lemma incl_comp_add_set_intersection : forall (D E:Ensemble At), Included (comp_add_set (Intersection D E)) (Intersection (comp_add_set D) (comp_add_set E)). intros D E. unfold comp_add_set. red. intros x h1. destruct h1 as [x h1 | x h2]. destruct h1 as [x h1 h2]. constructor. left; auto with sets. left; auto with sets. destruct h2 as [x h2]. subst. destruct h2 as [x h2 h3]. constructor. right. apply Im_intro with x; auto. right. apply Im_intro with x; auto. Qed. Lemma intersection_comp_add_set_eq_old : forall (E F:Ensemble At), Intersection (comp_add_set E) (comp_add_set F) = Union (comp_add_set (Intersection E F)) (comp_add_set (Intersection (comp_set E) F)). intros E F. unfold comp_add_set, comp_set. rewrite foil_union. rewrite im_intersection_inj. rewrite im_intersection_inj. rewrite im_im. assert (h1:(fun x : bt A => - - x) = id). apply functional_extensionality. intros. rewrite <- doub_neg; auto. rewrite h1. rewrite <- im_id. pose proof (Union_commutative (Intersection (Im E (Bcomp (Bc A))) F) (Intersection E (Im F (Bcomp (Bc A))))) as h2. unfold bt, At in h2. unfold bt, At. rewrite h2. rewrite <- Union_associative. pose proof (Union_commutative (Intersection (Im E (Bcomp (Bc A))) (Im F (Bcomp (Bc A)))) (Union (Intersection E (Im F (Bcomp (Bc A)))) (Intersection (Im E (Bcomp (Bc A))) F))) as h3. unfold bt, At in h3. rewrite h3. rewrite <- Union_associative. reflexivity. apply inj_comp. apply inj_comp. Qed. (*This loosely is a set that doesn't contain any pair of complements.*) Definition dual_free_set (E:Ensemble At) : Prop := forall x:At, Inn E x -> ~Inn E (- x). Definition dual_freed_set (E:Ensemble At) := [x:At | Inn E x /\ ~Inn E (- x)]. Lemma dual_freed_set_compat : forall (E:Ensemble At), dual_free_set (dual_freed_set E). intro E. red. intros x h1. intro h2. destruct h1 as [h1]. destruct h2 as [h2]. tauto. Qed. Definition comp_set_dual_freed_set : forall (E:Ensemble At), comp_set (dual_freed_set E) = dual_freed_set (comp_set E). intro E. apply Extensionality_Ensembles. red; split; red; intros x h1. destruct h1 as [x h1]. subst. destruct h1 as [h1]. destruct h1 as [h1 h2]. constructor. split. apply Im_intro with x; auto. rewrite <- doub_neg. intro h3. destruct h3 as [x h3]. subst. rewrite <- doub_neg in h2. contradiction. destruct h1 as [h1]. destruct h1 as [h1 h2]. destruct h1 as [x h1]. subst. rewrite <- doub_neg in h2. apply Im_intro with x. constructor. split; auto. contradict h2. apply Im_intro with (- x); auto. rewrite <- doub_neg. reflexivity. reflexivity. Qed. Lemma dual_freed_set_disjoint_comp : forall (E:Ensemble At), disjoint (dual_freed_set E) (comp_set (dual_freed_set E)). intro E. red. apply Extensionality_Ensembles; red; split; auto with sets. red; intros x h1. destruct h1 as [x h1 h2]. destruct h1 as [h1]. destruct h1 as [h1 h3]. destruct h2 as [x h2]. subst. rewrite <- doub_neg in h3. destruct h2 as [h2]. destruct h2; contradiction. Qed. Lemma dual_free_set_comp_iff : forall (E:Ensemble At), dual_free_set E <-> dual_free_set (comp_set E). intro E. split. intro h1. red in h1. red. intros x h2. destruct h2 as [x h2]. subst. apply h1 in h2. contradict h2. rewrite <- doub_neg in h2. destruct h2. subst. rewrite <- doub_neg. assumption. intro h1. red in h1. red. intros x h2. assert (h3:Inn (comp_set E) (- x)). unfold comp_set. apply Im_intro with x; auto. apply h1 in h3. rewrite <- doub_neg in h3. contradict h3. apply Im_intro with (- x); auto. rewrite <- doub_neg. reflexivity. Qed. Lemma gen_ens_comp_add_set : forall {B:Bool_Alg} (E:Ensemble (Btype (Bc B))), Gen_Ens (comp_add_set E) = Gen_Ens E. intros B E. pose proof (closed_gen_ens _ E) as h0. pose proof (gen_ens_includes _ E) as hin. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. apply (h1 (Gen_Ens E)). constructor. split. red. intros b h2. destruct h2 as [b h2l | b h2r]. apply hin. assumption. destruct h2r as [b h2r]. subst. apply comp_closed; auto. assumption. red. intros x h1. pose proof (incl_comp_add_set E) as h2. pose proof (gen_ens_preserves_inclusion _ _ h2) as h3. auto with sets. Qed. (*This theorem says that every element in a subalgebra generated by E can be written as a finite join of finite meets of elements and complements of elements from E.*) Theorem gen_ens_eq : forall (E:Ensemble At), Gen_Ens E = [x : At | exists (S:Ensemble At) (pfs:Finite S), x = plus_set S pfs /\ forall s:At, Inn S s -> exists (R:Ensemble At) (pfr:Finite R), Included R (comp_add_set E) /\ s = times_set R pfr]. intro E. pose (fun (F : Ensemble At) (pf: Finite F) (pfinc:Included F E) => Gen_Ens F) as BF_set. pose (fun (F: Ensemble At) (pf:Finite F) => (non_zero_el_prod_maps F)) as non_zero_el_prod_maps_f. assert (hfin: forall (F:Ensemble At) (pf:Finite F), Finite (non_zero_el_prod_maps F)). intros. apply non_zero_el_prod_maps_fin; auto. assert (h1:forall (F:Ensemble At) (pf:Finite F) (pfinc:Included F E), BF_set F pf pfinc = [a:At | exists X (pfinc2:Included X (non_zero_el_prod_maps_f F pf)), a = plus_set (Im X el_prod) (finite_image _ _ _ el_prod (Finite_downward_closed _ _ (hfin F pf) _ pfinc2))]). intros F h1 h2. pose proof (normal_form' _ h1) as h3. unfold BF_set. destruct h3 as [h3l h3r]. apply Extensionality_Ensembles. red. split. red. intros x h4. specialize (h3r (exist _ _ h4)). destruct h3r as [X h3r]. red in h3r. destruct h3r as [h3a h3b]. destruct h3a as [h3a h3c]. constructor. exists X. exists h3a. simpl in h3c. rewrite <- h3c. unfold plus_subset_non_zero_el_prod_maps. apply plus_set_functional. reflexivity. red. intros x h4. destruct h4 as [h4]. destruct h4 as [S h5]. destruct h5 as [h5 h6]. assert (h7:forall f, Inn S f -> Inn (Gen_Ens F) (el_prod f)). intros f h7. apply el_prod_in_gen; auto. rewrite h6. pose proof (closed_gen_ens _ F) as h8. apply plus_set_closed. assumption. red. intros x' h9. destruct h9 as [G h9]. subst. apply h7. assumption. simpl in BF_set. pose proof (family_included_subs_directed_eq_B' E) as h2. destruct h2 as [h2l h2r]. rewrite <- h2r. apply Extensionality_Ensembles. red. split. red. intros C h3. destruct h3 as [C x h4 h5]. destruct h4 as [h4]. destruct h4 as [F h4]. destruct h4 as [h4a [h4b h4c]]. subst. pose proof (normal_form _ h4c) as h6. destruct h6 as [h6l h6r]. destruct (h6r (exist _ _ h5)) as [S h6b]. destruct h6b as [h6a h6b]. destruct h6a as [h6a h6c]. destruct h6c as [h6c h6d]. constructor. pose proof (f_equal (@proj1_sig _ _) h6d) as h7. simpl in h7. pose proof (plus_set_sub_compat _ _ _ S h6a (finite_image _ _ S (@proj1_sig _ _) h6a)) as h8. simpl in h7. simpl in h8. unfold Gen in h7. rewrite h7 in h8. exists (Im S (proj1_sig (P:=fun a : Btype (Bc A) => Inn (Gen_Ens F) a))). exists (finite_image _ _ S (@proj1_sig _ _) h6a). rewrite <- h8. split. reflexivity. intros a h9. destruct h9 as [a h9]. subst. specialize (h6c _ h9). rewrite h6l in h6c. destruct h6c as [h10 h11]. destruct h11 as [f h11]. unfold el_prod in h11. exists (Im F (fun i : At => eps i (f |-> i))). exists (finite_image _ _ _ _ h4c). split. pose proof (im_eps_fin_map_incl_comp_add_set f) as h12. pose proof (comp_add_set_preserves_inclusion _ _ h4b) as h13. auto with sets. unfold At in h11. unfold At. pose proof (proof_irrelevance _ h4c (fin_map_fin_dom f)); subst; auto. red. intros x h3. destruct h3 as [h3]. destruct h3 as [S h3]. destruct h3 as [h3 h4]. rewrite h2r. destruct h4 as [h4 h5]. subst. apply plus_set_closed; auto. apply closed_gen_ens. red. intros x h6. specialize (h5 _ h6). destruct h5 as [R h5]. destruct h5 as [h5 h7]. destruct h7 as [h7l h7r]. subst. assert (h8:Inn (Gen_Ens R) (times_set R h5)). pose proof (closed_gen_ens _ R) as h7. pose proof (gen_ens_includes _ R) as h8'. pose proof (times_set_closed _ R (Gen_Ens R) h5 h7 h8') as h10. assumption. pose proof (gen_ens_preserves_inclusion _ _ h7l) as h9. rewrite gen_ens_comp_add_set in h9. auto with sets. Qed. Lemma el_prod_fin_map_dom_subst_compat : forall {C C':Ensemble At} (a:Fin_map C signe mns) (pf:C = C'), el_prod a = el_prod (fin_map_dom_subst pf a). intros C C' a h1. subst. rewrite fin_map_dom_subst_eq_refl. reflexivity. Qed. Lemma el_prod_compose_fin_map_dom_subst_compat : forall {B:Bool_Alg} {C C':Ensemble (bt B)} (a:Fin_map C signe mns) (g:bt B->bt A) (pf:C = C'), el_prod_compose g a = el_prod_compose g (fin_map_dom_subst pf a). intros B C C' a g h1. subst. rewrite fin_map_dom_subst_eq_refl. reflexivity. Qed. End NormalForm. Lemma im_only_has_sames_eps : forall {B:Bool_Alg} {n:nat} (f:nat->bt B) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pfo:only_has_sames' (map_seg (fun_to_seg f n)) (seg n) F), let F' := only_has_sames_map_seg_to_fin_map f F pfo in im_in_ens (seg_fun_to_ens (fun (i : nat) (pf : i < length (map_seg (fun_to_seg f n))) => eps (nth_lt (map_seg (fun_to_seg f n)) i pf) (F |-> i))) = Im (Im (seg n) f) (fun i : Btype (Bc B) => eps i (F' |-> i)). intros B n f F h1 F'. rewrite im_im. apply Extensionality_Ensembles; red; split; red; intros i h2. destruct h2 as [i h2]; subst. destruct h2 as [h2]. pose proof h2 as h2'. rewrite length_map_seg in h2'. econstructor. constructor. apply h2'. rewrite seg_fun_to_ens_compat. rewrite nth_lt_map_seg'. rewrite fun_to_seg_compat. f_equal. unfold F'. rewrite only_has_sames_map_seg_to_fin_map_compat; auto. destruct h2 as [i h2]; subst. destruct h2 as [h2]. econstructor. rewrite seg_fun_to_ens_compat. f_equal. rewrite nth_lt_map_seg', fun_to_seg_compat. reflexivity. unfold F'. rewrite only_has_sames_map_seg_to_fin_map_compat; auto. Grab Existential Variables. rewrite length_map_seg. constructor; auto. Qed. Lemma only_has_sames_map_el_prod : forall {B:Bool_Alg} {n:nat} (f:nat->bt B) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pfo:only_has_sames' (map_seg (fun_to_seg f n)) (seg n) F), let F' := only_has_sames_map_seg_to_fin_map f F pfo in el_prod_l _ F = el_prod _ F'. intros B n f F h1 F'. unfold el_prod_l, el_prod. rewrite times_list_compat. gen0. rewrite list_to_set_map_seg. intro h2. apply f_equal_dep. apply im_only_has_sames_eps. Qed. Lemma im_only_has_sames_eps_compose : forall {A B:Bool_Alg} {n:nat} (f:nat->bt B) (g:bt B->bt A) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pfo:only_has_sames' (map_seg (fun_to_seg f n)) (seg n) F), let F' := only_has_sames_map_seg_to_fin_map f F pfo in im_in_ens (seg_fun_to_ens (fun (i : nat) (pf : i < length (map_seg (fun_to_seg f n))) => eps (g (nth_lt (map_seg (fun_to_seg f n)) i pf)) (F |-> i))) = Im (Im (seg n) f) (fun i : bt B => eps (g i) (F' |-> i)). intros A B n f g F h1 F'. apply Extensionality_Ensembles; red; split; red; intros i h2. destruct h2 as [i h2]; subst. pose proof h2 as h2'. rewrite length_map_seg in h2'. destruct h2' as [h3]. rewrite im_im. apply Im_intro with i. constructor; auto. rewrite seg_fun_to_ens_compat. rewrite nth_lt_map_seg', fun_to_seg_compat. f_equal. unfold F'. rewrite only_has_sames_map_seg_to_fin_map_compat; auto. rewrite im_im in h2. destruct h2 as [j h2]; subst. destruct h2 as [h2]. econstructor. rewrite seg_fun_to_ens_compat. rewrite nth_lt_map_seg', fun_to_seg_compat. unfold F'. f_equal. apply only_has_sames_map_seg_to_fin_map_compat; auto. Grab Existential Variables. rewrite length_map_seg. constructor; auto. Qed. Lemma only_has_sames_map_el_prod_compose : forall {A B:Bool_Alg} (n:nat) (f:nat->bt B) (g:bt B->bt A) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pf:only_has_sames' (map_seg (fun_to_seg f n)) (seg n) F), let F' := only_has_sames_map_seg_to_fin_map f F pf in el_prod_l_compose _ F g = el_prod_compose _ g F'. intros A B n f g F h1 F'. unfold el_prod_l_compose, el_prod_compose. rewrite times_list_compat. gen0. rewrite list_to_set_map_seg. intro h2. apply f_equal_dep. apply im_only_has_sames_eps_compose. Qed. Lemma im_only_has_sames_im_im_lind_singularize : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfle:m <= length l) (F:Fin_map (seg m) signe mns), only_has_sames l m pfle (seg m) F -> im_in_ens (seg_fun_to_ens (fun (i : nat) (pf : i < length (firstn m l)) => eps (nth_lt (firstn m l) i pf) (fin_map_seg_transfer F (eq_sym (length_firstn l m pfle)) |-> i))) = Im (im_im_lind_singularize ba_eq_dec l m pfle) (fun i : Btype (Bc B) => eps i (fin_map_seg_to_im_in_ens ba_eq_dec F l pfle |-> i)). intros B l m h1 F h2. destruct h2 as [h2a h2b]. unfold seg_fun_to_ens. apply Extensionality_Ensembles; red; split; red; intros x h3. destruct h3 as [x h3]; subst; destruct h3 as [h3]. unfold im_im_lind_singularize. rewrite im_im_in_ens. econstructor. rewrite fin_map_seg_transfer_compat'; auto. f_equal. symmetry. erewrite nth_lt_functional3. erewrite <- lind_singularize_compat. rewrite nth_lt_firstn_eq. reflexivity. erewrite (fin_map_seg_to_im_in_ens_compat ba_eq_dec F l h1). match goal with |- _ = F |-> ?x' => destruct (nat_eq_dec x x') as [h4 | h4] end. rewrite h4 at 1; auto. eapply h2b. pose proof h3 as h3'. rewrite length_firstn in h3'. constructor; auto. assumption. constructor. apply lt_inv_inv_im_im_lind_singularize. assumption. rewrite inv_inv_im_im_lind_singularize_compat. unfold lind_singularize. symmetry. erewrite nth_lt_functional3. rewrite lind_compat. apply nth_lt_functional3. unfold im_im_lind_singularize in h3. rewrite im_im_in_ens in h3. destruct h3 as [i h3]; subst. pose proof h3 as h4. destruct h4 as [i h4]; subst. simpl in h3. pose proof h4 as h4'. destruct h4'. econstructor. simpl. f_equal. unfold lind_singularize. erewrite nth_lt_functional3. rewrite lind_compat. rewrite nth_lt_firstn_eq. apply nth_lt_functional3. erewrite (fin_map_seg_to_im_in_ens_compat ba_eq_dec F l h1). rewrite fin_map_seg_transfer_compat'. symmetry. match goal with |- _ = F |-> ?x' => destruct (nat_eq_dec i x') as [h5 | h5] end. rewrite h5 at 1; auto. eapply h2b. assumption. constructor. apply lt_inv_inv_im_im_lind_singularize. assumption. rewrite inv_inv_im_im_lind_singularize_compat. unfold lind_singularize. symmetry. erewrite nth_lt_functional3. rewrite lind_compat. apply nth_lt_functional3. rewrite length_firstn; auto. Grab Existential Variables. assumption. unfold im_im_lind_singularize. econstructor. reflexivity. constructor; rewrite length_firstn; auto. pose proof h3 as h3'. rewrite length_firstn in h3'; auto. unfold im_im_lind_singularize. econstructor. reflexivity. Grab Existential Variables. unfold im_lind_singularize. econstructor. apply f_equal_dep. reflexivity. Grab Existential Variables. constructor. rewrite length_firstn in h3; auto. Qed. Lemma only_has_sames_el_prod_fin_map_seg_to_im_in_ens : forall {B:Bool_Alg} (l:list (bt B)) (m:nat) (pfm:m <= length l) (F:Fin_map (seg m) signe mns), only_has_sames l m pfm (seg m) F -> el_prod _ (fin_map_seg_to_im_in_ens ba_eq_dec F l pfm) = el_prod_l_firstn _ pfm F. intros B l m h1 F h2. unfold el_prod, el_prod_l_firstn, el_prod_l. rewrite times_list_compat. symmetry. gen0. rewrite list_to_set_map_seg. intro h3. apply f_equal_dep. apply im_only_has_sames_im_im_lind_singularize; auto. Qed. Lemma im_only_has_sames_im_im_lind_singularize_compose : forall {B A:Bool_Alg} (lb:list (bt B)) (la:list (bt A)) (m:nat) (pfm:m <= length lb) (pfe:length lb = length la) (F:Fin_map (seg m) signe mns) (pfle:linds ba_eq_dec lb = linds ba_eq_dec la), let f := proj1_sig_fun (two_lists_to_sig_fun ba_eq_dec ba_eq_dec lb la pfle) in let pfm' := le_eq_trans _ _ _ pfm pfe in only_has_sames la m pfm' (seg m) F -> im_in_ens (seg_fun_to_ens (fun (i : nat) (pf : i < length (firstn m la)) => eps (nth_lt (firstn m la) i pf) (fin_map_seg_transfer F (eq_sym (length_firstn la m pfm')) |-> i))) = Im (im_im_lind_singularize ba_eq_dec lb m pfm) (fun i : bt B => eps (sig_fun_app f 0 i) (fin_map_seg_to_im_in_ens ba_eq_dec F lb pfm |-> i)). intros B A lb la m hm he F hle f h1 h2. unfold im_im_lind_singularize. rewrite im_im_in_ens. destruct h2 as [h2 h3]. unfold seg_fun_to_ens. apply Extensionality_Ensembles; red; split; intros i h4. destruct h4 as [i h4]; subst. pose proof h4 as h4'. rewrite length_firstn in h4'; auto. destruct h4' as [h4']. match goal with |- Inn (im_in_ens ?g) _ => pose proof (im_in_ens_intro g) as h5 end. specialize (h5 (lind_singularize ba_eq_dec lb i (lt_le_trans i m (length lb) h4' hm))). hfirst h5. unfold im_lind_singularize. econstructor. apply f_equal_dep. reflexivity. specialize (h5 hf (eps (nth_lt (firstn m la) i (seg_in i (length (firstn m la)) h4)) (fin_map_seg_transfer F (eq_sym (length_firstn la m h1)) |-> i))). hfirst h5. unfold sig_fun_app. lr_match_goal; fold hlr; fold hlr; destruct hlr as [hl | hr]. unfold f. rewrite proj1_sig_fun_compat. f_equal. unfold lind_singularize. generalize hl. unfold lind_singularize. erewrite nth_lt_functional3. rewrite lind_compat. intro h7. symmetry. erewrite f_equal_dep. rewrite two_lists_to_sig_fun_compat. simpl. rewrite nth_lt_firstn_eq. apply nth_lt_functional3. apply nth_lt_functional3. rewrite fin_map_seg_transfer_compat'. unfold lind_singularize. erewrite nth_lt_functional3. rewrite lind_compat. pose proof (fin_map_seg_to_im_in_ens_compat ba_eq_dec F lb hm (nth_lt _ _ (lt_le_trans _ _ _ h4' hm))) as h6. hfirst h6. unfold im_im_lind_singularize. match goal with |- Inn (im_in_ens ?g) _ => pose proof (im_in_ens_intro g) as h7 end. specialize (h7 (lind_singularize ba_eq_dec lb i (lt_le_trans _ _ _ h4' hm)) hf (nth_lt _ _ (lt_le_trans _ _ _ h4' hm))). hfirst h7. unfold lind_singularize. erewrite nth_lt_functional3. erewrite lind_compat at 1. reflexivity. specialize (h7 hf0); auto. specialize (h6 hf0). rewrite h6 at 1. match goal with |- _ = F |-> ?r => destruct (nat_eq_dec i r) as [h8 | h8] end. rewrite h8 at 1; auto. eapply h3. constructor; auto. constructor. apply lt_inv_inv_im_im_lind_singularize. assumption. symmetry. pose proof hle as h10. rewrite linds_eq_iff' in h10. simpl in h10. erewrite nth_lt_functional3. symmetry. erewrite nth_lt_functional3. erewrite <- h10 at 1. rewrite inv_inv_im_im_lind_singularize_compat. apply nth_lt_functional3. rewrite length_firstn; auto. contradict hr. rewrite <- list_to_set_in_iff. rewrite (list_singularize_in_iff). apply in_nth_lt. apply h5; auto. destruct h4 as [i h4]; subst. pose proof h4 as h4'. destruct h4' as [i h5]; subst. simpl in h4, h5. destruct h5 as [h5]. simpl. econstructor. unfold sig_fun_app. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. unfold f. f_equal. rewrite proj1_sig_fun_compat. unfold lind_singularize. generalize hl. unfold im_lind_singularize, lind_singularize. erewrite nth_lt_functional3. rewrite lind_compat. pose proof (two_lists_to_sig_fun_compat ba_eq_dec ba_eq_dec lb la hle i (lt_le_trans i m (length lb) (seg_in i m (intro_characteristic_sat (fun m0 : nat => m0 < m) i h5)) hm)) as h6. simpl in h6. intro h7. erewrite f_equal_dep. rewrite h6 at 1. simpl. rewrite nth_lt_firstn_eq. apply nth_lt_functional3. reflexivity. rewrite (fin_map_seg_to_im_in_ens_compat ba_eq_dec F lb hm) at 1. rewrite fin_map_seg_transfer_compat'. match goal with |- F |-> ?r = _ => destruct (nat_eq_dec i r) as [h8 | h8] end. symmetry. rewrite h8 at 1; auto. symmetry. eapply h3. constructor; auto. constructor. apply lt_inv_inv_im_im_lind_singularize. assumption. pose proof hle as h10. rewrite linds_eq_iff' in h10. simpl in h10. erewrite nth_lt_functional3. symmetry. erewrite nth_lt_functional3. rewrite <- h10. rewrite inv_inv_im_im_lind_singularize_compat. unfold lind_singularize. erewrite nth_lt_functional3. rewrite lind_compat. apply nth_lt_functional3. rewrite length_firstn; auto. contradict hr. rewrite <- list_to_set_in_iff. rewrite list_singularize_in_iff. apply in_nth_lt. Grab Existential Variables. eapply lt_le_trans. apply h5. assumption. assumption. apply lt_inv_inv_im_im_lind_singularize. assumption. econstructor. simpl. apply nth_lt_functional3. constructor. rewrite length_firstn; auto. eapply lt_le_trans. apply h4'. assumption. symmetry; auto. apply lt_inv_inv_im_im_lind_singularize. assumption. eapply lt_le_trans. apply h4'. assumption. constructor; auto. Grab Existential Variables. assumption. Qed. Lemma only_has_sames_el_prod_compose_fin_map_seg_to_im_in_ens : forall {B A:Bool_Alg} (lb:list (bt B)) (la:list (bt A)) (m:nat) (pfm:m <= length lb) (pfe:length lb = length la) (F:Fin_map (seg m) signe mns) (pfle:linds ba_eq_dec lb = linds ba_eq_dec la), let f := proj1_sig_fun (two_lists_to_sig_fun ba_eq_dec ba_eq_dec lb la pfle) in let pfm' := le_eq_trans _ _ _ pfm pfe in only_has_sames la m pfm' (seg m) F -> el_prod_compose _ (sig_fun_app f 0) (fin_map_seg_to_im_in_ens ba_eq_dec F lb pfm) = el_prod_l_firstn _ pfm' F. intros B A lb la m hm he F hle f h1 h2. unfold el_prod_compose, el_prod_l_firstn, el_prod_l. rewrite times_list_compat. apply f_equal_dep. rewrite list_to_set_map_seg. symmetry. apply im_only_has_sames_im_im_lind_singularize_compose; auto. Qed. Lemma gen_ens_subalg_eq : forall {B:Bool_Alg} {C:Ensemble (bt B)} (pf:alg_closed C), let D:=Subalg C pf in forall (E:Ensemble (bt D)), im_proj1_sig (Gen_Ens E) = Gen_Ens (im_proj1_sig E). intros B C h0 D E. do 2 rewrite gen_ens_eq. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. subst. destruct h1 as [h1]. destruct h1 as [S [h2 [h3 h4]]]. subst. constructor. exists (im_proj1_sig S). exists (finite_image _ _ _ _ h2). split. rewrite (plus_set_sub_compat B C h0 S h2). reflexivity. intros b h5. destruct h5 as [b h5]. subst. destruct b as [b h6]. simpl. specialize (h4 _ h5). destruct h4 as [R [h4a [h4b h4c]]]. exists (im_proj1_sig R). exists (finite_image _ _ _ _ h4a). split. red. red in h4b. intros x h7. destruct h7 as [x h7]. subst. destruct x as [x h8]. simpl. specialize (h4b _ h7). inversion h4b as [y h9 | y h10]; subst. clear h4b. left. apply Im_intro with (exist (fun x0 : Btype (Bc B) => Inn C x0) x h8); auto. right. inversion h10 as [y h11 ? h12]. subst. destruct y as [y h13]. unfold im_proj1_sig. rewrite im_im. apply Im_intro with (exist (fun a : Btype (Bc B) => Inn C a) y h13); auto. simpl. simpl in h12. apply exist_injective in h12. subst. unfold Bcomp_sub. simpl. reflexivity. rewrite (times_set_sub_compat B C h0 R h4a). pose proof (f_equal (@proj1_sig _ _) h4c) as h7. simpl in h7. assumption. red. intros x h1. destruct h1 as [h1]. destruct h1 as [S [h2 [h3 h4]]]. subst. assert (h5:Included [x:(bt B) | exists (S0 : Ensemble (bt B)) (pfs:Finite S0), x = plus_set S0 pfs /\ forall (s:(bt B)), Inn S0 s -> exists (R:Ensemble (bt B)) (pfr:Finite R), Included R (comp_add_set (im_proj1_sig E)) /\ s = times_set R pfr] C). red. intros x h5. destruct h5 as [h5]. destruct h5 as [R [h5 [h6 h7]]]. subst. apply plus_set_closed; auto. red. intros x h8. specialize (h7 _ h8). destruct h7 as [R' [h9 [h10 h11]]]. subst. apply times_set_closed; auto. intros r h11. specialize (h10 _ h11). destruct h10 as [r h10 | r h12]. destruct h10 as [r h12 ? h13]. subst. apply proj2_sig. destruct h12 as [r h13 ? h14]. subst. apply comp_closed; auto. destruct h13 as [r h13]. subst. apply proj2_sig. assert (h6: [x : Btype (Bc D) | exists (S0 : Ensemble (Btype (Bc D))) (pfs : Finite S0), x = plus_set S0 pfs /\ (forall s : Btype (Bc D), Inn S0 s -> exists (R : Ensemble (Btype (Bc D))) (pfr : Finite R), Included R (comp_add_set E) /\ s = times_set R pfr)] = im_proj2_sig _ h5). apply Extensionality_Ensembles. red. split. red. intros x h6. destruct x as [x h7]. assert (h8:Inn [x0 : bt B | exists (S0 : Ensemble (bt B)) (pfs : Finite S0), x0 = plus_set S0 pfs /\ (forall s : bt B, Inn S0 s -> exists (R : Ensemble (bt B)) (pfr : Finite R), Included R (comp_add_set (im_proj1_sig E)) /\ s = times_set R pfr)] x). constructor. destruct h6 as [h6]. destruct h6 as [R [h6 [h8 h9]]]. exists (im_proj1_sig R). exists (finite_image _ _ _ _ h6). rewrite (plus_set_sub_compat B C h0 R h6). split. pose proof (f_equal (@proj1_sig _ _) h8) as h10. simpl in h10. assumption. intros s h10. destruct h10 as [s h10]. subst. specialize (h9 _ h10). destruct h9 as [R' [h11 [h12 h13]]]. subst. exists (im_proj1_sig R'). exists (finite_image _ _ _ _ h11). split. red. intros r h13. destruct h13 as [r h13]. subst. specialize (h12 _ h13). destruct h12 as [r h12a | r h12b]. left. apply Im_intro with r; auto. destruct h12b as [r h12b]. subst. right. unfold im_proj1_sig. rewrite im_im. apply Im_intro with r; auto. rewrite (times_set_sub_compat B C h0 R' h11). reflexivity. apply Im_intro with (exist _ _ h8). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. red. intros x h6. destruct h6 as [x h6]. subst. clear h6. destruct x as [x h6]. simpl. destruct h6 as [h6]. destruct h6 as [R [h6 [h7 h8]]]. subst. constructor. assert (h9:Included R C). red. intros r h9. specialize (h8 _ h9). destruct h8 as [R' [h10 [h11 h12]]]. subst. apply times_set_closed; auto. red. intros r h13. specialize (h11 _ h13). destruct h11 as [r h14 | r h15]. destruct h14 as [r h14]. subst. apply proj2_sig. destruct h15 as [r h15]. subst. destruct h15 as [r h15]. subst. apply comp_closed; auto. apply proj2_sig. exists (im_proj2_sig _ h9). assert (h10:Finite (im_proj2_sig R h9)). eapply finite_image. rewrite <- finite_full_sig_iff. assumption. exists h10. split. apply proj1_sig_injective. simpl. pose proof (plus_set_sub_compat B C h0 (im_proj2_sig R h9) h10 (finite_image _ _ _ _ h10)) as h11. unfold D. rewrite <- h11. apply plus_set_functional. unfold im_proj2_sig. rewrite im_im. simpl. rewrite <- im_full_sig_proj1_sig. reflexivity. intros s h11. destruct s as [s h12]. inversion h11 as [s' h13 ? h14]. subst. apply exist_injective in h14. subst. clear h11. destruct s' as [s' h14]. simpl in h12. specialize (h8 _ h14). simpl. destruct h8 as [R' [h15 [h16 h17]]]. subst. assert (h17:Included R' C). red. intros x h18. specialize (h16 _ h18). destruct h16 as [x h16a | x h16b]. destruct h16a as [x h16a]. subst. apply proj2_sig. destruct h16b as [x h16b]. subst. destruct h16b as [x h16b]. subst. apply comp_closed; auto. apply proj2_sig. exists (im_proj2_sig _ h17). assert (h18:Finite (im_proj2_sig R' h17)). eapply finite_image. rewrite <- finite_full_sig_iff. assumption. exists h18. split. red. intros x h19. destruct h19 as [x h19]. subst. clear h19. destruct x as [x h19]. simpl. specialize (h16 _ h19). destruct h16 as [x h16a | x h16b]. destruct h16a as [x h16a]. subst. left. rewrite unfold_sig in h16a. pose proof (proof_irrelevance _ (h17 (proj1_sig x) h19) (proj2_sig x)) as h20. subst. rewrite <- h20 in h16a at 1. assumption. destruct h16b as [x h16b]. subst. destruct h16b as [x h16b]. subst. right. apply Im_intro with x. assumption. apply proj1_sig_injective. simpl. unfold Bcomp_sub. reflexivity. apply proj1_sig_injective. simpl. pose proof (times_set_sub_compat B C h0 (im_proj2_sig R' h17) h18 (finite_image _ _ _ _ h18)) as h19. unfold D. rewrite <- h19. apply times_set_functional. unfold im_proj2_sig. rewrite im_im. simpl. rewrite <- im_full_sig_proj1_sig. reflexivity. rewrite h6 at 1. rewrite <- im_proj1_sig_undoes_im_proj2_sig at 1. constructor. exists S. exists h2. split; auto. Qed. Section NormalForm'. Variable A:Bool_Alg. Let At := Btype (Bc A). Lemma gen_ens_closed_add_plus_set_aux : forall (B:Ensemble At) (r:At), alg_closed B -> forall (D:Ensemble At) (pf:Finite D), Included D [d : At | exists p q : At, Inn B p /\ Inn B q /\ d = p * r + q * - r] -> exists p q : At, Inn B p /\ Inn B q /\ plus_set D pf = p * r + q * - r. intros B r h1 D h2. pose proof (finite_set_list_no_dup _ h2) as h3. destruct h3 as [l h3]. destruct h3 as [h3 h4]. revert h4 h3. revert h2. revert D r h1. revert B. induction l as [|a l h5]. intros B D r h1 h2 h4 h3. simpl in h3. subst. intros. exists 0. exists 0. split. apply zero_closed; auto. split. apply zero_closed; auto. rewrite plus_set_empty'. rewrite comm_prod. rewrite zero_prod. rewrite comm_prod. rewrite zero_prod. rewrite zero_sum. reflexivity. intros B D r h2 h3 h4 h1. pose proof (no_dup_cons _ _ h4) as h7. pose proof (no_dup_cons_nin _ _ h4) as h8. pose proof (subtract_preserves_finite _ a h3) as h9. specialize (h5 B (Subtract D a) r h2 h9 h7). subst. simpl in h5. rewrite list_to_set_in_iff in h8. pose proof (sub_add_same_nin _ _ h8) as h10. specialize (h5 h10). simpl in h9. simpl in h3. simpl. assert (h11:Finite (list_to_set l)). eapply Finite_downward_closed. apply h3. auto with sets. pose proof (subsetT_eq_compat _ _ _ _ h9 h11 h10) as h12. dependent rewrite -> h12 in h5. intros h13. assert (h14:Included (list_to_set l) [d : At | exists p q : At, Inn B p /\ Inn B q /\ d = p * r + q * - r]). auto with sets. specialize (h5 h14). destruct h5 as [p h5]. destruct h5 as [q h5]. destruct h5 as [h5l [h5r h5c]]. specialize (h13 a). assert (h15:Inn (Add (list_to_set l) a) a). right. constructor. specialize (h13 h15). destruct h13 as [h13]. destruct h13 as [p' h13]. destruct h13 as [q' h13]. destruct h13 as [h13a [h13b h13c]]. assert (h16:plus_set (Add (list_to_set l) a) h3 = a + (plus_set (list_to_set l) h11)). pose proof (plus_set_add' (list_to_set l) h11 a h3) as h16. unfold At in h16. unfold At. rewrite h16 at 1. reflexivity. rewrite h13c in h16 at 2. rewrite h5c in h16. rewrite <- assoc_sum in h16. rewrite (assoc_sum _ (q' * -r) (p*r) (q*-r)) in h16. rewrite (comm_sum _ (q' * -r) (p*r)) in h16. rewrite <- (assoc_sum _ (p*r) (q' * -r) (q * -r)) in h16. rewrite <- (dist_sum_r q' q (-r)) in h16. rewrite assoc_sum in h16. rewrite <- dist_sum_r in h16. exists (p' + p). exists (q' + q). split. apply plus_closed; auto. split. apply plus_closed; auto. assumption. Qed. Lemma gen_ens_closed_add : forall B:Ensemble At, alg_closed B -> forall r:At, Gen_Ens (Add B r) = [x:At | exists p q, Inn B p /\ Inn B q /\ x = p*r + q*-r]. intros B h1 r. pose proof (gen_ens_eq _ (Add B r)) as h2. assert (h3':forall (C:Ensemble At) (pf:Finite C), Included C (comp_add_set B) -> Inn B (times_set C pf)). intros C h3 h4. apply times_set_closed; auto. red. intros x h5. specialize (h4 _ h5). destruct h4 as [x h4l | x h4r]. assumption. destruct h4r as [x h4r]. subst. apply comp_closed; auto. assert (heq:Gen_Ens (Add B r) = [x:At | exists D (pf:Finite D), x = plus_set _ pf /\ (Included D (Union B (Union [a:At | exists p, Inn B p /\ a = p * r] [a:At | exists q, Inn B q /\ a = q * -r])))]). rewrite h2. apply Extensionality_Ensembles. red. split. red. intros x h3. constructor. destruct h3 as [h3]. destruct h3 as [S h3]. destruct h3 as [h3 h4]. destruct h4 as [h4 h5]. exists S. exists h3. split; auto. red. intros s h6. specialize (h5 _ h6). destruct h5 as [R h5]. destruct h5 as [h5 h7]. destruct h7 as [h7 h8]. subst. pose proof (comp_add_set_add B r) as h8. unfold At in h8. unfold At in h7. simpl in h8. simpl in h7. rewrite h8 in h7. apply incl_union_setminus in h7. pose proof (decompose_int_setminus R (Couple r (-r))) as h9. assert (h10:Included (Intersection R (Couple r (-r))) R). auto with sets. pose proof (Finite_downward_closed _ _ h5 _ h10) as h11. assert (h12:Included (Setminus R (Couple r (-r))) R). apply setminus_inc. pose proof (Finite_downward_closed _ _ h5 _ h12) as h13. pose proof (times_set_union _ _ h11 h13) as h14. pose proof (subsetT_eq_compat _ _ _ _ h5 (Union_preserves_Finite (Btype (Bc A)) (Intersection R (Couple r (- r))) (Setminus R (Couple r (- r))) h11 h13) h9) as h15. dependent rewrite -> h15. unfold At in h14. unfold At. rewrite h14 at 1. assert (h16:Included (Intersection R (Couple r (-r))) (Couple r (-r))). auto with sets. apply incl_couple_inv in h16. destruct h16 as [h16a | [h16b | [h16c | h16d]]]. pose proof (Empty_is_finite At) as h17. pose proof (subsetT_eq_compat _ _ _ _ h11 h17 h16a) as h18. unfold At in h18. unfold At. dependent rewrite -> h18. rewrite times_set_empty'. rewrite comm_prod. rewrite one_prod. left. apply times_set_closed; auto. apply incl_comp_add_set_closed; auto with sets. pose proof (Singleton_is_finite _ r) as h17. pose proof (subsetT_eq_compat _ _ _ _ h11 h17 h16b) as h18. unfold At in h18. unfold At. dependent rewrite -> h18. rewrite times_set_sing'. rewrite comm_prod. right. left. constructor. exists (times_set (Setminus R (Couple r (- r))) h13); auto. pose proof (Singleton_is_finite _ (-r)) as h17. pose proof (subsetT_eq_compat _ _ _ _ h11 h17 h16c) as h18. unfold At in h18. unfold At. dependent rewrite -> h18. rewrite times_set_sing'. rewrite comm_prod. right. right. constructor. exists (times_set (Setminus R (Couple r (- r))) h13); auto. pose proof (finite_couple r (-r)) as h17. pose proof (subsetT_eq_compat _ _ _ _ h11 h17 h16d) as h18. unfold At in h18. unfold At. dependent rewrite -> h18. rewrite times_set_couple'. rewrite comp_prod. rewrite comm_prod. rewrite zero_prod. left. apply zero_closed; auto. red. intros x h4. destruct h4 as [h4]. destruct h4 as [D h4]. destruct h4 as [h4 h5]. destruct h5 as [h5l h5r]. constructor. exists D. exists h4. split; auto. intros d h6. apply h5r in h6. destruct h6 as [d h6a | d h6b]. exists (Singleton d). exists (Singleton_is_finite _ d). split. red. intros d' h7. destruct h7. left. left. assumption. rewrite times_set_sing'. reflexivity. destruct h6b as [d h6b | d h6c]. destruct h6b as [h6b]. destruct h6b as [p h6b]. destruct h6b as [h7 h8]. exists (Couple p r). exists (finite_couple p r). split. red. intros a h9. destruct h9. left. left. assumption. left. right. constructor. subst. rewrite times_set_couple'. reflexivity. destruct h6c as [h6c]. destruct h6c as [q h6c]. destruct h6c as [h7 h8]. exists (Couple q (-r)). exists (finite_couple q (-r)). split. red. intros a h9. destruct h9. left. left. assumption. right. apply Im_intro with r. right. constructor. reflexivity. subst. rewrite times_set_couple'. reflexivity. assert (h4:Included (Union B (Union [a : At | exists p : At, Inn B p /\ a = p * r] [a : At | exists q : At, Inn B q /\ a = q * - r])) [x : At | exists p q : At, Inn B p /\ Inn B q /\ x = p * r + q * - r]). red. intros x h4. destruct h4 as [x h4a | x h4b]. constructor. exists x. exists x. repeat split; auto. rewrite <- dist_sum. rewrite comp_sum. rewrite one_prod. reflexivity. destruct h4b as [x h4b | x h4c]. destruct h4b as [h4b]. destruct h4b as [a h4b]. destruct h4b as [h5 h6]. constructor. exists a. exists 0. split; auto. split. apply zero_closed; auto. rewrite (comm_prod _ 0 (-r)). rewrite zero_prod. rewrite zero_sum. assumption. destruct h4c as [h4c]. destruct h4c as [a h4c]. destruct h4c as [h5 h6]. constructor. exists 0. exists a. repeat split; auto. apply zero_closed; auto. rewrite (comm_prod _ 0 r). rewrite zero_prod. rewrite comm_sum. rewrite zero_sum. assumption. assert (h5:Included (Gen_Ens (Add B r)) [x : At | exists (D : Ensemble (Btype (Bc A))) (pf : Finite D), x = plus_set D pf /\ Included D [d:At | exists p q : At, Inn B p /\ Inn B q /\ d = p * r + q * - r]]). rewrite heq. red. intros x h5. destruct h5 as [h5]. destruct h5 as [D h5]. destruct h5 as [h5 h6]. destruct h6 as [h6a h6b]. constructor. exists D. exists h5. split; auto. intros d h7. assert (h8:Inn [x : At | exists p q : At, Inn B p /\ Inn B q /\ x = p * r + q * - r] d). auto with sets. destruct h8 as [h8]. constructor. assumption. apply Extensionality_Ensembles. red. split. Focus 2. red. intros x h6. destruct h6 as [h6]. destruct h6 as [p h6]. destruct h6 as [q h6]. destruct h6 as [h6a [h6b h6c]]. subst. apply plus_closed. apply closed_gen_ens. apply times_closed. apply closed_gen_ens. apply gen_ens_includes. left. assumption. apply gen_ens_includes. right. constructor. apply times_closed. apply closed_gen_ens. apply gen_ens_includes. left. assumption. apply comp_closed. apply closed_gen_ens. apply gen_ens_includes. right. constructor. red. intros x h6. apply h5 in h6. destruct h6 as [h6]. destruct h6 as [D h6]. destruct h6 as [h6 h7]. destruct h7 as [h7 h8]. constructor. subst. apply gen_ens_closed_add_plus_set_aux; auto. Qed. Lemma in_gen_ens_iff : forall (E:Ensemble At) (x:At), Inn (Gen_Ens E) x <-> exists (F:Ensemble At), Finite F /\ Included F (comp_add_set E) /\ Inn (Gen_Ens F) x. intros E x. split. intros h1. rewrite gen_ens_eq in h1. destruct h1 as [h1]. destruct h1 as [S [h1 [h2 h3]]]. subst. revert h3. revert E. pose proof (finite_set_list_no_dup _ h1) as h2. revert h1. destruct h2 as [l h2]. destruct h2 as [h2 h3]. subst. induction l as [|a l h4]. intros. simpl. rewrite plus_set_empty'. exists (Empty_set _). split. apply Empty_is_finite. split; auto with sets. rewrite gen_ens_empty. left. simpl. intros h1 E h2. pose proof (no_dup_cons_nin _ _ h3) as h5. pose proof (no_dup_cons _ _ h3) as h6. assert (h7:Finite (list_to_set l)). eapply Finite_downward_closed. apply h1. auto with sets. assert (h8: forall s : Btype (Bc A), Inn (list_to_set l) s -> exists (R : Ensemble (Btype (Bc A))) (pfr : Finite R), Included R (comp_add_set E) /\ s = times_set R pfr). intros s h9. specialize (h2 s (Add_intro1 _ _ _ _ h9)). destruct h2 as [R [h2a [h2b]]]. subst. exists R. exists h2a. split; auto. specialize (h4 h6 h7 E h8). destruct h4 as [F [h4a [h4b h4c]]]. pose proof (h2 a (Add_intro2 _ _ a)) as h9. destruct h9 as [R [h9 [h10 h11]]]. subst. exists (Union R F). split; auto. apply Union_preserves_Finite; auto. split. auto with sets. pose proof (closed_gen_ens _ (Union R F)) as h11. assert (h12:Finite (list_to_set l)). apply list_to_set_finite. assert (h14: Inn (Gen_Ens (Union R F)) (times_set R h9)). rewrite gen_ens_union. apply gen_ens_includes. left. apply times_set_closed. apply closed_gen_ens. apply gen_ens_includes. assert (h13:Inn (Gen_Ens (Union R F)) (plus_set (list_to_set l) h12)). rewrite gen_ens_union. apply gen_ens_includes; auto. right. pose proof (proof_irrelevance _ h7 h12); subst; auto. pose proof (plus_closed _ _ h11 (times_set R h9) (plus_set (list_to_set l) h12) h14 h13) as h15. erewrite plus_set_add'. apply h15. intro h1. destruct h1 as [F [h1 [h2 h3]]]. rewrite <- gen_ens_comp_add_set. pose proof (gen_ens_preserves_inclusion _ _ h2) as h4. auto with sets. Qed. End NormalForm'. Arguments eps_map [A] [E] _ _. Arguments eps_map_compose [A] [B] [E] _ _ _. Arguments fin_map_seg_to_firstn [A] [m] [l] _ _. Arguments el_prod_l [A] [l] _. Arguments el_prod_l_firstn [A] [l] [m] _ _. Section CompleteAndRegularSubalgebras. Variable B:Bool_Alg. Let Bt := Btype (Bc B). Inductive complete_subalg (A:Ensemble Bt) : Prop := | complete_subalg_intro : alg_closed A -> forall (pf:set_complete B), (forall (C:Ensemble Bt), Included C A -> Inn A (sup_set_complete C pf)) -> complete_subalg A. Lemma complete_subalg_complete_supalg : forall (A:Ensemble Bt), complete_subalg A -> set_complete B. intros A h1. destruct h1; auto. Qed. Lemma complete_subalg_cont_inf : forall (A:Ensemble Bt) (pf:complete_subalg A), forall (C:Ensemble Bt) (p:Bt), Included C A -> Inn A (inf_set_complete C (complete_subalg_complete_supalg _ pf)). intros A h1 C p h2. destruct h1 as [h3 h4 h5]. pose proof (incl_comp_set_closed _ _ h3 h2) as h6. pose proof (h5 _ h6) as h7. pose proof (sup_set_complete_compat (comp_set C) h4) as h8. rewrite doub_neg in h8. rewrite <-inf_sup_compat_iff in h8. unfold inf_set_complete. destruct Description.constructive_definite_description as [q h10]. simpl. pose proof (inf_unq _ _ _ h8 h10). subst. apply comp_closed; auto. Qed. Inductive regular_subalg (A:Ensemble Bt) : Prop := | regular_subalg_intro : (forall (pfc:alg_closed A) (C:Ensemble Bt) (pfi:Included C A), (forall (p:Bt) (pfp:Inn A p), sup (B:=Subalg _ pfc) (subset_sig _ _ pfi) (exist _ _ pfp) -> sup C p)) -> regular_subalg A. Lemma regular_subalg_iff : forall (A:Ensemble Bt) (pfc:alg_closed A), regular_subalg A <-> (forall (E:Ensemble Bt) (pfi:Included E A), sup (B:=Subalg _ pfc) (subset_sig _ _ pfi) (exist _ _ (one_closed _ _ pfc)) -> sup E 1). intros A h1. split. intros h2. destruct h2 as [h2]. intros E h5 h6. specialize (h2 h1 _ h5 1 (one_closed B A h1) h6). assumption. intros h2. constructor. intros hs E0 h3 p h4 h5. pose (Subalg _ hs) as A'. assert (h6:Included (Add E0 (-p)) A). red. intros x h6. destruct h6 as [x h6l | x h6r]. auto with sets. destruct h6r. apply comp_closed; auto. assert (h7:sup (B:=A') (subset_sig _ _ h6) (exist _ _ (one_closed _ _ h1))). pose proof (comp_sum A' (exist _ _ h4)) as h7. pose proof (@sup_add_assoc A' (subset_sig _ _ h3) ((Bcomp (Bc A') (exist _ _ h4))) _ h5) as h8. rewrite h7 in h8. simpl in h8. assert (h9: (Add (subset_sig E0 A h3) (exist (Inn A) (Bcomp_sub B A (exist (Inn A) p h4)) (C_c B A hs (exist (Inn A) p h4)))) = subset_sig (Add E0 (-p)) A h6). apply Extensionality_Ensembles. red. split. red. intros x h9. destruct h9 as [x h9l | x h9r]. rewrite subset_sig_compat. rewrite subset_sig_compat in h9l. left. assumption. destruct h9r. rewrite subset_sig_compat. simpl. right. constructor. red. intros x h9. rewrite subset_sig_compat in h9. inversion h9 as [x' h10 | x' h11]. subst. left. rewrite subset_sig_compat. assumption. subst. inversion h11 as [h12]. right. destruct x as [x h13]. simpl in h12. subst. simpl. unfold Bcomp_sub. simpl. pose proof (proof_irrelevance _ h13 (C_c B A hs (exist (Inn A) p h4))); subst. constructor. rewrite <- h9. pose proof (proof_irrelevance _ (one_closed B A h1) (O_c B A hs)) as h10. rewrite h10. assumption. specialize (h2 _ h6). unfold A' in h2. unfold A' in h7. pose proof (proof_irrelevance _ hs h1); subst. specialize (h2 h7). pose proof (@set_infnt_distr_1_sup B _ p _ h2) as h8. rewrite Im_add in h8. rewrite comp_prod in h8. rewrite one_prod in h8. assert (h9: Im E0 (fun x : Btype (Bc B) => p * x) = E0). apply Extensionality_Ensembles. red. split. red. intros x h9. destruct h9 as [x h9]. subst. pose proof (le_sup _ _ _ h5 (exist _ _ (h3 _ h9))) as h10. assert (h11:Inn (subset_sig E0 A h3) (exist (Inn A) x (h3 _ h9))). rewrite subset_sig_compat. simpl. assumption. specialize (h10 h11). rewrite subset_sig_compat in h11. simpl in h11. red in h10. rewrite eq_ord in h10. simpl in h10. unfold Btimes_sub in h10. simpl in h10. apply exist_injective in h10. rewrite comm_prod. rewrite h10. assumption. red. intros x h9. pose proof (le_sup _ _ _ h5 (exist _ _ (h3 _ h9))) as h10. assert (h11: Inn (subset_sig E0 A h3) (exist (Inn A) x (h3 x h9))). rewrite subset_sig_compat. simpl. assumption. specialize (h10 h11). red in h10. rewrite eq_ord in h10. simpl in h10. unfold Btimes_sub in h10. simpl in h10. apply exist_injective in h10. rewrite <- h10. rewrite comm_prod. apply Im_intro with x; auto. rewrite h9 in h8. apply sup_subtract_zero in h8. destruct (classic (Inn E0 0)) as [h9' | h10']. rewrite sub_add_same_in in h8. rewrite sup_subtract_zero. assumption. assumption. rewrite sub_add_same_nin in h8; auto. Qed. End CompleteAndRegularSubalgebras. Section Examples. Variable fos:Field_of_Sets. Definition psa_fos := psa (Xt fos). Let Bt := Btype (Bc psa_fos). Lemma fos_closed : @alg_closed psa_fos (F fos). constructor. red. intros A B. destruct A. destruct B. unfold Bplus_sub. simpl. apply (Union_closed fos); assumption. red. intros A B. destruct A. destruct B. unfold Btimes_sub. simpl. apply (Int_closed fos); assumption. red. simpl. apply full_in_F. red. simpl. apply empty_in_F. red. intro A. destruct A as [? h1]. unfold Bcomp_sub. simpl. apply (Comp_closed fos). assumption. Qed. Definition fos_sub_psa := @Subalg psa_fos (F fos) fos_closed. Variable Xt:Type. Definition fin_cof_fos' := fin_cof_fos Xt. Lemma fin_cof_closed : @alg_closed (psa Xt) (F_fc Xt). constructor. red. intros A B. destruct A. destruct B. unfold Bplus_sub. simpl. apply (Union_closed_fc); assumption. red. intros A B. destruct A. destruct B. unfold Btimes_sub. simpl. apply (Int_closed_fc); assumption. red. simpl. apply full_in_F_fc. red. simpl. apply empty_in_F_fc. red. intro A. destruct A. unfold Bcomp_sub. simpl. apply (Comp_closed_fc); assumption. Qed. Definition fin_cof_sub_psa := @Subalg (psa Xt) (F_fc Xt) fin_cof_closed. (* ". . . consider the field A of finite and cofinite subsets of a set X. We shall prove that it is generated in P(X) by the set E of one- element (that is, singleton) subsets of X." *) Lemma fin_cof_gen_singletons' : @Gen (psa Xt) (SingF (Full_set Xt)) = fin_cof_sub_psa. unfold fin_cof_sub_psa. unfold Gen. pose (@SA_cont_A (psa Xt) (SingF (Full_set Xt))) as F. unfold SingF in F. pose (int_saf_A F) as S. simpl. pose fin_cof_closed as h1. pose proof (P_c _ _ h1) as h2. assert (h3:Included (SingF (Full_set Xt)) (F_fc Xt)). red. intros A h4. inversion h4 as [x ? ? h5]. rewrite h5. pose proof (Singleton_is_finite _ x) as h6. constructor. left; assumption. assert (h4:int_saf_A (@SA_cont_A (psa Xt) (SingF (Full_set Xt))) = (F_fc Xt)). apply Extensionality_Ensembles. red. split. red. intros sing h5. inversion h5 as [? h6]. apply h6. constructor. split. red. intros sing'. unfold Bt in sing'. simpl in sing'. intro h7. inversion h7 as [y ? ? h8]. rewrite h8. constructor. left. apply Singleton_is_finite. assumption. red. intros S' h4. pose proof (union_singF_eq S') as h5. inversion h4 as [h6]. assert (hfin : forall fin:(Btype (Bc (psa Xt))), Inn (F_fc Xt) fin -> Finite fin -> Inn S fin). intros fin h4' h6'. pose proof (union_singF_eq fin) as h5'. induction h6' as [|fin' h7 h8 x h9]. constructor. intros E h6''. inversion h6'' as [h7]. destruct h7 as [? h8]. pose proof (Z_c _ _ h8) as h9. red in h9. simpl in h9. assumption. unfold SingF in h5'. pose proof (Im_add _ _ fin' x (Singleton (U:=Xt))) as h10. unfold SingF. rewrite h10 in h5'. pose proof (family_union_add (Im fin' (Singleton (U:=Xt))) (Singleton x)) as h11. rewrite h11 in h5'. rewrite h5'. pose proof (union_singF_eq fin') as h12. unfold SingF in h12. assert (h13:Inn (F_fc Xt) fin'). constructor. left; assumption. pose proof (h8 h13 h12) as h14. rewrite <- h12. pose (Im (Full_set Xt) (Singleton (U:=Xt))) as J. assert (h15:Inn (int_saf_A (@SA_cont_A (psa Xt) J)) (Singleton x)). assert (h16:Inn J (Singleton x)). unfold J. apply Im_intro with x. apply Full_intro. trivial. unfold SA_cont_A in F. pose (saf_A F) as EF. unfold Bt in EF. simpl in EF. unfold int_saf_A. assert (h17:forall (G:Family Xt), Inn EF G -> Included J G). intros G h18. inversion h18 as [h19]. destruct h19 as [h19l]. unfold J. assumption. constructor. unfold Bt. simpl. intros G h18. unfold EF in h17. pose proof (h17 G h18) as h19. auto with sets. pose proof (saf_closed F) as h20. constructor. intros S0 h21. pose proof (h20 S0 h21) as h22. pose proof (P_c _ _ h22) as h23. red in h23. unfold SubBtype in h23. unfold Bplus_sub in h23. unfold Bt in h23. simpl in h23. assert (h24: Inn S0 (Singleton x)). inversion h15 as [? h16]. apply (h16 S0 h21). assert (h25: Inn S0 fin'). inversion h14. apply (H S0 h21). apply (h23 (exist _ (Singleton x) h24) (exist _ fin' h25)). destruct h6 as [h6l | h6r]. apply hfin; assumption. assert (h7:Inn (F_fc Xt) (Comp S')). constructor. left; assumption. pose proof (hfin _ h7 h6r) as h8. pose proof (int_saf_alg_closed F) as h9. pose proof (C_c _ _ h9) as h10. red in h10. unfold SubBtype in h10. unfold Bcomp_sub in h10. unfold Bt in h10. simpl in h10. pose proof (h10 (exist _ (Comp S') h8)) as h11. simpl in h11. pose proof (Complement_Complement _ S') as h12. rewrite <- h12. assumption. unfold Gen_Ens. unfold SA_cont_A in h4. unfold int_saf_A in h4. simpl in h4. apply subalg_functional. assumption. Qed. End Examples. (*Bool_alg_p analogues*) Section ParametricAnalogues. Section AlgClosed_p. Variable T:Type. Variable Bp:Bool_Alg_p T. Let Btp := Btype_p T (Bc_p T Bp). Variable A_p':Ensemble T. Hypothesis incl_subalg : Included A_p' (A_p T (Bc_p T Bp)). Definition SubBtype_p := sig_set A_p'. (*This converts A_p' (an [Ensemble T]) to an [Ensemble Btype (Bc (ba_conv Bp))].*) Definition ba_conv_und_subalg := (ba_conv_set (Im (full_sig A_p') (fun x=>exist _ (proj1_sig x) (incl_subalg _ (proj2_sig x))))). Definition SubBtype_p_conv (x:SubBtype_p) : SubBtype _ ba_conv_und_subalg. unfold SubBtype, ba_conv. simpl. unfold SubBtype_p in x. pose (exist _ (proj1_sig x) (incl_subalg _ (proj2_sig x))) as x'. assert (h1:Inn ba_conv_und_subalg x'). apply Im_intro with x. constructor. reflexivity. refine (exist _ _ h1). Defined. Lemma SubBtype_p_conv_inj : forall (x y:SubBtype_p), SubBtype_p_conv x = SubBtype_p_conv y -> x = y. intros x y. destruct x as [x h1]. destruct y as [y h2]. unfold SubBtype_p_conv. intro h3. apply exist_injective in h3. apply exist_injective in h3. simpl in h3. apply proj1_sig_injective. simpl. assumption. Qed. Definition Bplus_sub_p (x y:SubBtype_p) : Btp. destruct x as [x h1]. destruct y as [y h2]. pose proof (incl_subalg _ h1) as h3. pose proof (incl_subalg _ h2) as h4. pose (exist _ _ h3) as x'. pose (exist _ _ h4) as y'. refine (x' %+ y'). Defined. Lemma Bplus_sub_p_eq : forall (x y:SubBtype_p), Bplus_sub_p x y = Bplus_sub _ _ (SubBtype_p_conv x) (SubBtype_p_conv y). intros x y. destruct x as [x h1]. destruct y as [y h2]. unfold Bplus_sub_p, Bplus_sub. simpl. reflexivity. Qed. Definition Btimes_sub_p (x y:SubBtype_p) : Btp. destruct x as [x h1]. destruct y as [y h2]. pose proof (incl_subalg _ h1) as h3. pose proof (incl_subalg _ h2) as h4. pose (exist _ _ h3) as x'. pose (exist _ _ h4) as y'. refine (x' %* y'). Defined. Lemma Btimes_sub_p_eq : forall (x y:SubBtype_p), Btimes_sub_p x y = Btimes_sub _ _ (SubBtype_p_conv x) (SubBtype_p_conv y). intros x y. destruct x as [x h1]. destruct y as [y h2]. unfold Btimes_sub_p, Btimes_sub. simpl. reflexivity. Qed. Definition Bcomp_sub_p (x:SubBtype_p) : Btp. destruct x as [x h1]. pose proof (incl_subalg _ h1) as h3. pose (exist _ _ h3) as x'. refine (%-x'). Defined. Lemma Bcomp_sub_p_eq : forall (x:SubBtype_p), Bcomp_sub_p x = Bcomp_sub _ _ (SubBtype_p_conv x). intro x. destruct x as [x h1]. unfold Bcomp_sub_p, Bcomp_sub. simpl. reflexivity. Qed. Notation "x %+|' y" := (Bplus_sub_p x y) (at level 50, left associativity). Notation "x %*|' y" := (Btimes_sub_p x y) (at level 40, left associativity). Notation "%-|' x" := (Bcomp_sub_p x) (at level 30). Definition Plus_closed_sub_p : Prop := (forall (x y:SubBtype_p), Inn A_p' (proj1_sig (x %+|' y))). Lemma Plus_closed_sub_p_iff : Plus_closed_sub_p <-> Plus_closed_sub ba_conv_und_subalg. split. intro h1. red. intros x y. unfold ba_conv_set. unfold ba_conv_type. unfold transfer_dep. unfold eq_rect_r. simpl. red in h1. destruct x as [x h2]. destruct y as [y h3]. destruct h2 as [x h2]. destruct h3 as [y h3]. subst. specialize (h1 x y). simpl. apply Im_intro with (exist _ _ h1). constructor. simpl. unfold Bplus_sub. simpl. apply proj1_sig_injective. simpl. destruct x as [x h4]. destruct y as [y h5]. simpl. reflexivity. intro h1. red. intros x y. red in h1. unfold SubBtype in h1. destruct x as [x h2]. destruct y as [y h3]. simpl. pose (exist _ _ (incl_subalg _ h2)) as x'. pose (exist _ _ (incl_subalg _ h3)) as y'. assert (h4:Inn (ba_conv_set (Im (full_sig A_p') (fun x : sig_set A_p' => exist (Inn (A_p T (Bc_p T Bp))) (proj1_sig x) (incl_subalg (proj1_sig x) (proj2_sig x))))) x'). apply Im_intro with (exist _ _ h2). constructor. simpl. unfold x'. reflexivity. assert (h5:Inn (ba_conv_set (Im (full_sig A_p') (fun x : sig_set A_p' => exist (Inn (A_p T (Bc_p T Bp))) (proj1_sig x) (incl_subalg (proj1_sig x) (proj2_sig x))))) y'). apply Im_intro with (exist _ _ h3). constructor. simpl. unfold y'. reflexivity. specialize (h1 (exist _ _ h4) (exist _ _ h5)). inversion h1 as [z h7 h8 h9]. subst. unfold Bplus_sub in h9. simpl in h9. unfold x', y' in h9. rewrite h9. simpl. apply proj2_sig. Qed. Definition Times_closed_sub_p : Prop := (forall (x y:SubBtype_p), Inn A_p' (proj1_sig (x %*|' y))). Lemma Times_closed_sub_p_iff : Times_closed_sub_p <-> Times_closed_sub ba_conv_und_subalg. split. intro h1. red. intros x y. unfold ba_conv_set. unfold ba_conv_type. unfold transfer_dep. unfold eq_rect_r. simpl. red in h1. destruct x as [x h2]. destruct y as [y h3]. destruct h2 as [x h2]. destruct h3 as [y h3]. subst. specialize (h1 x y). simpl. apply Im_intro with (exist _ _ h1). constructor. simpl. unfold Bplus_sub. simpl. apply proj1_sig_injective. simpl. destruct x as [x h4]. destruct y as [y h5]. simpl. reflexivity. intro h1. red. intros x y. red in h1. unfold SubBtype in h1. destruct x as [x h2]. destruct y as [y h3]. simpl. pose (exist _ _ (incl_subalg _ h2)) as x'. pose (exist _ _ (incl_subalg _ h3)) as y'. assert (h4:Inn (ba_conv_set (Im (full_sig A_p') (fun x : sig_set A_p' => exist (Inn (A_p T (Bc_p T Bp))) (proj1_sig x) (incl_subalg (proj1_sig x) (proj2_sig x))))) x'). apply Im_intro with (exist _ _ h2). constructor. simpl. unfold x'. reflexivity. assert (h5:Inn (ba_conv_set (Im (full_sig A_p') (fun x : sig_set A_p' => exist (Inn (A_p T (Bc_p T Bp))) (proj1_sig x) (incl_subalg (proj1_sig x) (proj2_sig x))))) y'). apply Im_intro with (exist _ _ h3). constructor. simpl. unfold y'. reflexivity. specialize (h1 (exist _ _ h4) (exist _ _ h5)). inversion h1 as [z h7 h8 h9]. subst. unfold Btimes_sub in h9. simpl in h9. unfold x', y' in h9. rewrite h9. simpl. apply proj2_sig. Qed. Definition One_closed_sub_p : Prop := Inn (A_p') (proj1_sig (Bone_p T (Bc_p T Bp))). Lemma One_closed_sub_p_iff : One_closed_sub_p <-> One_closed_sub ba_conv_und_subalg. split. intro h1. red in h1. pose (exist _ _ h1) as one. apply Im_intro with one. constructor. unfold one. simpl. apply proj1_sig_injective. simpl. reflexivity. intro h1. red in h1. inversion h1 as [one h2 ? h4]. subst. red. rewrite ba_conv_one. rewrite h4. simpl. apply proj2_sig. Qed. Definition Zero_closed_sub_p : Prop := Inn (A_p') (proj1_sig (Bzero_p T (Bc_p T Bp))). Lemma Zero_closed_sub_p_iff : Zero_closed_sub_p <-> Zero_closed_sub ba_conv_und_subalg. split. intro h1. red in h1. pose (exist _ _ h1) as zero. apply Im_intro with zero. constructor. unfold zero. simpl. apply proj1_sig_injective. simpl. reflexivity. intro h1. red in h1. inversion h1 as [zero h2 ? h4]. subst. red. rewrite ba_conv_zero. rewrite h4. simpl. apply proj2_sig. Qed. Definition Comp_closed_sub_p : Prop := (forall x:SubBtype_p, Inn A_p' (proj1_sig (%-|' x))). Lemma Comp_closed_sub_p_iff : Comp_closed_sub_p <-> Comp_closed_sub ba_conv_und_subalg. split. intro h1. red. intros x. unfold ba_conv_set. unfold ba_conv_type. unfold transfer_dep. unfold eq_rect_r. simpl. red in h1. destruct x as [x h2]. destruct h2 as [x h2]. subst. specialize (h1 x). simpl. apply Im_intro with (exist _ _ h1). constructor. simpl. unfold Bcomp_sub. simpl. apply proj1_sig_injective. simpl. destruct x as [x h4]. simpl. reflexivity. intro h1. red. intros x. red in h1. unfold SubBtype in h1. destruct x as [x h2]. simpl. pose (exist _ _ (incl_subalg _ h2)) as x'. assert (h4:Inn (ba_conv_set (Im (full_sig A_p') (fun x : sig_set A_p' => exist (Inn (A_p T (Bc_p T Bp))) (proj1_sig x) (incl_subalg (proj1_sig x) (proj2_sig x))))) x'). apply Im_intro with (exist _ _ h2). constructor. simpl. unfold x'. reflexivity. specialize (h1 (exist _ _ h4)). inversion h1 as [z h7 h8 h9]. subst. unfold Bcomp_sub in h9. simpl in h9. unfold x' in h9. rewrite h9. simpl. apply proj2_sig. Qed. Definition Plus_closed_p'' (x y:Btp) : Plus_closed_sub_p -> Inn ba_conv_und_subalg x -> Inn ba_conv_und_subalg y -> Inn ba_conv_und_subalg (x%+y). rewrite Plus_closed_sub_p_iff. apply (@Plus_closed'' (ba_conv Bp)). Qed. Definition Times_closed_p'' (x y:Btp) : Times_closed_sub_p -> Inn ba_conv_und_subalg x -> Inn ba_conv_und_subalg y -> Inn ba_conv_und_subalg (x%*y). rewrite Times_closed_sub_p_iff. apply (@Times_closed'' (ba_conv Bp)). Qed. Definition Comp_closed_p'' (b:Btp) : Comp_closed_sub_p -> Inn ba_conv_und_subalg b -> Inn ba_conv_und_subalg (%-b). rewrite Comp_closed_sub_p_iff. apply (@Comp_closed'' (ba_conv Bp)). Qed. Record alg_closed_p : Prop := { P_c_p : Plus_closed_sub_p; T_c_p : Times_closed_sub_p; O_c_p : One_closed_sub_p; Z_c_p : Zero_closed_sub_p; C_c_p : Comp_closed_sub_p}. Lemma alg_closed_p_iff : alg_closed_p <-> alg_closed ba_conv_und_subalg. split. intro h1. destruct h1 as [h1a h1b h1c h1d h1e]. rewrite Plus_closed_sub_p_iff in h1a. rewrite Times_closed_sub_p_iff in h1b. rewrite One_closed_sub_p_iff in h1c. rewrite Zero_closed_sub_p_iff in h1d. rewrite Comp_closed_sub_p_iff in h1e. constructor; auto. intro h1. destruct h1 as [h1a h1b h1c h1d h1e]. rewrite <- Plus_closed_sub_p_iff in h1a. rewrite <- Times_closed_sub_p_iff in h1b. rewrite <- One_closed_sub_p_iff in h1c. rewrite <- Zero_closed_sub_p_iff in h1d. rewrite <- Comp_closed_sub_p_iff in h1e. constructor; auto. Qed. Variable Acp : alg_closed_p. Definition Bc_p' := Build_Bconst_p T A_p' (Full_set SubBtype_p) (fun (x y:SubBtype_p) => exist (Inn A_p') (proj1_sig (x %+|' y)) (P_c_p Acp x y)) (fun (x y:SubBtype_p) => exist (Inn A_p') (proj1_sig (x %*|' y)) (T_c_p Acp x y)) (exist (Inn A_p') (proj1_sig %1) (O_c_p Acp)) (exist (Inn A_p') (proj1_sig %0) (Z_c_p Acp)) (fun (x:SubBtype_p) => (exist (Inn A_p') (proj1_sig (%-|' x)) (C_c_p Acp x))). Definition Bc_p_conv' := Bc' _ _ (iff1 alg_closed_p_iff Acp). Section Bc_p''. Infix "%+|" := (Bplus_p T Bc_p') (at level 50, left associativity). Infix "%*|" := (Btimes_p T Bc_p') (at level 40, left associativity). Notation "%0|" := (Bzero_p T Bc_p'). Notation "%1|" := (Bone_p T Bc_p'). Notation "%-| x" := ((Bcomp_p T Bc_p') x) (at level 30). Lemma bplus_p_bc_p_eq' : forall (x y:SubBtype_p), SubBtype_p_conv (x %+| y) = Bplus Bc_p_conv' (SubBtype_p_conv x) (SubBtype_p_conv y). intros x y. unfold SubBtype_p_conv. simpl. apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. unfold Bplus_sub. simpl. simpl. unfold Bplus_sub_p. destruct x as [x h1]. destruct y as [y h2]. simpl. reflexivity. Qed. Lemma btimes_p_bc_p_eq' : forall (x y:SubBtype_p), SubBtype_p_conv (x %*| y) = Btimes Bc_p_conv' (SubBtype_p_conv x) (SubBtype_p_conv y). intros x y. unfold SubBtype_p_conv. simpl. apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. unfold Btimes_sub. simpl. simpl. unfold Btimes_sub_p. destruct x as [x h1]. destruct y as [y h2]. simpl. reflexivity. Qed. Lemma bzero_p_bc_p_eq' : SubBtype_p_conv %0| = Bzero Bc_p_conv'. unfold SubBtype_p_conv. apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma bone_p_bc_p_eq' : SubBtype_p_conv %1| = Bone Bc_p_conv'. unfold SubBtype_p_conv. apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma bcomp_p_bc_p_eq' : forall (x:SubBtype_p), SubBtype_p_conv (%-| x) = Bcomp Bc_p_conv' (SubBtype_p_conv x). intro x. unfold SubBtype_p_conv. simpl. apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. unfold Bcomp_sub. simpl. simpl. unfold Bcomp_sub_p. destruct x as [x h1]. simpl. reflexivity. Qed. Lemma assoc_sum_p' : forall n m p:SubBtype_p, n %+| (m %+| p) = n %+| m %+| p. intros n m p. pose proof (assoc_sum' _ ba_conv_und_subalg (iff1 alg_closed_p_iff Acp) (SubBtype_p_conv n) (SubBtype_p_conv m) (SubBtype_p_conv p)) as h1. apply SubBtype_p_conv_inj. do 2 rewrite bplus_p_bc_p_eq'. rewrite h1 at 1. do 2 rewrite bplus_p_bc_p_eq'. reflexivity. Qed. Lemma assoc_prod_p' : forall n m p:SubBtype_p, n %*| (m %*| p) = n %*| m %*| p. intros n m p. pose proof (assoc_prod' _ ba_conv_und_subalg (iff1 alg_closed_p_iff Acp) (SubBtype_p_conv n) (SubBtype_p_conv m) (SubBtype_p_conv p)) as h1. apply SubBtype_p_conv_inj. do 2 rewrite btimes_p_bc_p_eq'. rewrite h1 at 1. do 2 rewrite btimes_p_bc_p_eq'. reflexivity. Qed. Lemma comm_sum_p' : forall n m:SubBtype_p, n %+| m = m %+| n. intros n m. pose proof (comm_sum' _ ba_conv_und_subalg (iff1 alg_closed_p_iff Acp) (SubBtype_p_conv n) (SubBtype_p_conv m)) as h1. apply SubBtype_p_conv_inj. rewrite bplus_p_bc_p_eq'. rewrite h1 at 1. rewrite bplus_p_bc_p_eq'. reflexivity. Qed. Lemma comm_prod_p' : forall n m:SubBtype_p, n %*| m = m %*| n. intros n m. pose proof (comm_prod' _ ba_conv_und_subalg (iff1 alg_closed_p_iff Acp) (SubBtype_p_conv n) (SubBtype_p_conv m)) as h1. apply SubBtype_p_conv_inj. rewrite btimes_p_bc_p_eq'. rewrite h1 at 1. rewrite btimes_p_bc_p_eq'. reflexivity. Qed. Lemma abs_sum_p' : forall n m:SubBtype_p, n %+| (n %*| m) = n. intros n m. pose proof (abs_sum' _ ba_conv_und_subalg (iff1 alg_closed_p_iff Acp) (SubBtype_p_conv n) (SubBtype_p_conv m)) as h1. apply SubBtype_p_conv_inj. rewrite bplus_p_bc_p_eq'. rewrite btimes_p_bc_p_eq'. rewrite h1 at 1. reflexivity. Qed. Lemma abs_prod_p': forall n m:SubBtype_p, n %*| (n %+| m) = n. intros n m. pose proof (abs_prod' _ ba_conv_und_subalg (iff1 alg_closed_p_iff Acp) (SubBtype_p_conv n) (SubBtype_p_conv m)) as h1. apply SubBtype_p_conv_inj. rewrite btimes_p_bc_p_eq'. rewrite bplus_p_bc_p_eq'. rewrite h1 at 1. reflexivity. Qed. Lemma dist_sum_p' : forall n m p:SubBtype_p, p %*| (n %+| m) = p %*| n %+ p %*| m. intros n m p. pose proof (dist_sum' _ ba_conv_und_subalg (iff1 alg_closed_p_iff Acp) (SubBtype_p_conv n) (SubBtype_p_conv m) (SubBtype_p_conv p)) as h1. apply SubBtype_p_conv_inj. rewrite btimes_p_bc_p_eq'. rewrite bplus_p_bc_p_eq'. rewrite h1 at 1. rewrite bplus_p_bc_p_eq'. rewrite btimes_p_bc_p_eq'. rewrite btimes_p_bc_p_eq'. reflexivity. Qed. Lemma dist_prod_p': forall n m p:SubBtype_p, p %+| (n %*| m) = (p %+| n) %*| (p %+| m). intros n m p. pose proof (dist_prod' _ ba_conv_und_subalg (iff1 alg_closed_p_iff Acp) (SubBtype_p_conv n) (SubBtype_p_conv m) (SubBtype_p_conv p)) as h1. apply SubBtype_p_conv_inj. rewrite bplus_p_bc_p_eq'. rewrite btimes_p_bc_p_eq'. rewrite h1 at 1. rewrite btimes_p_bc_p_eq'. rewrite bplus_p_bc_p_eq'. rewrite bplus_p_bc_p_eq'. reflexivity. Qed. Lemma comp_sum_p': forall n:SubBtype_p, n %+| (%-| n) = %1. intro n. pose proof (comp_sum' _ ba_conv_und_subalg (iff1 alg_closed_p_iff Acp) (SubBtype_p_conv n)) as h1. apply SubBtype_p_conv_inj. rewrite bplus_p_bc_p_eq'. rewrite bcomp_p_bc_p_eq'. rewrite h1 at 1. rewrite bone_p_bc_p_eq'. reflexivity. Qed. Lemma comp_prod_p': forall n:SubBtype_p, n %*| (%-| n) = %0. intro n. pose proof (comp_prod' _ ba_conv_und_subalg (iff1 alg_closed_p_iff Acp) (SubBtype_p_conv n)) as h1. apply SubBtype_p_conv_inj. rewrite btimes_p_bc_p_eq'. rewrite bcomp_p_bc_p_eq'. rewrite h1 at 1. rewrite bzero_p_bc_p_eq'. reflexivity. Qed. Lemma und_set_p' : (BS_p T Bc_p') = Full_set (Btype_p T Bc_p'). simpl. unfold SubBtype_p. unfold Btype_p. simpl. reflexivity. Qed. Definition Subalg_p := Build_Bool_Alg_p T Bc_p' und_set_p' assoc_sum_p' assoc_prod_p' comm_sum_p' comm_prod_p' abs_sum_p' abs_prod_p' dist_sum_p' dist_prod_p' comp_sum_p' comp_prod_p'. Lemma two_ops_imp_times_p : Comp_closed_sub_p -> Plus_closed_sub_p -> Times_closed_sub_p. intros h1 h2. rewrite Comp_closed_sub_p_iff in h1. rewrite Plus_closed_sub_p_iff in h2. pose proof (two_ops_imp_times _ _ h1 h2) as h3. rewrite <- Times_closed_sub_p_iff in h3. assumption. Qed. Lemma two_ops_imp_plus_p : Comp_closed_sub_p -> Times_closed_sub_p -> Plus_closed_sub_p. intros h1 h2. rewrite Comp_closed_sub_p_iff in h1. rewrite Times_closed_sub_p_iff in h2. pose proof (two_ops_imp_plus _ _ h1 h2) as h3. rewrite <- Plus_closed_sub_p_iff in h3. assumption. Qed. Definition ba_conv_subalg : Bool_Alg. rewrite alg_closed_p_iff in Acp. refine (Subalg _ Acp). Defined. Lemma im_proj1_sig_ba_conv_subalg_eq : im_proj1_sig (ba_ens (ba_conv Subalg_p)) = im_proj1_sig (im_proj1_sig (ba_ens ba_conv_subalg)). apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. subst. destruct x as [x h2]. simpl. unfold Subalg_p in h2. simpl in h2. assert (h3:Inn ba_conv_und_subalg (exist _ _ (incl_subalg _ h2))). apply Im_intro with (exist _ _ h2). constructor. apply proj1_sig_injective; simpl. reflexivity. apply Im_intro with (exist _ _ (incl_subalg _ h2)). unfold ba_conv_subalg. apply Im_intro with (exist _ _ h3). constructor. simpl. reflexivity. simpl. reflexivity. red. intros x h1. destruct h1 as [x h1]. subst. destruct x as [x h2]. simpl. inversion h1 as [x' h3 ? h4]. subst. destruct x' as [x' h5]. destruct h5 as [x' h5]. subst. simpl in h4. pose proof (f_equal (@proj1_sig _ _) h4) as h4'. clear h4. simpl in h4'. subst. apply Im_intro with x'. unfold ba_ens, ba_conv, Subalg_p. simpl. unfold bt. simpl. constructor. reflexivity. Qed. End Bc_p''. End AlgClosed_p. Arguments alg_closed_p [T] [Bp] _ _. Arguments alg_closed_p_iff [T] [Bp] _ _. Arguments Subalg_p [T] _ _ _ _. Arguments Plus_closed_sub_p [T] [Bp] _ _. Arguments Times_closed_sub_p [T] [Bp] _ _. Arguments Comp_closed_sub_p [T] [Bp] _ _. Arguments Zero_closed_sub_p [T] [Bp] _. Arguments One_closed_sub_p [T] [Bp] _. Arguments ba_conv_und_subalg [T] [Bp] _ _ _. Lemma incl_ens_btp : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)), Included (im_proj1_sig E) (ba_p_ens Bp). intros T Bp E. red. intros x h1. destruct h1 as [x h1]. subst. apply proj2_sig. Qed. Lemma in_ba_conv_und_subalg_iff : forall {T:Type} (Bp:Bool_Alg_p T) (E:Ensemble T) (pf:Included E (ba_p_ens Bp)) (x:btp Bp), Inn (ba_conv_und_subalg E pf) x <-> Inn E (proj1_sig x). intros T Bp E h1 x. split. intro h2. destruct h2 as [x h2]. subst. simpl. apply proj2_sig. intro h2. apply Im_intro with (exist _ _ h2). constructor. simpl. destruct x as [x h3]. simpl. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma ba_conv_und_subalg_im_proj1_sig : forall {T:Type} {Bp:Bool_Alg_p T} (S:Ensemble (btp Bp)) (pf:Included (im_proj1_sig S) (ba_p_ens Bp)), ba_conv_und_subalg (im_proj1_sig S) pf = S. intros T Bp S h1. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. clear h2. destruct x as [x h2]. simpl. destruct h2 as [x h2]. subst. simpl. destruct x as [x h3]. simpl. generalize (h1 x (Im_intro (sig_set (A_p T (Bc_p T Bp))) T S (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h3) h2 x eq_refl)). intro h4. pose proof (proof_irrelevance _ h4 h3); subst; auto. red; intros x h2. unfold ba_conv_und_subalg. assert (h3:Inn (im_proj1_sig S) (proj1_sig x)). apply Im_intro with x; auto. apply Im_intro with (exist _ _ h3). constructor. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma alg_closed_p_ba_p_ens_refl : forall {T:Type} (Bp:Bool_Alg_p T), alg_closed_p (ba_p_ens Bp) (inclusion_reflexive _). intros T Bp. constructor. red. intros. unfold Bplus_sub_p. destruct x as [x h1]. destruct y as [y h2]. apply in_ba_p_ens_plus. destruct x as [x h1]. destruct y as [y h2]. apply in_ba_p_ens_times. red. apply in_ba_p_ens_one. red. apply in_ba_p_ens_zero. red. intros. destruct x as [x h1]. apply in_ba_p_ens_comp. Qed. Lemma subalg_functional_p : forall {T:Type} {Bp:Bool_Alg_p T} (A A':Ensemble T) (pfeq:A = A') (pfia:Included A (A_p T (Bc_p T Bp))) (pfia':Included A' (A_p T (Bc_p T Bp))), forall (pfa:alg_closed_p A pfia) (pfa':alg_closed_p A' pfia'), Subalg_p _ A pfia pfa = Subalg_p _ A' pfia' pfa'. intros T Bp A A' h1 h2 h3 h4 h5. subst. pose proof (proof_irrelevance _ h2 h3); subst. pose proof (proof_irrelevance _ h4 h5); subst; auto. Qed. Lemma in_sup_alg_p : forall {T:Type} {Bp:Bool_Alg_p T} {A:Ensemble T} {pfa:Included A (ba_p_ens Bp)} {pfcl:alg_closed_p _ pfa} (x:Btype_p T (Bc_p T (Subalg_p _ _ pfa pfcl))), Inn (ba_p_ens Bp) (proj1_sig x). intros T Bp A h1 h2 x. destruct x as [x h3]. simpl. simpl in h3. apply (h1 _ h3). Qed. Section Closed_p. Variable T:Type. Variable Bp:Bool_Alg_p T. Let Btp := Btype_p T (Bc_p T Bp). Lemma plus_closed_p : forall (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), alg_closed_p S pf -> forall (x y:Btp) (pfx:Inn (ba_conv_und_subalg _ pf) x) (pfy:Inn (ba_conv_und_subalg _ pf) y), Inn (ba_conv_und_subalg _ pf) (x %+ y). intros S h1 h2 x y h3 h4. rewrite (alg_closed_p_iff _ h1) in h2. pose proof (plus_closed _ _ h2) as h5. specialize (h5 x y h3 h4). assumption. Qed. Lemma times_closed_p : forall (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), alg_closed_p S pf -> forall (x y:Btp) (pfx:Inn (ba_conv_und_subalg _ pf) x) (pfy:Inn (ba_conv_und_subalg _ pf) y), Inn (ba_conv_und_subalg _ pf) (x %* y). intros S h1 h2 x y h3 h4. rewrite (alg_closed_p_iff _ h1) in h2. pose proof (times_closed _ _ h2) as h5. specialize (h5 x y h3 h4). assumption. Qed. Lemma comp_closed_p : forall (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), alg_closed_p S pf -> forall (x:Btp) (pfx:Inn (ba_conv_und_subalg _ pf) x), Inn (ba_conv_und_subalg _ pf) (%-x). intros S h1 h2 x h3. rewrite (alg_closed_p_iff _ h1) in h2. pose proof (comp_closed _ _ h2) as h5. specialize (h5 x h3). assumption. Qed. Lemma zero_closed_p : forall (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), alg_closed_p S pf -> Inn (ba_conv_und_subalg _ pf) (%0). intros S h1 h2. rewrite (alg_closed_p_iff _ h1) in h2. pose proof (zero_closed _ _ h2) as h3. assumption. Qed. Lemma one_closed_p : forall (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), alg_closed_p S pf -> Inn (ba_conv_und_subalg _ pf) (%1). intros S h1 h2. rewrite (alg_closed_p_iff _ h1) in h2. pose proof (one_closed _ _ h2) as h3. assumption. Qed. Lemma plus_closed_p' : forall (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), alg_closed_p S pf -> forall (x y :T) (pfx:Inn S x) (pfy:Inn S y), Inn S (proj1_sig ((exist _ _ (pf _ pfx)) %+ (exist _ _ (pf _ pfy)))). intros S h1 h2 x y h3 h4. destruct h2 as [h2a]. red in h2a. specialize (h2a (exist _ _ h3) (exist _ _ h4)). unfold Bplus_sub_p in h2a. assumption. Qed. Lemma times_closed_p' : forall (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), alg_closed_p S pf -> forall (x y :T) (pfx:Inn S x) (pfy:Inn S y), Inn S (proj1_sig ((exist _ _ (pf _ pfx)) %* (exist _ _ (pf _ pfy)))). intros S h1 h2 x y h3 h4. destruct h2 as [? h2a]. red in h2a. specialize (h2a (exist _ _ h3) (exist _ _ h4)). unfold Btimes_sub_p in h2a. assumption. Qed. Lemma comp_closed_p' : forall (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), alg_closed_p S pf -> forall (x:T) (pfx:Inn S x), Inn S (proj1_sig (%-(exist _ _ (pf _ pfx)))). intros S h1 h2 x h3. destruct h2 as [? ? ? ? h2a]. simpl in h2a. red in h2a. specialize (h2a (exist _ _ h3)). unfold Bcomp_sub_p in h2a. assumption. Qed. Lemma one_closed_p' : forall (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), alg_closed_p S pf -> Inn S (proj1_sig (Bone_p _ (Bc_p _ Bp))). intros S h1 h2. destruct h2 as [? ? h2a]. red in h2a. assumption. Qed. Lemma zero_closed_p' : forall (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), alg_closed_p S pf -> Inn S (proj1_sig (Bzero_p _ (Bc_p _ Bp))). intros S h1 h2. destruct h2 as [? ? ? h2a]. red in h2a. assumption. Qed. Lemma closed_fun2_plus : forall (Cp:Bool_Alg_p T) (pf:Included (ba_p_ens Bp) (ba_p_ens Cp)), alg_closed_p _ pf -> closed_fun2 (Bplus_p _ (Bc_p _ Cp)) pf. intros Cp h0 h1. red. intros x y h2 h3. destruct h1 as [h1]. red in h1. pose proof (h1 (exist _ _ h2) (exist _ _ h3)) as h4. simpl in h4. assumption. Qed. Lemma closed_fun2_times : forall (Cp:Bool_Alg_p T) (pf:Included (ba_p_ens Bp) (ba_p_ens Cp)), alg_closed_p _ pf -> closed_fun2 (Btimes_p _ (Bc_p _ Cp)) pf. intros Cp h0 h1. red. intros x y h2 h3. destruct h1 as [? h1]. red in h1. pose proof (h1 (exist _ _ h2) (exist _ _ h3)) as h4. simpl in h4. assumption. Qed. Lemma closed_fun_comp : forall (Cp:Bool_Alg_p T) (pf:Included (ba_p_ens Bp) (ba_p_ens Cp)), alg_closed_p _ pf -> closed_fun (Bcomp_p _ (Bc_p _ Cp)) pf. intros Cp h0 h1. red. intros x h2. destruct h1 as [? ? ? ? h1]. red in h1. pose proof (h1 (exist _ _ h2)) as h4. simpl in h4. assumption. Qed. Lemma alg_closed_subst_iff_p : forall (E F:Ensemble T), E = F -> forall (pfe:Included E (ba_p_ens Bp)) (pff:Included F (ba_p_ens Bp)), (alg_closed_p E pfe <-> alg_closed_p F pff). intros E F h1 h2 h3. subst. pose proof (proof_irrelevance _ h2 h3); subst; tauto. Qed. End Closed_p. Arguments closed_fun2_plus [T] _ _ _ _ _ _ _ _. Arguments closed_fun2_times [T] _ _ _ _ _ _ _ _. Arguments closed_fun_comp [T] _ _ _ _ _ _. Section Subalg_p. Lemma ba_p_ens_subalg_p_compat : forall {T:Type} (Bp:Bool_Alg_p T) (E:Ensemble T) (pfi:Included E (ba_p_ens Bp)) (pfac:alg_closed_p _ pfi), ba_p_ens (Subalg_p _ _ pfi pfac) = E. intros T Bp E h1 h2. apply Extensionality_Ensembles. red. split. red. intros x h3. unfold Subalg_p, ba_p_ens in h3. simpl in h3. assumption. red. intros x h3. unfold Subalg_p, ba_p_ens. simpl. assumption. Qed. Definition subalg_of_p {T:Type} (Ap Bp:Bool_Alg_p T) := exists (pfin:Included (ba_p_ens Ap) (ba_p_ens Bp)) (pfac:alg_closed_p _ pfin), Ap = Subalg_p Bp _ pfin pfac. Lemma subalg_of_p_incl : forall {T:Type} (Ap Bp:Bool_Alg_p T), subalg_of_p Ap Bp -> Included (ba_p_ens Ap) (ba_p_ens Bp). intros T Ap Bp h1. destruct h1; auto. Qed. Lemma subalg_of_p_compat : forall {T:Type} (Bp:Bool_Alg_p T) (S:Ensemble T) (pfi:Included S (ba_p_ens Bp)) (pfa:alg_closed_p _ pfi), subalg_of_p (Subalg_p _ S pfi pfa) Bp. intros; red. ex_goal. rewrite ba_p_ens_subalg_p_compat; auto. exists hex. ex_goal . generalize hex. rewrite ba_p_ens_subalg_p_compat; intro; auto. pose proof (proof_irrelevance _ hex0 pfi); subst; auto. exists hex0. generalize hex hex0. rewrite ba_p_ens_subalg_p_compat; intros; apply f_equal_dep; apply proof_irrelevance. Qed. Lemma refl_subalg_of_p : forall (T:Type), Reflexive (@subalg_of_p T). intro T. red. intro B. red. exists (inclusion_reflexive _). exists (alg_closed_p_ba_p_ens_refl _). apply bc_inj_p. assert (h1:A_p T (Bc_p T B) = A_p T (Bc_p T (Subalg_p B (ba_p_ens B) (inclusion_reflexive (ba_p_ens B)) (alg_closed_p_ba_p_ens_refl B)))). simpl. reflexivity. pose proof (proof_irrelevance _ (sig_set_eq (A_p T (Bc_p T B)) (ba_p_ens B) h1) (eq_refl _)) as h2. apply (bconst_ext_p _ _ h1); simpl. simpl in h1. rewrite h2, transfer_dep_r_eq_refl. apply und_set_p. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite transfer_dep_r_fun2_eq. rewrite h2. unfold transfer_r. unfold eq_rect_r. simpl. destruct x as [x h3], y as [y h4]. apply proj1_sig_injective. simpl. simpl. pose proof (proof_irrelevance _ h3 (inclusion_reflexive (ba_p_ens B) x h3)) as h5. pose proof (proof_irrelevance _ h4 (inclusion_reflexive (ba_p_ens B) y h4)) as h6. rewrite h5, h6. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. rewrite h2, transfer_dep_r_eq_refl. apply functional_extensionality. intro x. apply functional_extensionality. intro y. unfold transfer_r. unfold eq_rect_r. simpl. destruct x as [x h3], y as [y h4]. apply proj1_sig_injective. simpl. unfold transfer. unfold eq_rect_r. simpl. pose proof (proof_irrelevance _ h3 (inclusion_reflexive (ba_p_ens B) x h3)) as h5. pose proof (proof_irrelevance _ h4 (inclusion_reflexive (ba_p_ens B) y h4)) as h6. rewrite h5, h6. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. rewrite h2. rewrite transfer_dep_r_eq_refl. apply proj1_sig_injective. simpl. reflexivity. rewrite h2. rewrite transfer_dep_r_eq_refl. apply proj1_sig_injective. simpl. reflexivity. apply functional_extensionality. intro x. rewrite transfer_dep_r_fun1_eq. rewrite h2. unfold transfer_r. unfold eq_rect_r. simpl. destruct x as [x h3]. apply proj1_sig_injective. simpl. simpl. unfold transfer. unfold eq_rect_r. simpl. pose proof (proof_irrelevance _ h3 (inclusion_reflexive (ba_p_ens B) x h3)) as h5. rewrite h5. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma trans_subalg_of_p : forall (T:Type), Transitive (@subalg_of_p T). intro T. red. intros Ap Bp Cp h1 h2. red in h1, h2. destruct h1 as [h3 [h4 h5]], h2 as [h6 [h7 h8]]. red. exists (Inclusion_is_transitive _ _ _ _ h3 h6). destruct h4 as [h4a h4b h4c h4d h4e], h7 as [h7a h7b h7c h7d h7e]. assert (h9:alg_closed_p (ba_p_ens Ap) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6)). constructor. red. intros x y. unfold Bplus_sub_p. pose proof (h4a x y) as h9. unfold Bplus_sub_p in h9. destruct x as [x h11], y as [y h12]. pose proof (ba_p_subst_plus _ _ h8 _ _ (h3 x h11) (h3 y h12)) as h13. simpl in h13. assert (h14: Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h6 {| P_c_p := h7a; T_c_p := h7b; O_c_p := h7c; Z_c_p := h7d; C_c_p := h7e |})) x). rewrite <- h8. apply (h3 x h11). assert (h15 : Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h6 {| P_c_p := h7a; T_c_p := h7b; O_c_p := h7c; Z_c_p := h7d; C_c_p := h7e |})) y). rewrite <- h8. apply (h3 y h12). specialize (h13 h14 h15). rewrite h13 in h9 at 1. pose proof (proof_irrelevance _ (h6 x h14) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6 x h11)) as h16. pose proof (proof_irrelevance _ (h6 y h15) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6 y h12)) as h17. rewrite h16, h17 in h9. assumption. red. intros x y. pose proof (h4b x y) as h9. unfold Btimes_sub_p in h9. destruct x as [x h11], y as [y h12]. pose proof (ba_p_subst_times _ _ h8 _ _ (h3 x h11) (h3 y h12)) as h13. simpl in h13. assert (h14: Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h6 {| P_c_p := h7a; T_c_p := h7b; O_c_p := h7c; Z_c_p := h7d; C_c_p := h7e |})) x). rewrite <- h8. apply (h3 x h11). assert (h15 : Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h6 {| P_c_p := h7a; T_c_p := h7b; O_c_p := h7c; Z_c_p := h7d; C_c_p := h7e |})) y). rewrite <- h8. apply (h3 y h12). specialize (h13 h14 h15). rewrite h13 in h9 at 1. pose proof (proof_irrelevance _ (h6 x h14) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6 x h11)) as h16. pose proof (proof_irrelevance _ (h6 y h15) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6 y h12)) as h17. rewrite h16, h17 in h9. assumption. red. red in h4c, h7c. pose proof (ba_p_subst_one _ _ h8) as h9. pose proof h4c as h4c'. rewrite h9 in h4c'. pose proof (ba_p_subst_one _ _ h5) as h10. simpl in h10, h4c'. rewrite h8 in h10. assumption. red. red in h4d, h7d. pose proof (ba_p_subst_zero _ _ h8) as h9. pose proof h4d as h4d'. rewrite h9 in h4d'. pose proof (ba_p_subst_zero _ _ h5) as h10. simpl in h10, h4d'. rewrite h8 in h10. assumption. red. intros. unfold Bcomp_sub_p. destruct x as [x h11]. red in h4e, h7e. pose proof (h4e (exist _ _ h11)) as h12. simpl in h12. pose proof (ba_p_subst_comp _ _ h8 _ (h3 x h11)) as h13. simpl in h13. assert (h14:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h6 {| P_c_p := h7a; T_c_p := h7b; O_c_p := h7c; Z_c_p := h7d; C_c_p := h7e |})) x). rewrite <- h8. apply (h3 x h11). specialize (h13 h14). rewrite h13 in h12 at 1. pose proof (proof_irrelevance _ (h6 x h14) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6 x h11)) as h15. rewrite h15 in h12; auto. exists h9. apply bc_inj_p. assert (h10:A_p T (Bc_p T Ap) = A_p T (Bc_p T (Subalg_p Cp (ba_p_ens Ap) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6) h9))). simpl. reflexivity. pose proof (proof_irrelevance _ (sig_set_eq (A_p T (Bc_p T Ap)) (ba_p_ens Ap) h10) (eq_refl _)) as h0. apply (bconst_ext_p _ _ h10); simpl; rewrite h0; try rewrite transfer_dep_r_eq_refl. apply und_set_p. apply functional_extensionality. intro x. apply functional_extensionality. intro y. apply proj1_sig_injective. simpl. destruct x as [x h11], y as [y h12]. red in h4a, h7a. unfold Bplus_sub_p. simpl. pose proof (ba_p_subst_plus _ _ h5 _ _ h11 h12) as h13. simpl in h13. assert (h14: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h3 {| P_c_p := h4a; T_c_p := h4b; O_c_p := h4c; Z_c_p := h4d; C_c_p := h4e |})) x). rewrite <- h5. assumption. assert (h15 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h3 {| P_c_p := h4a; T_c_p := h4b; O_c_p := h4c; Z_c_p := h4d; C_c_p := h4e |})) y). rewrite <- h5. assumption. specialize (h13 h14 h15). rewrite h13 at 1. pose proof (ba_p_subst_plus _ _ h8 _ _ (h3 x h14) (h3 y h15)) as h16. simpl in h16. assert (h17 : Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h6 {| P_c_p := h7a; T_c_p := h7b; O_c_p := h7c; Z_c_p := h7d; C_c_p := h7e |})) x). rewrite <- h8. apply (h3 x h14). assert (h18 : Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h6 {| P_c_p := h7a; T_c_p := h7b; O_c_p := h7c; Z_c_p := h7d; C_c_p := h7e |})) y). rewrite <- h8. apply (h3 y h15). specialize (h16 h17 h18). rewrite h16 at 1. pose proof (proof_irrelevance _ (h6 x h17) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6 x h11)) as h19. pose proof (proof_irrelevance _ (h6 y h18) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6 y h12)) as h20. rewrite h19, h20. reflexivity. apply functional_extensionality. intro x. apply functional_extensionality. intro y. apply proj1_sig_injective. simpl. destruct x as [x h11], y as [y h12]. red in h4a, h7a. unfold Bplus_sub_p. simpl. pose proof (ba_p_subst_times _ _ h5 _ _ h11 h12) as h13. simpl in h13. assert (h14: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h3 {| P_c_p := h4a; T_c_p := h4b; O_c_p := h4c; Z_c_p := h4d; C_c_p := h4e |})) x). rewrite <- h5. assumption. assert (h15 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h3 {| P_c_p := h4a; T_c_p := h4b; O_c_p := h4c; Z_c_p := h4d; C_c_p := h4e |})) y). rewrite <- h5. assumption. specialize (h13 h14 h15). rewrite h13 at 1. pose proof (ba_p_subst_times _ _ h8 _ _ (h3 x h14) (h3 y h15)) as h16. simpl in h16. assert (h17 : Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h6 {| P_c_p := h7a; T_c_p := h7b; O_c_p := h7c; Z_c_p := h7d; C_c_p := h7e |})) x). rewrite <- h8. apply (h3 x h14). assert (h18 : Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h6 {| P_c_p := h7a; T_c_p := h7b; O_c_p := h7c; Z_c_p := h7d; C_c_p := h7e |})) y). rewrite <- h8. apply (h3 y h15). specialize (h16 h17 h18). rewrite h16 at 1. pose proof (proof_irrelevance _ (h6 x h17) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6 x h11)) as h19. pose proof (proof_irrelevance _ (h6 y h18) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6 y h12)) as h20. rewrite h19, h20. reflexivity. apply proj1_sig_injective. simpl. rewrite h5. simpl. rewrite h8. simpl. reflexivity. apply proj1_sig_injective. simpl. rewrite h5. simpl. rewrite h8. simpl. reflexivity. apply functional_extensionality. intro x. destruct x as [x h11]. simpl. apply proj1_sig_injective. simpl. pose proof (ba_p_subst_comp _ _ h5 _ h11) as h12. simpl in h12. assert (h13: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h3 {| P_c_p := h4a; T_c_p := h4b; O_c_p := h4c; Z_c_p := h4d; C_c_p := h4e |})) x). rewrite <- h5. assumption. specialize (h12 h13). rewrite h12 at 1. pose proof (ba_p_subst_comp _ _ h8 _ (h3 x h13)) as h14. assert (h15: Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h6 {| P_c_p := h7a; T_c_p := h7b; O_c_p := h7c; Z_c_p := h7d; C_c_p := h7e |})) x). rewrite <- h8. apply (h3 x h13). specialize (h14 h15). rewrite h14 at 1. simpl. pose proof (proof_irrelevance _ (h6 x h15) (Inclusion_is_transitive T (ba_p_ens Ap) (ba_p_ens Bp) (ba_p_ens Cp) h3 h6 x h11)) as h16. rewrite h16. reflexivity. Qed. Lemma alg_closed_p_subalg_p_of : forall {T:Type} (Bp Cp:Bool_Alg_p T) (S:Ensemble T) (pfb:Included S (ba_p_ens Bp)) (pfc:Included S (ba_p_ens Cp)), subalg_of_p Bp Cp -> alg_closed_p S pfb -> alg_closed_p S pfc. intros T Bp Cp S h1 h2 h3 h4. red in h3. destruct h3 as [h3 [h5 h6]]. destruct h4 as [h4a h4b h4c h4d h4e]. constructor. red. intros x y. destruct x as [x h7], y as [y h8]. specialize (h4a (exist _ _ h7) (exist _ _ h8)). assert (h10: Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h3 h5)) x). rewrite <- h6. apply h1; auto. assert (h11:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h3 h5)) y). rewrite <- h6. apply h1; auto. pose proof (ba_p_subst_plus _ _ h6 _ _ (h1 _ h7) (h1 _ h8) h10 h11) as h9. unfold Bplus_sub_p in h4a. unfold Bplus_sub_p. rewrite h9 in h4a at 1. simpl in h4a. pose proof (proof_irrelevance _ (h3 x h10) (h2 x h7)) as h12. pose proof (proof_irrelevance _ (h3 y h11) (h2 y h8)) as h13. rewrite h12, h13 in h4a; auto. red. intros x y. destruct x as [x h7], y as [y h8]. specialize (h4b (exist _ _ h7) (exist _ _ h8)). assert (h10: Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h3 h5)) x). rewrite <- h6. apply h1; auto. assert (h11:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h3 h5)) y). rewrite <- h6. apply h1; auto. pose proof (ba_p_subst_times _ _ h6 _ _ (h1 _ h7) (h1 _ h8) h10 h11) as h9. unfold Btimes_sub_p in h4b. unfold Btimes_sub_p. rewrite h9 in h4b at 1. simpl in h4b. pose proof (proof_irrelevance _ (h3 x h10) (h2 x h7)) as h12. pose proof (proof_irrelevance _ (h3 y h11) (h2 y h8)) as h13. rewrite h12, h13 in h4b. assumption. red. red in h4c. pose proof (ba_p_subst_one _ _ h6) as h7. rewrite <- h7 at 1. assumption. red. red in h4d. pose proof (ba_p_subst_zero _ _ h6) as h7. rewrite <- h7 at 1. assumption. red. intro x. destruct x as [x h7]. specialize (h4e (exist _ _ h7)). assert (h10: Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h3 h5)) x). rewrite <- h6. apply h1; auto. pose proof (ba_p_subst_comp _ _ h6 _ (h1 _ h7) h10) as h9. unfold Bcomp_sub_p in h4b. unfold Bcomp_sub_p. rewrite h9 in h4e at 1. simpl in h4e. pose proof (proof_irrelevance _ (h3 x h10) (h2 x h7)) as h12. rewrite h12 in h4e; auto. Qed. Lemma subalg_of_p_alg_closed_subalg_compat : forall {T:Type} (Bp Cp:Bool_Alg_p T), subalg_of_p Bp Cp -> forall (B:Ensemble (bt (ba_conv Bp))) (C:Ensemble (bt (ba_conv Cp))) (pfb:alg_closed B) (pfc:alg_closed C) (S:Ensemble (bt (Subalg B pfb))) (S':Ensemble (bt (Subalg C pfc))), im_proj1_sig (im_proj1_sig S) = im_proj1_sig (im_proj1_sig S') -> (alg_closed S <-> alg_closed S'). intros T Bp Cp h1 B C h2 h3 S S' h4. split. intro h5. destruct h5 as [h5a h5b h5c h5d h5e]. red in h5a, h5b, h5c, h5d, h5e. constructor; red. intros x y. destruct x as [x h6], y as [y h7]. assert (h8:Inn (im_proj1_sig (im_proj1_sig S')) (proj1_sig (proj1_sig x))). destruct x as [x h8]. simpl. destruct x as [x h9]. simpl. assert (h10:Inn (im_proj1_sig S') (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) x h9)). apply Im_intro with (exist (fun a : Btype (Bc (ba_conv Cp)) => Inn C a) (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) x h9) h8). assumption. apply proj1_sig_injective. simpl. reflexivity. apply Im_intro with (exist _ _ h9). assumption. simpl. reflexivity. rewrite <- h4 in h8. assert (h8':Inn (im_proj1_sig (im_proj1_sig S')) (proj1_sig (proj1_sig y))). destruct y as [y h8']. simpl. destruct y as [y h9']. simpl. assert (h10:Inn (im_proj1_sig S') (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) y h9')). apply Im_intro with (exist (fun a : Btype (Bc (ba_conv Cp)) => Inn C a) (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) y h9') h8'). assumption. apply proj1_sig_injective. simpl. reflexivity. apply Im_intro with (exist _ _ h9'). assumption. simpl. reflexivity. rewrite <- h4 in h8'. inversion h8 as [x' h10 ? h11]. subst. inversion h8' as [y' h10' ? h11']. subst. destruct h10 as [x' h10]. subst. destruct h10' as [y' h10']. subst. specialize (h5a (exist _ _ h10) (exist _ _ h10')). destruct x as [x h20], x' as [x' h21], y as [y h22], y' as [y' h23]. destruct x as [x h24], x' as [x' h25], y as [y h26], y' as [y' h27]. simpl in h8, h8', h11, h11'. subst. assert (h12:Inn (im_proj1_sig S) (proj1_sig (Bplus_sub (Subalg B h2) S (exist (Inn S) (exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h25) h21) h10) (exist (Inn S) (exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) y' h27) h23) h10')))). eapply Im_intro. apply h5a. reflexivity. assert (h13:Inn (im_proj1_sig (im_proj1_sig S)) (proj1_sig (proj1_sig (Bplus_sub (Subalg B h2) S (exist (Inn S) (exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h25) h21) h10) (exist (Inn S) (exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) y' h27) h23) h10'))))). eapply Im_intro. apply h12. reflexivity. rewrite h4 in h13. inversion h13 as [z h14 ? h15]. subst. destruct z as [z h16]. simpl in h15. subst. inversion h14 as [z h17 ? h18]. subst. destruct z as [z h19]. simpl in h18. subst. subst. subst. unfold Bplus_sub, Subalg. simpl. unfold Bplus_sub. simpl. unfold Bplus_sub in h17. simpl in h17. unfold Bplus_sub in h17. simpl in h17. assert (h28: exist (fun x : sig_set (A_p T (Bc_p T Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) (proj1_sig (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h25 %+ exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) y' h27)) h16) h19 = exist (Inn C) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h24 %+ exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) y' h26) (P_c (ba_conv Cp) C h3 (exist (fun a : sig_set (A_p T (Bc_p T Cp)) => Inn C a) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h24) h20) (exist (fun a : sig_set (A_p T (Bc_p T Cp)) => Inn C a) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) y' h26) h22))). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. red in h1. destruct h1 as [h1a [h1b h1c]]. assert (h30:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) x'). rewrite <- h1c. assumption. assert (h31:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) y'). rewrite <- h1c. assumption. pose proof (ba_p_subst_plus _ _ h1c _ _ h25 h27 h30 h31) as h29. rewrite h29 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. reflexivity. apply proj1_sig_injective. reflexivity. rewrite <- h28 at 1. assumption. intros x y. destruct x as [x h6], y as [y h7]. assert (h8:Inn (im_proj1_sig (im_proj1_sig S')) (proj1_sig (proj1_sig x))). destruct x as [x h8]. simpl. destruct x as [x h9]. simpl. assert (h10:Inn (im_proj1_sig S') (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) x h9)). apply Im_intro with (exist (fun a : Btype (Bc (ba_conv Cp)) => Inn C a) (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) x h9) h8). assumption. apply proj1_sig_injective. simpl. reflexivity. apply Im_intro with (exist _ _ h9). assumption. simpl. reflexivity. rewrite <- h4 in h8. assert (h8':Inn (im_proj1_sig (im_proj1_sig S')) (proj1_sig (proj1_sig y))). destruct y as [y h8']. simpl. destruct y as [y h9']. simpl. assert (h10:Inn (im_proj1_sig S') (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) y h9')). apply Im_intro with (exist (fun a : Btype (Bc (ba_conv Cp)) => Inn C a) (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) y h9') h8'). assumption. apply proj1_sig_injective. simpl. reflexivity. apply Im_intro with (exist _ _ h9'). assumption. simpl. reflexivity. rewrite <- h4 in h8'. inversion h8 as [x' h10 ? h11]. subst. inversion h8' as [y' h10' ? h11']. subst. destruct h10 as [x' h10]. subst. destruct h10' as [y' h10']. subst. specialize (h5b (exist _ _ h10) (exist _ _ h10')). destruct x as [x h20], x' as [x' h21], y as [y h22], y' as [y' h23]. destruct x as [x h24], x' as [x' h25], y as [y h26], y' as [y' h27]. simpl in h8, h8', h11, h11'. subst. assert (h12:Inn (im_proj1_sig S) (proj1_sig (Btimes_sub (Subalg B h2) S (exist (Inn S) (exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h25) h21) h10) (exist (Inn S) (exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) y' h27) h23) h10')))). eapply Im_intro. apply h5b. reflexivity. assert (h13:Inn (im_proj1_sig (im_proj1_sig S)) (proj1_sig (proj1_sig (Btimes_sub (Subalg B h2) S (exist (Inn S) (exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h25) h21) h10) (exist (Inn S) (exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) y' h27) h23) h10'))))). eapply Im_intro. apply h12. reflexivity. rewrite h4 in h13. inversion h13 as [z h14 ? h15]. subst. destruct z as [z h16]. simpl in h15. subst. inversion h14 as [z h17 ? h18]. subst. destruct z as [z h19]. simpl in h18. subst. subst. subst. unfold Btimes_sub, Subalg. simpl. unfold Btimes_sub. simpl. unfold Btimes_sub in h17. simpl in h17. unfold Btimes_sub in h17. simpl in h17. assert (h28: exist (fun x : sig_set (A_p T (Bc_p T Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) (proj1_sig (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h25 %* exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) y' h27)) h16) h19 = exist (Inn C) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h24 %* exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) y' h26) (T_c (ba_conv Cp) C h3 (exist (fun a : sig_set (A_p T (Bc_p T Cp)) => Inn C a) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h24) h20) (exist (fun a : sig_set (A_p T (Bc_p T Cp)) => Inn C a) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) y' h26) h22))). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. red in h1. destruct h1 as [h1a [h1b h1c]]. assert (h30:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) x'). rewrite <- h1c. assumption. assert (h31:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) y'). rewrite <- h1c. assumption. pose proof (ba_p_subst_times _ _ h1c _ _ h25 h27 h30 h31) as h29. rewrite h29 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. reflexivity. apply proj1_sig_injective. reflexivity. rewrite <- h28 at 1. assumption. destruct h1 as [h1a [h1b h1c]]. pose proof (ba_p_subst_one _ _ h1c) as h6. simpl. simpl in h5c. assert (h7:Inn (im_proj1_sig S) %1). apply Im_intro with (exist (Inn B) %1 (O_c (ba_conv Bp) B h2)). assumption. simpl. reflexivity. assert (h8:Inn (im_proj1_sig (im_proj1_sig S)) (proj1_sig (Bone_p _ (Bc_p _ Bp)))). apply Im_intro with %1. assumption. reflexivity. rewrite h4 in h8 at 1. rewrite h6 in h8. inversion h8 as [a h9 ? h10]. clear h1c. subst. destruct a as [a h11]. subst. simpl in h10. subst. inversion h9 as [a h12 ? h13]. subst. destruct a as [a h14]. simpl in h13. subst. assert (h15: exist (Inn C) %1 (O_c (ba_conv Cp) C h3) = exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) (proj1_sig %1) h11) h14). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. reflexivity. rewrite h15 at 1. assumption. destruct h1 as [h1a [h1b h1c]]. pose proof (ba_p_subst_zero _ _ h1c) as h6. simpl. simpl in h5d. assert (h7:Inn (im_proj1_sig S) %0). apply Im_intro with (exist (Inn B) %0 (Z_c (ba_conv Bp) B h2)). assumption. simpl. reflexivity. assert (h8:Inn (im_proj1_sig (im_proj1_sig S)) (proj1_sig (Bzero_p _ (Bc_p _ Bp)))). apply Im_intro with %0. assumption. reflexivity. rewrite h4 in h8 at 1. rewrite h6 in h8. inversion h8 as [a h9 ? h10]. clear h1c. subst. destruct a as [a h11]. subst. simpl in h10. subst. inversion h9 as [a h12 ? h13]. subst. destruct a as [a h14]. simpl in h13. subst. assert (h15: exist (Inn C) %0 (Z_c (ba_conv Cp) C h3) = exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) (proj1_sig %0) h11) h14). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. reflexivity. rewrite h15 at 1. assumption. intro x. destruct x as [x h6]. assert (h8:Inn (im_proj1_sig (im_proj1_sig S')) (proj1_sig (proj1_sig x))). destruct x as [x h8]. simpl. destruct x as [x h9]. simpl. assert (h10:Inn (im_proj1_sig S') (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) x h9)). apply Im_intro with (exist (fun a : Btype (Bc (ba_conv Cp)) => Inn C a) (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) x h9) h8). assumption. apply proj1_sig_injective. simpl. reflexivity. apply Im_intro with (exist _ _ h9). assumption. simpl. reflexivity. rewrite <- h4 in h8. inversion h8 as [x' h10 ? h11]. subst. destruct h10 as [x' h10]. subst. specialize (h5e (exist _ _ h10)). destruct x as [x h20], x' as [x' h21]. destruct x as [x h24], x' as [x' h25]. simpl in h8, h11. subst. assert (h12:Inn (im_proj1_sig S) (proj1_sig (Bcomp_sub (Subalg B h2) S (exist (Inn S) (exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h25) h21) h10)))). eapply Im_intro. apply h5e. reflexivity. assert (h13:Inn (im_proj1_sig (im_proj1_sig S)) (proj1_sig (proj1_sig (Bcomp_sub (Subalg B h2) S (exist (Inn S) (exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h25) h21) h10))))). eapply Im_intro. apply h12. reflexivity. rewrite h4 in h13. inversion h13 as [z h14 ? h15]. subst. destruct z as [z h16]. simpl in h15. subst. inversion h14 as [z h17 ? h18]. subst. destruct z as [z h19]. simpl in h18. subst. unfold Bcomp_sub, Subalg. simpl. unfold Bcomp_sub. simpl. unfold Btimes_sub in h17. simpl in h17. unfold Bcomp_sub in h17. simpl in h17. assert (h28:exist (fun x : sig_set (A_p T (Bc_p T Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) (proj1_sig (%-exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h25)) h16) h19 = exist (Inn C) (%-exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h24) (C_c (ba_conv Cp) C h3 (exist (fun a : sig_set (A_p T (Bc_p T Cp)) => Inn C a) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h24) h20))). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. red in h1. destruct h1 as [h1a [h1b h1c]]. assert (h30:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) x'). rewrite <- h1c. assumption. pose proof (ba_p_subst_comp _ _ h1c _ h25 h30) as h29. rewrite h29 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. reflexivity. rewrite <- h28 at 1. assumption. intro h5. destruct h5 as [h5a h5b h5c h5d h5e]. red in h5a, h5b, h5c, h5d, h5e. constructor; red. intros x y. destruct x as [x h6], y as [y h7]. assert (h8:Inn (im_proj1_sig (im_proj1_sig S)) (proj1_sig (proj1_sig x))). destruct x as [x h8]. simpl. destruct x as [x h9]. simpl. assert (h10:Inn (im_proj1_sig S) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h9)). apply Im_intro with (exist (fun a : Btype (Bc (ba_conv Bp)) => Inn B a) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h9) h8). assumption. apply proj1_sig_injective. simpl. reflexivity. apply Im_intro with (exist _ _ h9). assumption. simpl. reflexivity. rewrite h4 in h8. assert (h8':Inn (im_proj1_sig (im_proj1_sig S)) (proj1_sig (proj1_sig y))). destruct y as [y h8']. simpl. destruct y as [y h9']. simpl. assert (h10:Inn (im_proj1_sig S) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) y h9')). apply Im_intro with (exist (fun a : Btype (Bc (ba_conv Bp)) => Inn B a) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) y h9') h8'). assumption. apply proj1_sig_injective. simpl. reflexivity. apply Im_intro with (exist _ _ h9'). assumption. simpl. reflexivity. rewrite h4 in h8'. inversion h8 as [x' h10 ? h11]. subst. inversion h8' as [y' h10' ? h11']. subst. destruct h10 as [x' h10]. subst. destruct h10' as [y' h10']. subst. specialize (h5a (exist _ _ h10) (exist _ _ h10')). destruct x as [x h20], x' as [x' h21], y as [y h22], y' as [y' h23]. destruct x as [x h24], x' as [x' h25], y as [y h26], y' as [y' h27]. simpl in h8, h8', h11, h11'. subst. assert (h12:Inn (im_proj1_sig S') (proj1_sig (Bplus_sub (Subalg C h3) S' (exist (Inn S') (exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h25) h21) h10) (exist (Inn S') (exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) y' h27) h23) h10')))). eapply Im_intro. apply h5a. reflexivity. assert (h13:Inn (im_proj1_sig (im_proj1_sig S')) (proj1_sig (proj1_sig (Bplus_sub (Subalg C h3) S' (exist (Inn S') (exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h25) h21) h10) (exist (Inn S') (exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) y' h27) h23) h10'))))). eapply Im_intro. apply h12. reflexivity. rewrite <- h4 in h13. inversion h13 as [z h14 ? h15]. subst. destruct z as [z h16]. simpl in h15. subst. inversion h14 as [z h17 ? h18]. subst. destruct z as [z h19]. simpl in h18. subst. subst. subst. unfold Bplus_sub, Subalg. simpl. unfold Bplus_sub. simpl. unfold Bplus_sub in h17. simpl in h17. unfold Bplus_sub in h17. simpl in h17. assert (h28: exist (fun x : sig_set (A_p T (Bc_p T Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) (proj1_sig (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h25 %+ exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) y' h27)) h16) h19 = exist (Inn B) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h24 %+ exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) y' h26) (P_c (ba_conv Bp) B h2 (exist (fun a : sig_set (A_p T (Bc_p T Bp)) => Inn B a) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h24) h20) (exist (fun a : sig_set (A_p T (Bc_p T Bp)) => Inn B a) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) y' h26) h22))). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. red in h1. destruct h1 as [h1a [h1b h1c]]. assert (h30:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) x'). rewrite <- h1c. assumption. assert (h31:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) y'). rewrite <- h1c. assumption. pose proof (ba_p_subst_plus _ _ (eq_sym h1c) _ _ h30 h31 h24 h26) as h29. rewrite <- h29 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. reflexivity. apply proj1_sig_injective. reflexivity. rewrite <- h28 at 1. assumption. intros x y. destruct x as [x h6], y as [y h7]. assert (h8:Inn (im_proj1_sig (im_proj1_sig S)) (proj1_sig (proj1_sig x))). destruct x as [x h8]. simpl. destruct x as [x h9]. simpl. assert (h10:Inn (im_proj1_sig S) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h9)). apply Im_intro with (exist (fun a : Btype (Bc (ba_conv Bp)) => Inn B a) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h9) h8). assumption. apply proj1_sig_injective. simpl. reflexivity. apply Im_intro with (exist _ _ h9). assumption. simpl. reflexivity. rewrite h4 in h8. assert (h8':Inn (im_proj1_sig (im_proj1_sig S)) (proj1_sig (proj1_sig y))). destruct y as [y h8']. simpl. destruct y as [y h9']. simpl. assert (h10:Inn (im_proj1_sig S) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) y h9')). apply Im_intro with (exist (fun a : Btype (Bc (ba_conv Bp)) => Inn B a) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) y h9') h8'). assumption. apply proj1_sig_injective. simpl. reflexivity. apply Im_intro with (exist _ _ h9'). assumption. simpl. reflexivity. rewrite h4 in h8'. inversion h8 as [x' h10 ? h11]. subst. inversion h8' as [y' h10' ? h11']. subst. destruct h10 as [x' h10]. subst. destruct h10' as [y' h10']. subst. specialize (h5b (exist _ _ h10) (exist _ _ h10')). destruct x as [x h20], x' as [x' h21], y as [y h22], y' as [y' h23]. destruct x as [x h24], x' as [x' h25], y as [y h26], y' as [y' h27]. simpl in h8, h8', h11, h11'. subst. assert (h12:Inn (im_proj1_sig S') (proj1_sig (Btimes_sub (Subalg C h3) S' (exist (Inn S') (exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h25) h21) h10) (exist (Inn S') (exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) y' h27) h23) h10')))). eapply Im_intro. apply h5b. reflexivity. assert (h13:Inn (im_proj1_sig (im_proj1_sig S')) (proj1_sig (proj1_sig (Btimes_sub (Subalg C h3) S' (exist (Inn S') (exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h25) h21) h10) (exist (Inn S') (exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) y' h27) h23) h10'))))). eapply Im_intro. apply h12. reflexivity. rewrite <- h4 in h13. inversion h13 as [z h14 ? h15]. subst. destruct z as [z h16]. simpl in h15. subst. inversion h14 as [z h17 ? h18]. subst. destruct z as [z h19]. simpl in h18. subst. subst. subst. unfold Btimes_sub, Subalg. simpl. unfold Btimes_sub. simpl. unfold Bplus_sub in h17. simpl in h17. unfold Btimes_sub in h17. simpl in h17. assert (h28: exist (fun x : sig_set (A_p T (Bc_p T Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) (proj1_sig (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h25 %* exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) y' h27)) h16) h19 = exist (Inn B) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h24 %* exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) y' h26) (T_c (ba_conv Bp) B h2 (exist (fun a : sig_set (A_p T (Bc_p T Bp)) => Inn B a) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h24) h20) (exist (fun a : sig_set (A_p T (Bc_p T Bp)) => Inn B a) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) y' h26) h22))). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. red in h1. destruct h1 as [h1a [h1b h1c]]. assert (h30:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) x'). rewrite <- h1c. assumption. assert (h31:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) y'). rewrite <- h1c. assumption. pose proof (ba_p_subst_times _ _ (eq_sym h1c) _ _ h30 h31 h24 h26) as h29. rewrite <- h29 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. reflexivity. apply proj1_sig_injective. reflexivity. rewrite <- h28 at 1. assumption. destruct h1 as [h1a [h1b h1c]]. pose proof (ba_p_subst_one _ _ h1c) as h6. simpl. simpl in h5c. assert (h7:Inn (im_proj1_sig S') %1). apply Im_intro with (exist (Inn C) %1 (O_c (ba_conv Cp) C h3)). assumption. simpl. reflexivity. assert (h8:Inn (im_proj1_sig (im_proj1_sig S')) (proj1_sig (Bone_p _ (Bc_p _ Cp)))). apply Im_intro with %1. assumption. reflexivity. rewrite <- h4 in h8 at 1. rewrite <- h6 in h8 at 1. inversion h8 as [a h9 ? h10]. clear h1c. subst. destruct a as [a h11]. subst. simpl in h10. subst. inversion h9 as [a h12 ? h13]. subst. destruct a as [a h14]. simpl in h13. subst. assert (h15: exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) (proj1_sig %1) h11) h14 = exist (Inn B) %1 (O_c (ba_conv Bp) B h2)). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. reflexivity. rewrite <- h15 at 1. assumption. destruct h1 as [h1a [h1b h1c]]. pose proof (ba_p_subst_zero _ _ h1c) as h6. simpl. simpl in h5d. assert (h7:Inn (im_proj1_sig S') %0). apply Im_intro with (exist (Inn C) %0 (Z_c (ba_conv Cp) C h3)). assumption. simpl. reflexivity. assert (h8:Inn (im_proj1_sig (im_proj1_sig S')) (proj1_sig (Bzero_p _ (Bc_p _ Cp)))). apply Im_intro with %0. assumption. reflexivity. rewrite <- h4 in h8 at 1. rewrite <- h6 in h8 at 1. inversion h8 as [a h9 ? h10]. clear h1c. subst. destruct a as [a h11]. subst. simpl in h10. subst. inversion h9 as [a h12 ? h13]. subst. destruct a as [a h14]. simpl in h13. subst. assert (h15: exist (fun x : Btype (Bc (ba_conv Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) (proj1_sig %0) h11) h14 = exist (Inn B) %0 (Z_c (ba_conv Bp) B h2)). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. reflexivity. rewrite <- h15 at 1. assumption. intro x. destruct x as [x h6]. assert (h8:Inn (im_proj1_sig (im_proj1_sig S)) (proj1_sig (proj1_sig x))). destruct x as [x h8]. simpl. destruct x as [x h9]. simpl. assert (h10:Inn (im_proj1_sig S) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h9)). apply Im_intro with (exist (fun a : Btype (Bc (ba_conv Bp)) => Inn B a) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h9) h8). assumption. apply proj1_sig_injective. simpl. reflexivity. apply Im_intro with (exist _ _ h9). assumption. simpl. reflexivity. rewrite h4 in h8. inversion h8 as [x' h10 ? h11]. subst. destruct h10 as [x' h10]. subst. specialize (h5e (exist _ _ h10)). destruct x as [x h20], x' as [x' h21]. destruct x as [x h24], x' as [x' h25]. simpl in h8, h11. subst. assert (h12:Inn (im_proj1_sig S') (proj1_sig (Bcomp_sub (Subalg C h3) S' (exist (Inn S') (exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h25) h21) h10)))). eapply Im_intro. apply h5e. reflexivity. assert (h13:Inn (im_proj1_sig (im_proj1_sig S')) (proj1_sig (proj1_sig (Bcomp_sub (Subalg C h3) S' (exist (Inn S') (exist (fun x : Btype (Bc (ba_conv Cp)) => Inn C x) (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h25) h21) h10))))). eapply Im_intro. apply h12. reflexivity. rewrite <- h4 in h13. inversion h13 as [z h14 ? h15]. subst. destruct z as [z h16]. simpl in h15. subst. inversion h14 as [z h17 ? h18]. subst. destruct z as [z h19]. simpl in h18. subst. unfold Bcomp_sub, Subalg. simpl. unfold Bcomp_sub. simpl. unfold Btimes_sub in h17. simpl in h17. unfold Bcomp_sub in h17. simpl in h17. assert (h28: exist (fun x : sig_set (A_p T (Bc_p T Bp)) => Inn B x) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) (proj1_sig (%-exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) x' h25)) h16) h19 = exist (Inn B) (%-exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h24) (C_c (ba_conv Bp) B h2 (exist (fun a : sig_set (A_p T (Bc_p T Bp)) => Inn B a) (exist (fun x : T => Inn (A_p T (Bc_p T Bp)) x) x' h24) h20))). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. red in h1. destruct h1 as [h1a [h1b h1c]]. assert (h30:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) x'). rewrite <- h1c. assumption. pose proof (ba_p_subst_comp _ _ (eq_sym h1c) _ h30 h24) as h29. rewrite <- h29 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. reflexivity. rewrite <- h28 at 1. assumption. Qed. Lemma sig_set_im_proj1_sig_ba_ens_ba_conv_eq : forall {T:Type} (Bp:Bool_Alg_p T), sig_set (im_proj1_sig (ba_ens (ba_conv Bp))) = bt (ba_conv Bp). intros T Bp. unfold bt, ba_conv, ba_ens, bt. simpl. f_equal. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. subst. apply proj2_sig. red. intros x h1. apply Im_intro with (exist _ _ h1). constructor. simpl. reflexivity. Defined. Lemma ba_conv_und_subalg_im_proj1_Sig_ba_ens_ba_conv_eq : forall {T:Type} (Bp Cp:Bool_Alg_p T) (hin:Included (im_proj1_sig (ba_ens (ba_conv Bp))) (ba_p_ens Cp)) (hin':Included (im_proj1_sig (ba_ens (ba_conv Bp))) (im_proj1_sig (ba_ens (ba_conv Cp)))), ba_conv_und_subalg (im_proj1_sig (ba_ens (ba_conv Bp))) hin = transfer_dep (sig_set_im_proj1_sig_ba_ens_ba_conv_eq Cp) (im_proj2_sig (im_proj1_sig (ba_ens (ba_conv Bp))) hin'). intros T Bp Cp h1 h2. apply Extensionality_Ensembles. red. split. red. intros x h5. destruct h5 as [x h5]. subst. clear h5. destruct x as [x h5]. simpl. destruct h5 as [x h5]. subst. destruct x as [x h6]. simpl. rewrite <- (transfer_undoes_transfer_r (sig_set_im_proj1_sig_ba_ens_ba_conv_eq Cp)). rewrite <- transfer_in. unfold im_proj1_sig, im_proj2_sig. assert (h7:Inn (Im (ba_ens (ba_conv Bp)) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0))) ((proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h6)))). simpl. eapply Im_intro. apply h5. simpl. reflexivity. simpl in h7. apply Im_intro with (exist _ _ h7). constructor. simpl. apply proj1_sig_injective. simpl. pose (bt (ba_conv Cp)) as test. unfold bt in test. simpl in test. pose proof (sig_set_im_proj1_sig_ba_ens_ba_conv_eq Cp) as h8. unfold bt in h8. simpl in h8. assert (h9:im_proj1_sig (ba_ens (ba_conv Cp)) = A_p T (Bc_p T Cp)). unfold im_proj1_sig. unfold ba_conv, ba_ens, bt. simpl. fold (full_sig (A_p T (Bc_p T Cp))). rewrite <- im_full_sig_proj1_sig at 1. reflexivity. pose proof (transfer_r_sig_set_eq _ _ h9 h8 (exist (Inn (A_p T (Bc_p T Cp))) x (h1 x (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (ba_ens (ba_conv Bp)) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h6) h5 x eq_refl)))) as h10. pose proof (proof_irrelevance _ h8 (sig_set_im_proj1_sig_ba_ens_ba_conv_eq Cp)); subst. simpl in h10. simpl. rewrite h10 at 1. simpl. reflexivity. red. intros x h3. rewrite <- (transfer_undoes_transfer_r (sig_set_im_proj1_sig_ba_ens_ba_conv_eq Cp)) in h3. rewrite <- transfer_in in h3. inversion h3 as [a h4 ? h5]. subst. clear h4. destruct a as [a h4]. simpl in h5. destruct h4 as [a h4]. subst. destruct a as [a h6]. simpl in h5. pose proof (f_equal (@proj1_sig _ _) h5) as h7. clear h5. simpl in h7. unfold ba_conv_und_subalg, im_proj1_sig. unfold ba_conv_set, ba_conv_type. rewrite transfer_dep_eq_refl. assert (h8:Inn (Im (ba_ens (ba_conv Bp)) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0))) ((proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) a h6))). simpl. subst. eapply Im_intro. apply h4. simpl. reflexivity. simpl in h8. apply Im_intro with (exist _ _ h8). constructor. simpl. apply proj1_sig_injective. simpl. subst. destruct x as [x h9]. simpl. pose proof (sig_set_im_proj1_sig_ba_ens_ba_conv_eq Cp) as h11. unfold bt in h11. simpl in h11. assert (h10:im_proj1_sig (ba_ens (ba_conv Cp)) = A_p T (Bc_p T Cp)). unfold im_proj1_sig. unfold ba_conv, ba_ens, bt. simpl. fold (full_sig (A_p T (Bc_p T Cp))). rewrite <- im_full_sig_proj1_sig at 1. reflexivity. pose proof (transfer_r_sig_set_eq _ _ h10 h11 (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) x h9)) as h13. pose proof (proof_irrelevance _ h11 (sig_set_im_proj1_sig_ba_ens_ba_conv_eq Cp)); subst. rewrite h13 at 1; auto. Qed. Inductive subalg_p_bc_compat {T:Type} (Bp Cp:Bool_Alg_p T) (pf:Included (ba_p_ens Bp) (ba_p_ens Cp)) (pfac:alg_closed_p (ba_p_ens Bp) pf) : Prop := | subalg_p_bc_compat_intro : Bplus_p T (Bc_p T Bp) = restriction_sig2' (Bplus_p T (Bc_p T Cp)) (ba_p_ens Bp) pf (closed_fun2_plus _ _ pf pfac) -> Btimes_p T (Bc_p T Bp) = restriction_sig2' (Btimes_p T (Bc_p T Cp)) (ba_p_ens Bp) pf (closed_fun2_times _ _ pf pfac) -> proj1_sig (Bzero_p _ (Bc_p _ Bp)) = proj1_sig (Bzero_p _ (Bc_p _ Cp)) -> proj1_sig (Bone_p _ (Bc_p _ Bp)) = proj1_sig (Bone_p _ (Bc_p _ Cp)) -> Bcomp_p T (Bc_p T Bp) = restriction_sig' (Bcomp_p T (Bc_p T Cp)) (ba_p_ens Bp) pf (closed_fun_comp _ _ pf pfac) -> subalg_p_bc_compat Bp Cp pf pfac. Lemma bs_p_eq_transfer_dep_r_sig_set_eq_full_sig : forall {T:Type} (Bp Cp:Bool_Alg_p T) (pfi:Included (ba_p_ens Bp) (ba_p_ens Cp)) (pfac:alg_closed_p (ba_p_ens Bp) pfi) (pfeq: ba_p_ens Bp = ba_p_ens (Subalg_p Cp (ba_p_ens Bp) pfi pfac)), BS_p T (Bc_p T Bp) = transfer_dep_r (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T (Subalg_p Cp (ba_p_ens Bp) pfi pfac))) pfeq) (BS_p T (Bc_p T (Subalg_p Cp (ba_p_ens Bp) pfi pfac))). intros T Bp Cp h1 h2 h3. simpl. apply Extensionality_Ensembles. red. split. red. intros x h4. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Bp)) (ba_p_ens Bp) h3)). rewrite <- transfer_in_r. constructor. red. intros. apply in_bs_p. Qed. Lemma subalg_of_p_iff : forall {T:Type} (Bp Cp:Bool_Alg_p T), subalg_of_p Bp Cp <-> exists (pf:Included (ba_p_ens Bp) (ba_p_ens Cp)) (pfac:alg_closed_p (ba_p_ens Bp) pf), subalg_p_bc_compat _ _ pf pfac. intros T Bp Cp. split. intro h1. red in h1. destruct h1 as [h1 [h2 h3]]. exists h1. exists h2. constructor. apply functional_extensionality. intro x. apply functional_extensionality. intro y. unfold restriction_sig2'. apply proj1_sig_injective. simpl. destruct x as [x h4], y as [y h5]. simpl. assert (h7:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1 h2)) x). unfold ba_p_ens, Subalg_p. simpl. assumption. assert (h8:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1 h2)) y). unfold ba_p_ens, Subalg_p. simpl. assumption. pose proof (ba_p_subst_plus _ _ h3 _ _ h4 h5 h7 h8) as h6. simpl in h6. rewrite h6 at 1. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. apply functional_extensionality. intro x. apply functional_extensionality. intro y. unfold restriction_sig2'. apply proj1_sig_injective. simpl. destruct x as [x h4], y as [y h5]. simpl. assert (h7:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1 h2)) x). unfold ba_p_ens, Subalg_p. simpl. assumption. assert (h8:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1 h2)) y). unfold ba_p_ens, Subalg_p. simpl. assumption. pose proof (ba_p_subst_times _ _ h3 _ _ h4 h5 h7 h8) as h6. simpl in h6. rewrite h6 at 1. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. pose proof (ba_p_subst_zero _ _ h3) as h4. assumption. pose proof (ba_p_subst_one _ _ h3) as h4. assumption. apply functional_extensionality. intro x. unfold restriction_sig2'. apply proj1_sig_injective. simpl. destruct x as [x h4]. simpl. assert (h7:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1 h2)) x). unfold ba_p_ens, Subalg_p. simpl. assumption. pose proof (ba_p_subst_comp _ _ h3 _ h4 h7) as h6. simpl in h6. rewrite h6 at 1. f_equal. f_equal. apply proj1_sig_injective; auto. intro h1. destruct h1 as [h1 [h2 h3]]. destruct h3 as [h3a h3b h3c h3d h3e]. red. exists h1, h2. apply bc_inj_p. assert (h6: A_p T (Bc_p T Bp) = A_p T (Bc_p T (Subalg_p Cp (ba_p_ens Bp) h1 h2))). unfold Subalg_p. simpl. reflexivity. pose proof (bconst_ext_p (Bc_p T Bp) (Bc_p T (Subalg_p Cp (ba_p_ens Bp) h1 h2)) h6) as h5. apply h5; clear h5. unfold Subalg_p. simpl. erewrite bs_p_eq_transfer_dep_r_sig_set_eq_full_sig. reflexivity. rewrite h3a. apply functional_extensionality. intro x. apply functional_extensionality. intro y. simpl. unfold restriction_sig2'. apply proj1_sig_injective. f_equal. apply proj1_sig_injective. simpl. f_equal. rewrite transfer_dep_r_fun2_eq. rewrite (transfer_r_sig_set_eq _ _ h6 (sig_set_eq (A_p T (Bc_p T Bp)) (ba_p_ens Bp) h6)) at 1. simpl. do 2 rewrite (transfer_sig_set_eq _ _ h6 (sig_set_eq (A_p T (Bc_p T Bp)) (ba_p_ens Bp) h6)). simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. reflexivity. rewrite h3b. apply functional_extensionality. intro x. apply functional_extensionality. intro y. simpl. unfold restriction_sig2'. apply proj1_sig_injective. f_equal. apply proj1_sig_injective. simpl. f_equal. rewrite transfer_dep_r_fun2_eq. rewrite (transfer_r_sig_set_eq _ _ h6 (sig_set_eq (A_p T (Bc_p T Bp)) (ba_p_ens Bp) h6)) at 1. simpl. do 2 rewrite (transfer_sig_set_eq _ _ h6 (sig_set_eq (A_p T (Bc_p T Bp)) (ba_p_ens Bp) h6)). simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. reflexivity. simpl. symmetry. pose proof (proof_irrelevance _ (sig_set_eq (A_p T (Bc_p T Bp)) (ba_p_ens Bp) h6) (eq_refl _)) as h9. rewrite h9. rewrite transfer_dep_r_eq_refl. apply proj1_sig_injective. simpl. symmetry. assumption. pose proof (proof_irrelevance _ (sig_set_eq (A_p T (Bc_p T Bp)) (ba_p_ens Bp) h6) (eq_refl _)) as h9. simpl; rewrite h9, transfer_dep_r_eq_refl. apply proj1_sig_injective. simpl. auto. rewrite h3e. apply functional_extensionality. intro x. simpl. unfold restriction_sig'. apply proj1_sig_injective. f_equal. apply proj1_sig_injective. simpl. rewrite transfer_dep_r_fun1_eq. rewrite (transfer_r_sig_set_eq _ _ h6 (sig_set_eq (A_p T (Bc_p T Bp)) (ba_p_ens Bp) h6)) at 1. simpl. rewrite (transfer_sig_set_eq _ _ h6 (sig_set_eq (A_p T (Bc_p T Bp)) (ba_p_ens Bp) h6)). simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma subalg_of_p_eq : forall {T:Type} (Ap Bp:Bool_Alg_p T), subalg_of_p Bp Ap -> ba_p_ens Bp = ba_p_ens Ap -> Bp = Ap. intros T Ap Bp h1 h2. destruct h1 as [h1a [h1b h1c]]. pose proof (f_equal (@ba_p_ens _) h1c) as h1c'. pose proof (bs_p_eq_transfer_dep_r_sig_set_eq_full_sig _ _ h1a h1b h1c') as h3. apply bc_inj_p. unfold Subalg_p in h3. simpl in h3. assert (h5: BS_p T (Bc_p T Bp) = transfer_dep_r (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2) (BS_p T (Bc_p T Ap))). apply Extensionality_Ensembles. red. split. red. intros x h5. destruct x as [x h6]. unfold ba_p_ens in h2. pose proof h6 as h6'. rewrite h2 in h6'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2) (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h6)). rewrite <- transfer_in_r. rewrite (transfer_sig_set_eq _ _ h2 (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)). simpl. rewrite und_set_p. constructor. red. intros x h5. destruct x as [x h6]. rewrite und_set_p. constructor. pose proof (bconst_ext_p (Bc_p T Bp) (Bc_p T Ap) h2 h5) as h4. apply h4. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite <- transfer_fun2_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2) x) at 2. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2) y) at 2. rewrite transfer_fun2_r_eq'. do 2 rewrite transfer_undoes_transfer_r. rewrite (transfer_r_sig_set_eq _ _ h2 (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)). apply proj1_sig_injective. simpl. do 2 rewrite (transfer_sig_set_eq _ _ h2 (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)). destruct x as [x h6], y as [y h7]. simpl. assert (h9:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) x). rewrite <- h1c. assumption. assert (h10 : Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) y). rewrite <- h1c. assumption. pose proof (ba_p_subst_plus _ _ h1c _ _ h6 h7 h9 h10) as h8. rewrite h8 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite <- transfer_fun2_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2) x) at 2. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2) y) at 2. rewrite transfer_fun2_r_eq'. do 2 rewrite transfer_undoes_transfer_r. rewrite (transfer_r_sig_set_eq _ _ h2 (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)). apply proj1_sig_injective. simpl. do 2 rewrite (transfer_sig_set_eq _ _ h2 (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)). destruct x as [x h6], y as [y h7]. simpl. assert (h9:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) x). rewrite <- h1c. assumption. assert (h10 : Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) y). rewrite <- h1c. assumption. pose proof (ba_p_subst_times _ _ h1c _ _ h6 h7 h9 h10) as h8. rewrite h8 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. rewrite transfer_dep_r_transfer_r_compat. unfold id. pose (proof_irrelevance _ (f_equal (fun x:Type => x) (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)) (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)) as h6. rewrite h6 at 1. rewrite (transfer_r_sig_set_eq _ _ h2 (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)). apply proj1_sig_injective. simpl. pose proof (ba_p_subst_one _ _ h1c) as h7. rewrite h7 at 1. reflexivity. rewrite transfer_dep_r_transfer_r_compat. unfold id. pose (proof_irrelevance _ (f_equal (fun x:Type => x) (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)) (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)) as h6. rewrite h6 at 1. rewrite (transfer_r_sig_set_eq _ _ h2 (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)). apply proj1_sig_injective. simpl. pose proof (ba_p_subst_zero _ _ h1c) as h7. rewrite h7 at 1. reflexivity. apply functional_extensionality. intro x. rewrite <- transfer_fun_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2) x) at 2. rewrite transfer_fun_r_eq'. rewrite transfer_undoes_transfer_r. rewrite (transfer_r_sig_set_eq _ _ h2 (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)). apply proj1_sig_injective. simpl. rewrite (transfer_sig_set_eq _ _ h2 (sig_set_eq (A_p T (Bc_p T Bp)) (A_p T (Bc_p T Ap)) h2)). destruct x as [x h6]. simpl. assert (h9:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) x). rewrite <- h1c. assumption. pose proof (ba_p_subst_comp _ _ h1c _ h6 h9) as h8. rewrite h8 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. Qed. Lemma subalg_p_alg_closed_lift_sig_type : forall {T:Type} (Bp Cp:Bool_Alg_p T), subalg_of_p Bp Cp -> exists (pf:Included (ba_p_ens Bp) (ba_p_ens Cp)), alg_closed (im_proj2_sig _ pf) (B:=ba_conv Cp). intros T Bp Cp h1. red in h1. destruct h1 as [h1 [h2 h3]]. exists h1. assert (h5:Included (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1 h2)) (ba_p_ens Cp)). red. intros x h5. rewrite <- (in_ba_p_ens_eq _ _ _ h3) in h5 at 1. simpl in h5. apply h1; auto. pose proof (subsetT_eq_compat _ (fun D => Included (ba_p_ens D) (ba_p_ens Cp)) _ _ h1 h5 h3) as h6. pose proof h2 as h2'. rewrite alg_closed_p_iff in h2'. assumption. Qed. Lemma times_set_sub_compat_p : forall {T:Type} (Bp Cp:Bool_Alg_p T) (pfs:subalg_of_p Cp Bp) (D:Ensemble (btp Bp)) (D':Ensemble (btp Cp)) (pfd:Finite D) (pfd':Finite D'), im_proj1_sig D = im_proj1_sig D' -> proj1_sig (times_set_p D pfd)= proj1_sig (times_set_p D' pfd'). intros T Bp Cp h1 D D' h2 h3 h4. pose proof (finite_set_list_no_dup _ h3) as h5. destruct h5 as [l [h6 h7]]. subst. revert h4 h3 h7 h2. revert D h1. induction l as [|a l h8]. intros D h1 h2 h3 h4 h5. simpl in h2. unfold im_proj1_sig in h2. rewrite image_empty in h2 at 1. apply empty_image in h2. subst. simpl. do 2 rewrite times_set_empty_p'. destruct h1 as [h1a [h1b h1c]]. pose proof (ba_p_subst_one _ _ h1c) as h6. rewrite h6 at 1. reflexivity. intros D h1 h2 h3 h4 h5. pose proof (no_dup_cons_nin _ _ h4) as h6. apply no_dup_cons in h4. pose proof h1 as h1'. destruct h1' as [h1a [h1b h1c]]. destruct a as [a h9]. pose proof (h1a _ h9) as h10. assert (h11:im_proj1_sig (Subtract D (exist (Inn (ba_p_ens Bp)) a h10)) = im_proj1_sig (list_to_set l)). rewrite (im_proj1_sig_subtract_eq D (exist _ _ h10)) at 1. rewrite h2 at 1. simpl. rewrite (im_proj1_sig_add_eq (list_to_set l) (exist _ _ h9)) at 1. simpl. rewrite sub_add_same_nin. reflexivity. intro h11. destruct h11 as [a h11 ? h12]. rewrite <- list_to_set_in_iff in h11. contradict h6. destruct a as [a h13]. simpl in h12. clear h1c. subst. pose proof (proof_irrelevance _ h9 h13); subst; auto. pose proof (subtract_preserves_finite _ (exist _ _ h10) h5) as h12. specialize (h8 (Subtract D (exist _ _ h10)) h1 h11 (list_to_set_finite _) h4 h12). simpl. simpl in h3. rewrite (times_set_add_p' _ (list_to_set_finite _) (exist _ _ h9) h3) at 1. simpl. assert (h13: Included (Singleton (exist (Inn (ba_p_ens Bp)) a (h1a a h9))) D). red. intros x h14. destruct h14. rewrite <- in_im_proj1_sig_iff. simpl. assert (h15: @im_proj1_sig T (@Inn T (@ba_p_ens T Bp)) D = @im_proj1_sig T (A_p T (Bc_p T Bp)) D ). f_equal. rewrite h15 at 1. rewrite h2 at 1. apply Im_intro with (exist _ _ h9). rewrite <- list_to_set_in_iff. simpl. left. apply proj1_sig_injective. simpl. reflexivity. simpl. reflexivity. pose proof (decompose_setminus_inc D (Singleton (exist _ _ (h1a _ h9))) h13) as h14. pose proof (proof_irrelevance _ h10 (h1a _ h9)) as h16. rewrite <- h16 in h14. assert (h15:Finite (Singleton (exist (Inn (ba_p_ens Bp)) a h10))). apply Singleton_is_finite. unfold Subtract in h12. pose proof (subsetT_eq_compat _ (fun S=>Finite S) _ _ h5 (Union_preserves_Finite _ _ _ h15 h12) h14) as h17. dependent rewrite -> h17. rewrite (times_set_union_p' _ _ h15 h12 (Union_preserves_Finite {x | Inn (ba_p_ens Bp) x} (Singleton (exist (Inn (ba_p_ens Bp)) a h10)) (Setminus D (Singleton (exist (Inn (ba_p_ens Bp)) a h10))) h15 h12)) at 1. rewrite times_set_sing_p' at 1. f_equal. assert (h19: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h1a h1b)) a). rewrite <- h1c. assumption. assert (h20 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h1a h1b)) (proj1_sig (times_set_p (list_to_set l) (list_to_set_finite l)))). rewrite <- h1c. apply proj2_sig. pose proof (ba_p_subst_times _ _ h1c a (proj1_sig (times_set_p (list_to_set l) (list_to_set_finite l))) h9 (@proj2_sig _ _ _) h19 h20) as h18. rewrite <- unfold_sig in h18. rewrite h18 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. rewrite <- h8 at 1. f_equal. Qed. Lemma plus_set_sub_compat_p : forall {T:Type} (Bp Cp:Bool_Alg_p T) (pfs:subalg_of_p Cp Bp) (D:Ensemble (btp Bp)) (D':Ensemble (btp Cp)) (pfd:Finite D) (pfd':Finite D'), im_proj1_sig D = im_proj1_sig D' -> proj1_sig (plus_set_p D pfd)= proj1_sig (plus_set_p D' pfd'). intros T Bp Cp h1 D D' h2 h3 h4. pose proof (finite_set_list_no_dup _ h3) as h5. destruct h5 as [l [h6 h7]]. subst. revert h4 h3 h7 h2. revert D h1. induction l as [|a l h8]. intros D h1 h2 h3 h4 h5. simpl in h2. unfold im_proj1_sig in h2. rewrite image_empty in h2 at 1. apply empty_image in h2. subst. simpl. do 2 rewrite plus_set_empty_p'. destruct h1 as [h1a [h1b h1c]]. pose proof (ba_p_subst_zero _ _ h1c) as h6. rewrite h6 at 1. reflexivity. intros D h1 h2 h3 h4 h5. pose proof (no_dup_cons_nin _ _ h4) as h6. apply no_dup_cons in h4. pose proof h1 as h1'. destruct h1' as [h1a [h1b h1c]]. destruct a as [a h9]. pose proof (h1a _ h9) as h10. assert (h11:im_proj1_sig (Subtract D (exist (Inn (ba_p_ens Bp)) a h10)) = im_proj1_sig (list_to_set l)). rewrite (im_proj1_sig_subtract_eq D (exist _ _ h10)) at 1. rewrite h2 at 1. simpl. rewrite (im_proj1_sig_add_eq (list_to_set l) (exist _ _ h9)) at 1. simpl. rewrite sub_add_same_nin. reflexivity. intro h11. destruct h11 as [a h11 ? h12]. rewrite <- list_to_set_in_iff in h11. contradict h6. destruct a as [a h13]. simpl in h12. clear h1c. subst. pose proof (proof_irrelevance _ h9 h13); subst; auto. pose proof (subtract_preserves_finite _ (exist _ _ h10) h5) as h12. specialize (h8 (Subtract D (exist _ _ h10)) h1 h11 (list_to_set_finite _) h4 h12). simpl. simpl in h3. rewrite (plus_set_add_p' _ (list_to_set_finite _) (exist _ _ h9) h3) at 1. simpl. assert (h13: Included (Singleton (exist (Inn (ba_p_ens Bp)) a (h1a a h9))) D). red. intros x h14. destruct h14. rewrite <- in_im_proj1_sig_iff. simpl. assert (h15: @im_proj1_sig T (@Inn T (@ba_p_ens T Bp)) D = @im_proj1_sig T (A_p T (Bc_p T Bp)) D ). f_equal. rewrite h15 at 1. rewrite h2 at 1. apply Im_intro with (exist _ _ h9). rewrite <- list_to_set_in_iff. simpl. left. apply proj1_sig_injective. simpl. reflexivity. simpl. reflexivity. pose proof (decompose_setminus_inc D (Singleton (exist _ _ (h1a _ h9))) h13) as h14. pose proof (proof_irrelevance _ h10 (h1a _ h9)) as h16. rewrite <- h16 in h14. assert (h15:Finite (Singleton (exist (Inn (ba_p_ens Bp)) a h10))). apply Singleton_is_finite. unfold Subtract in h12. pose proof (subsetT_eq_compat _ (fun S=>Finite S) _ _ h5 (Union_preserves_Finite _ _ _ h15 h12) h14) as h17. dependent rewrite -> h17. rewrite (plus_set_union_p' _ _ h15 h12 (Union_preserves_Finite {x | Inn (ba_p_ens Bp) x} (Singleton (exist (Inn (ba_p_ens Bp)) a h10)) (Setminus D (Singleton (exist (Inn (ba_p_ens Bp)) a h10))) h15 h12)) at 1. rewrite plus_set_sing_p' at 1. f_equal. assert (h19: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h1a h1b)) a). rewrite <- h1c. assumption. assert (h20 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h1a h1b)) (proj1_sig (plus_set_p (list_to_set l) (list_to_set_finite l)))). rewrite <- h1c. apply proj2_sig. pose proof (ba_p_subst_plus _ _ h1c a (proj1_sig (plus_set_p (list_to_set l) (list_to_set_finite l))) h9 (@proj2_sig _ _ _) h19 h20) as h18. rewrite <- unfold_sig in h18. rewrite h18 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. rewrite <- h8 at 1. f_equal. Qed. (*This defintion is similar to [subalg_of_p] but uses a "forall" instead of "exists", and is useful in cases where you want to destruct [subalg_of_p] in an interactive function.*) Definition subalg_of_p' {T:Type} (Ap Bp:Bool_Alg_p T) := forall (pfin : Included (ba_p_ens Ap) (ba_p_ens Bp)) (pfac : alg_closed_p (ba_p_ens Ap) pfin), Ap = Subalg_p Bp (ba_p_ens Ap) pfin pfac. Lemma subalg_of_p_impl_prime : forall {T:Type} (Ap Bp:Bool_Alg_p T), subalg_of_p Ap Bp -> subalg_of_p' Ap Bp. intros T Ap Bp h1. red in h1. red. intros h2 h3. destruct h1 as [h1 [h4 h5]]. pose proof (proof_irrelevance _ h1 h2) as h6. assert (h7 : Subalg_p Bp (ba_p_ens Ap) h2 h3 = Subalg_p Bp (ba_p_ens Ap) h1 h4). clear h5. subst. pose proof (proof_irrelevance _ h3 h4); subst; auto. rewrite h7; auto. Qed. End Subalg_p. Section Closed_Finite_p. Lemma plus_list_closed_p : forall {T:Type} {Bp:Bool_Alg_p T} (S:Ensemble (btp Bp)) (l:list (btp Bp)), alg_closed_p _ (incl_ens_btp S) -> Included (list_to_set l) S -> Inn S (plus_list_p l). intros T Bp S l h1 h2. rewrite alg_closed_p_iff in h1. rewrite ba_conv_und_subalg_im_proj1_sig in h1. assert (h3:Included (list_to_set (ba_conv_list l)) (ba_conv_set S)). rewrite list_to_set_ba_conv_list. rewrite <- incl_ba_conv_set_iff. assumption. pose proof (plus_list_closed _ (ba_conv_set S) _ h1 h3) as h4. unfold ba_conv_set, ba_conv_list in h4. unfold ba_conv_type in h4. do 2 rewrite transfer_dep_eq_refl in h4. rewrite plus_list_p_eq. unfold ba_conv_list, ba_conv_type. rewrite transfer_dep_eq_refl. assumption. Qed. Lemma times_list_closed_p : forall {T:Type} {Bp:Bool_Alg_p T} (S:Ensemble (btp Bp)) (l:list (btp Bp)), alg_closed_p _ (incl_ens_btp S) -> Included (list_to_set l) S -> Inn S (times_list_p l). intros T Bp S l h1 h2. rewrite alg_closed_p_iff in h1. rewrite ba_conv_und_subalg_im_proj1_sig in h1. assert (h3:Included (list_to_set (ba_conv_list l)) (ba_conv_set S)). rewrite list_to_set_ba_conv_list. rewrite <- incl_ba_conv_set_iff. assumption. pose proof (times_list_closed _ (ba_conv_set S) _ h1 h3) as h4. unfold ba_conv_set, ba_conv_list in h4. unfold ba_conv_type in h4. do 2 rewrite transfer_dep_eq_refl in h4. rewrite times_list_p_eq. unfold ba_conv_list, ba_conv_type. rewrite transfer_dep_eq_refl. assumption. Qed. Lemma plus_set_closed_p : forall {T:Type} {Bp:Bool_Alg_p T} (R S:Ensemble (btp Bp)) (pf:Finite R), alg_closed_p _ (incl_ens_btp S) -> Included R S -> Inn S (plus_set_p R pf). intros T Bp R S h1 h2 h3. rewrite alg_closed_p_iff in h2. rewrite ba_conv_und_subalg_im_proj1_sig in h2. rewrite incl_ba_conv_set_iff in h3. pose proof (plus_set_closed _ (ba_conv_set R) (ba_conv_set S) h1 h2 h3) as h4. rewrite plus_set_p_eq. assumption. Qed. Lemma times_set_closed_p : forall {T:Type} {Bp:Bool_Alg_p T} (R S:Ensemble (btp Bp)) (pf:Finite R), alg_closed_p _ (incl_ens_btp S) -> Included R S -> Inn S (times_set_p R pf). intros T Bp R S h1 h2 h3. rewrite alg_closed_p_iff in h2. rewrite ba_conv_und_subalg_im_proj1_sig in h2. rewrite incl_ba_conv_set_iff in h3. pose proof (times_set_closed _ (ba_conv_set R) (ba_conv_set S) h1 h2 h3) as h4. rewrite times_set_p_eq. assumption. Qed. End Closed_Finite_p. Section Gen_Ens_p. Definition S_clo_cont_A_p {T:Type} {B:Bool_Alg_p T} (A:Ensemble T) (pfa:Included A (ba_p_ens B)) := [S:Ensemble T | Included A S /\ exists pfs : Included S (ba_p_ens B), alg_closed_p S pfs]. Lemma S_clo_cont_A_p_eq : forall {T:Type} {Bp:Bool_Alg_p T} (A:Ensemble T) (pfa:Included A (ba_p_ens Bp)), S_clo_cont_A_p A pfa = Im (S_clo_cont_A (B:=ba_conv Bp) (im_proj2_sig A pfa)) im_proj1_sig. intros T Bp A h1. apply Extensionality_Ensembles. red. split. red. intros S h2. destruct h2 as [h2]. destruct h2 as [h2 h3]. destruct h3 as [h3 h4]. apply Im_intro with (im_proj2_sig S h3). constructor. split. red. intros x h5. destruct h5 as [x h5]. subst. destruct x as [x h6]. simpl. clear h5. apply Im_intro with (exist _ _ (h2 _ h6)). constructor. f_equal. apply proof_irrelevance. rewrite alg_closed_p_iff in h4. assumption. rewrite <- im_proj1_sig_undoes_im_proj2_sig. reflexivity. red. intros S h2. destruct h2 as [S h2]. subst. destruct h2 as [h2]. destruct h2 as [h2 h3]. constructor. split. red in h2. red. intros x h4. specialize (h2 (exist _ _ (h1 _ h4))). assert (h5: Inn (im_proj2_sig A h1) (exist (Inn (ba_p_ens Bp)) x (h1 x h4)) ). apply Im_intro with (exist _ _ h4). constructor. apply proj1_sig_injective. reflexivity. specialize (h2 h5). apply Im_intro with (exist _ _ (h1 x h4)). assumption. reflexivity. assert (h4: Included (im_proj1_sig S) (ba_p_ens Bp)). red; intros x h4. destruct h4 as [x h4]. subst. apply proj2_sig. exists h4. rewrite alg_closed_p_iff. rewrite ba_conv_und_subalg_im_proj1_sig. assumption. Qed. Definition Gen_Ens_p {T:Type} {Bp:Bool_Alg_p T} (A:Ensemble T) (pf:Included A (ba_p_ens Bp)) : Ensemble T := FamilyIntersection (S_clo_cont_A_p A pf). Lemma incl_gen_ens_p : forall {T:Type} {Bp:Bool_Alg_p T} (A:Ensemble T) (pf:Included A (ba_p_ens Bp)), Included (Gen_Ens_p _ pf) (ba_p_ens Bp). intros T Bp A h1. red. intros x h2. destruct h2 as [x h2]. apply h2. constructor. split; auto. exists (inclusion_reflexive _). apply alg_closed_p_ba_p_ens_refl. Qed. Definition gen_ens_p_functional : forall {T:Type} {Bp:Bool_Alg_p T} (A A':Ensemble T) (pfa:Included A (ba_p_ens Bp)) (pfa':Included A' (ba_p_ens Bp)), A = A' -> Gen_Ens_p _ pfa = Gen_Ens_p _ pfa'. intros T Bp A A' h1 h2 h3. subst. assert (h4:h1 = h2). apply proof_irrelevance. subst. reflexivity. Qed. Lemma gen_ens_minimal_p : forall {T:Type} {Bp:Bool_Alg_p T} (A S:Ensemble T) (pf:Included A (ba_p_ens Bp)), let F := (S_clo_cont_A_p A pf) in Inn F S -> Included (Gen_Ens_p A pf) S. intros T Bp A S h1 F h0. unfold Gen_Ens_p. red. intros x h2. inversion h2 as [? h3]. subst. apply h3. assumption. Qed. Lemma gen_ens_minimal_p' : forall {T:Type} {Bp:Bool_Alg_p T} (A S:Ensemble T) (pfa:Included A (ba_p_ens Bp)) (pfs:Included S (ba_p_ens Bp)), alg_closed_p S pfs -> Included A S -> Included (Gen_Ens_p A pfa) S. intros T Bp A S h1 h2 h3 h4. apply gen_ens_minimal_p. constructor. split; auto. assert (h5:Included S (ba_p_ens Bp)). red; intros. apply h2; auto. exists h5. assert (h5 = h2). apply proof_irrelevance. subst. assumption. Qed. Lemma gen_ens_ba_conv_und_subalg : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)), Gen_Ens (ba_conv_und_subalg E pf) = ba_conv_und_subalg (Gen_Ens_p E pf) (incl_gen_ens_p E pf). intros T Bp E h1. apply Extensionality_Ensembles. red. split. red. intros x h2. pose proof (in_ba_conv_und_subalg_iff _ _ (incl_gen_ens_p _ h1) x) as h0. rewrite h0. destruct h2 as [x h2]. destruct x as [x h3]. simpl. simpl in h0. constructor. intros S h4. rewrite S_clo_cont_A_p_eq in h4. destruct h4 as [S h4]. subst. apply Im_intro with (exist _ _ h3). apply h2. assumption. simpl. reflexivity. red. intros x h2. pose proof (in_ba_conv_und_subalg_iff _ _ (incl_gen_ens_p _ h1)) as h0. rewrite h0 in h2. destruct x as [x hx]. simpl in h2. destruct h2 as [x h2]. constructor. intros S h3. specialize (h2 (im_proj1_sig S)). rewrite S_clo_cont_A_p_eq in h2. assert (h4:Inn (Im (S_clo_cont_A (im_proj2_sig E h1) (B:=ba_conv Bp)) im_proj1_sig) (im_proj1_sig S)). apply Im_intro with S. assumption. reflexivity. specialize (h2 h4). destruct h2 as [x h2]. subst. destruct x as [x h5]. simpl. simpl in hx. assert (hx = h5). apply proof_irrelevance. subst. assumption. Qed. Lemma closed_gen_ens_p : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)), alg_closed_p (Gen_Ens_p E pf) (incl_gen_ens_p _ pf). intros T Bp E h1. pose proof closed_gen_ens. pose proof (closed_gen_ens _ (ba_conv_und_subalg E h1)) as h2. rewrite gen_ens_ba_conv_und_subalg in h2. rewrite <- alg_closed_p_iff in h2. assumption. Qed. Lemma closed_gen_ens_p' : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pfi:Included E (ba_p_ens Bp)) (pfincl: Included (Gen_Ens_p _ pfi) (ba_p_ens Bp)), alg_closed_p (Gen_Ens_p E pfi) pfincl. intros T Bp E h1 h2. pose proof (closed_gen_ens_p _ h1) as h3. assert (h4:h2 = (incl_gen_ens_p _ h1)). apply proof_irrelevance. subst. assumption. Qed. Lemma gen_ens_p_eq : forall {T:Type} {Bp:Bool_Alg_p T} (A:Ensemble T) (pf:Included A (ba_p_ens Bp)), Gen_Ens_p A pf = im_proj1_sig (Gen_Ens (im_proj2_sig _ pf) (B:=ba_conv Bp)). intros T Bp A h1. apply Extensionality_Ensembles. red. split. red. intros x h2. apply Im_intro with (exist _ _ (incl_gen_ens_p _ _ _ h2)). destruct h2 as [x h2]. constructor. intros S h3. pose proof (h2 (im_proj1_sig S)) as h4. assert (h5:Inn (S_clo_cont_A_p A h1) (im_proj1_sig S)). rewrite S_clo_cont_A_p_eq. apply Im_intro with S. assumption. reflexivity. specialize (h4 h5). destruct h4 as [x h4]. subst. destruct x as [x h6]. simpl. simpl in h2. assert (h6 = incl_gen_ens_p A h1 x (family_intersection_intro T (S_clo_cont_A_p A h1) x h2)). apply proof_irrelevance. subst. assumption. simpl. reflexivity. red. intros x h2. destruct h2 as [x h2]. subst. destruct h2 as [x h2]. constructor. intros S h3. rewrite S_clo_cont_A_p_eq in h3. destruct h3 as [S h3]. subst. specialize (h2 _ h3). apply Im_intro with x; auto. Qed. Lemma gen_ens_p_eq' : forall {T:Type} {Bp:Bool_Alg_p T} (A:Ensemble T) (pf:Included A (ba_p_ens Bp)), Gen_Ens_p A pf = im_proj1_sig (Gen_Ens (im_proj2_sig _ (incl_gen_ens_p _ pf)) (B:=ba_conv Bp)). intros T Bp A h1. apply Extensionality_Ensembles. red. split. red. intros x h2. pose proof h2 as h2'. apply incl_gen_ens_p in h2'. apply Im_intro with (exist _ _ h2'). constructor. intros S h3. destruct h3 as [h3]. destruct h3 as [h3l h3r]. apply h3l. apply Im_intro with (exist _ _ h2). constructor. apply proj1_sig_injective; simpl; reflexivity. reflexivity. red. intros x h2. destruct h2 as [x h2]. subst. destruct h2 as [x h2]. destruct x as [x h3]. simpl. specialize (h2 (ba_conv_und_subalg _ (incl_gen_ens_p _ h1))). assert (h4: Inn (S_clo_cont_A (B:=ba_conv Bp) (im_proj2_sig (Gen_Ens_p A h1) (incl_gen_ens_p A h1))) (ba_conv_und_subalg (Gen_Ens_p A h1) (incl_gen_ens_p A h1))). constructor. split. apply inclusion_reflexive. rewrite <- alg_closed_p_iff. apply closed_gen_ens_p. specialize (h2 h4). inversion h2 as [x' h5 S h6]. clear h4. subst. apply exist_injective in h6. subst. apply proj2_sig. Qed. Definition Gen_p {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)) : Bool_Alg_p T := Subalg_p _ _ (incl_gen_ens_p _ pf) (closed_gen_ens_p _ pf). Lemma ba_p_ens_gen_p_eq : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)), ba_p_ens (Gen_p E pf) = Gen_Ens_p E pf. intros T Bp E h1. unfold Gen_p. unfold Gen_Ens_p. unfold ba_p_ens, Subalg_p. simpl. reflexivity. Qed. Lemma gen_p_subalg_of_p_compat : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)), subalg_of_p (Gen_p _ pf) Bp. intros T Bp E h1. red. assert (h2:Included (ba_p_ens (Gen_p E h1)) (ba_p_ens Bp)). rewrite ba_p_ens_gen_p_eq. apply incl_gen_ens_p. exists h2. assert (h3: alg_closed_p (ba_p_ens (Gen_p E h1)) h2). assert (h3:Included (Gen_Ens_p _ h1) (ba_p_ens Bp)). apply incl_gen_ens_p. pose proof (subsetT_eq_compat _ (fun S=>Included S (ba_p_ens Bp)) _ _ h2 h3 (ba_p_ens_gen_p_eq _ h1)) as h4. dependent rewrite -> h4. assert (h5:h3 = incl_gen_ens_p _ h1). apply proof_irrelevance. subst. apply closed_gen_ens_p. exists h3. unfold Gen_p. pose proof (ba_p_ens_subalg_p_compat _ _ (incl_gen_ens_p E h1) (closed_gen_ens_p E h1)) as h4. f_equal. apply subalg_functional_p. rewrite h4. reflexivity. Qed. Lemma gen_ens_includes_p : forall {T:Type} {Bp:Bool_Alg_p T} (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), Included S (Gen_Ens_p S pf). intros T Bp S h1. pose proof (gen_ens_includes (ba_conv Bp) (ba_conv_und_subalg S h1)) as h2. red in h2. red. intros x h3. specialize (h2 (exist _ _ (h1 _ h3))). assert (h4: Inn (ba_conv_und_subalg S h1) (exist (Inn (ba_p_ens Bp)) x (h1 x h3))). apply Im_intro with (exist _ _ h3). constructor. simpl. f_equal. specialize (h2 h4). clear h4. rewrite gen_ens_ba_conv_und_subalg in h2. inversion h2 as [x' h4 S' h5]. subst. apply exist_injective in h5. subst. apply proj2_sig. Qed. Lemma gen_ens_preserves_inclusion_p : forall {T:Type} {Bp:Bool_Alg_p T} (S U: Ensemble T) (pfs:Included S (ba_p_ens Bp)) (pfu:Included U (ba_p_ens Bp)), Included S U -> Included (Gen_Ens_p S pfs) (Gen_Ens_p U pfu). intros T Bp S U h1 h2 h3. pose proof (gen_ens_preserves_inclusion (ba_conv_und_subalg S h1) (ba_conv_und_subalg U h2)) as h4. assert (h5:Included (ba_conv_und_subalg S h1) (ba_conv_und_subalg U h2)). red. intros x h5. destruct x as [x h6]. inversion h5 as [x' h7 S' h8]. apply exist_injective in h8. subst. clear h7. destruct x' as [x' h7]. pose proof (h3 _ h7) as h8. simpl. apply Im_intro with (exist _ _ h8). constructor. apply proj1_sig_injective. simpl. reflexivity. specialize (h4 h5). do 2 rewrite gen_ens_p_eq. assert (h6: Included (im_proj1_sig (Gen_Ens (ba_conv_und_subalg S h1))) (im_proj1_sig (Gen_Ens (ba_conv_und_subalg U h2)))). red. intros x h6. destruct h6 as [x h6]. subst. apply h4 in h6. apply Im_intro with x. assumption. reflexivity. assumption. Qed. Lemma in_s_clo_cont_a_p_gen_ens_p: forall {T:Type} {Bp:Bool_Alg_p T} (A:Ensemble T) (pf:Included A (ba_p_ens Bp)), Inn (S_clo_cont_A_p A pf) (Gen_Ens_p A pf). intros T Bp A h1. constructor. split. apply gen_ens_includes_p. exists (incl_gen_ens_p _ _). apply closed_gen_ens_p. Qed. Definition two_p {T:Type} (Bp:Bool_Alg_p T) := (Couple (proj1_sig (Bzero_p T (Bc_p T Bp))) (proj1_sig (Bone_p T (Bc_p T Bp)))). Lemma two_p_eq : forall {T:Type} (Bp:Bool_Alg_p T), two_p Bp = im_proj1_sig (Couple (Bzero (Bc (ba_conv Bp))) (Bone (Bc (ba_conv Bp)))). intros T Bp. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1. apply Im_intro with %0. left. reflexivity. apply Im_intro with %1. right. reflexivity. red. intros x h1. destruct h1 as [x h1]. subst. destruct h1. left. right. Qed. Lemma incl_two_p : forall {T:Type} (Bp:Bool_Alg_p T), Included (two_p Bp) (ba_p_ens Bp). intros T Bp. red. intros x h1. destruct h1; apply proj2_sig. Qed. Lemma ba_conv_und_subalg_two_p : forall {T:Type} (Bp:Bool_Alg_p T), ba_conv_und_subalg (two_p Bp) (incl_two_p Bp) = Couple 0 1. intros T Bp. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct x as [x h2]. rewrite in_ba_conv_und_subalg_iff in h1. simpl in h1. destruct h1. simpl. rewrite (unfold_sig _ %0) at 1. assert (h3:h2 = proj2_sig %0). apply proof_irrelevance. subst. left. simpl. rewrite (unfold_sig _ %1) at 1. assert (h3:h2 = proj2_sig %1). apply proof_irrelevance. subst. right. red. intros x h1. rewrite in_ba_conv_und_subalg_iff. destruct h1. left. right. Qed. Lemma alg_closed_two_p : forall {T:Type} (Bp:Bool_Alg_p T), alg_closed_p (two_p Bp) (incl_two_p Bp). intros T Bp. pose proof (alg_closed_two (ba_conv Bp)) as h1. rewrite alg_closed_p_iff. rewrite ba_conv_und_subalg_two_p. assumption. Qed. Definition two_bap {T:Type} (Bp:Bool_Alg_p T) := Subalg_p Bp _ (incl_two_p Bp) (alg_closed_two_p Bp). Lemma ba_p_ens_two_bap : forall {T:Type} (Bp:Bool_Alg_p T), ba_p_ens (two_bap Bp) = two_p Bp. unfold two_p, two_bap; intros; auto. Qed. Lemma ba_p_ens_two_bap' : forall {T:Type} {Bp: Bool_Alg_p T}, ba_p_ens (two_bap Bp) = im_proj1_sig (Couple (Bzero_p T (Bc_p T Bp)) (Bone_p T (Bc_p T Bp))). intros T Bp. rewrite ba_p_ens_two_bap. unfold two_p, im_proj1_sig. rewrite im_couple. reflexivity. Qed. Lemma zero_in_two_bap : forall {T:Type} (Bp : Bool_Alg_p T), Inn (ba_p_ens (two_bap Bp)) (proj1_sig (Bzero_p T (Bc_p T Bp))). intros. rewrite ba_p_ens_two_bap. left; auto. Qed. Lemma one_in_two_bap : forall {T:Type} (Bp : Bool_Alg_p T), Inn (ba_p_ens (two_bap Bp)) (proj1_sig (Bone_p T (Bc_p T Bp))). intros. rewrite ba_p_ens_two_bap. right; auto. Qed. Lemma subalg_of_p_two_bap : forall {T:Type} (Bp:Bool_Alg_p T), subalg_of_p (two_bap Bp) Bp. unfold two_bap; intros; apply subalg_of_p_compat. Qed. Lemma two_bap_dec : forall {T:Type} {Bp:Bool_Alg_p T} (x:btp (two_bap Bp)), {x = %0} + {x = %1}. intros T Bp x. destruct (ba_eq_dec_p x (bzero_p (two_bap Bp))) as [h1 | h1]. left; auto. right. destruct x as [x hx]. destruct hx. apply proj1_sig_injective; simpl. unfold bzero_p in h1. simpl in h1. contradict h1. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. Qed. Lemma gen_ens_closed_eq_p : forall {T:Type} {Bp:Bool_Alg_p T} (S:Ensemble T) (pf:Included S (ba_p_ens Bp)), alg_closed_p S pf -> Gen_Ens_p S pf = S. intros T Bp S h1 h2. rewrite alg_closed_p_iff in h2. apply gen_ens_closed_eq in h2. rewrite gen_ens_p_eq. rewrite h2 at 1. pose proof (im_proj1_sig_undoes_im_proj2_sig _ _ h1) as h3. rewrite <- h3 at 1. reflexivity. Qed. Lemma gen_ens_empty_p : forall {T:Type} (Bp:Bool_Alg_p T), Gen_Ens_p (Bp:=Bp) (Empty_set T) (incl_empty _) = two_p Bp. intros T Bp. pose proof (gen_ens_empty (ba_conv Bp)) as h1. rewrite gen_ens_p_eq. pose proof (gen_ens_ba_conv_und_subalg _ (incl_empty (ba_p_ens Bp))) as h2. rewrite h2 at 1. pose proof (im_proj1_sig_undoes_im_proj2_sig _ _ (incl_gen_ens_p (Empty_set T) (incl_empty (ba_p_ens Bp)))) as h3. rewrite <- h3 at 1. rewrite gen_ens_p_eq. assert (h4:Empty_set _ = im_proj2_sig (Empty_set T) (incl_empty (ba_p_ens Bp))). apply Extensionality_Ensembles; red; split; auto with sets. red. intros x h4. destruct h4 as [x h4]. destruct x. contradiction. rewrite <- h4. rewrite h1 at 1. rewrite two_p_eq. reflexivity. Qed. Theorem gen_empty_p : forall {T:Type} (Bp:Bool_Alg_p T), Gen_p (Empty_set T) (incl_empty (ba_p_ens Bp)) = two_bap Bp. intros T Bp. unfold Gen_p, two_bap. f_equal. pose proof (gen_ens_empty_p Bp) as h1. pose proof (subsetT_eq_compat _ (fun S=>Included S (ba_p_ens Bp)) _ _ (incl_gen_ens_p (Empty_set T) (incl_empty (ba_p_ens Bp))) (incl_two_p Bp) h1) as h2. generalize (closed_gen_ens_p (Empty_set T) (incl_empty (ba_p_ens Bp))). dependent rewrite -> h2. intro h3. f_equal. apply proof_irrelevance. Qed. Lemma two_incl_gen_ens_p : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)), Included (two_p Bp) (Gen_Ens_p E pf). intros T Bp E h0. assert (h1:Included (Empty_set _) E). auto with sets. pose proof (gen_ens_preserves_inclusion_p _ _ (incl_empty (ba_p_ens Bp)) h0 h1 ) as h2. rewrite gen_ens_empty_p in h2 at 1. assumption. Qed. Lemma not_incl_gen_ens_neq_0_p : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (x:btp Bp), ~Inn (Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)) (proj1_sig x) -> x <> %0. intros T Bp E x h1 h2. subst. pose proof (two_incl_gen_ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)) as h2. unfold two_p in h2. auto with sets. Qed. Lemma not_incl_gen_ens_neq_0_p' : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)) (x:T), ~Inn (Gen_Ens_p E pf) x-> x <> proj1_sig (Bzero_p T (Bc_p T Bp)). intros T Bp E he x h1 h2. subst. pose proof (two_incl_gen_ens_p E he) as h2. unfold two_p in h2. auto with sets. Qed. Lemma not_incl_gen_ens_neq_1_p : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (x:btp Bp), ~Inn (Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)) (proj1_sig x) -> x <> %1. intros T Bp E x h1 h2. subst. pose proof (two_incl_gen_ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)) as h2. unfold two_p in h2. auto with sets. Qed. Lemma not_incl_gen_ens_neq_1_p' : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)) (x:T), ~Inn (Gen_Ens_p E pf) x-> x <> proj1_sig (Bone_p T (Bc_p T Bp)). intros T Bp E he x h1 h2. subst. pose proof (two_incl_gen_ens_p E he) as h2. unfold two_p in h2. auto with sets. Qed. Lemma inhabited_gen_ens_p : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)), Inhabited (Gen_Ens_p E pf). intros T Bp E h0. pose proof (two_incl_gen_ens_p E h0) as h1. assert (h2:Inn (two_p Bp) (proj1_sig (Bzero_p T (Bc_p T Bp)))). constructor. apply Inhabited_intro with (proj1_sig (Bzero_p T (Bc_p T Bp))). auto with sets. Qed. Lemma bt_subalg_ba_conv_und_subalg_eq : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pfi:Included E (ba_p_ens Bp)) (pfc:alg_closed (ba_conv_und_subalg E pfi)), bt (Subalg (ba_conv_und_subalg E pfi) pfc) = sig_set (im_proj2_sig E pfi). intros T Bp E h1 h2. unfold bt, Subalg. simpl. unfold SubBtype, sig_set. unfold ba_conv. simpl. assert (h3:forall a:sig_set (ba_p_ens Bp), Inn (ba_conv_und_subalg E h1) a <-> Inn (im_proj2_sig E h1) a). intro a. destruct a. tauto. apply iff_pred_sig_eq; auto. rewrite iff_pred_iff; tauto. Qed. Lemma im_proj1_sig_gen_ens_im_proj2_sig_im_proj2_sig : forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)) (pf':Included (im_proj2_sig E pf) (full_sig (ba_p_ens Bp))), im_proj1_sig (Gen_Ens (im_proj2_sig (im_proj2_sig E pf) pf') (B:=@Subalg (@ba_conv T Bp) (ba_ens (@ba_conv T Bp)) (alg_closed_ba_ens (@ba_conv T Bp)))) = Gen_Ens (im_proj2_sig E pf) (B:=ba_conv Bp). intros T Bp E h1 h2. apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [x h3]. subst. destruct h3 as [x h3]. destruct x as [x h4]. destruct x as [x h5]. simpl. constructor. intros S h6. assert (h7:forall x:Btype (Bc (ba_conv Bp)), Inn (ba_ens (ba_conv Bp)) x). intro a. destruct a as [a h7]. constructor. pose (Im S (fun x => (exist _ _ (h7 x)))) as S'. specialize (h3 S'). assert (h8: Inn (S_clo_cont_A (im_proj2_sig (im_proj2_sig E h1) h2) (B:=@Subalg (@ba_conv T Bp) (ba_ens (@ba_conv T Bp)) (alg_closed_ba_ens (@ba_conv T Bp)))) S'). destruct h6 as [h6]. destruct h6 as [h6a h6b]. constructor. split. red. intros a h8. destruct h8 as [a h8]. subst. clear h8. destruct a as [a h8]. destruct h8 as [a h8]. subst. simpl. destruct a as [a h9]. simpl. assert (h10:Inn (im_proj2_sig E h1) (exist _ _ (h1 _ h9))). eapply Im_intro. apply h8. apply proj1_sig_injective. simpl. reflexivity. apply h6a in h10. unfold S'. apply Im_intro with (exist (Inn (ba_p_ens Bp)) a (h1 a h9)). assumption. apply proj1_sig_injective. simpl. apply proj1_sig_injective; auto. rewrite <- alg_closed_improper_subalg. apply h6b. apply Extensionality_Ensembles. red. split. red. intros a h8. unfold S', im_proj1_sig. rewrite im_im. apply Im_intro with a; auto. red. intros a h8. destruct h8 as [a h8]. subst. destruct h8 as [a h8]. subst. simpl. assumption. specialize (h3 h8). inversion h3 as [a h9 ? h10]. subst. apply exist_injective in h10. subst. assumption. red. intros x h3. assert (h7:forall x:Btype (Bc (ba_conv Bp)), Inn (ba_ens (ba_conv Bp)) x). intro a. destruct a as [a h7]. constructor. pose (exist _ _ (h7 x)) as x'. assert (h8:Inn (Gen_Ens (im_proj2_sig (im_proj2_sig E h1) h2) (B:= (@Subalg (@ba_conv T Bp) (ba_ens (@ba_conv T Bp)) (alg_closed_ba_ens (@ba_conv T Bp))))) x'). destruct h3 as [x h3]. constructor. intros S h4. specialize (h3 (im_proj1_sig S)). assert (h8: Inn (S_clo_cont_A (im_proj2_sig E h1) (B:=ba_conv Bp)) (im_proj1_sig S)). destruct h4 as [h4]. destruct h4 as [h4 h5]. constructor. split. red. intros a h6. destruct h6 as [a h6]. subst. clear h6. destruct a as [a h6]. simpl. assert (h8:Inn (im_proj2_sig E h1) (exist _ _ (h1 _ h6))). assert (h9:Inn (full_sig E) (exist _ _ h6)). constructor. eapply Im_intro. apply h9. apply proj1_sig_injective. simpl. reflexivity. assert (h9: Inn (im_proj2_sig (im_proj2_sig E h1) h2) (exist _ _ (h2 _ h8))). assert (h9:Inn (full_sig (im_proj2_sig E h1)) (exist _ _ h8)). constructor. eapply Im_intro. apply h9. apply proj1_sig_injective. simpl. reflexivity. apply h4 in h9. apply Im_intro with (exist (Inn (full_sig (ba_p_ens Bp))) (exist (Inn (ba_p_ens Bp)) a (h1 a h6)) (h2 (exist (Inn (ba_p_ens Bp)) a (h1 a h6)) h8)). assumption. apply proj1_sig_injective. simpl. reflexivity. rewrite alg_closed_improper_subalg. apply h5. reflexivity. specialize (h3 h8). destruct h3 as [x h3]. assert (h9:x = x'). unfold x'. apply proj1_sig_injective. simpl. rewrite H. reflexivity. rewrite <- h9. assumption. apply Im_intro with x'. assumption. unfold x'. simpl. reflexivity. Qed. Lemma gen_ens_p_subalg_of_p : forall {T:Type} {Bp Cp:Bool_Alg_p T}, subalg_of_p Bp Cp -> forall (E:Ensemble T) (pfb:Included E (ba_p_ens Bp)) (pfc:Included E (ba_p_ens Cp)), Gen_Ens_p E pfb = Gen_Ens_p E pfc. intros T Bp Cp h1 E h2 h3. rewrite (gen_ens_p_eq E h3). pose proof h1 as h0. rewrite subalg_of_p_iff in h1. destruct h1 as [h1a [h1b h1c]]. pose proof h1b as h4. rewrite alg_closed_p_iff in h4. assert (h7:Included (im_proj2_sig _ h3) (im_proj2_sig _ h1a)). red. intros x h7. destruct h7 as [x h7]. subst. clear h7. destruct x as [x h7]. simpl. apply Im_intro with (exist _ _ (h2 _ h7)). constructor. apply proj1_sig_injective. simpl. reflexivity. pose (im_proj2_sig (im_proj2_sig _ h3) h7) as E'. pose proof (gen_ens_subalg_eq h4 E') as h5. unfold E' in h5. rewrite <- im_proj1_sig_undoes_im_proj2_sig in h5 at 1. rewrite <- h5 at 1. assert (h8:Included (im_proj2_sig E h2) (full_sig (ba_p_ens Bp))). red; intros; constructor. assert (he:im_proj1_sig (im_proj1_sig (Gen_Ens (im_proj2_sig (im_proj2_sig E h3) h7) (B:=@Subalg (@ba_conv T Cp) (@ba_conv_und_subalg T Cp (@ba_p_ens T Bp) h1a) h4))) = im_proj1_sig ( im_proj1_sig (Gen_Ens (im_proj2_sig (im_proj2_sig E h2)h8) (B:=Subalg _ (alg_closed_ba_ens (ba_conv Bp)))))). apply Extensionality_Ensembles. red. split. red. intros x h9. destruct h9 as [x h9]. subst. destruct h9 as [x h9]. subst. destruct x as [x h10]. simpl. destruct h10 as [x h10]. subst. inversion h9 as [a h11 q]. subst. simpl. clear h9. apply Im_intro with x. apply Im_intro with (exist _ _ h10). constructor. intros S h13. simpl in S. unfold SubBtype in S. simpl in h11. unfold SubBtype in h11. simpl in h11. pose proof h1a as h1'. rewrite ba_p_ens_eq in h1'. assert (h14:forall d: {a : Btype (Bc (ba_conv Bp)) | Inn (ba_ens (ba_conv Bp)) a}, Inn (ba_p_ens Cp) (proj1_sig (proj1_sig d))). intro d. destruct d as [d h15]. destruct d as [d h16]. pose proof h16 as h16'. apply h1a in h16'. simpl. assumption. assert (h15:forall d:{a:Btype (Bc (ba_conv Bp)) | Inn (ba_ens (ba_conv Bp)) a}, Inn (ba_conv_und_subalg (ba_p_ens Bp) h1a) (exist _ _ (h14 d))). intro d. unfold ba_conv_und_subalg. unfold ba_conv_set, ba_conv_type. rewrite transfer_dep_eq_refl. destruct d as [d h15]. apply Im_intro with d. constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. pose (Im S (fun d => (exist _ _ (h15 d)))) as S'. specialize (h11 S'). assert (h16:Inn (S_clo_cont_A (B:=@Subalg (@ba_conv T Cp) (@ba_conv_und_subalg T Cp (@ba_p_ens T Bp) h1a) h4) (im_proj2_sig (im_proj2_sig E h3) h7)) S'). destruct h13 as [h13]. destruct h13 as [h13a h13b]. constructor. split. red. intros a h16. destruct h16 as [a h16]. subst. clear h16. destruct a as [a h16]. simpl. destruct h16 as [a h16]. subst. destruct a as [a h17]. simpl. pose (exist _ _ (h2 _ h17)) as a'. assert (h18:Inn (im_proj2_sig E h2) a'). apply Im_intro with (exist _ _ h17). constructor. unfold a'. apply proj1_sig_injective. simpl. reflexivity. pose (exist _ _ (h8 _ h18)) as a''. specialize (h13a a''). assert (h19: Inn (im_proj2_sig (im_proj2_sig E h2) h8) a'' ). unfold a''. apply Im_intro with (exist _ _ h18). constructor. apply proj1_sig_injective. simpl. reflexivity. specialize (h13a h19). unfold S'. apply Im_intro with a''. assumption. apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. reflexivity. assert (h20 :im_proj1_sig (im_proj1_sig S) = im_proj1_sig (im_proj1_sig S')). apply Extensionality_Ensembles. red. split. red. intros t h21. destruct h21 as [t h21]. subst. destruct h21 as [t h21]. subst. destruct t as [t h22]. destruct t as [t h23]. simpl. apply Im_intro with (exist _ _ (h1a _ h23)). unfold S'. unfold im_proj1_sig. rewrite im_im. simpl. apply Im_intro with (exist _ _ h22). assumption. apply proj1_sig_injective. simpl. reflexivity. simpl. reflexivity. red. intros t h21. destruct h21 as [t h21]. subst. destruct h21 as [t h21]. subst. destruct t as [t h22]. destruct h22 as [t h22]. subst. simpl. destruct t as [t h23]. simpl. simpl in h21. apply Im_intro with (exist _ _ h23). apply Im_intro with (exist _ _ h22). inversion h21 as [a h24 ? h25]. subst. apply exist_injective in h25. apply exist_injective in h25. subst. assert (h25:a = (exist (Inn (full_sig (ba_p_ens Bp))) (exist (fun x0 : T => Inn (ba_p_ens Bp) x0) (proj1_sig (proj1_sig a)) h23) h22)). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. reflexivity. rewrite <- h25 at 1. assumption. simpl. apply proj1_sig_injective. simpl. reflexivity. simpl. reflexivity. erewrite <- subalg_of_p_alg_closed_subalg_compat. apply h13b. assumption. assumption. specialize (h11 h16). inversion h11 as [a h17 ? h18]. subst. clear h11. apply exist_injective in h18. apply exist_injective in h18. destruct x as [x h19]. simpl in h18. subst. assert (h20:a = exist (Inn (full_sig (ba_p_ens Bp))) (exist (fun x : T => Inn (ba_p_ens Bp) x) (proj1_sig (proj1_sig a)) h19) h10). apply proj1_sig_injective; apply proj1_sig_injective; auto. rewrite <- h20 at 1. assumption. simpl. reflexivity. reflexivity. red. intros x h9. destruct h9 as [x h9]. subst. destruct h9 as [x h9]. subst. destruct h9 as [x h9]. destruct x as [x h10]. destruct x as [x h11]. simpl. apply Im_intro with (exist _ _ (h1a _ h11)). assert (h12:Inn (ba_conv_und_subalg (ba_p_ens Bp) h1a) (exist _ _ (h1a _ h11))). apply Im_intro with (exist _ _ h11). constructor. apply proj1_sig_injective. simpl. reflexivity. assert (h14:Inn (Gen_Ens (im_proj2_sig (im_proj2_sig E h3) h7) (B:=(@Subalg (@ba_conv T Cp) (@ba_conv_und_subalg T Cp (@ba_p_ens T Bp) h1a) h4))) (exist _ _ h12)). constructor. intros S h13. assert (h14:forall x:(SubBtype (ba_conv Cp) (ba_conv_und_subalg (ba_p_ens Bp) h1a)), Inn (ba_p_ens Bp) (proj1_sig (proj1_sig x))). intro a. destruct a as [a h14]. simpl. destruct h14 as [a h14]. subst. clear h14. destruct a as [a h14]. simpl. assumption. assert (h15:forall x:(SubBtype (ba_conv Cp) (ba_conv_und_subalg (ba_p_ens Bp) h1a)), Inn (ba_ens (ba_conv Bp)) (exist _ _ (h14 x))). intro a. destruct a as [a h15]. simpl. destruct h15 as [a h15]. subst. destruct a as [a h16]. simpl. pose proof h16 as h16'. rewrite ba_p_ens_eq in h16'. destruct h16' as [a h16']. subst. assert (h17:a = exist (Inn (ba_p_ens Bp)) (proj1_sig a) (h14 (exist (fun a0 : sig_set (A_p T (Bc_p T Cp)) => Inn (ba_conv_und_subalg (ba_p_ens Bp) h1a) a0) (exist (Inn (A_p T (Bc_p T Cp))) (proj1_sig a) (h1a (proj1_sig a) h16)) (Im_intro (sig_set (ba_p_ens Bp)) {x | Inn (A_p T (Bc_p T Cp)) x} (full_sig (ba_p_ens Bp)) (fun x0 : sig_set (ba_p_ens Bp) => exist (Inn (A_p T (Bc_p T Cp))) (proj1_sig x0) (h1a (proj1_sig x0) (proj2_sig x0))) (exist (fun x0 : T => Inn (ba_p_ens Bp) x0) (proj1_sig a) h16) h15 (exist (Inn (A_p T (Bc_p T Cp))) (proj1_sig a) (h1a (proj1_sig a) h16)) eq_refl)))). apply proj1_sig_injective. simpl. reflexivity. rewrite <- h17 at 1. assumption. pose (Im S (fun a => (exist _ _ (h15 a)))) as S'. specialize (h9 S'). assert (h16: Inn (S_clo_cont_A (im_proj2_sig (im_proj2_sig E h2) h8) (B:=(@Subalg (@ba_conv T Bp) (ba_ens (@ba_conv T Bp)) (alg_closed_ba_ens (@ba_conv T Bp))))) S'). destruct h13 as [h13]. destruct h13 as [h13a h13b]. constructor. split. red. intros a h16. destruct h16 as [a h16]. subst. clear h16. destruct a as [a h16]. simpl. destruct h16 as [a h16]. subst. destruct a as [a h17]. simpl. pose (exist _ _ (h3 _ h17)) as a'. assert (h18:Inn (im_proj2_sig E h3) a'). eapply Im_intro. apply h16. unfold a'. apply proj1_sig_injective. simpl. reflexivity. pose (exist _ _ (h7 _ h18)) as a''. assert (h19:Inn (im_proj2_sig (im_proj2_sig E h3) h7) a''). eapply Im_intro. assert (h19:Inn (full_sig (im_proj2_sig E h3)) (exist _ _ h18)). constructor. apply h19. unfold a''. apply proj1_sig_injective. simpl. reflexivity. apply h13a in h19. unfold S'. apply Im_intro with a''; auto. apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. reflexivity. erewrite subalg_of_p_alg_closed_subalg_compat. apply h13b. assumption. apply Extensionality_Ensembles. red. split. red. intros a h16. destruct h16 as [a h16]. subst. destruct h16 as [a h16]. subst. destruct h16 as [a h16]. subst. simpl. destruct a as [a h17]. simpl. destruct h17 as [a h17]. subst. destruct a as [a h18]. simpl. simpl in h16. apply Im_intro with (exist _ _ (h1a _ h18)). assert (h19:Inn (ba_conv_und_subalg (ba_p_ens Bp) h1a) (exist _ _ (h1a _ h18))). eapply Im_intro. apply h17. apply proj1_sig_injective. simpl. reflexivity. apply Im_intro with (exist _ _ h19). assert (h20: exist (fun a0 : sig_set (A_p T (Bc_p T Cp)) => Inn (ba_conv_und_subalg (ba_p_ens Bp) h1a) a0) (exist (Inn (A_p T (Bc_p T Cp))) a (h1a a h18)) (Im_intro (sig_set (ba_p_ens Bp)) {x | Inn (A_p T (Bc_p T Cp)) x} (full_sig (ba_p_ens Bp)) (fun x0 : sig_set (ba_p_ens Bp) => exist (Inn (A_p T (Bc_p T Cp))) (proj1_sig x0) (h1a (proj1_sig x0) (proj2_sig x0))) (exist (fun x0 : T => Inn (ba_p_ens Bp) x0) a h18) h17 (exist (Inn (A_p T (Bc_p T Cp))) a (h1a a h18)) eq_refl) = exist (Inn (ba_conv_und_subalg (ba_p_ens Bp) h1a)) (exist (Inn (ba_p_ens Cp)) a (h1a a h18)) h19). apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. reflexivity. rewrite <- h20 at 1. assumption. apply proj1_sig_injective. simpl. reflexivity. simpl. reflexivity. red. intros a h16. destruct h16 as [a h16]. subst. destruct h16 as [a h16]. subst. destruct a as [a h17]. simpl. destruct h17 as [a h17]. subst. destruct a as [a h18]. simpl. simpl in h16. unfold S'. unfold im_proj1_sig. rewrite im_im. rewrite im_im. simpl. apply Im_intro with (exist (fun x0 : sig_set (A_p T (Bc_p T Cp)) => Inn (ba_conv_und_subalg (ba_p_ens Bp) h1a) x0) (exist (Inn (A_p T (Bc_p T Cp))) a (h1a a h18)) (Im_intro (sig_set (ba_p_ens Bp)) {x | Inn (A_p T (Bc_p T Cp)) x} (full_sig (ba_p_ens Bp)) (fun x0 : sig_set (ba_p_ens Bp) => exist (Inn (A_p T (Bc_p T Cp))) (proj1_sig x0) (h1a (proj1_sig x0) (proj2_sig x0))) (exist (fun x0 : T => Inn (ba_p_ens Bp) x0) a h18) h17 (exist (Inn (A_p T (Bc_p T Cp))) a (h1a a h18)) eq_refl)). assumption. simpl. reflexivity. specialize (h9 h16). inversion h9 as [a h17 ? h18]. subst. clear h9. apply exist_injective in h18. apply exist_injective in h18. subst. assert (h21:a = (exist (Inn (ba_conv_und_subalg (ba_p_ens Bp) h1a)) (exist (Inn (ba_p_ens Cp)) (proj1_sig (proj1_sig a)) (h1a (proj1_sig (proj1_sig a)) h11)) h12)). destruct a as [a h19]. simpl. apply proj1_sig_injective. simpl. destruct a as [a h20]. apply proj1_sig_injective. simpl. reflexivity. rewrite <- h21 at 1. assumption. apply Im_intro with (exist (Inn (ba_conv_und_subalg (ba_p_ens Bp) h1a)) (exist (Inn (ba_p_ens Cp)) x (h1a x h11)) h12). assumption. apply proj1_sig_injective. simpl. reflexivity. simpl. reflexivity. rewrite he at 1. rewrite gen_ens_p_eq. f_equal. symmetry. apply im_proj1_sig_gen_ens_im_proj2_sig_im_proj2_sig. Qed. Lemma gen_p_subalg_of_p : forall {T:Type} {Bp Cp:Bool_Alg_p T}, subalg_of_p Bp Cp -> forall (E:Ensemble T) (pfb:Included E (ba_p_ens Bp)) (pfc:Included E (ba_p_ens Cp)), Gen_p E pfb = Gen_p E pfc. intros T Bp Cp h1 E h2 h3. pose proof h1 as h1'. destruct h1' as [h1a [h1b h1c]]. apply bc_inj_p. pose proof (gen_ens_p_subalg_of_p h1 E h2 h3) as h4. do 2 rewrite <- ba_p_ens_gen_p_eq in h4. apply (bconst_ext_p _ _ h4). apply Extensionality_Ensembles. red. split. red. intros x h5. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4) x). rewrite <- transfer_in_r. constructor. red. intros x h5. constructor. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite <- transfer_fun2_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4) x) at 2. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4) y) at 2. rewrite transfer_fun2_r_compat'. simpl. apply proj1_sig_injective. simpl. destruct x as [x h5], y as [y h6]. simpl. rewrite (transfer_r_sig_set_eq _ _ h4 (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4)). simpl. do 2 rewrite (transfer_sig_set_eq _ _ h4 (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4)). simpl. assert (h8: Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) x). rewrite <- h1c. apply (incl_gen_ens_p E h2 x h5). assert (h9 : Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) y). rewrite <- h1c. apply (incl_gen_ens_p E h2 y h6). pose proof (ba_p_subst_plus _ _ h1c _ _ (incl_gen_ens_p E h2 x h5) (incl_gen_ens_p E h2 y h6) h8 h9) as h7. rewrite h7 at 1. unfold Subalg_p. simpl. f_equal. f_equal. apply proj1_sig_injective. reflexivity. apply proj1_sig_injective. reflexivity. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite <- transfer_fun2_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4) x) at 2. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4) y) at 2. rewrite transfer_fun2_r_compat'. simpl. apply proj1_sig_injective. simpl. destruct x as [x h5], y as [y h6]. simpl. rewrite (transfer_r_sig_set_eq _ _ h4 (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4)). simpl. do 2 rewrite (transfer_sig_set_eq _ _ h4 (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4)). simpl. assert (h8: Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) x). rewrite <- h1c. apply (incl_gen_ens_p E h2 x h5). assert (h9 : Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) y). rewrite <- h1c. apply (incl_gen_ens_p E h2 y h6). pose proof (ba_p_subst_times _ _ h1c _ _ (incl_gen_ens_p E h2 x h5) (incl_gen_ens_p E h2 y h6) h8 h9) as h7. rewrite h7 at 1. unfold Subalg_p. simpl. f_equal. f_equal. apply proj1_sig_injective. reflexivity. apply proj1_sig_injective. reflexivity. symmetry. simpl. pose proof (ba_p_subst_one _ _ h1c) as h5. apply proj1_sig_injective. simpl. rewrite transfer_dep_r_id_transfer_r_compat. rewrite (transfer_r_sig_set_eq _ _ h4 (sig_set_eq (Gen_Ens_p E h2) (Gen_Ens_p E h3) h4)). simpl. rewrite h5. reflexivity. symmetry. simpl. pose proof (ba_p_subst_zero _ _ h1c) as h5. apply proj1_sig_injective. simpl. rewrite transfer_dep_r_id_transfer_r_compat. rewrite (transfer_r_sig_set_eq _ _ h4 (sig_set_eq (Gen_Ens_p E h2) (Gen_Ens_p E h3) h4)). simpl. rewrite h5. reflexivity. apply functional_extensionality. intro x. rewrite <- transfer_fun_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4) x) at 2. rewrite transfer_fun_r_compat'. simpl. apply proj1_sig_injective. simpl. destruct x as [x h5]. rewrite (transfer_r_sig_set_eq _ _ h4 (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4)). simpl. rewrite (transfer_sig_set_eq _ _ h4 (sig_set_eq (A_p T (Bc_p T (Gen_p E h2))) (A_p T (Bc_p T (Gen_p E h3))) h4)). simpl. assert (h8: Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h1a h1b)) x). rewrite <- h1c. apply (incl_gen_ens_p E h2 x h5). pose proof (ba_p_subst_comp _ _ h1c _ (incl_gen_ens_p E h2 x h5) h8) as h7. rewrite h7 at 1. unfold Subalg_p. simpl. f_equal. f_equal. apply proj1_sig_injective. reflexivity. Qed. Lemma subalg_of_p_gen_p_gen_p : forall {T:Type} (Bp:Bool_Alg_p T) (E F:Ensemble T) (pfi:Included E F) (pfe:Included E (ba_p_ens Bp)) (pff:Included F (ba_p_ens Bp)), subalg_of_p (Gen_p E pfe) (Gen_p F pff). intros T Bp E F h0 h1 h2. red. assert (h3:Included (ba_p_ens (Gen_p E h1)) (ba_p_ens (Gen_p F h2))). do 2 rewrite ba_p_ens_gen_p_eq. apply gen_ens_preserves_inclusion_p; auto. exists h3. assert (h4: alg_closed_p (ba_p_ens (Gen_p E h1)) h3). assert (h8:subalg_of_p (Gen_p F h2) Bp). apply gen_p_subalg_of_p_compat. assert (h5:Included (Gen_Ens_p E h1) (ba_p_ens (Gen_p F h2))). pose proof (ba_p_ens_gen_p_eq E h1) as h4. rewrite <- h4. assumption. assert (h9:Included E (ba_p_ens (Gen_p F h2))). rewrite ba_p_ens_gen_p_eq. pose proof (gen_ens_includes_p F h2) as h10. auto with sets. pose proof (gen_ens_p_subalg_of_p h8 E h9 h1) as h10. assert (h11:ba_p_ens (Gen_p E h1) = Gen_Ens_p E h9). rewrite ba_p_ens_gen_p_eq. rewrite h10. reflexivity. assert (h12:Included (Gen_Ens_p E h9) (ba_p_ens (Gen_p F h2))). rewrite <- h11. assumption. pose proof (subsetT_eq_compat _ (fun S => Included S (ba_p_ens (Gen_p F h2))) _ _ h12 h3 h10) as h13. dependent rewrite <- h13. assert (h14:h12 = incl_gen_ens_p E h9). apply proof_irrelevance. subst. apply closed_gen_ens_p. exists h4. apply bc_inj_p. assert (h5: A_p T (Bc_p T (Gen_p E h1)) = A_p T (Bc_p T (Subalg_p (Gen_p F h2) (ba_p_ens (Gen_p E h1)) h3 h4))). unfold Bc_p, Subalg_p. simpl. rewrite ba_p_ens_gen_p_eq. reflexivity. apply (bconst_ext_p (Bc_p T (Gen_p E h1)) (Bc_p T (Subalg_p (Gen_p F h2) (ba_p_ens (Gen_p E h1)) h3 h4)) h5). apply Extensionality_Ensembles. red. split. red. intros x h6. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h1))) (A_p T (Bc_p T (Gen_p E h1))) h5) x). rewrite <- transfer_in_r. constructor. red. intros x h6. constructor. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite <- transfer_fun2_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h1))) (A_p T (Bc_p T (Subalg_p (Gen_p F h2) (ba_p_ens (Gen_p E h1)) h3 h4))) h5) x). rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h1))) (A_p T (Bc_p T (Subalg_p (Gen_p F h2) (ba_p_ens (Gen_p E h1)) h3 h4))) h5) y). rewrite transfer_fun2_r_compat'. simpl. do 2 rewrite transfer_r_undoes_transfer. apply proj1_sig_injective. simpl. destruct x as [x h5'], y as [y h6]. simpl. rewrite (transfer_r_sig_set_eq _ _ h5 (sig_set_eq (Gen_Ens_p E h1) (ba_p_ens (Gen_p E h1)) h5)). simpl. do 2 rewrite (transfer_sig_set_eq _ _ h5 (sig_set_eq (Gen_Ens_p E h1) (ba_p_ens (Gen_p E h1)) h5)). simpl. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite <- transfer_fun2_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h1))) (A_p T (Bc_p T (Subalg_p (Gen_p F h2) (ba_p_ens (Gen_p E h1)) h3 h4))) h5) x). rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h1))) (A_p T (Bc_p T (Subalg_p (Gen_p F h2) (ba_p_ens (Gen_p E h1)) h3 h4))) h5) y). rewrite transfer_fun2_r_compat'. simpl. do 2 rewrite transfer_r_undoes_transfer. apply proj1_sig_injective. simpl. destruct x as [x h5'], y as [y h6]. simpl. rewrite (transfer_r_sig_set_eq _ _ h5 (sig_set_eq (Gen_Ens_p E h1) (ba_p_ens (Gen_p E h1)) h5)). simpl. do 2 rewrite (transfer_sig_set_eq _ _ h5 (sig_set_eq (Gen_Ens_p E h1) (ba_p_ens (Gen_p E h1)) h5)). simpl. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. rewrite transfer_dep_r_id_transfer_r_compat. rewrite (transfer_r_sig_set_eq _ _ h5 (sig_set_eq (A_p T (Bc_p T (Gen_p E h1))) (A_p T (Bc_p T (Subalg_p (Gen_p F h2) (ba_p_ens (Gen_p E h1)) h3 h4))) h5)). apply proj1_sig_injective. simpl. reflexivity. rewrite transfer_dep_r_id_transfer_r_compat. rewrite (transfer_r_sig_set_eq _ _ h5 (sig_set_eq (A_p T (Bc_p T (Gen_p E h1))) (A_p T (Bc_p T (Subalg_p (Gen_p F h2) (ba_p_ens (Gen_p E h1)) h3 h4))) h5)). apply proj1_sig_injective. simpl. reflexivity. apply functional_extensionality. intro x. rewrite <- transfer_fun_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T (Gen_p E h1))) (A_p T (Bc_p T (Subalg_p (Gen_p F h2) (ba_p_ens (Gen_p E h1)) h3 h4))) h5) x). rewrite transfer_fun_r_compat'. simpl. rewrite transfer_r_undoes_transfer. apply proj1_sig_injective. simpl. destruct x as [x h5']. rewrite (transfer_r_sig_set_eq _ _ h5 (sig_set_eq (Gen_Ens_p E h1) (ba_p_ens (Gen_p E h1)) h5)). simpl. rewrite (transfer_sig_set_eq _ _ h5 (sig_set_eq (Gen_Ens_p E h1) (ba_p_ens (Gen_p E h1)) h5)). simpl. f_equal. f_equal. apply proj1_sig_injective; auto. Qed. Lemma gen_p_functional : forall {T:Type} (B:Bool_Alg_p T) (E E':Ensemble T) (pfeq:E = E') (pf:Included E (ba_p_ens B)) (pf':Included E' (ba_p_ens B)), Gen_p E pf = Gen_p E' pf'. intros; subst; auto. assert (pf=pf'). apply proof_irrelevance. subst. reflexivity. Qed. Lemma gen_p_subalg_of_p_eq : forall {T:Type} (Bp Cp:Bool_Alg_p T) (pf:subalg_of_p Cp Bp), Gen_p (ba_p_ens Cp) (subalg_of_p_incl _ _ pf) = Cp. intros T Bp Cp h1. pose proof h1 as h1'. destruct h1 as [h1a [h1b h1c]]. apply subalg_of_p_eq. pose proof (gen_p_subalg_of_p_compat _ (inclusion_reflexive (ba_p_ens Cp))) as h2. pose proof (gen_p_subalg_of_p h1' _ (inclusion_reflexive (ba_p_ens Cp)) h1a) as h3. assert (h4:h1a = (subalg_of_p_incl Cp Bp (ex_intro (fun pfin : Included (ba_p_ens Cp) (ba_p_ens Bp) => exists pfac : alg_closed_p (ba_p_ens Cp) pfin, Cp = Subalg_p Bp (ba_p_ens Cp) pfin pfac) h1a (ex_intro (fun pfac : alg_closed_p (ba_p_ens Cp) h1a => Cp = Subalg_p Bp (ba_p_ens Cp) h1a pfac) h1b h1c)))). apply proof_irrelevance. rewrite <- h4 at 1. rewrite <- h3. apply gen_p_subalg_of_p_compat. rewrite ba_p_ens_gen_p_eq. apply gen_ens_closed_eq_p. eapply alg_closed_p_subalg_p_of. apply h1'. apply alg_closed_p_ba_p_ens_refl. Qed. Lemma subalg_of_p_gen_p2 : forall {T:Type} {Bp Cp:Bool_Alg_p T} (pf:Included (ba_p_ens Cp) (ba_p_ens Bp)), subalg_of_p Cp Bp -> subalg_of_p Cp (Gen_p (ba_p_ens Cp) pf). intros T Bp Cp h1 h2. pose proof h2 as h2'. destruct h2 as [h2a [h2b h2c]]. red. assert (h3:Included (ba_p_ens Cp) (ba_p_ens (Gen_p (ba_p_ens Cp) h1))). pose proof (gen_ens_includes_p _ h2a) as h3. rewrite ba_p_ens_gen_p_eq. assumption. exists h3. assert (h4:alg_closed_p (ba_p_ens Cp) h3). constructor; simpl. red; intros x y. unfold Bplus_sub_p, Gen_p, ba_p_ens. destruct x as [x h4], y as [y h5]. simpl. assert (h6: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h2a h2b)) x). rewrite <- h2c; auto. assert (h7 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h2a h2b)) y). rewrite <- h2c; auto. pose proof (ba_p_subst_plus _ _ h2c _ _ h4 h5 h6 h7) as h8. simpl in h8. assert (h9:h2a x h6 = incl_gen_ens_p (A_p T (Bc_p T Cp)) h1 x (h3 x h4)). apply proof_irrelevance. assert (h10:h2a y h7 = incl_gen_ens_p (A_p T (Bc_p T Cp)) h1 y (h3 y h5)). apply proof_irrelevance. rewrite h9, h10 in h8 at 1. clear h9 h10. rewrite <- h8. apply proj2_sig. red; intros x y. unfold Btimes_sub_p, Gen_p, ba_p_ens. destruct x as [x h4], y as [y h5]. simpl. assert (h6: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h2a h2b)) x). rewrite <- h2c; auto. assert (h7 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h2a h2b)) y). rewrite <- h2c; auto. pose proof (ba_p_subst_times _ _ h2c _ _ h4 h5 h6 h7) as h8. simpl in h8. assert (h9:h2a x h6 = incl_gen_ens_p (A_p T (Bc_p T Cp)) h1 x (h3 x h4)). apply proof_irrelevance. assert (h10:h2a y h7 = incl_gen_ens_p (A_p T (Bc_p T Cp)) h1 y (h3 y h5)). apply proof_irrelevance. rewrite h9, h10 in h8 at 1. clear h9 h10. rewrite <- h8. apply proj2_sig. red. pose proof (ba_p_subst_one _ _ h2c) as h4. rewrite <- h4 at 1. apply proj2_sig. red. pose proof (ba_p_subst_zero _ _ h2c) as h4. rewrite <- h4 at 1. apply proj2_sig. red; intros x. unfold Bcomp_sub_p, Gen_p, ba_p_ens. destruct x as [x h4]. simpl. assert (h6: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h2a h2b)) x). rewrite <- h2c; auto. pose proof (ba_p_subst_comp _ _ h2c _ h4 h6) as h8. simpl in h8. assert (h9:h2a x h6 = incl_gen_ens_p (A_p T (Bc_p T Cp)) h1 x (h3 x h4)). apply proof_irrelevance. rewrite h9 in h8 at 1. clear h9. rewrite <- h8. apply proj2_sig. exists h4. pose proof (gen_p_subalg_of_p_eq _ _ h2') as h5. assert (h6:h1 = subalg_of_p_incl Cp Bp h2'). apply proof_irrelevance. rewrite <- h6 in h5 at 1. apply bc_inj_p. simpl. assert (h8: A_p T (Bc_p T Cp) = A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)). simpl. reflexivity. pose proof (bconst_ext_p (Bc_p T Cp) (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4) h8) as h9. apply h9. clear h9. apply Extensionality_Ensembles. red; split; red; intros x h9. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8) x). rewrite <- transfer_in_r. constructor. apply in_bs_p. apply functional_extensionality. intro x. apply functional_extensionality. intro y. destruct x as [x h10], y as [y h11]. rewrite <- transfer_fun2_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8) (exist _ _ h10)) at 2. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8) (exist _ _ h11)) at 2. simpl. pose proof (transfer_fun2_r_compat' (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8) (Bplus_p T (Bc_p' T (@Gen_p T Bp (@ba_p_ens T Cp) h1) (@ba_p_ens T Cp) h3 h4)) (@transfer (@sig_set T (A_p T (Bc_p T Cp))) (@sig_set T (A_p T (Bc_p T Cp))) (@sig_set_eq T (A_p T (Bc_p T Cp)) (A_p T (Bc_p T Cp)) h8) (@exist T (@Inn T (A_p T (Bc_p T Cp))) x h10))) (@transfer (@sig_set T (A_p T (Bc_p T Cp))) (@sig_set T (A_p T (Bc_p T Cp))) (@sig_set_eq T (A_p T (Bc_p T Cp)) (A_p T (Bc_p T Cp)) h8) (@exist T (@Inn T (A_p T (Bc_p T Cp))) y h11)) as h12. rewrite h12 at 1. rewrite (transfer_r_sig_set_eq _ _ h8 (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8)). apply proj1_sig_injective. simpl. do 2 rewrite (transfer_sig_set_eq _ _ h8 (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8)). simpl. assert (h14: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h2a h2b)) x). rewrite <- h2c; auto. assert (h15: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h2a h2b)) y). rewrite <- h2c; auto. pose proof (ba_p_subst_plus _ _ h2c _ _ h10 h11 h14 h15) as h13. simpl in h13. rewrite h13 at 1. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. apply functional_extensionality. intro x. apply functional_extensionality. intro y. destruct x as [x h10], y as [y h11]. rewrite <- transfer_fun2_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8) (exist _ _ h10)) at 2. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8) (exist _ _ h11)) at 2. simpl. pose proof (transfer_fun2_r_compat' (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8) (Btimes_p T (Bc_p' T (@Gen_p T Bp (@ba_p_ens T Cp) h1) (@ba_p_ens T Cp) h3 h4)) (@transfer (@sig_set T (A_p T (Bc_p T Cp))) (@sig_set T (A_p T (Bc_p T Cp))) (@sig_set_eq T (A_p T (Bc_p T Cp)) (A_p T (Bc_p T Cp)) h8) (@exist T (@Inn T (A_p T (Bc_p T Cp))) x h10))) (@transfer (@sig_set T (A_p T (Bc_p T Cp))) (@sig_set T (A_p T (Bc_p T Cp))) (@sig_set_eq T (A_p T (Bc_p T Cp)) (A_p T (Bc_p T Cp)) h8) (@exist T (@Inn T (A_p T (Bc_p T Cp))) y h11)) as h12. rewrite h12 at 1. rewrite (transfer_r_sig_set_eq _ _ h8 (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8)). apply proj1_sig_injective. simpl. do 2 rewrite (transfer_sig_set_eq _ _ h8 (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8)). simpl. assert (h14: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h2a h2b)) x). rewrite <- h2c; auto. assert (h15: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h2a h2b)) y). rewrite <- h2c; auto. pose proof (ba_p_subst_times _ _ h2c _ _ h10 h11 h14 h15) as h13. simpl in h13. rewrite h13 at 1. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. rewrite transfer_dep_r_id_transfer_r_compat. rewrite (transfer_r_sig_set_eq _ _ h8 (@sig_set_eq T (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (@Gen_p T Bp (@ba_p_ens T Cp) h1) (@ba_p_ens T Cp) h3 h4)) h8)). apply proj1_sig_injective; simpl. pose proof (ba_p_subst_one _ _ h2c) as h10. rewrite h10 at 1. reflexivity. rewrite transfer_dep_r_id_transfer_r_compat. rewrite (transfer_r_sig_set_eq _ _ h8 (@sig_set_eq T (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (@Gen_p T Bp (@ba_p_ens T Cp) h1) (@ba_p_ens T Cp) h3 h4)) h8)). apply proj1_sig_injective; simpl. pose proof (ba_p_subst_zero _ _ h2c) as h10. rewrite h10 at 1. reflexivity. apply functional_extensionality. intro x. destruct x as [x h10]. rewrite <- transfer_fun_r_transfer_dep_r_compat'. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8) (exist _ _ h10)) at 2. simpl. pose proof (transfer_fun_r_compat' (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8) (Bcomp_p T (Bc_p' T (@Gen_p T Bp (@ba_p_ens T Cp) h1) (@ba_p_ens T Cp) h3 h4)) (@transfer (@sig_set T (A_p T (Bc_p T Cp))) (@sig_set T (A_p T (Bc_p T Cp))) (@sig_set_eq T (A_p T (Bc_p T Cp)) (A_p T (Bc_p T Cp)) h8) (@exist T (@Inn T (A_p T (Bc_p T Cp))) x h10))) as h12. rewrite h12 at 1. rewrite (transfer_r_sig_set_eq _ _ h8 (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8)). apply proj1_sig_injective. simpl. rewrite (transfer_sig_set_eq _ _ h8 (sig_set_eq (A_p T (Bc_p T Cp)) (A_p T (Bc_p' T (Gen_p (ba_p_ens Cp) h1) (ba_p_ens Cp) h3 h4)) h8)). simpl. assert (h14: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h2a h2b)) x). rewrite <- h2c; auto. pose proof (ba_p_subst_comp _ _ h2c _ h10 h14) as h13. simpl in h13. rewrite h13 at 1. f_equal. f_equal. apply proj1_sig_injective; auto. Qed. Lemma gen_p_in_gen_p_eq: forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)) (pf':Included E (ba_p_ens (Gen_p E pf))), Gen_p E pf' = Gen_p E pf. intros T Bp E h1 h2. pose proof (gen_p_subalg_of_p_compat _ h1) as h3. apply (gen_p_subalg_of_p h3 _ h2 h1). Qed. Lemma gen_ens_p_in_gen_ens_p_eq: forall {T:Type} {Bp:Bool_Alg_p T} (E:Ensemble T) (pf:Included E (ba_p_ens Bp)) (pf':Included E (ba_p_ens (Gen_p E pf))), Gen_Ens_p E pf' = Gen_Ens_p E pf. intros T Bp E h1 h2. rewrite <- ba_p_ens_gen_p_eq. rewrite gen_p_in_gen_p_eq; auto. Qed. Definition comp_add_set_p {T:Type} {Ap:Bool_Alg_p T} (E:Ensemble (btp Ap)) := Union E (Im E (Bcomp_p T (Bc_p T Ap))). Definition comp_add_set_p_eq : forall {T:Type} {Ap:Bool_Alg_p T} (E:Ensemble (btp Ap)), comp_add_set_p E = comp_add_set (ba_conv_set E). intros T Ap E. unfold comp_add_set, comp_add_set, ba_conv_set, ba_conv_type. rewrite transfer_dep_eq_refl. reflexivity. Qed. Lemma incl_comp_add_set_p : forall {T:Type} {Ap:Bool_Alg_p T} (D:Ensemble (btp Ap)), Included D (comp_add_set_p D). intros T Ap D. red. intros x h1. left. assumption. Qed. Lemma incl_comp_add_set_closed_p : forall {T:Type} {Ap:Bool_Alg_p T} (D E:Ensemble (btp Ap)), alg_closed_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E) -> Included D (comp_add_set_p E) -> Included D E. intros T Ap D E h1 h2. red. red in h2. intros x h3. specialize (h2 _ h3). destruct h2 as [x h2l | x h2r]. assumption. destruct h2r as [x h4]. subst. assert (h5: Inn (im_proj1_sig E) (proj1_sig x)). rewrite in_im_proj1_sig_iff. assumption. pose proof (comp_closed_p' _ _ _ _ h1 (proj1_sig x) h5) as h6. inversion h6 as [x' h7 q h9 h10]. subst. clear h6. apply proj1_sig_injective in h9. subst. assert (h8:%- exist _ _ (incl_im_proj1_sig_ba_p_ens E (proj1_sig x) h5) = %- x). f_equal. apply proj1_sig_injective; auto. rewrite <- h8. assumption. Qed. Lemma comp_add_set_add_p : forall {T:Type} {Ap:Bool_Alg_p T} (D:Ensemble (btp Ap)) (a:btp Ap), comp_add_set_p (Add D a) = Union (comp_add_set_p D) (Couple a (%-a)). intros T Ap D a. rewrite comp_add_set_p_eq. rewrite comp_add_set_p_eq at 1. unfold ba_conv_set, ba_conv_type. do 2 rewrite transfer_dep_eq_refl. pose proof (@comp_add_set_add (ba_conv Ap) D a) as h1. rewrite h1 at 1. reflexivity. Qed. Lemma comp_add_set_preserves_inclusion_p : forall {T:Type} {Ap:Bool_Alg_p T} (D E:Ensemble (btp Ap)), Included D E -> Included (comp_add_set_p D) (comp_add_set_p E). intros T Ap D E h1. rewrite comp_add_set_p_eq. rewrite comp_add_set_p_eq at 1. apply (@comp_add_set_preserves_inclusion (ba_conv Ap)). assumption. Qed. Lemma gen_ens_comp_add_set_p : forall {T:Type} {Ap:Bool_Alg_p T} (E:Ensemble (btp Ap)), Gen_Ens_p (im_proj1_sig (comp_add_set_p E)) (incl_im_proj1_sig_ba_p_ens (comp_add_set_p E)) = Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E). intros T Ap E. do 2 rewrite gen_ens_p_eq. f_equal. unfold btp, ba_p_ens. pose proof (im_proj2_sig_undoes_im_proj1_sig' _ E (incl_im_proj1_sig_ba_p_ens E)) as h1. pose proof (im_proj2_sig_undoes_im_proj1_sig' _ (comp_add_set_p E) (incl_im_proj1_sig_ba_p_ens (comp_add_set_p E))) as h2. rewrite <- h1. rewrite <- h2. apply (@gen_ens_comp_add_set (ba_conv Ap)). Qed. Lemma finite_fin_gen_p : forall {T:Type} (Bp:Bool_Alg_p T) (E:Ensemble T) (pf:Included E (ba_p_ens Bp)), Finite E -> Finite (Gen_Ens_p E pf). intros T Bp E h1 h2. rewrite gen_ens_p_eq. apply finite_image; auto. assert (h3:Finite (im_proj2_sig E h1)). apply finite_image. rewrite <- finite_full_sig_iff. assumption. apply (finite_fin_gen (ba_conv Bp) _ h3). Qed. Lemma gen_p_eq_gen_ens_p_eq : forall {T:Type} {Bp:Bool_Alg_p T} (E F:Ensemble T) (pfe:Included E (ba_p_ens Bp)) (pff:Included F (ba_p_ens Bp)), Gen_p E pfe = Gen_p F pff -> Gen_Ens_p E pfe = Gen_Ens_p F pff. intros T Bp E F h1 h2 h3. do 2 rewrite <- ba_p_ens_gen_p_eq. f_equal. assumption. Qed. Lemma gen_p_eq_gen_ens_p_eq_in2 : forall {T:Type} {Bp:Bool_Alg_p T} (E F:Ensemble T) (pfe: Included E (ba_p_ens Bp)) (pff: Included F (ba_p_ens Bp)) (pff': Included F (ba_p_ens (Gen_p F pff))), Gen_p E pfe = Gen_p F pff' -> Gen_Ens_p E pfe = Gen_Ens_p F pff'. intros T Bp E F h1 h2 h3 h4. rewrite (gen_p_eq_gen_ens_p_eq _ _ _ h2). do 2 rewrite <- ba_p_ens_gen_p_eq. f_equal. rewrite gen_p_in_gen_p_eq. reflexivity. rewrite h4. rewrite gen_p_in_gen_p_eq. reflexivity. Qed. End Gen_Ens_p. Section NormalForm_p. Variable T:Type. Variable Ap:Bool_Alg_p T. Let Atp := btp Ap. Definition eps_p {T':Type} {Bp:Bool_Alg_p T'} (b:(btp Bp)) (sgn:sign) := if sgn then b else %-b. Definition eps_p' {T':Type} {Bp:Bool_Alg_p T'} (pr:(btp Bp)*sign) := eps_p (fst pr) (snd pr). Lemma eps_eps'_compat_p : forall {T':Type} {Bp:Bool_Alg_p T'} (x:(btp Bp)), eps_p x = (fun y:sign => eps_p' (x, y)). intro x. unfold eps_p, eps_p'. simpl. unfold eps_p. reflexivity. Qed. Lemma eps_p_eq : forall {T':Type} {Bp:Bool_Alg_p T'} (b:(btp Bp)) (sgn:sign), eps_p b sgn = eps (ba_conv_elt b) sgn. intros; auto. Qed. Lemma eps_p_eq' : forall {T':Type} {Bp:Bool_Alg_p T'} (pr:(btp Bp)*sign), eps_p' pr = eps' (ba_conv_elt (fst pr), snd pr). intros; auto. Qed. Lemma eps_p_eq0 : forall {T':Type} {Bp:Bool_Alg_p T'} (b:(btp Bp)) (sgn:sign), eps_p b sgn = @eps (ba_conv Bp) b sgn. intros; auto. Qed. Lemma eps_inj_p : forall {T':Type} {Bp:Bool_Alg_p T'} (x y:btp Bp) (s s':sign), eps_p x s = eps_p y s' -> (x = y /\ s = s') \/ (x = %-y /\ s = --s'). intros T' Bp x y s s' h1. do 2 rewrite eps_p_eq in h1. apply (@eps_inj (ba_conv Bp)) in h1; auto. Qed. Lemma eps_opp_p : forall {T':Type} {Bp:Bool_Alg_p T'} (x:btp Bp) (s:sign), eps_p x (--s) = %- (eps_p x s). unfold sign_flip; intros T' Bp x s; destruct s; simpl; auto. rewrite <- doub_neg_p; auto. Qed. Lemma eps_covers_p : forall {T':Type} {Bp:Bool_Alg_p T'} (E:Ensemble (btp Bp)) (pf:Finite E) (x:btp Bp), plus_set_p (Im signe (eps_p x)) (finite_image _ _ signe (eps_p x) (signe_finite)) = %1. intros T' Bp E h1 x. pose proof (eps_covers (ba_conv_set E) h1 (ba_conv_elt x)) as h2. rewrite plus_set_p_eq. assumption. Qed. Definition eps_map_p {E:Ensemble Atp} (f:Fin_map E signe mns) : Atp->Atp := (fun i:Atp => eps_p i (f |-> i)). Lemma eps_map_p_eq : forall {E:Ensemble Atp} (f:Fin_map E signe mns), eps_map_p f = eps_map (ba_conv_fin_map_dom f). intros; auto. Qed. Lemma eps_maps_dec_p : forall {E:Ensemble Atp} (f g:Fin_map E signe mns) (x:Atp), Inn E x -> {eps_map_p f x = eps_map_p g x} + {eps_map_p f x = %-(eps_map_p g x)}. intros E f g x h1. pose proof (eps_maps_dec _ (ba_conv_fin_map_dom f) (ba_conv_fin_map_dom g) (ba_conv_elt x) h1) as h2. assumption. Qed. Definition eps_map_compose_p {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (f:Fin_map E signe mns) (g:(btp Bp)->Atp) : (btp Bp) -> Atp := (fun i:(btp Bp) => eps_p (g i) (f |-> i)). Lemma eps_map_compose_p_eq : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (f:Fin_map E signe mns) (g:(btp Bp)->Atp), eps_map_compose_p f g = eps_map_compose (ba_conv_fin_map_dom f) (ba_conv_fun1 g) (A:=ba_conv Ap). intros; auto. Qed. Lemma eps_maps_compose_dec_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (f f':Fin_map E signe mns) (g:(btp Bp)->Atp) (x:(btp Bp)), Inn E x -> {eps_map_compose_p f g x = eps_map_compose_p f' g x} + {eps_map_compose_p f g x = %-(eps_map_compose_p f' g x)}. intros Bp E f f' g x h1. pose proof (eps_maps_compose_dec (ba_conv Ap) (ba_conv_fin_map_dom f) (ba_conv_fin_map_dom f') (ba_conv_fun1 g) (ba_conv_elt x) h1) as h2. assumption. Qed. Definition im_eps_p {E:Ensemble Atp} (f:Fin_map E signe mns) : Ensemble Atp := Im E (eps_map_p f). Lemma im_eps_p_eq : forall {E:Ensemble Atp} (f:Fin_map E signe mns), im_eps_p f = im_eps _ (ba_conv_fin_map_dom f). intros; auto. Qed. Definition im_eps_compose_p {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (f:Fin_map E signe mns) (g:(btp Bp)->Atp) : Ensemble Atp := Im E (eps_map_compose_p f g). Lemma im_eps_compose_p_eq : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (f:Fin_map E signe mns) (g:(btp Bp)->Atp), im_eps_compose_p f g = im_eps_compose (ba_conv Ap) (ba_conv_fin_map_dom f) (ba_conv_fun1 g). intros; auto. Qed. Lemma im_eps_fin_map_incl_comp_add_set_p : forall {D:Ensemble Atp} (f:Fin_map D signe mns), Included (Im D (fun i => eps_p i (f |-> i))) (comp_add_set_p D). intros D f. rewrite comp_add_set_p_eq. pose proof (im_eps_fin_map_incl_comp_add_set _ (ba_conv_fin_map_dom f)) as h1. assert (h2: Im (ba_conv_set D) (fun i : Btype (Bc (ba_conv Ap)) => eps i (ba_conv_fin_map_dom f |-> i)) = Im D (fun i : Atp => eps_p i (f |-> i))). apply im_ext_in. intros x h2. rewrite eps_p_eq. reflexivity. rewrite <- h2. assumption. Qed. Lemma finite_im_eps_p : forall {E:Ensemble Atp} (f:Fin_map E signe mns), Finite (im_eps_p f). intros E f. apply finite_image. apply (fin_map_fin_dom f). Qed. Definition finite_im_eps_compose_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (f:Fin_map E signe mns) (g:(btp Bp)->Atp), Finite (im_eps_compose_p f g). intros B E f g. apply finite_image. apply (fin_map_fin_dom f). Qed. Lemma im_eps_im_eps_compose_compat_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:sig_set E->Atp) (f : Fin_map (Im (full_sig E) g) signe mns) (pf:Finite E), im_eps_p f = im_eps_compose_p (fin_map_im_full_sig_eq ba_eq_dec_p f pf %0) (sig_fun_app g %0). intros Bp E g f h1. pose proof (im_eps_im_eps_compose_compat (ba_conv Ap) (ba_conv_sig_fun1 g) (ba_conv_fin_map_dom f) h1) as h2. rewrite (type_eq_dec_eq ba_eq_dec_p (@ba_eq_dec (ba_conv Bp))) at 1. assumption. Qed. Section HasCompsP. Inductive has_comps_p {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) : Prop := | has_comps_intro_p : forall (c d:nat) (pfc:c < m) (pfd:d < m), Finite P -> Inn P c -> Inn P d -> c <> d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = nth_lt l d (lt_le_trans _ _ _ pfd pfm) -> a |-> c <> a |-> d -> has_comps_p _ _ pfm P a. Definition has_comps_p' {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) : Prop := has_comps_p l (length l) (le_refl _) P a. Inductive has_sames_p {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) : Prop := | has_sames_intro_p : forall (c d:nat) (pfc:c < m) (pfd:d < m), Finite P -> Inn P c -> Inn P d -> c <> d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = nth_lt l d (lt_le_trans _ _ _ pfd pfm) -> a |-> c = a |-> d -> has_sames_p _ _ pfm P a. Definition has_sames_p' {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) : Prop := has_sames_p l (length l) (le_refl _) P a. Inductive has_comps_elt_p {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns) : Prop := | has_comps_elt_intro_p : Finite P -> forall (c d:nat) (pfc:c < m) (pfd:d < m), c <> d -> Inn P c -> Inn P d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = x -> nth_lt l d (lt_le_trans _ _ _ pfd pfm) = x -> a |-> c <> a |-> d -> has_comps_elt_p _ _ pfm P x a. Definition has_comps_elt_p' {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns) : Prop := has_comps_elt_p l (length l) (le_refl _) P x a. Inductive has_sames_elt_p {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns) : Prop := | has_sames_elt_intro_p : Finite P -> forall (c d:nat) (pfc:c < m) (pfd:d < m), c <> d -> Inn P c -> Inn P d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = x -> nth_lt l d (lt_le_trans _ _ _ pfd pfm) = x -> a |-> c = a |-> d -> has_sames_elt_p _ _ pfm P x a. Definition has_sames_elt_p' {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns) : Prop := has_sames_elt_p l (length l) (le_refl _) P x a. Lemma has_comps_p_functional : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m n:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) (pfem:m = n), let pfn := eq_le_trans _ _ _ (eq_sym pfem) pfm in let a' := fin_map_seg_transfer a pfem in has_comps_p l m pfm P a -> has_comps_p l n pfn P a'. simpl; intros Bp l m n h1 P a h2 h3; subst. rewrite fin_map_seg_transfer_eq_refl. simpl. assert (h4:eq_le_trans _ _ _ eq_refl h1 = h1). apply proof_irrelevance. rewrite h4; auto. Qed. Lemma has_comps_p_functional' : forall {Bp:Bool_Alg_p T} (l l':list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) (pfel:l = l'), let pfell := length_functional _ _ pfel in let a' := fin_map_seg_transfer a pfell in has_comps_p' l P a -> has_comps_p' l' P a'. simpl; intros; subst; simpl. assert (h1:eq_refl = length_functional l' l' eq_refl). apply proof_irrelevance. rewrite <- h1. rewrite fin_map_seg_transfer_eq_refl; auto. Qed. Lemma has_sames_p_functional : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m n:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) (pfem:m = n), let pfn := eq_le_trans _ _ _ (eq_sym pfem) pfm in let a' := fin_map_seg_transfer a pfem in has_sames_p l m pfm P a -> has_sames_p l n pfn P a'. simpl; intros Bp l m n h1 P a h2 h3; subst. rewrite fin_map_seg_transfer_eq_refl. simpl. assert (h4:eq_le_trans _ _ _ eq_refl h1 = h1). apply proof_irrelevance. rewrite h4; auto. Qed. Lemma has_sames_p_functional' : forall {Bp:Bool_Alg_p T} (l l':list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) (pfel:l = l'), let pfell := length_functional _ _ pfel in let a' := fin_map_seg_transfer a pfell in has_sames_p' l P a -> has_sames_p' l' P a'. simpl; intros; subst; simpl. assert (h1:eq_refl = length_functional l' l' eq_refl). apply proof_irrelevance. rewrite <- h1. rewrite fin_map_seg_transfer_eq_refl; auto. Qed. Lemma has_comps_p_functional_rev : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m n:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) (pfem:m = n), let pfn := eq_le_trans _ _ _ (eq_sym pfem) pfm in let a' := fin_map_seg_transfer a pfem in has_comps_p l n pfn P a' -> has_comps_p l m pfm P a. simpl; intros Bp l m n h1 P a h2 h3; subst. rewrite fin_map_seg_transfer_eq_refl in h3. simpl. assert (h4:eq_le_trans _ _ _ eq_refl h1 = h1). apply proof_irrelevance. rewrite <- h4; auto. Qed. Lemma has_comps_p_functional_rev' : forall {Bp:Bool_Alg_p T} (l l':list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) (pfel:l = l'), let pfell := length_functional _ _ pfel in let a' := fin_map_seg_transfer a pfell in has_comps_p' l' P a' -> has_comps_p' l P a. simpl; intros; subst; simpl. assert (h1:eq_refl = length_functional l' l' eq_refl). apply proof_irrelevance. rewrite <- h1 in H. rewrite <- fin_map_seg_transfer_eq_refl; auto. Qed. Lemma has_sames_p_functional_rev : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m n:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) (pfem:m = n), let pfn := eq_le_trans _ _ _ (eq_sym pfem) pfm in let a' := fin_map_seg_transfer a pfem in has_sames_p l n pfn P a' -> has_sames_p l m pfm P a. simpl; intros Bp l m n h1 P a h2 h3; subst. rewrite fin_map_seg_transfer_eq_refl in h3. simpl in h3. assert (h4:eq_le_trans _ _ _ eq_refl h1 = h1). apply proof_irrelevance. rewrite <- h4; auto. Qed. Lemma has_sames_p_functional_rev' : forall {Bp:Bool_Alg_p T} (l l':list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) (pfel:l = l'), let pfell := length_functional _ _ pfel in let a' := fin_map_seg_transfer a pfell in has_sames_p' l' P a' -> has_sames_p' l P a. simpl; intros; subst; simpl. assert (h1:eq_refl = length_functional l' l' eq_refl). apply proof_irrelevance. rewrite <- h1 in H. rewrite fin_map_seg_transfer_eq_refl in H; auto. Qed. Lemma has_comps_elt_finite_p : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns), has_comps_elt_p l m pfm P x a -> Finite P. intros ? ? ? ? ? ? ? h1. destruct h1; auto. Qed. Lemma has_comps_elt_finite_p' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns), has_comps_elt_p' l P x a -> Finite P. intros ? ? ? ? ? h1; red in h1. eapply has_comps_elt_finite_p in h1; auto. Qed. Lemma has_sames_elt_finite_p : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns), has_sames_elt_p l m pfm P x a -> Finite P. intros ? ? ? ? ? ? ? h1. destruct h1; auto. Qed. Lemma has_sames_elt_finite_p' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns), has_sames_elt_p' l P x a -> Finite P. intros ? ? ? ? ? h1; red in h1. eapply has_sames_elt_finite_p in h1; auto. Qed. Lemma has_comps_elt_p_functional : forall {Bp:Bool_Alg_p T} (l l':list (btp Bp)) (m:nat) (x x':btp Bp) (pfm: m <= length l) (pfm': m <= length l') (P:Ensemble nat) (a:Fin_map (seg m) signe mns), l = l' -> x = x' -> has_comps_elt_p _ _ pfm P x a -> has_comps_elt_p _ _ pfm' P x a. intros; subst; auto. assert (pfm = pfm'). apply proof_irrelevance. subst; auto. Qed. Lemma has_comps_elt_p_functional' : forall {Bp:Bool_Alg_p T} (l l':list (btp Bp)) (x x':btp Bp) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), l = l' -> x = x' -> has_comps_elt_p' _ P x a -> has_comps_elt_p' _ P x' a. intros; subst; auto. Qed. Lemma has_sames_elt_p_functional : forall {Bp:Bool_Alg_p T} (l l':list (btp Bp)) (m:nat) (x x':btp Bp) (pfm: m <= length l) (pfm': m <= length l') (P:Ensemble nat) (a:Fin_map (seg m) signe mns), l = l' -> x = x' -> has_sames_elt_p _ _ pfm P x a -> has_sames_elt_p _ _ pfm' P x' a. intros; subst; auto. assert (pfm = pfm'). apply proof_irrelevance. subst; auto. Qed. Lemma has_sames_elt_p_functional' : forall {Bp:Bool_Alg_p T} (l l':list (btp Bp)) (x x':btp Bp) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), l = l' -> x = x' -> has_sames_elt_p' _ P x a -> has_sames_elt_p' _ P x' a. intros; subst; auto. Qed. Lemma has_comps_elt_p_iff_b1_b2 : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns), has_comps_elt_p l m pfm P x a <-> has_comps_elt (ba_conv_list l) m pfm P (ba_conv_elt x) a. intros Bp l m h1 P x a; split; intro h2. destruct h2. econstructor. assumption. apply H0. assumption. assumption. unfold ba_conv_list, ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. gen0. rewrite transfer_dep_eq_refl. intro h6. unfold ba_conv_elt. erewrite nth_lt_functional3 in H3. apply H3. unfold ba_conv_list, ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. gen0. rewrite transfer_dep_eq_refl. intro h6. unfold ba_conv_elt. erewrite nth_lt_functional3 in H4. apply H4. assumption. destruct h2. econstructor. assumption. apply H0. assumption. assumption. apply H3. apply H4. assumption. Grab Existential Variables. assumption. assumption. Qed. Lemma has_comps_elt_p_iff_b1_b2' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns), has_comps_elt_p' l P x a <-> has_comps_elt' (ba_conv_list l) P (ba_conv_elt x) a. unfold has_comps_elt'; intros. pose proof (has_comps_elt_p_iff_b1_b2 l (length l) (le_refl _) P x a) as h1. unfold has_comps_elt_p'; auto. Qed. Lemma has_comps_elt_p_iff_b1 : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns), has_comps_elt_p l m pfm P x a <-> has_comps_elt (ba_conv_list l) m pfm P x a. intros Bp l m h1 P x a. pose proof (has_comps_elt_p_iff_b1_b2 l m h1 P x a) as h2. rewrite h2; tauto. Qed. Lemma has_comps_elt_p_iff_b1' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns), has_comps_elt_p' l P x a <-> has_comps_elt' (ba_conv_list l) P x a. unfold has_comps_elt'; intros. pose proof (has_comps_elt_p_iff_b1_b2 l (length l) (le_refl _) P x a) as h1. unfold has_comps_elt_p'; auto. Qed. Lemma has_comps_elt_p_iff_b2 : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns), has_comps_elt_p l m pfm P x a <-> has_comps_elt (B := ba_conv Bp) l m pfm P (ba_conv_elt x) a. intros; apply has_comps_elt_p_iff_b1. Qed. Lemma has_comps_elt_p_iff_b2' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns), has_comps_elt_p' l P x a <-> has_comps_elt' (B := ba_conv Bp) l P (ba_conv_elt x) a. intros; apply has_comps_elt_p_iff_b1. Qed. Lemma has_comps_elt_p_iff : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns), has_comps_elt_p l m pfm P x a <-> has_comps_elt (B := ba_conv Bp) l m pfm P x a. intros; apply has_comps_elt_p_iff_b2. Qed. Lemma has_comps_elt_p_iff' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns), has_comps_elt_p' l P x a <-> has_comps_elt' (B := ba_conv Bp) l P x a. intros; apply has_comps_elt_p_iff_b2. Qed. Lemma has_sames_elt_p_iff_b1_b2 : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns), has_sames_elt_p l m pfm P x a <-> has_sames_elt (ba_conv_list l) m pfm P (ba_conv_elt x) a. intros Bp l m h1 P x a; split; intro h2. destruct h2. econstructor. assumption. apply H0. assumption. assumption. unfold ba_conv_list, ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. gen0. rewrite transfer_dep_eq_refl. intro h6. unfold ba_conv_elt. erewrite nth_lt_functional3 in H3. apply H3. unfold ba_conv_list, ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. gen0. rewrite transfer_dep_eq_refl. intro h6. unfold ba_conv_elt. erewrite nth_lt_functional3 in H4. apply H4. assumption. destruct h2. econstructor. assumption. apply H0. assumption. assumption. apply H3. apply H4. assumption. Grab Existential Variables. assumption. assumption. Qed. Lemma has_sames_elt_p_iff_b1_b2' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns), has_sames_elt_p' l P x a <-> has_sames_elt' (ba_conv_list l) P (ba_conv_elt x) a. unfold has_comps_elt'; intros. pose proof (has_sames_elt_p_iff_b1_b2 l (length l) (le_refl _) P x a) as h1. unfold has_sames_elt_p'; auto. Qed. Lemma has_sames_elt_p_iff_b1 : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns), has_sames_elt_p l m pfm P x a <-> has_sames_elt (ba_conv_list l) m pfm P x a. intros Bp l m h1 P x a. pose proof (has_sames_elt_p_iff_b1_b2 l m h1 P x a) as h2. rewrite h2; tauto. Qed. Lemma has_sames_elt_p_iff_b1' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns), has_sames_elt_p' l P x a <-> has_sames_elt' (ba_conv_list l) P x a. unfold has_comps_elt'; intros. pose proof (has_sames_elt_p_iff_b1_b2 l (length l) (le_refl _) P x a) as h1. unfold has_sames_elt_p'; auto. Qed. Lemma has_sames_elt_p_iff_b2 : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns), has_sames_elt_p l m pfm P x a <-> has_sames_elt (B := ba_conv Bp) l m pfm P (ba_conv_elt x) a. intros; apply has_sames_elt_p_iff_b1. Qed. Lemma has_sames_elt_p_iff_b2' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns), has_sames_elt_p' l P x a <-> has_sames_elt' (B := ba_conv Bp) l P (ba_conv_elt x) a. intros; apply has_sames_elt_p_iff_b1. Qed. Lemma has_sames_elt_p_iff : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns), has_sames_elt_p l m pfm P x a <-> has_sames_elt (B := ba_conv Bp) l m pfm P x a. intros; apply has_sames_elt_p_iff_b2. Qed. Lemma has_sames_elt_p_iff' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns), has_sames_elt_p' l P x a <-> has_sames_elt' (B := ba_conv Bp) l P x a. intros; apply has_sames_elt_p_iff_b2. Qed. Lemma has_comps_p_iff : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), let pfm' := le_length_ba_conv_list _ _ pfm in has_comps_p _ _ pfm P a <-> has_comps (ba_conv_list l) _ pfm' P a. intros Bp l m hm P a. split. intro h1. destruct h1 as [c d hc hd h0 h1 h2 h3 h4 h5]. econstructor. assumption. apply h1. apply h2. assumption. do 2 rewrite nth_lt_ba_conv_list in h4. erewrite nth_lt_functional3. rewrite h4 at 1. apply nth_lt_functional3. assumption. intro h1. destruct h1 as [c d hc hd h0 h1 h2 h3 h4 h5]. econstructor. assumption. apply h1. apply h2. assumption. do 2 rewrite nth_lt_ba_conv_list. erewrite nth_lt_functional3. rewrite h4 at 1. apply nth_lt_functional3. assumption. Grab Existential Variables. assumption. assumption. assumption. assumption. Qed. Lemma has_comps_p_iff' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps_p' _ P a <-> has_comps' (ba_conv_list l) P a. unfold has_comps_p', has_comps'; intros. rewrite has_comps_p_iff. assert (h1:le_length_ba_conv_list _ _ (le_refl (length l)) = le_refl (length (ba_conv_list l))). apply proof_irrelevance. rewrite h1. tauto. Qed. Lemma has_comps_p_iff0 : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps_p _ _ pfm P a <-> @has_comps (ba_conv Bp) l _ pfm P a. intros; rewrite has_comps_p_iff. assert (pfm = le_length_ba_conv_list l m pfm). apply proof_irrelevance. rewrite <- H; tauto. Qed. Lemma has_comps_p_iff0' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps_p' _ P a <-> @has_comps' (ba_conv Bp) l P a. unfold has_comps_p', has_comps'; intros; rewrite has_comps_p_iff0; tauto. Qed. Lemma has_sames_p_iff : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), let pfm' := le_length_ba_conv_list _ _ pfm in has_sames_p _ _ pfm P a <-> has_sames (ba_conv_list l) _ pfm' P a. intros Bp l m hm P a. split. intro h1. destruct h1 as [c d hc hd h0 h1 h2 h3 h4 h5]. econstructor. assumption. apply h1. apply h2. assumption. do 2 rewrite nth_lt_ba_conv_list in h4. erewrite nth_lt_functional3. rewrite h4 at 1. apply nth_lt_functional3. assumption. intro h1. destruct h1 as [c d hc hd h0 h1 h2 h3 h4 h5]. econstructor. assumption. apply h1. apply h2. assumption. do 2 rewrite nth_lt_ba_conv_list. erewrite nth_lt_functional3. rewrite h4 at 1. apply nth_lt_functional3. assumption. Grab Existential Variables. assumption. assumption. assumption. assumption. Qed. Lemma has_sames_p_iff' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames_p' _ P a <-> has_sames' (ba_conv_list l) P a. unfold has_sames_p', has_comps'; intros. rewrite has_sames_p_iff. assert (h1:le_length_ba_conv_list _ _ (le_refl (length l)) = le_refl (length (ba_conv_list l))). apply proof_irrelevance. rewrite h1. tauto. Qed. Lemma has_sames_p_iff0 : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames_p _ _ pfm P a <-> @has_sames (ba_conv Bp) l _ pfm P a. intros; rewrite has_sames_p_iff. assert (pfm = le_length_ba_conv_list l m pfm). apply proof_irrelevance. rewrite <- H; tauto. Qed. Lemma has_sames_p_iff0' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames_p' _ P a <-> @has_sames' (ba_conv Bp) l P a. unfold has_sames_p', has_sames'; intros; rewrite has_sames_p_iff0; tauto. Qed. Lemma has_comps_p_linds_occ_impl : forall {Bp:Bool_Alg_p T} (la:list (btp Bp)) (m:nat) (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (x:btp Bp), has_comps_p la m pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m la) x)) F -> has_comps_p la m pfm (seg m) F. intros Bp la m h1 F x h2. pose proof (has_comps_linds_occ_impl (ba_conv_list la) m (le_length_ba_conv_list _ _ h1) F) as h3. rewrite linds_occ_ba_conv_list in h2. rewrite has_comps_p_iff in h2. specialize (h3 x h2). rewrite has_comps_p_iff; auto. Qed. Lemma has_comps_p_linds_occ_impl' : forall {Bp:Bool_Alg_p T} (la:list (btp Bp)) (F:Fin_map (seg (length la)) signe mns) (x:btp Bp), has_comps_p' la (list_to_set (linds_occ ba_eq_dec_p la x)) F -> has_comps_p' la (seg (length la)) F. unfold has_comps_p'; intros; eapply has_comps_p_linds_occ_impl. rewrite firstn_length. apply H. Qed. Lemma has_comps_p_S_linds_occ_in : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns), has_comps_p l (S m) pfsm (list_to_set (linds_occ ba_eq_dec_p (firstn (S m) l) (nth_lt l m pfsm))) F -> In (nth_lt l m pfsm) (firstn m l). intros Bp l m h1 F h2. rewrite has_comps_p_iff in h2. pose proof (has_comps_S_linds_occ_in (ba_conv_list l) m (le_length_ba_conv_list _ _ h1) F) as h3. rewrite nth_lt_ba_conv_list in h2. rewrite linds_occ_ba_conv_list in h2. erewrite nth_lt_functional3 in h2. specialize (h3 h2). erewrite nth_lt_functional3. apply h3. Qed. Lemma has_comps_elt_p_dec : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg m) signe mns), Finite P -> {has_comps_elt_p l m pfm P x a} + {~has_comps_elt_p l m pfm P x a}. intros Bp l m h1 P x a h2. pose proof (has_comps_elt_dec (ba_conv_list l) m h1 P x a h2) as h3. destruct h3 as [h3 | h3]. left. pose proof (has_comps_elt_p_iff l m h1 P x a) as h4. rewrite <- h4 in h3; auto. right. contradict h3. pose proof (has_comps_elt_p_iff l m h1 P x a) as h4. rewrite h4 in h3; auto. Qed. Lemma has_comps_elt_p_dec' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (C:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns), Finite C -> {has_comps_elt_p' l C x a} + {~has_comps_elt_p' l C x a}. intros; unfold has_comps_elt_p'; apply has_comps_elt_p_dec; auto. Qed. Lemma has_comps_elt_p_S : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfsm:S m <= length l) (P:Ensemble nat) (x:btp Bp) (a:Fin_map (seg (S m)) signe mns), let pfm := le_trans _ _ _ (le_n_Sn m) pfsm in has_comps_elt_p l m pfm P x (fin_map_seg_rest a m (lt_n_Sn _)) -> has_comps_elt_p l (S m) pfsm P x a. intros Bp l m h1 P x a h2 h3. pose proof (has_comps_elt_S (ba_conv_list l) m h1 P x a) as h4. simpl in h4; rewrite has_comps_elt_p_iff in h3; auto. specialize (h4 h3); auto. rewrite has_comps_elt_p_iff_b2; auto. Qed. Lemma has_comps_p_finite : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps_p l m pfm P a -> Finite P. intros ? ? ? ? ? ? h1. destruct h1; auto. Qed. Lemma has_comps_p_finite' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps_p' l P a -> Finite P. intros; eapply has_comps_p_finite. apply H. Qed. Lemma has_sames_p_finite : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames_p l m pfm P a -> Finite P. intros ? ? ? ? ? ? h1. destruct h1; auto. Qed. Lemma has_sames_p_finite' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames_p' l P a -> Finite P. intros; eapply has_sames_p_finite. apply H. Qed. Lemma has_comps_p_gt_0 : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps_p _ _ pfm P a -> 0 < m. intros; eapply has_comps_gt_0. erewrite <- has_comps_p_iff0 at 1. apply H. Qed. Lemma has_comps_gt_0_p' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps_p' _ P a -> l <> nil. intros Bp l P a h1. destruct l as [|c le]. inversion h1. simpl in pfc; omega. intro; subst; discriminate. Qed. Lemma has_sames_p_gt_0 : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames_p _ _ pfm P a -> 0 < m. intros; eapply has_sames_gt_0. erewrite <- has_sames_p_iff0 at 1. apply H. Qed. Lemma has_sames_gt_0_p' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames_p' _ P a -> l <> nil. intros Bp l P a h1. destruct l as [|c le]. inversion h1. simpl in pfc; omega. intro; subst; discriminate. Qed. Lemma has_comps_p_inh : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps_p _ _ pfm P a -> Inhabited P. intros B l m hm P a h1. destruct h1. econstructor. apply H0. Qed. Lemma has_comps_p_inh' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps_p' _ P a -> Inhabited P. intros ? ? ? ? h1; red in h1; apply has_comps_p_inh in h1; auto. Qed. Lemma has_sames_p_inh : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames_p _ _ pfm P a -> Inhabited P. intros B l m hm P a h1. destruct h1. econstructor. apply H0. Qed. Lemma has_sames_p_inh' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames_p' _ P a -> Inhabited P. intros ? ? ? ? h1; red in h1; apply has_sames_p_inh in h1; auto. Qed. Lemma has_comps_p_card : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps_p _ _ pfm P a -> 1 < card_fun1 P. intros Bp l m hm P a h1. destruct h1. eapply in2_card. apply H2. assumption. assumption. assumption. Qed. Lemma has_comps_p_card' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps_p' _ P a -> 1 < card_fun1 P. unfold has_comps_p'; intros; eapply has_comps_p_card. apply H. Qed. Lemma has_sames_p_card : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames_p _ _ pfm P a -> 1 < card_fun1 P. intros Bp l m hm P a h1. destruct h1. eapply in2_card. apply H2. assumption. assumption. assumption. Qed. Lemma has_sames_p_card' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames_p' _ P a -> 1 < card_fun1 P. unfold has_sames_p'; intros; eapply has_sames_p_card. apply H. Qed. Lemma has_comps_iff_p : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps_p _ _ pfm P a <-> Finite P /\ (exists (x:btp Bp) (r:nat) (pfr:r < m), let pfr := lt_le_trans _ _ _ pfr pfm in Inn P r /\ nth_lt l r pfr = x /\ has_comps_elt_p _ _ pfm P x a). intros Bp l m h1 P a. pose proof (has_comps_iff (B:=ba_conv Bp) h1 P a) as h2. rewrite has_comps_p_iff0, h2; auto. intuition. destruct H4 as [x [r [h5 [h6 [h7 h8]]]]]. exists x, r, h5. intuition. rewrite has_comps_elt_p_iff; auto. destruct H3 as [x [r [h5 [h6 [h7 h8]]]]]. exists x, r, h5. intuition. rewrite <- has_comps_elt_p_iff; auto. Qed. Lemma has_comps_iff_p' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps_p' _ P a <-> Finite P /\ (exists (x:btp Bp) (r:nat) (pfr:r < length l), Inn P r /\ nth_lt l r pfr = x /\ has_comps_elt_p' _ P x a). unfold has_comps_p'. intros; rewrite has_comps_iff_p. intuition. destruct H1 as [x [i [h3 [h4 [h5 h6]]]]]. exists x, i, h3. split; auto. split; auto. rewrite <- h5. f_equal; apply proof_irrelevance. destruct H1 as [x [r [h2 [h3 [h4 h5]]]]]. subst. exists (nth_lt l r h2), _, h2. split; auto. split. apply nth_lt_functional3. unfold has_comps_elt_p' in h5; auto. Qed. Lemma has_sames_iff_p : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames_p _ _ pfm P a <-> Finite P /\ (exists (x:btp Bp) (r:nat) (pfr:r < m), let pfr := lt_le_trans _ _ _ pfr pfm in Inn P r /\ nth_lt l r pfr = x /\ has_sames_elt_p _ _ pfm P x a). intros Bp l m h1 P a. pose proof (has_sames_iff (B:=ba_conv Bp) h1 P a) as h2. rewrite has_sames_p_iff0, h2. intuition. destruct H4 as [x [r [h5 [h6 [h7 h8]]]]]. exists x, r, h5. intuition. rewrite has_sames_elt_p_iff; auto. destruct H3 as [x [r [h5 [h6 [h7 h8]]]]]. exists x, r, h5. intuition. rewrite <- has_sames_elt_p_iff; auto. Qed. Lemma has_sames_iff_p' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames_p' _ P a <-> Finite P /\ (exists (x:btp Bp) (r:nat) (pfr:r < length l), Inn P r /\ nth_lt l r pfr = x /\ has_sames_elt_p' _ P x a). unfold has_sames_p'. intros; rewrite has_sames_iff_p. intuition. destruct H1 as [x [i [h3 [h4 [h5 h6]]]]]. exists x, i, h3. split; auto. split; auto. rewrite <- h5. f_equal; apply proof_irrelevance. destruct H1 as [x [r [h2 [h3 [h4 h5]]]]]. subst. exists (nth_lt l r h2), _, h2. split; auto. split. apply nth_lt_functional3. unfold has_sames_elt_p' in h5; auto. Qed. Lemma has_comps_p_incl : forall {l:list Atp} {m:nat} (pfm:m <= length l) (C C':Ensemble nat) (F:Fin_map (seg m) signe mns), Included C C' -> Finite C' -> has_comps_p _ _ pfm C F -> has_comps_p _ _ pfm C' F. intros l m hm C C' F h1 hf h2. destruct h2 as [r s hr hs h0 h4 h5 h6 h7]. apply h1 in h4. apply h1 in h5. econstructor. assumption. apply h4. apply h5. assumption. apply h7. assumption. Qed. Lemma has_comps_p_incl' : forall {l:list Atp} (C C':Ensemble nat) (F:Fin_map (seg (length l)) signe mns), Included C C' -> Finite C' -> has_comps_p' _ C F -> has_comps_p' _ C' F. unfold has_comps_p'; intros; eapply has_comps_p_incl; auto. apply H. assumption. Qed. Lemma has_comps_p_rest_impl : forall {l:list Atp} (m:nat) (pfm : m <= length l) (pfsm : S m <= length l) (C:Ensemble nat) (F:Fin_map (seg (S m)) signe mns), has_comps_p _ _ pfm C (fin_map_seg_rest F m (lt_n_Sn m)) -> has_comps_p _ _ pfsm (Add C m) F. intros l m h1 h2 C F h3. inversion h3 as [r s hr hs hf h4 h5 h6 h7 h8]. clear h3. do 2 rewrite fin_map_seg_rest_compat in h8; auto. simpl in hr, hs. econstructor; simpl. apply Add_preserves_Finite; auto. left. apply h4. left. apply h5. assumption. simpl in h7. erewrite nth_lt_functional3. rewrite h7 at 1. apply nth_lt_functional3. assumption. Grab Existential Variables. simpl; apply lt_S; auto. simpl; apply lt_S; auto. Qed. Lemma has_comps_p_rest_impl_final : forall {l:list Atp} (m:nat) (pfm : m <= length l) (pfsm : S m <= length l) (C:Ensemble nat) (F:Fin_map (seg (S m)) signe mns), Finite C -> has_comps_p _ _ pfm C (fin_map_seg_rest F m (lt_n_Sn m)) -> has_comps_p _ _ pfsm C F. intros l m h1 h2 C F hc h3. destruct h3 as [r s hr hs hf h4 h5 h6 h7 h8]. simpl in hr, hs. do 2 rewrite fin_map_seg_rest_compat in h8; auto. econstructor; simpl. assumption. apply h4. apply h5. assumption. simpl in h7. erewrite nth_lt_functional3. rewrite h7 at 1. apply nth_lt_functional3. assumption. Grab Existential Variables. simpl. apply lt_S; auto. simpl. apply lt_S; auto. Qed. Lemma nin_has_comps_p_impl_removelast : forall {l:list Atp} {m:nat} (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns), ~In (nth_lt l m pfsm) (firstn m l) -> forall x:Atp, In x (firstn m l) -> has_comps_p l (S m) pfsm (list_to_set (linds_occ ba_eq_dec_p (firstn (S m) l) x)) F -> let pfm := le_length_removelast l m pfsm in has_comps_p (removelast l) m pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m (removelast l)) x)) (fin_map_seg_rest F m (lt_n_Sn _)). intros l m h1 F h2 x h3 h4 h5. rewrite has_comps_p_iff0, linds_occ_ba_conv_list in h4. rewrite has_comps_p_iff0, linds_occ_ba_conv_list. apply (nin_has_comps_impl_removelast (A:=ba_conv Ap) h1 F h2 x h3 h4). Qed. Lemma has_comps_p_removelast_impl : forall {l:list Atp} (m:nat) (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns) (x:Atp), let pfm := le_length_removelast l m pfsm in In x (firstn m l) -> ~ In (nth_lt l m pfsm) (firstn m l) -> has_comps_p (removelast l) m pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m l) x)) (fin_map_seg_rest F m (lt_n_Sn m)) -> has_comps_p l (S m) pfsm (list_to_set (linds_occ ba_eq_dec_p (firstn (S m) l) x)) F. intros l m h1 F x h2 hi hn h3. rewrite has_comps_p_iff0, linds_occ_ba_conv_list' in h3. rewrite has_comps_p_iff0, linds_occ_ba_conv_list'. apply (@has_comps_removelast_impl (ba_conv Ap) l m h1 F x hi hn h3). Qed. Lemma has_comps_p_removelast_impl' : forall (l:list Atp) (m:nat) (pfm:m <= length (removelast l)) (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns), has_comps_p (removelast l) m pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m (removelast l)) (nth_lt l m pfsm))) (fin_map_seg_rest F m (lt_n_Sn m)) -> has_comps_p l (S m) pfsm (list_to_set (linds_occ ba_eq_dec_p (firstn (S m) l) (nth_lt l m pfsm))) F. intros l m h1 h2 F h3. rewrite has_comps_p_iff0, linds_occ_ba_conv_list' in h3. rewrite has_comps_p_iff0, linds_occ_ba_conv_list'. apply (@has_comps_removelast_impl' (ba_conv Ap) l m h1 h2 F h3). Qed. Lemma has_comps_p_removelast_impl'' : forall (l:list Atp) (m:nat) (pfm:m <= length (removelast l)) (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns) (x:Atp), has_comps_p (removelast l) m pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m (removelast l)) x)) (fin_map_seg_rest F m (lt_n_Sn m)) -> has_comps_p l (S m) pfsm (list_to_set (linds_occ ba_eq_dec_p (firstn (S m) l) x)) F. intros l m h1 h2 F x h3. rewrite has_comps_p_iff0, linds_occ_ba_conv_list' in h3. rewrite has_comps_p_iff0, linds_occ_ba_conv_list'. apply (@has_comps_removelast_impl'' (ba_conv Ap) l m h1 h2 F x h3). Qed. Lemma has_comps_p_removelast_iff : forall (l:list Atp) (m:nat) (pfm: m <= length l) (pfm':m <= length (removelast l)) (x:Atp) (F:Fin_map (seg m) signe mns), has_comps_p (removelast l) m pfm' (list_to_set (linds_occ ba_eq_dec_p (firstn m l) x)) F <-> has_comps_p l m pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m l) x)) F. intros l m h1 h2 x F. pose proof (@has_comps_removelast_iff (ba_conv Ap) l m h1 h2 x F) as h4. rewrite has_comps_p_iff0, linds_occ_ba_conv_list'. unfold Atp in h4; unfold Atp. rewrite h4. rewrite has_comps_p_iff0; tauto. Qed. Lemma impl_has_comps_p_removelast_neq : forall (l:list Atp) (m:nat) (pfm:m <= length (removelast l)) (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns) (x:Atp), x <> nth_lt l m pfsm -> has_comps_p l (S m) pfsm (list_to_set (linds_occ ba_eq_dec_p (firstn (S m) l) x)) F -> has_comps_p (removelast l) m pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m (removelast l)) x)) (fin_map_seg_rest F m (lt_n_Sn m)). intros l m h1 h2 F x h3 h4. rewrite has_comps_p_iff0, linds_occ_ba_conv_list' in h4. rewrite has_comps_p_iff0, linds_occ_ba_conv_list'. apply (@impl_has_comps_removelast_neq (ba_conv Ap) l m h1 h2 F x h3 h4). Qed. Lemma has_comps_p_dec : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (C:Ensemble nat) (F:Fin_map (seg m) signe mns), Finite C -> {has_comps_p _ _ pfm C F} + {~ has_comps_p _ _ pfm C F}. intros Bp l m h1 C F h2. pose proof h1 as h1'. pose proof (@has_comps_dec (ba_conv Bp) (ba_conv_list l) m (le_length_ba_conv_list _ _ h1) F C h2) as h3. destruct h3 as [h3 | h3]. left. rewrite has_comps_p_iff; auto. right. rewrite has_comps_p_iff; auto. Qed. Lemma has_comps_p_dec' : forall {Bp: Bool_Alg_p T} {l:list (btp Bp)} (F:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> {has_comps_p' _ P F} + {~has_comps_p' _ P F}. unfold has_comps_p'; intros; apply has_comps_p_dec; auto. Qed. Lemma has_sames_p_dec : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (C:Ensemble nat) (F:Fin_map (seg m) signe mns), Finite C -> {has_sames_p _ _ pfm C F} + {~ has_sames_p _ _ pfm C F}. intros Bp l m h1 C F h2. pose proof h1 as h1'. pose proof (@has_sames_dec (ba_conv Bp) (ba_conv_list l) m (le_length_ba_conv_list _ _ h1) F C h2) as h3. destruct h3 as [h3 | h3]. left. rewrite has_sames_p_iff; auto. right. rewrite has_sames_p_iff; auto. Qed. Lemma has_sames_p_dec' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (F:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> {has_sames_p' _ P F} + {~has_sames_p' _ P F}. unfold has_sames_p'; intros; apply has_sames_p_dec; auto. Qed. Lemma has_comps_p_lt_impl : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m k:nat) (pfm : m <= length l) (pfk : k < m) (C:Ensemble nat) (F:Fin_map (seg m) signe mns), Finite C -> let pfk' := lt_le_trans _ _ _ pfk pfm in let pfk'' := lt_le_weak _ _ pfk' in has_comps_p _ _ pfm C F -> has_comps_p _ _ pfk'' C (fin_map_seg_rest F k pfk) \/ (exists (r s:nat) (pfr:r < k) (pfks:k <= s) (pfs:s < m), let pfr' := lt_le_trans _ _ _ (lt_trans _ _ _ (lt_le_trans _ _ _ pfr pfks) pfs) pfm in let pfs' := lt_le_trans _ _ _ pfs pfm in Inn C r /\ Inn C s /\ nth_lt l r pfr' = nth_lt l s pfs' /\ F |-> r <> F |-> s) \/ (exists (r s:nat) (pfr:k <= r) (pfr:r < s) (pfs:s < m), let pfr' := lt_le_trans _ _ _ (lt_trans _ _ _ pfr pfs) pfm in let pfs' := lt_le_trans _ _ _ pfs pfm in Inn C r /\ Inn C s /\ nth_lt l r pfr' = nth_lt l s pfs' /\ F |-> r <> F |-> s). intros Bp l m k hm hk C F hf hk' hk'' h3. destruct h3 as [r s h6 h7 h0 h8 h9 h10 h11 h12]. destruct (lt_le_dec r k) as [h13 | h13]; subst. destruct (lt_le_dec s k) as [h14 | h14]; subst. left. econstructor. assumption. apply h8. apply h9. assumption. erewrite nth_lt_functional3. rewrite h11 at 1. apply nth_lt_functional3. rewrite fin_map_seg_rest_compat. rewrite fin_map_seg_rest_compat; auto. assumption. right. left. exists r, s, h13, h14, h7. simpl. repeat split; auto. erewrite nth_lt_functional3. unfold Atp. rewrite h11 at 1. apply nth_lt_functional3. destruct (lt_le_dec s k) as [h14 | h14]. right. left. exists s, r, h14, h13, h6. simpl. apply neq_sym in h12. repeat split; auto. erewrite nth_lt_functional3. unfold Atp. rewrite h11 at 1. apply nth_lt_functional3. right. right. destruct (lt_le_dec r s) as [h15 | h15]. exists r, s, h13, h15, h7. simpl. repeat split; auto. erewrite nth_lt_functional3. unfold Atp. rewrite h11 at 1. apply nth_lt_functional3. exists s, r, h14. ex_goal. omega. exists hex, h6. simpl. repeat split; auto. erewrite nth_lt_functional3. unfold Atp. rewrite h11 at 1. apply nth_lt_functional3. Grab Existential Variables. omega. omega. assumption. assumption. Qed. Lemma has_comps_p_S_impl : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (m:nat) (pfm : m <= length l) (pfsm : S m <= length l) (C:Ensemble nat) (F:Fin_map (seg (S m)) signe mns), Finite C -> has_comps_p _ _ pfsm C F -> (has_comps_p _ _ pfm C (fin_map_seg_rest F m (lt_n_Sn m))) \/ (Inn C m /\ (exists (r:nat) (pfr:r < m), Inn C r /\ nth_lt l r (lt_le_trans _ _ _ pfr pfm) = nth_lt l m (lt_le_trans _ _ _ (lt_n_Sn m) pfsm) /\ F |-> r <> F |-> m)). intros Bp l m h1 h2 C F hf h3. rewrite has_comps_p_iff0 in h3. rewrite has_comps_p_iff0. apply (@has_comps_S_impl (ba_conv Bp) l m h1 h2 C F hf h3). Qed. Lemma has_comps_p_S_impl_dec : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (m:nat) (pfm : m <= length l) (pfsm : S m <= length l) (C:Ensemble nat) (F:Fin_map (seg (S m)) signe mns), Finite C -> has_comps_p _ _ pfsm C F -> {has_comps_p _ _ pfm C (fin_map_seg_rest F m (lt_n_Sn m))} + {Inn C m /\ (exists (r:nat) (pfr:r < m), Inn C r /\ nth_lt l r (lt_le_trans _ _ _ pfr pfm) = nth_lt l m (lt_le_trans _ _ _ (lt_n_Sn m) pfsm) /\ F |-> r <> F |-> m)}. intros Bp l m h1 h2 C F hf h3. destruct (has_comps_p_dec h1 C (fin_map_seg_rest F m (lt_n_Sn m)) hf) as [h4 | h4]. left; auto. right. pose proof (has_comps_p_S_impl m h1 h2 C F hf h3) as h5. destruct h5; try contradiction; auto. Qed. Lemma has_comps_p_last_dec : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (m:nat) (pfm : m <= length l) (pfsm : S m <= length l) (C:Ensemble nat) (pfc:Finite C) (F:Fin_map (seg (S m)) signe mns), {Inn C m /\ (exists (r:nat) (pfr:r < m), Inn C r /\ nth_lt l r (lt_le_trans _ _ _ pfr pfm) = nth_lt l m (lt_le_trans _ _ _ (lt_n_Sn m) pfsm) /\ F |-> r <> F |-> m)} + {~ (Inn C m /\ (exists (r:nat) (pfr:r < m), Inn C r /\ nth_lt l r (lt_le_trans _ _ _ pfr pfm) = nth_lt l m (lt_le_trans _ _ _ (lt_n_Sn m) pfsm) /\ F |-> r <> F |-> m))}. intros Bp l m h1 h2 C hfc F. destruct (in_fin_ens_dec nat_eq_dec _ hfc m) as [h3 | h3]. pose (fun r pfr => Inn C r /\ nth_lt l r (lt_le_trans _ _ _ pfr h1) = nth_lt l m (lt_le_trans _ _ _ (lt_n_Sn _) h2) /\ F |-> r <> F |-> m) as R. pose proof (lt_ex_dep_nat_dec R) as h4. hfirst h4. red; unfold R. intros r h5. destruct (in_fin_ens_dec nat_eq_dec _ hfc r) as [h6 | h6]. destruct (ba_eq_dec_p (nth_lt l r (lt_le_trans r m (length l) h5 h1)) (nth_lt l m (lt_le_trans m (S m) (length l) (lt_n_Sn m) h2))) as [h7 | h7]. destruct (sign_eq_dec (F |-> r) (F |-> m)) as [h8 | h8]. right. intro h9. destruct h9 as [? [? ?]]. contradiction. left. repeat split; auto. right. intro h9. destruct h9 as [? [? ?]]. contradiction. right. intro h9. destruct h9 as [? [? ?]]. contradiction. specialize (h4 hf). unfold R in h4. destruct h4 as [h4 | h4]. left. split; auto. right. intro h9. destruct h9 as [? [? ?]]. contradict h4. exists x; auto. right. intro h4. destruct h4; contradiction. Qed. Lemma has_comps_p_add_impl : forall {Bp:Bool_Alg_p T} (l : list (btp Bp)) (m c: nat) (pfm : m <= length l) (pfc:c < length l) (C : Ensemble nat) (F : Fin_map (seg m) signe mns), Finite C -> has_comps_p _ _ pfm (Add C c) F -> has_comps_p _ _ pfm C F \/ (exists (r : nat) (pfr : r < m), Inn C r /\ nth_lt l r (lt_le_trans r m (length l) pfr pfm) = nth_lt l c pfc /\ F |-> r <> F |-> c). intros Bp l m c h1 h2 C F h3 h4. inversion h4. simpl in pfc, pfd, H3, H4. destruct H0 as [r h5 | r h5], H1 as [d h6 | d h6]. left. econstructor. assumption. apply h5. apply h6. assumption. simpl. apply H3. assumption. destruct h6. right. exists r, pfc. split; auto. split; auto. unfold Atp in H3. unfold Atp. rewrite H3 at 1. apply nth_lt_functional3. destruct h5. right. exists d, pfd. split; auto. split; auto. unfold Atp in H3. unfold Atp. rewrite <- H3 at 1. apply nth_lt_functional3. destruct h5, h6. contradict H2; auto. Qed. Lemma has_comps_p_add_list_to_set : forall {Bp:Bool_Alg_p T} (l : list (btp Bp)) (m c: nat) (pfm : m <= length l) (pfc:c < m) (ln : list nat) (F : Fin_map (seg m) signe mns), In ln (linds ba_eq_dec_p (firstn m l)) -> has_comps_p _ _ pfm (Add (list_to_set ln) c) F -> has_comps_p _ _ pfm (list_to_set ln) F. intros Bp l m c h1 h2 ln F h3 h4. pose proof (has_comps_p_add_impl l m c h1) as h5. hfirst h5. eapply lt_le_trans. apply h2. assumption. specialize (h5 hf (list_to_set _) F (list_to_set_finite _) h4). destruct h5 as [| h5]; auto. destruct h5 as [r [h5 [h6 [h7 h8]]]]. rewrite <-list_to_set_in_iff in h6. pose proof (nth_lt_eq_in_linds ba_eq_dec_p (firstn m l) ln r c) as h9. hfirst h9. rewrite length_firstn; auto. specialize (h9 hf0). hfirst h9. rewrite length_firstn; auto. specialize (h9 hf1 h6). hfirst h9. do 2 rewrite nth_lt_firstn_eq. erewrite nth_lt_functional3. rewrite h7 at 1. apply nth_lt_functional3. specialize (h9 hf2 h3). econstructor. apply list_to_set_finite. rewrite <- list_to_set_in_iff. apply h6. rewrite <- list_to_set_in_iff. apply h9. intro; subst. contradict h8; auto. erewrite nth_lt_functional3. unfold Atp in h7. unfold Atp. rewrite h7 at 1. apply nth_lt_functional3. assumption. Grab Existential Variables. simpl. assumption. simpl. assumption. Qed. Lemma has_comps_p_add_list_to_set' : forall {Bp:Bool_Alg_p T} (l : list (btp Bp)) (c: nat) (pfc:c < length l) (ln : list nat) (F : Fin_map (seg (length l)) signe mns), In ln (linds ba_eq_dec_p l) -> has_comps_p' _ (Add (list_to_set ln) c) F -> has_comps_p' _ (list_to_set ln) F. unfold has_comps_p'; intros Bp l c h1 ln F h2 h3. eapply has_comps_p_add_list_to_set. apply h1. rewrite firstn_length; auto. assumption. Qed. Lemma has_comps_dup_p : forall {Bp:Bool_Alg_p T} (m:nat) (l:list (btp Bp)) (pfm:m <= length l) (C:Ensemble nat) (a:Fin_map (seg m) signe mns), has_comps_p _ _ pfm C a -> ~NoDup l. intros Bp m l pfm C a h1. destruct h1 as [r s h1 h2 h3]. eapply nth_lt_dup'. apply H1. apply H2. Qed. Lemma has_comps_dup_p' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (C:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_comps_p' _ C a -> ~NoDup l. unfold has_comps_p'; intros Bp l C a h1. eapply has_comps_dup_p. apply h1. Qed. Lemma has_sames_dup_p : forall {Bp:Bool_Alg_p T} (m:nat) (l:list (btp Bp)) (pfm:m <= length l) (C:Ensemble nat) (a:Fin_map (seg m) signe mns), has_sames_p _ _ pfm C a -> ~NoDup l. intros Bp m l pfm C a h1. destruct h1 as [r s h1 h2 h3]. eapply nth_lt_dup'. apply H1. apply H2. Qed. Lemma has_sames_dup_p' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (C:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), has_sames_p' _ C a -> ~NoDup l. unfold has_sames_p'; intros Bp l C a h1. eapply has_sames_dup_p. apply h1. Qed. Lemma not_has_comps_p_or_sames_p_nin_iff : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (a:Fin_map (seg m) signe mns) (P:Ensemble nat), Finite P -> (~has_comps_p _ _ pfm P a /\ ~has_sames_p _ _ pfm P a <-> (forall (r s:nat) (pfr:r < m) (pfs:s < m), r <> s -> let pfr' := lt_le_trans _ _ _ pfr pfm in let pfs' := lt_le_trans _ _ _ pfs pfm in nth_lt l r pfr' = nth_lt l s pfs' -> ~Inn P r \/ ~Inn P s)). intros Bp l m h1 a P h2. rewrite has_comps_p_iff0, has_sames_p_iff0. apply (@not_has_comps_or_sames_nin_iff (ba_conv Bp) (ba_conv_list l) m h1 a P h2). Qed. Lemma not_has_comps_p_or_sames_p_nin_iff' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (a:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> (~has_comps_p' _ P a /\ ~has_sames_p' _ P a <-> (forall (r s:nat) (pfr:r < length l) (pfs:s < length l), r <> s -> nth_lt l r pfr = nth_lt l s pfs -> ~Inn P r \/ ~Inn P s)). unfold has_comps_p', has_sames_p'; intros; rewrite not_has_comps_p_or_sames_p_nin_iff. simpl. intuition. eapply H0; auto. erewrite nth_lt_functional3. rewrite H2 at 1. apply nth_lt_functional3. eapply H0; auto. apply H2. assumption. Grab Existential Variables. assumption. assumption. Qed. Lemma not_has_comps_or_sames_elt_p_nin_iff : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (x:btp Bp) (a:Fin_map (seg m) signe mns) (P:Ensemble nat), Finite P -> (~has_comps_elt_p _ _ pfm P x a /\ ~has_sames_elt_p _ _ pfm P x a <-> (forall (r s: nat) (pfr: r < m) (pfs: s < m), let pfr' := lt_le_trans r m (length l) pfr pfm in let pfs' := lt_le_trans s m (length l) pfs pfm in r <> s -> nth_lt l r pfr' = x -> nth_lt l s pfs' = x -> ~ Inn P r \/ ~Inn P s)). intros Bp l m h1 x a P h2. pose proof (@not_has_comps_or_sames_elt_nin_iff (ba_conv Bp) (ba_conv_list l) m h1 x a P h2) as h3. pose proof (has_comps_elt_p_iff l m h1 P x a) as h4. rewrite h4. pose proof (has_sames_elt_p_iff l m h1 P x a) as h5. rewrite h5. intuition. Qed. Lemma not_has_comps_or_sames_elt_p_nin_iff' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> (~has_comps_elt_p' _ P x a /\ ~has_sames_elt_p' _ P x a <-> (forall (r s: nat) (pfr: r < length l) (pfs: s < length l), r <> s -> nth_lt l r pfr = x -> nth_lt l s pfs = x -> ~Inn P r \/ ~Inn P s)). unfold has_comps_elt_p', has_sames_elt_p'; intros Bp l x a P h1. erewrite not_has_comps_or_sames_elt_p_nin_iff; auto. simpl. intuition; subst. eapply H; auto. apply nth_lt_functional3. rewrite <- H2. apply nth_lt_functional3. eapply H; auto. rewrite <- H2. apply nth_lt_functional3. Grab Existential Variables. assumption. assumption. assumption. Qed. Lemma has_p_dec : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (a:Fin_map (seg m) signe mns) (P:Ensemble nat), Finite P -> {has_comps_p _ _ pfm P a} + {has_sames_p _ _ pfm P a} + {forall (r s:nat) (pfr:r < m) (pfs:s < m), r <> s -> let pfr' := lt_le_trans _ _ _ pfr pfm in let pfs' := lt_le_trans _ _ _ pfs pfm in nth_lt l r pfr' = nth_lt l s pfs' -> ~Inn P r \/ ~Inn P s}. intros Bp l m h1 a P h2. pose proof (@has_dec (ba_conv Bp) (ba_conv_list l) m h1 a P h2) as h3. destruct h3 as [[h3 | h3] | h3]. left; left. rewrite has_comps_p_iff0; auto. left; right. rewrite has_sames_p_iff0; auto. right; auto. Qed. Lemma has_p_dec' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (a:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> {has_comps_p' _ P a} + {has_sames_p' _ P a} + {forall (r s:nat) (pfr:r < length l) (pfs:s < length l), r <> s -> nth_lt l r pfr = nth_lt l s pfs -> ~Inn P r \/ ~Inn P s}. unfold has_comps_p', has_sames_p'; intros Bp l a P h1. pose proof (has_p_dec l (length l) (le_refl _) a P h1) as h2. intuition. right. intros r s h2 h3 h4 h5. specialize (b _ _ h2 h3 h4). hfirst b. erewrite nth_lt_functional3. rewrite h5 at 1. apply nth_lt_functional3. apply b; auto. Qed. Lemma has_elt_p_dec : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (x:btp Bp) (a:Fin_map (seg m) signe mns) (P:Ensemble nat), Finite P -> {has_comps_elt_p _ _ pfm P x a} + {has_sames_elt_p _ _ pfm P x a} + {(forall (r s: nat) (pfr: r < m) (pfs: s < m), let pfr' := lt_le_trans r m (length l) pfr pfm in let pfs' := lt_le_trans s m (length l) pfs pfm in r <> s -> nth_lt l r pfr' = x -> nth_lt l s pfs' = x -> ~ Inn P r \/ ~Inn P s)}. intros Bp l m h1 x a P h2. pose proof (@has_elt_dec (ba_conv Bp) (ba_conv_list l) m h1 (ba_conv_elt x) a P h2) as h3. destruct h3 as [[h3 | h3] | h3]. left; left. pose proof (has_comps_elt_p_iff l m h1 P x a) as h4. rewrite h4; auto. left; right. pose proof (has_sames_elt_p_iff l m h1 P x a) as h4. rewrite h4; auto. right; auto. Qed. Lemma has_elt_p_dec' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (x:btp Bp) (a:Fin_map (seg (length l)) signe mns) (P:Ensemble nat), Finite P -> {has_comps_elt_p' _ P x a} + {has_sames_elt_p' _ P x a} + {(forall (r s: nat) (pfr: r < length l) (pfs: s < length l), r <> s -> nth_lt l r pfr = x -> nth_lt l s pfs = x -> ~ Inn P r \/ ~Inn P s)}. unfold has_comps_elt_p'; intros Bp l x a P h1. pose proof (has_elt_p_dec l (length l) (le_refl _) x a P h1) as h2. destruct h2 as [[h2 | h2] | h2]. left; auto. left; right; auto. right. intros; subst; eapply h2; auto. apply nth_lt_functional3. rewrite <- H1. apply nth_lt_functional3. Grab Existential Variables. assumption. assumption. Qed. Lemma not_has_comps_elt_const_p : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (x:btp Bp) (F:Fin_map (seg m) signe mns), Finite P -> ~has_comps_elt_p _ _ pfm P x F -> forall (c d:nat) (pfc:c < m) (pfd: d < m), let pfc' := lt_le_trans _ _ _ pfc pfm in let pfd' := lt_le_trans _ _ _ pfd pfm in c <> d -> Inn P c -> Inn P d -> nth_lt l c pfc' = x -> nth_lt l d pfd' = x -> F |-> c = F |-> d. intros Bp l m h1 P x F h2 h3 c d hc hd h4 h5 h6 h7 h8 h9 h10; subst. pose proof (has_comps_elt_p_iff l m h1 P (nth_lt l c h4) F) as h11. rewrite h11 in h3. pose proof @not_has_comps_elt_const. apply (@not_has_comps_elt_const (ba_conv Bp) (ba_conv_list l) m h1 P (nth_lt l c h4) F h2 h3 c d hc hd h6 h7 h8). apply nth_lt_functional3. rewrite <- h10. apply nth_lt_functional3. Qed. Lemma not_has_comps_elt_const_p' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (x:btp Bp) (F:Fin_map (seg (length l)) signe mns), Finite P -> ~has_comps_elt_p' _ P x F -> forall (c d:nat) (pfc:c < length l) (pfd: d < length l), c <> d -> Inn P c -> Inn P d -> nth_lt l c pfc = x -> nth_lt l d pfd = x -> F |-> c = F |-> d. intros. subst. eapply not_has_comps_elt_const_p. apply H. apply H0. assumption. assumption. assumption. apply nth_lt_functional3. rewrite <- H5. apply nth_lt_functional3. Grab Existential Variables. assumption. assumption. Qed. (*Uses [NNPP], there [LEM].*) Lemma not_has_comps_nth_eq_differ : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m c d:nat) (pfm:m <= length l) (pfc:c < length (firstn m l)) (pfd:d < length (firstn m l)) (P:Ensemble nat) (F:Fin_map (seg m) signe mns), Finite P -> ~has_comps_p _ _ pfm P F -> Inn P c -> Inn P d -> nth_lt (firstn m l) c pfc = nth_lt (firstn m l) d pfd -> F |-> c = F |-> d. intros Bp l m c d h1 h2 h3 P F h4 h5 h6 h7 h8. pose proof h2 as h2'. pose proof h3 as h3'. rewrite length_firstn in h2'; auto. rewrite length_firstn in h3'; auto. destruct (nat_eq_dec c d) as [hn | hn]; subst; auto. apply NNPP. contradict h5. econstructor. assumption. apply h6. apply h7. assumption. simpl. pose proof (nth_lt_firstn_eq l c m h2) as h10. simpl in h10. symmetry in h10. unfold Atp in h10, h8. unfold Atp. simpl. simpl in h10, h8. erewrite nth_lt_functional3 in h10. rewrite h10 at 1. rewrite h8 at 1. rewrite nth_lt_firstn_eq. apply nth_lt_functional3. assumption. Grab Existential Variables. simpl. assumption. simpl. assumption. Qed. Lemma not_has_comps_nth_eq_differ' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (c d:nat) (pfc:c < length l) (pfd:d < length l) (P:Ensemble nat) (F:Fin_map (seg (length l)) signe mns), Finite P -> ~has_comps_p' _ P F -> Inn P c -> Inn P d -> nth_lt l c pfc = nth_lt l d pfd -> F |-> c = F |-> d. intros; eapply not_has_comps_nth_eq_differ. apply H. apply H0. assumption. assumption. do 2 erewrite nth_lt_firstn_eq. erewrite nth_lt_functional3. rewrite H3 at 1. apply nth_lt_functional3. Grab Existential Variables. rewrite length_firstn; auto. rewrite length_firstn; auto. Qed. Lemma has_comps_p_subtract : forall {Bp:Bool_Alg_p T} (l:list (btp Bp) ) (E:Ensemble nat) (m:nat) (pf:m <= length l) (F:Fin_map (seg m) signe mns), has_comps_p _ _ pf E F -> has_comps_p _ _ pf (Subtract E m) F. intros Bp l E m h1 F h2. destruct h2. econstructor. apply subtract_preserves_finite; auto. constructor. apply H0. simpl in pfc. rewrite in_sing_iff. omega. constructor. apply H1. rewrite in_sing_iff. simpl in pfd. omega. assumption. simpl. simpl in H3. apply H3. assumption. Qed. (*Returns a list of all elements that are the rightmost index of a pair of indices at which [l] has complements.*) Fixpoint has_comps_p_to_list_ind {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} {P:Ensemble nat} : forall {a:Fin_map (seg m) signe mns} {pfm:m <= length l}, has_comps_p _ _ pfm P a -> list nat := match m with | 0 => fun _ _ _ => nil | S m' => fun a pfsm' pfh => let pff := has_comps_p_finite _ _ _ P a pfh in let pfm' := le_trans _ _ _ (le_n_Sn _) pfsm' in let pfhld := has_comps_p_last_dec m' pfm' pfsm' P pff a in let pfhd := has_comps_p_S_impl_dec m' pfm' pfsm' P a pff pfh in match pfhld with | left pfe => match pfhd with | left pfph => m' :: has_comps_p_to_list_ind pfph | right _ => m' :: nil end | right pfne => let pfph := dec_not_r pfhd pfne in has_comps_p_to_list_ind pfph end end. Lemma in_has_comps_p_to_list_ind_lt : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} {P:Ensemble nat} {a:Fin_map (seg m) signe mns} {pfm:m <= length l} {pfh:has_comps_p _ _ pfm P a} (n:nat), In n (has_comps_p_to_list_ind pfh) -> n < m. intros Bp l m. induction m as [|m ih]; simpl. intros; contradiction. intros P a h1 h2 n h3. lr_match h3; fold hlr in h3; destruct hlr as [hl | hr]. lr_match h3; fold hlr in h3; destruct hlr as [hl' | hr']. destruct h3 as [|h3]; subst; auto with arith. apply ih in h3; auto with arith. destruct h3; subst; auto with arith. contradiction. apply ih in h3; auto with arith. Qed. Lemma in_has_comps_p_to_list_ind_impl : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} {P:Ensemble nat} {a:Fin_map (seg m) signe mns} {pfm:m <= length l}, Finite P -> forall {pfh:has_comps_p _ _ pfm P a} (n:nat) (pfn:In n (has_comps_p_to_list_ind pfh)), let pfinlt := in_has_comps_p_to_list_ind_lt _ pfn in let pfinlt' := lt_le_trans _ _ _ pfinlt pfm in let a' := fin_map_seg_rest a n pfinlt in has_comps_elt_p _ _ pfm P (nth_lt l n pfinlt') a. intros Bp l m. induction m as [|m ih]; simpl. intros; contradiction. intros P a h1 hf h2 n h3. lr_match h3; fold hlr in h3. pose proof (has_comps_p_last_dec m (le_trans m (S m) (length l) (le_n_Sn m) h1) h1 P (has_comps_p_finite l (S m) h1 P a h2) a) as h4. destruct h4 as [h4 | h4]. assert (h5:hlr = left h4). destruct hlr. f_equal. apply proof_irrelevance. contradiction. pose proof h3 as h3'. rewrite h5 in h3'. lr_match h3'; fold hlr0 in h3'; destruct hlr0 as [h6 | h6]. unfold hlr in h3, h5. clear hlr. destruct h3' as [|h7]; subst. destruct h4 as [h4 [r [h8 [h9 [h10 h11]]]]]. assert (h8':r <> n). omega. econstructor. assumption. apply h8'. assumption. assumption. symmetry. erewrite nth_lt_functional3. rewrite <- h10 at 1. apply nth_lt_functional3. apply nth_lt_functional3. assumption. specialize (ih P (fin_map_seg_rest a m (lt_n_Sn _)) (le_trans _ _ _ (le_n_Sn _) h1) hf h6 _ h7). simpl in ih. apply has_comps_elt_p_S in ih. erewrite nth_lt_functional3. apply ih. destruct h3'. unfold hlr in h3, h5. clear hlr. subst. destruct h6 as [h6 [r [h7 [h8 [h9 h10]]]]]. assert (h7':r <> n). omega. econstructor. assumption. apply h7'. assumption. assumption. erewrite nth_lt_functional3. unfold Atp in h9. unfold Atp. rewrite h9 at 1. apply nth_lt_functional3. apply nth_lt_functional3. assumption. contradiction. assert (h5:hlr = right h4). destruct hlr. contradiction. f_equal; apply proof_irrelevance. pose proof h3 as h3'. rewrite h5 in h3'. specialize (ih _ _ _ hf _ _ h3'). simpl in ih. apply has_comps_elt_p_S in ih. erewrite nth_lt_functional3. apply ih. Grab Existential Variables. auto with arith. auto with arith. auto with arith. auto with arith. Qed. Inductive only_has_comps_p {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) : Prop := | only_has_comps_p_intro : Finite P -> (forall (c d:nat) (pfc:c < m) (pfd:d < m), Inn P c -> Inn P d -> c <> d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = nth_lt l d (lt_le_trans _ _ _ pfd pfm) -> a |-> c <> a |-> d) -> only_has_comps_p l m pfm P a. Definition only_has_comps_p' {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) : Prop := only_has_comps_p l (length l) (le_refl _) P a. Inductive only_has_sames_p {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns) : Prop := | only_has_sames_p_intro : Finite P -> (forall (c d:nat) (pfc:c < m) (pfd:d < m), Inn P c -> Inn P d -> c <> d -> nth_lt l c (lt_le_trans _ _ _ pfc pfm) = nth_lt l d (lt_le_trans _ _ _ pfd pfm) -> a |-> c = a |-> d) -> only_has_sames_p l m pfm P a. Definition only_has_sames_p' {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns) : Prop := only_has_sames_p l (length l) (le_refl _) P a. Lemma only_has_comps_p_iff : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), let pfm' := le_length_ba_conv_list _ _ pfm in only_has_comps_p _ _ pfm P a <-> only_has_comps (ba_conv_list l) _ pfm' P a. intros; split; intro h1; destruct h1; econstructor. assumption. intros; eapply H0; auto. erewrite nth_lt_functional3 in H4. rewrite <- nth_lt_ba_conv_list in H4. rewrite H4. erewrite nth_lt_functional3. rewrite <- nth_lt_ba_conv_list. apply nth_lt_functional3. assumption. intros; eapply H0; auto. erewrite nth_lt_functional3. rewrite <-nth_lt_ba_conv_list. unfold Atp, btp in H4; unfold Atp, btp. rewrite H4 at 1. symmetry. erewrite nth_lt_functional3. rewrite <-nth_lt_ba_conv_list. apply nth_lt_functional3. Grab Existential Variables. eapply lt_le_trans. apply pfd. assumption. assumption. assumption. eapply lt_le_trans. apply pfd. assumption. assumption. assumption. Qed. Lemma only_has_comps_p_iff' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), only_has_comps_p' _ P a <-> only_has_comps' (ba_conv_list l) P a. unfold only_has_comps_p', only_has_comps'; intros; rewrite only_has_comps_p_iff. match goal with |- only_has_comps _ _ ?pf' _ _ <-> only_has_comps _ _ ?pf'' _ _ => assert (h1:pf' = pf'') end. apply proof_irrelevance. rewrite h1; tauto. Qed. Lemma only_has_comps_p_iff0 : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), only_has_comps_p _ _ pfm P a <-> @only_has_comps (ba_conv Bp) l _ pfm P a. intros; split; intro h2; destruct h2; constructor; auto. Qed. Lemma only_has_comps_p_iff0' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), only_has_comps_p' _ P a <-> @only_has_comps' (ba_conv Bp) l P a. unfold only_has_comps_p', only_has_comps'; intros; apply only_has_comps_p_iff0. Qed. Lemma only_has_sames_p_iff : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), let pfm' := le_length_ba_conv_list _ _ pfm in only_has_sames_p _ _ pfm P a <-> only_has_sames (ba_conv_list l) _ pfm' P a. intros; split; intro h1; destruct h1; econstructor. assumption. intros; eapply H0; auto. erewrite nth_lt_functional3 in H4. rewrite <- nth_lt_ba_conv_list in H4. rewrite H4. erewrite nth_lt_functional3. rewrite <- nth_lt_ba_conv_list. apply nth_lt_functional3. assumption. intros; eapply H0; auto. erewrite nth_lt_functional3. rewrite <-nth_lt_ba_conv_list. unfold Atp, btp in H4; unfold Atp, btp. rewrite H4 at 1. symmetry. erewrite nth_lt_functional3. rewrite <-nth_lt_ba_conv_list. apply nth_lt_functional3. Grab Existential Variables. eapply lt_le_trans. apply pfd. assumption. assumption. assumption. eapply lt_le_trans. apply pfd. assumption. assumption. assumption. Qed. Lemma only_has_sames_p_iff' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), only_has_sames_p' _ P a <-> only_has_sames' (ba_conv_list l) P a. unfold only_has_sames_p', only_has_sames'; intros; rewrite only_has_sames_p_iff. match goal with |- only_has_sames _ _ ?pf' _ _ <-> only_has_sames _ _ ?pf'' _ _ => assert (h1:pf' = pf'') end. apply proof_irrelevance. rewrite h1; tauto. Qed. Lemma only_has_sames_p_iff0 : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), only_has_sames_p _ _ pfm P a <-> @only_has_sames (ba_conv Bp) l _ pfm P a. intros; split; intro h2; destruct h2; constructor; auto. Qed. Lemma only_has_sames_p_iff0' : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), only_has_sames_p' _ P a <-> @only_has_sames' (ba_conv Bp) l P a. unfold only_has_sames_p', only_has_sames'; intros; apply only_has_sames_p_iff0. Qed. Lemma only_has_comps_p_compat : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), Finite P -> (only_has_comps_p l m pfm P a <-> ~has_sames_p l m pfm P a). intros; rewrite only_has_comps_p_iff0, has_sames_p_iff0; apply only_has_comps_compat; auto. Qed. Lemma only_has_comps_p_compat' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), Finite P -> (only_has_comps_p' l P a <-> ~has_sames_p' l P a). unfold only_has_comps_p', has_sames_p'; intros; apply only_has_comps_p_compat; auto. Qed. Lemma only_has_sames_p_compat : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), Finite P -> (only_has_sames_p l m pfm P a <-> ~has_comps_p l m pfm P a). intros; rewrite only_has_sames_p_iff0, has_comps_p_iff0; apply only_has_sames_compat; auto. Qed. Lemma only_has_sames_p_compat' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), Finite P -> (only_has_sames_p' l P a <-> ~has_comps_p' l P a). unfold only_has_sames_p', has_comps_p'; intros; apply only_has_sames_p_compat; auto. Qed. Lemma only_has_comps_p_finite : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), only_has_comps_p l m pfm P a -> Finite P. intros ? ? ? ? ? ? h1; destruct h1; auto. Qed. Lemma only_has_comps_p_finite' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), only_has_comps_p' l P a -> Finite P. intros ? ? ? ? h1; destruct h1; auto. Qed. Lemma only_has_sames_p_finite : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), only_has_sames_p l m pfm P a -> Finite P. intros ? ? ? ? ? ? h1; destruct h1; auto. Qed. Lemma only_has_sames_p_finite' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), only_has_sames_p' l P a -> Finite P. intros ? ? ? ? h1; destruct h1; auto. Qed. Lemma only_has_comps_p_dec : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), Finite P -> {only_has_comps_p l m pfm P a} + {has_sames_p l m pfm P a}. intros Bp l m h1 P a h2. pose proof (@only_has_comps_dec (ba_conv Bp) l m h1 P a h2) as h3. destruct h3 as [h3 | h3]. left. rewrite only_has_comps_p_iff0; auto. right. rewrite has_sames_p_iff. eapply has_sames_functional'. reflexivity. apply h3. Qed. Lemma only_has_comps_p_dec' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), Finite P -> {only_has_comps_p' l P a} + {has_sames_p' l P a}. unfold only_has_comps_p', has_sames_p'; intros; apply only_has_comps_p_dec; auto. Qed. Lemma only_has_sames_p_dec : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfm:m <= length l) (P:Ensemble nat) (a:Fin_map (seg m) signe mns), Finite P -> {only_has_sames_p l m pfm P a} + {has_comps_p l m pfm P a}. intros Bp l m h1 P a h2. pose proof (@only_has_sames_dec (ba_conv Bp) l m h1 P a h2) as h3. destruct h3 as [h3 | h3]. left. rewrite only_has_sames_p_iff0; auto. right. rewrite has_comps_p_iff. eapply has_comps_functional'. reflexivity. apply h3. Qed. Lemma only_has_sames_p_dec' : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (P:Ensemble nat) (a:Fin_map (seg (length l)) signe mns), Finite P -> {only_has_sames_p' l P a} + {has_comps_p' l P a}. unfold only_has_sames_p', has_comps_p'; intros; apply only_has_sames_p_dec; auto. Qed. Lemma linds_eq_determines_has_comps_p1 : forall {Bp:Bool_Alg_p T} {B:Bool_Alg} (l:list (btp Bp)) (l':list (bt B)) (P:Ensemble nat) (m:nat) (pfm:m <= length l) (pfe:length l = length l') (a:Fin_map (seg m) signe mns), Finite P -> let pfm' := le_eq_trans _ _ _ pfm pfe in linds ba_eq_dec_p (firstn m l) = linds ba_eq_dec (firstn m l') -> (has_comps_p l m pfm P a <-> has_comps l' m pfm' P a). intros Bp B l l' P m h1 h2 a h3 h4 h5. rewrite has_comps_p_iff0. rewrite <- (firstn_ba_conv_list l) in h5. rewrite ba_eq_dec_p_eq in h5. pose proof (linds_eq_determines_has_comps (ba_conv_list l) l' P m h1 h2 a h3 h5) as h6. simpl in h6. split; intro h7. rewrite h6 in h7. eapply has_comps_functional' in h7. apply h7. reflexivity. eapply has_comps_functional' in h7. rewrite <- h6 in h7. assumption. reflexivity. Qed. Lemma linds_eq_determines_has_sames_p1 : forall {Bp:Bool_Alg_p T} {B:Bool_Alg} (l:list (btp Bp)) (l':list (bt B)) (P:Ensemble nat) (m:nat) (pfm:m <= length l) (pfe:length l = length l') (a:Fin_map (seg m) signe mns), Finite P -> let pfm' := le_eq_trans _ _ _ pfm pfe in linds ba_eq_dec_p (firstn m l) = linds ba_eq_dec (firstn m l') -> (has_sames_p l m pfm P a <-> has_sames l' m pfm' P a). intros Bp B l l' P m h1 h2 a h3 h4 h5. rewrite has_sames_p_iff0. rewrite <- (firstn_ba_conv_list l) in h5. rewrite ba_eq_dec_p_eq in h5. pose proof (linds_eq_determines_has_sames (ba_conv_list l) l' P m h1 h2 a h3 h5) as h6. simpl in h6. split; intro h7. rewrite h6 in h7. eapply has_sames_functional' in h7. apply h7. reflexivity. eapply has_sames_functional' in h7. rewrite <- h6 in h7. assumption. reflexivity. Qed. Lemma linds_eq_determines_only_has_sames_p1 : forall {Bp:Bool_Alg_p T} {B:Bool_Alg} (l:list (btp Bp)) (l':list (bt B)) (P:Ensemble nat) (m:nat) (pfm:m <= length l) (pfe:length l = length l') (a:Fin_map (seg m) signe mns), Finite P -> linds ba_eq_dec_p (firstn m l) = linds ba_eq_dec (firstn m l') -> let pfm' := le_eq_trans _ _ _ pfm pfe in (only_has_sames_p l m pfm P a <-> only_has_sames l' m pfm' P a). intros Bp B l l' P m hm he a h1 h2 h3. rewrite only_has_sames_p_compat, only_has_sames_compat; auto. pose proof (linds_eq_determines_has_comps_p1 l l' P m hm he a h1 h2) as h4. fold h3 in h4. tauto. Qed. Lemma linds_eq_determines_has_comps_p : forall {Bp:Bool_Alg_p T} (l l':list (btp Bp)) (P:Ensemble nat) (m:nat) (pfm:m <= length l) (pfe:length l = length l') (a:Fin_map (seg m) signe mns), Finite P -> let pfm' := le_eq_trans _ _ _ pfm pfe in linds ba_eq_dec_p (firstn m l) = linds ba_eq_dec_p (firstn m l') -> (has_comps_p l m pfm P a <-> has_comps_p l' m pfm' P a). intros Bp l l' P m h1 h2 a h3 h4 h5. rewrite has_comps_p_iff0. rewrite has_comps_p_iff0. rewrite <- (firstn_ba_conv_list l) in h5. rewrite <- (firstn_ba_conv_list l') in h5. rewrite ba_eq_dec_p_eq in h5. pose proof (linds_eq_determines_has_comps (ba_conv_list l) (ba_conv_list l') P m h1 h2 a h3 h5) as h6. simpl in h6. split; intro h7. rewrite h6 in h7. eapply has_comps_functional' in h7. apply h7. reflexivity. eapply has_comps_functional' in h7. rewrite <- h6 in h7. assumption. reflexivity. Qed. Lemma linds_eq_determines_has_sames_p : forall {Bp:Bool_Alg_p T} (l l':list (btp Bp)) (P:Ensemble nat) (m:nat) (pfm:m <= length l) (pfe:length l = length l') (a:Fin_map (seg m) signe mns), Finite P -> let pfm' := le_eq_trans _ _ _ pfm pfe in linds ba_eq_dec_p (firstn m l) = linds ba_eq_dec_p (firstn m l') -> (has_sames_p l m pfm P a <-> has_sames_p l' m pfm' P a). intros Bp l l' P m h1 h2 a h3 h4 h5. rewrite has_sames_p_iff0. rewrite has_sames_p_iff0. rewrite <- (firstn_ba_conv_list l) in h5. rewrite <- (firstn_ba_conv_list l') in h5. rewrite ba_eq_dec_p_eq in h5. pose proof (linds_eq_determines_has_sames (ba_conv_list l) (ba_conv_list l') P m h1 h2 a h3 h5) as h6. simpl in h6. split; intro h7. rewrite h6 in h7. eapply has_sames_functional' in h7. apply h7. reflexivity. eapply has_sames_functional' in h7. rewrite <- h6 in h7. assumption. reflexivity. Qed. Lemma only_has_sames_p_has_comps_distinct_linds' : forall {Bp:Bool_Alg_p T} {B:Bool_Alg} (la:list (btp Bp)) (lb:list (bt B)) (n:nat) (pfe:length la = length lb) (F:Fin_map (seg (length la)) signe mns), only_has_sames_p' la (seg (length la)) F -> let F' := fin_map_seg_transfer F pfe in has_comps' lb (seg (length la)) F' -> linds ba_eq_dec_p la <> linds ba_eq_dec lb. intros Bp B la lb n h1 F h3 F' h4 h5. rewrite only_has_sames_p_compat' in h3. pose proof (linds_eq_determines_has_comps_p1 la lb (seg (length la)) (length la) (le_refl _) h1 F (finite_seg _)) as h6. simpl in h6. hfirst h6. rewrite firstn_length, h1, firstn_length; auto. specialize (h6 hf). rewrite h6 in h3. unfold has_comps' in h4. eapply has_comps_transfer_rev in h4. specialize (h3 h4); auto. apply finite_seg. Qed. Lemma only_has_comps_p_has_sames_distinct_linds' : forall {Bp:Bool_Alg_p T} {B:Bool_Alg} (la:list (btp Bp)) (lb:list (bt B)) (n:nat) (pfe:length la = length lb) (F:Fin_map (seg (length la)) signe mns), only_has_comps_p' la (seg (length la)) F -> let F' := fin_map_seg_transfer F pfe in has_sames' lb (seg (length la)) F' -> linds ba_eq_dec_p la <> linds ba_eq_dec lb. intros Bp B la lb n h1 F h3 F' h4 h5. rewrite only_has_comps_p_compat' in h3. pose proof (linds_eq_determines_has_sames_p1 la lb (seg (length la)) (length la) (le_refl _) h1 F (finite_seg _)) as h6. simpl in h6. hfirst h6. rewrite firstn_length, h1, firstn_length; auto. specialize (h6 hf). rewrite h6 in h3. unfold has_sames' in h4. eapply has_sames_transfer_rev in h4. specialize (h3 h4); auto. apply finite_seg. Qed. Lemma only_has_sames_map_seg_to_fin_map_p_ex : forall {Bp:Bool_Alg_p T} {n:nat} (f:nat->btp Bp) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns), only_has_sames_p' (map_seg (fun_to_seg f n)) (seg n) F -> exists! F':Fin_map (Im (seg n) f) signe mns, forall i:nat, i < n -> F' |-> (f i) = F |-> i. intros Bp n f F h1. pose proof (@only_has_sames_map_seg_to_fin_map_ex (ba_conv Bp) n f F) as h2. hfirst h2. rewrite <- only_has_sames_p_iff0'; auto. apply h2; auto. Qed. Definition only_has_sames_map_seg_to_fin_map_p {Bp:Bool_Alg_p T} {n:nat} (f:nat->btp Bp) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pf:only_has_sames_p' (map_seg (fun_to_seg f n)) (seg n) F) : Fin_map (Im (seg n) f) signe mns := proj1_sig (constructive_definite_description _ (only_has_sames_map_seg_to_fin_map_p_ex f F pf)). Lemma only_has_sames_map_seg_to_fin_map_p_compat : forall {Bp:Bool_Alg_p T} {n:nat} (f:nat->btp Bp) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pf:only_has_sames_p' (map_seg (fun_to_seg f n)) (seg n) F), let F' := only_has_sames_map_seg_to_fin_map_p f F pf in forall i:nat, i < n -> F' |-> (f i) = F |-> i. simpl; unfold only_has_sames_map_seg_to_fin_map_p; intros; destruct constructive_definite_description; auto. Qed. Lemma only_has_sames_map_seg_to_fin_map_p_eq : forall {Bp:Bool_Alg_p T} {n:nat} (f:nat->btp Bp) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pf:only_has_sames_p' (map_seg (fun_to_seg f n)) (seg n) F), let pf' := iff1 (only_has_sames_p_iff0' _ _) pf in only_has_sames_map_seg_to_fin_map_p f F pf = @only_has_sames_map_seg_to_fin_map (ba_conv Bp) _ f F pf'. intros Bp n f F h1 h2. apply fin_map_ext. intros x h3. destruct h3 as [x h3]; subst. destruct h3 as [h3]. rewrite only_has_sames_map_seg_to_fin_map_p_compat, (@only_has_sames_map_seg_to_fin_map_compat (ba_conv Bp)); auto. Qed. Lemma fin_map_im_seg_to_seg_length_p_ex : forall {Bp:Bool_Alg_p T} {n:nat} (f:nat->btp Bp) (F:Fin_map (Im (seg n) f) signe mns), exists! F':Fin_map (seg (length (map_seg (fun i (_:i < n) => f i)))) signe mns, forall i (pf:i < n), F' |-> i = F |-> f i. intros Bp n f F. pose proof (@fin_map_im_seg_to_seg_length_ex (ba_conv Bp) n f F) as h2. apply h2; auto. Qed. Definition fin_map_im_seg_to_seg_length_p {Bp:Bool_Alg_p T} {n:nat} (f:nat->btp Bp) (F:Fin_map (Im (seg n) f) signe mns) : Fin_map (seg (length (map_seg (fun i (_:i < n) => f i)))) signe mns := proj1_sig (constructive_definite_description _ (fin_map_im_seg_to_seg_length_p_ex f F)). Lemma fin_map_im_seg_to_seg_length_p_compat : forall {Bp:Bool_Alg_p T} {n:nat} (f:nat->btp Bp) (F:Fin_map (Im (seg n) f) signe mns), let F' := fin_map_im_seg_to_seg_length_p f F in forall i (pf:i < n), F' |-> i = F |-> f i. simpl; unfold fin_map_im_seg_to_seg_length_p; intros; destruct constructive_definite_description; auto. Qed. End HasCompsP. Definition el_prod_p {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp) } (f:Fin_map E signe mns) : btp Bp. pose (fun i => eps_p i (f |-> i)) as p. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ E p h1) as h2. refine (times_set_p _ h2). Defined. Lemma el_prod_p_eq : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (f:Fin_map E signe mns), el_prod_p f = el_prod _ (ba_conv_fin_map_dom f). intros Bp E f. unfold el_prod_p, el_prod, ba_conv_fin_map_dom. rewrite times_set_p_eq. unfold ba_conv_elt, ba_conv_set. unfold ba_conv_type. rewrite transfer_eq_refl. f_equal. Qed. Definition el_sum_p {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (f:Fin_map E signe mns) : btp Bp. pose (fun i => eps_p i (f |-> i)) as p. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ E p h1) as h2. refine (plus_set_p _ h2). Defined. Lemma el_sum_p_eq : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (f:Fin_map E signe mns), el_sum_p f = el_sum _ (ba_conv_fin_map_dom f). intros Bp E f. unfold el_sum_p, el_sum, ba_conv_fin_map_dom. rewrite plus_set_p_eq. unfold ba_conv_elt, ba_conv_set. unfold ba_conv_type. rewrite transfer_eq_refl. f_equal. Qed. Definition el_prod_compose_p {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:(btp Bp)->Atp) (f:Fin_map E signe mns) : Atp. pose (fun i:(btp Bp) => eps_p (g i) (f |-> i)) as p. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ E p h1) as h2. refine (times_set_p _ h2). Defined. Lemma el_prod_compose_p_eq : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:(btp Bp)->Atp) (f:Fin_map E signe mns), el_prod_compose_p g f = el_prod_compose (ba_conv Ap) (ba_conv_fun1 g) f. intros Bp E g f. unfold el_prod_compose_p, el_prod_compose, ba_conv_fun1. rewrite times_set_p_eq. unfold ba_conv_elt, ba_conv_set. unfold ba_conv_type. rewrite transfer_eq_refl. f_equal. Qed. Definition el_prod_compose_p' {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:(btp Bp)->Atp) (f:Fin_map E signe mns) : Atp. pose (fun i:sig_set E => g (proj1_sig i)) as g'. pose (fun i:sig_set E => eps_p (g' i) (f |-> (proj1_sig i))) as p. pose proof (fin_map_fin_dom f) as h1. rewrite finite_full_sig_iff in h1. pose proof (finite_image _ _ (full_sig E) p h1) as h2. refine (times_set_p _ h2). Defined. Lemma el_prod_compose_p_eq' : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:(btp Bp)->Atp) (f:Fin_map E signe mns), el_prod_compose_p' g f = el_prod_compose' (ba_conv Ap) (ba_conv_fun1 g) (ba_conv_fin_map_dom f). intros Bp E g f. unfold el_prod_compose_p', el_prod_compose', ba_conv_fun1, ba_conv_fin_map_dom. rewrite times_set_p_eq. unfold ba_conv_elt. unfold ba_conv_type. rewrite transfer_eq_refl. f_equal. Qed. Definition el_prod_compose_p1 {A:Bool_Alg} {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:(btp Bp)->bt A) (f:Fin_map E signe mns) : bt A. pose (fun i:(btp Bp) => eps (g i) (f |-> i)) as p. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ E p h1) as h2. refine (times_set _ h2). Defined. Definition el_prod_compose_p1_eq : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:(btp Bp)->bt A) (f:Fin_map E signe mns), el_prod_compose_p1 g f = el_prod_compose _ (ba_conv_fun1 g) (ba_conv_fin_map_dom f). intros A Bp E g f. unfold el_prod_compose_p1, el_prod_compose. apply times_set_functional. f_equal. Qed. Lemma el_prod_compose_p1_eq0 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:(btp Bp)->bt A) (f:Fin_map E signe mns), el_prod_compose_p1 g f = @el_prod_compose A (ba_conv Bp) _ g f. intros; auto. Qed. Lemma el_prod_compose_p1_eq' : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:(btp Bp)->Atp) (f:Fin_map E signe mns), el_prod_compose_p g f = el_prod_compose_p1 (ba_conv_fun2 g) f. intros Bp E g f. unfold el_prod_compose_p, el_prod_compose_p1. rewrite times_set_p_eq. unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. unfold ba_conv_set, ba_conv_type. generalize (finite_image (btp Bp) (btp Ap) E (fun i : btp Bp => eps_p (g i) (f |-> i)) (fin_map_fin_dom f)). intro h0. assert (h1:times_set (transfer_dep eq_refl (Im E (fun i : btp Bp => eps_p (g i) (f |-> i)))) h0 (B:=ba_conv Ap) = times_set (Im E (fun i : btp Bp => eps_p (g i) (f |-> i))) h0 (B:=ba_conv Ap)). apply times_set_functional. rewrite transfer_dep_eq_refl. reflexivity. rewrite h1 at 1. apply (times_set_functional (B:=ba_conv Ap)). reflexivity. Qed. Lemma el_prod_compose_el_prod_compose'_compat_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:(btp Bp)->Atp), el_prod_compose_p g = el_prod_compose_p' (E:=E) g. intros B E g. pose proof (el_prod_compose_el_prod_compose'_compat (E:=(ba_conv_set E)) (ba_conv Ap) (ba_conv_fun1 g)) as h1. unfold ba_conv_set, ba_conv_type in h1. rewrite transfer_dep_eq_refl in h1. apply functional_extensionality. intro x. rewrite el_prod_compose_p_eq, el_prod_compose_p_eq'. rewrite h1. reflexivity. Qed. Definition el_prod_l_compose_p {Bp:Bool_Alg_p T} {l:list (btp Bp)} (F:Fin_map (seg (length l)) signe mns) (g:btp Bp->Atp) := times_list_p (map_seg (fun i (pf:i < length l) => eps_p (g (nth_lt l i pf)) (F |-> i))). Lemma el_prod_l_compose_p_eq : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (F:Fin_map (seg (length l)) signe mns) (g:btp Bp->Atp), el_prod_l_compose_p F g = @el_prod_l_compose (ba_conv Ap) (ba_conv Bp) _ F g. intros. unfold el_prod_l_compose_p, el_prod_l_compose. rewrite times_list_p_eq; auto. Qed. Definition el_prod_l_compose_p1 {B:Bool_Alg} {l:list Atp} (F:Fin_map (seg (length l)) signe mns) (g:Atp->bt B) := times_list (map_seg (fun i (pf:i < length l) => eps (g (nth_lt l i pf)) (F |-> i))). Lemma el_prod_l_compose_p1_eq : forall {B:Bool_Alg} {l:list Atp} (F:Fin_map (seg (length l)) signe mns) (g:Atp->bt B), el_prod_l_compose_p1 F g = @el_prod_l_compose B (ba_conv Ap) _ F g. intros; auto. Qed. Definition el_prod_l_compose_p2 {B:Bool_Alg} {l:list (bt B)} (F:Fin_map (seg (length l)) signe mns) (g:bt B->Atp) := times_list_p (map_seg (fun i (pf:i < length l) => eps_p (g (nth_lt l i pf)) (F |-> i))). Lemma el_prod_l_compose_p2_eq : forall {B:Bool_Alg} {l:list (bt B)} (F:Fin_map (seg (length l)) signe mns) (g:bt B->Atp), el_prod_l_compose_p2 F g = @el_prod_l_compose (ba_conv Ap) B _ F g. intros. unfold el_prod_l_compose_p2, el_prod_l_compose. rewrite times_list_p_eq; auto. Qed. Definition el_sum_compose_p {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:(btp Bp)->Atp) (f:Fin_map E signe mns) : Atp. pose (fun i:(btp Bp) => eps_p (g i) (f |-> i)) as p. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ E p h1) as h2. refine (plus_set_p _ h2). Defined. Lemma el_sum_compose_p_eq : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:(btp Bp)->Atp) (f:Fin_map E signe mns), el_sum_compose_p g f = el_sum_compose (ba_conv Ap) (ba_conv_fun1 g) f. intros Bp E g f. unfold el_sum_compose_p, el_sum_compose, ba_conv_fun1. rewrite plus_set_p_eq. unfold ba_conv_elt, ba_conv_set. unfold ba_conv_type. rewrite transfer_eq_refl. f_equal. Qed. Lemma el_prod_el_prod_compose_compat_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} {g:sig_set E->Atp} (f:Fin_map (Im (full_sig E) g) signe mns) (pf:Finite E), el_prod_p f = el_prod_compose_p (sig_fun_app g %0) (fin_map_im_full_sig_eq ba_eq_dec_p f pf %0). intros Bp E g f h1. assert (h2:exists! f':Fin_map (Im (full_sig (ba_conv_set E)) (ba_conv_sig_fun1 g)) signe mns, forall y, Inn (Im (full_sig E) g) y -> f |-> y = f' |-> y). assert (h2:Im (full_sig (ba_conv_set E)) (ba_conv_sig_fun1 g) = Im (full_sig E) g). unfold ba_conv_set, ba_conv_type. apply im_ext_in. intros x h2. reflexivity. rewrite h2. exists f. red; split; auto. intros f' h3. apply fin_map_ext. assumption. pose (proj1_sig (constructive_definite_description _ h2)) as f'. pose proof (el_prod_el_prod_compose_compat (ba_conv Ap) f' h1) as h3. rewrite el_prod_p_eq, el_prod_compose_p_eq. unfold f' in h3. destruct constructive_definite_description as [f'' h4]. simpl in h3. simpl in f'. assert (h5:f'' = ba_conv_fin_map_dom f). apply fin_map_ext. intros y h5. rewrite h4. reflexivity. destruct h5 as [y h5]. subst. apply Im_intro with y. constructor. reflexivity. subst. unfold btp, bt in h3. unfold bt, Btype_p, Atp, btp. simpl in h3. simpl. rewrite h3 at 1. f_equal. f_equal. apply type_eq_dec_eq. Qed. Lemma el_sum_el_sum_compose_compat_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} {g:sig_set E->Atp} (f:Fin_map (Im (full_sig E) g) signe mns) (pf:Finite E), el_sum_p f = el_sum_compose_p (sig_fun_app g %0) (fin_map_im_full_sig_eq ba_eq_dec_p f pf %0). intros Bp E g f h1. assert (h2:exists! f':Fin_map (Im (full_sig (ba_conv_set E)) (ba_conv_sig_fun1 g)) signe mns, forall y, Inn (Im (full_sig E) g) y -> f |-> y = f' |-> y). assert (h2:Im (full_sig (ba_conv_set E)) (ba_conv_sig_fun1 g) = Im (full_sig E) g). unfold ba_conv_set, ba_conv_type. apply im_ext_in. intros x h2. reflexivity. rewrite h2. exists f. red; split; auto. intros f' h3. apply fin_map_ext. assumption. pose (proj1_sig (constructive_definite_description _ h2)) as f'. pose proof (el_sum_el_sum_compose_compat (ba_conv Ap) f' h1) as h3. rewrite el_sum_p_eq, el_sum_compose_p_eq. unfold f' in h3. destruct constructive_definite_description as [f'' h4]. simpl in h3. simpl in f'. assert (h5:f'' = ba_conv_fin_map_dom f). apply fin_map_ext. intros y h5. rewrite h4. reflexivity. destruct h5 as [y h5]. subst. apply Im_intro with y. constructor. reflexivity. subst. unfold btp, bt in h3. unfold bt, Btype_p, Atp, btp. simpl in h3. simpl. rewrite h3 at 1. f_equal. f_equal. apply type_eq_dec_eq. Qed. Lemma el_prod_el_prod_compose_compat_p' : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (f : Fin_map E signe mns) (g:sig_set E->Atp) (f': Fin_map (Im (full_sig E) g) signe mns), (forall b:(btp Bp), Inn E b -> f' |-> ((g, (Bzero_p T (Bc_p T Ap))) ||-> b) = f |-> b) -> el_prod_compose_p (sig_fun_app g %0) f = el_prod_p f'. intros Bp E f g f' h1. assert (h2:exists! f'':Fin_map (Im (full_sig (ba_conv_set E)) (ba_conv_sig_fun1 g)) signe mns, forall y, Inn (Im (full_sig E) g) y -> f' |-> y = f'' |-> y). assert (h2:Im (full_sig (ba_conv_set E)) (ba_conv_sig_fun1 g) = Im (full_sig E) g). unfold ba_conv_set, ba_conv_type. apply im_ext_in. intros x h2. reflexivity. rewrite h2. exists f'. red; split; auto. intros f'' h3. apply fin_map_ext. assumption. pose (proj1_sig (constructive_definite_description _ h2)) as f''. assert (hf:f'' = f'). unfold f''. destruct constructive_definite_description as [f''' h6]. simpl. simpl in f''. apply fin_map_ext. intros y h7. rewrite h6. reflexivity. destruct h7 as [y h7]. subst. apply Im_intro with y. constructor. reflexivity. assert (h4: (forall b : bt (ba_conv Bp), Inn (ba_conv_set E) b -> f'' |-> ((ba_conv_sig_fun1 g, (Bzero_p T (Bc_p T Ap))) ||-> b) = ba_conv_fin_map_dom f |-> b) ). intros b h5. specialize (h1 b h5). rewrite <- h1 at 1. f_equal. apply hf. pose proof (el_prod_el_prod_compose_compat' (ba_conv Ap) (ba_conv_fin_map_dom f) (ba_conv_sig_fun1 g) f'' h4) as h3. rewrite el_prod_compose_p_eq, el_prod_p_eq. unfold btp, bt in h3. unfold bt, Btype_p, Atp, btp. simpl in h3. simpl. rewrite h3 at 1. f_equal. apply hf. Qed. Section ElProdLP. Definition el_prod_l_p {Bp:Bool_Alg_p T} {l:list (btp Bp)} (F:Fin_map (seg (length l)) signe mns) := times_list_p (map_seg (fun i (pf:i < length l) => eps_p (nth_lt l i pf) (F |-> i))). Lemma el_prod_l_p_eq : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (F:Fin_map (seg (length l)) signe mns), el_prod_l_p F = @el_prod_l (ba_conv Bp) l F. intros; unfold el_prod_l_p, el_prod_l. rewrite times_list_p_eq; auto. Qed. Definition el_prod_l_firstn_p {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pf:m <= length l) (F:Fin_map (seg m) signe mns) := let F' := fin_map_seg_transfer F (eq_sym (length_firstn l m pf)) in el_prod_l_p F'. Lemma el_prod_l_firstn_p_eq : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pf:m <= length l) (F:Fin_map (seg m) signe mns), el_prod_l_firstn_p pf F = @el_prod_l_firstn (ba_conv Bp) l m pf F. intros; unfold el_prod_l_firstn_p, el_prod_l_firstn; apply el_prod_l_p_eq. Qed. Lemma el_prod_l_p0 : forall (Bp:Bool_Alg_p T) (F:Fin_map (seg (length nil)) signe mns), el_prod_l_p F = bone_p Bp. intro; unfold el_prod_l_p; auto. Qed. Lemma el_prod_l_firstn_p0 : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (F:Fin_map (seg 0) signe mns), el_prod_l_firstn_p (le_0_n (length l)) F = bone_p _. intros; unfold el_prod_l_firstn_p, el_prod_l_p; auto. Qed. Lemma has_comps_el_prod_l_firstn_p : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pfle:m <= length l) (F:Fin_map (seg m) signe mns), has_comps_p l m pfle (seg m) F -> el_prod_l_firstn_p pfle F = %0. intros Bp l m h1 F h2. rewrite el_prod_l_firstn_p_eq, ba_conv_zero. rewrite has_comps_p_iff0 in h2. apply (@has_comps_el_prod_l_firstn (ba_conv Bp) l m h1 F h2); auto. Qed. Lemma el_prod_l_firstn_p_0 : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm: m <= length l) (F:Fin_map (seg m) signe mns) (x:btp Bp), In x (firstn m l) -> let P := list_to_set (linds_occ ba_eq_dec_p (firstn m l) x) in has_comps_p _ _ pfm P F -> el_prod_l_firstn_p pfm F = %0. intros Bp l m h1 F x h2 P h3. rewrite el_prod_l_firstn_p_eq, ba_conv_zero. unfold P in h3. rewrite has_comps_p_iff0 in h3. apply (@el_prod_l_firstn_0 (ba_conv Bp) l m h1 F x h2); auto. rewrite linds_occ_ba_conv_list' in h3; auto. Qed. Lemma el_prod_l_p_0 : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (F:Fin_map (seg (length l)) signe mns) (x:btp Bp), In x l -> let P := list_to_set (linds_occ ba_eq_dec_p l x) in has_comps_p' _ P F -> el_prod_l_p F = %0. intros. rewrite el_prod_l_p_eq. erewrite el_prod_l_0. reflexivity. apply H. rewrite has_comps_p_iff0' in H0; auto. unfold P in H0. rewrite linds_occ_ba_conv_list in H0; auto. Qed. Lemma has_comps_el_prod_p_l0 : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (F:Fin_map (seg (length l)) signe mns), has_comps_p' l (seg (length l)) F -> el_prod_l_p F = %0. intros Bp l f h1. rewrite has_comps_p_iff' in h1. rewrite el_prod_l_p_eq. erewrite has_comps_el_prod_l0. reflexivity. assumption. Qed. Lemma el_prod_l_firstn_p_S : forall {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pf:S m <= length l) (F:Fin_map (seg (S m)) signe mns), let pf' := lt_le_trans _ _ _ (lt_n_Sn m) pf in let pf'' := lt_le_weak _ _ pf' in el_prod_l_firstn_p pf F = (eps_p (nth_lt l m pf') (F |-> m)) %* (el_prod_l_firstn_p pf'' (fin_map_seg_rest F _ (lt_n_Sn m))). intros Bp l m h1 F; simpl. do 2 rewrite el_prod_l_firstn_p_eq. rewrite eps_p_eq. apply (@el_prod_l_firstn_S (ba_conv Bp) l m h1 F). Qed. Lemma el_prod_l_firstn_p_removelast_compat : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm: m <= length l) (pfm':m <= length (removelast l)) (F:Fin_map (seg (S m)) signe mns), el_prod_l_firstn_p pfm (fin_map_seg_rest F m (lt_n_Sn _)) = el_prod_l_firstn_p pfm' (fin_map_seg_rest F m (lt_n_Sn _)). intros Bp l m h1 h2 F. do 2 rewrite el_prod_l_firstn_p_eq. apply (@el_prod_l_firstn_removelast_compat (ba_conv Bp) l m h1 h2 F). Qed. Lemma in_last_el_prod_l_firstn_p_dec : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfsm:S m <= length l) (F:Fin_map (seg (S m)) signe mns), let pfm := le_trans _ _ _ (le_n_Sn _) pfsm in let x := nth_lt l m pfsm in let P := list_to_set (linds_occ ba_eq_dec_p (firstn (S m) l) x) in In x (firstn m l) -> Finite P -> {has_comps_elt_p _ _ pfsm P (nth_lt l m pfsm) F /\ el_prod_l_firstn_p pfsm F = %0} + {~has_comps_elt_p _ _ pfsm P (nth_lt l m pfsm) F /\ el_prod_l_firstn_p pfsm F = el_prod_l_firstn_p pfm (fin_map_seg_rest F _ (lt_n_Sn _))}. intros Bp l m h1 F h2 x P h3 hf. unfold P in hf. unfold P. rewrite linds_occ_ba_conv_list' in hf. do 2 rewrite el_prod_l_firstn_p_eq. pose proof (in_last_el_prod_l_firstn_dec (ba_conv Bp) h1 F h3 hf) as h4. destruct h4 as [h4 | h4]. left. rewrite has_comps_elt_p_iff, linds_occ_ba_conv_list'; auto. right. rewrite has_comps_elt_p_iff, linds_occ_ba_conv_list'; auto. Qed. Definition fun_for_el_prod_l_firstn_p {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (x:btp Bp) (pf:In x (firstn m la)) : btp Bp := if (has_comps_p_dec pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m la) x)) F (list_to_set_finite _)) then bzero_p _ else eps_p x (F |-> lind ba_eq_dec_p _ _ pf). Definition fun_for_el_prod_l_p {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns) (x:btp Bp) (pf:In x la) : btp Bp := if (has_comps_p_dec' F (list_to_set (linds_occ ba_eq_dec_p la x)) (list_to_set_finite _)) then bzero_p _ else eps_p x (F |-> lind ba_eq_dec_p _ _ pf). Lemma fun_for_el_prod_l_p_eq : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns) (x:btp Bp) (pf:In x la), fun_for_el_prod_l_p F x pf = @fun_for_el_prod_l (ba_conv Bp) _ F x pf. intros; unfold fun_for_el_prod_l, fun_for_el_prod_l_p. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. reflexivity. rewrite has_comps_p_iff', linds_occ_ba_conv_list in hl. contradiction. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. rewrite has_comps_p_iff', linds_occ_ba_conv_list in hr. contradiction. rewrite eps_p_eq; repeat f_equal. apply ba_eq_dec_p_eq. Qed. Lemma fun_for_el_prod_l_firstn_p_eq : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (x:btp Bp) (pf:In x (firstn m la)), fun_for_el_prod_l_firstn_p pfm F x pf = @fun_for_el_prod_l_firstn (ba_conv Bp) la m pfm F x pf. intros; unfold fun_for_el_prod_l_firstn_p, fun_for_el_prod_l_firstn. lr_if_goal; fold hlr; destruct hlr as [hl | hr]; lr_if_goal; fold hlr; destruct hlr as [hl' | hr']; auto. rewrite linds_occ_ba_conv_list, has_comps_p_iff0 in hl. contradiction. rewrite linds_occ_ba_conv_list, has_comps_p_iff0 in hr. contradiction. rewrite eps_p_eq; auto. repeat f_equal. apply ba_eq_dec_p_eq. Qed. Lemma fun_for_el_prod_l_firstn_p_functional : forall {Bp:Bool_Alg_p T} {la la':list (btp Bp)} {m:nat} (pfm:m <= length la) (pfm':m <= length la') (F:Fin_map (seg m) signe mns) (x x':btp Bp) (pf:In x (firstn m la)) (pf':In x' (firstn m la')), x = x' -> la = la' -> fun_for_el_prod_l_firstn_p pfm F x pf = fun_for_el_prod_l_firstn_p pfm' F x' pf'. intros; subst; repeat f_equal; try apply proof_irrelevance. Qed. Lemma fun_for_el_prod_l_p_functional : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns) (x:btp Bp) (pf:In x la) (x':btp Bp) (pf':In x' la), x = x' -> fun_for_el_prod_l_p F x pf = fun_for_el_prod_l_p F x' pf'. intros; subst; repeat f_equal; apply proof_irrelevance. Qed. Lemma fun_for_el_prod_l_p_functional' : forall {Bp:Bool_Alg_p T} {la la':list (btp Bp)} (F:Fin_map (seg (length la)) signe mns) (x:btp Bp) (pf:In x la) (x':btp Bp) (pf':In x' la') (pfl:la = la'), x = x' -> let F' := fin_map_seg_transfer F (f_equal (@length (btp Bp)) pfl) in fun_for_el_prod_l_p F x pf = fun_for_el_prod_l_p F' x' pf'. simpl; intros; subst; repeat f_equal. simpl. rewrite fin_map_seg_transfer_eq_refl; auto. apply proof_irrelevance. Qed. (*This domain is [linds _(firstn m la)].*) Definition fun_for_el_prod_l_linds_firstn_p {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (ln:list nat) (pf:In ln (linds ba_eq_dec_p (firstn m la))) : btp Bp := if (has_comps_p_dec pfm (list_to_set ln) F (list_to_set_finite _)) then Bzero_p T (Bc_p T Bp) else let (x, pf) := constructive_definite_description _ (iff1 (in_linds_iff _ _ _) pf) in let (pf', _) := pf in eps_p x (F |-> lind ba_eq_dec_p _ _ pf'). (*This domain is [linds _ la].*) Definition fun_for_el_prod_l_linds_p {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns) (ln:list nat) (pf:In ln (linds ba_eq_dec_p la)) : btp Bp := if (has_comps_p_dec' F (list_to_set ln) (list_to_set_finite _)) then Bzero_p _ (Bc_p _ Bp) else let (x, pf) := constructive_definite_description _ (iff1 (in_linds_iff _ _ _) pf) in let (pf', _) := pf in eps_p x (F |-> lind ba_eq_dec_p _ _ pf'). Lemma fun_for_el_prod_l_linds_firstn_p_functional : forall {Bp:Bool_Alg_p T} {la la':list (btp Bp)} {m:nat} (pfm:m <= length la) (pfm':m <= length la') (F:Fin_map (seg m) signe mns) (ln ln':list nat) (pf:In ln (linds ba_eq_dec_p (firstn m la))) (pf':In ln' (linds ba_eq_dec_p (firstn m la'))), la = la' -> ln = ln' -> fun_for_el_prod_l_linds_firstn_p pfm F ln pf = fun_for_el_prod_l_linds_firstn_p pfm' F ln' pf'. intros; subst; repeat f_equal; try apply proof_irrelevance. Qed. Lemma fun_for_el_prod_l_linds_p_functional : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns) (ln ln':list nat) (pf:In ln (linds ba_eq_dec_p la)) (pf':In ln' (linds ba_eq_dec_p la)), ln = ln' -> fun_for_el_prod_l_linds_p F ln pf = fun_for_el_prod_l_linds_p F ln' pf'. intros; subst; repeat f_equal; try apply proof_irrelevance. Qed. Lemma fun_for_el_prod_l_linds_p_functional' : forall {Bp:Bool_Alg_p T} {la la':list (btp Bp)} (F:Fin_map (seg (length la)) signe mns) (ln ln':list nat) (pf:In ln (linds ba_eq_dec_p la)) (pf':In ln' (linds ba_eq_dec_p la')) (pfe:la = la'), ln = ln' -> let F' := fin_map_seg_transfer F (f_equal (@length (btp Bp)) pfe) in fun_for_el_prod_l_linds_p F ln pf = fun_for_el_prod_l_linds_p F' ln' pf'. simpl; intros; subst; repeat f_equal; try apply proof_irrelevance; simpl; rewrite fin_map_seg_transfer_eq_refl; auto. Qed. Lemma fun_for_el_prod_l_linds_firstn_p_eq : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (ln:list nat) (pf:In ln (linds ba_eq_dec_p (firstn m la))), let pf' := in_subst_l _ _ ln (linds_ba_conv_list' (firstn m la)) pf in fun_for_el_prod_l_linds_firstn_p pfm F ln pf = @fun_for_el_prod_l_linds_firstn (ba_conv Bp) la m pfm F ln pf'. intros Bp la m h1 F ln h2; simpl. unfold fun_for_el_prod_l_linds_firstn_p, fun_for_el_prod_l_linds_firstn. lr_if_goal; fold hlr; destruct hlr as [hl | hr]; lr_if_goal; fold hlr; destruct hlr as [hl' | hr']; auto. rewrite has_comps_p_iff0 in hl. contradiction. rewrite has_comps_p_iff0 in hr. contradiction. destruct constructive_definite_description as [x h3]. destruct constructive_definite_description as [x' h4]. destruct h3 as [h3 h5], h4 as [h4 h6]; subst. rewrite ba_eq_dec_p_eq in h6. apply inj_linds_occ in h6; subst. rewrite eps_p_eq. repeat f_equal. apply ba_eq_dec_p_eq. apply proof_irrelevance. assumption. assumption. Qed. Lemma fun_for_el_prod_l_linds_p_eq : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns) (ln:list nat) (pf:In ln (linds ba_eq_dec_p la)), let pf' := in_subst_l _ _ ln (linds_ba_conv_list' la) pf in fun_for_el_prod_l_linds_p F ln pf = @fun_for_el_prod_l_linds (ba_conv Bp) la F ln pf'. intros Bp la F ln h2; simpl. unfold fun_for_el_prod_l_linds_p, fun_for_el_prod_l_linds. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. reflexivity. rewrite has_comps_p_iff' in hl. contradiction. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. rewrite has_comps_p_iff' in hr; contradiction. destruct constructive_definite_description as [x h3]. destruct constructive_definite_description as [x' h4]. destruct h3 as [h3 h5], h4 as [h4 h6]; subst. rewrite ba_eq_dec_p_eq in h6. rewrite eps_p_eq. apply inj_linds_occ in h6; subst. repeat f_equal. apply ba_eq_dec_p_eq. apply proof_irrelevance. assumption. assumption. Qed. Lemma fun_for_el_prod_l_firstn_p_has : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (x:btp Bp) (pf:In x (firstn m la)), has_comps_p _ _ pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m la) x)) F -> fun_for_el_prod_l_firstn_p pfm F x pf = %0. unfold fun_for_el_prod_l_firstn_p; intros; lr_if_goal; fold hlr; destruct hlr; try contradiction; auto. Qed. Lemma fun_for_el_prod_l_p_has : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns) (x:btp Bp) (pf:In x la), has_comps_p' _ (list_to_set (linds_occ ba_eq_dec_p la x)) F -> fun_for_el_prod_l_p F x pf = %0. intros; rewrite fun_for_el_prod_l_p_eq. rewrite fun_for_el_prod_l_has. reflexivity. rewrite ba_eq_dec_p_eq, has_comps_p_iff' in H; auto. Qed. Lemma fun_for_el_prod_l_firstn_p_not_has : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (x:btp Bp) (pf:In x (firstn m la)), ~ has_comps_p _ _ pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m la) x)) F -> fun_for_el_prod_l_firstn_p pfm F x pf = eps_p x (F |-> lind ba_eq_dec_p _ _ pf). unfold fun_for_el_prod_l_firstn_p; intros; lr_if_goal; fold hlr; destruct hlr; try contradiction; auto. Qed. Lemma fun_for_el_prod_l_p_not_has : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (F:Fin_map (seg (length la)) signe mns) (x:btp Bp) (pf:In x la), ~ has_comps_p' _ (list_to_set (linds_occ ba_eq_dec_p la x)) F -> fun_for_el_prod_l_p F x pf = eps_p x (F |-> lind ba_eq_dec_p _ _ pf). intros; rewrite fun_for_el_prod_l_p_eq, eps_p_eq, fun_for_el_prod_l_not_has; repeat f_equal. rewrite ba_eq_dec_p_eq; auto. rewrite has_comps_p_iff', ba_eq_dec_p_eq in H; auto. Qed. Lemma fun_for_el_prod_l_firstn_p_S_neq : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {n:nat} (pfn:S n <= length la) (F:Fin_map (seg (S n)) signe mns) (x:btp Bp) (pfins:In x (firstn (S n) la)) (pfin: In x (firstn n la)), x <> nth_lt la n pfn -> let pfm' := lt_le_weak _ _ (lt_le_trans _ _ _ (lt_n_Sn _) pfn) in fun_for_el_prod_l_firstn_p pfn F x pfins = fun_for_el_prod_l_firstn_p pfm' (fin_map_seg_rest F n (lt_n_Sn _)) x pfin. intros Bp la n h1 F x h2 h3 hn hle. pose proof (@fun_for_el_prod_l_firstn_S_neq (ba_conv Bp) la n h1 F x h2 h3 hn) as h4. simpl in h4. do 2 rewrite fun_for_el_prod_l_firstn_p_eq. rewrite h4; auto. Qed. Lemma fun_for_el_prod_l_firstn_p_S_eq_has_comps : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {n:nat} (pfn:S n <= length la) (F:Fin_map (seg (S n)) signe mns) (pfins:In (nth_lt la n pfn) (firstn (S n) la)) (pfin: In (nth_lt la n pfn) (firstn n la)), let pfn' := lt_le_weak _ _ (lt_le_trans _ _ _ (lt_n_Sn _) pfn) in has_comps_p _ _ pfn' (list_to_set (linds_occ ba_eq_dec_p (firstn n la) (nth_lt la n pfn))) (fin_map_seg_rest F n (lt_n_Sn _)) -> fun_for_el_prod_l_firstn_p pfn F _ pfins = %0. intros Bp la n h1 F h2 h3 h4 hc. pose proof (@fun_for_el_prod_l_firstn_S_eq_has_comps (ba_conv Bp) la n h1 F h2 h3) as h5. simpl in h5. rewrite has_comps_p_iff0, linds_occ_ba_conv_list' in hc. rewrite fun_for_el_prod_l_firstn_p_eq. apply(h5 hc). Qed. Lemma fun_for_el_prod_l_firstn_p_S_eq_not_has_comps : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {n:nat} (pfn:S n <= length la) (F:Fin_map (seg (S n)) signe mns) (pfins:In (nth_lt la n pfn) (firstn (S n) la)) (pfin: In (nth_lt la n pfn) (firstn n la)), let pfn' := lt_le_weak _ _ (lt_le_trans _ _ _ (lt_n_Sn _) pfn) in ~has_comps_p _ _ pfn' (list_to_set (linds_occ ba_eq_dec_p (firstn n la) (nth_lt la n pfn))) (fin_map_seg_rest F n (lt_n_Sn _)) -> ~has_comps_p _ _ pfn (list_to_set (linds_occ ba_eq_dec_p (firstn (S n) la) (nth_lt la n pfn))) F -> fun_for_el_prod_l_firstn_p pfn F _ pfins = fun_for_el_prod_l_firstn_p pfn' (fin_map_seg_rest F n (lt_n_Sn _)) _ pfin. intros Bp la n h1 F h2 h3 h4 h5 h8. do 2 rewrite fun_for_el_prod_l_firstn_p_eq. rewrite has_comps_p_iff0, linds_occ_ba_conv_list' in h5. rewrite has_comps_p_iff0, linds_occ_ba_conv_list' in h8. apply (@fun_for_el_prod_l_firstn_S_eq_not_has_comps (ba_conv Bp) la n h1 F h2 h3); auto. Qed. (*This iterates over [m] and [list_singularize la].*) Definition list_for_el_prod_l_firstn_p {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) : list (btp Bp) := map_in_dep (fun_in_dep_incl _ _ (fun_for_el_prod_l_firstn_p pfm F) (incl_list_singularize ba_eq_dec_p _)). Lemma list_for_el_prod_l_firstn_p_eq : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns), list_for_el_prod_l_firstn_p pfm F = @list_for_el_prod_l_firstn (ba_conv Bp) _ _ pfm F. intros; unfold list_for_el_prod_l_firstn_p, list_for_el_prod_l_firstn. unfold Atp, btp. rewrite ba_eq_dec_p_eq. apply map_in_dep_ext. intros. apply functional_extensionality_dep. intro. unfold fun_in_dep_incl. rewrite <- fun_for_el_prod_l_firstn_p_eq. reflexivity. Qed. (*This iterates over [list_singularize la].*) Definition list_for_el_prod_l_p {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns) : list (btp Bp) := map_in_dep (fun_in_dep_incl _ _ (fun_for_el_prod_l_p F) (incl_list_singularize ba_eq_dec_p _)). Lemma list_for_el_prod_l_p_eq : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns), list_for_el_prod_l_p F = @list_for_el_prod_l (ba_conv Bp) _ F. intros; unfold list_for_el_prod_l_p, list_for_el_prod_l, Atp, btp. rewrite ba_eq_dec_p_eq. apply map_in_dep_ext. intros. apply functional_extensionality_dep. intro. unfold fun_in_dep_incl. rewrite <- fun_for_el_prod_l_p_eq; auto. Qed. (*This iterates over [linds (firstn m la)].*) Definition list_for_el_prod_l_linds_firstn_p {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) : list (btp Bp) := map_in_dep (fun_for_el_prod_l_linds_firstn_p pfm F). Lemma list_for_el_prod_l_linds_firstn_p_eq : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns), list_for_el_prod_l_linds_firstn_p pfm F = @list_for_el_prod_l_linds_firstn (ba_conv Bp) _ _ pfm F. intros Bp la m h1 F. unfold list_for_el_prod_l_linds_firstn_p, list_for_el_prod_l_linds_firstn. erewrite map_in_dep_functional. apply map_in_dep_ext. intros x h2. apply functional_extensionality. intro h3. rewrite transfer_fun_in_dep_compat. rewrite fun_for_el_prod_l_linds_firstn_p_eq. f_equal. apply proof_irrelevance. Grab Existential Variables. apply linds_ba_conv_list. Qed. (*This iterates over [linds la].*) Definition list_for_el_prod_l_linds_p {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns) : list (btp Bp) := map_in_dep (fun_for_el_prod_l_linds_p F). Lemma list_for_el_prod_l_linds_p_eq : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns), list_for_el_prod_l_linds_p F = @list_for_el_prod_l_linds (ba_conv Bp) _ F. intros Bp la F. unfold list_for_el_prod_l_linds_p, list_for_el_prod_l_linds. erewrite map_in_dep_functional. apply map_in_dep_ext. intros x h1. apply functional_extensionality. intro h3. rewrite transfer_fun_in_dep_compat, fun_for_el_prod_l_linds_p_eq. f_equal. apply proof_irrelevance. Grab Existential Variables. rewrite ba_eq_dec_p_eq; auto. Qed. Lemma list_for_el_prod_l_to_firstn_p : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} (F:Fin_map (seg (length la)) signe mns), list_for_el_prod_l_p F = list_for_el_prod_l_firstn_p (le_refl _) F. intros; rewrite list_for_el_prod_l_p_eq, list_for_el_prod_l_firstn_p_eq, list_for_el_prod_l_to_firstn; auto. Qed. Lemma in_list_for_el_prod_l_firstn_p_iff : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns) (y:btp Bp), In y (list_for_el_prod_l_firstn_p pfm F) <-> exists (x:btp Bp) (pfx:In x (firstn m la)), y = fun_for_el_prod_l_firstn_p pfm F x pfx. intros Bp la m h1 F y. pose proof (@in_list_for_el_prod_l_firstn_iff (ba_conv Bp) la m h1 F y) as h2. rewrite list_for_el_prod_l_firstn_p_eq. unfold Atp, btp. unfold Atp, btp in h2. rewrite h2. split; intro h3; destruct h3 as [x [h4 h5]]; subst; exists x, h4; rewrite fun_for_el_prod_l_firstn_p_eq; auto. Qed. Lemma list_for_el_prod_l_firstn_p_eq_list_for_el_prod_l_linds_firstn_p : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfm:m <= length la) (F:Fin_map (seg m) signe mns), list_for_el_prod_l_firstn_p pfm F = list_for_el_prod_l_linds_firstn_p pfm F. intros Bp la m h1 F. rewrite list_for_el_prod_l_firstn_p_eq, list_for_el_prod_l_linds_firstn_p_eq. rewrite list_for_el_prod_l_firstn_eq_list_for_el_prod_l_linds_firstn; auto. Qed. Lemma incl_list_to_set_list_for_el_prod_l_firstn_p_S : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfsm:S m <= length la) (F:Fin_map (seg (S m)) signe mns), let pfm := le_trans _ _ _ (le_n_Sn _) pfsm in Included (list_to_set (list_for_el_prod_l_firstn_p pfsm F)) (list_to_set (list_for_el_prod_l_firstn_p pfm (fin_map_seg_rest F m (lt_n_Sn m)) ++ (fun_for_el_prod_l_firstn_p pfsm F _ (in_nth_lt_firstn la _ _ pfsm (lt_n_Sn _))) :: nil)). intros Bp la m h1 F h2. do 2 rewrite list_for_el_prod_l_firstn_p_eq. eapply inclusion_transitive. apply (@incl_list_to_set_list_for_el_prod_l_firstn_S (ba_conv Bp)). apply eq_incl. repeat f_equal. rewrite fun_for_el_prod_l_firstn_p_eq; auto. Qed. Lemma list_to_set_list_for_el_prod_l_firstn_p_S_eq_has_has : forall {Bp:Bool_Alg_p T} {la:list (btp Bp)} {m:nat} (pfsm:S m <= length la) (F:Fin_map (seg (S m)) signe mns), let pfm := le_trans _ _ _ (le_n_Sn _) pfsm in has_comps_p _ _ pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m la) (nth_lt la m pfsm))) (fin_map_seg_rest F m (lt_n_Sn _)) -> has_comps_p _ _ pfsm (list_to_set (linds_occ ba_eq_dec_p (firstn (S m) la) (nth_lt la m pfsm))) F -> list_to_set (list_for_el_prod_l_firstn_p pfsm F) = list_to_set (list_for_el_prod_l_firstn_p pfm (fin_map_seg_rest F m (lt_n_Sn m))). intros Bp la m h1 F hle h2 h3. pose proof (@list_to_set_list_for_el_prod_l_firstn_S_eq_has_has (ba_conv Bp) la m h1 F) as h4. simpl in h4. rewrite list_for_el_prod_l_firstn_p_eq. unfold Atp, btp, Btype_p in h4. unfold Atp, btp, Btype_p. rewrite h4. rewrite list_for_el_prod_l_firstn_p_eq; auto. rewrite has_comps_p_iff0, linds_occ_ba_conv_list in h2; auto. simpl in h3. rewrite has_comps_p_iff0, linds_occ_ba_conv_list in h3; auto. Qed. Lemma has_comps_elt_in0_list_for_el_prod_l_firstn_p_S : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm: m <= length l) (F:Fin_map (seg m) signe mns) (x:btp Bp), has_comps_elt_p l m pfm (list_to_set (linds_occ ba_eq_dec_p (firstn m l) x)) x F -> In %0 (list_for_el_prod_l_firstn_p pfm F). intros Bp l m h1 F x h2. rewrite has_comps_elt_p_iff , linds_occ_ba_conv_list' in h2. pose proof (has_comps_elt_in0_list_for_el_prod_l_firstn (ba_conv Bp) h1 F x h2) as h3. rewrite list_for_el_prod_l_firstn_p_eq; auto. Qed. Lemma nin_list_for_el_prod_l_firstn_p_S : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfsm: S m <= length l) (F:Fin_map (seg (S m)) signe mns), let pfm := le_length_removelast _ _ pfsm in ~ In (nth_lt l m pfsm) (firstn m l) -> list_for_el_prod_l_firstn_p pfsm F = list_for_el_prod_l_firstn_p pfm (fin_map_seg_rest F _ (lt_n_Sn _)) ++ eps_p (nth_lt l m pfsm) (F |-> m) :: nil. intros Bp l m h1 F h2 h3. pose proof (nin_list_for_el_prod_l_firstn_S (ba_conv Bp) h1 F) as h4. simpl in h4. specialize (h4 h3). do 2 rewrite list_for_el_prod_l_firstn_p_eq. rewrite h4; auto. Qed. Lemma in_not_has_comps_elt_p_list_for_el_prod_l_firstn_p_S : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfsm: S m <= length l) (F:Fin_map (seg (S m)) signe mns), let pfm := le_length_removelast _ _ pfsm in In (nth_lt l m pfsm) (firstn m l) -> ~has_comps_elt_p l (S m) pfsm (list_to_set (linds_occ ba_eq_dec_p (firstn (S m) l) (nth_lt l m pfsm))) (nth_lt l m pfsm) F -> list_to_set (list_for_el_prod_l_firstn_p pfsm F) = list_to_set (list_for_el_prod_l_firstn_p pfm (fin_map_seg_rest F _ (lt_n_Sn _))). intros Bp l m h1 F h2 h3 h4. do 2 rewrite list_for_el_prod_l_firstn_p_eq. rewrite has_comps_elt_p_iff, linds_occ_ba_conv_list in h4. apply (in_not_has_comps_elt_list_for_el_prod_l_firstn_S (ba_conv Bp)); auto. Qed. (*This is the goal of all this preamble, proving that we can access the elemetary product corresponding to each 2-valued map on list indices, with a product of a _real_ Boolean list, and not some composite of list and map. Yay!*) Lemma el_prod_l_firstn_eq_p : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} {m:nat} (pfm:m <= length l) (F:Fin_map (seg m) signe mns), el_prod_l_firstn_p pfm F = times_list_p (list_for_el_prod_l_firstn_p pfm F). intros Bp l m h1 F. rewrite times_list_p_eq, el_prod_l_firstn_p_eq, list_for_el_prod_l_firstn_p_eq, el_prod_l_firstn_eq; auto. Qed. Lemma el_prod_l_eq_p : forall {Bp:Bool_Alg_p T} {l:list (btp Bp)} (F:Fin_map (seg (length l)) signe mns), el_prod_l_p F = times_list_p (list_for_el_prod_l_p F). intros; rewrite el_prod_l_p_eq, list_for_el_prod_l_p_eq, el_prod_l_eq, times_list_p_eq; auto. Qed. Lemma el_prod_compose_p1_sig_fun_app_proj_eq_p1 : forall {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} (i_E : nat -> sig_set E) (n:nat), surjective i_E -> forall (g: sig_set (Im (seg n) (fun i : nat => proj1_sig (i_E i)) : Ensemble (btp Bp)) -> bt A) (a':Fin_map (Im (seg n) (fun i : nat => proj1_sig (i_E i))) signe mns) (pf:Finite (Im (seg n) (fun i : nat => proj1_sig (i_E i)))), let g' := (fun x:sig_set E => match (in_fin_ens_dec ba_eq_dec_p _ pf (proj1_sig x)) with | left P => g (exist _ _ P) | right _ => 0 end) in el_prod_compose_p1 (sig_fun_app g 0) a' = 0 <-> el_prod_compose_p1 (sig_fun_app g' 0) a' = 0. intros Bp A E i_E n h1 g a' h2 k. unfold k. unfold el_prod_compose_p1. do 2 erewrite times_set_im_im. split; intro h3. eapply times_set_zero_incl. apply h3. red; intros x h4. destruct h4 as [x h4]; subst. econstructor. apply h4. f_equal. unfold sig_fun_app. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. repeat f_equal. apply proof_irrelevance. contradict hr''. econstructor. apply h4. simpl. reflexivity. contradict hr'. apply proj2_sig. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. contradiction. reflexivity. reflexivity. eapply times_set_zero_incl. apply h3. red; intros x h4. destruct h4 as [x h4]; subst. econstructor. apply h4. f_equal. unfold sig_fun_app. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. repeat f_equal. apply proof_irrelevance. contradict hr''. econstructor. apply h4. simpl. reflexivity. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. contradiction. reflexivity. contradict hr. apply proj2_sig. Grab Existential Variables. apply finite_seg. apply finite_seg. Qed. Lemma el_prod_compose_p1_sig_fun_app_im_seg : forall {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} (i_E:nat->sig_set E) (n:nat) (F:Ensemble (btp Bp)) (pff:Finite F) (g : sig_set F -> bt A) (a : Fin_map (Im (seg n) (fun i : nat => proj1_sig (i_E i))) signe mns), Included F E -> surjective i_E -> (el_prod_compose_p1 (sig_fun_app (fun x : sig_set (Im (seg n) (fun i : nat => proj1_sig (i_E i))) => match in_fin_ens_dec ba_eq_dec_p F pff (proj1_sig x) with | left P => g (exist (Inn F) (proj1_sig x) P) | right _ => 0 end) 0) a = 0 <-> el_prod_compose_p1 (sig_fun_app (fun x : sig_set E => match in_fin_ens_dec ba_eq_dec_p F pff (proj1_sig x) with | left P => g (exist (Inn F) (proj1_sig x) P) | right _ => 0 end) 0) a = 0). intros Bp A E i_E n F h1 g a h2 h3. unfold el_prod_compose_p1. do 2 erewrite times_set_im_im. split; intro h4. eapply times_set_zero_incl. apply h4. red; intros x h5. destruct h5 as [x h5]; subst. econstructor. apply h5. repeat f_equal. unfold sig_fun_app. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]. repeat f_equal. apply proof_irrelevance. simpl in hl', hr0; contradiction. simpl in hl'. contradict hr''. apply h2; auto. simpl in hr'. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]. simpl in hl0. contradiction. simpl in hr0. reflexivity. reflexivity. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]. simpl in hl0. contradict hr. econstructor. apply h5. reflexivity. reflexivity. reflexivity. eapply times_set_zero_incl. apply h4. red; intros x h5. destruct h5 as [x h5]; subst. econstructor. apply h5. repeat f_equal. unfold sig_fun_app. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]. repeat f_equal. apply proof_irrelevance. simpl in hl', hr0; contradiction. simpl in hl'. contradict hr''. econstructor. apply h5. reflexivity. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]. simpl in hl0, hr'. contradiction. reflexivity. reflexivity. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. simpl in hl''. contradict hr. apply h2; auto. reflexivity. reflexivity. Grab Existential Variables. apply finite_seg. apply finite_seg. Qed. Lemma el_prod_compose_p1_sig_fun_app_comp : forall {Bp:Bool_Alg_p T} {A:Bool_Alg} {E F:Ensemble (btp Bp)} (a':Fin_map F signe mns) (g:sig_set F -> bt A) (pff:Finite F), Included F E -> (el_prod_compose_p1 (sig_fun_app (fun x : sig_set E => match in_fin_ens_dec ba_eq_dec_p F pff (proj1_sig x) with | left P => g (exist (Inn F) (proj1_sig x) P) | right _ => 0 end) 0) a' = 0 <-> el_prod_compose_p1 (sig_fun_app g 0) a' = 0). intros Bp A E F a' g h1 hi. split; intro h2. unfold el_prod_compose_p1 in h2. unfold el_prod_compose_p1. eapply times_set_zero_incl. apply h2. red; intros x h3. destruct h3 as [x h3]; subst. econstructor. apply h3. repeat f_equal. unfold sig_fun_app. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. repeat f_equal. apply proof_irrelevance. simpl in hl'. contradiction. simpl in hr'. contradiction. contradict hr. apply hi; auto. unfold el_prod_compose_p1 in h2. unfold el_prod_compose_p1. eapply times_set_zero_incl. apply h2. red; intros x h3. destruct h3 as [x h3]; subst. econstructor. apply h3. repeat f_equal. unfold sig_fun_app. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. repeat f_equal. apply proof_irrelevance. simpl in hl'. contradiction. simpl in hr'. contradict hr'. apply hi; auto. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. lr_match_goal; fold hlr; destruct hlr as [hl'' | hr'']. simpl in hl''. contradiction. reflexivity. reflexivity. Qed. Lemma im_only_has_sames_eps_p : forall {Bp:Bool_Alg_p T} {n:nat} (f:nat->btp Bp) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pfo:only_has_sames_p' (map_seg (fun_to_seg f n)) (seg n) F), let F' := only_has_sames_map_seg_to_fin_map_p f F pfo in im_in_ens (seg_fun_to_ens (fun (i : nat) (pf : i < length (map_seg (fun_to_seg f n))) => eps_p (nth_lt (map_seg (fun_to_seg f n)) i pf) (F |-> i))) = Im (Im (seg n) f) (fun i => eps_p i (F' |-> i)). intros Bp n f F h1 F'. pose proof (@im_only_has_sames_eps (ba_conv Bp) n f F) as h2. hfirst h2. rewrite <- only_has_sames_p_iff0'. apply h1. specialize (h2 hf). simpl in h2. rewrite eq_ba_conv_set_iff'. apply Extensionality_Ensembles; red; split; red; intros j h3. destruct h3 as [j h3]; subst. rewrite seg_fun_to_ens_compat. rewrite eps_p_eq. pose proof h3 as h3'. destruct h3' as [h3']. rewrite length_map_seg in h3'. pose proof (in_im_in_ens (fun (i : nat) (pf : i < length (map_seg (fun_to_seg f n))) => @eps (ba_conv Bp) (nth_lt (map_seg (fun_to_seg f n)) i pf) (F |-> i)) j) as h4. hfirst h4. red. rewrite length_map_seg. omega. specialize (h4 hf0). simpl in h4. unfold seg_fun_to_ens in h2. rewrite nth_lt_map_seg', fun_to_seg_compat in h4. terml h2. let P := type of h4 in match P with Inn ?S' _ => pose S' as S end. rewrite nth_lt_map_seg', fun_to_seg_compat. assert (h5:x = S). unfold x, S. apply Extensionality_Ensembles; red; split; red; intros a h5. destruct h5 as [a h5]; subst. rewrite nth_lt_map_seg', fun_to_seg_compat. econstructor. rewrite nth_lt_map_seg', fun_to_seg_compat. reflexivity. destruct h5 as [a h5]; subst. econstructor. do 2 rewrite nth_lt_map_seg', fun_to_seg_compat. reflexivity. fold x S in h2, h4. termr h2. match goal with |- Inn (ba_conv_set' ?E') _ => pose E' as E end. fold E. assert (h6:E = y). unfold E, y. do 2 rewrite im_im. apply im_ext_in. intros i h6. destruct h6. rewrite eps_p_eq. f_equal. unfold F'. rewrite only_has_sames_map_seg_to_fin_map_p_eq. repeat f_equal. apply proof_irrelevance. rewrite h6. fold y in h2. rewrite <- h2. rewrite h5; auto. destruct h3 as [j h3]; subst. rewrite eps_p_eq. pose proof h3 as h3'. destruct h3' as [j h3']; subst. destruct h3' as [h3']. pose proof (Im_intro _ _ (@Im nat (bt (ba_conv Bp)) (seg n) f) (fun i : sig_set (A_p T (Bc_p T Bp)) => @eps (ba_conv Bp) i (only_has_sames_map_seg_to_fin_map (ba_conv_fun2 f) F hf |-> i)) _ h3 _ eq_refl) as h4. termr h2. let P := type of h4 in match P with Inn ?S' _ => pose S' as S end. assert (h5:y = S); auto. fold y S in h2, h4. terml h2. match goal with |- Inn (ba_conv_set' ?E') _ => pose E' as E end. fold E. assert (h6:E = x); auto. rewrite h6. fold x in h2. rewrite h2. rewrite h5. unfold F'; auto. rewrite only_has_sames_map_seg_to_fin_map_p_compat; auto. rewrite <- ba_conv_elt_compat'; auto. rewrite only_has_sames_map_seg_to_fin_map_compat in h4; auto. Grab Existential Variables. rewrite length_map_seg. constructor. red in h5. rewrite length_map_seg in h5. omega. rewrite length_map_seg in h5. red. rewrite length_map_seg. destruct h5; auto. Qed. Lemma only_has_sames_map_el_prod_p : forall {Bp:Bool_Alg_p T} {n:nat} (f:nat->btp Bp) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pfo:only_has_sames_p' (map_seg (fun_to_seg f n)) (seg n) F), let F' := only_has_sames_map_seg_to_fin_map_p f F pfo in el_prod_l_p F = el_prod_p F'. intros Bp n f F h1 F'. pose proof (@only_has_sames_map_el_prod (ba_conv Bp) _ f F) as h2. hfirst h2. rewrite <- only_has_sames_p_iff0'. apply h1. specialize (h2 hf). simpl in h2. rewrite el_prod_l_p_eq. rewrite el_prod_p_eq; auto. unfold Atp, btp, bt, Btype_p in h2. unfold Atp, btp, bt, Btype_p. simpl in h2; simpl. unfold ba_conv_fun2 in h2. rewrite h2. f_equal. apply fin_map_ext. intros x h3. destruct h3 as [x h3]; subst. destruct h3 as [h3]. unfold F'. rewrite only_has_sames_map_seg_to_fin_map_compat; auto. unfold ba_conv_fin_map_dom. pose proof (only_has_sames_map_seg_to_fin_map_p_compat f F h1) as h4. simpl in h4. symmetry. apply h4; auto. Qed. Lemma only_has_sames_map_el_prod_compose_p : forall {Bp:Bool_Alg_p T} {n:nat} (f:nat->btp Bp) (g:btp Bp->Atp) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pfo:only_has_sames_p' (map_seg (fun_to_seg f n)) (seg n) F), let F' := only_has_sames_map_seg_to_fin_map_p f F pfo in el_prod_l_compose_p F g = el_prod_compose_p g F'. intros Bp n f g F h1 F'. pose proof (@only_has_sames_map_el_prod_compose (ba_conv Ap) (ba_conv Bp) n f g F) as h3. simpl in h3. hfirst h3. rewrite <- only_has_sames_p_iff0'. apply h1. specialize (h3 hf). rewrite el_prod_l_compose_p_eq, el_prod_compose_p_eq. unfold el_prod_l_compose, el_prod_compose. unfold el_prod_l_compose, el_prod_compose in h3. let P := type of h3 in match P with times_list ?l' = _ => pose l' as l end. match goal with |- times_list ?l' = _ => pose l' as l0 end. assert (h4:l = l0); auto. fold l0. fold l in h3. rewrite h4 in h3. rewrite h3. apply f_equal_dep. do 2 rewrite im_im. apply im_ext_in. intros i h5. rewrite only_has_sames_map_seg_to_fin_map_compat. f_equal. unfold F'. pose proof (only_has_sames_map_seg_to_fin_map_p_compat f F h1) as h6. simpl in h6. destruct h5 as [h5]. specialize (h6 _ h5); auto. destruct h5; auto. Qed. Lemma only_has_sames_map_el_prod_compose_p1 : forall {B:Bool_Alg} {n:nat} (f:nat->Atp) (g:Atp->bt B) (F:Fin_map (seg (length (map_seg (fun_to_seg f n)))) signe mns) (pfo:only_has_sames_p' (map_seg (fun_to_seg f n)) (seg n) F), let F' := only_has_sames_map_seg_to_fin_map_p f F pfo in el_prod_l_compose_p1 F g = el_prod_compose_p1 g F'. intros B n f g F h1 F'. pose proof (@only_has_sames_map_el_prod_compose B (ba_conv Ap) n f g F) as h2. hfirst h2. rewrite <- only_has_sames_p_iff0'; auto. specialize (h2 hf). rewrite el_prod_l_compose_p1_eq. simpl in h2. rewrite el_prod_compose_p1_eq0. unfold el_prod_l_compose, el_prod_compose. unfold el_prod_l_compose, el_prod_compose in h2. unfold F'. let P := type of h2 in match P with times_list ?l' = _ => pose l' as l end. match goal with |- times_list ?l' = _ => pose l' as l0 end. fold l0. fold l in h2. assert (h4:l = l0); auto. rewrite <- h4. rewrite h2. apply f_equal_dep. do 2 rewrite im_im. apply im_ext_in. intros i h5. destruct h5 as [h5]. pose proof (@only_has_sames_map_seg_to_fin_map_compat (ba_conv Ap) n f F hf) as h7. simpl in h7. specialize (h7 i h5). rewrite h7. pose proof (@only_has_sames_map_seg_to_fin_map_p_compat Ap n f F h1) as h8. simpl in h8. specialize (h8 i h5). clear h7. f_equal; auto. Qed. Lemma only_has_sames_el_prod_fin_map_seg_to_im_in_ens_p : forall (l:list Atp) (m:nat) (pfm:m <= length l) (F:Fin_map (seg m) signe mns), only_has_sames_p l m pfm (seg m) F -> el_prod_p (fin_map_seg_to_im_in_ens ba_eq_dec_p F l pfm) = el_prod_l_firstn_p pfm F. intros l m h1 F h2. pose proof (@only_has_sames_el_prod_fin_map_seg_to_im_in_ens (ba_conv Ap) l m h1 F) as h3. hfirst h3. rewrite <- only_has_sames_p_iff0; auto. specialize (h3 hf). rewrite el_prod_p_eq. rewrite ba_eq_dec_p_eq. rewrite h3 at 1. rewrite el_prod_l_firstn_p_eq; auto. Qed. Lemma only_has_sames_el_prod_compose_fin_map_seg_to_im_in_ens_p1 : forall {A:Bool_Alg} (l:list Atp) (la:list (bt A)) (m:nat) (pfm:m <= length l) (pfe:length l = length la) (F:Fin_map (seg m) signe mns) (pfle:linds ba_eq_dec_p l = linds ba_eq_dec la), let f := proj1_sig_fun (two_lists_to_sig_fun ba_eq_dec_p ba_eq_dec l la pfle) in let pfm' := le_eq_trans _ _ _ pfm pfe in only_has_sames la m pfm' (seg m) F -> el_prod_compose_p1 (sig_fun_app f 0) (fin_map_seg_to_im_in_ens ba_eq_dec_p F l pfm) = el_prod_l_firstn pfm' F. intros A l la m ha h2 F h3 f h5 h6. pose proof (@only_has_sames_el_prod_compose_fin_map_seg_to_im_in_ens (ba_conv Ap) A l la m ha h2 F) as h7. simpl in h7. pose proof h3 as h3'. rewrite ba_eq_dec_p_eq in h3'. specialize (h7 h3' h6). unfold f. generalize h3. rewrite ba_eq_dec_p_eq. intro h4. rewrite el_prod_compose_p1_eq0. assert (h4 = h3'). apply proof_irrelevance. subst. rewrite h7 at 1. reflexivity. Qed. Lemma el_prod_compose_sig_fun_app_only_has_sames_p1 : forall {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} {D:Ensemble (bt A)} (i_E:nat->sig_set E) (i_D:nat->sig_set D) (pfse:surjective i_E) (pfsd:surjective i_D) (n:nat) (Fb : Fin_map (seg (length (map_seg (fun (i : nat) (_ : i < n) => proj1_sig (i_E i))))) signe mns) (pfo: only_has_sames_p' (map_seg (fun_to_seg (proj1_sig_fun i_E) n)) (seg n) Fb), fun_linds_eq_nat (sig_set_eq_dec ba_eq_dec_p E) (sig_set_eq_dec ba_eq_dec D) i_E i_D -> el_prod_compose_p1 (sig_fun_app (surj_nat_to_fun_sig i_E i_D pfse) 0) (only_has_sames_map_seg_to_fin_map_p (proj1_sig_fun i_E) Fb pfo) = el_prod_l (fin_map_seg_transfer Fb (eq_trans (length_map_seg (fun (i : nat) (_ : i < n) => proj1_sig (i_E i))) (eq_sym (length_map_seg (fun (i : nat) (_ : i < n) => proj1_sig (i_D i)))))). intros Bp A E D i_E i_D hse hsd n Fb h3 hfl. unfold el_prod_compose_p1, el_prod_l. gen0. rewrite im_im. intro h5. rewrite times_list_compat. symmetry. gen0. rewrite list_to_set_map_seg. intro h6. apply times_set_functional. apply Extensionality_Ensembles; red; split; red; intros x h7. destruct h7 as [i h7]; subst. destruct h7 as [h7]. pose proof h7 as h7'. rewrite length_map_seg in h7'. apply Im_intro with i. constructor; auto. rewrite seg_fun_to_ens_compat. rewrite nth_lt_map_seg'. f_equal. unfold sig_fun_app. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. pose proof (surj_nat_to_fun_sig_compat ba_eq_dec_p ba_eq_dec i_E i_D hse hsd hfl i) as h8; auto. unfold proj1_sig_fun. assert (hl = proj2_sig _). apply proof_irrelevance. subst. rewrite <- unfold_sig; auto. contradict hr. unfold proj1_sig_fun. apply proj2_sig. pose proof (fin_map_seg_transfer_compat Fb (eq_trans (length_map_seg (fun (i0 : nat) (_ : i0 < n) => proj1_sig (i_E i0))) (eq_sym (length_map_seg (fun (i0 : nat) (_ : i0 < n) => proj1_sig (i_D i0)))))) as h8. simpl in h8. specialize (h8 (exist (fun c => c < length (map_seg (fun i (_:i < n) => proj1_sig (i_D i)))) _ h7)). simpl in h8. simpl. unfold proj1_sig_fun. let P := type of h8 in match P with fin_map_seg_transfer Fb ?pf' |-> i = _ => pose pf' as pf1 end. match goal with |- fin_map_seg_transfer Fb ?pf' |-> i = _ => pose pf' as pf2 end. fold pf1. fold pf2 in h8. assert (h9: pf1 = pf2). apply proof_irrelevance. rewrite h9. rewrite h8 at 1. pose proof (only_has_sames_map_seg_to_fin_map_p_compat (fun x => proj1_sig (i_E x)) Fb h3 i h7') as h10. simpl in h10. rewrite h10 at 1. f_equal. rewrite segT_transfer_r_compat. simpl; auto. destruct h7 as [x h7]; subst. destruct h7 as [h7]. pose proof (in_seg x _ (lt_eq_trans _ _ _ h7 (eq_sym (length_map_seg (fun i (_:i < n) => proj1_sig (i_D i)))))) as hi. apply im_in_ens_intro with x hi. rewrite seg_fun_to_ens_compat. rewrite nth_lt_map_seg'. f_equal. unfold sig_fun_app. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. assert (hl = proj2_sig _). apply proof_irrelevance. subst. unfold proj1_sig_fun. rewrite <- unfold_sig. pose proof (surj_nat_to_fun_sig_compat ba_eq_dec_p ba_eq_dec i_E i_D hse hsd hfl x) as h8. rewrite h8. reflexivity. contradict hr. unfold proj1_sig_fun. apply proj2_sig. rewrite only_has_sames_map_seg_to_fin_map_p_compat. assert (h8:x = proj1_sig (exist (fun c => c < length (map_seg (fun i (_:i < n) => proj1_sig (i_D i)))) _ (seg_in _ _ hi))); auto. rewrite h8 at 2. rewrite fin_map_seg_transfer_compat. rewrite segT_transfer_r_compat. simpl. reflexivity. assumption. Qed. Lemma el_prod_l_has_comps_map_seg : forall {Bp:Bool_Alg_p T} {A:Bool_Alg} (E:Ensemble (btp Bp)) (D:Ensemble (bt A)) (i_E:nat->sig_set E) (i_D:nat->sig_set D) (n:nat) (F : Fin_map (seg (length (map_seg (fun (i : nat) (_ : i < n) => proj1_sig (i_E i))))) signe mns), has_comps_p' (map_seg (fun_to_seg (proj1_sig_fun i_E) n)) (seg n) F -> el_prod_l_p F = %0. intros Bp A E D i_E i_D n F h4. pose proof (length_map_seg (fun_to_seg (proj1_sig_fun i_E) n)) as h5. rewrite <- h5 in h4 at 3. apply (has_comps_el_prod_p_l0 _ _ h4). Qed. Lemma el_prod_l_has_comps_map_seg_transfer : forall {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} {D:Ensemble (bt A)} (i_E:nat->sig_set E) (i_D:nat->sig_set D) (n:nat) (F:Fin_map (seg (length (map_seg (fun (i : nat) (_ : i < n) => proj1_sig (i_E i))))) signe mns) (pfse:surjective i_E) (pfsd:surjective i_D) (pffe:fun_linds_eq_nat (sig_set_eq_dec ba_eq_dec_p E) (sig_set_eq_dec ba_eq_dec D) i_E i_D) (pfh:has_comps_p' (map_seg (fun_to_seg (proj1_sig_fun i_E) n)) (seg n) F), el_prod_l (fin_map_seg_transfer F (eq_trans (length_map_seg (fun (i : nat) (_ : i < n) => proj1_sig (i_E i))) (eq_sym (length_map_seg (fun (i : nat) (_ : i < n) => proj1_sig (i_D i)))))) = 0. intros Bp A E D i_E i_D n F h1 h2 h3 h4. pose proof (length_map_seg (fun_to_seg (proj1_sig_fun i_E) n)) as h5. rewrite <- h5 in h4 at 3. pose proof (has_comps_el_prod_p_l0 _ _ h4) as h6. rewrite el_prod_l_p_eq in h6. inversion h4. do 2 rewrite nth_lt_map_seg', fun_to_seg_compat, proj1_sig_fun_compat in H3. rewrite length_map_seg in pfc, pfd. pose proof (in_map_seg (fun_to_seg (proj1_sig_fun i_E) n) c pfc) as h8. rewrite fun_to_seg_compat, proj1_sig_fun_compat in h8. pose proof (lt_eq_trans _ _ _ pfc (eq_sym (length_map_seg (fun i (_:i < n) => i_D i)))) as h9. pose proof (lt_eq_trans _ _ _ pfd (eq_sym (length_map_seg (fun i (_:i < n) => i_D i)))) as h10. unfold el_prod_l. pose proof (times_list_opps). match goal with |- times_list ?lm = _ => pose proof (times_list_opps _ lm (eps (proj1_sig (i_D c)) (F |-> c)) (eps (proj1_sig (i_D d)) (F |-> d))) as h12 end. hfirst h12. rewrite in_map_seg_iff. exists c. ex_goal. rewrite length_map_seg; auto. exists hex. rewrite nth_lt_map_seg'. rewrite fin_map_seg_transfer_compat'; auto. specialize (h12 hf). hfirst h12. rewrite in_map_seg_iff. exists d. ex_goal. rewrite length_map_seg; auto. exists hex. rewrite nth_lt_map_seg'. rewrite fin_map_seg_transfer_compat'; auto. specialize (h12 hf0). pose proof (fun_linds_eq_nat_linds_occ (sig_set_eq_dec ba_eq_dec_p E) (sig_set_eq_dec ba_eq_dec D) _ _ h3 c n pfc) as h13. pose proof (fun_linds_eq_nat_linds_occ (sig_set_eq_dec ba_eq_dec_p E) (sig_set_eq_dec ba_eq_dec D) _ _ h3 d n pfd) as h14. pose proof H3 as h15. apply proj1_sig_injective in h15. rewrite <- h15 in h14. rewrite h13 in h14. apply inj_linds_occ in h14. hfirst h12. rewrite h14. pose proof H4 as h16. apply neq_sym in h16. apply sign_opp in h16. unfold fun_to_seg, proj1_sig_fun in h16. rewrite h16. rewrite eps_opp. rewrite <- doub_neg; auto. specialize (h12 hf1); auto. rewrite in_map_seg_iff. exists c, pfc. unfold fun_to_seg; auto. rewrite in_map_seg_iff. exists d, pfd. unfold fun_to_seg; auto. Qed. Lemma el_prod_compose_sig_fun_app_p1 : forall {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} {D:Ensemble (bt A)} (i_E:nat->sig_set E) (i_D:nat->sig_set D) (pfse:surjective i_E) (pfsd:surjective i_D) (n:nat) (pffl:fun_linds_eq_nat (sig_set_eq_dec ba_eq_dec_p E) (sig_set_eq_dec ba_eq_dec D) i_E i_D) (F:Fin_map (Im (seg n) (fun i : nat => proj1_sig (i_E i))) signe mns), let F' := fin_map_im_seg_to_seg_length_p (proj1_sig_fun i_E) F in let g := surj_nat_to_fun_sig i_E i_D pfse in el_prod_compose_p1 (sig_fun_app g 0) F = el_prod_l (fin_map_seg_transfer F' (eq_trans (length_map_seg (fun (i : nat) (_ : i < n) => proj1_sig (i_E i))) (eq_sym (length_map_seg (fun (i : nat) (_ : i < n) => proj1_sig (i_D i)))))). intros Bp A E D i_E i_D hse hsd n hfl F F' g. unfold el_prod_compose_p1. unfold el_prod_l. rewrite times_list_compat. apply times_set_functional. rewrite im_im. rewrite list_to_set_map_seg. apply Extensionality_Ensembles; red; split; red; intros x h7. destruct h7 as [i h7]; subst. pose proof h7 as h7'. rewrite <- (length_map_seg (fun (i : nat) (_ : i < n) => proj1_sig_fun i_D i)) in h7. econstructor. rewrite seg_fun_to_ens_compat. rewrite nth_lt_map_seg'. rewrite fin_map_seg_transfer_compat. rewrite segT_transfer_r_compat; simpl. unfold sig_fun_app, g, F'. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. pose proof (surj_nat_to_fun_sig_compat ba_eq_dec_p ba_eq_dec i_E i_D hse hsd hfl i) as h8. assert (h9:hl = proj2_sig _). apply proof_irrelevance. subst. rewrite <- unfold_sig. rewrite h8. repeat f_equal. destruct h7 as [h7]. pose proof (f_equal (@proj1_sig _ (fun m : nat => m < @length (bt A) (@map_seg (bt A) n (fun (i0 : nat) (_ : i0 < n) => @proj1_sig (bt A) (fun x : bt A => Inn D x) (i_D i0))))) (eq_refl (exist (fun j => j < length (map_seg (fun (i0 : nat) (_ : i0 < n) => proj1_sig_fun i_D i0))) _ h7))) as h9. rewrite h9 at 1. f_equal. apply f_equal_dep. reflexivity. simpl. rewrite fin_map_im_seg_to_seg_length_p_compat; auto. destruct h7'; auto. contradict hr. apply proj2_sig. destruct h7 as [i h7]; subst. apply Im_intro with i. rewrite length_map_seg in h7; auto. rewrite seg_fun_to_ens_compat. unfold sig_fun_app. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. rewrite nth_lt_map_seg'. unfold g. pose proof (surj_nat_to_fun_sig_compat ba_eq_dec_p ba_eq_dec i_E i_D hse hsd hfl i) as h8. assert (h9:hl = proj2_sig _). apply proof_irrelevance. subst. rewrite <- unfold_sig. rewrite h8 at 1. pose proof (fin_map_seg_transfer_compat F' (eq_trans (length_map_seg (fun (i0 : nat) (_ : i0 < n) => proj1_sig (i_E i0))) (eq_sym (length_map_seg (fun (i0 : nat) (_ : i0 < n) => proj1_sig (i_D i0)))))) as h9. simpl in h9. destruct h7 as [h7]. specialize (h9 (exist (fun c => c < (length (map_seg (fun (i0 : nat) (_ : i0 < n) => proj1_sig (i_D i0))))) _ h7)). simpl in h9. rewrite h9 at 1. rewrite segT_transfer_r_compat. simpl. unfold F'. rewrite fin_map_im_seg_to_seg_length_p_compat. unfold proj1_sig_fun; auto. pose proof h7 as h10. rewrite length_map_seg in h10; auto. contradict hr. apply proj2_sig. Grab Existential Variables. simpl. rewrite length_map_seg. destruct h7'; auto. simpl; auto. Qed. Lemma el_prod_l_fin_map_im_seg_to_seg_length_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (i_E:nat->sig_set E) (n:nat) (F:Fin_map (Im (seg n) (fun i : nat => proj1_sig (i_E i))) signe mns), let F' := fin_map_im_seg_to_seg_length_p (proj1_sig_fun i_E) F in el_prod_p F = el_prod_l_p F'. intros Bp E i_E n F F'. unfold F', el_prod_p, el_prod_l_p. rewrite times_list_compat_p. apply f_equal_dep. rewrite im_im. rewrite list_to_set_map_seg. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]; subst. destruct h1 as [h1]. econstructor. rewrite seg_fun_to_ens_compat. rewrite nth_lt_map_seg'. rewrite fin_map_im_seg_to_seg_length_p_compat. unfold proj1_sig_fun. reflexivity. assumption. destruct h1 as [x h1]; subst. destruct h1 as [h1]. pose proof h1 as h1'. rewrite length_map_seg in h1'. apply Im_intro with x. constructor; auto. rewrite seg_fun_to_ens_compat. rewrite nth_lt_map_seg'. rewrite fin_map_im_seg_to_seg_length_p_compat. reflexivity. assumption. Grab Existential Variables. constructor. rewrite length_map_seg; auto. Qed. End ElProdLP. Lemma el_prod_le_ai_p : forall {E:Ensemble Atp} (a:Fin_map E signe mns) (i:Atp), Inn E i -> le_p (el_prod_p a) (eps_p i (a |-> i)). intros E a i h1. unfold el_prod_p. apply le_times_set_p. apply Im_intro with i; auto. Qed. Lemma ai_le_el_sum_p : forall {E:Ensemble Atp} (a:Fin_map E signe mns) (i:Atp), Inn E i -> le_p (eps_p i (a |-> i)) (el_sum_p a). intros E a i h1. unfold el_sum_p. apply le_plus_set_p. apply Im_intro with i; auto. Qed. Lemma non_zero_el_prod_inj_p : forall {E:Ensemble Atp} (f g:Fin_map E signe mns), el_prod_p f <> %0 -> el_prod_p f = el_prod_p g -> im_eps_p f = im_eps_p g. intros E f g h1 h2. rewrite im_eps_p_eq. rewrite (im_eps_p_eq g). rewrite el_prod_p_eq in h1. rewrite el_prod_p_eq in h2. rewrite (el_prod_p_eq g) in h2. rewrite ba_conv_zero in h1. pose proof (non_zero_el_prod_inj _ (ba_conv_fin_map_dom f) (ba_conv_fin_map_dom g) h1 h2) as h3. rewrite h3 at 1. reflexivity. Qed. Lemma el_prod_disjoint_p : forall (E:Ensemble Atp) (a b:Fin_map E signe mns), a <> b -> (el_prod_p a) %* (el_prod_p b) = %0. intros E a b h1. pose proof (el_prod_disjoint _ _ (ba_conv_fin_map_dom a) (ba_conv_fin_map_dom b) h1) as h2. rewrite el_prod_p_eq. rewrite (el_prod_p_eq b). rewrite h2 at 1. reflexivity. Qed. Lemma el_prod_compose_le_ai_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (a:Fin_map E signe mns) (g:(btp Bp)->Atp) (i:(btp Bp)), Inn E i -> le_p (el_prod_compose_p g a) (eps_p (g i) (a |-> i)). intros B E a g i h1. unfold el_prod_compose_p. apply le_times_set_p. apply Im_intro with i; auto. Qed. Lemma ai_le_el_sum_compose_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (a:Fin_map E signe mns) (g:(btp Bp)->Atp) (i:(btp Bp)), Inn E i -> le_p (eps_p (g i) (a |-> i)) (el_sum_compose_p g a). intros B E a g i h1. unfold el_sum_compose_p. apply le_plus_set_p. apply Im_intro with i; auto. Qed. Lemma el_prod_compose_disjoint_p : forall {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (g:(btp Bp)->Atp) (a b:Fin_map E signe mns), a <> b -> (el_prod_compose_p g a) %* (el_prod_compose_p g b) = %0. intros Bp E g a b h1. pose proof (el_prod_compose_disjoint (ba_conv Ap) _ (ba_conv_fun1 g) (ba_conv_fin_map_dom a) (ba_conv_fin_map_dom b) h1) as h2. rewrite el_prod_compose_p_eq. rewrite (el_prod_compose_p_eq g b). rewrite h2 at 1. reflexivity. Qed. Lemma non_zero_el_prod_compose_inj_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (f f':Fin_map E signe mns) (g:(btp Bp)->Atp), el_prod_compose_p g f <> %0 -> el_prod_compose_p g f = el_prod_compose_p g f' -> im_eps_compose_p f g = im_eps_compose_p f' g. intros Bp E f f' g h1 h2. rewrite im_eps_compose_p_eq. rewrite (im_eps_compose_p_eq f'). rewrite el_prod_compose_p_eq in h1. rewrite el_prod_compose_p_eq in h2. rewrite (el_prod_compose_p_eq g f') in h2. rewrite ba_conv_zero in h1. pose proof (non_zero_el_prod_compose_inj (ba_conv Ap) (ba_conv_fin_map_dom f) (ba_conv_fin_map_dom f') (ba_conv_fun1 g) h1 h2) as h3. rewrite h3 at 1. reflexivity. Qed. Lemma non_zero_el_prod_compose_constant_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:btp Bp->Atp) (f:Fin_map E signe mns), el_prod_compose_p g f <> %0 -> forall (x x':btp Bp), Inn E x -> Inn E x' -> g x = g x' -> f |-> x = f |-> x'. intros Bp E g f h1 x x' h2 h3 h4. rewrite el_prod_compose_p_eq in h1. pose proof (non_zero_el_prod_compose_constant (ba_conv Ap) (ba_conv_fun1 g) (ba_conv_fin_map_dom f) h1 x x' h2 h3 h4) as h5. assumption. Qed. Lemma el_prod_disjoint_p' : forall (E:Ensemble Atp) (X Y:Ensemble (Fin_map E signe mns)), Inhabited [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] -> (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_p (fst x) %* el_prod_p (snd x))) = Singleton %0. intros E X Y h1. pose (Im X (ba_conv_fin_map_dom (def:=mns))) as X'. pose (Im Y (ba_conv_fin_map_dom (def:=mns))) as Y'. assert (h2:Inhabited [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X' Y') x /\ fst x <> snd x]). destruct h1 as [f h1]. destruct f as [fx fy]. apply Inhabited_intro with (ba_conv_fin_map_dom fx, ba_conv_fin_map_dom fy). constructor. simpl. inversion h1 as [h2]. simpl in h2. destruct h2 as [h2a h2b]. destruct h2a as [h2a]. simpl in h2a; destruct h2a as [h2a h2c]. split. constructor; simpl. split. apply Im_intro with fx; auto. apply Im_intro with fy; auto. assumption. pose proof (el_prod_disjoint' (ba_conv Ap) (ba_conv_set E) (Im X (ba_conv_fin_map_dom (def:=mns))) (Im Y (ba_conv_fin_map_dom (def:=mns))) h2) as h3. rewrite ba_conv_zero. rewrite <- h3 at 1. f_equal. apply Extensionality_Ensembles. red. split. red. intros f h4. destruct h4 as [h4]. destruct f as [fx fy]. destruct h4 as [h4 h5]. simpl in h5. constructor. simpl. split; auto. constructor; simpl. destruct h4 as [h4]. simpl in h4; destruct h4 as [h4a h4b]. split. apply Im_intro with fx; auto. apply Im_intro with fy; auto. red. intros f h4. destruct h4 as [h4]. destruct h4 as [h4a h4b]. destruct h4a as [h4a]. destruct h4a as [h4a h4c]. destruct f as [fx fy]. simpl in h4a, h4b, h4c. constructor; simpl. destruct h4a as [fx h4a]. subst. destruct h4c as [fy h4c]. subst. constructor. constructor; auto. assumption. apply functional_extensionality. intro f. destruct f as [fx fy]. simpl. rewrite el_prod_p_eq. rewrite (el_prod_p_eq fy). reflexivity. Qed. Lemma el_prod_compose_disjoint_p' : forall {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (g:(btp Bp)->Atp) (X Y:Ensemble (Fin_map E signe mns)), Inhabited [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] -> (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_compose_p g (fst x) %* el_prod_compose_p g (snd x))) = Singleton %0. intros Bp E g X Y h1. pose (Im X (ba_conv_fin_map_dom (def:=mns))) as X'. pose (Im Y (ba_conv_fin_map_dom (def:=mns))) as Y'. assert (h2:Inhabited [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X' Y') x /\ fst x <> snd x]). destruct h1 as [f h1]. destruct f as [fx fy]. apply Inhabited_intro with (ba_conv_fin_map_dom fx, ba_conv_fin_map_dom fy). constructor. simpl. inversion h1 as [h2]. simpl in h2. destruct h2 as [h2a h2b]. destruct h2a as [h2a]. simpl in h2a; destruct h2a as [h2a h2c]. split. constructor; simpl. split. apply Im_intro with fx; auto. apply Im_intro with fy; auto. assumption. pose proof (el_prod_compose_disjoint' (ba_conv Ap) (ba_conv_set E) (ba_conv_fun1 g) (Im X (ba_conv_fin_map_dom (def:=mns))) (Im Y (ba_conv_fin_map_dom (def:=mns))) h2) as h3. rewrite ba_conv_zero. rewrite <- h3 at 1. f_equal. apply Extensionality_Ensembles. red. split. red. intros f h4. destruct h4 as [h4]. destruct f as [fx fy]. destruct h4 as [h4 h5]. simpl in h5. constructor. simpl. split; auto. constructor; simpl. destruct h4 as [h4]. simpl in h4; destruct h4 as [h4a h4b]. split. apply Im_intro with fx; auto. apply Im_intro with fy; auto. red. intros f h4. destruct h4 as [h4]. destruct h4 as [h4a h4b]. destruct h4a as [h4a]. destruct h4a as [h4a h4c]. destruct f as [fx fy]. simpl in h4a, h4b, h4c. constructor; simpl. destruct h4a as [fx h4a]. subst. destruct h4c as [fy h4c]. subst. constructor. constructor; auto. assumption. apply functional_extensionality. intro f. destruct f as [fx fy]. simpl. rewrite el_prod_compose_p_eq. rewrite (el_prod_compose_p_eq (ba_conv_fun1 g) fy) at 1. reflexivity. Qed. Lemma plus_set_el_prod_disjoint_p : forall (E:Ensemble Atp) (X Y:Ensemble (Fin_map E signe mns)) (pf:Finite [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x]), plus_set_p (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_p (fst x) %* el_prod_p (snd x))) (finite_image _ _ _ _ pf) = %0. intros E X Y h1. pose (Im X (ba_conv_fin_map_dom (def:=mns))) as X'. pose (Im Y (ba_conv_fin_map_dom (def:=mns))) as Y'. assert (h0: [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X' Y') x /\ fst x <> snd x] = Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun pr=> (ba_conv_fin_map_dom (fst pr), ba_conv_fin_map_dom (snd pr)))). apply Extensionality_Ensembles. red. split. red. intros pr h2. apply Im_intro with pr; auto. destruct h2 as [h2]. destruct h2 as [h2a h2b]. destruct h2a as [h2a]. destruct h2a as [h2a h2c]. destruct pr as [fx fy]. simpl in h2a, h2b, h2c. destruct h2a as [fx h2a]. subst. destruct h2c as [fy h2c]. subst. constructor; simpl. split. constructor. simpl. split; auto. assumption. apply surjective_pairing. red. intros f h2. destruct h2 as [pr h2]. subst. destruct h2 as [h2]. destruct h2 as [h2a h2b]. destruct h2a as [h2a]. destruct h2a as [h2a h2c]. destruct pr as [fx fy]. simpl in h2a, h2b, h2c. simpl. constructor; simpl. split. constructor; simpl. split. apply Im_intro with fx; auto. apply Im_intro with fy; auto. assumption. assert (h2:Finite [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X' Y') x /\ fst x <> snd x]). rewrite h0. apply finite_image; auto. pose proof (plus_set_el_prod_disjoint (ba_conv Ap) E X' Y' h2) as h3. rewrite plus_set_p_eq. rewrite ba_conv_zero. rewrite <- h3. unfold ba_conv_elt. unfold ba_conv_type. rewrite transfer_eq_refl. unfold ba_conv_set. unfold ba_conv_type. apply (plus_set_functional (B:=ba_conv Ap)). rewrite transfer_dep_eq_refl. f_equal. unfold Atp, bt. unfold Atp, bt in h0. rewrite h0 at 1. pose proof (im_id [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x]) as h4. unfold Atp, bt in h4. unfold Atp, bt. rewrite h4 at 1. f_equal. apply functional_extensionality. intro x. unfold ba_conv_fin_map_dom. unfold id. apply surjective_pairing. apply functional_extensionality. intro f. destruct f as [fx fy]. simpl. rewrite el_prod_p_eq. rewrite (el_prod_p_eq fy). reflexivity. Qed. Lemma plus_set_el_prod_disjoint_p' : forall (E:Ensemble Atp) (X Y:Ensemble (Fin_map E signe mns)) (pf:Finite [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x]) (pfi:Finite (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_p (fst x) %* el_prod_p (snd x)))), plus_set_p (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_p (fst x) %* el_prod_p (snd x))) pfi = %0. intros E X Y h1 h2. pose proof (plus_set_el_prod_disjoint_p _ X Y h1) as h3. rewrite <- h3. apply plus_set_functional_p. reflexivity. Qed. Lemma plus_set_el_prod_compose_disjoint_p : forall {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (g:(btp Bp)->Atp) (X Y:Ensemble (Fin_map E signe mns)) (pf:Finite [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x]), plus_set_p (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_compose_p g (fst x) %* el_prod_compose_p g (snd x))) (finite_image _ _ _ _ pf) = %0. intros Bp E g X Y h1. pose (Im X (ba_conv_fin_map_dom (def:=mns))) as X'. pose (Im Y (ba_conv_fin_map_dom (def:=mns))) as Y'. assert (h0: [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X' Y') x /\ fst x <> snd x] = Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun pr=> (ba_conv_fin_map_dom (fst pr), ba_conv_fin_map_dom (snd pr)))). apply Extensionality_Ensembles. red. split. red. intros pr h2. apply Im_intro with pr; auto. destruct h2 as [h2]. destruct h2 as [h2a h2b]. destruct h2a as [h2a]. destruct h2a as [h2a h2c]. destruct pr as [fx fy]. simpl in h2a, h2b, h2c. destruct h2a as [fx h2a]. subst. destruct h2c as [fy h2c]. subst. constructor; simpl. split. constructor. simpl. split; auto. assumption. apply surjective_pairing. red. intros f h2. destruct h2 as [pr h2]. subst. destruct h2 as [h2]. destruct h2 as [h2a h2b]. destruct h2a as [h2a]. destruct h2a as [h2a h2c]. destruct pr as [fx fy]. simpl in h2a, h2b, h2c. simpl. constructor; simpl. split. constructor; simpl. split. apply Im_intro with fx; auto. apply Im_intro with fy; auto. assumption. assert (h2:Finite [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X' Y') x /\ fst x <> snd x]). rewrite h0. apply finite_image; auto. pose proof (plus_set_el_prod_compose_disjoint (ba_conv Ap) (ba_conv_set E) (ba_conv_fun1 g) X' Y' h2) as h3. rewrite plus_set_p_eq. rewrite ba_conv_zero. rewrite <- h3. unfold ba_conv_elt. unfold ba_conv_type. rewrite transfer_eq_refl. unfold ba_conv_set. unfold ba_conv_type. apply (plus_set_functional (B:=ba_conv Ap)). rewrite transfer_dep_eq_refl. f_equal. unfold Atp, bt. unfold Atp, bt in h0. rewrite h0 at 1. pose proof (im_id [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x]) as h4. unfold Atp, bt in h4. unfold Atp, bt. rewrite h4 at 1. f_equal. apply functional_extensionality. intro x. unfold ba_conv_fin_map_dom. unfold id. apply surjective_pairing. apply functional_extensionality. intro f. destruct f as [fx fy]. simpl. rewrite el_prod_compose_p_eq. rewrite (el_prod_compose_p_eq g fy). reflexivity. Qed. Lemma plus_set_el_prod_compose_disjoint_p' : forall {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (g:(btp Bp)->Atp) (X Y:Ensemble (Fin_map E signe mns)) (pf:Finite [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x]) (pfi:Finite (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_compose_p g (fst x) %* el_prod_compose_p g (snd x)))), plus_set_p (Im [x : Fin_map E signe mns * Fin_map E signe mns | Inn (cart_prod X Y) x /\ fst x <> snd x] (fun x : Fin_map E signe mns * Fin_map E signe mns => el_prod_compose_p g (fst x) %* el_prod_compose_p g (snd x))) pfi = %0. intros Bp E g X Y h1 h2. pose proof (plus_set_el_prod_compose_disjoint_p _ g X Y h1) as h3. rewrite <- h3. apply plus_set_functional_p. reflexivity. Qed. Lemma el_prod_covers_p : forall (E:Ensemble Atp) (pf:Finite E), plus_set_p (Im (Full_set (Fin_map E signe mns)) el_prod_p) (finite_image _ _ _ _ (finite_fin_maps _ _ mns pf signe_finite)) = %1. intros E h1. pose proof (el_prod_covers _ (ba_conv_set E) h1) as h2. rewrite ba_conv_one. rewrite <- h2. rewrite plus_set_p_eq. unfold ba_conv_elt, ba_conv_set, ba_conv_type. rewrite transfer_eq_refl. apply (plus_set_functional (B:=ba_conv Ap)). rewrite transfer_dep_eq_refl. f_equal. apply functional_extensionality. intro x. rewrite el_prod_p_eq. reflexivity. Qed. Lemma el_prod_covers_p' : forall (E:Ensemble Atp) (pf:Finite E) (pf:Finite (Im (Full_set (Fin_map E signe mns)) el_prod_p)), plus_set_p (Im (Full_set (Fin_map E signe mns)) el_prod_p) pf = %1. intros E h1 h2. pose proof (el_prod_covers_p _ h1) as h3. rewrite <- h3. apply plus_set_functional_p. reflexivity. Qed. Lemma el_prod_ex_non_zero_el_prod_compose_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:btp Bp->Atp) (a:Fin_map E signe mns), el_prod_compose_p g a <> %0 -> exists a':Fin_map (Im E g) signe mns, el_prod_p a' = el_prod_compose_p g a. intros Bp E g a h1. rewrite el_prod_compose_p_eq, ba_conv_zero in h1. pose proof (el_prod_ex_non_zero_el_prod_compose (ba_conv Ap) (ba_conv_fun1 g) (ba_conv_fin_map_dom a) h1) as h2. destruct h2 as [a' h2]. assert (h3:Im E g = Im (ba_conv_set E) (ba_conv_fun1 g)). apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [x h3]. subst. apply Im_intro with x; auto. red. intros x h3. destruct h3 as [x h3]. subst. apply Im_intro with x; auto. rewrite h3. exists a'. rewrite el_prod_p_eq, el_prod_compose_p_eq. unfold ba_conv_fin_map_dom. unfold ba_conv_set, bt, ba_conv_type, Atp in h2. unfold ba_conv_set, bt, ba_conv_type, Atp. rewrite h2 at 1. reflexivity. Qed. Lemma non_zero_el_prod_compose_ex_el_prod_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)}, Finite E -> forall (g:btp Bp->Atp) (a:Fin_map (Im E g) signe mns), el_prod_p a <> %0 -> exists a':Fin_map E signe mns, el_prod_compose_p g a' = el_prod_p a. intros Bp E h1 g a h2. rewrite el_prod_p_eq, ba_conv_zero in h2. assert (h3:Included (Im (ba_conv_set E) (ba_conv_fun1 g)) (Im E g)). red. intros x h4. destruct h4 as [x h4]. subst. apply Im_intro with x; auto. pose proof (@non_zero_el_prod_compose_ex_el_prod (ba_conv Ap) (ba_conv Bp) _ h1 (ba_conv_fun g) (ba_conv_fin_map_dom (restriction a h3))) as h4. assert (h5:el_prod (ba_conv Ap) (ba_conv_fin_map_dom (restriction a h3)) <> 0). intro h5. rewrite <- h5 in h2. contradict h2. f_equal. apply fin_map_ext. intros x h6. unfold ba_conv_fin_map_dom. rewrite restriction_compat. reflexivity. destruct h6 as [x h6]. subst. apply Im_intro with x; auto. specialize (h4 h5). destruct h4 as [a' h4]. exists a'. rewrite el_prod_compose_p_eq, el_prod_p_eq. rewrite h4 at 1. unfold ba_conv_fin_map_dom. f_equal. apply fin_map_ext. intros x h6. rewrite restriction_compat; auto. Qed. Lemma el_prod_compose_covers_p : forall {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (pf:Finite E) (g:(btp Bp)->Atp), plus_set_p (Im (Full_set (Fin_map E signe mns)) (el_prod_compose_p g)) (finite_image _ _ _ _ (finite_fin_maps _ _ mns pf signe_finite)) = %1. intros Bp E h1 g. rewrite plus_set_p_eq. unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. pose proof (el_prod_compose_covers (ba_conv Ap) (ba_conv_set E) h1 (ba_conv_fun g)) as h2. rewrite ba_conv_one. rewrite <- h2. apply (plus_set_functional (B:=ba_conv Ap)). unfold ba_conv_set, ba_conv_type. rewrite transfer_dep_eq_refl. f_equal. apply functional_extensionality. intros f. rewrite el_prod_compose_p_eq. reflexivity. Qed. Definition non_zero_el_prod_maps_p (E:Ensemble Atp) := [a:Fin_map E signe mns | el_prod_p a <> %0]. Lemma non_zero_el_prod_maps_p_eq : forall (E:Ensemble Atp), non_zero_el_prod_maps_p E = non_zero_el_prod_maps (ba_conv Ap) (ba_conv_set E). intro E. unfold non_zero_el_prod_maps_p, non_zero_el_prod_maps. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [h1]. constructor. rewrite el_prod_p_eq in h1. assumption. red. intros x h1. destruct h1 as [h1]. constructor. rewrite el_prod_p_eq. assumption. Qed. Lemma non_zero_el_prod_maps_fin_p : forall (E:Ensemble Atp), Finite E -> Finite (non_zero_el_prod_maps_p E). intros E h1. assert (h2:Included (non_zero_el_prod_maps_p E) (Full_set (Fin_map E signe mns))). red. intros; constructor. pose proof (finite_fin_maps _ _ mns h1 signe_finite) as h3. apply (Finite_downward_closed _ _ h3 _ h2). Qed. Definition non_zero_el_prod_compose_maps_p {Bp:Bool_Alg_p T} (g:(btp Bp)->Atp) (E:Ensemble (btp Bp)) := [a:Fin_map E signe mns | el_prod_compose_p g a <> %0]. Lemma non_zero_el_prod_compose_maps_p_eq : forall {Bp:Bool_Alg_p T} (g:(btp Bp)->Atp) (E:Ensemble (btp Bp)), non_zero_el_prod_compose_maps_p g E = non_zero_el_prod_compose_maps (ba_conv Ap) (ba_conv_fun g) (ba_conv_set E). intros Bp g E. unfold non_zero_el_prod_compose_maps_p, non_zero_el_prod_compose_maps. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [h1]. constructor. rewrite el_prod_compose_p_eq in h1. assumption. red. intros x h1. destruct h1 as [h1]. constructor. rewrite el_prod_compose_p_eq. assumption. Qed. Lemma non_zero_el_prod_compose_maps_fin_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (g:btp Bp->Atp), Finite E -> Finite (non_zero_el_prod_compose_maps_p g E). intros B E g h1. assert (h2:Included (non_zero_el_prod_compose_maps_p g E) (Full_set (Fin_map E signe mns))). red. intros; constructor. pose proof (finite_fin_maps _ _ mns h1 signe_finite) as h3. apply (Finite_downward_closed _ _ h3 _ h2). Qed. Definition plus_subset_non_zero_el_prod_maps_p (E:Ensemble Atp) (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_maps_p E)) : Atp := plus_set_p (Im S el_prod_p) (finite_image _ _ _ el_prod_p (Finite_downward_closed _ _ (non_zero_el_prod_maps_fin_p _ pfe) _ pfi)). Lemma plus_subset_non_zero_el_prod_maps_p_eq : forall (E:Ensemble Atp) (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_maps_p E)), plus_subset_non_zero_el_prod_maps_p E pfe S pfi = plus_subset_non_zero_el_prod_maps _ (ba_conv_set E) pfe (ba_conv_ens_fin_map_dom S) (eq_rect _ _ (iff1 (incl_ba_conv_ens_fin_map_dom_iff (def:=mns) _ _ ) pfi) _ (non_zero_el_prod_maps_p_eq _ )). intros E h1 S h2. unfold plus_subset_non_zero_el_prod_maps_p, plus_subset_non_zero_el_prod_maps, ba_conv_set, ba_conv_ens_fin_map_dom. simpl. rewrite plus_set_p_eq. apply (plus_set_functional (B:=ba_conv Ap)). unfold ba_conv_set, ba_conv_type. erewrite transfer_dep_eq_refl. unfold eq_rect_r. simpl. f_equal. apply functional_extensionality. intro f. rewrite el_prod_p_eq. reflexivity. Qed. Definition plus_subset_non_zero_el_prod_compose_maps_p {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (pfe:Finite E) (g:(btp Bp)->Atp) (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_compose_maps_p g E)) : Atp := plus_set_p (Im S (el_prod_compose_p g)) (finite_image _ _ _ (el_prod_compose_p g) (finite_fin_map_ens S pfe signe_finite)). Lemma plus_subset_non_zero_el_prod_compose_maps_p_eq : forall {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (pfe:Finite E) (g:(btp Bp)->Atp) (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_compose_maps_p g E)), plus_subset_non_zero_el_prod_compose_maps_p E pfe g S pfi = plus_subset_non_zero_el_prod_compose_maps (ba_conv Ap) (ba_conv_set E) pfe (ba_conv_fun g) (ba_conv_ens_fin_map_dom S) (eq_rect _ _ (iff1 (incl_ba_conv_ens_fin_map_dom_iff (def:=mns) _ _ ) pfi) _ (non_zero_el_prod_compose_maps_p_eq _ _ )). intros Bp E h1 g S h2. unfold plus_subset_non_zero_el_prod_compose_maps_p, plus_subset_non_zero_el_prod_compose_maps, ba_conv_set, ba_conv_ens_fin_map_dom. simpl. rewrite plus_set_p_eq. apply (plus_set_functional (B:=ba_conv Ap)). unfold ba_conv_set, ba_conv_type. erewrite transfer_dep_eq_refl. unfold eq_rect_r. simpl. f_equal. apply functional_extensionality. intro f. rewrite el_prod_compose_p_eq. reflexivity. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_maps_preserves_int_p : forall (E:Ensemble Atp) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_maps_p E) -> Included Y (non_zero_el_prod_maps_p E) -> Included (Intersection X Y) (non_zero_el_prod_maps_p E). intros E h1 X Y h2 h3. red. intros a h4. destruct h4 as [a h4l h4r]. constructor. red in h3. specialize (h3 _ h4r). destruct h3. assumption. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_maps_preserves_union_p : forall (E:Ensemble Atp) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_maps_p E) -> Included Y (non_zero_el_prod_maps_p E) -> Included (Union X Y) (non_zero_el_prod_maps_p E). intros E h1 X Y h2 h3. red. intros a h4. destruct h4 as [a h4l | a h4r]. red in h2. specialize (h2 _ h4l). assumption. red in h3. specialize (h3 _ h4r). assumption. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp_p : forall (E:Ensemble Atp) (pfe:Finite E) (X:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_maps_p E) -> Included (Setminus (non_zero_el_prod_maps_p E) X) (non_zero_el_prod_maps_p E). intros E h1 X h2. red. intros a h4. destruct h4 as [h4l]. assumption. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_maps_preserves_setminus_p : forall (E:Ensemble Atp) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_maps_p E) -> Included (Setminus X Y) (non_zero_el_prod_maps_p E). intros E h1 X Y h2. red. intros a h3. destruct h3 as [h3a h3b]. auto with sets. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_maps_preserves_symdiff_full_p : forall (E:Ensemble Atp) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_maps_p E) -> Included Y (non_zero_el_prod_maps_p E) -> Included (Symdiff_full (non_zero_el_prod_maps_p E) X Y) (non_zero_el_prod_maps_p E). intros E h1 X Y h2 h3. red. intros a h4. destruct h4 as [a h4l | a h4r]. destruct h4l as [a h4a h4b]. auto with sets. destruct h4r as [a h4a h4b]. auto with sets. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_union_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (pfe:Finite E) (g:(btp Bp)->Atp) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_compose_maps_p g E) -> Included Y (non_zero_el_prod_compose_maps_p g E) -> Included (Union X Y) (non_zero_el_prod_compose_maps_p g E). intros B E h1 g X Y h2 h3. red. intros a h4. destruct h4 as [a h4l | a h4r]. red in h2. specialize (h2 _ h4l). assumption. red in h3. specialize (h3 _ h4r). assumption. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_int_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (pfe:Finite E) (g:(btp Bp)->Atp) (X Y:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_compose_maps_p g E) -> Included Y (non_zero_el_prod_compose_maps_p g E) -> Included (Intersection X Y) (non_zero_el_prod_compose_maps_p g E). intros B E h1 g X Y h2 h3. red. intros a h4. destruct h4 as [a h4l h4r]. constructor. red in h3. specialize (h3 _ h4r). destruct h3. assumption. Qed. Lemma inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_comp_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (pfe:Finite E) (g:(btp Bp)->Atp) (X:Ensemble (Fin_map E signe mns)), Included X (non_zero_el_prod_compose_maps_p g E) -> Included (Setminus (non_zero_el_prod_compose_maps_p g E) X) (non_zero_el_prod_compose_maps_p g E). intros B E h1 g X h2. red. intros a h4. destruct h4 as [h4l]. assumption. Qed. Lemma plus_subset_non_zero_el_prod_maps_empty_p : forall (E:Ensemble Atp) (pfe:Finite E), plus_subset_non_zero_el_prod_maps_p E pfe _ (incl_empty _) = %0. intros E h1. rewrite plus_subset_non_zero_el_prod_maps_p_eq. rewrite ba_conv_zero. rewrite <- (plus_subset_non_zero_el_prod_maps_empty (ba_conv Ap) _ h1). f_equal. apply proof_irrelevance. Qed. Lemma le_plus_subset_non_zero_el_prod_maps_p : forall (E:Ensemble Atp) (pfe:Finite E) (X:Ensemble (Fin_map E signe mns)) (pf:Included X (non_zero_el_prod_maps_p E)) (a:Fin_map E signe mns), Inn X a -> le_p (el_prod_p a) (plus_subset_non_zero_el_prod_maps_p E pfe _ pf). intros E h1 X h2 a h3. unfold plus_subset_non_zero_el_prod_maps_p. apply le_plus_set_p. apply Im_intro with a; auto. Qed. Lemma plus_subset_non_zero_el_prod_maps_empty_rev_p : forall (E:Ensemble Atp) (pfe:Finite E) (X:Ensemble (Fin_map E signe mns)) (pf:Included X (non_zero_el_prod_maps_p E)), plus_subset_non_zero_el_prod_maps_p E pfe _ pf = %0 -> X = Empty_set _. intros E h1 X h2 h3. assert (h5: Included (ba_conv_ens_fin_map_dom X) (non_zero_el_prod_maps (ba_conv Ap) E)). red. intros f h5. pose proof h2 as h2'. rewrite non_zero_el_prod_maps_p_eq in h2'. apply h2'. assumption. assert (h6: plus_subset_non_zero_el_prod_maps (ba_conv Ap) E h1 (ba_conv_ens_fin_map_dom X) h5 = 0). rewrite plus_subset_non_zero_el_prod_maps_p_eq in h3. rewrite ba_conv_zero in h3. rewrite <- h3. f_equal. apply proof_irrelevance. pose proof (plus_subset_non_zero_el_prod_maps_empty_rev (ba_conv Ap) _ h1 (ba_conv_ens_fin_map_dom X) h5 h6) as h4. assumption. Qed. Lemma plus_set_el_prod_zero_p : forall (E:Ensemble Atp) (pf:Finite _), plus_set_p (Im [x : Fin_map E signe mns | el_prod_p x = %0] el_prod_p) pf = %0. intros E h1. rewrite plus_set_p_eq. assert (h3: Im [x : Fin_map E signe mns | el_prod_p x = %0] el_prod_p = Im [x : Fin_map E signe mns | el_prod (ba_conv Ap) x = 0] (el_prod (ba_conv Ap))). apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. destruct h2 as [h2]. rewrite el_prod_p_eq in h2. apply Im_intro with x. constructor. assumption. rewrite el_prod_p_eq. reflexivity. red. intros x h2. destruct h2 as [x h2]. subst. destruct h2 as [h2]. apply Im_intro with x. constructor. rewrite el_prod_p_eq. assumption. rewrite el_prod_p_eq. reflexivity. assert (h2:Finite (Im [x : Fin_map E signe mns | el_prod (ba_conv Ap) x = 0] (el_prod (ba_conv Ap)))). rewrite <- h3. assumption. pose proof (plus_set_el_prod_zero _ (ba_conv_set E) h2) as h4. rewrite ba_conv_zero. rewrite <- h4. unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. apply (plus_set_functional (B:=ba_conv Ap)). unfold ba_conv_set, ba_conv_type. rewrite transfer_dep_eq_refl. f_equal. apply Extensionality_Ensembles. red. split. red. intros x h5. destruct h5 as [h5]. constructor. rewrite el_prod_p_eq in h5. assumption. red. intros x h5. destruct h5 as [h5]. constructor. rewrite el_prod_p_eq. assumption. apply functional_extensionality. intro x. rewrite el_prod_p_eq. reflexivity. Qed. Lemma plus_set_el_prod_compose_zero_p : forall {Bp:Bool_Alg_p T} (g:btp Bp->Atp) (E:Ensemble (btp Bp)) (pf:Finite _), plus_set_p (Im [x : Fin_map E signe mns | el_prod_compose_p g x = %0] (el_prod_compose_p g)) pf = %0. intros Bp g E h1. rewrite plus_set_p_eq. assert (h3: Im [x : Fin_map E signe mns | el_prod_compose_p g x = %0] (el_prod_compose_p g) = Im [x : Fin_map E signe mns | el_prod_compose (ba_conv Ap) (ba_conv_fun g) x = 0] (el_prod_compose (ba_conv Ap) (ba_conv_fun g))). apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. destruct h2 as [h2]. rewrite el_prod_compose_p_eq in h2. apply Im_intro with x. constructor. assumption. rewrite el_prod_compose_p_eq. reflexivity. red. intros x h2. destruct h2 as [x h2]. subst. destruct h2 as [h2]. apply Im_intro with x. constructor. rewrite el_prod_compose_p_eq. assumption. rewrite el_prod_compose_p_eq. reflexivity. assert (h2:Finite (Im [x : Fin_map E signe mns | el_prod_compose (ba_conv Ap) (B:=ba_conv Bp) (ba_conv_fun g) x = 0] (el_prod_compose _ (ba_conv_fun g)))). rewrite <- h3. assumption. pose proof (plus_set_el_prod_compose_zero _ (ba_conv_fun g) (ba_conv_set E) h2) as h4. rewrite ba_conv_zero. rewrite <- h4. unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. apply (plus_set_functional (B:=ba_conv Ap)). unfold ba_conv_set, ba_conv_type. rewrite transfer_dep_eq_refl. f_equal. apply Extensionality_Ensembles. red. split. red. intros x h5. destruct h5 as [h5]. constructor. rewrite el_prod_compose_p_eq in h5. assumption. red. intros x h5. destruct h5 as [h5]. constructor. rewrite el_prod_compose_p_eq. assumption. apply functional_extensionality. intro x. rewrite el_prod_compose_p_eq. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_maps_full_p : forall (E:Ensemble Atp) (pfe:Finite E), plus_subset_non_zero_el_prod_maps_p E pfe _ (inclusion_reflexive _) = %1. intros E h1. rewrite plus_subset_non_zero_el_prod_maps_p_eq. rewrite ba_conv_one. pose proof (plus_subset_non_zero_el_prod_maps_full (ba_conv Ap) E h1) as h2. rewrite <- h2 at 1. generalize (eq_rect (ba_conv_ens_fin_map_dom (non_zero_el_prod_maps_p E)) (Included (ba_conv_ens_fin_map_dom (non_zero_el_prod_maps_p E))) (iff1 (incl_ba_conv_ens_fin_map_dom_iff (non_zero_el_prod_maps_p E) (non_zero_el_prod_maps_p E)) (inclusion_reflexive (non_zero_el_prod_maps_p E))) (non_zero_el_prod_maps (ba_conv Ap) (ba_conv_set (ba_conv_set E))) (non_zero_el_prod_maps_p_eq (ba_conv_set E))). generalize (inclusion_reflexive (non_zero_el_prod_maps (ba_conv Ap) E)). rewrite non_zero_el_prod_maps_p_eq. intros h3 h4. assert (h3 = h4). apply proof_irrelevance. subst. f_equal. Qed. Lemma plus_subset_non_zero_el_prod_maps_int_p : forall (E:Ensemble Atp) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_maps_p E)) (pfy:Included Y (non_zero_el_prod_maps_p E)), (plus_subset_non_zero_el_prod_maps_p E pfe _ pfx) %* (plus_subset_non_zero_el_prod_maps_p E pfe _ pfy) = plus_subset_non_zero_el_prod_maps_p E pfe (Intersection X Y) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_int_p _ pfe X Y pfx pfy). intros E h1 X Y h2 h3. pose proof (iff1 (incl_ba_conv_ens_fin_map_dom_iff _ _) h2) as h4. pose proof (iff1 (incl_ba_conv_ens_fin_map_dom_iff _ _) h3) as h5. rewrite non_zero_el_prod_maps_p_eq in h4, h5. pose proof (plus_subset_non_zero_el_prod_maps_int (ba_conv Ap) _ h1 _ _ h4 h5) as h6. rewrite plus_subset_non_zero_el_prod_maps_p_eq. rewrite (plus_subset_non_zero_el_prod_maps_p_eq E h1 Y h3). rewrite (plus_subset_non_zero_el_prod_maps_p_eq E h1 (Intersection X Y) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_int_p E h1 X Y h2 h3)). assert (h7:inclusion_plus_subset_non_zero_el_prod_maps_preserves_int (ba_conv Ap) E h1 (ba_conv_ens_fin_map_dom X) (ba_conv_ens_fin_map_dom Y) h4 h5 = (eq_rect (ba_conv_ens_fin_map_dom (non_zero_el_prod_maps_p E)) (Included (ba_conv_ens_fin_map_dom (Intersection X Y))) (iff1 (incl_ba_conv_ens_fin_map_dom_iff (Intersection X Y) (non_zero_el_prod_maps_p E)) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_int_p E h1 X Y h2 h3)) (non_zero_el_prod_maps (ba_conv Ap) (ba_conv_set (ba_conv_set E))) (non_zero_el_prod_maps_p_eq (ba_conv_set E)))). apply proof_irrelevance. rewrite <- h7. unfold btp, bt, Atp in h6. unfold Atp, btp, bt. simpl in h6. simpl. rewrite <- h6 at 1. f_equal. f_equal. apply proof_irrelevance. f_equal. apply proof_irrelevance. Qed. Lemma plus_subset_non_zero_el_prod_maps_union_p : forall (E:Ensemble Atp) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_maps_p E)) (pfy:Included Y (non_zero_el_prod_maps_p E)), (plus_subset_non_zero_el_prod_maps_p E pfe _ pfx) %+ (plus_subset_non_zero_el_prod_maps_p E pfe _ pfy) = plus_subset_non_zero_el_prod_maps_p E pfe (Union X Y) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_union_p _ pfe X Y pfx pfy). intros E h1 X Y h2 h3. unfold plus_subset_non_zero_el_prod_maps_p. generalize (finite_image (Fin_map E signe mns) Atp (Union X Y) el_prod_p (Finite_downward_closed (Fin_map E signe mns) (non_zero_el_prod_maps_p E) (non_zero_el_prod_maps_fin_p E h1) (Union X Y) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_union_p E h1 X Y h2 h3))). rewrite <- plus_set_union_p. intro h4. apply plus_set_functional_p. rewrite im_union. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_maps_comp_p : forall (E:Ensemble Atp) (pfe:Finite E) (X:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_maps_p E)), %- (plus_subset_non_zero_el_prod_maps_p E pfe _ pfx) = plus_subset_non_zero_el_prod_maps_p E pfe _ (inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp_p _ pfe X pfx). intros E h1 X h2. pose proof (incl_ba_conv_ens_fin_map_dom_iff X (non_zero_el_prod_maps_p E)) as h3. unfold Atp in h2. pose proof h2 as h2'. rewrite h3 in h2'. rewrite non_zero_el_prod_maps_p_eq in h2'. pose proof (plus_subset_non_zero_el_prod_maps_comp (ba_conv Ap) _ h1 _ h2') as h4. rewrite ba_conv_comp. rewrite plus_subset_non_zero_el_prod_maps_p_eq. rewrite (plus_subset_non_zero_el_prod_maps_p_eq E h1). unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. assert (h5:h2' = (eq_rect (ba_conv_ens_fin_map_dom (non_zero_el_prod_maps_p E)) (Included (ba_conv_ens_fin_map_dom X)) (iff1 (incl_ba_conv_ens_fin_map_dom_iff X (non_zero_el_prod_maps_p E)) h2) (non_zero_el_prod_maps (ba_conv Ap) (ba_conv_set (ba_conv_set E))) (non_zero_el_prod_maps_p_eq (ba_conv_set E)))). apply proof_irrelevance. subst. rewrite h4 at 1. generalize (inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp (ba_conv Ap) E h1 (ba_conv_ens_fin_map_dom X) (eq_rect (ba_conv_ens_fin_map_dom (non_zero_el_prod_maps_p E)) (Included (ba_conv_ens_fin_map_dom X)) (iff1 (incl_ba_conv_ens_fin_map_dom_iff X (non_zero_el_prod_maps_p E)) h2) (non_zero_el_prod_maps (ba_conv Ap) (ba_conv_set (ba_conv_set E))) (non_zero_el_prod_maps_p_eq (ba_conv_set E)))). generalize (eq_rect (ba_conv_ens_fin_map_dom (non_zero_el_prod_maps_p E)) (Included (ba_conv_ens_fin_map_dom (Setminus (non_zero_el_prod_maps_p E) X))) (iff1 (incl_ba_conv_ens_fin_map_dom_iff (Setminus (non_zero_el_prod_maps_p E) X) (non_zero_el_prod_maps_p E)) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp_p E h1 X h2)) (non_zero_el_prod_maps (ba_conv Ap) (ba_conv_set (ba_conv_set E))) (non_zero_el_prod_maps_p_eq (ba_conv_set E))). rewrite non_zero_el_prod_maps_p_eq. intros h5 h6. f_equal. apply proof_irrelevance. Qed. Lemma plus_subset_non_zero_el_prod_maps_setminus_p : forall (E:Ensemble Atp) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_maps_p E)) (pfy:Included Y (non_zero_el_prod_maps_p E)), (plus_subset_non_zero_el_prod_maps_p E pfe _ pfx) %* %- (plus_subset_non_zero_el_prod_maps_p E pfe _ pfy) = plus_subset_non_zero_el_prod_maps_p E pfe (Intersection X (Setminus (non_zero_el_prod_maps_p E) Y)) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_int_p _ pfe X _ pfx (inclusion_plus_subset_non_zero_el_prod_maps_preserves_comp_p _ pfe Y pfy)). intros E h1 X Y h2 h3. rewrite <- plus_subset_non_zero_el_prod_maps_int_p. rewrite <- plus_subset_non_zero_el_prod_maps_comp_p. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_maps_symdiff_full_p : forall (E:Ensemble Atp) (pfe:Finite E) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_maps_p E)) (pfy:Included Y (non_zero_el_prod_maps_p E)), (plus_subset_non_zero_el_prod_maps_p E pfe _ pfx) /%\ (plus_subset_non_zero_el_prod_maps_p E pfe _ pfy) = plus_subset_non_zero_el_prod_maps_p E pfe (Symdiff_full (non_zero_el_prod_maps_p E) X Y) (inclusion_plus_subset_non_zero_el_prod_maps_preserves_symdiff_full_p _ pfe X Y pfx pfy). intros E h1 X Y h2 h3. unfold sym_diff_p. do 2 rewrite plus_subset_non_zero_el_prod_maps_comp_p. do 2 rewrite plus_subset_non_zero_el_prod_maps_int_p. rewrite plus_subset_non_zero_el_prod_maps_union_p. unfold Symdiff_full. unfold Setminus_full. unfold plus_subset_non_zero_el_prod_maps. apply plus_set_functional_p. reflexivity. Qed. Definition non_zero_el_prod_maps_of_set_p {E:Ensemble Atp} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)) := Intersection S (non_zero_el_prod_maps_p E). Lemma non_zero_el_prod_maps_of_set_p_eq : forall {E:Ensemble Atp} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)), non_zero_el_prod_maps_of_set_p pfe S = non_zero_el_prod_maps_of_set (ba_conv Ap) pfe (ba_conv_ens_fin_map_dom S). intros E h1 S. unfold non_zero_el_prod_maps_of_set_p, non_zero_el_prod_maps_of_set. rewrite non_zero_el_prod_maps_p_eq. reflexivity. Qed. Lemma non_zero_el_prod_maps_of_set_eq_p : forall {E:Ensemble Atp} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)), non_zero_el_prod_maps_of_set_p pfe S = [f:Fin_map E signe mns | Inn S f /\ el_prod_p f <> %0]. intros E h1 S. pose proof (non_zero_el_prod_maps_of_set_eq (ba_conv Ap) h1 (ba_conv_ens_fin_map_dom S)) as h2. rewrite non_zero_el_prod_maps_of_set_p_eq. rewrite h2 at 1. apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [h3]. destruct h3 as [h3 h4]. constructor. rewrite ba_conv_zero. rewrite el_prod_p_eq. split; auto. red. intros x h3. destruct h3 as [h3]. destruct h3 as [h3 h4]. constructor. rewrite ba_conv_zero in h4. rewrite el_prod_p_eq in h4. split; auto. Qed. Lemma non_zero_el_prod_maps_of_set_decompose_p : forall {E:Ensemble Atp} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)), S = Union (non_zero_el_prod_maps_of_set_p pfe S) [f:Fin_map E signe mns | Inn S f /\ el_prod_p f = %0]. intros E h1 S. apply Extensionality_Ensembles. red. split. red. intros f h2. destruct (classic (el_prod_p f = %0)) as [h3 | h4]. right. constructor. split; auto. left. constructor; auto. constructor; auto. red. intros f h2. destruct h2 as [f h2 | f h3]. destruct h2; auto. destruct h3 as [h3]. destruct h3; auto. Qed. Lemma incl_non_zero_el_prod_maps_of_set_p : forall {E:Ensemble Atp} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)), Included (non_zero_el_prod_maps_of_set_p pfe S) (non_zero_el_prod_maps_p E). intros E h1 S. unfold non_zero_el_prod_maps_of_set_p. auto with sets. Qed. Lemma finite_non_zero_el_prod_maps_of_set_p : forall {E:Ensemble Atp} (pfe:Finite E) (S:Ensemble (Fin_map E signe mns)), Finite (non_zero_el_prod_maps_of_set_p pfe S). intros E h1 S. eapply Finite_downward_closed. apply (non_zero_el_prod_maps_fin_p _ h1). apply incl_non_zero_el_prod_maps_of_set_p. Qed. Definition non_zero_el_prod_compose_maps_of_set_p {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (pfe:Finite E) (g:btp Bp->Atp) (S:Ensemble (Fin_map E signe mns)) := Intersection S (non_zero_el_prod_compose_maps_p g E). Lemma non_zero_el_prod_compose_maps_of_set_p_eq : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (pfe:Finite E) (g:btp Bp->Atp) (S:Ensemble (Fin_map E signe mns)), non_zero_el_prod_compose_maps_of_set_p pfe g S = non_zero_el_prod_compose_maps_of_set (ba_conv Ap) (E:=ba_conv_set E) pfe (ba_conv_fun g) (ba_conv_ens_fin_map_dom S). intros Bp E h1 g S. unfold non_zero_el_prod_compose_maps_of_set_p, non_zero_el_prod_compose_maps_of_set. rewrite non_zero_el_prod_compose_maps_p_eq. reflexivity. Qed. Lemma non_zero_el_prod_compose_maps_of_set_eq_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (pfe:Finite E) (g:btp Bp->Atp) (S:Ensemble (Fin_map E signe mns)), non_zero_el_prod_compose_maps_of_set_p pfe g S = [f:Fin_map E signe mns | Inn S f /\ el_prod_compose_p g f <> %0]. intros B E h1 g S. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [f h2 h3]. destruct h3 as [h3]. constructor. split; auto. red. intros x h2. destruct h2 as [h2]. destruct h2 as [h2a h2b]. constructor; auto. constructor. assumption. Qed. Lemma non_zero_el_prod_compose_maps_of_set_decompose_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (pfe:Finite E) (g:btp Bp->Atp) (S:Ensemble (Fin_map E signe mns)), S = Union (non_zero_el_prod_compose_maps_of_set_p pfe g S) [f:Fin_map E signe mns | Inn S f /\ el_prod_compose_p g f = %0]. intros B E h1 g S. apply Extensionality_Ensembles. red. split. red. intros f h2. destruct (classic (el_prod_compose_p g f = %0)) as [h3 | h4]. right. constructor. split; auto. left. constructor; auto. constructor; auto. red. intros f h2. destruct h2 as [f h2 | f h3]. destruct h2; auto. destruct h3 as [h3]. destruct h3; auto. Qed. Lemma incl_non_zero_el_prod_compose_maps_of_set_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (pfe:Finite E) (g:btp Bp->Atp) (S:Ensemble (Fin_map E signe mns)), Included (non_zero_el_prod_compose_maps_of_set_p pfe g S) (non_zero_el_prod_compose_maps_p g E). intros B E h1 g S. unfold non_zero_el_prod_compose_maps_of_set_p. auto with sets. Qed. Lemma finite_non_zero_el_prod_compose_maps_of_set_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (pfe:Finite E) (g:btp Bp->Atp) (S:Ensemble (Fin_map E signe mns)), Finite (non_zero_el_prod_compose_maps_of_set_p pfe g S). intros B E h1 g S. eapply Finite_downward_closed. apply (non_zero_el_prod_compose_maps_fin_p g h1). apply incl_non_zero_el_prod_compose_maps_of_set_p. Qed. Lemma plus_set_zero_el_prod_maps_of_set_p : forall {E:Ensemble Atp}, Finite E -> forall (S:Ensemble (Fin_map E signe mns)) (pf:Finite (Im [f : Fin_map E signe mns | Inn S f /\ el_prod_p f = %0] el_prod_p)), plus_set_p _ pf = %0. intros E h1 S h0. rewrite plus_set_p_eq. pose proof h0 as h2'. assert (h3: Im [x : Fin_map E signe mns | Inn S x /\ el_prod_p x = %0] el_prod_p = Im [x : Fin_map E signe mns | Inn S x /\ el_prod (ba_conv Ap) x = 0] (el_prod (ba_conv Ap))). apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. destruct h2 as [h2]. rewrite el_prod_p_eq in h2. apply Im_intro with x. constructor. assumption. rewrite el_prod_p_eq. reflexivity. red. intros x h2. destruct h2 as [x h2]. subst. destruct h2 as [h2]. apply Im_intro with x. constructor. rewrite el_prod_p_eq. assumption. rewrite el_prod_p_eq. reflexivity. rewrite h3 in h2'. pose proof (plus_set_zero_el_prod_maps_of_set (ba_conv Ap) h1 S h2') as h4. rewrite ba_conv_zero. rewrite <- h4. unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. apply (plus_set_functional (B:=ba_conv (Ap))). unfold ba_conv_set, ba_conv_type. rewrite transfer_dep_eq_refl. f_equal. apply Extensionality_Ensembles. red. split. red. intros x h5. destruct h5 as [h5]. constructor. rewrite el_prod_p_eq in h5. assumption. red. intros x h5. destruct h5 as [h5]. constructor. rewrite el_prod_p_eq. rewrite ba_conv_zero. assumption. apply functional_extensionality. intro x. rewrite el_prod_p_eq; auto. Qed. Lemma plus_subset_el_prod_compose_maps_eq_same_non_zero_p : forall {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (pfe:Finite E) (g:btp Bp->Atp) (S:Ensemble (Fin_map E signe mns)), plus_set_p (Im S (el_prod_compose_p g)) (finite_image _ _ _ (el_prod_compose_p g) (finite_fin_map_ens S pfe signe_finite)) = plus_set_p (Im (non_zero_el_prod_compose_maps_of_set_p pfe g S) (el_prod_compose_p g)) (finite_image _ _ _ (el_prod_compose_p g) (finite_non_zero_el_prod_compose_maps_of_set_p pfe g S)). intros Bp E h1 g S. rewrite plus_set_p_eq. rewrite plus_set_p_eq. f_equal. pose proof (plus_subset_el_prod_compose_maps_eq_same_non_zero (ba_conv Ap) (B:=(ba_conv Bp)) (E:=ba_conv_set E) h1 (ba_conv_fun g) (ba_conv_ens_fin_map_dom S)) as h2. assert (h4:Im S (el_prod_compose_p g) = Im (ba_conv_ens_fin_map_dom S) (el_prod_compose (ba_conv Ap) (ba_conv_fun g))). f_equal. apply functional_extensionality. intro x. rewrite el_prod_compose_p_eq. reflexivity. generalize (finite_image (Fin_map E signe mns) Atp S (el_prod_compose_p g) (finite_fin_map_ens S h1 signe_finite)). rewrite h4. intro h5. assert (h5 = finite_image (Fin_map (ba_conv_set E) signe mns) (Btype (Bc (ba_conv Ap))) (ba_conv_ens_fin_map_dom S) (el_prod_compose (ba_conv Ap) (ba_conv_fun g)) (finite_fin_map_ens (ba_conv_ens_fin_map_dom S) h1 signe_finite)). apply proof_irrelevance. subst. rewrite h2 at 1. apply plus_set_functional. unfold ba_conv_set, ba_conv_type. rewrite transfer_dep_eq_refl. f_equal. rewrite non_zero_el_prod_compose_maps_of_set_p_eq. reflexivity. apply functional_extensionality. intro f. rewrite el_prod_compose_p_eq. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_empty_p : forall {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (pfe:Finite E) (g:(btp Bp)->Atp), plus_subset_non_zero_el_prod_compose_maps_p E pfe g _ (incl_empty _) = %0. intros. unfold plus_subset_non_zero_el_prod_compose_maps_p. generalize ((finite_image (Fin_map E signe mns) Atp (Empty_set (Fin_map E signe mns)) (el_prod_compose_p g) (finite_fin_map_ens (Empty_set (Fin_map E signe mns)) pfe signe_finite))). pose proof (image_empty (Fin_map E signe mns) _ (el_prod_compose_p g)) as h1. rewrite h1. intro. apply plus_set_empty_p'. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_full_p : forall {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (pfe:Finite E) (g:(btp Bp)->Atp), plus_subset_non_zero_el_prod_compose_maps_p E pfe g _ (inclusion_reflexive _) = %1. intros Bp E h1 g. pose proof (plus_subset_non_zero_el_prod_compose_maps_full (ba_conv Ap) (ba_conv_set E) h1 (ba_conv_fun g)) as h2. rewrite ba_conv_one. rewrite <- h2. rewrite plus_subset_non_zero_el_prod_compose_maps_p_eq. generalize (eq_rect (ba_conv_ens_fin_map_dom (non_zero_el_prod_compose_maps_p g E)) (Included (ba_conv_ens_fin_map_dom (non_zero_el_prod_compose_maps_p g E))) (iff1 (incl_ba_conv_ens_fin_map_dom_iff (non_zero_el_prod_compose_maps_p g E) (non_zero_el_prod_compose_maps_p g E)) (inclusion_reflexive (non_zero_el_prod_compose_maps_p g E))) (non_zero_el_prod_compose_maps (ba_conv Ap) (ba_conv_fun g) (ba_conv_set (ba_conv_set E))) (non_zero_el_prod_compose_maps_p_eq g (ba_conv_set E))). rewrite non_zero_el_prod_compose_maps_p_eq. intro h3. f_equal. apply proof_irrelevance. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_int_p : forall {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (pfe:Finite E) (g:(btp Bp)->Atp) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_compose_maps_p g E)) (pfy:Included Y (non_zero_el_prod_compose_maps_p g E)), plus_subset_non_zero_el_prod_compose_maps_p E pfe g (Intersection X Y) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_int_p pfe g X Y pfx pfy) = plus_subset_non_zero_el_prod_compose_maps_p E pfe g X pfx %* plus_subset_non_zero_el_prod_compose_maps_p E pfe g Y pfy. intros Bp E h1 g X Y h2 h3. pose proof h2 as h2'. pose proof h3 as h3'. apply (iff1 (incl_ba_conv_ens_fin_map_dom_iff X (non_zero_el_prod_compose_maps_p g E))) in h2'. apply (iff1 (incl_ba_conv_ens_fin_map_dom_iff Y (non_zero_el_prod_compose_maps_p g E))) in h3'. rewrite non_zero_el_prod_compose_maps_p_eq in h2', h3'. pose proof (plus_subset_non_zero_el_prod_compose_maps_int (ba_conv Ap) (ba_conv_set E) h1 (ba_conv_fun g) _ _ h2' h3') as h4. do 2 rewrite plus_subset_non_zero_el_prod_compose_maps_p_eq. generalize (eq_rect (ba_conv_ens_fin_map_dom (non_zero_el_prod_compose_maps_p g E)) (Included (ba_conv_ens_fin_map_dom (Intersection X Y))) (iff1 (incl_ba_conv_ens_fin_map_dom_iff (Intersection X Y) (non_zero_el_prod_compose_maps_p g E)) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_int_p h1 g X Y h2 h3)) (non_zero_el_prod_compose_maps (ba_conv Ap) (ba_conv_fun g) (ba_conv_set (ba_conv_set E))) (non_zero_el_prod_compose_maps_p_eq g (ba_conv_set E))). intro h5. assert (h5= inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_int (ba_conv Ap) h1 (ba_conv_fun g) (ba_conv_ens_fin_map_dom X) (ba_conv_ens_fin_map_dom Y) h2' h3' (B:=ba_conv Bp)). apply proof_irrelevance. subst. rewrite h4 at 1. rewrite ba_conv_times. f_equal. rewrite plus_subset_non_zero_el_prod_compose_maps_p_eq. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_union_p : forall {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (pfe:Finite E) (g:(btp Bp)->Atp) (X Y:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_compose_maps_p g E)) (pfy:Included Y (non_zero_el_prod_compose_maps_p g E)), plus_subset_non_zero_el_prod_compose_maps_p E pfe g (Union X Y) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_union_p pfe g X Y pfx pfy) = plus_subset_non_zero_el_prod_compose_maps_p E pfe g X pfx %+ plus_subset_non_zero_el_prod_compose_maps_p E pfe g Y pfy. intros Bp E h1 g X Y h2 h3. pose proof h2 as h2'. pose proof h3 as h3'. apply (iff1 (incl_ba_conv_ens_fin_map_dom_iff X (non_zero_el_prod_compose_maps_p g E))) in h2'. apply (iff1 (incl_ba_conv_ens_fin_map_dom_iff Y (non_zero_el_prod_compose_maps_p g E))) in h3'. rewrite non_zero_el_prod_compose_maps_p_eq in h2', h3'. pose proof (plus_subset_non_zero_el_prod_compose_maps_union (ba_conv Ap) (ba_conv_set E) h1 (ba_conv_fun g) _ _ h2' h3') as h4. do 2 rewrite plus_subset_non_zero_el_prod_compose_maps_p_eq. generalize (eq_rect (ba_conv_ens_fin_map_dom (non_zero_el_prod_compose_maps_p g E)) (Included (ba_conv_ens_fin_map_dom (Union X Y))) (iff1 (incl_ba_conv_ens_fin_map_dom_iff (Union X Y) (non_zero_el_prod_compose_maps_p g E)) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_union_p h1 g X Y h2 h3)) (non_zero_el_prod_compose_maps (ba_conv Ap) (ba_conv_fun g) (ba_conv_set (ba_conv_set E))) (non_zero_el_prod_compose_maps_p_eq g (ba_conv_set E))). intro h5. assert (h5= inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_union (ba_conv Ap) h1 (ba_conv_fun g) (ba_conv_ens_fin_map_dom X) (ba_conv_ens_fin_map_dom Y) h2' h3' (B:=ba_conv Bp)). apply proof_irrelevance. subst. rewrite h4 at 1. rewrite ba_conv_plus. f_equal. rewrite plus_subset_non_zero_el_prod_compose_maps_p_eq. reflexivity. Qed. Lemma plus_subset_non_zero_el_prod_compose_maps_comp_p : forall {Bp:Bool_Alg_p T} (E:Ensemble (btp Bp)) (pfe:Finite E) (g:(btp Bp)->Atp) (X:Ensemble (Fin_map E signe mns)) (pfx:Included X (non_zero_el_prod_compose_maps_p g E)), plus_subset_non_zero_el_prod_compose_maps_p E pfe g (Setminus (non_zero_el_prod_compose_maps_p g E) X) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_comp_p pfe g X pfx) = %-plus_subset_non_zero_el_prod_compose_maps_p E pfe g X pfx. intros Bp E h1 g X h2. pose proof h2 as h2'. apply (iff1 (incl_ba_conv_ens_fin_map_dom_iff X (non_zero_el_prod_compose_maps_p g E))) in h2'. rewrite non_zero_el_prod_compose_maps_p_eq in h2'. pose proof (plus_subset_non_zero_el_prod_compose_maps_comp (ba_conv Ap) (ba_conv_set E) h1 (ba_conv_fun g) _ h2') as h4. do 2 rewrite plus_subset_non_zero_el_prod_compose_maps_p_eq. generalize (eq_rect (ba_conv_ens_fin_map_dom (non_zero_el_prod_compose_maps_p g E)) (Included (ba_conv_ens_fin_map_dom (Setminus (non_zero_el_prod_compose_maps_p g E) X))) (iff1 (incl_ba_conv_ens_fin_map_dom_iff (Setminus (non_zero_el_prod_compose_maps_p g E) X) (non_zero_el_prod_compose_maps_p g E)) (inclusion_plus_subset_non_zero_el_prod_compose_maps_preserves_comp_p h1 g X h2)) (non_zero_el_prod_compose_maps (ba_conv Ap) (ba_conv_fun g) (ba_conv_set (ba_conv_set E))) (non_zero_el_prod_compose_maps_p_eq g (ba_conv_set E))). intro h5. pose proof (non_zero_el_prod_compose_maps_p_eq g E) as h6. simpl in h6, h5. pose proof h5 as h5'. rewrite h6 in h5'. pose proof (f_equal (fun S => (ba_conv_ens_fin_map_dom (Setminus S X))) h6) as h7. simpl in h7. pose proof (subsetT_eq_compat _ (fun S=>Included S (non_zero_el_prod_compose_maps (ba_conv Ap) (ba_conv_fun g) (ba_conv_set (ba_conv_set E)))) _ _ h5 h5' h7) as h8. dependent rewrite -> h8. rewrite h4 at 1. rewrite ba_conv_comp. f_equal. Qed. Lemma incl_im_proj1_sig_ex_ens_fin_map_plus_subset_non_zero_el_prod_maps_p : forall (E:Ensemble Atp) (pfe:Finite E), Included (im_proj1_sig [a : Atp | exists (S : Ensemble (Fin_map E signe mns)) (pfi : Included S (non_zero_el_prod_maps_p E)), a = plus_subset_non_zero_el_prod_maps_p E pfe S pfi]) (A_p T (Bc_p T Ap)). intros E h0. red. intros x h2. destruct h2 as [x h2]. subst. apply proj2_sig. Qed. Lemma closed_all_plus_subsets_non_zero_el_prod_maps_p : forall (E:Ensemble Atp) (pfe:Finite E), alg_closed_p (im_proj1_sig [a:Atp | exists (S:Ensemble (Fin_map E signe mns)) (pfi:Included S (non_zero_el_prod_maps_p E)), a = plus_subset_non_zero_el_prod_maps_p E pfe S pfi]) (incl_im_proj1_sig_ex_ens_fin_map_plus_subset_non_zero_el_prod_maps_p _ pfe). intros E h1. assert (h3: [a : Atp | exists (S : Ensemble (Fin_map E signe mns)) (pfi : Included S (non_zero_el_prod_maps_p E)), a = plus_subset_non_zero_el_prod_maps_p E h1 S pfi] = [a : (bt (ba_conv Ap)) | exists (S : Ensemble (Fin_map E signe mns)) (pfi : Included S (non_zero_el_prod_maps (ba_conv Ap) E)), a = plus_subset_non_zero_el_prod_maps (ba_conv Ap) E h1 S pfi]). apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [h3]. destruct h3 as [S [h3 h4]]. subst. constructor. pose proof h3 as h3'. rewrite non_zero_el_prod_maps_p_eq in h3'. exists S, h3'. rewrite plus_subset_non_zero_el_prod_maps_p_eq. f_equal. apply proof_irrelevance. red. intros x h3. destruct h3 as [h3]. destruct h3 as [S [h3 h4]]. subst. constructor. pose proof h3 as h3'. rewrite <- (non_zero_el_prod_maps_p_eq E) in h3' at 1. exists S, h3'. rewrite plus_subset_non_zero_el_prod_maps_p_eq. f_equal. apply proof_irrelevance. rewrite alg_closed_p_iff. pose proof (f_equal (fun S => (im_proj1_sig S)) h3) as h4. simpl in h4. assert (h2':Included (im_proj1_sig [a : bt (ba_conv Ap) | exists (S : Ensemble (Fin_map E signe mns)) (pfi : Included S (non_zero_el_prod_maps (ba_conv Ap) E)), a = plus_subset_non_zero_el_prod_maps (ba_conv Ap) E h1 S pfi]) (A_p T (Bc_p T Ap))). rewrite <- h4. apply (incl_im_proj1_sig_ex_ens_fin_map_plus_subset_non_zero_el_prod_maps_p E h1). pose proof (subsetT_eq_compat _ (fun S => Included S (A_p T (Bc_p T Ap))) _ _ (incl_im_proj1_sig_ex_ens_fin_map_plus_subset_non_zero_el_prod_maps_p E h1) h2' h4) as h5. dependent rewrite -> h5. rewrite ba_conv_und_subalg_im_proj1_sig. apply closed_all_plus_subsets_non_zero_el_prod_maps. Qed. Lemma el_prod_sing_inc_p : forall (E:Ensemble Atp) (pf:Finite E) (a:Fin_map E signe mns), el_prod_p a <> %0 -> Included (Singleton a) (non_zero_el_prod_maps_p E). intros E h1 a h2. red. intros a' h3. destruct h3; subst. unfold non_zero_el_prod_maps_p. constructor. assumption. Qed. Lemma el_prod_sing_p : forall (E:Ensemble Atp) (pf:Finite E) (a:Fin_map E signe mns) (pfel:el_prod_p a <> %0), el_prod_p a = plus_subset_non_zero_el_prod_maps_p _ pf (Singleton a) (el_prod_sing_inc_p _ pf _ pfel). intros E h1 a h2. unfold plus_subset_non_zero_el_prod_maps_p. assert (h3:Im (Singleton a) el_prod_p = (Singleton (el_prod_p a))). apply Extensionality_Ensembles. red. split. red. intros x' h13. inversion h13. subst. destruct H. subst. constructor. red. intros x' h13. destruct h13. apply Im_intro with a. constructor. reflexivity. symmetry. unfold Atp, btp in h3. unfold Atp, btp. gen0. rewrite h3. intro h4. rewrite plus_set_sing_p'. reflexivity. Qed. Lemma el_prod_compose_p1_transfer_fun_sig_compat : forall {B:Bool_Alg} {E E':Ensemble Atp} (a:Fin_map E signe mns) (pf:E = E') (g:sig_set E -> bt B), el_prod_compose_p1 (sig_fun_app (transfer_fun_sig pf g) 0) a = el_prod_compose_p1 (sig_fun_app g 0) a. intros B E E' a h1 g. subst. rewrite transfer_fun_sig_eq_refl. reflexivity. Qed. Theorem gen_ens_eq_p : forall (E:Ensemble Atp) (pf:Included (im_proj1_sig E) (ba_p_ens Ap)), Gen_Ens_p (im_proj1_sig E) pf = [x : T | exists (S:Ensemble Atp) (pfs:Finite S), x = (proj1_sig (plus_set_p S pfs)) /\ forall s:btp Ap, Inn S s -> exists (R:Ensemble Atp) (pfr:Finite R), Included R (comp_add_set_p E) /\ s = times_set_p R pfr]. intros E h1. rewrite gen_ens_p_eq. pose proof (im_proj2_sig_undoes_im_proj1_sig' _ _ h1) as h2. rewrite <- h2 at 1. clear h2. pose proof (gen_ens_eq (ba_conv Ap) E) as h2. rewrite h2 at 1. clear h2. apply Extensionality_Ensembles. red; split; red; intros x h2. destruct h2 as [x' h2]. subst. destruct h2 as [h2]. destruct h2 as [S [h2 [h3 h4]]]. subst. constructor. exists S. exists h2. split. rewrite plus_set_p_eq. reflexivity. intros s h5. specialize (h4 _ h5). destruct h4 as [R [h4 [h6 h7]]]. subst. exists R. exists h4. split; auto. rewrite times_set_p_eq. reflexivity. destruct h2 as [h2]. destruct h2 as [S [h2 [h3 h4]]]. subst. apply Im_intro with (plus_set_p S h2); auto. constructor. exists S. exists h2. split. rewrite plus_set_p_eq. reflexivity. intros s h5. specialize (h4 _ h5). destruct h4 as [R [h4 [h6 h7]]]. subst. exists R. exists h4. split; auto. rewrite times_set_p_eq. reflexivity. Qed. Lemma incl_im_full_set_finc_ba_p_ens : forall (E:Ensemble Atp), Included (im_proj1_sig E) (ba_p_ens Ap) -> forall (pfinc:forall F : finc E, Included (im_proj1_sig (proj1_sig F)) (ba_p_ens Ap)), Included (FamilyUnion (Im (Full_set (finc E)) (fun x : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x)) (pfinc x))))) (ba_p_ens Ap). intros E h1 h2 . red. intros x h11. destruct h11 as [S x h11 h12]. destruct h11 as [S h11]. subst. clear h11. destruct h12 as [x h12]. apply h12. constructor. split. apply h2. exists (inclusion_reflexive _). apply alg_closed_p_ba_p_ens_refl. Qed. Lemma closed_incl_im_full_set_finc_ba_p_ens : forall (E:Ensemble Atp) (pf:Included (im_proj1_sig E) (ba_p_ens Ap)) (pfinc:forall F : finc E, Included (im_proj1_sig (proj1_sig F)) (ba_p_ens Ap)), alg_closed_p (FamilyUnion (Im (Full_set (finc E)) (fun x : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x)) (pfinc x))))) (incl_im_full_set_finc_ba_p_ens E pf pfinc). intros E h1 h2. constructor; simpl. red. intros x y. destruct x as [x h3], y as [y h4]. simpl. destruct h3 as [C x h3], h4 as [D y h4]. simpl. destruct h3 as [C h3], h4 as [D h4]; subst. apply family_union_intro with (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig (finc_union _ C D))) (h2 (finc_union _ C D)))). apply Im_intro with (finc_union _ C D). constructor. reflexivity. simpl. assert (h5:Included (im_proj1_sig (proj1_sig C)) (ba_p_ens Ap)). red. intros a h5. destruct h5 as [a h5]. subst. apply proj2_sig. assert (h6:Included (im_proj1_sig (proj1_sig D)) (ba_p_ens Ap)). red. intros a h6. destruct h6 as [a h6]. subst. apply proj2_sig. pose proof (finc_union_incl_l _ C D) as h7. pose proof (finc_union_incl_r _ C D) as h8. assert (h9:Included (im_proj1_sig (proj1_sig (finc_union Atp C D))) (ba_p_ens Ap)). red. intros a ha. destruct ha as [a ha]. subst. apply proj2_sig. unfold Atp, btp, Btype_p in h7, h8. pose proof (incl_im_proj1_sig_iff (proj1_sig C) (proj1_sig (finc_union Atp C D))) as hiffc. pose proof (incl_im_proj1_sig_iff (proj1_sig D) (proj1_sig (finc_union Atp C D))) as hiffd. unfold Atp in hiffc, hiffd. rewrite <- hiffc in h7. rewrite <- hiffd in h8. pose proof (closed_gen_ens_p _ (h2 (finc_union _ C D))) as h10. pose proof (gen_ens_preserves_inclusion_p _ _ h5 h9 h7) as h11. pose proof (gen_ens_preserves_inclusion_p _ _ h6 h9 h8) as h12. assert (h13:Inn (Gen_Ens_p (im_proj1_sig (proj1_sig (finc_union Atp C D))) (h2 (finc_union Atp C D))) x). apply h11; auto. assert (h14:Inn (Gen_Ens_p (im_proj1_sig (proj1_sig (finc_union Atp C D))) (h2 (finc_union Atp C D))) y). apply h12; auto. rewrite ba_p_ens_gen_p_eq. pose proof (plus_closed_p' _ _ _ _ h10 _ _ h13 h14) as h15. assert (h16: proj1_sig (exist (Inn (A_p T (Bc_p T Ap))) x (incl_im_full_set_finc_ba_p_ens E h1 h2 x (family_union_intro T (Im (Full_set (finc E)) (fun x0 : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x0)) (h2 x0)))) (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig C)) (h2 C))) x (Im_intro (finc E) (Ensemble T) (Full_set (finc E)) (fun x0 : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x0)) (h2 x0))) C h3 (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig C)) (h2 C))) eq_refl) i)) %+ exist (Inn (A_p T (Bc_p T Ap))) y (incl_im_full_set_finc_ba_p_ens E h1 h2 y (family_union_intro T (Im (Full_set (finc E)) (fun x0 : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x0)) (h2 x0)))) (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig D)) (h2 D))) y (Im_intro (finc E) (Ensemble T) (Full_set (finc E)) (fun x0 : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x0)) (h2 x0))) D h4 (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig D)) (h2 D))) eq_refl) i0))) = (proj1_sig (exist (Inn (ba_p_ens Ap)) x (incl_gen_ens_p (im_proj1_sig (proj1_sig (finc_union Atp C D))) (h2 (finc_union Atp C D)) x h13) %+ exist (Inn (ba_p_ens Ap)) y (incl_gen_ens_p (im_proj1_sig (proj1_sig (finc_union Atp C D))) (h2 (finc_union Atp C D)) y h14)))). f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. rewrite h16 at 1. clear h16. assumption. red. intros x y. destruct x as [x h3], y as [y h4]. simpl. destruct h3 as [C x h3], h4 as [D y h4]. simpl. destruct h3 as [C h3], h4 as [D h4]; subst. apply family_union_intro with (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig (finc_union _ C D))) (h2 (finc_union _ C D)))). apply Im_intro with (finc_union _ C D). constructor. reflexivity. simpl. assert (h5:Included (im_proj1_sig (proj1_sig C)) (ba_p_ens Ap)). red. intros a h5. destruct h5 as [a h5]. subst. apply proj2_sig. assert (h6:Included (im_proj1_sig (proj1_sig D)) (ba_p_ens Ap)). red. intros a h6. destruct h6 as [a h6]. subst. apply proj2_sig. pose proof (finc_union_incl_l _ C D) as h7. pose proof (finc_union_incl_r _ C D) as h8. assert (h9:Included (im_proj1_sig (proj1_sig (finc_union Atp C D))) (ba_p_ens Ap)). red. intros a ha. destruct ha as [a ha]. subst. apply proj2_sig. unfold Atp, btp, Btype_p in h7, h8. pose proof (incl_im_proj1_sig_iff (proj1_sig C) (proj1_sig (finc_union Atp C D))) as hiffc. pose proof (incl_im_proj1_sig_iff (proj1_sig D) (proj1_sig (finc_union Atp C D))) as hiffd. unfold Atp in hiffc, hiffd. rewrite <- hiffc in h7. rewrite <- hiffd in h8. pose proof (closed_gen_ens_p _ (h2 (finc_union _ C D))) as h10. pose proof (gen_ens_preserves_inclusion_p _ _ h5 h9 h7) as h11. pose proof (gen_ens_preserves_inclusion_p _ _ h6 h9 h8) as h12. assert (h13:Inn (Gen_Ens_p (im_proj1_sig (proj1_sig (finc_union Atp C D))) (h2 (finc_union Atp C D))) x). apply h11; auto. assert (h14:Inn (Gen_Ens_p (im_proj1_sig (proj1_sig (finc_union Atp C D))) (h2 (finc_union Atp C D))) y). apply h12; auto. rewrite ba_p_ens_gen_p_eq. pose proof (times_closed_p' _ _ _ _ h10 _ _ h13 h14) as h15. assert (h16: proj1_sig (exist (Inn (A_p T (Bc_p T Ap))) x (incl_im_full_set_finc_ba_p_ens E h1 h2 x (family_union_intro T (Im (Full_set (finc E)) (fun x0 : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x0)) (h2 x0)))) (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig C)) (h2 C))) x (Im_intro (finc E) (Ensemble T) (Full_set (finc E)) (fun x0 : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x0)) (h2 x0))) C h3 (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig C)) (h2 C))) eq_refl) i)) %* exist (Inn (A_p T (Bc_p T Ap))) y (incl_im_full_set_finc_ba_p_ens E h1 h2 y (family_union_intro T (Im (Full_set (finc E)) (fun x0 : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x0)) (h2 x0)))) (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig D)) (h2 D))) y (Im_intro (finc E) (Ensemble T) (Full_set (finc E)) (fun x0 : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x0)) (h2 x0))) D h4 (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig D)) (h2 D))) eq_refl) i0))) = (proj1_sig (exist (Inn (ba_p_ens Ap)) x (incl_gen_ens_p (im_proj1_sig (proj1_sig (finc_union Atp C D))) (h2 (finc_union Atp C D)) x h13) %* exist (Inn (ba_p_ens Ap)) y (incl_gen_ens_p (im_proj1_sig (proj1_sig (finc_union Atp C D))) (h2 (finc_union Atp C D)) y h14)))). f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. rewrite h16 at 1. clear h16. assumption. red. assert (h3:Included (Empty_set _) E). auto with sets. pose (exist (fun S:Ensemble Atp => Finite S /\ Included S E) (Empty_set Atp) (conj (@Empty_is_finite _) h3)) as FE. apply family_union_intro with (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig FE)) (h2 FE))). apply Im_intro with FE; auto. constructor. unfold FE. simpl. pose proof (gen_empty_p Ap) as h4. assert (h5: Gen_p (im_proj1_sig (Empty_set Atp)) (h2 (exist (fun S : Ensemble Atp => Finite S /\ Included S E) (Empty_set Atp) (conj (Empty_is_finite Atp) h3))) = Gen_p (Empty_set T) (incl_empty (ba_p_ens Ap))). apply gen_p_functional. unfold im_proj1_sig. rewrite image_empty. reflexivity. rewrite h5. rewrite ba_p_ens_gen_p_eq. rewrite gen_ens_empty_p. right. red. assert (h3:Included (Empty_set _) E). auto with sets. pose (exist (fun S:Ensemble Atp => Finite S /\ Included S E) (Empty_set Atp) (conj (@Empty_is_finite _) h3)) as FE. apply family_union_intro with (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig FE)) (h2 FE))). apply Im_intro with FE; auto. constructor. unfold FE. simpl. pose proof (gen_empty_p Ap) as h4. assert (h5: Gen_p (im_proj1_sig (Empty_set Atp)) (h2 (exist (fun S : Ensemble Atp => Finite S /\ Included S E) (Empty_set Atp) (conj (Empty_is_finite Atp) h3))) = Gen_p (Empty_set T) (incl_empty (ba_p_ens Ap))). apply gen_p_functional. unfold im_proj1_sig. rewrite image_empty. reflexivity. rewrite h5. rewrite ba_p_ens_gen_p_eq. rewrite gen_ens_empty_p. left. red. intros x. destruct x as [x h3]. simpl. destruct h3 as [C x h3]. simpl. destruct h3 as [C h3]. subst. apply family_union_intro with (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig C)) (h2 C))). apply Im_intro with C. constructor. reflexivity. simpl. assert (h5:Included (im_proj1_sig (proj1_sig C)) (ba_p_ens Ap)). red. intros a h5. destruct h5 as [a h5]. subst. apply proj2_sig. pose proof (closed_gen_ens_p _ (h2 C)) as h10. rewrite ba_p_ens_gen_p_eq at 1. pose proof (comp_closed_p' _ _ _ _ h10 _ i) as h15. assert (h16: proj1_sig (%-exist (Inn (ba_p_ens Ap)) x (incl_gen_ens_p (im_proj1_sig (proj1_sig C)) (h2 C) x i)) = proj1_sig (%-exist (Inn (A_p T (Bc_p T Ap))) x (incl_im_full_set_finc_ba_p_ens E h1 h2 x (family_union_intro T (Im (Full_set (finc E)) (fun x0 : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x0)) (h2 x0)))) (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig C)) (h2 C))) x (Im_intro (finc E) (Ensemble T) (Full_set (finc E)) (fun x0 : finc E => ba_p_ens (Gen_p (im_proj1_sig (proj1_sig x0)) (h2 x0))) C h3 (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig C)) (h2 C))) eq_refl) i)))). f_equal. f_equal. apply proj1_sig_injective; auto. rewrite <- h16 at 1. clear h16. assumption. Qed. Lemma el_prod_fin_map_dom_subst_compat_p : forall {C C':Ensemble Atp} (a:Fin_map C signe mns) (pf:C = C'), el_prod_p a = el_prod_p (fin_map_dom_subst pf a). intros C C' a h1. subst. rewrite fin_map_dom_subst_eq_refl. reflexivity. Qed. Lemma el_prod_compose_fin_map_dom_subst_compat_p1 : forall {B:Bool_Alg} {C C':Ensemble Atp} (a:Fin_map C signe mns) (g:btp Ap->bt B) (pf:C = C'), el_prod_compose_p1 g a = el_prod_compose_p1 g (fin_map_dom_subst pf a). intros B C C' a g h1. subst. rewrite fin_map_dom_subst_eq_refl. reflexivity. Qed. Lemma el_prod_compose_fin_map_dom_subst_compat_p : forall {Bp:Bool_Alg_p T} {C C':Ensemble (btp Bp)} (a:Fin_map C signe mns) (g:btp Bp-> btp Ap) (pf:C = C'), el_prod_compose_p g a = el_prod_compose_p g (fin_map_dom_subst pf a). intros B C C' a g h1. subst. rewrite fin_map_dom_subst_eq_refl. reflexivity. Qed. End NormalForm_p. Lemma el_prod_in_gen_p : forall {T:Type} {B:Bool_Alg_p T} {E:Ensemble (btp B)} (F:Fin_map E signe mns), Inn (Gen_Ens_p (im_proj1_sig E) (incl_ens_btp E)) (proj1_sig (el_prod_p _ F)). intros T B E F. pose proof (el_prod_in_gen _ _ (ba_conv_fin_map_dom F)) as h1. rewrite gen_ens_p_eq. rewrite <- im_proj2_sig_undoes_im_proj1_sig' at 1. apply Im_intro with (el_prod_p _ F). rewrite el_prod_p_eq. assumption. reflexivity. Qed. Lemma el_prod_compat_fin_map_eq_base : forall {T:Type} {Ap Bp:Bool_Alg_p T}, subalg_of_p Ap Bp -> forall (D:Ensemble (btp Ap)) (D':Ensemble (btp Bp)), im_proj1_sig D = im_proj1_sig D' -> forall (a:Fin_map D signe mns) (a':Fin_map D' signe mns) (pf:forall x : btp Ap, Inn (ba_p_ens Bp) (proj1_sig x)), (forall x:btp Ap, Inn D x -> a |-> x = a' |-> (exist _ _ (pf x))) -> proj1_sig (el_prod_p _ a) = proj1_sig (el_prod_p _ a'). intros T Ap Bp h1 D D' h0 a a' h2 h3. unfold el_prod_p. pose proof (fin_map_fin_dom a) as h4. pose proof (fin_map_fin_dom a') as h5. symmetry in h0. assert (h0':im_proj1_sig (Im D (fun i : btp Ap => eps_p i (a |-> i))) = im_proj1_sig (Im D' (fun i : btp Bp => eps_p i (a' |-> i)))). unfold im_proj1_sig. do 2 rewrite im_im. apply Extensionality_Ensembles. red; split; red; intros x h6. destruct h6 as [x h6]. subst. assert (h7:Inn (im_proj1_sig D) (proj1_sig x)). apply Im_intro with x; auto. rewrite <- h0 in h7. inversion h7 as [x' h8 q h9]. subst. apply Im_intro with x'; auto. unfold eps_p. rewrite h3 at 1. assert (h10:exist _ _ (h2 x) = x'). apply proj1_sig_injective; auto. rewrite h10 at 1. destruct (a' |-> x'); auto. destruct h1 as [h1a [h1b h1c]]. assert (h11: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h1a h1b)) (proj1_sig x)). rewrite <- h1c. apply proj2_sig. pose proof (ba_p_subst_comp _ _ h1c _ (proj2_sig x) h11) as h12. rewrite (unfold_sig _ x), (unfold_sig _ x'). rewrite h12 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. assumption. destruct h6 as [x h6]. subst. assert (h7:Inn (im_proj1_sig D') (proj1_sig x)). apply Im_intro with x; auto. rewrite h0 in h7. inversion h7 as [x' h8 q h9]. subst. apply Im_intro with x'; auto. unfold eps_p. rewrite h3 at 1. assert (h10:exist _ _ (h2 x') = x). apply proj1_sig_injective; auto. rewrite h10 at 1. destruct (a' |-> x); auto. destruct h1 as [h1a [h1b h1c]]. assert (h11: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h1a h1b)) (proj1_sig x')). rewrite <- h1c. apply proj2_sig. rewrite (unfold_sig _ x), (unfold_sig _ x'). pose proof (ba_p_subst_comp _ _ h1c _ (proj2_sig x') h11) as h12. rewrite h12 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. assumption. symmetry in h0'. pose proof (times_set_sub_compat_p _ _ h1 _ _ (finite_image (btp Bp) (btp Bp) D' (fun i : btp Bp => eps_p i (a' |-> i)) (fin_map_fin_dom a')) (finite_image (btp Ap) (btp Ap) D (fun i : btp Ap => eps_p i (a |-> i)) (fin_map_fin_dom a)) h0') as h6. rewrite h6 at 1. reflexivity. Qed. (*Note, oftentimes using this, you will use the argument [sig_fun_app g' 0] for what is listed as [g'], utilizing a different function [g':sig_set E'->bt A].*) Lemma el_prod_compose_compat_fin_map_eq_base : forall {T:Type} {Ap Bp:Bool_Alg_p T} (A:Bool_Alg), subalg_of_p Ap Bp -> forall {E:Ensemble (btp Bp)} (g:sig_set E -> bt A) (g':btp Ap->bt A) (D:Ensemble (btp Ap)) (D':Ensemble (btp Bp)), im_proj1_sig D = im_proj1_sig D' -> forall (pfinc:Included D' E) (pf : forall x : btp Ap, Inn (ba_p_ens Bp) (proj1_sig x)), (forall (x:btp Ap) (pfin:Inn D x), g' x = (sig_fun_app g 0) (exist _ _ (pf x))) -> forall (a:Fin_map D signe mns) (a':Fin_map D' signe mns), (forall x:btp Ap, Inn D x -> a |-> x = a' |-> (exist _ _ (pf x))) -> el_prod_compose_p1 _ (sig_fun_app g 0) a' = el_prod_compose_p1 _ g' a. intros T Ap Bp A h1 E g g' D D' h2 h3 h4 h5 a a' h6. unfold el_prod_compose_p1. apply times_set_functional. apply Extensionality_Ensembles. red; split; red; intros x h7. destruct h7 as [x h7]. subst. assert (h8:Inn (im_proj1_sig D') (proj1_sig x)). apply Im_intro with x; auto. rewrite <- h2 in h8. inversion h8 as [x' h9 q h10]. subst. apply Im_intro with x'; auto. unfold eps. destruct x as [x h11], x' as [x' h12]; simpl. simpl in h7, h8, h9, h10. subst. rewrite h6; auto. simpl. assert (h10:exist _ _ h11 = exist _ _ (h4 (exist (fun x : T => Inn (A_p T (Bc_p T Ap)) x) x' h12))). apply proj1_sig_injective; auto. rewrite h10 at 1. simpl. destruct (a' |-> exist (Inn (ba_p_ens Bp)) x' (h4 (exist (fun x : T => Inn (A_p T (Bc_p T Ap)) x) x' h12))). rewrite h5. f_equal. apply proj1_sig_injective; auto. assumption. rewrite h5. f_equal. f_equal. apply proj1_sig_injective; auto. assumption. destruct h7 as [x h7]. subst. assert (h8:Inn (im_proj1_sig D) (proj1_sig x)). apply Im_intro with x; auto. rewrite h2 in h8. inversion h8 as [x' h9 q h10]. subst. apply Im_intro with x'; auto. unfold eps. destruct x as [x h11], x' as [x' h12]; simpl. simpl in h7, h8, h9, h10. subst. rewrite h6; auto. simpl. assert (h10:exist _ _ h12 = exist _ _ (h4 (exist (fun x : T => Inn (A_p T (Bc_p T Ap)) x) x' h11))). apply proj1_sig_injective; auto. rewrite h10 at 1. simpl. destruct (a' |-> exist (Inn (ba_p_ens Bp)) x' (h4 (exist (fun x : T => Inn (A_p T (Bc_p T Ap)) x) x' h11))). rewrite h5. f_equal. apply proj1_sig_injective; auto. assumption. rewrite h5. f_equal. f_equal. apply proj1_sig_injective; auto. assumption. Qed. Lemma el_prod_compose_decompose_restriction_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} (g:btp Bp->bt A) (E F:Ensemble (btp Bp)) (pf:Included F E) (a:Fin_map E signe mns), el_prod_compose_p1 _ g a = el_prod_compose_p1 _ g (restriction a pf) * el_prod_compose_p1 _ g (restriction a (setminus_inc E F)). intros T A Bp g E F h1 a. do 3 rewrite el_prod_compose_p1_eq. rewrite (el_prod_compose_decompose_restriction _ (ba_conv_fun1 g) _ _ h1). reflexivity. Qed. Section SubLift. (*This is for lifting the type of an element of a subalgebra to the type of the superalgebra. It only uses the inclusion property of being a subalgebra, not the structure.*) Definition sub_lift {T:Type} {Ap Bp:Bool_Alg_p T} (pf:subalg_of_p Bp Ap) (x:btp Bp) : btp Ap. pose proof (subalg_of_p_incl _ _ pf) as h1. refine (exist _ _ (h1 _ (proj2_sig x))). Defined. Lemma sub_lift_plus_eq : forall {T:Type} {Ap Bp:Bool_Alg_p T} (pf:subalg_of_p Bp Ap) (x y:btp Bp), sub_lift pf (x %+ y) = (sub_lift pf x) %+ (sub_lift pf y). intros T Ap Bp h1 x y. unfold sub_lift. destruct h1 as [h1a [h1b h1c]]. apply proj1_sig_injective. simpl. assert (h3:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) (proj1_sig x)). rewrite <- h1c. apply proj2_sig. assert (h4 : Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) (proj1_sig y)). rewrite <- h1c. apply proj2_sig. pose proof (ba_p_subst_plus _ _ h1c (proj1_sig x) (proj1_sig y) (proj2_sig x) (proj2_sig y) h3 h4) as h2. simpl in h2. do 2 rewrite <- unfold_sig in h2. rewrite h2 at 1. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. Qed. Lemma sub_lift_times_eq : forall {T:Type} {Ap Bp:Bool_Alg_p T} (pf:subalg_of_p Bp Ap) (x y:btp Bp), sub_lift pf (x %* y) = (sub_lift pf x) %* (sub_lift pf y). intros T Ap Bp h1 x y. unfold sub_lift. destruct h1 as [h1a [h1b h1c]]. apply proj1_sig_injective. simpl. assert (h3:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) (proj1_sig x)). rewrite <- h1c. apply proj2_sig. assert (h4 : Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) (proj1_sig y)). rewrite <- h1c. apply proj2_sig. pose proof (ba_p_subst_times _ _ h1c (proj1_sig x) (proj1_sig y) (proj2_sig x) (proj2_sig y) h3 h4) as h2. simpl in h2. do 2 rewrite <- unfold_sig in h2. rewrite h2 at 1. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. Qed. Lemma sub_lift_comp_eq : forall {T:Type} {Ap Bp:Bool_Alg_p T} (pf:subalg_of_p Bp Ap) (x:btp Bp), sub_lift pf (%- x) = %- (sub_lift pf x). intros T Ap Bp h1 x. unfold sub_lift. destruct h1 as [h1a [h1b h1c]]. apply proj1_sig_injective. simpl. assert (h3:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) (proj1_sig x)). rewrite <- h1c. apply proj2_sig. pose proof (ba_p_subst_comp _ _ h1c (proj1_sig x) (proj2_sig x) h3) as h2. simpl in h2. rewrite <- unfold_sig in h2. rewrite h2 at 1. f_equal. f_equal. apply proj1_sig_injective; auto. Qed. Lemma sub_lift_zero_eq : forall {T:Type} {Ap Bp:Bool_Alg_p T} (pf:subalg_of_p Bp Ap), sub_lift pf (%0) = %0. intros T Ap Bp h1. unfold sub_lift. apply proj1_sig_injective. simpl. destruct h1 as [h1a [h1b h1c]]. pose proof (ba_p_subst_zero _ _ h1c) as h2. rewrite h2 at 1. reflexivity. Qed. Lemma sub_lift_one_eq : forall {T:Type} {Ap Bp:Bool_Alg_p T} (pf:subalg_of_p Bp Ap), sub_lift pf (%1) = %1. intros T Ap Bp h1. unfold sub_lift. apply proj1_sig_injective. simpl. destruct h1 as [h1a [h1b h1c]]. pose proof (ba_p_subst_one _ _ h1c) as h2. rewrite h2 at 1. reflexivity. Qed. Lemma le_p_sub_lift_iff : forall {T:Type} {Ap Bp:Bool_Alg_p T} (pf:subalg_of_p Bp Ap) (x y:btp Bp), le_p x y <-> le_p (sub_lift pf x) (sub_lift pf y). intros T Ap Bp h1 x y. split. intro h2. red in h2; red. unfold sub_lift. apply proj1_sig_injective. red in h1. destruct h1 as [h1a [h1b h1c]]. simpl. assert (h4:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) (proj1_sig x)). rewrite <- h1c. apply proj2_sig. assert (h5:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) (proj1_sig y)). rewrite <- h1c. apply proj2_sig. pose proof (ba_p_subst_plus _ _ h1c (proj1_sig x) (proj1_sig y) (proj2_sig x) (proj2_sig y) h4 h5) as h6. simpl in h6. assert (h7:(exist (Inn (A_p T (Bc_p T Ap))) (proj1_sig x) (h1a (proj1_sig x) h4) %+ exist (Inn (A_p T (Bc_p T Ap))) (proj1_sig y) (h1a (proj1_sig y) h5)) = (exist (Inn (ba_p_ens Ap)) (proj1_sig x) (subalg_of_p_incl Bp Ap (ex_intro (fun pfin : Included (ba_p_ens Bp) (ba_p_ens Ap) => exists pfac : alg_closed_p (ba_p_ens Bp) pfin, Bp = Subalg_p Ap (ba_p_ens Bp) pfin pfac) h1a (ex_intro (fun pfac : alg_closed_p (ba_p_ens Bp) h1a => Bp = Subalg_p Ap (ba_p_ens Bp) h1a pfac) h1b h1c)) (proj1_sig x) (proj2_sig x)) %+ exist (Inn (ba_p_ens Ap)) (proj1_sig y) (subalg_of_p_incl Bp Ap (ex_intro (fun pfin : Included (ba_p_ens Bp) (ba_p_ens Ap) => exists pfac : alg_closed_p (ba_p_ens Bp) pfin, Bp = Subalg_p Ap (ba_p_ens Bp) pfin pfac) h1a (ex_intro (fun pfac : alg_closed_p (ba_p_ens Bp) h1a => Bp = Subalg_p Ap (ba_p_ens Bp) h1a pfac) h1b h1c)) (proj1_sig y) (proj2_sig y)))). f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. rewrite <- h7 at 1. rewrite <- h6 at 1. clear h7. clear h6. do 2 rewrite <- unfold_sig. f_equal. assumption. intro h2. destruct h1 as [h1a [h1b h1c]]. red in h2. red. unfold sub_lift in h2. assert (h4:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) (proj1_sig x)). rewrite <- h1c. apply proj2_sig. assert (h5:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Bp) h1a h1b)) (proj1_sig y)). rewrite <- h1c. apply proj2_sig. pose proof (ba_p_subst_plus _ _ h1c (proj1_sig x) (proj1_sig y) (proj2_sig x) (proj2_sig y) h4 h5) as h6. simpl in h6. assert (h7:(exist (Inn (A_p T (Bc_p T Ap))) (proj1_sig x) (h1a (proj1_sig x) h4) %+ exist (Inn (A_p T (Bc_p T Ap))) (proj1_sig y) (h1a (proj1_sig y) h5)) = (exist (Inn (ba_p_ens Ap)) (proj1_sig x) (subalg_of_p_incl Bp Ap (ex_intro (fun pfin : Included (ba_p_ens Bp) (ba_p_ens Ap) => exists pfac : alg_closed_p (ba_p_ens Bp) pfin, Bp = Subalg_p Ap (ba_p_ens Bp) pfin pfac) h1a (ex_intro (fun pfac : alg_closed_p (ba_p_ens Bp) h1a => Bp = Subalg_p Ap (ba_p_ens Bp) h1a pfac) h1b h1c)) (proj1_sig x) (proj2_sig x)) %+ exist (Inn (ba_p_ens Ap)) (proj1_sig y) (subalg_of_p_incl Bp Ap (ex_intro (fun pfin : Included (ba_p_ens Bp) (ba_p_ens Ap) => exists pfac : alg_closed_p (ba_p_ens Bp) pfin, Bp = Subalg_p Ap (ba_p_ens Bp) pfin pfac) h1a (ex_intro (fun pfac : alg_closed_p (ba_p_ens Bp) h1a => Bp = Subalg_p Ap (ba_p_ens Bp) h1a pfac) h1b h1c)) (proj1_sig y) (proj2_sig y)))). f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. rewrite <- h7 in h2 at 1. pose proof (f_equal (@proj1_sig _ _) h2) as h2'. simpl in h2'. rewrite <- h6 in h2' at 1. apply proj1_sig_injective. do 2 rewrite <- unfold_sig in h2'. assumption. Qed. Inductive le_p_sub {T:Type} {Ap Bp:Bool_Alg_p T} (b:btp Bp) (a:btp Ap): Prop := | le_p_sub_intro : forall (pf:subalg_of_p Bp Ap), le_p (sub_lift pf b) a -> le_p_sub b a. Inductive le_p_sub_r {T:Type} {Ap Bp:Bool_Alg_p T} (a:btp Ap) (b:btp Bp) : Prop := | le_p_sub_intro_r : forall (pf:subalg_of_p Bp Ap), le_p a (sub_lift pf b) -> le_p_sub_r a b. Lemma zero_min_sub : forall {T:Type} {Ap Bp:Bool_Alg_p T}, subalg_of_p Bp Ap -> forall (a:btp Ap), le_p_sub (Bzero_p _ (Bc_p _ Bp)) a. intros T Ap Bp h1 a. assert (h3:le_p (sub_lift h1 %0) a). rewrite sub_lift_zero_eq. apply zero_min_p. apply (le_p_sub_intro %0 a h1 h3). Qed. Lemma zero_min_sub_r : forall {T:Type} {Ap Bp:Bool_Alg_p T}, subalg_of_p Bp Ap -> forall (b:btp Bp), le_p_sub_r (Bzero_p _ (Bc_p _ Ap)) b. intros T Ap Bp h1 b. assert (h3:le_p %0 (sub_lift h1 b)). rewrite <- (sub_lift_zero_eq h1). rewrite <- le_p_sub_lift_iff. apply zero_min_p. apply (le_p_sub_intro_r %0 b h1 h3). Qed. Lemma one_max_sub : forall {T:Type} {Ap Bp:Bool_Alg_p T}, subalg_of_p Bp Ap -> forall (b:btp Bp), le_p_sub b (Bone_p _ (Bc_p _ Ap)). intros T Ap Bp h1 b. assert (h3:le_p (sub_lift h1 b) %1). rewrite <- (sub_lift_one_eq h1). rewrite <- le_p_sub_lift_iff. apply one_max_p. apply (le_p_sub_intro b %1 h1 h3). Qed. Lemma one_max_sub_r : forall {T:Type} {Ap Bp:Bool_Alg_p T}, subalg_of_p Bp Ap -> forall (a:btp Ap), le_p_sub_r a (Bone_p _ (Bc_p _ Bp)). intros T Ap Bp h1 a. assert (h3:le_p a (sub_lift h1 (Bone_p _ (Bc_p _ Bp)))). rewrite (sub_lift_one_eq h1). apply one_max_p. apply (le_p_sub_intro_r a %1 h1 h3). Qed. End SubLift. Lemma degen_subalg_of_compat_p : forall {T:Type} (Ap Bp: Bool_Alg_p T), subalg_of_p Ap Bp -> (degen_p Ap <-> degen_p Bp). intros T Ap Bp h1. destruct h1 as [h1 [h2 h3]]. pose proof (ba_p_subst_zero _ _ h3) as h5. pose proof (ba_p_subst_one _ _ h3) as h6. unfold degen_p. simpl in h5, h6. split. intro h4. apply proj1_sig_injective. rewrite <- h5 at 1. rewrite <- h6 at 1. rewrite h4. reflexivity. intro h4. apply proj1_sig_injective. rewrite h5 at 1. rewrite h6 at 1. rewrite h4. reflexivity. Qed. End ParametricAnalogues. Arguments alg_closed_p [T] [Bp] _ _. Arguments Subalg_p [T] _ _ _ _. Arguments el_prod_l_p [T] [Bp] [l] _. Arguments el_prod_l_firstn_p [T] [Bp] [l] [m] _ _. Section DirectedFamilyOfAlgebras. Variable T:Type. Definition directed (S:fam_ba_p T) := forall Ap Bp:Bool_Alg_p T, Inn S Ap -> Inn S Bp -> exists Cp:Bool_Alg_p T, Inn S Cp /\ subalg_of_p Ap Cp /\ subalg_of_p Bp Cp. Lemma directed_times_ex : forall {S:fam_ba_p T}, directed S -> forall (a b:sig_set (FamilyUnion (Im S (@ba_p_ens T)))), exists! c:sig_set (FamilyUnion (Im S (@ba_p_ens T))), forall (Ap Bp:Bool_Alg_p T) (pfina:Inn (ba_p_ens Ap) (proj1_sig a)) (pfinb:Inn (ba_p_ens Bp) (proj1_sig b)), Inn S Ap -> Inn S Bp -> exists (Cp:Bool_Alg_p T) (pfinca:Included (ba_p_ens Ap) (ba_p_ens Cp)) (pfincb:Included (ba_p_ens Bp) (ba_p_ens Cp)) (pfaca:alg_closed_p _ pfinca) (pfacb:alg_closed_p _ pfincb), Inn (ba_p_ens Cp) (proj1_sig c) /\ Ap = Subalg_p _ _ pfinca pfaca /\ Bp = Subalg_p _ _ pfincb pfacb /\ Inn S Cp /\ proj1_sig c = proj1_sig (exist _ _ (pfinca _ pfina) %* exist _ _ (pfincb _ pfinb)). intros S h1 a b. destruct a as [a h2], b as [b h3]. destruct h2 as [A a h2 h4], h3 as [B b h3 h5]. destruct h2 as [Ap h2], h3 as [Bp h3]. subst. simpl. red in h1. pose proof (h1 _ _ h2 h3) as h6. destruct h6 as [Cp [h6 [h7 h8]]]. red in h7, h8. destruct h7 as [h7 [h9 h10]], h8 as [h8 [h11 h12]]. pose (exist _ _ (h7 _ h4) %* exist _ _ (h8 _ h5)) as c'. assert (h13:Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) (proj1_sig c')). unfold c'. apply family_union_intro with (ba_p_ens Cp). apply Im_intro with Cp. assumption. reflexivity. apply proj2_sig. exists (exist _ _ h13). red. split. intros A B h14 h15 h16 h17. pose proof (h1 _ _ h6 h16) as h18. pose proof (h1 _ _ h6 h17) as h19. destruct h18 as [C [h18a [h18b h18c]]], h19 as [D [h19a [h19b h19c]]]. pose proof (h1 _ _ h18a h19a) as h20. destruct h20 as [Ep [h20a [h20b h20c]]]. pose proof h18b as h18b'. pose proof h19b as h19b'. pose proof h18c as h18c'. pose proof h19c as h19c'. pose proof h20b as h20b'. pose proof h20c as h20c'. red in h18c, h19c, h20b, h20c. destruct h18c as [h21a [h21b h21c]], h20b as [h23a [h23b h23c]], h19c as [h22a [h22b h22c]], h20c as [h24a [h24b h24c]]. exists Ep. pose proof (Inclusion_is_transitive _ _ _ _ h21a h23a) as h25. pose proof (Inclusion_is_transitive _ _ _ _ h22a h24a) as h26. pose proof (trans_subalg_of_p _ _ _ _ h18c' h20b') as h27. pose proof (trans_subalg_of_p _ _ _ _ h19c' h20c') as h28. destruct h27 as [h27a [h27b h27c]], h28 as [h28a [h28b h28c]]. exists h27a, h28a, h27b, h28b. simpl. repeat split. apply h24a. destruct h19b as [h29a [h29b h29c]]. apply h29a. apply proj2_sig. assumption. assumption. assumption. unfold c'. red in h19b. destruct h19b as [h29 [h30 h31]]. pose proof (ba_p_subst_times _ _ h31 _ _ (h7 a h4) (h8 b h5)) as h32. assert (h33:Inn (ba_p_ens (Subalg_p D (ba_p_ens Cp) h29 h30)) a). rewrite <- h31. apply (h7 a h4). assert (h34 : Inn (ba_p_ens (Subalg_p D (ba_p_ens Cp) h29 h30)) b). rewrite <- h31. apply (h8 b h5). specialize (h32 h33 h34). rewrite h32 at 1. simpl. pose proof (ba_p_subst_times _ _ h24c _ _ (h29 a h33) (h29 b h34)) as h35. assert (h36:Inn (ba_p_ens (Subalg_p Ep (ba_p_ens D) h24a h24b)) a). rewrite <- h24c. apply (h29 a h33). assert (h37 : Inn (ba_p_ens (Subalg_p Ep (ba_p_ens D) h24a h24b)) b). rewrite <- h24c. apply (h29 b h34). specialize (h35 h36 h37). rewrite h35 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. intros c h14. specialize (h14 _ _ h4 h5 h2 h3). destruct h14 as [Dp [h15 [h16 [h17 [h18 [h19 [h20 [h21 [h22 h23]]]]]]]]]. apply proj1_sig_injective. simpl. rewrite h23 at 1. unfold c'. pose proof (h1 _ _ h6 h22) as h24. destruct h24 as [D [h24 [h25 h26]]]. red in h25, h26. destruct h25 as [h25a [h25b h25c]], h26 as [h26a [h26b h26c]]. pose proof (ba_p_subst_times _ _ h25c _ _ (h7 a h4) (h8 b h5)) as h27. assert (h28: Inn (ba_p_ens (Subalg_p D (ba_p_ens Cp) h25a h25b)) a). rewrite <- h25c. apply (h7 a h4). assert (h29 : Inn (ba_p_ens (Subalg_p D (ba_p_ens Cp) h25a h25b)) b). rewrite <- h25c. apply (h8 b h5). specialize (h27 h28 h29). rewrite h27 at 1. simpl. pose proof (ba_p_subst_times _ _ h26c _ _ (h15 a h4) (h16 b h5)) as h30. assert (h31:Inn (ba_p_ens (Subalg_p D (ba_p_ens Dp) h26a h26b)) a). rewrite <- h26c. apply (h15 a h4). assert (h32 : Inn (ba_p_ens (Subalg_p D (ba_p_ens Dp) h26a h26b)) b). rewrite <- h26c. apply (h16 b h5). specialize (h30 h31 h32). rewrite h30 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. Qed. Definition directed_times {S:fam_ba_p T} (pf:directed S) : sig_set (FamilyUnion (Im S (@ba_p_ens T))) -> sig_set (FamilyUnion (Im S (@ba_p_ens T))) -> sig_set (FamilyUnion (Im S (@ba_p_ens T))) := (fun a b => proj1_sig (constructive_definite_description _ (directed_times_ex pf a b))). Lemma directed_times_compat : forall {S:fam_ba_p T} (pf:directed S) (a b:sig_set (FamilyUnion (Im S (@ba_p_ens T)))), let c:=directed_times pf a b in forall (Ap Bp:Bool_Alg_p T) (pfina:Inn (ba_p_ens Ap) (proj1_sig a)) (pfinb:Inn (ba_p_ens Bp) (proj1_sig b)), Inn S Ap -> Inn S Bp -> exists (Cp:Bool_Alg_p T) (pfinca:Included (ba_p_ens Ap) (ba_p_ens Cp)) (pfincb:Included (ba_p_ens Bp) (ba_p_ens Cp)) (pfaca:alg_closed_p _ pfinca) (pfacb:alg_closed_p _ pfincb), Inn (ba_p_ens Cp) (proj1_sig c) /\ Ap = Subalg_p _ _ pfinca pfaca /\ Bp = Subalg_p _ _ pfincb pfacb /\ Inn S Cp /\ proj1_sig c = proj1_sig (exist _ _ (pfinca _ pfina) %* exist _ _ (pfincb _ pfinb)). intros S pf a b c. unfold directed_times in c. destruct constructive_definite_description as [c' h2]. simpl in c. simpl. assumption. Qed. Lemma directed_plus_ex : forall {S:fam_ba_p T}, directed S -> forall (a b:sig_set (FamilyUnion (Im S (@ba_p_ens T)))), exists! c:sig_set (FamilyUnion (Im S (@ba_p_ens T))), forall (Ap Bp:Bool_Alg_p T) (pfina:Inn (ba_p_ens Ap) (proj1_sig a)) (pfinb:Inn (ba_p_ens Bp) (proj1_sig b)), Inn S Ap -> Inn S Bp -> exists (Cp:Bool_Alg_p T) (pfinca:Included (ba_p_ens Ap) (ba_p_ens Cp)) (pfincb:Included (ba_p_ens Bp) (ba_p_ens Cp)) (pfaca:alg_closed_p _ pfinca) (pfacb:alg_closed_p _ pfincb), Inn (ba_p_ens Cp) (proj1_sig c) /\ Ap = Subalg_p _ _ pfinca pfaca /\ Bp = Subalg_p _ _ pfincb pfacb /\ Inn S Cp /\ proj1_sig c = proj1_sig (exist _ _ (pfinca _ pfina) %+ exist _ _ (pfincb _ pfinb)). intros S h1 a b. destruct a as [a h2], b as [b h3]. destruct h2 as [A a h2 h4], h3 as [B b h3 h5]. destruct h2 as [Ap h2], h3 as [Bp h3]. subst. simpl. red in h1. pose proof (h1 _ _ h2 h3) as h6. destruct h6 as [Cp [h6 [h7 h8]]]. red in h7, h8. destruct h7 as [h7 [h9 h10]], h8 as [h8 [h11 h12]]. pose (exist _ _ (h7 _ h4) %+ exist _ _ (h8 _ h5)) as c'. assert (h13:Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) (proj1_sig c')). unfold c'. apply family_union_intro with (ba_p_ens Cp). apply Im_intro with Cp. assumption. reflexivity. apply proj2_sig. exists (exist _ _ h13). red. split. intros A B h14 h15 h16 h17. pose proof (h1 _ _ h6 h16) as h18. pose proof (h1 _ _ h6 h17) as h19. destruct h18 as [C [h18a [h18b h18c]]], h19 as [D [h19a [h19b h19c]]]. pose proof (h1 _ _ h18a h19a) as h20. destruct h20 as [Ep [h20a [h20b h20c]]]. pose proof h18b as h18b'. pose proof h19b as h19b'. pose proof h18c as h18c'. pose proof h19c as h19c'. pose proof h20b as h20b'. pose proof h20c as h20c'. red in h18c, h19c, h20b, h20c. destruct h18c as [h21a [h21b h21c]], h20b as [h23a [h23b h23c]], h19c as [h22a [h22b h22c]], h20c as [h24a [h24b h24c]]. exists Ep. pose proof (Inclusion_is_transitive _ _ _ _ h21a h23a) as h25. pose proof (Inclusion_is_transitive _ _ _ _ h22a h24a) as h26. pose proof (trans_subalg_of_p _ _ _ _ h18c' h20b') as h27. pose proof (trans_subalg_of_p _ _ _ _ h19c' h20c') as h28. destruct h27 as [h27a [h27b h27c]], h28 as [h28a [h28b h28c]]. exists h27a, h28a, h27b, h28b. simpl. repeat split. apply h24a. destruct h19b as [h29a [h29b h29c]]. apply h29a. apply proj2_sig. assumption. assumption. assumption. unfold c'. red in h19b. destruct h19b as [h29 [h30 h31]]. pose proof (ba_p_subst_plus _ _ h31 _ _ (h7 a h4) (h8 b h5)) as h32. assert (h33:Inn (ba_p_ens (Subalg_p D (ba_p_ens Cp) h29 h30)) a). rewrite <- h31. apply (h7 a h4). assert (h34 : Inn (ba_p_ens (Subalg_p D (ba_p_ens Cp) h29 h30)) b). rewrite <- h31. apply (h8 b h5). specialize (h32 h33 h34). rewrite h32 at 1. simpl. pose proof (ba_p_subst_plus _ _ h24c _ _ (h29 a h33) (h29 b h34)) as h35. assert (h36:Inn (ba_p_ens (Subalg_p Ep (ba_p_ens D) h24a h24b)) a). rewrite <- h24c. apply (h29 a h33). assert (h37 : Inn (ba_p_ens (Subalg_p Ep (ba_p_ens D) h24a h24b)) b). rewrite <- h24c. apply (h29 b h34). specialize (h35 h36 h37). rewrite h35 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. intros c h14. specialize (h14 _ _ h4 h5 h2 h3). destruct h14 as [Dp [h15 [h16 [h17 [h18 [h19 [h20 [h21 [h22 h23]]]]]]]]]. apply proj1_sig_injective. simpl. rewrite h23 at 1. unfold c'. pose proof (h1 _ _ h6 h22) as h24. destruct h24 as [D [h24 [h25 h26]]]. red in h25, h26. destruct h25 as [h25a [h25b h25c]], h26 as [h26a [h26b h26c]]. pose proof (ba_p_subst_plus _ _ h25c _ _ (h7 a h4) (h8 b h5)) as h27. assert (h28: Inn (ba_p_ens (Subalg_p D (ba_p_ens Cp) h25a h25b)) a). rewrite <- h25c. apply (h7 a h4). assert (h29 : Inn (ba_p_ens (Subalg_p D (ba_p_ens Cp) h25a h25b)) b). rewrite <- h25c. apply (h8 b h5). specialize (h27 h28 h29). rewrite h27 at 1. simpl. pose proof (ba_p_subst_plus _ _ h26c _ _ (h15 a h4) (h16 b h5)) as h30. assert (h31:Inn (ba_p_ens (Subalg_p D (ba_p_ens Dp) h26a h26b)) a). rewrite <- h26c. apply (h15 a h4). assert (h32 : Inn (ba_p_ens (Subalg_p D (ba_p_ens Dp) h26a h26b)) b). rewrite <- h26c. apply (h16 b h5). specialize (h30 h31 h32). rewrite h30 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. Qed. Definition directed_plus {S:fam_ba_p T} (pf:directed S) : sig_set (FamilyUnion (Im S (@ba_p_ens T))) -> sig_set (FamilyUnion (Im S (@ba_p_ens T))) -> sig_set (FamilyUnion (Im S (@ba_p_ens T))) := (fun a b => proj1_sig (constructive_definite_description _ (directed_plus_ex pf a b))). Lemma directed_plus_compat : forall {S:fam_ba_p T} (pf:directed S) (a b:sig_set (FamilyUnion (Im S (@ba_p_ens T)))), let c:=directed_plus pf a b in forall (Ap Bp:Bool_Alg_p T) (pfina:Inn (ba_p_ens Ap) (proj1_sig a)) (pfinb:Inn (ba_p_ens Bp) (proj1_sig b)), Inn S Ap -> Inn S Bp -> exists (Cp:Bool_Alg_p T) (pfinca:Included (ba_p_ens Ap) (ba_p_ens Cp)) (pfincb:Included (ba_p_ens Bp) (ba_p_ens Cp)) (pfaca:alg_closed_p _ pfinca) (pfacb:alg_closed_p _ pfincb), Inn (ba_p_ens Cp) (proj1_sig c) /\ Ap = Subalg_p _ _ pfinca pfaca /\ Bp = Subalg_p _ _ pfincb pfacb /\ Inn S Cp /\ proj1_sig c = proj1_sig (exist _ _ (pfinca _ pfina) %+ exist _ _ (pfincb _ pfinb)). intros S pf a b c. unfold directed_plus in c. destruct constructive_definite_description as [c' h2]. simpl in c. simpl. assumption. Qed. Lemma directed_comp_ex : forall {S:fam_ba_p T}, directed S -> forall a:sig_set (FamilyUnion (Im S (@ba_p_ens T))), exists! c:sig_set (FamilyUnion (Im S (@ba_p_ens T))), forall (Ap:Bool_Alg_p T) (pfina:Inn (ba_p_ens Ap) (proj1_sig a)), Inn S Ap -> exists (Cp:Bool_Alg_p T), Inn (ba_p_ens Cp) (proj1_sig c) /\ subalg_of_p Ap Cp /\ Inn S Cp /\ proj1_sig c = proj1_sig (%- exist _ _ pfina). intros S h1 a. destruct a as [a h2]. destruct h2 as [A a h2 h4]. destruct h2 as [A h2]. subst. pose (%- exist _ _ h4) as c'. assert (h12:Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) (proj1_sig c')). apply family_union_intro with (ba_p_ens A). apply Im_intro with A. assumption. reflexivity. apply in_ba_p_ens_comp. inversion h12 as [C ? hc hd]. subst. destruct hc as [Cp he]. subst. exists (exist _ _ h12). simpl. red. split. intros Ap h0 h0'. simpl. pose proof (h1 _ _ h0' he) as hf. destruct hf as [Dp [hg [hh hi]]]. exists Dp. repeat split; auto. red in hi. destruct hi as [hi [hj hk]]. apply hi; auto. unfold c'. pose proof (h1 _ _ h2 h0') as hj. destruct hj as [Ep [hk [hl hm]]]. red in hl, hm. destruct hl as [hla [hlb hlc]], hm as [hma [hmb hmc]]. pose proof (ba_p_subst_comp _ _ hlc _ h4) as hn. assert (ho:Inn (ba_p_ens (Subalg_p Ep (ba_p_ens A) hla hlb)) a). rewrite <- hlc. auto. specialize (hn ho). pose proof (ba_p_subst_comp _ _ hmc _ h0) as hp. assert (hq: Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap) hma hmb)) a). rewrite <- hmc. auto. specialize (hp hq). rewrite hn, hp. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. intros c h13. specialize (h13 _ h4 h2). destruct h13 as [Dp [h14 [h15 [h16 h17]]]]. apply proj1_sig_injective. simpl. rewrite h17 at 1. unfold c'. reflexivity. Qed. Definition directed_comp {S:fam_ba_p T} (pf:directed S) : sig_set (FamilyUnion (Im S (@ba_p_ens T))) -> sig_set (FamilyUnion (Im S (@ba_p_ens T))) := (fun a => proj1_sig (constructive_definite_description _ (directed_comp_ex pf a))). Lemma directed_comp_compat : forall {S:fam_ba_p T} (pf:directed S), forall a:sig_set (FamilyUnion (Im S (@ba_p_ens T))), let c := directed_comp pf a in forall (Ap:Bool_Alg_p T) (pfina:Inn (ba_p_ens Ap) (proj1_sig a)), Inn S Ap -> exists (Cp:Bool_Alg_p T), Inn (ba_p_ens Cp) (proj1_sig c) /\ subalg_of_p Ap Cp /\ Inn S Cp /\ proj1_sig c = proj1_sig (%- exist _ _ pfina). intros. unfold directed_comp in c. destruct constructive_definite_description as [c' h1]. simpl in c. unfold c. apply h1; auto. Qed. Lemma directed_zero_ex : forall {S:fam_ba_p T}, directed S -> Inhabited S -> exists! zr:sig_set (FamilyUnion (Im S (@ba_p_ens T))), forall (Ap:Bool_Alg_p T) (pfinas:Inn S Ap), exists (Cp:Bool_Alg_p T), Inn S Cp /\ Inn (ba_p_ens Cp) (proj1_sig zr) /\ subalg_of_p Ap Cp /\ proj1_sig (Bzero_p T (Bc_p T Ap)) = proj1_sig zr. intros S h1 h2. destruct h2 as [Cp h2]. assert (h3:Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) (proj1_sig (Bzero_p T (Bc_p T Cp)))). apply family_union_intro with (ba_p_ens Cp). apply Im_intro with Cp. assumption. reflexivity. apply proj2_sig. exists (exist _ _ h3). red. split. intros Ap h4. red in h1. specialize (h1 _ _ h4 h2). destruct h1 as [Dp [h1a [h1b h1c]]]. destruct h1c as [h5 [h6 h7]]. exists Dp. repeat split; simpl; auto. pose proof (ba_p_subst_zero _ _ h7) as h8. rewrite h8 at 1. simpl. apply in_ba_p_ens_zero. destruct h1b as [h9 [h10 h11]]. pose proof (ba_p_subst_zero _ _ h7) as h8. rewrite h8 at 1. pose proof (ba_p_subst_zero _ _ h11) as h12. rewrite h12 at 1. reflexivity. intros zr h4. apply proj1_sig_injective. specialize (h4 _ h2). destruct h4 as [Dp [h5 [h6 [h7 h8]]]]. simpl. rewrite h8. reflexivity. Qed. Definition directed_zero {S:fam_ba_p T} (pf:directed S) (pfinh:Inhabited S) := (proj1_sig (constructive_definite_description _ (directed_zero_ex pf pfinh))). Lemma directed_zero_compat : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let zr := directed_zero pfd pfinh in forall (Ap:Bool_Alg_p T) (pfinas:Inn S Ap), exists (Cp:Bool_Alg_p T), Inn S Cp /\ Inn (ba_p_ens Cp) (proj1_sig zr) /\ subalg_of_p Ap Cp /\ proj1_sig (Bzero_p T (Bc_p T Ap)) = proj1_sig zr. intros S h1 h2 zr. unfold directed_zero in zr. destruct constructive_definite_description as [zr' h3]. simpl in zr. unfold zr. assumption. Qed. Lemma directed_one_ex : forall {S:fam_ba_p T}, directed S -> Inhabited S -> exists! one:sig_set (FamilyUnion (Im S (@ba_p_ens T))), forall (Ap:Bool_Alg_p T) (pfinas:Inn S Ap), exists (Cp:Bool_Alg_p T), Inn S Cp /\ Inn (ba_p_ens Cp) (proj1_sig one) /\ subalg_of_p Ap Cp /\ proj1_sig (Bone_p T (Bc_p T Ap)) = proj1_sig one. intros S h1 h2. destruct h2 as [Cp h2]. assert (h3:Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) (proj1_sig (Bone_p T (Bc_p T Cp)))). apply family_union_intro with (ba_p_ens Cp). apply Im_intro with Cp. assumption. reflexivity. apply proj2_sig. exists (exist _ _ h3). red. split. intros Ap h4. red in h1. specialize (h1 _ _ h4 h2). destruct h1 as [Dp [h1a [h1b h1c]]]. destruct h1c as [h5 [h6 h7]]. exists Dp. repeat split; simpl; auto. pose proof (ba_p_subst_one _ _ h7) as h8. rewrite h8 at 1. simpl. apply in_ba_p_ens_one. destruct h1b as [h9 [h10 h11]]. pose proof (ba_p_subst_one _ _ h7) as h8. rewrite h8 at 1. pose proof (ba_p_subst_one _ _ h11) as h12. rewrite h12 at 1. reflexivity. intros zr h4. apply proj1_sig_injective. specialize (h4 _ h2). destruct h4 as [Dp [h5 [h6 [h7 h8]]]]. simpl. rewrite h8. reflexivity. Qed. Definition directed_one {S:fam_ba_p T} (pf:directed S) (pfinh:Inhabited S) := (proj1_sig (constructive_definite_description _ (directed_one_ex pf pfinh))). Lemma directed_one_compat : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let one := directed_one pfd pfinh in forall (Ap:Bool_Alg_p T) (pfinas:Inn S Ap), exists (Cp:Bool_Alg_p T), Inn S Cp /\ Inn (ba_p_ens Cp) (proj1_sig one) /\ subalg_of_p Ap Cp /\ proj1_sig (Bone_p T (Bc_p T Ap)) = proj1_sig one. intros S h1 h2 one. unfold directed_one in one. destruct constructive_definite_description as [one' h3]. simpl in one. unfold one. assumption. Qed. Definition directed_bcp {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S) := Build_Bconst_p _ (FamilyUnion (Im S (@ba_p_ens T))) (full_sig (FamilyUnion (Im S (@ba_p_ens T)))) (directed_plus pfd) (directed_times pfd) (directed_one pfd pfinh) (directed_zero pfd pfinh) (directed_comp pfd). Lemma directed_und_set_p : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let Bc_p := directed_bcp pfd pfinh in BS_p T Bc_p = Full_set (Btype_p T Bc_p). intros S h1 h2 Bc_p. unfold Bc_p. simpl. reflexivity. Qed. Lemma directed_assoc_sum_p : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let Bc_p := directed_bcp pfd pfinh in forall n m p:(Btype_p T Bc_p), n %+ (m %+ p) = n %+ m %+ p. intros S h1 h2 Bc_p n m p. pose proof (directed_plus_compat h1 m p) as hmp. pose proof (directed_plus_compat h1 n (m %+ p)) as hn'mp. pose proof (directed_plus_compat h1 n m) as hnm. pose proof (directed_plus_compat h1 (n %+ m) p) as hnm'p. simpl in hmp, hn'mp, hnm, hnm'p. destruct n as [n h3]. simpl in h3. destruct m as [m h4]. simpl in h4. destruct p as [p h9]. simpl in h9. destruct h3 as [N n h5 h6]. destruct h5 as [Np h5]. subst. destruct h4 as [M m h7 h8]. destruct h7 as [Mp h7]. subst. destruct h9 as [P p h10 h11]. destruct h10 as [Pp h10]. subst. simpl. specialize (hnm _ _ h6 h8 h5 h7). simpl in hnm. specialize (hmp _ _ h8 h11 h7 h10). simpl in hmp. destruct hnm as [Ap [h13 [h14 [h15 [h16 [h17 [h18 [h19 [h20 h21]]]]]]]]]. destruct hmp as [Ap' [h13' [h14' [h15' [h16' [h17' [h18' [h19' [h20' h21']]]]]]]]]. specialize (hn'mp _ _ h6 h17' h5 h20'). simpl in hn'mp. specialize (hnm'p _ _ h17 h11 h20 h10). simpl in hnm'p. destruct hn'mp as [Ap'' [h13'' [h14'' [h15'' [h16'' [h17'' [h18'' [h19'' [h20'' h21'']]]]]]]]]. destruct hnm'p as [Ap''' [h13''' [h14''' [h15''' [h16''' [h17''' [h18''' [h19''' [h20''' h21''']]]]]]]]]. simpl in h21, h21', h21'', h21'''. simpl. apply (proj1_sig_injective (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x)). rewrite h21'', h21''' at 1. generalize (h14'' (proj1_sig (directed_plus h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) m (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Mp) m (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Mp h7 (ba_p_ens Mp) eq_refl) h8)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) p (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Pp) p (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Pp h10 (ba_p_ens Pp) eq_refl) h11)))) h17'). intro h22. generalize (h13''' (proj1_sig (directed_plus h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) n (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Np) n (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Np h5 (ba_p_ens Np) eq_refl) h6)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) m (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Mp) m (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Mp h7 (ba_p_ens Mp) eq_refl) h8)))) h17) as h23. intro h23. assert (h24:Inn (ba_p_ens Ap'') (proj1_sig (exist (Inn (ba_p_ens Ap')) m (h13' m h8) %+ exist (Inn (ba_p_ens Ap')) p (h14' p h11)))). rewrite <- h21'. assumption. pose proof (subsetT_eq_compat _ _ _ _ h22 h24 h21') as h25. dependent rewrite -> h25. pose proof (ba_p_subst_plus _ _ h19'' _ _ (h13' m h8) (h14' p h11)) as h26. assert (h27: Inn (ba_p_ens (Subalg_p Ap'' (ba_p_ens Ap') h14'' h16'')) m). rewrite <- h19''. apply (h13' m h8). assert (h28 : Inn (ba_p_ens (Subalg_p Ap'' (ba_p_ens Ap') h14'' h16'')) p). rewrite <- h19''. apply (h14' p h11). specialize (h26 h27 h28). assert (h29:Inn (ba_p_ens Ap'') (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Ap'' (ba_p_ens Ap') h14'' h16''))) m h27 %+ exist (Inn (ba_p_ens (Subalg_p Ap'' (ba_p_ens Ap') h14'' h16''))) p h28))). apply in_sup_alg_p. pose proof (subsetT_eq_compat _ _ _ _ h24 h29 h26) as h30. dependent rewrite -> h30. simpl. rewrite simpl_sig. assert (h31:Inn (ba_p_ens Ap''') (proj1_sig (exist (Inn (ba_p_ens Ap)) n (h13 n h6) %+ exist (Inn (ba_p_ens Ap)) m (h14 m h8)))). rewrite <- h21. assumption. pose proof (subsetT_eq_compat _ _ _ _ h23 h31 h21) as h32. dependent rewrite -> h32. pose proof (ba_p_subst_plus _ _ h18''' _ _ (h13 n h6) (h14 m h8)) as h33. assert (h34:Inn (ba_p_ens (Subalg_p Ap''' (ba_p_ens Ap) h13''' h15''')) n). rewrite <- h18'''. apply (h13 n h6). assert (h35 : Inn (ba_p_ens (Subalg_p Ap''' (ba_p_ens Ap) h13''' h15''')) m). rewrite <- h18'''. apply (h14 m h8). specialize (h33 h34 h35). assert (h36:Inn (ba_p_ens Ap''') (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Ap''' (ba_p_ens Ap) h13''' h15'''))) n h34 %+ exist (Inn (ba_p_ens (Subalg_p Ap''' (ba_p_ens Ap) h13''' h15'''))) m h35))). apply in_sup_alg_p. pose proof (subsetT_eq_compat _ _ _ _ h31 h36 h33) as h37. dependent rewrite -> h37. simpl. rewrite simpl_sig. rewrite assoc_sum_p at 1. pose proof (h1 _ _ h20'' h20''') as h38. destruct h38 as [Dp [h38 [h39 h40]]]. destruct h39 as [h39a [h39b h39c]]. destruct h40 as [h40a [h40b h40c]]. pose proof (ba_p_subst_plus3 _ _ h39c _ _ _ (h13'' n h6) (h14'' m h27) (h14'' p h28)) as h41. assert (h42: Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap'') h39a h39b)) n). rewrite <- h39c. apply (h13'' n h6). assert (h43 : Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap'') h39a h39b)) m). rewrite <- h39c. apply (h14'' m h27). assert (h44 : Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap'') h39a h39b)) p). rewrite <- h39c. apply (h14'' p h28). specialize (h41 h42 h43 h44). rewrite h41 at 1. pose proof (ba_p_subst_plus3 _ _ h40c _ _ _ (h13''' n h34) (h13''' m h35) (h14''' p h11)) as h45. assert (h46: Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap''') h40a h40b)) n). rewrite <- h40c. apply (h13''' n h34). assert (h47 : Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap''') h40a h40b)) m). rewrite <- h40c. apply (h13''' m h35). assert (h48 : Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap''') h40a h40b)) p). rewrite <- h40c. apply (h14''' p h11). specialize (h45 h46 h47 h48). rewrite h45 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma directed_assoc_prod_p : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let Bc_p := directed_bcp pfd pfinh in forall n m p:(Btype_p T Bc_p), n %* (m %* p) = n %* m %* p. intros S h1 h2 Bc_p n m p. pose proof (directed_times_compat h1 m p) as hmp. pose proof (directed_times_compat h1 n (m %* p)) as hn'mp. pose proof (directed_times_compat h1 n m) as hnm. pose proof (directed_times_compat h1 (n %* m) p) as hnm'p. simpl in hmp, hn'mp, hnm, hnm'p. destruct n as [n h3]. simpl in h3. destruct m as [m h4]. simpl in h4. destruct p as [p h9]. simpl in h9. destruct h3 as [N n h5 h6]. destruct h5 as [Np h5]. subst. destruct h4 as [M m h7 h8]. destruct h7 as [Mp h7]. subst. destruct h9 as [P p h10 h11]. destruct h10 as [Pp h10]. subst. simpl. specialize (hnm _ _ h6 h8 h5 h7). simpl in hnm. specialize (hmp _ _ h8 h11 h7 h10). simpl in hmp. destruct hnm as [Ap [h13 [h14 [h15 [h16 [h17 [h18 [h19 [h20 h21]]]]]]]]]. destruct hmp as [Ap' [h13' [h14' [h15' [h16' [h17' [h18' [h19' [h20' h21']]]]]]]]]. specialize (hn'mp _ _ h6 h17' h5 h20'). simpl in hn'mp. specialize (hnm'p _ _ h17 h11 h20 h10). simpl in hnm'p. destruct hn'mp as [Ap'' [h13'' [h14'' [h15'' [h16'' [h17'' [h18'' [h19'' [h20'' h21'']]]]]]]]]. destruct hnm'p as [Ap''' [h13''' [h14''' [h15''' [h16''' [h17''' [h18''' [h19''' [h20''' h21''']]]]]]]]]. simpl in h21, h21', h21'', h21'''. simpl. apply (proj1_sig_injective (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x)). rewrite h21'', h21''' at 1. generalize (h14'' (proj1_sig (directed_times h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) m (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Mp) m (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Mp h7 (ba_p_ens Mp) eq_refl) h8)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) p (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Pp) p (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Pp h10 (ba_p_ens Pp) eq_refl) h11)))) h17'). intro h22. generalize (h13''' (proj1_sig (directed_times h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) n (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Np) n (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Np h5 (ba_p_ens Np) eq_refl) h6)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) m (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Mp) m (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Mp h7 (ba_p_ens Mp) eq_refl) h8)))) h17) as h23. intro h23. assert (h24:Inn (ba_p_ens Ap'') (proj1_sig (exist (Inn (ba_p_ens Ap')) m (h13' m h8) %* exist (Inn (ba_p_ens Ap')) p (h14' p h11)))). rewrite <- h21'. assumption. pose proof (subsetT_eq_compat _ _ _ _ h22 h24 h21') as h25. dependent rewrite -> h25. pose proof (ba_p_subst_times _ _ h19'' _ _ (h13' m h8) (h14' p h11)) as h26. assert (h27: Inn (ba_p_ens (Subalg_p Ap'' (ba_p_ens Ap') h14'' h16'')) m). rewrite <- h19''. apply (h13' m h8). assert (h28 : Inn (ba_p_ens (Subalg_p Ap'' (ba_p_ens Ap') h14'' h16'')) p). rewrite <- h19''. apply (h14' p h11). specialize (h26 h27 h28). assert (h29:Inn (ba_p_ens Ap'') (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Ap'' (ba_p_ens Ap') h14'' h16''))) m h27 %* exist (Inn (ba_p_ens (Subalg_p Ap'' (ba_p_ens Ap') h14'' h16''))) p h28))). apply in_sup_alg_p. pose proof (subsetT_eq_compat _ _ _ _ h24 h29 h26) as h30. dependent rewrite -> h30. simpl. rewrite simpl_sig. assert (h31:Inn (ba_p_ens Ap''') (proj1_sig (exist (Inn (ba_p_ens Ap)) n (h13 n h6) %* exist (Inn (ba_p_ens Ap)) m (h14 m h8)))). rewrite <- h21. assumption. pose proof (subsetT_eq_compat _ _ _ _ h23 h31 h21) as h32. dependent rewrite -> h32. pose proof (ba_p_subst_times _ _ h18''' _ _ (h13 n h6) (h14 m h8)) as h33. assert (h34:Inn (ba_p_ens (Subalg_p Ap''' (ba_p_ens Ap) h13''' h15''')) n). rewrite <- h18'''. apply (h13 n h6). assert (h35 : Inn (ba_p_ens (Subalg_p Ap''' (ba_p_ens Ap) h13''' h15''')) m). rewrite <- h18'''. apply (h14 m h8). specialize (h33 h34 h35). assert (h36:Inn (ba_p_ens Ap''') (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Ap''' (ba_p_ens Ap) h13''' h15'''))) n h34 %* exist (Inn (ba_p_ens (Subalg_p Ap''' (ba_p_ens Ap) h13''' h15'''))) m h35))). apply in_sup_alg_p. pose proof (subsetT_eq_compat _ _ _ _ h31 h36 h33) as h37. dependent rewrite -> h37. simpl. rewrite simpl_sig. rewrite assoc_prod_p at 1. pose proof (h1 _ _ h20'' h20''') as h38. destruct h38 as [Dp [h38 [h39 h40]]]. destruct h39 as [h39a [h39b h39c]]. destruct h40 as [h40a [h40b h40c]]. pose proof (ba_p_subst_times3 _ _ h39c _ _ _ (h13'' n h6) (h14'' m h27) (h14'' p h28)) as h41. assert (h42: Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap'') h39a h39b)) n). rewrite <- h39c. apply (h13'' n h6). assert (h43 : Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap'') h39a h39b)) m). rewrite <- h39c. apply (h14'' m h27). assert (h44 : Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap'') h39a h39b)) p). rewrite <- h39c. apply (h14'' p h28). specialize (h41 h42 h43 h44). rewrite h41 at 1. pose proof (ba_p_subst_times3 _ _ h40c _ _ _ (h13''' n h34) (h13''' m h35) (h14''' p h11)) as h45. assert (h46: Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap''') h40a h40b)) n). rewrite <- h40c. apply (h13''' n h34). assert (h47 : Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap''') h40a h40b)) m). rewrite <- h40c. apply (h13''' m h35). assert (h48 : Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Ap''') h40a h40b)) p). rewrite <- h40c. apply (h14''' p h11). specialize (h45 h46 h47 h48). rewrite h45 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma directed_comm_sum_p : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let Bc_p := directed_bcp pfd pfinh in forall n m:(Btype_p T Bc_p), n %+ m = m %+ n. intros S h1 h2 Bc_p n m. pose proof (directed_plus_compat h1 n m) as hnm. pose proof (directed_plus_compat h1 m n) as hmn. destruct n as [n h3]. simpl in h3. destruct m as [m h4]. simpl in h4. destruct h3 as [N n h5 h6]. destruct h5 as [Np h5]. subst. destruct h4 as [M m h7 h8]. destruct h7 as [Mp h7]. subst. specialize (hnm _ _ h6 h8 h5 h7). simpl in hnm. specialize (hmn _ _ h8 h6 h7 h5). simpl in hmn. destruct hnm as [Ap [h13 [h14 [h15 [h16 [h17 [h18 [h19 [h20 h21]]]]]]]]]. destruct hmn as [Ap' [h13' [h14' [h15' [h16' [h17' [h18' [h19' [h20' h21']]]]]]]]]. simpl in h21, h21'. simpl. apply (proj1_sig_injective (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x)). rewrite h21, h21' at 1. rewrite comm_sum_p. pose proof (h1 _ _ h20 h20') as h22. destruct h22 as [Bp [h22 [h23 h24]]]. destruct h23 as [h23a [h23b h23c]], h24 as [h24a [h24b h24c]]. pose proof (ba_p_subst_plus _ _ h23c _ _ (h14 m h8) (h13 n h6)) as h25. assert (h26 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h23a h23b)) m). rewrite <- h23c. apply (h14 m h8). assert (h27 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h23a h23b)) n). rewrite <- h23c. apply (h13 n h6). specialize (h25 h26 h27). rewrite h25 at 1. pose proof (ba_p_subst_plus _ _ h24c _ _ (h13' m h8) (h14' n h6)) as h28. assert (h29: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap') h24a h24b)) m). rewrite <- h24c. apply (h13' m h8). assert (h30 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap') h24a h24b)) n). rewrite <- h24c. apply (h14' n h6). specialize (h28 h29 h30). rewrite h28 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma directed_comm_prod_p : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let Bc_p := directed_bcp pfd pfinh in forall n m:(Btype_p T Bc_p), n %* m = m %* n. intros S h1 h2 Bc_p n m. pose proof (directed_times_compat h1 n m) as hnm. pose proof (directed_times_compat h1 m n) as hmn. destruct n as [n h3]. simpl in h3. destruct m as [m h4]. simpl in h4. destruct h3 as [N n h5 h6]. destruct h5 as [Np h5]. subst. destruct h4 as [M m h7 h8]. destruct h7 as [Mp h7]. subst. specialize (hnm _ _ h6 h8 h5 h7). simpl in hnm. specialize (hmn _ _ h8 h6 h7 h5). simpl in hmn. destruct hnm as [Ap [h13 [h14 [h15 [h16 [h17 [h18 [h19 [h20 h21]]]]]]]]]. destruct hmn as [Ap' [h13' [h14' [h15' [h16' [h17' [h18' [h19' [h20' h21']]]]]]]]]. simpl in h21, h21'. simpl. apply (proj1_sig_injective (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x)). rewrite h21, h21' at 1. rewrite comm_prod_p. pose proof (h1 _ _ h20 h20') as h22. destruct h22 as [Bp [h22 [h23 h24]]]. destruct h23 as [h23a [h23b h23c]], h24 as [h24a [h24b h24c]]. pose proof (ba_p_subst_times _ _ h23c _ _ (h14 m h8) (h13 n h6)) as h25. assert (h26 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h23a h23b)) m). rewrite <- h23c. apply (h14 m h8). assert (h27 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h23a h23b)) n). rewrite <- h23c. apply (h13 n h6). specialize (h25 h26 h27). rewrite h25 at 1. pose proof (ba_p_subst_times _ _ h24c _ _ (h13' m h8) (h14' n h6)) as h28. assert (h29: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap') h24a h24b)) m). rewrite <- h24c. apply (h13' m h8). assert (h30 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap') h24a h24b)) n). rewrite <- h24c. apply (h14' n h6). specialize (h28 h29 h30). rewrite h28 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma directed_abs_sum_p : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let Bc_p := directed_bcp pfd pfinh in forall n m:(Btype_p T Bc_p), n %+ (n %* m) = n. intros S h1 h2 Bc_p n m. pose proof (directed_times_compat h1 n m) as hnm. pose proof (directed_plus_compat h1 n (n %* m)) as hn'nm. simpl in hnm, hn'nm. destruct n as [n h3]. simpl in h3. destruct m as [m h4]. simpl in h4. destruct h3 as [N n h5 h6]. destruct h5 as [Np h5]. subst. destruct h4 as [M m h7 h8]. destruct h7 as [Mp h7]. subst. simpl. specialize (hnm _ _ h6 h8 h5 h7). simpl in hnm. destruct hnm as [Ap [h13 [h14 [h15 [h16 [h17 [h18 [h19 [h20 h21]]]]]]]]]. specialize (hn'nm _ _ h6 h17 h5 h20). simpl in hn'nm. destruct hn'nm as [Ap' [h13' [h14' [h15' [h16' [h17' [h18' [h19' [h20' h21']]]]]]]]]. simpl in h21, h21'. simpl. apply (proj1_sig_injective (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x)). rewrite h21' at 1. generalize (h14' (proj1_sig (directed_times h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) n (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Np) n (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Np h5 (ba_p_ens Np) eq_refl) h6)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) m (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Mp) m (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Mp h7 (ba_p_ens Mp) eq_refl) h8)))) h17). intro h22. simpl. assert (h23:Inn (ba_p_ens Ap') (proj1_sig (exist (Inn (ba_p_ens Ap)) n (h13 n h6) %* exist (Inn (ba_p_ens Ap)) m (h14 m h8)))). rewrite <- h21. assumption. pose proof (subsetT_eq_compat _ _ _ _ h22 h23 h21) as h24. dependent rewrite -> h24. pose proof (h1 _ _ h20 h20') as h25. destruct h25 as [Bp [h25 [h26 h27]]]. destruct h26 as [h26a [h26b h26c]], h27 as [h27a [h27b h27c]]. pose proof (ba_p_subst_times _ _ h26c _ _ (h13 n h6) (h14 m h8)) as h28. assert (h29 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b)) n). rewrite <- h26c. apply (h13 n h6). assert (h30 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b)) m). rewrite <- h26c. apply (h14 m h8). specialize (h28 h29 h30). assert (h31:Inn (ba_p_ens Ap') (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b))) n h29 %* exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b))) m h30))). rewrite <- h28. assumption. pose proof (subsetT_eq_compat _ _ _ _ h23 h31 h28) as h32. dependent rewrite -> h32. simpl. pose proof (ba_p_subst_plus _ _ h27c _ _ (h13' n h6) h31) as h33. assert (h34: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap') h27a h27b)) n). rewrite <- h27c. apply (h13' n h6). assert (h35 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap') h27a h27b)) (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b))) n h29 %* exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b))) m h30))). rewrite <- h27c. apply h31. specialize (h33 h34 h35). rewrite h33 at 1. simpl. rewrite simpl_sig. assert (h36:h27a n h34 = h26a n h29). apply proof_irrelevance. rewrite h36. rewrite abs_sum_p at 1. simpl. reflexivity. Qed. Lemma directed_abs_prod_p : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let Bc_p := directed_bcp pfd pfinh in forall n m:(Btype_p T Bc_p), n %* (n %+ m) = n. intros S h1 h2 Bc_p n m. pose proof (directed_plus_compat h1 n m) as hnm. pose proof (directed_times_compat h1 n (n %+ m)) as hn'nm. simpl in hnm, hn'nm. destruct n as [n h3]. simpl in h3. destruct m as [m h4]. simpl in h4. destruct h3 as [N n h5 h6]. destruct h5 as [Np h5]. subst. destruct h4 as [M m h7 h8]. destruct h7 as [Mp h7]. subst. simpl. specialize (hnm _ _ h6 h8 h5 h7). simpl in hnm. destruct hnm as [Ap [h13 [h14 [h15 [h16 [h17 [h18 [h19 [h20 h21]]]]]]]]]. specialize (hn'nm _ _ h6 h17 h5 h20). simpl in hn'nm. destruct hn'nm as [Ap' [h13' [h14' [h15' [h16' [h17' [h18' [h19' [h20' h21']]]]]]]]]. simpl in h21, h21'. simpl. apply (proj1_sig_injective (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x)). rewrite h21' at 1. generalize (h14' (proj1_sig (directed_plus h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) n (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Np) n (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Np h5 (ba_p_ens Np) eq_refl) h6)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) m (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Mp) m (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Mp h7 (ba_p_ens Mp) eq_refl) h8)))) h17). intro h22. simpl. assert (h23:Inn (ba_p_ens Ap') (proj1_sig (exist (Inn (ba_p_ens Ap)) n (h13 n h6) %+ exist (Inn (ba_p_ens Ap)) m (h14 m h8)))). rewrite <- h21. assumption. pose proof (subsetT_eq_compat _ _ _ _ h22 h23 h21) as h24. dependent rewrite -> h24. pose proof (h1 _ _ h20 h20') as h25. destruct h25 as [Bp [h25 [h26 h27]]]. destruct h26 as [h26a [h26b h26c]], h27 as [h27a [h27b h27c]]. pose proof (ba_p_subst_plus _ _ h26c _ _ (h13 n h6) (h14 m h8)) as h28. assert (h29 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b)) n). rewrite <- h26c. apply (h13 n h6). assert (h30 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b)) m). rewrite <- h26c. apply (h14 m h8). specialize (h28 h29 h30). assert (h31:Inn (ba_p_ens Ap') (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b))) n h29 %+ exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b))) m h30))). rewrite <- h28. assumption. pose proof (subsetT_eq_compat _ _ _ _ h23 h31 h28) as h32. dependent rewrite -> h32. simpl. pose proof (ba_p_subst_times _ _ h27c _ _ (h13' n h6) h31) as h33. assert (h34: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap') h27a h27b)) n). rewrite <- h27c. apply (h13' n h6). assert (h35 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap') h27a h27b)) (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b))) n h29 %+ exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h26a h26b))) m h30))). rewrite <- h27c. apply h31. specialize (h33 h34 h35). rewrite h33 at 1. simpl. rewrite simpl_sig. assert (h36:h27a n h34 = h26a n h29). apply proof_irrelevance. rewrite h36. rewrite abs_prod_p at 1. simpl. reflexivity. Qed. Lemma directed_dist_sum_p : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let Bc_p := directed_bcp pfd pfinh in forall n m p:(Btype_p T Bc_p), p%*(n %+ m) = p %* n %+ p %* m. intros S h1 h2 Bc_p n m p. pose proof (directed_plus_compat h1 n m) as hnm. pose proof (directed_times_compat h1 p (n %+ m)) as hp'nm. pose proof (directed_times_compat h1 p n) as hpn. pose proof (directed_times_compat h1 p m) as hpm. pose proof (directed_plus_compat h1 (p %* n) (p %* m)) as hpn'pm. simpl in hnm, hp'nm, hpn, hpm, hpn'pm. destruct n as [n h3]. simpl in h3. destruct m as [m h4]. simpl in h4. destruct p as [p h9]. simpl in h9. destruct h3 as [N n h5 h6]. destruct h5 as [Np h5]. subst. destruct h4 as [M m h7 h8]. destruct h7 as [Mp h7]. subst. destruct h9 as [P p h10 h11]. destruct h10 as [Pp h10]. subst. simpl. specialize (hnm _ _ h6 h8 h5 h7). simpl in hnm. specialize (hpn _ _ h11 h6 h10 h5). simpl in hpn. specialize (hpm _ _ h11 h8 h10 h7). simpl in hpm. destruct hnm as [Ap [h13 [h14 [h15 [h16 [h17 [h18 [h19 [h20 h21]]]]]]]]]. destruct hpn as [Ap' [h13' [h14' [h15' [h16' [h17' [h18' [h19' [h20' h21']]]]]]]]]. destruct hpm as [Ap'' [h13'' [h14'' [h15'' [h16'' [h17'' [h18'' [h19'' [h20'' h21'']]]]]]]]]. specialize (hp'nm _ _ h11 h17 h10 h20). simpl in hp'nm. specialize (hpn'pm _ _ h17' h17'' h20' h20''). simpl in hpn'pm. destruct hp'nm as [Ap''' [h13''' [h14''' [h15''' [h16''' [h17''' [h18''' [h19''' [h20''' h21''']]]]]]]]]. destruct hpn'pm as [Ap'''' [h13'''' [h14'''' [h15'''' [h16'''' [h17'''' [h18'''' [h19'''' [h20'''' h21'''']]]]]]]]]. simpl in h21, h21', h21'', h21''', h21''''. simpl. apply (proj1_sig_injective (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x)). rewrite h21''', h21'''' at 1. generalize (h14''' (proj1_sig (directed_plus h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) n (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Np) n (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Np h5 (ba_p_ens Np) eq_refl) h6)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) m (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Mp) m (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Mp h7 (ba_p_ens Mp) eq_refl) h8)))) h17). intro h22. generalize (h13'''' (proj1_sig (directed_times h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) p (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Pp) p (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Pp h10 (ba_p_ens Pp) eq_refl) h11)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) n (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Np) n (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Np h5 (ba_p_ens Np) eq_refl) h6)))) h17'). intro h23. generalize (h14'''' (proj1_sig (directed_times h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) p (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Pp) p (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Pp h10 (ba_p_ens Pp) eq_refl) h11)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) m (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Mp) m (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Mp h7 (ba_p_ens Mp) eq_refl) h8)))) h17''). intro h24. assert (h25:Inn (ba_p_ens Ap''') (proj1_sig (exist (Inn (ba_p_ens Ap)) n (h13 n h6) %+ exist (Inn (ba_p_ens Ap)) m (h14 m h8)))). rewrite <- h21. apply h14'''. assumption. pose proof (subsetT_eq_compat _ _ _ _ h22 h25 h21) as h26. dependent rewrite -> h26. assert (h27:Inn (ba_p_ens Ap'''') (proj1_sig (exist (Inn (ba_p_ens Ap')) p (h13' p h11) %* exist (Inn (ba_p_ens Ap')) n (h14' n h6)))). rewrite <- h21'. apply h13''''. assumption. pose proof (subsetT_eq_compat _ _ _ _ h23 h27 h21') as h28. dependent rewrite -> h28. assert (h29:Inn (ba_p_ens Ap'''') ( proj1_sig (exist (Inn (ba_p_ens Ap'')) p (h13'' p h11) %* exist (Inn (ba_p_ens Ap'')) m (h14'' m h8)))). rewrite <- h21''. apply h14''''. assumption. pose proof (subsetT_eq_compat _ _ _ _ h24 h29 h21'') as h30. dependent rewrite -> h30. pose proof (h1 _ _ h20 h20''') as h31. destruct h31 as [Bp [h32 [h33 h34]]]. destruct h33 as [h33a [h33b h33c]], h34 as [h34a [h34b h34c]]. pose proof (ba_p_subst_times _ _ h34c _ _ (h13''' p h11) h25) as h35. assert (h36:Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap''') h34a h34b)) p). rewrite <- h34c. apply (h13''' p h11). assert (h37 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap''') h34a h34b)) (proj1_sig (exist (Inn (ba_p_ens Ap)) n (h13 n h6) %+ exist (Inn (ba_p_ens Ap)) m (h14 m h8)))). rewrite <- h34c. assumption. specialize (h35 h36 h37). rewrite h35 at 1. pose proof (ba_p_subst_plus _ _ h33c _ _ (h13 n h6) (h14 m h8)) as h38. assert (h39: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h33a h33b)) n). rewrite <- h33c. apply (h13 n h6). assert (h40 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h33a h33b)) m). rewrite <- h33c. apply (h14 m h8). specialize (h38 h39 h40). assert (h41:Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap''') h34a h34b)) (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h33a h33b))) n h39 %+ exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h33a h33b))) m h40))). rewrite <- h38. rewrite <- h34c. assumption. pose proof (subsetT_eq_compat _ _ _ _ h37 h41 h38) as h42. dependent rewrite -> h42. simpl. rewrite simpl_sig. rewrite dist_sum_p. pose proof (h1 _ _ h20' h20'''') as h43. destruct h43 as [Cp [h44 [h45 h46]]]. pose proof (h1 _ _ h20'' h20'''') as h47. destruct h47 as [Dp [h47 [h48 h49]]]. pose proof (h1 _ _ h44 h47) as h50. destruct h50 as [Ep [h51 [h52 h53]]]. pose proof (trans_subalg_of_p _ _ _ _ h45 h52) as h54. pose proof (trans_subalg_of_p _ _ _ _ h48 h53) as h55. pose proof (trans_subalg_of_p _ _ _ _ h46 h52) as h56. destruct h54 as [h54a [h54b h54c]]. destruct h55 as [h55a [h55b h55c]]. destruct h56 as [h56a [h56b h56c]]. pose proof (ba_p_subst_plus _ _ h56c _ _ h27 h29) as h57. assert (ha:Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'''') h56a h56b)) (proj1_sig (exist (Inn (ba_p_ens Ap')) p (h13' p h11) %* exist (Inn (ba_p_ens Ap')) n (h14' n h6)))). rewrite <- h56c. assumption. assert (hb : Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'''') h56a h56b)) (proj1_sig (exist (Inn (ba_p_ens Ap'')) p (h13'' p h11) %* exist (Inn (ba_p_ens Ap'')) m (h14'' m h8)))). rewrite <- h56c. assumption. specialize (h57 ha hb). rewrite h57 at 1. pose proof (ba_p_subst_times _ _ h54c _ _ (h13' p h11) (h14' n h6)) as h58. assert (h59: Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap') h54a h54b)) p). rewrite <- h54c. apply (h13' p h11). assert (h60 : Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap') h54a h54b)) n). rewrite <- h54c. apply (h14' n h6). specialize (h58 h59 h60). simpl. assert (h61:Inn (ba_p_ens Ep) (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap') h54a h54b))) p h59 %* exist (Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap') h54a h54b))) n h60))). apply in_sup_alg_p. pose proof (subsetT_eq_compat _ _ _ _ (h56a (proj1_sig (exist (Inn (ba_p_ens Ap')) p (h13' p h11) %* exist (Inn (ba_p_ens Ap')) n (h14' n h6))) ha) h61 h58) as h62. dependent rewrite -> h62. pose proof (ba_p_subst_times _ _ h55c _ _ (h13'' p h11) (h14'' m h8)) as h63. assert (h64:Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'') h55a h55b)) p). rewrite <- h55c. apply (h13'' p h11). assert (h65 : Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'') h55a h55b)) m). rewrite <- h55c. apply (h14'' m h8). specialize (h63 h64 h65). assert (h66:Inn (ba_p_ens Ep) (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'') h55a h55b))) p h64 %* exist (Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'') h55a h55b))) m h65))). apply in_sup_alg_p. pose proof (subsetT_eq_compat _ _ _ _ (h56a (proj1_sig (exist (Inn (ba_p_ens Ap'')) p (h13'' p h11) %* exist (Inn (ba_p_ens Ap'')) m (h14'' m h8))) hb) h66 h63) as h67. dependent rewrite -> h67. simpl. do 2 rewrite simpl_sig. pose proof (h1 _ _ h32 h51) as h68. destruct h68 as [Fp [h68 [h69 h70]]]. red in h69, h70. destruct h69 as [h69a [h69b h69c]], h70 as [h70a [h70b h70c]]. pose proof (ba_p_subst_times_plus_times _ _ h69c _ _ _ _ (h34a p h36) (h33a n h39) (h34a p h36) (h33a m h40)) as h71. assert (h72:Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Bp) h69a h69b)) p). rewrite <- h69c. apply (h34a p h36). assert (h73 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Bp) h69a h69b)) n). rewrite <- h69c. apply (h33a n h39). assert (h74 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Bp) h69a h69b)) p). rewrite <- h69c. apply (h34a p h36). assert (h75 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Bp) h69a h69b)) m). rewrite <- h69c. apply (h33a m h40). specialize (h71 h72 h73 h74 h75). rewrite h71 at 1. pose proof (ba_p_subst_times_plus_times _ _ h70c _ _ _ _ (h54a p h59) (h54a n h60) (h55a p h64) (h55a m h65)) as h76. assert (h77:Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Ep) h70a h70b)) p). rewrite <- h70c. apply (h54a p h59). assert (h78 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Ep) h70a h70b)) n). rewrite <- h70c. apply (h54a n h60). assert (h79 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Ep) h70a h70b)) p). rewrite <- h70c. apply (h55a p h64). assert (h80 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Ep) h70a h70b)) m). rewrite <- h70c. apply (h55a m h65). specialize (h76 h77 h78 h79 h80). rewrite h76 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma directed_dist_prod_p : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let Bc_p := directed_bcp pfd pfinh in forall n m p:(Btype_p T Bc_p), p%+(n %* m) = (p %+ n) %* (p %+ m). intros S h1 h2 Bc_p n m p. pose proof (directed_times_compat h1 n m) as hnm. pose proof (directed_plus_compat h1 p (n %* m)) as hp'nm. pose proof (directed_plus_compat h1 p n) as hpn. pose proof (directed_plus_compat h1 p m) as hpm. pose proof (directed_times_compat h1 (p %+ n) (p %+ m)) as hpn'pm. simpl in hnm, hp'nm, hpn, hpm, hpn'pm. destruct n as [n h3]. simpl in h3. destruct m as [m h4]. simpl in h4. destruct p as [p h9]. simpl in h9. destruct h3 as [N n h5 h6]. destruct h5 as [Np h5]. subst. destruct h4 as [M m h7 h8]. destruct h7 as [Mp h7]. subst. destruct h9 as [P p h10 h11]. destruct h10 as [Pp h10]. subst. simpl. specialize (hnm _ _ h6 h8 h5 h7). simpl in hnm. specialize (hpn _ _ h11 h6 h10 h5). simpl in hpn. specialize (hpm _ _ h11 h8 h10 h7). simpl in hpm. destruct hnm as [Ap [h13 [h14 [h15 [h16 [h17 [h18 [h19 [h20 h21]]]]]]]]]. destruct hpn as [Ap' [h13' [h14' [h15' [h16' [h17' [h18' [h19' [h20' h21']]]]]]]]]. destruct hpm as [Ap'' [h13'' [h14'' [h15'' [h16'' [h17'' [h18'' [h19'' [h20'' h21'']]]]]]]]]. specialize (hp'nm _ _ h11 h17 h10 h20). simpl in hp'nm. specialize (hpn'pm _ _ h17' h17'' h20' h20''). simpl in hpn'pm. destruct hp'nm as [Ap''' [h13''' [h14''' [h15''' [h16''' [h17''' [h18''' [h19''' [h20''' h21''']]]]]]]]]. destruct hpn'pm as [Ap'''' [h13'''' [h14'''' [h15'''' [h16'''' [h17'''' [h18'''' [h19'''' [h20'''' h21'''']]]]]]]]]. simpl in h21, h21', h21'', h21''', h21''''. simpl. apply (proj1_sig_injective (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x)). rewrite h21''', h21'''' at 1. generalize (h14''' (proj1_sig (directed_times h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) n (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Np) n (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Np h5 (ba_p_ens Np) eq_refl) h6)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) m (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Mp) m (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Mp h7 (ba_p_ens Mp) eq_refl) h8)))) h17). intro h22. generalize (h13'''' (proj1_sig (directed_plus h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) p (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Pp) p (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Pp h10 (ba_p_ens Pp) eq_refl) h11)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) n (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Np) n (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Np h5 (ba_p_ens Np) eq_refl) h6)))) h17'). intro h23. generalize (h14'''' (proj1_sig (directed_plus h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) p (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Pp) p (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Pp h10 (ba_p_ens Pp) eq_refl) h11)) (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) m (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Mp) m (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Mp h7 (ba_p_ens Mp) eq_refl) h8)))) h17''). intro h24. assert (h25:Inn (ba_p_ens Ap''') (proj1_sig (exist (Inn (ba_p_ens Ap)) n (h13 n h6) %* exist (Inn (ba_p_ens Ap)) m (h14 m h8)))). rewrite <- h21. apply h14'''. assumption. pose proof (subsetT_eq_compat _ _ _ _ h22 h25 h21) as h26. dependent rewrite -> h26. assert (h27:Inn (ba_p_ens Ap'''') (proj1_sig (exist (Inn (ba_p_ens Ap')) p (h13' p h11) %+ exist (Inn (ba_p_ens Ap')) n (h14' n h6)))). rewrite <- h21'. apply h13''''. assumption. pose proof (subsetT_eq_compat _ _ _ _ h23 h27 h21') as h28. dependent rewrite -> h28. assert (h29:Inn (ba_p_ens Ap'''') ( proj1_sig (exist (Inn (ba_p_ens Ap'')) p (h13'' p h11) %+ exist (Inn (ba_p_ens Ap'')) m (h14'' m h8)))). rewrite <- h21''. apply h14''''. assumption. pose proof (subsetT_eq_compat _ _ _ _ h24 h29 h21'') as h30. dependent rewrite -> h30. pose proof (h1 _ _ h20 h20''') as h31. destruct h31 as [Bp [h32 [h33 h34]]]. destruct h33 as [h33a [h33b h33c]], h34 as [h34a [h34b h34c]]. pose proof (ba_p_subst_plus _ _ h34c _ _ (h13''' p h11) h25) as h35. assert (h36:Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap''') h34a h34b)) p). rewrite <- h34c. apply (h13''' p h11). assert (h37 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap''') h34a h34b)) (proj1_sig (exist (Inn (ba_p_ens Ap)) n (h13 n h6) %* exist (Inn (ba_p_ens Ap)) m (h14 m h8)))). rewrite <- h34c. assumption. specialize (h35 h36 h37). rewrite h35 at 1. pose proof (ba_p_subst_times _ _ h33c _ _ (h13 n h6) (h14 m h8)) as h38. assert (h39: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h33a h33b)) n). rewrite <- h33c. apply (h13 n h6). assert (h40 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h33a h33b)) m). rewrite <- h33c. apply (h14 m h8). specialize (h38 h39 h40). assert (h41:Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap''') h34a h34b)) (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h33a h33b))) n h39 %* exist (Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ap) h33a h33b))) m h40))). rewrite <- h38. rewrite <- h34c. assumption. pose proof (subsetT_eq_compat _ _ _ _ h37 h41 h38) as h42. dependent rewrite -> h42. simpl. rewrite simpl_sig. rewrite dist_prod_p. pose proof (h1 _ _ h20' h20'''') as h43. destruct h43 as [Cp [h44 [h45 h46]]]. pose proof (h1 _ _ h20'' h20'''') as h47. destruct h47 as [Dp [h47 [h48 h49]]]. pose proof (h1 _ _ h44 h47) as h50. destruct h50 as [Ep [h51 [h52 h53]]]. pose proof (trans_subalg_of_p _ _ _ _ h45 h52) as h54. pose proof (trans_subalg_of_p _ _ _ _ h48 h53) as h55. pose proof (trans_subalg_of_p _ _ _ _ h46 h52) as h56. destruct h54 as [h54a [h54b h54c]]. destruct h55 as [h55a [h55b h55c]]. destruct h56 as [h56a [h56b h56c]]. pose proof (ba_p_subst_times _ _ h56c _ _ h27 h29) as h57. assert (ha:Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'''') h56a h56b)) (proj1_sig (exist (Inn (ba_p_ens Ap')) p (h13' p h11) %+ exist (Inn (ba_p_ens Ap')) n (h14' n h6)))). rewrite <- h56c. assumption. assert (hb : Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'''') h56a h56b)) (proj1_sig (exist (Inn (ba_p_ens Ap'')) p (h13'' p h11) %+ exist (Inn (ba_p_ens Ap'')) m (h14'' m h8)))). rewrite <- h56c. assumption. specialize (h57 ha hb). rewrite h57 at 1. pose proof (ba_p_subst_plus _ _ h54c _ _ (h13' p h11) (h14' n h6)) as h58. assert (h59: Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap') h54a h54b)) p). rewrite <- h54c. apply (h13' p h11). assert (h60 : Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap') h54a h54b)) n). rewrite <- h54c. apply (h14' n h6). specialize (h58 h59 h60). simpl. assert (h61:Inn (ba_p_ens Ep) (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap') h54a h54b))) p h59 %+ exist (Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap') h54a h54b))) n h60))). apply in_sup_alg_p. pose proof (subsetT_eq_compat _ _ _ _ (h56a (proj1_sig (exist (Inn (ba_p_ens Ap')) p (h13' p h11) %+ exist (Inn (ba_p_ens Ap')) n (h14' n h6))) ha) h61 h58) as h62. dependent rewrite -> h62. pose proof (ba_p_subst_plus _ _ h55c _ _ (h13'' p h11) (h14'' m h8)) as h63. assert (h64:Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'') h55a h55b)) p). rewrite <- h55c. apply (h13'' p h11). assert (h65 : Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'') h55a h55b)) m). rewrite <- h55c. apply (h14'' m h8). specialize (h63 h64 h65). assert (h66:Inn (ba_p_ens Ep) (proj1_sig (exist (Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'') h55a h55b))) p h64 %+ exist (Inn (ba_p_ens (Subalg_p Ep (ba_p_ens Ap'') h55a h55b))) m h65))). apply in_sup_alg_p. pose proof (subsetT_eq_compat _ _ _ _ (h56a (proj1_sig (exist (Inn (ba_p_ens Ap'')) p (h13'' p h11) %+ exist (Inn (ba_p_ens Ap'')) m (h14'' m h8))) hb) h66 h63) as h67. dependent rewrite -> h67. simpl. do 2 rewrite simpl_sig. pose proof (h1 _ _ h32 h51) as h68. destruct h68 as [Fp [h68 [h69 h70]]]. red in h69, h70. destruct h69 as [h69a [h69b h69c]], h70 as [h70a [h70b h70c]]. pose proof (ba_p_subst_plus_times_plus _ _ h69c _ _ _ _ (h34a p h36) (h33a n h39) (h34a p h36) (h33a m h40)) as h71. assert (h72:Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Bp) h69a h69b)) p). rewrite <- h69c. apply (h34a p h36). assert (h73 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Bp) h69a h69b)) n). rewrite <- h69c. apply (h33a n h39). assert (h74 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Bp) h69a h69b)) p). rewrite <- h69c. apply (h34a p h36). assert (h75 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Bp) h69a h69b)) m). rewrite <- h69c. apply (h33a m h40). specialize (h71 h72 h73 h74 h75). rewrite h71 at 1. pose proof (ba_p_subst_plus_times_plus _ _ h70c _ _ _ _ (h54a p h59) (h54a n h60) (h55a p h64) (h55a m h65)) as h76. assert (h77:Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Ep) h70a h70b)) p). rewrite <- h70c. apply (h54a p h59). assert (h78 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Ep) h70a h70b)) n). rewrite <- h70c. apply (h54a n h60). assert (h79 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Ep) h70a h70b)) p). rewrite <- h70c. apply (h55a p h64). assert (h80 : Inn (ba_p_ens (Subalg_p Fp (ba_p_ens Ep) h70a h70b)) m). rewrite <- h70c. apply (h55a m h65). specialize (h76 h77 h78 h79 h80). rewrite h76 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma directed_comp_sum_p : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let Bc_p := directed_bcp pfd pfinh in forall n:(Btype_p T Bc_p), n %+ (%- n) = %1. intros S h1 h2 Bc_p n. pose proof (directed_plus_compat h1 n (%- n)) as hn. pose proof (directed_comp_compat h1 n) as hn'. pose proof (directed_one_compat h1 h2) as ho. destruct n as [n h3]. simpl in h3. destruct h3 as [N n h5 h6]. destruct h5 as [Np h5]. subst. simpl in hn, hn'. specialize (hn' _ h6 h5). destruct hn' as [Cp [h7 [h8 [h9 h10]]]]. apply proj1_sig_injective. simpl. specialize (hn _ _ h6 h7 h5 h9). destruct hn as [Ap [h13 [h14 [h15 [h16 [h17 [h18 [h19 [h20 h21]]]]]]]]]. rewrite h21 at 1. generalize (h14 (proj1_sig (directed_comp h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) n (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Np) n (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Np h5 (ba_p_ens Np) eq_refl) h6)))) h7). intro h22. assert (h23:Inn (ba_p_ens Ap) (proj1_sig (%-exist (Inn (ba_p_ens Np)) n h6))). rewrite <- h10. assumption. pose proof (subsetT_eq_compat _ _ _ _ h22 h23 h10) as h24. dependent rewrite -> h24. pose proof (ba_p_subst_comp _ _ h18 _ h6) as h25. assert (h26:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Np) h13 h15)) n). rewrite <- h18. assumption. specialize (h25 h26). assert (h27:Inn (ba_p_ens Ap) (proj1_sig (%-exist (Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Np) h13 h15))) n h26))). apply in_sup_alg_p. pose proof (subsetT_eq_compat _ _ _ _ h23 h27 h25) as h28. dependent rewrite -> h28. simpl. rewrite simpl_sig. assert (h29:h13 n h6 = h13 n h26). apply proof_irrelevance. rewrite h29. rewrite comp_sum_p at 1. specialize (ho Ap h20). destruct ho as [Dp [ho1 [ho2 [ho3 ho4]]]]. rewrite <- ho4 at 1. reflexivity. Qed. Lemma directed_comp_prod_p : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S), let Bc_p := directed_bcp pfd pfinh in forall n:(Btype_p T Bc_p), n %* (%- n) = %0. intros S h1 h2 Bc_p n. pose proof (directed_times_compat h1 n (%- n)) as hn. pose proof (directed_comp_compat h1 n) as hn'. pose proof (directed_zero_compat h1 h2) as ho. destruct n as [n h3]. simpl in h3. destruct h3 as [N n h5 h6]. destruct h5 as [Np h5]. subst. simpl in hn, hn'. specialize (hn' _ h6 h5). destruct hn' as [Cp [h7 [h8 [h9 h10]]]]. apply proj1_sig_injective. simpl. specialize (hn _ _ h6 h7 h5 h9). destruct hn as [Ap [h13 [h14 [h15 [h16 [h17 [h18 [h19 [h20 h21]]]]]]]]]. rewrite h21 at 1. generalize (h14 (proj1_sig (directed_comp h1 (exist (fun x : T => Inn (FamilyUnion (Im S (ba_p_ens (T:=T)))) x) n (family_union_intro T (Im S (ba_p_ens (T:=T))) (ba_p_ens Np) n (Im_intro (Bool_Alg_p T) (Ensemble T) S (ba_p_ens (T:=T)) Np h5 (ba_p_ens Np) eq_refl) h6)))) h7). intro h22. assert (h23:Inn (ba_p_ens Ap) (proj1_sig (%-exist (Inn (ba_p_ens Np)) n h6))). rewrite <- h10. assumption. pose proof (subsetT_eq_compat _ _ _ _ h22 h23 h10) as h24. dependent rewrite -> h24. pose proof (ba_p_subst_comp _ _ h18 _ h6) as h25. assert (h26:Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Np) h13 h15)) n). rewrite <- h18. assumption. specialize (h25 h26). assert (h27:Inn (ba_p_ens Ap) (proj1_sig (%-exist (Inn (ba_p_ens (Subalg_p Ap (ba_p_ens Np) h13 h15))) n h26))). apply in_sup_alg_p. pose proof (subsetT_eq_compat _ _ _ _ h23 h27 h25) as h28. dependent rewrite -> h28. simpl. rewrite simpl_sig. assert (h29:h13 n h6 = h13 n h26). apply proof_irrelevance. rewrite h29. rewrite comp_prod_p at 1. specialize (ho Ap h20). destruct ho as [Dp [ho1 [ho2 [ho3 ho4]]]]. rewrite <- ho4 at 1. reflexivity. Qed. Definition directed_ba_p {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S) := Build_Bool_Alg_p T (directed_bcp pfd pfinh) (directed_und_set_p pfd pfinh) (directed_assoc_sum_p pfd pfinh) (directed_assoc_prod_p pfd pfinh) (directed_comm_sum_p pfd pfinh) (directed_comm_prod_p pfd pfinh) (directed_abs_sum_p pfd pfinh) (directed_abs_prod_p pfd pfinh) (directed_dist_sum_p pfd pfinh) (directed_dist_prod_p pfd pfinh) (directed_comp_sum_p pfd pfinh) (directed_comp_prod_p pfd pfinh). Lemma directed_ba_p_subalg : forall {S:fam_ba_p T} (pfd:directed S) (pfinh:Inhabited S) (Bp:Bool_Alg_p T), Inn S Bp -> subalg_of_p Bp (directed_ba_p pfd pfinh). intros S h1 h2 Bp h3. assert (h4: Included (A_p T (Bc_p T Bp)) (ba_p_ens (directed_ba_p h1 h2))). unfold directed_ba_p. unfold ba_p_ens. simpl. red. intros x h4. apply family_union_intro with (ba_p_ens Bp). apply Im_intro with Bp. assumption. reflexivity. assumption. exists h4. assert (h5:alg_closed_p (A_p T (Bc_p T Bp)) h4). assert (h5:Comp_closed_sub_p _ _ _ h4). red. intro z. destruct z as [z h5]. unfold directed_ba_p. unfold Bcomp_sub_p. simpl. pose proof (directed_comp_compat h1 (exist (Inn (FamilyUnion (Im S (ba_p_ens (T:=T))))) z (h4 z h5))) as h6. simpl in h6. specialize (h6 _ h5 h3). destruct h6 as [Cp [h8 [h9 [h10 h11]]]]. rewrite h11. apply in_ba_p_ens_comp. assert (h6:Plus_closed_sub_p _ _ _ h4). red. intros x y. destruct x as [x h6], y as [y h7]. unfold directed_ba_p. unfold Bplus_sub_p. simpl. pose proof (directed_plus_compat h1 (exist (Inn (FamilyUnion (Im S (ba_p_ens (T:=T))))) x (h4 x h6)) (exist (Inn (FamilyUnion (Im S (ba_p_ens (T:=T))))) y (h4 y h7))) as h8. simpl in h8. specialize (h8 _ _ h6 h7 h3 h3). destruct h8 as [Cp [h8 [h9 [h10 [h11 [h12 [h13 [h14 [h15 h16]]]]]]]]]. rewrite h16. pose ((exist _ _ h6) %+ (exist _ _ h7)) as z. pose proof (proj2_sig z) as h17. simpl in h17. unfold z in h17. pose proof (ba_p_subst_plus _ _ h13 _ _ h6 h7) as h18. assert (h19: Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h8 h10)) x). rewrite <- h13. assumption. assert (h20 : Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h8 h10)) y). rewrite <- h13. assumption. specialize (h18 h19 h20). rewrite h18 in h17 at 1. simpl in h17. assert (h21:h19 = h6). apply proof_irrelevance. assert (h22:h20 = h7). apply proof_irrelevance. assert (h23:h8 = h9). apply proof_irrelevance. rewrite h21, h22, h23 in h17. rewrite h23. assumption. constructor; auto. apply two_ops_imp_times_p; auto. red. unfold directed_ba_p. simpl. pose proof (directed_one_compat h1 h2) as h7. simpl in h7. specialize (h7 _ h3). destruct h7 as [Cp [h8 [h9 [h10 h11]]]]. rewrite <- h11. apply in_ba_p_ens_one. red. unfold directed_ba_p. simpl. pose proof (directed_zero_compat h1 h2) as h7. simpl in h7. specialize (h7 _ h3). destruct h7 as [Cp [h8 [h9 [h10 h11]]]]. rewrite <- h11. apply in_ba_p_ens_zero. exists h5. apply bc_inj_p. simpl. assert (h6:A_p T (Bc_p T Bp) = A_p T (Bc_p' T (directed_ba_p h1 h2) (ba_p_ens Bp) h4 h5)). simpl. reflexivity. assert (h7: (sig_set_eq (A_p T (Bc_p T Bp)) (ba_p_ens Bp) h6) = eq_refl). apply proof_irrelevance. apply (bconst_ext_p _ _ h6); simpl. unfold SubBtype_p. rewrite h7. rewrite transfer_dep_r_eq_refl. rewrite und_set_p. simpl. unfold Btype_p. reflexivity. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite transfer_dep_r_fun2_eq. simpl. rewrite h7. rewrite transfer_r_eq_refl. destruct x as [x h8], y as [y h9]. simpl. pose proof (directed_plus_compat h1 (exist (Inn (FamilyUnion (Im S (ba_p_ens (T:=T))))) x (h4 x h8)) (exist (Inn (FamilyUnion (Im S (ba_p_ens (T:=T))))) y (h4 y h9))) as h10. simpl in h10. specialize (h10 _ _ h8 h9 h3 h3). destruct h10 as [Cp [h8' [h9' [h10' [h11' [h12' [h13' [h14' [h15' h16']]]]]]]]]. apply proj1_sig_injective. simpl. rewrite h16'. pose proof (ba_p_subst_plus _ _ h13' _ _ h8 h9) as h17. assert (h18:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h8' h10')) x). rewrite <- h13'; auto. assert (h19 : Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h8' h10')) y). rewrite <- h13'; auto. specialize (h17 h18 h19). simpl in h17. rewrite h17 at 1. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite transfer_dep_r_fun2_eq. simpl. rewrite h7. rewrite transfer_r_eq_refl. destruct x as [x h8], y as [y h9]. simpl. pose proof (directed_times_compat h1 (exist (Inn (FamilyUnion (Im S (ba_p_ens (T:=T))))) x (h4 x h8)) (exist (Inn (FamilyUnion (Im S (ba_p_ens (T:=T))))) y (h4 y h9))) as h10. simpl in h10. specialize (h10 _ _ h8 h9 h3 h3). destruct h10 as [Cp [h8' [h9' [h10' [h11' [h12' [h13' [h14' [h15' h16']]]]]]]]]. apply proj1_sig_injective. simpl. rewrite h16'. pose proof (ba_p_subst_times _ _ h13' _ _ h8 h9) as h17. assert (h18:Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h8' h10')) x). rewrite <- h13'; auto. assert (h19 : Inn (ba_p_ens (Subalg_p Cp (ba_p_ens Bp) h8' h10')) y). rewrite <- h13'; auto. specialize (h17 h18 h19). simpl in h17. rewrite h17 at 1. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. rewrite h7. rewrite transfer_dep_r_eq_refl. apply proj1_sig_injective. simpl. pose proof (directed_one_compat h1 h2 _ h3) as h8. simpl in h8. destruct h8 as [Cp [h8 [h9 [h10 h11]]]]. assumption. rewrite h7. rewrite transfer_dep_r_eq_refl. apply proj1_sig_injective. simpl. pose proof (directed_zero_compat h1 h2 _ h3) as h8. simpl in h8. destruct h8 as [Cp [h8 [h9 [h10 h11]]]]. assumption. apply functional_extensionality. intro x. rewrite transfer_dep_r_fun1_eq. simpl. rewrite h7. rewrite transfer_r_eq_refl. destruct x as [x h8]. simpl. pose proof (directed_comp_compat h1 (exist (Inn (FamilyUnion (Im S (ba_p_ens (T:=T))))) x (h4 x h8))) as h10. simpl in h10. specialize (h10 _ h8 h3). destruct h10 as [Cp [h8' [h9' [h10' h11']]]]. apply proj1_sig_injective. simpl. rewrite h11'. f_equal. Qed. End DirectedFamilyOfAlgebras. Section DirectedFamilyOfAlgebras'. Lemma incl_ba_p_ens_directed_ba_p : forall {T:Type} {S:fam_ba_p T} (pfd:directed _ S) (pfinh:Inhabited S) (Bp:Bool_Alg_p T), (forall Cp:Bool_Alg_p T, Inn S Cp -> subalg_of_p Cp Bp) -> Included (ba_p_ens (directed_ba_p _ pfd pfinh)) (ba_p_ens Bp). intros T S pfd pfinh Bp h1. red. intros x h2. destruct h2 as [E x h2 h3]. destruct h2 as [Cp h2]. subst. specialize (h1 Cp h2). destruct h1 as [h1a [h1b h1c]]. apply h1a; auto. Qed. Lemma alg_closed_p_ba_p_ens_directed_ba_p : forall {T:Type} {S:fam_ba_p T} (pfd:directed _ S) (pfinh:Inhabited S) (Bp:Bool_Alg_p T) (pff: forall Cp:Bool_Alg_p T, Inn S Cp -> subalg_of_p Cp Bp), alg_closed_p (ba_p_ens (directed_ba_p _ pfd pfinh)) (incl_ba_p_ens_directed_ba_p pfd pfinh Bp pff). intros T S pfd pfinh Bp h1. constructor; red; simpl. intros x y. destruct x as [x h2], y as [y h3]. simpl. pose proof h2 as h2'. pose proof h3 as h3'. destruct h2 as [X x h2 hx], h3 as [Y y h3 hy]. destruct h2 as [BX h2 X h2a], h3 as [BY h3 Y h3a]. subst. unfold ba_p_ens, directed_ba_p. simpl. pose proof (h1 _ h2) as h4. pose proof (h1 _ h3) as h5. simpl. pose proof (directed_plus_compat _ pfd (exist _ _ h2') (exist _ _ h3') _ _ hx hy h2 h3) as h6. simpl in h6. destruct h6 as [Cp [h8 [h9 [h10 [h11 [h12 [h13 [h14 [h15 h16]]]]]]]]]. destruct h4 as [h4a [h4b h4c]], h5 as [h5a [h5b h5c]]. apply family_union_intro with (ba_p_ens Cp); auto. apply Im_intro with Cp; auto. pose proof (h1 _ h15) as h17. red in h17. destruct h17 as [h17a [h17b h17c]]. assert (h18: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h17a h17b)) x). rewrite <- h17c. apply h8; auto. assert (h19 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h17a h17b)) y). rewrite <- h17c. apply h9; auto. pose proof (ba_p_subst_plus _ _ h17c _ _ (h8 x hx) (h9 y hy) h18 h19) as h20. simpl in h20. assert (h21:exist _ _ (h17a x h18) = exist _ _ (incl_ba_p_ens_directed_ba_p pfd pfinh Bp h1 x (family_union_intro T (Im S (fun Bp0 : Bool_Alg_p T => A_p T (Bc_p T Bp0))) (A_p T (Bc_p T BX)) x (Im_intro (Bool_Alg_p T) (Ensemble T) S (fun Bp0 : Bool_Alg_p T => A_p T (Bc_p T Bp0)) BX h2 (A_p T (Bc_p T BX)) eq_refl) hx))). apply proj1_sig_injective; auto. assert (h22:exist _ _ (h17a y h19) = exist _ _ (incl_ba_p_ens_directed_ba_p pfd pfinh Bp h1 y (family_union_intro T (Im S (fun Bp0 : Bool_Alg_p T => A_p T (Bc_p T Bp0))) (A_p T (Bc_p T BY)) y (Im_intro (Bool_Alg_p T) (Ensemble T) S (fun Bp0 : Bool_Alg_p T => A_p T (Bc_p T Bp0)) BY h3 (A_p T (Bc_p T BY)) eq_refl) hy))). apply proj1_sig_injective; auto. rewrite <- h21 at 1. rewrite <- h22 at 1. unfold ba_p_ens. rewrite <- h20. apply proj2_sig. intros x y. destruct x as [x h2], y as [y h3]. simpl. pose proof h2 as h2'. pose proof h3 as h3'. destruct h2 as [X x h2 hx], h3 as [Y y h3 hy]. destruct h2 as [BX h2 X h2a], h3 as [BY h3 Y h3a]. subst. unfold ba_p_ens, directed_ba_p. simpl. pose proof (h1 _ h2) as h4. pose proof (h1 _ h3) as h5. simpl. pose proof (directed_plus_compat _ pfd (exist _ _ h2') (exist _ _ h3') _ _ hx hy h2 h3) as h6. simpl in h6. destruct h6 as [Cp [h8 [h9 [h10 [h11 [h12 [h13 [h14 [h15 h16]]]]]]]]]. destruct h4 as [h4a [h4b h4c]], h5 as [h5a [h5b h5c]]. apply family_union_intro with (ba_p_ens Cp); auto. apply Im_intro with Cp; auto. pose proof (h1 _ h15) as h17. red in h17. destruct h17 as [h17a [h17b h17c]]. assert (h18: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h17a h17b)) x). rewrite <- h17c. apply h8; auto. assert (h19 : Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h17a h17b)) y). rewrite <- h17c. apply h9; auto. pose proof (ba_p_subst_times _ _ h17c _ _ (h8 x hx) (h9 y hy) h18 h19) as h20. simpl in h20. assert (h21:exist _ _ (h17a x h18) = exist _ _ (incl_ba_p_ens_directed_ba_p pfd pfinh Bp h1 x (family_union_intro T (Im S (fun Bp0 : Bool_Alg_p T => A_p T (Bc_p T Bp0))) (A_p T (Bc_p T BX)) x (Im_intro (Bool_Alg_p T) (Ensemble T) S (fun Bp0 : Bool_Alg_p T => A_p T (Bc_p T Bp0)) BX h2 (A_p T (Bc_p T BX)) eq_refl) hx))). apply proj1_sig_injective; auto. assert (h22:exist _ _ (h17a y h19) = exist _ _ (incl_ba_p_ens_directed_ba_p pfd pfinh Bp h1 y (family_union_intro T (Im S (fun Bp0 : Bool_Alg_p T => A_p T (Bc_p T Bp0))) (A_p T (Bc_p T BY)) y (Im_intro (Bool_Alg_p T) (Ensemble T) S (fun Bp0 : Bool_Alg_p T => A_p T (Bc_p T Bp0)) BY h3 (A_p T (Bc_p T BY)) eq_refl) hy))). apply proj1_sig_injective; auto. rewrite <- h21 at 1. rewrite <- h22 at 1. unfold ba_p_ens. rewrite <- h20. apply proj2_sig. pose proof pfinh as pfinh'. destruct pfinh as [Cp h0]. pose proof (h1 _ h0) as h2. destruct h2 as [h2 [h3 h4]]. pose proof (ba_p_subst_one _ _ h4) as h5. rewrite <- h5 at 1. unfold ba_p_ens, directed_ba_p. simpl. apply family_union_intro with (ba_p_ens Cp); auto. apply Im_intro with Cp; auto. apply proj2_sig. pose proof pfinh as pfinh'. destruct pfinh as [Cp h0]. pose proof (h1 _ h0) as h2. destruct h2 as [h2 [h3 h4]]. pose proof (ba_p_subst_zero _ _ h4) as h5. rewrite <- h5 at 1. unfold ba_p_ens, directed_ba_p. simpl. apply family_union_intro with (ba_p_ens Cp); auto. apply Im_intro with Cp; auto. apply proj2_sig. intros x. destruct x as [x h2]. simpl. pose proof h2 as h2'. destruct h2 as [X x h2 hx]. destruct h2 as [BX h2 X h2a]. subst. unfold ba_p_ens, directed_ba_p. simpl. apply family_union_intro with (ba_p_ens BX). apply Im_intro with BX; auto. pose proof (h1 _ h2) as h4. red in h4. destruct h4 as [h4a [h4b h4c]]. assert (h18: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens BX) h4a h4b)) x). rewrite <- h4c; auto. pose proof (ba_p_subst_comp _ _ h4c _ hx h18) as h19. simpl in h19. assert (h20:exist _ _ (h4a x h18) = exist _ _ (incl_ba_p_ens_directed_ba_p pfd pfinh Bp h1 x (family_union_intro T (Im S (fun Bp0 : Bool_Alg_p T => A_p T (Bc_p T Bp0))) (A_p T (Bc_p T BX)) x (Im_intro (Bool_Alg_p T) (Ensemble T) S (fun Bp0 : Bool_Alg_p T => A_p T (Bc_p T Bp0)) BX h2 (A_p T (Bc_p T BX)) eq_refl) hx))). apply proj1_sig_injective; auto. rewrite h20 in h19 at 1. rewrite <- h19 at 1. apply proj2_sig. Qed. Lemma subalg_of_p_directed_ba_p : forall {T:Type} {Bp:Bool_Alg_p T} (S:fam_ba_p T) (pfd:directed _ S) (pfinh:Inhabited S), (forall Cp:Bool_Alg_p T, Inn S Cp -> subalg_of_p Cp Bp) -> subalg_of_p (directed_ba_p T pfd pfinh) Bp. intros T Bp S h1 h2 h3. pose proof (alg_closed_p_ba_p_ens_directed_ba_p h1 h2 _ h3) as h4. red. exists (incl_ba_p_ens_directed_ba_p h1 h2 Bp h3). exists h4. apply bc_inj_p. assert (h6:A_p T (Bc_p T (directed_ba_p T h1 h2)) = A_p T (Bc_p T (Subalg_p Bp (ba_p_ens (directed_ba_p T h1 h2)) (incl_ba_p_ens_directed_ba_p h1 h2 Bp h3) h4))). simpl. unfold directed_ba_p, ba_p_ens. simpl. f_equal. pose proof (bconst_ext_p (Bc_p T (directed_ba_p T h1 h2)) (Bc_p T (Subalg_p Bp (ba_p_ens (directed_ba_p T h1 h2)) (incl_ba_p_ens_directed_ba_p h1 h2 Bp h3) h4)) h6) as h5. apply h5; simpl. clear h5. rewrite <- transfer_dep_r_transfer_dep_r_sig_set_eq_compat. apply Extensionality_Ensembles. red; split; red; intros x h7. pose proof (transfer_in' (sig_set_eq _ _ h6) (Full_set (SubBtype_p T (ba_p_ens (directed_ba_p T h1 h2)))) x) as h8. pose proof (transfer_sig_set_eq_in _ _ h6 (Full_set (SubBtype_p T (ba_p_ens (directed_ba_p T h1 h2)))) x) as h9. rewrite <- h9 at 1. constructor. constructor. pose proof (transfer_fun2_r_transfer_dep_r_compat' (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6) (fun x0 y0 : SubBtype_p T (ba_p_ens (directed_ba_p T h1 h2)) => exist (Inn (ba_p_ens (directed_ba_p T h1 h2))) (proj1_sig (Bplus_sub_p T Bp (ba_p_ens (directed_ba_p T h1 h2)) (incl_ba_p_ens_directed_ba_p h1 h2 Bp h3) x0 y0)) (P_c_p T Bp (ba_p_ens (directed_ba_p T h1 h2)) (incl_ba_p_ens_directed_ba_p h1 h2 Bp h3) h4 x0 y0))) as h10. rewrite <- h10 at 1. apply functional_extensionality. intros x. apply functional_extensionality. intro y. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6) x) at 2. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6) y) at 2. rewrite transfer_fun2_r_compat'. rewrite (transfer_r_sig_set_eq _ _ h6 (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6)). simpl. apply proj1_sig_injective; simpl. clear h5 h10. do 2 rewrite (transfer_sig_set_eq _ _ h6). simpl. pose proof (directed_plus_compat T h1 x y) as h12. destruct x as [x h7], y as [y h8]. simpl. destruct h7 as [C x h7], h8 as [D y h8]. destruct h7 as [Cp h7 ? h10], h8 as [Dp h8 ? h11]. subst. simpl in h12. specialize (h12 _ _ i i0 h7 h8). destruct h12 as [Ep [h15 [h16 [h17 [h18 [h19 [h20 [h21 [h22 h23]]]]]]]]]. rewrite h23 at 1. pose proof (h3 _ h22) as h24. destruct h24 as [h24a [h24b h24c]]. assert (h26: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ep) h24a h24b)) x). rewrite <- h24c. apply h15; auto. assert (h27: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ep) h24a h24b)) y). rewrite <- h24c. apply h16; auto. pose proof (ba_p_subst_plus _ _ h24c _ _ (h15 x i) (h16 y i0) h26 h27) as h25. rewrite h25 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. pose proof (transfer_fun2_r_transfer_dep_r_compat' (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6) (fun x0 y0 : SubBtype_p T (ba_p_ens (directed_ba_p T h1 h2)) => exist (Inn (ba_p_ens (directed_ba_p T h1 h2))) (proj1_sig (Btimes_sub_p T Bp (ba_p_ens (directed_ba_p T h1 h2)) (incl_ba_p_ens_directed_ba_p h1 h2 Bp h3) x0 y0)) (T_c_p T Bp (ba_p_ens (directed_ba_p T h1 h2)) (incl_ba_p_ens_directed_ba_p h1 h2 Bp h3) h4 x0 y0))) as h10. rewrite <- h10 at 1. apply functional_extensionality. intros x. apply functional_extensionality. intro y. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6) x) at 2. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6) y) at 2. rewrite transfer_fun2_r_compat'. rewrite (transfer_r_sig_set_eq _ _ h6 (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6)). simpl. apply proj1_sig_injective; simpl. clear h5 h10. do 2 rewrite (transfer_sig_set_eq _ _ h6). simpl. pose proof (directed_times_compat T h1 x y) as h12. destruct x as [x h7], y as [y h8]. simpl. destruct h7 as [C x h7], h8 as [D y h8]. destruct h7 as [Cp h7 ? h10], h8 as [Dp h8 ? h11]. subst. simpl in h12. specialize (h12 _ _ i i0 h7 h8). destruct h12 as [Ep [h15 [h16 [h17 [h18 [h19 [h20 [h21 [h22 h23]]]]]]]]]. rewrite h23 at 1. pose proof (h3 _ h22) as h24. destruct h24 as [h24a [h24b h24c]]. assert (h26: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ep) h24a h24b)) x). rewrite <- h24c. apply h15; auto. assert (h27: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Ep) h24a h24b)) y). rewrite <- h24c. apply h16; auto. pose proof (ba_p_subst_times _ _ h24c _ _ (h15 x i) (h16 y i0) h26 h27) as h25. rewrite h25 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. clear h5. rewrite transfer_dep_r_transfer_r_compat. assert (h7:(f_equal id (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6)) = (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6)). apply proof_irrelevance. pose proof (O_c_p T Bp (ba_p_ens (directed_ba_p T h1 h2)) (incl_ba_p_ens_directed_ba_p h1 h2 Bp h3) h4) as h8. red in h8. pose proof (transfer_r_sig_set_eq _ _ h6 (sig_set_eq _ _ h6) (exist _ _ h8)) as h9. simpl in h9. rewrite <- h7 in h9. assert (h10:h8 = (O_c_p T Bp (ba_p_ens (directed_ba_p T h1 h2)) (incl_ba_p_ens_directed_ba_p h1 h2 Bp h3) h4)). apply proof_irrelevance. subst. rewrite h9 at 1. apply proj1_sig_injective. simpl. pose proof (directed_one_compat T h1 h2) as h10. destruct h2 as [Ap h2]. specialize (h10 _ h2). destruct h10 as [Cp [h11 [h12 [h13 h14]]]]. rewrite <- h14 at 1. red in h13. destruct h13 as [h13a [h13b h13c]]. pose proof (ba_p_subst_one _ _ h13c) as h15. rewrite h15 at 1. simpl. pose proof (h3 _ h11) as h16. destruct h16 as [h16 [h17 h18]]. pose proof (ba_p_subst_one _ _ h18) as h19. rewrite h19. reflexivity. rewrite transfer_dep_r_transfer_r_compat. assert (h7:(f_equal id (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6)) = (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6)). apply proof_irrelevance. pose proof (Z_c_p T Bp (ba_p_ens (directed_ba_p T h1 h2)) (incl_ba_p_ens_directed_ba_p h1 h2 Bp h3) h4) as h8. red in h8. pose proof (transfer_r_sig_set_eq _ _ h6 (sig_set_eq _ _ h6) (exist _ _ h8)) as h9. simpl in h9. rewrite <- h7 in h9. assert (h10:h8 = (Z_c_p T Bp (ba_p_ens (directed_ba_p T h1 h2)) (incl_ba_p_ens_directed_ba_p h1 h2 Bp h3) h4)). apply proof_irrelevance. subst. rewrite h9 at 1. apply proj1_sig_injective. simpl. pose proof (directed_zero_compat T h1 h2) as h10. destruct h2 as [Ap h2]. specialize (h10 _ h2). destruct h10 as [Cp [h11 [h12 [h13 h14]]]]. rewrite <- h14 at 1. red in h13. destruct h13 as [h13a [h13b h13c]]. pose proof (ba_p_subst_zero _ _ h13c) as h15. rewrite h15 at 1. simpl. pose proof (h3 _ h11) as h16. destruct h16 as [h16 [h17 h18]]. pose proof (ba_p_subst_zero _ _ h18) as h19. rewrite h19. reflexivity. pose proof (transfer_fun_r_transfer_dep_r_compat' (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6)) as h10. rewrite <- h10 at 1. apply functional_extensionality. intros x. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6) x) at 2. rewrite transfer_fun_r_compat'. rewrite (transfer_r_sig_set_eq _ _ h6 (sig_set_eq (FamilyUnion (Im S (ba_p_ens (T:=T)))) (ba_p_ens (directed_ba_p T h1 h2)) h6)). simpl. apply proj1_sig_injective; simpl. clear h5 h10. rewrite (transfer_sig_set_eq _ _ h6). simpl. pose proof (directed_comp_compat T h1 x) as h12. destruct x as [x h7]. simpl. destruct h7 as [C x h7]. destruct h7 as [Cp h7 ? h10]. subst. simpl in h12. specialize (h12 _ i h7). destruct h12 as [Ep [h15 [h16 [h17 h18]]]]. rewrite h18 at 1. pose proof (h3 _ h7) as h24. destruct h24 as [h24a [h24b h24c]]. assert (h26: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h24a h24b)) x). rewrite <- h24c. assumption. pose proof (ba_p_subst_comp _ _ h24c x i h26) as h25. rewrite h25 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. Qed. End DirectedFamilyOfAlgebras'. Arguments directed [T] _. Arguments directed_bcp [T] [S] _ _. Arguments directed_ba_p [T] [S] _ _.bool2-v0-3/FunctionAlgebrasBasics.v0000644000175000017500000006244313575574055020115 0ustar dwyckoff76dwyckoff76(*Copyright (C) 2019, Daniel Wyckoff*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import EqDec. Require Import NaryFunctions. Require Import Equality. Require Import SetUtilities. Require Import SetUtilities1_5. Require Import TypeUtilities. Require Import TypeUtilities2. Require Import ListUtilities. Require Import FunctionalExtensionality. Require Import LogicUtilities. Require Import ArithUtilities. Require Import FunctionProperties. Require Import FunctionProperties2. Require Import Permutations. (*Version 0.3 bool2 -- certified -- Version 8.4 pl4 Coq.*) (*This brief stray into function algebras follows the notation and text of Dietlinde Lau's "Function Algbebras on Finite Sets" closely, just as the Boolean portion follows Steven Givant's "Introduction to Boolean Algebras.", and the current focus, Partial Algebras, follows Reichel's "Structural Induction on Partial Algebras." *) (*The type of functions (operations) in the usual sense, of specified arity.*) (*All types of functions are assumed to be finite.*) Definition function (T:Type) (arity:nat) : Type := T^arity -> T. Definition function_transfer {T a b} (pf:a = b) (f:function T a) : function T b := transfer_dep pf f. Definition function_transfer_r {T a b} (pf:b = a) (f:function T a) : function T b := transfer_dep_r pf f. Lemma function_transfer_eq_refl : forall {T a} (f:function T a), function_transfer eq_refl f = f. unfold function_transfer; intros; rewrite transfer_dep_eq_refl; auto. Qed. Lemma function_transfer_r_eq_refl : forall {T a} (f:function T a), function_transfer_r eq_refl f = f. unfold function_transfer_r; intros; rewrite transfer_dep_r_eq_refl; auto. Qed. Lemma finite_function : forall {T} n, FiniteT T -> FiniteT (function T n). unfold function; intros. apply finite_exp; auto. apply finite_nprod; auto. Qed. Lemma function_eq_dec : forall {T} n, FiniteT T -> type_eq_dec (function T n). intros; red; intros; apply finite_eq_dec; apply finite_function; auto. Qed. Definition arb_function (T:Type) := {n:nat & (function T n)}. Definition set_of_functions (T:Type) := Ensemble (arb_function T). Definition all_functions (T:Type) : Ensemble (set_of_functions T) := Full_set _. Definition set_of_functions_n (T:Type) (n:nat) : Type := Ensemble (function T n). Definition all_functions_n (T:Type) (arity:nat) : set_of_functions_n T arity := Full_set _. Lemma finite_set_of_functions_n : forall {T n}, FiniteT T -> forall (E:set_of_functions_n T n), Finite E. intros; apply FiniteT_Finite; apply finite_function; auto. Qed. Lemma finite_all_functions_n : forall T n, FiniteT T -> Finite (all_functions_n T n). intros; apply finite_set_of_functions_n; auto. Qed. Lemma arb_function_eq_dec : forall {T}, FiniteT T -> type_eq_dec (arb_function T). intros T h0; red; intros f g. destruct f as [m f], g as [n g]. destruct (nat_eq_dec m n) as [|h1]; subst. pose proof (function_eq_dec n h0) as h1. red in h1. specialize (h1 f g). destruct h1; subst; intuition. right. contradict h1. apply existT_injective1 in h1; auto. Qed. Definition app0 {T} (f:function T 0) : T := f tt. Definition app1 {T} (f:function T 1) (x:T) : T := f (x, tt). Definition app2 {T} (f:function T 2) (x y:T) : T := f (x, (y, tt)). Definition app3 {T} (f:function T 3) (x y z:T) : T := f (x, (y, (z, tt))). Definition app4 {T} (f:function T 4) (x1 x2 x3 x4:T) : T := f (x1, (x2, (x3, (x4, tt)))). Definition app5 {T} (f:function T 5) (x1 x2 x3 x4 x5:T) : T := f (x1, (x2, (x3, (x4, (x5, tt))))). Definition app6 {T} (f:function T 6) (x1 x2 x3 x4 x5 x6:T) : T := f (x1, (x2, (x3, (x4, (x5, (x6, tt)))))). Definition app7 {T} (f:function T 7) (x1 x2 x3 x4 x5 x6 x7:T) : T := f (x1, (x2, (x3, (x4, (x5, (x6, (x7, tt))))))). Definition app8 {T} (f:function T 8) (x1 x2 x3 x4 x5 x6 x7 x8:T) : T := f (x1, (x2, (x3, (x4, (x5, (x6, (x7, (x8, tt)))))))). Definition app9 {T} (f:function T 9) (x1 x2 x3 x4 x5 x6 x7 x8 x9:T) : T := f (x1, (x2, (x3, (x4, (x5, (x6, (x7, (x8, (x9, tt))))))))). Definition Pnk_alg (n k:nat) : Type := function {m | m < k} n. Definition Pk_alg (k:nat) : Type := arb_function {m | m < k}. Definition PnA_alg (n:nat) (A:Type) : Type := function A n. Definition P_A (A:Type) : Type := arb_function A. Definition E2 : Type := {n | n < 2}. Definition to_E2 n (pf:n < 2) : E2 := exist (fun c => c < 2) n pf. Definition e2_0 : E2 := to_E2 0 (lt_0_Sn _). Definition e2_1 : E2 := to_E2 1 (lt_n_Sn _). Definition unary_to_function {T} (f:T->T) : function T 1 := fun x => f (fst x). Definition binary_to_function {T} (f:T->T->T) : function T 2 := fun x => f (fst x) (fst (snd x)). Definition ternary_to_function {T} (f:T->T->T->T) : function T 3 := fun x => f (fst x) (fst (snd x)) (fst (snd (snd x))). Definition unary_to_arb_function {T} (f:T->T) : arb_function T := existT _ _ (unary_to_function f). Definition binary_to_arb_function {T} (f:T->T->T) : arb_function T := existT _ _ (binary_to_function f). Definition ternary_to_arb_function {T} (f:T->T->T->T) : arb_function T := existT _ _ (ternary_to_function f). Definition function_to_arb {A n} (f:function A n) : arb_function A := existT _ n f. (*Digit or degree or arity of [arb_function]*) Definition ar {A} (f:arb_function A) : nat := projT1 f. Definition arbn {T} (f:arb_function T) : function T (ar f) := projT2 f. Definition tuple_trans {T} {m n} (pf:m = n) (x:T^m) : T^n := transfer_dep pf x. Lemma tuple_trans_eq_refl : forall {T} {n} (x:T^n), tuple_trans eq_refl x = x. unfold tuple_trans; intros; rewrite transfer_dep_eq_refl; auto. Qed. (*Extracts the function from an [arb_function]*) Definition funct {A} (f:arb_function A) : function A (ar f) := match f with | existT _ f => f end. Definition im_arb_function {A} (f:arb_function A) : Ensemble A := Im (Full_set _) (funct f). Lemma finite_im_arb_function : forall {A} (f:arb_function A), FiniteT A -> Finite (im_arb_function f). intros A f h1. unfold im_arb_function. apply finite_image. rewrite Finite_FiniteT_iff. apply finite_nprod; auto. Qed. Definition PnA_to_PA_ens {A n} (F:Ensemble (PnA_alg n A)) : Ensemble (P_A A) := Im F (fun f => function_to_arb f). Definition Fn_alg {A:Type} (F:Ensemble (P_A A)) (n:nat) : Ensemble (PnA_alg n A) := [f:PnA_alg n A | Inn F (function_to_arb f)]. Definition P_AB {A} (B:Ensemble A) : Ensemble (P_A A) := [f:P_A A | Included (im_arb_function f) B]. Definition P_kl k l : Ensemble (Pk_alg k) := [f:Pk_alg k | Included (im_proj1_sig (im_arb_function f)) (seg l)]. Definition P_A' A l : Ensemble (P_A A) := [f:P_A A | cardinal _ (im_arb_function f) l]. Definition P_k' n l := P_A' {m | m < n} l. Lemma arb_function_ext : forall {T} (f g:arb_function T) (pf:ar f = ar g), (forall (x:T^(ar f)), funct f x = funct g (tuple_trans pf x)) -> f = g. unfold ar, funct; intros T f g. destruct f as [m f], g as [n g]. intros h0 h1; subst; simpl in h0; subst. f_equal. apply functional_extensionality. intro x. specialize (h1 x). rewrite tuple_trans_eq_refl in h1; auto. Qed. Lemma function_ext : forall {T:Type} {m n} (f:function T m) (g:function T n) (pf:m = n), (forall x:T^m, f x = g (tuple_trans pf x)) -> f = function_transfer_r pf g. intros T m n f g h1 h2; subst. rewrite function_transfer_r_eq_refl. pose proof (arb_function_ext (existT _ _ f) (existT _ _ g)) as h3. simpl in h3. specialize (h3 eq_refl h2). eapply existT_inv in h3. rewrite transfer_dep_r_eq_refl in h3; auto. Grab Existential Variables. reflexivity. Qed. Definition agree_except {T n} (x y:T^n) i : Prop := forall j (pf:j < n), i <> j -> nth_prod j pf x = nth_prod j pf y. Inductive essential_var {T n} i (pfi:i < n) (f:function T n) : Prop := | essential_var_intro : forall (x y:T^n), agree_except x y i -> f x <> f y -> essential_var i pfi f. Definition depends {T} (f:arb_function T) i (pf:i < ar f) : Prop := essential_var i pf (funct f). Definition depends' {T n} (f:function T n) i (pf:i < n) : Prop := essential_var i pf f. Definition fictitious_var {T a} i (f:function T a) : Prop := forall (x y:T^a), agree_except x y i -> f x = f y. (*This uses LEM (NNPP).*) Lemma essential_iff : forall {T i n} (pf:i < n) (f:function T n), essential_var i pf f <-> ~fictitious_var i f. intros T i n h2 f. split; intro h4. destruct h4. intro h4. red in h4. specialize (h4 x y H). contradiction. unfold fictitious_var in h4. apply NNPP. contradict h4. intros x y h5. apply NNPP. contradict h4. econstructor. apply h5. assumption. Qed. Inductive essential {T} (f:arb_function T) : Prop := | essential_intro : forall i (pf:i < ar f), essential_var i pf (funct f) -> essential f. (*This takes a function from a FiniteT on [option T] and returns a function on [T], with values of [f] being translated one to the left, i.e. [f (Some %)] = [f_f_o %] -- values that map to None, are replaced by [def]*) Definition function_from_option {T n} (def:T) : function (option T) n -> 0 < n -> function T n := match n with | 0 => fun f pflt => fun _ => False_rect _ (lt_irrefl _ pflt) | S n' => fun f _ => fun x => let y := f (im_nprod (@Some T) x) in match y with | None => def | Some y' => y' end end. (*This takes a function from a FiniteT of cardinality [S n] and returns a function on [n], with values of [f] being translated one to the left, i.e. [f (Some %)] = [a_f_f_o %] -- values that map to None, are replaced by [def:T]*) Definition arb_function_from_option {T} (def:T) (f:arb_function (option T)) : 0 < ar f -> arb_function T := fun pf => existT _ (ar f) (function_from_option def (funct f) pf). (*Projection*) Definition e {T:Type} n i (pf:i < n) : function T n := fun x => nth_prod i pf x. (*Constant*) Definition c {T:Type} n (a:T) : function T n := fun _ => a. (*Boolean and*) Definition e2and : function {n:nat | n < 2} 2 := fun p => let (x, z) := p in let (y, _) := z in let (a, pfa) := x in let (b, pfb) := y in match a with | 0 => exist (fun n => n < 2) 0 (lt_0_Sn _) | S _ => match b with | 0 => exist (fun n => n < 2) 0 (lt_0_Sn _) | S _ => exist (fun n => n < 2) 1 (lt_n_Sn _) end end. (*Boolean and*) Definition e2or : function {n:nat | n < 2} 2 := fun p => let (x, z) := p in let (y, _) := z in let (a, pfa) := x in let (b, pfb) := y in match a with | 0 => match b with | 0 => exist (fun n => n < 2) 0 (lt_0_Sn _) | S _ => exist (fun n => n < 2) 1 (lt_n_Sn _) end | S _ => exist (fun n => n < 2) 1 (lt_n_Sn _) end. (*Boolean exclusive disjunction (+)*) Definition e2xor : function {n:nat | n < 2} 2 := fun p => let (x, z) := p in let (y, _) := z in let (a, pfa) := x in let (b, pfb) := y in match a with | 0 => match b with | 0 => exist (fun n => n < 2) 0 (lt_0_Sn _) | S _ => exist (fun n => n < 2) 1 (lt_n_Sn _) end | S _ => match b with | 0 => exist (fun n => n < 2) 1 (lt_n_Sn _) | S _ => exist (fun n => n < 2) 0 (lt_0_Sn _) end end. (*Boolean implication*) Definition e2impl : function {n:nat | n < 2} 2 := fun p => let (x, z) := p in let (y, _) := z in let (a, pfa) := x in let (b, pfb) := y in match a with | 0 => exist (fun n => n < 2) 1 (lt_n_Sn _) | S _ => match b with | 0 => exist (fun n => n < 2) 0 (lt_0_Sn _) | S _ => exist (fun n => n < 2) 1 (lt_n_Sn _) end end. Definition e2iff : function {n:nat | n < 2} 2 := fun p => let (x, z) := p in let (y, _) := z in let (a, pfa) := x in let (b, pfb) := y in match a with | 0 => match b with | 0 => exist (fun n => n < 2) 1 (lt_n_Sn _) | S _ => exist (fun n => n < 2) 0 (lt_0_Sn _) end | S _ => match b with | 0 => exist (fun n => n < 2) 0 (lt_0_Sn _) | S _ => exist (fun n => n < 2) 1 (lt_n_Sn _) end end. Definition e2not : function {n | n < 2} 1 := fun p => let (x, _) := p in let (a, _) := x in match a with | 0 => exist (fun n => n < 2) 1 (lt_n_Sn _) | S _ => exist (fun n => n < 2) 0 (lt_0_Sn _) end. (*These [e]-Boolean algebras don't require you to write [app2 [f]]*) Definition eand x y := app2 e2and x y. Definition eor x y := app2 e2or x y. Definition exor x y := app2 e2xor x y. Definition eiff x y := app2 e2iff x y. Definition eimpl x y := app2 e2impl x y. Definition enot x := app1 e2not x. (*Commutative Boolean Functions*) Inductive comm_boolean_function : function {n:nat | n < 2} 2 -> Prop := | boolean_and : comm_boolean_function e2and | boolean_or : comm_boolean_function e2or | boolean_xor : comm_boolean_function e2xor | boolean_iff : comm_boolean_function e2iff. (*The 5 boolean functions in P_2*) Inductive boolean_function : function {n:nat | n < 2} 2 -> Prop := | boolean_comm : forall f, comm_boolean_function f -> boolean_function f | boolean_impl : boolean_function e2impl. Lemma comm_boolean_function_assoc : forall (f:function {n | n < 2} 2) x y z, comm_boolean_function f -> app2 f (app2 f x y) z = app2 f x (app2 f y z). unfold app2; intros f x y z h1; revert x y z; simpl; destruct h1; intros x y z; destruct x as [x], y as [y], z as [z]; destruct x, y, z; simpl; auto. Qed. Lemma and_same : forall x, eand x x = x. unfold eand, app2; intro x; destruct x as [x]; destruct x; simpl; auto; f_equal; try apply proof_irrelevance. pose proof l as h1. apply lt_S_n in h1. apply lt1 in h1; subst; f_equal; apply proof_irrelevance. Qed. Lemma or_same : forall x, eor x x = x. unfold eor, app2; intro x; destruct x as [x]; destruct x; simpl; auto; f_equal; try apply proof_irrelevance. pose proof l as h1. apply lt_S_n in h1. apply lt1 in h1; subst. f_equal; apply proof_irrelevance. Qed. Lemma xor_same : forall x, exor x x = exist (fun c => c < 2) 0 (lt_0_Sn _). unfold exor, app2; intro x; destruct x as [x]; destruct x; simpl; auto; f_equal; try apply proof_irrelevance. Qed. Lemma iff_same : forall x, eiff x x = exist (fun c => c < 2) 1 (lt_n_Sn _). unfold eiff, app2; intro x; destruct x as [x]; destruct x; simpl; auto; f_equal; try apply proof_irrelevance. Qed. Lemma impl_same : forall x, eimpl x x = exist (fun c => c < 2) 1 (lt_n_Sn _). unfold eiff, app2; intro x; destruct x as [x]; destruct x; simpl; auto; f_equal; try apply proof_irrelevance. Qed. Lemma and0 : forall x, eand x e2_0 = e2_0. unfold eand, e2_0; intro x; destruct x as [x]; destruct x; simpl; auto. Qed. Lemma and1 : forall x, eand x e2_1 = x. unfold eand, app2, e2_1; intro x; destruct x as [x]; destruct x; simpl; auto; f_equal; try apply proof_irrelevance. pose proof (lt_S_n _ _ l) as h1. apply lt1 in h1; subst. f_equal; apply proof_irrelevance. Qed. Lemma or0 : forall x, eor x e2_0 = x. unfold eor, e2_0, app2; intro x; destruct x as [x]; destruct x; simpl; auto; f_equal; try apply proof_irrelevance. pose proof (lt_S_n _ _ l) as h1. apply lt1 in h1; subst. f_equal; apply proof_irrelevance. Qed. Lemma or1 : forall x, eor x e2_1 = e2_1. unfold eor, e2_1, app2; intro x; destruct x as [x]; destruct x; simpl; auto. Qed. Lemma comm_boolean_function_compat : forall (f:function _ _), comm_boolean_function f -> forall x y, app2 f x y = app2 f y x. unfold app2; intros f h1; destruct h1; simpl; intros x y; destruct x as [x], y as [y]; destruct x, y; auto. Qed. Lemma and_not : forall x, eand x (enot x) = exist (fun x => x < 2) 0 (lt_0_Sn _). intro x; destruct x as [x]; destruct x; auto. Qed. Lemma or_not : forall x, eor x (enot x) = exist (fun x => x < 2) 1 (lt_n_Sn _). intro x; destruct x as [x]; destruct x; auto. Qed. Lemma not_not : forall x, enot (enot x) = x. unfold enot, app1; intro x; destruct x as [x]; destruct x; simpl; f_equal. apply proof_irrelevance. pose (lt_S_n _ _ l) as h1. apply lt1 in h1; subst. f_equal. apply proof_irrelevance. Qed. Lemma demorgan_and : forall x y, enot (eand x y) = eor (enot x) (enot y). intros x y; destruct x as [x], y as [y]; destruct x, y; simpl; auto. Qed. Lemma demorgan_or : forall x y, enot (eor x y) = eand (enot x) (enot y). intros x y; destruct x as [x], y as [y]; destruct x, y; simpl; auto. Qed. Lemma impl_eq : forall x y, eimpl x y = eor (enot x) y. intros x y; destruct x as [x], y as [y]; destruct x, y; simpl; auto. Qed. Lemma not_impl : forall x y, enot (eimpl x y) = eand x (enot y). intros x y; destruct x as [x], y as [y]; destruct x, y; simpl; auto. Qed. Lemma distr_over_and : forall x y z, eor x (eand y z) = eand (eor x y) (eor x z). unfold eor, eand, app2; intros x y z; destruct x as [x], y as [y], z as [z]; destruct x, y, z; simpl; auto. Qed. Lemma distr_over_or : forall x y z, eand x (eor y z) = eor (eand x y) (eand x z). unfold eor, eand, app2; intros x y z; destruct x as [x], y as [y], z as [z]; destruct x, y, z; simpl; auto. Qed. Lemma absorb_over_and : forall x y, eor x (eand x y) = x. unfold eor, eand, app2; intros x y; destruct x as [x], y as [y]; destruct x, y; simpl; auto; f_equal; try apply proof_irrelevance. pose proof (lt_S_n _ _ l) as h1. apply lt1 in h1; subst. f_equal; apply proof_irrelevance. pose proof (lt_S_n _ _ l) as h1. pose proof (lt_S_n _ _ l0) as h2. apply lt1 in h1. apply lt1 in h2. subst. f_equal. apply proof_irrelevance. Qed. Lemma abosrb_over_or : forall x y, eand x (eor x y) = x. unfold eor, eand, app2; intros x y; destruct x as [x], y as [y]; destruct x, y; simpl; auto; f_equal; try apply proof_irrelevance. pose proof (lt_S_n _ _ l) as h1. apply lt1 in h1; subst. f_equal; apply proof_irrelevance. pose proof (lt_S_n _ _ l) as h1. pose proof (lt_S_n _ _ l0) as h2. apply lt1 in h1. apply lt1 in h2. subst. f_equal. apply proof_irrelevance. Qed. Section OrderOnFunctions. Definition lt_function {T n} (pfft:FiniteT T) (ht:T -> (FiniteT_card T pfft) @) (hf:T^n -> (FiniteT_card _ (finite_nprod _ n pfft)) @) (pffit:FiniteT_index pfft ht) (pffif:FiniteT_index (finite_nprod _ n pfft) hf) : Relation (function T n) := lt_fun (finite_nprod _ n pfft) hf pfft ht pffif pffit. End OrderOnFunctions. Section SuperpositionOperations. (*The unassumed premise that "[r < n] and t is surjective" and "[r > n] or [r = n]" and [t] is injective distinguishes edelta and enabla (resp.).*) (*The [s] prefix stands for "superposition"*) Definition sdelta {T} {r n} {t:seg_fun n nat} (pft:seg_fun_endo t r) (f:function T n) : function T r := fun x => f (nprod_identify pft x). Definition snabla {T} {u n} {q:seg_fun n nat} (pfq:seg_fun_endo q u) (f:function T n) : function T u := sdelta pfq f. Definition spi {T} {n} {s:seg_fun n nat} (pfs:seg_fun_endo s n) (f:function T n) : function T n := sdelta pfs f. Definition sstar {T} {m n i} (pf0:0 < i) (pfi:i < n) (f:function T n) (g:function T m) : function T (m + (pred n)) := fun (x:T^(m + (pred n))) => let tri := star_split x pf0 pfi in let fs := fst tri in let s := fst (snd tri) in let t := snd (snd tri) in let gx := g s in let pfe := S_pred_plus_minus _ _ pf0 pfi in let x' := transfer_nprod pfe (glue_nprod (glue_nprod_elt fs gx) t) in f x'. Inductive sup_op1 {T} : forall {m n}, (function T m -> function T n) -> Prop := | pi_intro : forall {n} (s:seg_fun n nat) (pfs:perm_seg s), let pfe := inj_seg_fun_endo_endo _ pfs in sup_op1 (spi pfe) | delta_intro : forall {r n}, r < n -> forall (t:seg_fun n nat) (pfs:surj_seg_fun_endo t r), let pfe := surj_seg_fun_endo_endo _ pfs in sup_op1 (sdelta pfe) | nabla_intro : forall {u n}, u > n -> forall (q:seg_fun n nat) (pfi:inj_seg_fun_endo q u), let pfe := inj_seg_fun_endo_endo _ pfi in sup_op1 (sdelta pfe). Inductive sup_op2 {T} : forall {m n}, (function T n -> function T m -> function T (m + (pred n))) -> Prop := | star_intro : forall {m n i} (f:function T n) (g:function T m) (pf0:0 < i) (pfi:i < n), sup_op2 (sstar pf0 pfi) (m := m). Inductive sup_op {T} : forall {U}, (U->Prop) -> Prop := | sup_op1_intro : forall m n, sup_op (@sup_op1 T m n) | sup_op2_intro : forall m n, sup_op (@sup_op2 T m n). End SuperpositionOperations. Section MalTsevOperations. Definition mzeta {T n} (f:function T n) : function T n := fun x => f (nprod_identify (seg_fun_endo_right_shift n) x). Definition mtau {T n} : function T n -> function T n := match n with | 0 => fun f => f | S n' => match n' with | 0 => fun f => f | S n'' => fun f x => let pf0 := lt_0_Sn (S n'') in let pf1 := lt_n_S _ _ (lt_0_Sn n'') in f (nprod_identify (transpose_seg_fun_endo pf0 pf1) x) end end. Definition mdelta_le {T n} (pfle:n <= 1) (f:function T n) : function T n := f. Definition mdelta_gt {T n} (pf1:1 < n) (f:function T n) : function T (pred n) := fun x => let pf0 := lt_S_pred _ _ pf1 in let pfe := eq_sym (S_pred _ _ pf1) in f (transfer_nprod pfe (nprod_identify (add_seg_to_identify_endo (pred n) pf0) x)). Definition mnabla0 {T} : function T 0 -> function T 0 := fun f => f. Definition mnabla {T n} (f:function T n) : function T (S n) := fun x => f (nprod_identify (tl_seg_for_fictitious_endo (S n)) x). Definition mstar {T m n} (pfn:0 < n) (pfm:0 < m) (f:function T n) (g:function T m) : function T (pred (m + n)) := fun x => let pfle := le_pred_sum_pos _ _ pfm pfn in let pr := split_nprod x _ pfle in let pfe := S_pred_sum_minus _ _ pfm pfn in f (transfer_nprod pfe (g (fst pr), snd pr)). Inductive mal_tsev_op1 {T} : forall {m n}, (function T m -> function T n) -> Prop := | zeta_intro : forall {n}, mal_tsev_op1 (@mzeta T n) | tau_intro : forall {n}, mal_tsev_op1 (@mtau T n) | mdelta_le_intro : forall {n} (pf:n <= 1), mal_tsev_op1 (mdelta_le pf) | mdelta_intro : forall {n} (pf:1 < n), mal_tsev_op1 (mdelta_gt pf) | nabla0_intro : mal_tsev_op1 mnabla0 | mnabla_intro : forall {n}, mal_tsev_op1 (@mnabla T n). Inductive mal_tsev_op2 {T} : forall {m n}, (function T n -> function T m -> function T (pred (m + n))) -> Prop := | mstar_intro : forall {m n} (pfm:0 < m) (pfn:0 < n), mal_tsev_op2 (mstar pfn pfm). Inductive mal_tsev_op {T} : forall {U}, (U->Prop) -> Prop := | mal_tsev_intro : forall m n, mal_tsev_op (@mal_tsev_op1 T m n) | mal_tsev_op2_intro : forall m n, mal_tsev_op (@mal_tsev_op2 T m n). Definition mzeta_to_arb {T} (f:arb_function T) : arb_function T := existT _ _ (mzeta (projT2 f)). Definition mtau_to_arb {T} (f:arb_function T) : arb_function T := existT _ _ (mtau (projT2 f)). Definition mdelta_to_arb {T} (f:arb_function T) : arb_function T. destruct f as [n f]. revert f. induction n as [|n ih]. intro f; refine (existT _ _ f). destruct n as [|n]. intro f; refine (existT _ _ f). intro f. refine (existT _ (S n) _). refine (mdelta_gt (lt_1_SSn _) f). Defined. Definition mnabla_to_arb {T} (f:arb_function T) : arb_function T := existT _ _ (mnabla (projT2 f)). Definition mstar_to_arb {T} (f:arb_function T) (g:arb_function T) : arb_function T. destruct f as [n f], g as [m g]. destruct n as [|n]. refine (existT _ _ f). destruct m as [|m]. refine (existT _ _ g). refine (existT _ _ (mstar (lt_0_Sn _) (lt_0_Sn _) f g)). Defined. End MalTsevOperations.bool2-v0-3/InfiniteOperations.v0000644000175000017500000013261113575574055017346 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014, Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Versiuon 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Export BoolAlgBasics. Require Import SetUtilities. Require Export Image. Require Import FunctionalExtensionality. Require Export Basics. Require Import DecidableDec. Require Import Description. Require Import LogicUtilities. Require Import TypeUtilities. (*This section uses the indexed function notation from Givant/Halmos; for ensemble notation, see the subsequent top-level section: "Section infinite_set"*) Definition Bcompi {It:Type} {B:Bool_Alg} (p: It->Btype (Bc B)) : It->(Btype (Bc B)) := fun (i:It) => -(p i). Notation "-. x" := (Bcompi x) (at level 30). Section infinite. Variable B:Bool_Alg. Let Bt := Btype (Bc B). Variable It : Type. Let IS := Full_set It. Lemma doub_neg_ind : forall (p : It->Bt), -.(-.p) = p. intros p. unfold Bcompi. cut ((fun i : It => - - p i) = (fun i : It => p i)). intro h. apply h. apply functional_extensionality. intro x. rewrite doub_neg. reflexivity. Qed. Definition SupGen {Gt:Type} (S:Ensemble Gt) (p:Gt->Bt) (b:Bt) : Prop := sup (Im S p) b. Definition InfGen {Gt:Type} (S:Ensemble Gt) (p:Gt->Bt) (b:Bt) : Prop := inf (Im S p) b. Definition Sup (p: It->Bt) (b:Bt) : Prop := SupGen IS p b. Definition Inf (p: It->Bt) (b:Bt) : Prop := InfGen IS p b. Lemma comp_ind : forall (p : It->Bt) (x:Bt), Im IS p (-x) <-> Im IS (-.p) x. intros p x. split; intro h. (* left *) inversion h; clear h; subst. pose proof (f_equal (@Bcomp _) H0) as h1. rewrite <- doub_neg in h1; subst. econstructor. apply H. unfold Bcompi; auto. (* right *) inversion h; clear h; subst. econstructor. apply H. unfold Bcompi; rewrite <- doub_neg; auto. Qed. Lemma infnt_de_mor1 : forall (p : It->Bt) (b:Bt), Sup p b -> Inf (-.p) (-b). intros p b h1. unfold Sup in h1. unfold SupGen in h1. unfold sup in h1. elim h1. intros h1l h1r. unfold ub in h1l. unfold Inf. unfold inf. split. unfold lb. (* -b is a lower bound for -.p *) intros t h2. cut (le (-t) b). rewrite doub_neg with (x := t) at 2. apply mono_comp. unfold In in h2. pose proof (comp_ind p t) as h0. unfold Bt in h0. unfold Bt in h2. rewrite <- h0 in h2. apply h1l with (s := -t). unfold In. assumption. (* next subgoal: -b is the greatest lower bound *) intros b'. unfold lb. unfold In. intro h3. rewrite doub_neg with (x := b'). apply mono_comp. apply h1r with (b' := -b'). unfold ub. unfold In. intros s. intro h4. rewrite doub_neg in h4. rewrite comp_ind in h4. specialize (h3 (-s) h4). eapply mono_comp in h3. rewrite <- doub_neg in h3; auto. Qed. (*Dual of the previous of course.*) Lemma infnt_de_mor2 : forall (p : It->Bt) (b : Bt), Inf p b -> Sup (-.p) (-b). intros p b h1. unfold Inf in h1. unfold InfGen in h1. unfold inf in h1. elim h1. intros h1l h1r. unfold lb in h1l. unfold Sup. unfold sup. split. unfold ub. (* -b is an upper bound for -.p*) intros t h2. cut (le b (-t)). rewrite doub_neg with (x :=t) at 2. apply mono_comp. unfold In in h2. pose proof (comp_ind p t) as h0. unfold Bt in h0. unfold Bt in h2. rewrite <- h0 in h2. apply h1l with (s := -t). unfold In. assumption. (* next suboal: -b is the greatest lower bound *) intros b'. unfold ub. unfold In. intro h3. rewrite doub_neg with (x := b'). apply mono_comp. apply h1r with (b' := -b'). unfold lb. unfold In. intros s. intro h4. rewrite doub_neg in h4. rewrite comp_ind in h4. specialize (h3 (-s) h4). eapply mono_comp in h3. rewrite <- doub_neg in h3; auto. Qed. Lemma infnt_de_mor3 : forall (p : It->Bt) (b : Bt), Sup (-.p) b -> Inf p (-b). intros p b h1. rewrite <- doub_neg_ind with (p := p). apply infnt_de_mor1 with (p := (-.p)). assumption. Qed. Lemma infnt_de_mor4 : forall (p : It->Bt) (b : Bt), Inf (-.p) b -> Sup p (-b). intros p b h1. rewrite <- doub_neg_ind with (p := p). apply infnt_de_mor2 with (p := (-.p)). assumption. Qed. Definition ind_complete : Prop := forall (p : It->Bt), (exists b_, Inf p b_) /\ (exists b', Sup p b'). Lemma ind_complete_sup : (forall (p : It->Bt), (exists b', Sup p b')) -> ind_complete. intros h1; red; intro p. split. specialize (h1 (-.p)) ; destruct h1 as [b]. exists (-b). apply infnt_de_mor3; auto. auto. Qed. Lemma ind_complete_inf : (forall (p : It->Bt), (exists b_, Inf p b_)) -> ind_complete. intros h1; red; intro p. split; auto. specialize (h1 (-.p)). destruct h1 as [b h1]. exists (-b). apply infnt_de_mor4; auto. Qed. Section infnt_assoc. Variable Jt:Type. Variable IIS:Ensemble (Ensemble It). Hypothesis covers : (FamilyUnion IIS) = IS. Lemma infnt_assoc_sup : (forall (p: It -> Bt) (q:Jt->Bt), (forall (S:Ensemble It), In IIS S -> (exists (j:Jt), SupGen S p (q j))) -> (forall (j:Jt), exists (S:Ensemble It), In IIS S /\ SupGen S p (q j)) -> (forall (q':Bt), SupGen (Full_set Jt) q q' -> Sup p q')). intros p q h1 h2' q'. intro h4'. destruct (ens_eq_dec _ IIS (Empty_set _)) as [h3 | h4]. rewrite h3 in covers. rewrite h3 in h2'. clear h3. assert (h3:Full_set Jt = Empty_set _). apply Extensionality_Ensembles; red; split; auto with sets. red. intros x ?. specialize (h2' x). destruct h2' as [S h2]. destruct h2; contradiction. rewrite h3 in h4'. unfold SupGen in h4'. rewrite image_empty in h4'. pose proof (sup_empty B) as h5. pose proof (sup_unq _ _ _ h4' h5) as heq. rewrite heq. red. red. rewrite empty_family_union in covers. subst. rewrite <- covers. rewrite image_empty. assumption. apply not_empty_Inhabited in h4. destruct h4 as [S' h4]. red. red. assert (h7: ub (Im IS p) q'). unfold ub. intros b h8. unfold In in h8. inversion h8. assert (h5:forall (i:It), (exists (S1:(Ensemble It)), (In IIS S1) /\ In S1 i)). intro i. assert (h6: In (FamilyUnion IIS) i). rewrite covers. constructor. inversion h6. exists S. tauto. assert (h10:exists S1 : Ensemble It, In IIS S1 /\ In S1 x). apply h5. elim h10. intros S1 h. assert (h11: exists j : Jt, SupGen S1 p (q j)). apply h1. apply h. elim h11. intros j h12. assert (h13: le b (q j)). apply h12. unfold In. apply Im_intro with (x := x) (y := b). apply h. unfold In. assumption. assert (h14: le (q j) q'). apply h4'. unfold In. apply Im_intro with (x := j). apply Full_intro. reflexivity. apply trans_le with (x := b) (y := (q j)) (z := q'). assumption. assumption. assert (h15: forall r:Bt, ub (Im IS p) r -> le q' r). intros r h16. assert (h17: forall (S: Ensemble It), In IIS S -> exists j : Jt, le (q j) r). intros S h30. assert (h31: exists j : Jt, SupGen S p (q j)). apply h1. apply h30. elim h31. intros j h32. assert (h21: le (q j) r). apply h32. unfold ub. intros b' h22. unfold In in h22. inversion h22. apply h16. unfold In. assert (h23: In IS x). unfold IS. apply Full_intro. apply Im_intro with (x := x). assumption. assumption. exists j. assumption. assert (h24: ub (Im (Full_set Jt) q) r). unfold ub. intros b h25. unfold In in h25. inversion h25. assert (h26: exists S : Ensemble It, In IIS S /\ SupGen S p (q x)). apply h2'. elim h26. intros S h27. elim h27. intros h27l h27r. rewrite <- H0 in h27r. apply h27r. unfold ub. intros b' h28. unfold In in h28. apply h16. inversion h28. rewrite H3. unfold In. apply Im_intro with (x:=x0). assert (h29: (FamilyUnion IIS = IS)). apply covers. unfold IS. apply Full_intro. reflexivity. apply h4'. apply h24. unfold Sup. unfold SupGen. unfold sup. split. assumption. apply h15. Qed. (*dual*) Lemma infnt_assoc_inf : (forall (p: It -> Bt) (q:Jt->Bt), (forall (S:Ensemble It), In IIS S -> (exists (j:Jt), InfGen S p (q j))) -> (forall (j:Jt), exists (S:Ensemble It), In IIS S /\ InfGen S p (q j)) -> (forall (q':Bt), InfGen (Full_set Jt) q q' -> Inf p q')). intros p q h1 h2' q' h4'. destruct (ens_eq_dec _ IIS (Empty_set _)) as [h3 | h4]. rewrite h3 in covers. rewrite h3 in h1. rewrite h3 in h2'. clear h3. assert (h3:Full_set Jt = Empty_set _). apply Extensionality_Ensembles; red; split; auto with sets. red. intros x ?. specialize (h2' x). destruct h2' as [S h2]. destruct h2; contradiction. rewrite h3 in h4'. unfold InfGen in h4'. rewrite image_empty in h4'. pose proof (inf_empty B) as h5. pose proof (inf_unq _ _ _ h4' h5). red. red. rewrite empty_family_union in covers. subst. rewrite <- covers. rewrite image_empty. assumption. apply not_empty_Inhabited in h4. destruct h4 as [S' h4]. red. red. assert (h7: lb (Im IS p) q'). unfold lb. intros b h8. unfold In in h8. inversion h8. assert (h5:forall (i:It), (exists (S1:(Ensemble It)), (In IIS S1) /\ In S1 i)). intro i. assert (h6: In (FamilyUnion IIS) i). rewrite covers. constructor. inversion h6. exists S. tauto. assert (h10:exists S1 : Ensemble It, In IIS S1 /\ In S1 x). apply h5. elim h10. intros S1 h. assert (h11: exists j : Jt, InfGen S1 p (q j)). apply h1. apply h. elim h11. intros j h12. assert (h13: le (q j) b). apply h12. unfold In. apply Im_intro with (x := x) (y := b). apply h. unfold In. assumption. assert (h14: le q' (q j)). apply h4'. unfold In. apply Im_intro with (x := j). apply Full_intro. reflexivity. apply trans_le with (x := q') (y := (q j)) (z := b). assumption. assumption. assert (h15: forall r:Bt, lb (Im IS p) r -> le r q'). intros r h16. assert (h17: forall (S: Ensemble It), In IIS S -> exists j : Jt, le r (q j)). intros S h30. assert (h31: exists j : Jt, InfGen S p (q j)). apply h1. apply h30. elim h31. intros j h32. assert (h21: le r (q j)). apply h32. unfold lb. intros b' h22. unfold In in h22. inversion h22. apply h16. unfold In. assert (h23: In IS x). unfold IS. apply Full_intro. apply Im_intro with (x := x). assumption. assumption. exists j. assumption. assert (h24: lb (Im (Full_set Jt) q) r). unfold lb. intros b h25. unfold In in h25. inversion h25. assert (h26: exists S : Ensemble It, In IIS S /\ InfGen S p (q x)). apply h2'. elim h26. intros S h27. elim h27. intros h27l h27r. rewrite <- H0 in h27r. apply h27r. unfold lb. intros b' h28. unfold In in h28. apply h16. inversion h28. rewrite H3. unfold In. apply Im_intro with (x:=x0). assert (h29: (FamilyUnion IIS = IS)). apply covers. unfold IS. apply Full_intro. reflexivity. apply h4'. apply h24. unfold Inf. unfold InfGen. unfold inf. split. assumption. apply h15. Qed. End infnt_assoc. Lemma infnt_distr_1_sup : forall {T:Type} (p:Bt) (q:T->Bt) (q':Bt), let pq := (compose (fun x:Bt => p*x) q) in SupGen (Full_set T) q q' -> SupGen (Full_set T) pq (p*q'). intros T p q q' pq h1. assert (h2:ub (Im (Full_set T) pq) (p*q')). unfold ub. intro s. unfold In. intro h3. inversion h3. unfold pq in H0. unfold compose in H0. rewrite H0. apply mono_prod. apply refl_le. apply h1. unfold In. apply Im_intro with (x:=x). apply H. reflexivity. assert (h4: forall r:Bt, ub (Im (Full_set T) pq) r -> le (p*q') r). intros r h5. unfold ub in h5. assert (h6: ub (Im (Full_set T) q) (r + (-p))). unfold ub. intros b h7. unfold In in h7. inversion h7. assert (h8: b = (p + (-p)) * b). rewrite comp_sum. rewrite comm_prod. rewrite one_prod. reflexivity. assert (h9: (p + (-p))*b = (p*b) + (-p)*b). apply dist_sum_r. assert (h10: b = (p*b) + (-p)*b). congruence. assert (h11:le (p*b) r). apply h5. unfold In. rewrite H0. apply Im_intro with (x := x). apply H. unfold pq. unfold compose. reflexivity. assert (h12:le (-p*b) (-p)). unfold le. apply eq_ord. assert (h13: (-p)*b*(-p) = (-p)*(-p)*b). rewrite <- assoc_prod. rewrite (comm_prod _ b (-p)). rewrite assoc_prod. reflexivity. assert (h14: (-p)*(-p) = (-p)). apply idem_prod. rewrite h14 in h13. apply h13. assert (h15 : le (p*b + (-p*b)) (r + -p)). apply mono_sum. assumption. assumption. rewrite h10. assumption. assert (h11: le q' (r+ (-p))). apply h1. assumption. assert (h12: le p p). apply refl_le. assert (h13: le (p*q') (p*(r+-p))). apply mono_prod. assumption. assumption. assert (h14: p*(r+ -p) = p*r + p*-p). apply dist_sum. assert (h15: p*r + p*-p = p*r). rewrite comp_prod. rewrite zero_sum. reflexivity. assert (h16: le (p*r) r). unfold le. apply eq_ord. assert (h18: r*r = r). apply idem_prod. assert (h19: p*r*r = p*(r*r)). rewrite assoc_prod. reflexivity. rewrite h18 in h19. assumption. rewrite <- h15 in h16. rewrite <- h14 in h16. apply trans_le with (x := (p*q')) (y := (p*(r+-p))) (z:=r). assumption. assumption. unfold Sup. unfold SupGen. unfold sup. split. assumption. assumption. Qed. Lemma infnt_distr_1_inf : forall {T:Type} (p:Bt) (q:T->Bt) (q':Bt), let pq := (compose (fun x:Bt => p+x) q) in InfGen (Full_set T) q q' -> InfGen (Full_set T) pq (p+q'). intros T p q q' pq h1. assert (h2:lb (Im (Full_set T) pq) (p+q')). unfold lb. intro s. unfold In. intro h3. inversion h3. unfold pq in H0. unfold compose in H0. rewrite H0. apply mono_sum. apply refl_le. apply h1. unfold In. apply Im_intro with (x:=x). apply H. reflexivity. assert (h4: forall r:Bt, lb (Im (Full_set T) pq) r -> le r (p+q')). intros r h5. unfold lb in h5. assert (h6: lb (Im (Full_set T) q) (r * (-p))). unfold ub. intros b h7. unfold In in h7. inversion h7. assert (h8: b = (p * (-p)) + b). rewrite comp_prod. rewrite comm_sum. rewrite zero_sum. reflexivity. assert (h9: (p * (-p))+b = (p+b) * ((-p)+b)). apply dist_prod_r. assert (h10: b = (p+b) * ((-p)+b)). congruence. assert (h11:le r (p+b)). apply h5. unfold In. rewrite H0. apply Im_intro with (x := x). apply H. unfold pq. unfold compose. reflexivity. assert (h12:le (-p) (-p+b)). unfold le. assert (h13: (-p)+((-p)+b) = ((-p)+(-p))+b). rewrite assoc_sum. reflexivity. assert (h14: (-p)+(-p) = (-p)). apply idem_sum. rewrite h14 in h13. apply h13. assert (h15 : le (r * -p) ((p+b) * (-p+b))). apply mono_prod. assumption. assumption. rewrite h10. assumption. assert (h11: le (r * (-p)) q'). apply h1. assumption. assert (h12: le p p). apply refl_le. assert (h13: le (p+(r*-p)) (p+q')). apply mono_sum. assumption. assumption. assert (h14: p+(r * -p) = (p+r) *( p+-p)). apply dist_prod. assert (h15: (p+r) * (p+-p) = p+r). rewrite comp_sum. rewrite one_prod. reflexivity. assert (h16: le r (p+r)). unfold le. assert (h18: r+r = r). apply idem_sum. assert (h19: r + (p + r) = (r + r) + p). rewrite (comm_sum _ p r). apply assoc_sum. rewrite h19. rewrite h18. rewrite (comm_sum _ p r). reflexivity. rewrite <- h15 in h16. rewrite <- h14 in h16. apply trans_le with (x := r) (y := (p+(r*-p))) (z:=p+q'). assumption. assumption. unfold Inf. unfold InfGen. unfold inf. split. assumption. assumption. Qed. End infinite. Arguments Bcompi [It] [B] _ _. Arguments SupGen [B] [Gt] _ _ _. Arguments InfGen [B] [Gt] _ _ _. Arguments Sup [B] [It] _ _. Arguments Inf [B] [It] _ _. Arguments comp_ind [B] _ _ _. Arguments infnt_de_mor1 [B] _ _ _ _. Arguments infnt_de_mor2 [B] _ _ _ _. Arguments infnt_de_mor3 [B] _ _ _ _. Arguments infnt_de_mor4 [B] _ _ _ _. Arguments ind_complete [B] _. Arguments ind_complete_sup [B] _ _ _. Arguments ind_complete_inf [B] _ _ _. Arguments infnt_assoc_sup [B] _ _ _ _ _ _ _ _ _ _. Arguments infnt_assoc_inf [B] _ _ _ _ _ _ _ _ _ _. Arguments infnt_distr_1_sup [B] _ _ _ _ _. Arguments infnt_distr_1_inf [B] _ _ _ _ _. (*This section contains much of [Section infinite], only it uses ensembles instead of indexed functions.*) Section infinite_set. Variable B:Bool_Alg. Let Bt := (Btype (Bc B)). Lemma sup_Sup_iff : forall (A:Ensemble Bt) (b:Bt), sup A b <-> @Sup _ (sig_set A) (@proj1_sig _ _) b. intros A b. assert (h2:(Im (full_sig A) (proj1_sig (P:=fun x : Bt => In A x))) = A). apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. destruct x as [x]. simpl. assumption. red. intros x h2. apply Im_intro with (exist _ _ h2). constructor. simpl. reflexivity. split. intro h1. red. red. unfold Bt in h2. unfold Bt. unfold full_sig in h2. rewrite h2. assumption. intro h1. red in h1. red in h1. unfold Bt in h2. unfold Bt in h1. unfold full_sig in h2. rewrite h2 in h1. assumption. Qed. Lemma inf_Inf_iff : forall (A:Ensemble Bt) (b:Bt), inf A b <-> @Inf _ (sig_set A) (@proj1_sig _ _) b. intros A b. assert (h2:(Im (full_sig A) (proj1_sig (P:=fun x : Bt => In A x))) = A). apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. destruct x as [x]. simpl. assumption. red. intros x h2. apply Im_intro with (exist _ _ h2). constructor. simpl. reflexivity. split. intro h1. red. red. unfold Bt in h2. unfold Bt. unfold full_sig in h2. rewrite h2. assumption. intro h1. red in h1. red in h1. unfold Bt in h2. unfold Bt in h1. unfold full_sig in h2. rewrite h2 in h1. assumption. Qed. Lemma set_infnt_de_mor1 : forall (A:Ensemble Bt) (b:Bt), sup A b -> inf (comp_set A) (-b). intros A b h1. pose proof (infnt_de_mor1 (sig_set A) (@proj1_sig _ _) b) as h2. rewrite <- sup_Sup_iff in h2. specialize (h2 h1). rewrite inf_Inf_iff. red. red. red. red in h2. red in h2. red in h2. destruct h2 as [h2l h2r]. split. red. red in h2l. intros s h3. destruct h3 as [s h3]. subst. destruct s as [s h4]. simpl. destruct h4 as [s h4]. subst. assert (h5: In (Im (full_sig A) (Bcompi (proj1_sig (P:=fun x : Bt => In A x)))) (-s)). apply Im_intro with (exist _ _ h4). constructor. unfold Bcompi. simpl. reflexivity. apply h2l; auto. intros b' h3. assert (h4:lb (Im (full_sig A) (Bcompi (proj1_sig (P:=fun x : Bt => In A x)))) b'). red. red in h3. intros s h4. destruct h4 as [s h4]. subst. unfold Bcompi. destruct s as [s h5]. simpl. apply h3. assert (h6:In (comp_set A) (-s)). apply Im_intro with s; auto. apply Im_intro with (exist _ _ h6). constructor. simpl. reflexivity. apply h2r; auto. Qed. Lemma set_infnt_de_mor2 : forall (A:Ensemble Bt) (b : Bt), inf A b -> sup (comp_set A) (-b). intros A b h1. pose proof (infnt_de_mor2 (sig_set A) (@proj1_sig _ _) b) as h2. rewrite <- inf_Inf_iff in h2. specialize (h2 h1). rewrite sup_Sup_iff. red. red. red. red in h2. red in h2. red in h2. destruct h2 as [h2l h2r]. split. red. red in h2l. intros s h3. destruct h3 as [s h3]. subst. destruct s as [s h4]. simpl. destruct h4 as [s h4]. subst. assert (h5: In (Im (full_sig A) (Bcompi (proj1_sig (P:=fun x : Bt => In A x)))) (-s)). apply Im_intro with (exist _ _ h4). constructor. unfold Bcompi. simpl. reflexivity. apply h2l; auto. intros b' h3. assert (h4:ub (Im (full_sig A) (Bcompi (proj1_sig (P:=fun x : Bt => In A x)))) b'). red. red in h3. intros s h4. destruct h4 as [s h4]. subst. unfold Bcompi. destruct s as [s h5]. simpl. apply h3. assert (h6:In (comp_set A) (-s)). apply Im_intro with s; auto. apply Im_intro with (exist _ _ h6). constructor. simpl. reflexivity. apply h2r; auto. Qed. Lemma set_infnt_de_mor3 : forall (A:Ensemble Bt) (b : Bt), sup (comp_set A) b -> inf A (-b). intros A b h1. rewrite <- comp_set_comp_set with (A:=A). apply set_infnt_de_mor1. assumption. Qed. Lemma set_infnt_de_mor4 : forall (A:Ensemble Bt) (b : Bt), inf (comp_set A) b -> sup A (-b). intros A b h1. rewrite <- comp_set_comp_set with (A:=A). apply set_infnt_de_mor2. assumption. Qed. Definition set_complete : Prop := forall A:Ensemble Bt, exists b' b_:Bt, sup A b' /\ inf A b_. Lemma sup_set_complete_ex : forall (A:Ensemble Bt), set_complete -> exists! p:Bt, sup A p. intros A h1. specialize (h1 A). destruct h1 as [b h1]. destruct h1 as [b' h1]. destruct h1 as [h1 h2]. exists b. red. split; auto. intros b'' h3. apply sup_unq with A; auto. Qed. Definition sup_set_complete (A:Ensemble Bt) (pf:set_complete) := proj1_sig (constructive_definite_description _ (sup_set_complete_ex A pf)). Lemma sup_set_complete_compat : forall (A:Ensemble Bt) (pf:set_complete), sup A (sup_set_complete A pf). intros A h1. unfold sup_set_complete. destruct constructive_definite_description as [a h2]. simpl. assumption. Qed. Lemma inf_set_complete_ex : forall (A:Ensemble Bt), set_complete -> exists! p:Bt, inf A p. intros A h1. specialize (h1 A). destruct h1 as [b h1]. destruct h1 as [b' h1]. destruct h1 as [h1 h2]. exists b'. red. split; auto. intros b'' h3. apply inf_unq with A; auto. Qed. Definition inf_set_complete (A:Ensemble Bt) (pf:set_complete) := proj1_sig (constructive_definite_description _ (inf_set_complete_ex A pf)). Lemma inf_set_complete_compat : forall (A:Ensemble Bt) (pf:set_complete), inf A (inf_set_complete A pf). intros A h1. unfold inf_set_complete. destruct constructive_definite_description as [a h2]. simpl. assumption. Qed. Lemma set_complete_sup : (forall A:Ensemble Bt, exists b', sup A b') -> set_complete. intros h1; red; intro A. pose proof (h1 (comp_set A)) as h1'. specialize (h1 A). destruct h1 as [b' h1]. destruct h1' as [b_ h1']. exists b', (-b_). split; auto. apply set_infnt_de_mor3; auto. Qed. Lemma set_complete_inf : (forall A:Ensemble Bt, (exists b_, inf A b_)) -> set_complete. intros h1; red; intro A. pose proof (h1 (comp_set A)) as h1'. specialize (h1 A). destruct h1' as [b' h2]. exists (-b'). destruct h1 as [b_ h1]. exists b_. split; auto. apply set_infnt_de_mor4; auto. Qed. Lemma set_infnt_assoc_sup : forall (F:Family Bt) (pf:forall A:Ensemble Bt, In F A -> exists p:Bt, sup A p), forall q:Bt, sup (Im (full_sig F) (fun A => (proj1_sig (constructive_definite_description _ (ex_sup_unq _ _ (pf _ (proj2_sig A))))))) q -> sup (FamilyUnion F) q. intros F h1 q h2. destruct (ens_eq_dec _ F (Empty_set _)) as [he | hne]. subst. rewrite empty_family_union. rewrite full_empty_sig_empty, image_empty in h2. unfold Bt. assumption. unfold full_sig, sig_set in h2. pose proof (full_sig_family_union_eq F) as h3. symmetry in h3. unfold full_sig in h3. term1 h2. redterm0 c. rename c0 into f. pose proof (infnt_assoc_sup _ _ _ h3 (@proj1_sig _ _) f) as h4. hfirst h4. intros S h6. destruct h6 as [h6]. exists (exist _ _ h6). simpl. unfold f. destruct constructive_definite_description as [x h8]. simpl. unfold SupGen. assumption. specialize (h4 hf). hfirst h4. intro S. destruct S as [S h7]. simpl. unfold f. destruct constructive_definite_description as [x h8]. simpl. pose proof family_union_intro. exists (Im (full_sig S) (fun x => (exist _ (proj1_sig x) (family_union_intro _ _ _ _ h7 (proj2_sig x))))). split. constructor. rewrite im_im. simpl. rewrite <- im_full_sig_proj1_sig. assumption. red. rewrite im_im. simpl. rewrite <- im_full_sig_proj1_sig. assumption. specialize (h4 hf0 _ h2). red in h4. red in h4. fold (full_sig (FamilyUnion F)) in h4. rewrite <- im_full_sig_proj1_sig in h4. assumption. Qed. Lemma set_infnt_assoc_inf : forall (F:Family Bt) (pf:forall A:Ensemble Bt, In F A -> exists p:Bt, inf A p), forall q:Bt, inf (Im (full_sig F) (fun A => (proj1_sig (constructive_definite_description _ (ex_inf_unq _ _ (pf _ (proj2_sig A))))))) q -> inf (FamilyUnion F) q. intros F h1 q h2. destruct (ens_eq_dec _ F (Empty_set _)) as [he | hne]. subst. rewrite empty_family_union. rewrite full_empty_sig_empty in h2. rewrite image_empty in h2. unfold Bt. assumption. unfold full_sig in h2. unfold sig_set in h2. pose proof (full_sig_family_union_eq F) as h3. symmetry in h3. unfold full_sig in h3. term1 h2. redterm0 c. rename c0 into f. pose proof (infnt_assoc_inf (sig_set (FamilyUnion F)) (sig_set F) _ h3 (@proj1_sig _ _) f) as h4. hfirst h4. intros S h6. destruct h6 as [h6]. exists (exist _ _ h6). simpl. unfold f. destruct constructive_definite_description as [x h8]. simpl. unfold InfGen. assumption. specialize (h4 hf). hfirst h4. intro S. destruct S as [S h7]. simpl. unfold f. destruct constructive_definite_description as [x h8]. simpl. pose proof family_union_intro. exists (Im (full_sig S) (fun x => (exist _ (proj1_sig x) (family_union_intro _ _ _ _ h7 (proj2_sig x))))). split. constructor. rewrite im_im. simpl. rewrite <- im_full_sig_proj1_sig. assumption. red. rewrite im_im. simpl. rewrite <- im_full_sig_proj1_sig. assumption. specialize (h4 hf0). specialize (h4 q). unfold InfGen in h4. specialize (h4 h2). red in h4. red in h4. fold (full_sig (FamilyUnion F)) in h4. rewrite <- im_full_sig_proj1_sig in h4. assumption. Qed. Lemma sup_add_assoc : forall (A:Ensemble Bt) (a p:Bt), sup A p -> sup (Add A a) (p+a). intros A a p h1. pose proof (set_infnt_assoc_sup (Couple A (Singleton a))) as h2. hfirst h2. intros C h3. destruct h3. exists p. assumption. exists a. apply sup_singleton. specialize (h2 hf (p + a)). hfirst h2. red; split; red. intros s h4. destruct h4 as [A' h4 s]. subst. destruct constructive_definite_description as [p' h6]. simpl. destruct A' as [A' h7]. destruct h7. simpl in h6. pose proof (sup_unq _ _ _ h1 h6). subst. apply le_plus. simpl in h6. pose proof (sup_singleton _ a) as h7. pose proof (sup_unq _ _ _ h6 h7). subst. rewrite comm_sum. apply le_plus. intros b' h8. red in h8. pose proof (h8 a) as h9. hfirst h9. assert (h10:In (Couple A (Singleton a)) (Singleton a)). right. apply Im_intro with (exist _ _ h10). constructor. simpl. destruct constructive_definite_description as [a' h12]. simpl. apply sup_unq with (Singleton a). apply sup_singleton. assumption. specialize (h9 hf0). pose proof (h8 p) as h11. hfirst h11. pose proof (hf A) as h13. hfirst h13. left. apply Im_intro with (exist _ _ hf1). constructor. simpl. destruct constructive_definite_description as [p' h14]. simpl. apply sup_unq with A; auto. specialize (h11 hf1). rewrite <- (idem_sum b'). apply mono_sum; auto. specialize (h2 hf0). rewrite couple_add_sing in h2. rewrite family_union_add in h2. rewrite family_union_sing in h2. unfold Add. rewrite Union_commutative. assumption. Qed. Lemma inf_add_assoc : forall (A:Ensemble Bt) (a p:Bt), inf A p -> inf (Add A a) (p*a). intros A a p h1. pose proof (set_infnt_assoc_inf (Couple A (Singleton a))) as h2. hfirst h2. intros C h3. destruct h3. exists p. assumption. exists a. apply inf_singleton. specialize (h2 hf ((p * a))). hfirst h2. red; split; red. intros s h4. destruct h4 as [A' h4 s]. subst. destruct constructive_definite_description as [p' h6]. simpl. destruct A' as [A' h7]. destruct h7. simpl in h6. pose proof (inf_unq _ _ _ h1 h6). subst. apply times_le. simpl in h6. pose proof (inf_singleton _ a) as h7. pose proof (inf_unq _ _ _ h6 h7). subst. rewrite comm_prod. apply times_le. intros b' h8. red in h8. pose proof (h8 a) as h9. hfirst h9. assert (h10:In (Couple A (Singleton a)) (Singleton a)). right. apply Im_intro with (exist _ _ h10). constructor. simpl. destruct constructive_definite_description as [a' h12]. simpl. apply inf_unq with (Singleton a). apply inf_singleton. assumption. specialize (h9 hf0). pose proof (h8 p) as h11. hfirst h11. pose proof (hf A) as h13. hfirst h13. left. apply Im_intro with (exist _ _ hf1). constructor. simpl. destruct constructive_definite_description as [p' h14]. simpl. apply inf_unq with A; auto. specialize (h11 hf1). rewrite <- (idem_prod b'). apply mono_prod; auto. specialize (h2 hf0). rewrite couple_add_sing in h2. rewrite family_union_add in h2. rewrite family_union_sing in h2. unfold Add. rewrite Union_commutative. assumption. Qed. Lemma set_infnt_distr_1_sup : forall (A:Ensemble Bt) (p q:Bt), sup A q -> sup (Im A (fun x => p*x)) (p*q). intros A p q h1. pose proof (infnt_distr_1_sup {x:Bt|In A x} p (@proj1_sig _ _) q) as h2; simpl in h2. hfirst h2. red. fold (sig_set A). fold (full_sig A). rewrite <- im_full_sig_proj1_sig. assumption. specialize (h2 hf). red in h2. simpl in h2. unfold compose in h2. term1 h2. gterm1. assert (h3:c = c0). unfold c, c0; apply Extensionality_Ensembles; red; split; red; intros x h4. destruct h4 as [x h4]. subst. apply Im_intro with (proj1_sig x). apply proj2_sig. reflexivity. destruct h4 as [x h4]. subst. apply Im_intro with (exist _ _ h4). constructor. simpl. reflexivity. unfold c, c0 in h3. rewrite h3 in h2; auto. Qed. Lemma set_infnt_distr_1_inf : forall (A:Ensemble Bt) (p q:Bt), inf A q -> inf (Im A (fun x => p+x)) (p+q). intros A p q h1. pose proof (infnt_distr_1_inf {x:Bt|In A x} p (@proj1_sig _ _) q) as h2; simpl in h2. hfirst h2. red. fold (sig_set A). fold (full_sig A). rewrite <- im_full_sig_proj1_sig; auto. specialize (h2 hf). red in h2; simpl in h2; unfold compose in h2. term1 h2. gterm1. assert (h3:c = c0). unfold c, c0. apply Extensionality_Ensembles; red; split; red; intros x h4. destruct h4 as [x h4]. subst. apply Im_intro with (proj1_sig x). apply proj2_sig. reflexivity. destruct h4 as [x h4]. subst. apply Im_intro with (exist _ _ h4). constructor. simpl. reflexivity. unfold c, c0 in h3. rewrite h3 in h2; auto. Qed. End infinite_set. (*Theorems*) Arguments sup_Sup_iff [B] _ _. Arguments inf_Inf_iff [B] _ _. Arguments set_infnt_de_mor1 [B] _ _ _. Arguments set_infnt_de_mor2 [B] _ _ _. Arguments set_infnt_de_mor3 [B] _ _ _. Arguments set_infnt_de_mor4 [B] _ _ _. Arguments sup_set_complete_ex [B] _ _. Arguments sup_set_complete_compat [B] _ _. Arguments inf_set_complete_ex [B] _ _. Arguments inf_set_complete_compat [B] _ _. Arguments set_complete_sup [B] _ _. Arguments set_complete_inf [B] _ _. Arguments set_infnt_assoc_sup [B] _ _ _ _. Arguments set_infnt_assoc_inf [B] _ _ _ _. Arguments sup_add_assoc [B] _ _ _ _. Arguments inf_add_assoc [B] _ _ _ _. Arguments set_infnt_distr_1_sup [B] _ _ _ _. Arguments set_infnt_distr_1_inf [B] _ _ _ _. (*Definitions*) Arguments sup_set_complete [B] _ _. Arguments inf_set_complete [B] _ _. Definition Bcompi_p {It T:Type} {Bp:Bool_Alg_p T} (p: It->Btype_p T (Bc_p T Bp)) : It->Btype_p T (Bc_p T Bp) := fun (i:It) => %- (p i). Notation "%~ x" := (Bcompi_p x) (at level 30). (*This is a section of dual versions of each of the theorems, this time with [Bool_Alg_p]s instead of [Bool_Alg]s. The below theorems are just wrappers.*) Section ParametricAnalogues. Variable T:Type. Variable Bp:Bool_Alg_p T. Let Btp := Btype_p T (Bc_p T Bp). Variable It : Type. Let IS := Full_set It. Definition ba_conv_ind {Tt:Type} (p:Tt->Btp) : Tt->(Btype (Bc (ba_conv Bp))) := transfer_dep (U:=fun T=>(Tt->T)) (ba_conv_type Bp) p. Lemma Bcompi_p_eq : forall (p:It->Btp), (ba_conv_ind (%~ p)) = Bcompi (ba_conv_ind p). intro p. unfold ba_conv_ind. unfold ba_conv_type. do 2 rewrite transfer_dep_eq_refl. unfold Bcompi. unfold Bcompi_p. reflexivity. Qed. Lemma Bcompi_p_eq' : forall p:It->Btp, Bcompi_p p = Bcompi (ba_conv_ind p). auto. Qed. Lemma doub_neg_ind_p : forall (p : It->Btp), %~(%~p) = p. intro p. do 2 rewrite Bcompi_p_eq'. apply (@doub_neg_ind (ba_conv Bp)). Qed. Definition SupGen_p {Gt:Type} (S:Ensemble Gt) (p:Gt->Btp) (b:Btp) : Prop := sup_p (Im S p) b. Definition InfGen_p {Gt:Type} (S:Ensemble Gt) (p:Gt->Btp) (b:Btp) : Prop := inf_p (Im S p) b. Definition Sup_p (p: It->Btp) (b:Btp) : Prop := SupGen_p IS p b. Definition Inf_p (p: It->Btp) (b:Btp) : Prop := InfGen_p IS p b. Lemma SupGen_p_iff : forall (Gt:Type) (S:Ensemble Gt) (p:Gt->Btp) (b:Btp), SupGen_p S p b <-> SupGen S (ba_conv_ind p) (ba_conv_elt b). intros Gt S p b. unfold SupGen_p. unfold SupGen. rewrite sup_p_iff. tauto. Qed. Definition InfGen_p_iff : forall (Gt:Type) (S:Ensemble Gt) (p:Gt->Btp) (b:Btp), InfGen_p S p b <-> InfGen S (ba_conv_ind p) (ba_conv_elt b). intros Gt S p b. unfold InfGen_p. unfold InfGen. rewrite inf_p_iff. tauto. Qed. Definition Sup_p_iff : forall (p: It->Btp) (b:Btp), Sup_p p b <-> Sup (ba_conv_ind p) (ba_conv_elt b). intros p b. unfold Sup_p. unfold Sup. rewrite SupGen_p_iff. tauto. Qed. Definition Inf_p_iff : forall (p: It->Btp) (b:Btp), Inf_p p b <-> Inf (ba_conv_ind p) (ba_conv_elt b). intros p b. unfold Inf_p. unfold Inf. rewrite InfGen_p_iff. tauto. Qed. Lemma comp_ind_p : forall (p : It->Btp) (x:Btp), Im IS p (%-x) <-> Im IS (%~p) x. intros p x. rewrite ba_conv_comp. rewrite Bcompi_p_eq'. apply (@comp_ind (ba_conv Bp)). Qed. Lemma infnt_de_mor1_p : forall (p : It->Btp) (b:Btp), Sup_p p b -> Inf_p (%~p) (%-b). intros p b. rewrite Sup_p_iff. rewrite Inf_p_iff. apply (@infnt_de_mor1 (ba_conv Bp)). Qed. Lemma infnt_de_mor2_p : forall (p : It->Btp) (b : Btp), Inf_p p b -> Sup_p (%~p) (%-b). intros p b. rewrite Sup_p_iff. rewrite Inf_p_iff. apply (@infnt_de_mor2 (ba_conv Bp)). Qed. Lemma infnt_de_mor3_p : forall (p : It->Btp) (b : Btp), Sup_p (%~p) b -> Inf_p p (%-b). intros p b. rewrite Sup_p_iff. rewrite Inf_p_iff. apply (@infnt_de_mor3 (ba_conv Bp)). Qed. Lemma infnt_de_mor4_p : forall (p : It->Btp) (b : Btp), Inf_p (%~p) b -> Sup_p p (%-b). intros p b. rewrite Sup_p_iff. rewrite Inf_p_iff. apply (@infnt_de_mor4 (ba_conv Bp)). Qed. Definition ind_complete_p : Prop := forall (p : It->Btp), (exists b_, Inf_p p b_) /\ (exists b', Sup_p p b'). Lemma ind_complete_p_iff : ind_complete_p <-> (@ind_complete (ba_conv Bp) It). unfold ind_complete_p. unfold ind_complete. split. intros h1 p. specialize (h1 p). assumption. intros h1 p. specialize (h1 p). assumption. Qed. Lemma ind_complete_sup_p : (forall (p : It->Btp), (exists b', Sup_p p b')) -> ind_complete_p. rewrite ind_complete_p_iff; intro; apply ind_complete_sup; auto. Qed. Lemma ind_complete_inf_p : (forall (p : It->Btp), (exists b_, Inf_p p b_)) -> ind_complete_p. rewrite ind_complete_p_iff; intro; apply ind_complete_inf; auto. Qed. Variable Jt:Type. Variable IIS:Ensemble (Ensemble It). Hypothesis covers : (FamilyUnion IIS) = IS. Lemma infnt_assoc_sup_p : (forall (p: It -> Btp) (q:Jt->Btp), (forall (S:Ensemble It), In IIS S -> (exists (j:Jt), SupGen_p S p (q j))) -> (forall (j:Jt), exists (S:Ensemble It), In IIS S /\ SupGen_p S p (q j)) -> (forall (q':Btp), SupGen_p (Full_set Jt) q q' -> Sup_p p q')). intros p q h1 h2 q' h3. rewrite Sup_p_iff. rewrite SupGen_p_iff in h3. eapply infnt_assoc_sup; auto. apply covers. intros S h4. specialize (h1 _ h4). destruct h1 as [j h1]. rewrite SupGen_p_iff in h1. exists j; auto. Qed. Lemma infnt_assoc_inf_p : (forall (p: It -> Btp) (q:Jt->Btp), (forall (S:Ensemble It), In IIS S -> (exists (j:Jt), InfGen_p S p (q j))) -> (forall (j:Jt), exists (S:Ensemble It), In IIS S /\ InfGen_p S p (q j)) -> (forall (q':Btp), InfGen_p (Full_set Jt) q q' -> Inf_p p q')). intros p q h1 h2 q' h3. rewrite Inf_p_iff. rewrite InfGen_p_iff in h3. eapply infnt_assoc_inf; auto. apply covers. intros S h4. specialize (h1 _ h4). destruct h1 as [j h1]. rewrite InfGen_p_iff in h1. exists j; auto. Qed. Lemma infnt_distr_1_sup_p : forall {T':Type} (p:Btp) (q:T'->Btp) (q':Btp), let pq := (compose (fun x:Btp => p%*x) q) in SupGen_p (Full_set T') q q' -> SupGen_p (Full_set T') pq (p%*q'). intros T' p q q' pq. do 2 rewrite SupGen_p_iff. rewrite ba_conv_times. apply (@infnt_distr_1_sup (ba_conv Bp)). Qed. Lemma infnt_distr_1_inf_p : forall {T':Type} (p:Btp) (q:T'->Btp) (q':Btp), let pq := (compose (fun x:Btp => p%+x) q) in InfGen_p (Full_set T') q q' -> InfGen_p (Full_set T') pq (p%+q'). intros T' p q q' pq. do 2 rewrite InfGen_p_iff. rewrite ba_conv_plus. apply (@infnt_distr_1_inf (ba_conv Bp)). Qed. (*ommitted sup_le and Sup', because I believe they're obsolete.*) End ParametricAnalogues. (*Theorems*) Arguments Bcompi_p_eq [T] [Bp] [It] _. Arguments Bcompi_p_eq' [T] [Bp] [It] _. Arguments doub_neg_ind_p [T] [Bp] [It] _. Arguments SupGen_p_iff [T] [Bp] [Gt] _ _ _. Arguments InfGen_p_iff [T] [Bp] [Gt] _ _ _. Arguments Sup_p_iff [T] [Bp] [It] _ _. Arguments Inf_p_iff [T] [Bp] [It] _ _. Arguments comp_ind_p [T] [Bp] [It] _ _. Arguments infnt_de_mor1_p [T] [Bp] [It] _ _ _. Arguments infnt_de_mor2_p [T] [Bp] [It] _ _ _. Arguments infnt_de_mor3_p [T] [Bp] [It] _ _ _. Arguments infnt_de_mor4_p [T] [Bp] [It] _ _ _. Arguments ind_complete_p_iff [T] _ _. Arguments ind_complete_sup_p [T] [Bp] [It] _ _. Arguments ind_complete_inf_p [T] [Bp] [It] _ _. Arguments infnt_assoc_sup_p [T] [Bp] [It] _ _ _ _ _ _ _ _ _. Arguments infnt_assoc_inf_p [T] [Bp] [It] _ _ _ _ _ _ _ _ _. Arguments infnt_distr_1_sup_p [T] [Bp] [T'] _ _ _ _. Arguments infnt_distr_1_inf_p [T] [Bp] [T'] _ _ _ _. (*Definitions*) Arguments ba_conv_ind [T] [Bp] [Tt] _ _. Arguments SupGen_p [T] [Bp] [Gt] _ _ _. Arguments InfGen_p [T] [Bp] [Gt] _ _ _. Arguments Sup_p [T] [Bp] [It] _ _. Arguments Inf_p [T] [Bp] [It] _ _. Arguments ind_complete_p [T] _ _. Section ParametricAnalogues'. Variable T:Type. Variable Bp:Bool_Alg_p T. Let Btp := Btype_p T (Bc_p T Bp). Lemma sup_Sup_iff_p : forall (A:Ensemble Btp) (b:Btp), sup_p A b <-> @Sup_p _ _ (sig_set A) (@proj1_sig _ _) b. intros A b. rewrite sup_p_iff. rewrite Sup_p_iff. apply (@sup_Sup_iff (ba_conv Bp)). Qed. Lemma inf_Inf_iff_p : forall (A:Ensemble Btp) (b:Btp), inf_p A b <-> @Inf_p _ _ (sig_set A) (@proj1_sig _ _) b. intros A b. rewrite inf_p_iff. rewrite Inf_p_iff. apply (@inf_Inf_iff (ba_conv Bp)). Qed. Lemma set_infnt_de_mor1_p : forall (A:Ensemble Btp) (b:Btp), sup_p A b -> inf_p (comp_set_p A) (%-b). intros A b. rewrite sup_p_iff. rewrite inf_p_iff. rewrite ba_conv_comp. rewrite comp_set_p_eq. apply (@set_infnt_de_mor1 (ba_conv Bp)). Qed. Lemma set_infnt_de_mor2_p : forall (A:Ensemble Btp) (b : Btp), inf_p A b -> sup_p (comp_set_p A) (%-b). intros A b. rewrite sup_p_iff. rewrite inf_p_iff. rewrite ba_conv_comp. rewrite comp_set_p_eq. apply (@set_infnt_de_mor2 (ba_conv Bp)). Qed. Lemma set_infnt_de_mor3_p : forall (A:Ensemble Btp) (b : Btp), sup_p (comp_set_p A) b -> inf_p A (%-b). intros A b. rewrite sup_p_iff. rewrite inf_p_iff. rewrite ba_conv_comp. rewrite comp_set_p_eq. apply (@set_infnt_de_mor3 (ba_conv Bp)). Qed. Lemma set_infnt_de_mor4_p : forall (A:Ensemble Btp) (b : Btp), inf_p (comp_set_p A) b -> sup_p A (%-b). intros A b. rewrite sup_p_iff. rewrite inf_p_iff. rewrite ba_conv_comp. rewrite comp_set_p_eq. apply (@set_infnt_de_mor4 (ba_conv Bp)). Qed. Definition set_complete_p : Prop := forall A:Ensemble Btp, exists b' b_:Btp, sup_p A b' /\ inf_p A b_. Lemma set_complete_p_iff : set_complete_p <-> set_complete (ba_conv Bp). unfold set_complete_p, set_complete. split. intros h1 A. specialize (h1 A). assumption. intros h2 A. specialize (h2 A). assumption. Qed. Lemma sup_set_complete_ex_p : forall (A:Ensemble Btp), set_complete_p -> exists! p:Btp, sup_p A p. intros A h1. rewrite set_complete_p_iff in h1. pose proof (@sup_set_complete_ex (ba_conv Bp) A h1) as h2. assumption. Qed. Definition sup_set_complete_p (A:Ensemble Btp) (pf:set_complete_p) := proj1_sig (constructive_definite_description _ (sup_set_complete_ex_p A pf)). Definition sup_set_complete_p_eq : forall (A:Ensemble Btp) (pf:set_complete_p), sup_set_complete_p A pf = sup_set_complete (ba_conv_set A) (iff1 set_complete_p_iff pf). intros A h1. unfold sup_set_complete_p. unfold sup_set_complete. destruct constructive_definite_description as [x h2]. destruct constructive_definite_description as [y h3]. simpl. rewrite sup_p_iff in h2. pose proof (sup_unq _ _ _ h2 h3) as h4. assumption. Qed. Lemma sup_set_complete_compat_p : forall (A:Ensemble Btp) (pf:set_complete_p), sup_p A (sup_set_complete_p A pf). intros A h1. rewrite sup_p_iff. rewrite sup_set_complete_p_eq. apply sup_set_complete_compat. Qed. Lemma inf_set_complete_ex_p : forall (A:Ensemble Btp), set_complete_p -> exists! p:Btp, inf_p A p. intros A h1. rewrite set_complete_p_iff in h1. pose proof (@inf_set_complete_ex (ba_conv Bp) A h1) as h2. assumption. Qed. Definition inf_set_complete_p (A:Ensemble Btp) (pf:set_complete_p) := proj1_sig (constructive_definite_description _ (inf_set_complete_ex_p A pf)). Definition inf_set_complete_p_eq : forall (A:Ensemble Btp) (pf:set_complete_p), inf_set_complete_p A pf = inf_set_complete (ba_conv_set A) (iff1 set_complete_p_iff pf). intros A h1. unfold inf_set_complete_p. unfold inf_set_complete. destruct constructive_definite_description as [x h2]. destruct constructive_definite_description as [y h3]. simpl. rewrite inf_p_iff in h2. pose proof (inf_unq _ _ _ h2 h3) as h4. assumption. Qed. Lemma inf_set_complete_compat_p : forall (A:Ensemble Btp) (pf:set_complete_p), inf_p A (inf_set_complete_p A pf). intros A h1. rewrite inf_p_iff. rewrite inf_set_complete_p_eq. apply inf_set_complete_compat. Qed. Lemma set_complete_sup_p : (forall A:Ensemble Btp, exists b', sup_p A b') -> set_complete_p. intros h1. rewrite set_complete_p_iff. apply set_complete_sup; auto. Qed. Lemma set_complete_inf_p : (forall A:Ensemble Btp, (exists b_, inf_p A b_)) -> set_complete_p. intros h1. rewrite set_complete_p_iff. apply set_complete_inf; auto. Qed. Lemma set_infnt_assoc_sup_p : forall (F:Family Btp) (pf:forall A:Ensemble Btp, In F A -> exists p:Btp, sup_p A p), forall q:Btp, sup_p (Im (full_sig F) (fun A => (proj1_sig (constructive_definite_description _ (ex_sup_unq_p _ _ _ (pf _ (proj2_sig A))))))) q -> sup_p (FamilyUnion F) q. intros F h1 q h2. rewrite sup_p_iff in h2. rewrite sup_p_iff. unfold ba_conv_set, ba_conv_type in h2. rewrite transfer_dep_eq_refl in h2. pose proof (set_infnt_assoc_sup F (B := ba_conv Bp) h1) as h3. apply h3. term1 h2. redterm0 c. gterm1. redterm0 c1. rename c0 into f; rename c2 into g. clear c1 c. assert (h4:f = g). unfold f, g; apply functional_extensionality; intro. destruct constructive_definite_description as [c h5], constructive_definite_description as [d h6]; simpl. eapply sup_unq_p. apply h5. assumption. fold g; rewrite <- h4; auto. Qed. Lemma set_infnt_assoc_inf_p : forall (F:Family Btp) (pf:forall A:Ensemble Btp, In F A -> exists p:Btp, inf_p A p), forall q:Btp, inf_p (Im (full_sig F) (fun A => (proj1_sig (constructive_definite_description _ (ex_inf_unq_p _ _ _ (pf _ (proj2_sig A))))))) q -> inf_p (FamilyUnion F) q. intros F h1 q h2. rewrite inf_p_iff in h2. rewrite inf_p_iff. unfold ba_conv_set, ba_conv_type in h2. rewrite transfer_dep_eq_refl in h2. apply (set_infnt_assoc_inf F h1 (B := ba_conv Bp)). term1 h2. gterm1. redterm0 c. redterm0 c0. assert (h3:c1 = c2). unfold c1, c2. apply functional_extensionality; intro; simpl. destruct constructive_definite_description as [e h3]. destruct constructive_definite_description as [f h4]. simpl. eapply inf_unq_p. apply h3. apply h4. fold c2; rewrite <- h3; auto. Qed. Lemma sup_add_assoc_p : forall (A:Ensemble Btp) (a p:Btp), sup_p A p -> sup_p (Add A a) (p%+a). intros A a p. do 2 rewrite sup_p_iff. apply sup_add_assoc. Qed. Lemma inf_add_assoc_p : forall (A:Ensemble Btp) (a p:Btp), inf_p A p -> inf_p (Add A a) (p%*a). intros A a p. do 2 rewrite inf_p_iff. apply inf_add_assoc. Qed. Lemma set_infnt_distr_1_sup_p : forall (A:Ensemble Btp) (p q:Btp), sup_p A q -> sup_p (Im A (fun x => p%*x)) (p%*q). intros A p q. do 2 rewrite sup_p_iff. apply set_infnt_distr_1_sup. Qed. Lemma set_infnt_distr_1_inf_p : forall (A:Ensemble Btp) (p q:Btp), inf_p A q -> inf_p (Im A (fun x => p%+x)) (p%+q). intros A p q. do 2 rewrite inf_p_iff. apply set_infnt_distr_1_inf. Qed. End ParametricAnalogues'. (*Theorems*) Arguments sup_Sup_iff_p [T] [Bp] _ _. Arguments inf_Inf_iff_p [T] [Bp] _ _. Arguments set_infnt_de_mor1_p [T] [Bp] _ _ _. Arguments set_infnt_de_mor2_p [T] [Bp] _ _ _. Arguments set_infnt_de_mor3_p [T] [Bp] _ _ _. Arguments set_infnt_de_mor4_p [T] [Bp] _ _ _. Arguments sup_set_complete_ex_p [T] [Bp] _ _. Arguments sup_set_complete_compat_p [T] [Bp] _ _. Arguments inf_set_complete_ex_p [T] [Bp] _ _. Arguments set_complete_sup_p [T] [Bp] _ _. Arguments set_complete_inf_p [T] [Bp] _ _. Arguments set_infnt_assoc_sup_p [T] [Bp] _ _ _ _. Arguments set_infnt_assoc_inf_p [T] [Bp] _ _ _ _. Arguments sup_add_assoc_p [T] [Bp] _ _ _ _. Arguments inf_add_assoc_p [T] [Bp] _ _ _ _. Arguments set_infnt_distr_1_sup_p [T] [Bp] _ _ _ _. Arguments set_infnt_distr_1_inf_p [T] [Bp] _ _ _ _. (*Definitions*) Arguments sup_set_complete_p [T] [Bp] _ _. Arguments inf_set_complete_p [T] [Bp] _ _.bool2-v0-3/BoolAlgBasics.v0000644000175000017500000040440313575574055016202 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-2019, Daniel Wyckoff.*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Export Ensembles. Require Export Constructive_sets. Require Export Setoid. Require Import List. Require Import EnsemblesImplicit. Require Import ImageImplicit. Require Import Description. Require Import SetUtilities. Require Import FiniteMaps. Require Import FunctionalExtensionality. Require Import TypeUtilities. Require Import LogicUtilities. Require Import DecidableDec. Require Import FunctionProperties. Require Import SetUtilities2. Require Import Infinite_sets. Require Import ListUtilities. Require Import Bool. Require Import EqDec. (*Version 0.3 Alpha bool2 -- certified -- Version 8.4 pl4 Coq.*) Delimit Scope ba_scope with Bool_Alg. Record Bconst : Type := {Btype : Type; BS : Ensemble Btype; Bplus : Btype -> Btype -> Btype; Btimes : Btype -> Btype -> Btype; Bone : Btype; Bzero : Btype; Bcomp : Btype -> Btype}. Infix "+" := (Bplus _) (at level 50, left associativity) : ba_scope. Infix "*" := (Btimes _) (at level 40, left associativity) : ba_scope. Notation "0" := (Bzero _) : ba_scope. Notation "1" := (Bone _) : ba_scope. Notation "- x" := ((Bcomp _) x) : ba_scope. Open Scope ba_scope. Record Bool_Alg : Type := {Bc : Bconst; und_set : (BS Bc) = Full_set (Btype Bc); assoc_sum : forall n m p:(Btype Bc), n + (m + p) = n + m + p; assoc_prod : forall n m p:(Btype Bc), n * (m * p) = n * m * p; comm_sum : forall n m:(Btype Bc), n + m = m + n; comm_prod : forall n m:(Btype Bc), n * m = m * n; abs_sum : forall n m:(Btype Bc), n + (n * m) = n; abs_prod: forall n m:(Btype Bc), n * (n + m) = n; dist_sum : forall n m p:(Btype Bc), p*(n + m) = p * n + p * m; dist_prod: forall n m p:(Btype Bc), p+(n * m) = (p+n) * (p+m); comp_sum: forall n:(Btype Bc), n + (- n) = 1; comp_prod: forall n:(Btype Bc), n * (- n) = 0}. Hint Resolve assoc_sum assoc_prod comm_sum comm_prod abs_sum abs_prod dist_sum dist_prod comp_sum comp_prod. (*Note, the reason there's a sepearate BS from (Full_set Btype) has to do with building specific Boolean alebras, when sometimes you don't know right away that your Boolean type's Full_set does indeed qualify as a Boolean algebra, which might only follow from proving each axiom. Or at least something like that happened ocne a long time ago!*) Lemma bconst_ext : forall (Bc Bc':Bconst) (pf:Btype Bc = Btype Bc'), BS Bc = transfer_dep_r pf (BS Bc') -> Bplus Bc = transfer_dep_r (U:=fun T=>T->T->T) pf (Bplus Bc') -> Btimes Bc = transfer_dep_r (U:=fun T=>T->T->T) pf (Btimes Bc') -> Bone Bc = transfer_dep_r (U:=id) pf (Bone Bc') -> Bzero Bc = transfer_dep_r (U:=id) pf (Bzero Bc') -> Bcomp Bc = transfer_dep_r (U:=fun T=>T->T) pf (Bcomp Bc') -> Bc = Bc'. intros Bc Bc' h1 h2 h3 h4 h5 h6 h7. destruct Bc, Bc'. simpl. simpl in h1, h2, h3, h4, h5, h6, h7. subst. do 6 rewrite transfer_dep_r_eq_refl. reflexivity. Qed. Definition bt (B:Bool_Alg) := Btype (Bc B). Definition ba_ens (B:Bool_Alg) := Full_set (bt B). Definition bzero (B:Bool_Alg) := Bzero (Bc B). Definition bone (B:Bool_Alg) := Bone (Bc B). (*Throughout the development there's a need for handling families of separate algebras or finding extensions of a given algebra, and for this purpose, [Bool_Alg] doesn't suffice, since [Ensemlbe Bool_Alg] yields a universe inconsistency error. I introduce a "dual" record pair Bconst_p / Bool_Alg_p, the p standing for parametric. The difference is that in these structures there is a parametric type T, which is necessary to compare elements in distinct algebras or taking extensions. Every statment about a Bool_Alg has a corresponding statement about Bool_Alg_p (but not vice versa).*) Delimit Scope ba_p_scope with Bool_Alg_p. Record Bconst_p (T:Type) : Type := {A_p : Ensemble T; Btype_p := sig_set A_p; BS_p : Ensemble Btype_p; Bplus_p : Btype_p -> Btype_p -> Btype_p; Btimes_p : Btype_p -> Btype_p -> Btype_p; Bone_p : Btype_p; Bzero_p : Btype_p; Bcomp_p : Btype_p -> Btype_p}. Infix "%+" := (Bplus_p _ _) (at level 50, left associativity) : ba_p_scope. Infix "%*" := (Btimes_p _ _) (at level 40, left associativity) : ba_p_scope. Notation "%0" := (Bzero_p _ _) : ba_p_scope. Notation "%1" := (Bone_p _ _) : ba_p_scope. Notation "%- x" := ((Bcomp_p _ _) x) (at level 30) : ba_p_scope. Open Scope ba_p_scope. Definition bconst_conv {T:Type} (Bc_p:Bconst_p T) : Bconst := Build_Bconst (sig_set (A_p T Bc_p)) (BS_p T Bc_p) (Bplus_p T Bc_p) (Btimes_p T Bc_p) (Bone_p T Bc_p) (Bzero_p T Bc_p) (Bcomp_p T Bc_p). Lemma bconst_ext_p : forall {T:Type} (Bcp Bcp':Bconst_p T) (pf:A_p T Bcp = A_p T Bcp'), BS_p T Bcp = transfer_dep_r (sig_set_eq _ _ pf) (BS_p T Bcp') -> Bplus_p T Bcp = transfer_dep_r (sig_set_eq _ _ pf) (U:=fun T=>T->T->T) (Bplus_p T Bcp') -> Btimes_p T Bcp = transfer_dep_r (sig_set_eq _ _ pf) (U:=fun T=>T->T->T) (Btimes_p T Bcp') -> Bone_p T Bcp = transfer_dep_r (U:=id) (sig_set_eq _ _ pf) (Bone_p T Bcp') -> Bzero_p T Bcp = transfer_dep_r (U:=id) (sig_set_eq _ _ pf) (Bzero_p T Bcp') -> Bcomp_p T Bcp = transfer_dep_r (U:=fun T=>T->T) (sig_set_eq _ _ pf) (Bcomp_p T Bcp') -> Bcp = Bcp'. intros T Bcp Bcp' h0 h1 h2 h3 h4 h5 h6. destruct Bcp, Bcp'. simpl. simpl in h0, h1. destruct h0. simpl in h1, h2, h3, h4, h5, h6. unfold sig_set_eq in h1, h2, h3, h4, h5, h6. unfold eq_ind_r in h1, h2, h3, h4, h5, h6. unfold eq_ind in h1, h2, h3, h4, h5, h6. simpl in h1, h2, h3, h4, h5, h6. rewrite transfer_dep_r_eq_refl in h1, h2, h3, h4, h5, h6. subst. reflexivity. Qed. Record Bool_Alg_p (T:Type) : Type := {Bc_p : Bconst_p T; und_set_p : (BS_p T Bc_p) = Full_set (Btype_p T Bc_p); assoc_sum_p : forall n m p:(Btype_p T Bc_p), n %+ (m %+ p) = n %+ m %+ p; assoc_prod_p : forall n m p:(Btype_p T Bc_p), n %* (m %* p) = n %* m %* p; comm_sum_p : forall n m:(Btype_p T Bc_p), n %+ m = m %+ n; comm_prod_p : forall n m:(Btype_p T Bc_p), n %* m = m %* n; abs_sum_p : forall n m:(Btype_p T Bc_p), n %+ (n %* m) = n; abs_prod_p : forall n m:(Btype_p T Bc_p), n %* (n %+ m) = n; dist_sum_p : forall n m p:(Btype_p T Bc_p), p%*(n %+ m) = p %* n %+ p %* m; dist_prod_p : forall n m p:(Btype_p T Bc_p), p%+(n %* m) = (p%+n) %* (p%+m); comp_sum_p : forall n:(Btype_p T Bc_p), n %+ (%- n) = %1; comp_prod_p : forall n:(Btype_p T Bc_p), n %* (%- n) = %0}. Definition ba_p_ens {T:Type} (Bp:Bool_Alg_p T) := A_p T (Bc_p T Bp). Definition btp {T:Type} (Bp:Bool_Alg_p T) := Btype_p T (Bc_p T Bp). Definition bzero_p {T:Type} (Bp:Bool_Alg_p T) := Bzero_p T (Bc_p T Bp). Definition bone_p {T:Type} (Bp:Bool_Alg_p T) := Bone_p T (Bc_p T Bp). (*Eventually a hundred years from now, someone might formalize the fact that Boolean algebras is a decidable first order theory, so it's a good idea to isolate decidable equality, and to begrudge LEM!*) Axiom ba_eq_dec : forall {B:Bool_Alg} (x y:bt B), {x = y} + {x <> y}. Lemma ba_eq_dec_sig : forall {B:Bool_Alg} {E:Ensemble (bt B)} (x y:sig_set E), {x = y} + {x <> y}. intros ? ? x y. destruct x as [x]; subst. destruct y as [y]; subst. destruct (ba_eq_dec x y) as [|he]; subst. left. apply proj1_sig_injective; auto. right. contradict he. apply exist_injective in he; auto. Qed. Lemma incl_im_proj1_sig_ba_p_ens : forall {T:Type} {Bp:Bool_Alg_p T} (S:Ensemble (btp Bp)), Included (im_proj1_sig S) (ba_p_ens Bp). intros T Bp S. red. intros x h1. destruct h1 as [x h1]. subst. apply proj2_sig. Qed. Definition ba_conv {T:Type} (B:Bool_Alg_p T) : Bool_Alg. pose (bconst_conv (Bc_p T B)) as Bc. refine (Build_Bool_Alg Bc (und_set_p T B) (assoc_sum_p T B) (assoc_prod_p T B) (comm_sum_p T B) (comm_prod_p T B) (abs_sum_p T B) (abs_prod_p T B) (dist_sum_p T B) (dist_prod_p T B) (comp_sum_p T B) (comp_prod_p T B)). Defined. Lemma ba_p_ens_eq : forall {T:Type} (Bp:Bool_Alg_p T), ba_p_ens Bp = im_proj1_sig (ba_ens (ba_conv Bp)). intros T Bp. unfold ba_p_ens, ba_ens. apply Extensionality_Ensembles. red. split. red. intros x h1. apply Im_intro with (exist _ _ h1). constructor. simpl. reflexivity. red. intros x h1. destruct h1 as [x h1]. subst. apply proj2_sig. Qed. Lemma ba_conv_type : forall {T:Type} (Bp:Bool_Alg_p T), btp Bp = bt (ba_conv Bp). auto. Defined. Definition ba_conv_elt {T:Type} {Bp:Bool_Alg_p T} (x:btp Bp) := transfer (ba_conv_type Bp) x. Definition ba_conv_elt' {T:Type} {Bp:Bool_Alg_p T} (x:btp Bp) : bt (ba_conv Bp) := x. Lemma ba_conv_elt_compat' : forall {T:Type} {Bp:Bool_Alg_p T} (x:btp Bp), ba_conv_elt' x = ba_conv_elt x. auto. Qed. Lemma ba_conv_elt_eq : forall {T:Type} {Bp:Bool_Alg_p T} (x:btp Bp), x = ba_conv_elt x. intros T B x. unfold ba_conv_type. unfold transfer. unfold eq_rect_r. simpl. reflexivity. Qed. Lemma ba_eq_dec_p : forall {T:Type} {Bp:Bool_Alg_p T} (x y:btp Bp), {x = y} + {x <> y}. intros T Bp x y. pose proof (ba_eq_dec (ba_conv_elt x) (ba_conv_elt y)) as h1. destruct h1; [left | right]; try constructor; auto. Defined. Lemma ba_eq_dec_p_eq : forall {T:Type} {Bp:Bool_Alg_p T}, @ba_eq_dec_p T Bp = @ba_eq_dec (@ba_conv T Bp). intros T Bp. unfold ba_eq_dec_p. apply functional_extensionality_dep. intro x. apply functional_extensionality_dep. intro y. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. pose proof hl as hl'. unfold ba_conv_elt, ba_conv_type in hl'. do 2 rewrite transfer_eq_refl in hl'. subst. destruct (@ba_eq_dec (ba_conv Bp) y y) as [h1 | h1]. f_equal. apply proof_irrelevance. contradict h1; auto. unfold ba_conv_elt, ba_conv_type in hr. pose proof hr as hr'. unfold ba_conv_elt, ba_conv_type in hr'. do 2 rewrite transfer_eq_refl in hr'. destruct (@ba_eq_dec (ba_conv Bp) x y) as [h1 | h1]. contradiction. f_equal. apply proof_irrelevance. Qed. Lemma ba_eq_dec_sig_p : forall {T:Type} {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} (x y:sig_set E), {x = y} + {x <> y}. intros ? ? ? x y. destruct x as [x]; subst. destruct y as [y]; subst. destruct (ba_eq_dec_p x y) as [|he]; subst. left. apply proj1_sig_injective; auto. right. contradict he. apply exist_injective in he; auto. Qed. Lemma ba_conv_plus : forall {T:Type} (Bp:Bool_Alg_p T) (x y:btp Bp), x %+ y = (ba_conv_elt x) + (ba_conv_elt y). simpl. intros T B x y. pose proof (ba_conv_elt_eq x) as h1. simpl in h1. pose proof (ba_conv_elt_eq y) as h2. simpl in h2. rewrite <- h1. rewrite <- h2. reflexivity. Qed. Lemma ba_conv_times : forall {T:Type} (Bp:Bool_Alg_p T) (x y:btp Bp), x %* y = (ba_conv_elt x) * (ba_conv_elt y). simpl. intros T B x y. pose proof (ba_conv_elt_eq x) as h1. simpl in h1. pose proof (ba_conv_elt_eq y) as h2. simpl in h2. rewrite <- h1. rewrite <- h2. reflexivity. Qed. Lemma ba_conv_comp : forall {T:Type} (Bp:Bool_Alg_p T) (x:btp Bp), %- x = -(ba_conv_elt x). simpl. intros T B x. pose proof (ba_conv_elt_eq x) as h1. simpl in h1. rewrite <- h1. reflexivity. Qed. Lemma ba_conv_one : forall {T:Type} (Bp:Bool_Alg_p T), Bone_p T (Bc_p T Bp) = Bone (Bc (ba_conv Bp)). simpl. auto. Qed. Lemma ba_conv_zero : forall {T:Type} (Bp:Bool_Alg_p T), Bzero_p T (Bc_p T Bp) = Bzero (Bc (ba_conv Bp)). simpl. auto. Qed. Definition ba_conv_set {T:Type} {Bp:Bool_Alg_p T} (A:Ensemble (btp Bp)) := (transfer_dep (ba_conv_type Bp) A). Definition ba_conv_set' {T:Type} {Bp:Bool_Alg_p T} (A:Ensemble (btp Bp)) : Ensemble (bt (ba_conv Bp)) := A. Lemma bc_subst : forall (Bc Bc':Bconst) (P:Bconst->Prop), Bc = Bc' -> (P Bc <-> P Bc'). intros; subst. tauto. Qed. Definition ba_conv_list {T:Type} {Bp:Bool_Alg_p T} (l:list (Btype_p T (Bc_p T Bp))) := transfer_dep (ba_conv_type Bp) l. Definition ba_conv_faml {T:Type} {Bp:Bool_Alg_p T} (l:faml (btp Bp)) := transfer_dep (U:=(fun V=>faml V)) (ba_conv_type Bp) l. Lemma length_ba_conv_list_eq : forall {T:Type} {Bp:Bool_Alg_p T} (l:list (btp Bp)), length (ba_conv_list l) = length l. auto. Qed. Lemma lt_length_ba_conv_list : forall {T:Type} {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat), m < length l -> m < length (ba_conv_list l). auto. Qed. Lemma le_length_ba_conv_list : forall {T:Type} {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat), m <= length l -> m <= length (ba_conv_list l). auto. Qed. Lemma nth_lt_ba_conv_list : forall {T:Type} {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat) (pf:m < length l), let pf' := lt_length_ba_conv_list _ _ pf in nth_lt l m pf = nth_lt (ba_conv_list l) m pf'. intros. apply nth_lt_functional3. Qed. Lemma firstn_ba_conv_list : forall {T:Type} {Bp:Bool_Alg_p T} (l:list (btp Bp)) (m:nat), firstn m (ba_conv_list l) = firstn m l. auto. Qed. Lemma count_occ_ba_conv_list : forall {T:Type} {Bp:Bool_Alg_p T} (l:list (btp Bp)) (x:btp Bp), count_occ ba_eq_dec_p l x = count_occ ba_eq_dec (ba_conv_list l) (ba_conv_elt x). intros. f_equal. apply ba_eq_dec_p_eq. Qed. Lemma count_occ_ba_conv_list' : forall {T:Type} {Bp:Bool_Alg_p T} (l:list (btp Bp)) (x:btp Bp), count_occ ba_eq_dec_p l x = count_occ (@ba_eq_dec (ba_conv Bp)) l x. intros; rewrite ba_eq_dec_p_eq; auto. Qed. Lemma linds_occ_ba_conv_list : forall {T:Type} {Bp:Bool_Alg_p T} (l:list (btp Bp)) (x:btp Bp), linds_occ ba_eq_dec_p l x = linds_occ ba_eq_dec (ba_conv_list l) (ba_conv_elt x). intros T Bp l x. destruct (in_dec ba_eq_dec_p x l) as [h1 | h1]. rewrite linds_occ_eq_iff. ex_goal. rewrite length_linds_occ. apply count_occ_ba_conv_list. exists hex. intros m h2. rewrite (nth_lt_linds_occ ba_eq_dec (ba_conv_list l) (ba_conv_elt x) h1). rewrite ba_eq_dec_p_eq. apply lind_occ_functional; auto. rewrite nin_linds_occ_eq_iff; auto. rewrite (nin_linds_occ_eq_iff ba_eq_dec (ba_conv_list l) (ba_conv_elt x)); auto. Grab Existential Variables. assumption. Qed. Lemma linds_occ_ba_conv_list' : forall {T:Type} {Bp:Bool_Alg_p T} (l:list (btp Bp)) (x:btp Bp), linds_occ ba_eq_dec_p l x = linds_occ (@ba_eq_dec (ba_conv Bp)) l x. intros; rewrite ba_eq_dec_p_eq; auto. Qed. Lemma linds_ba_conv_list : forall {T:Type} {Bp:Bool_Alg_p T} (l:list (btp Bp)), linds ba_eq_dec_p l = linds ba_eq_dec (ba_conv_list l). intros; unfold linds; f_equal. apply functional_extensionality. intro; apply linds_occ_ba_conv_list'. f_equal. apply ba_eq_dec_p_eq. Qed. Lemma linds_ba_conv_list' : forall {T:Type} {Bp:Bool_Alg_p T} (l:list (btp Bp)), linds ba_eq_dec_p l = linds (@ba_eq_dec (ba_conv Bp)) l. intros; rewrite ba_eq_dec_p_eq; auto. Qed. Definition ba_conv_fun {T:Type} {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp) : bt (ba_conv Ap)->bt (ba_conv Bp). intro x. refine (f x). Defined. Definition ba_conv_fun1 {T:Type} {Ap:Bool_Alg_p T} {T':Type} (f:btp Ap->T') : bt (ba_conv Ap)->T'. intro x. refine (f x). Defined. Definition ba_conv_fun2 {T':Type} {T:Type} {Ap:Bool_Alg_p T} (f:T'->btp Ap) : T'-> bt (ba_conv Ap). intro x. refine (f x). Defined. Definition ba_conv_sig_fun1 {T:Type} {Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)} {U:Type} (g:sig_set E->U) : sig_set (ba_conv_set E)->U. intro p. refine (g p). Defined. Definition bc_sig_set_conv {T:Type} (A:Ensemble T) (Bc:Bconst) (pf:Btype Bc = sig_set A) : Bconst_p T := Build_Bconst_p T A (transfer_dep pf (BS Bc)) (transfer_dep (U:=fun T=>T->T->T) pf (Bplus Bc)) (transfer_dep (U:=fun T=>T->T->T) pf (Btimes Bc)) (transfer pf (Bone Bc)) (transfer pf (Bzero Bc)) (transfer_dep (U:=fun T=>T->T) pf (Bcomp Bc)). Lemma ba_sig_set_conv_und_set : forall {T:Type} {A:Ensemble T} {B:Bool_Alg} {pf:bt B = sig_set A}, BS_p T (bc_sig_set_conv _ _ pf) = Full_set (Btype_p T (bc_sig_set_conv _ _ pf)). intros T A B h1. simpl. apply Extensionality_Ensembles. red. split. red. intros x h2. constructor. red. intros x h2. clear h2. unfold bc_sig_set_conv, Btype_p in x. simpl in x. rewrite <- (transfer_undoes_transfer_r h1). rewrite <- transfer_in. rewrite und_set. constructor. Qed. Lemma transfer_dep_fun_plus : forall {T:Type} {A:Ensemble T} {Bc:Bconst} {pf:Btype Bc = sig_set A}, transfer_dep (U:=fun T=>T->T->T) pf (Bplus Bc) = fun (x:sig_set A) (y:sig_set A) => transfer pf ((transfer_r pf x) + (transfer_r pf y)). intros T A Bc h1. destruct Bc. simpl in h1. subst. rewrite transfer_dep_eq_refl. simpl. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite transfer_eq_refl. do 2 rewrite transfer_r_eq_refl. reflexivity. Qed. Lemma transfer_dep_fun_times : forall {T:Type} {A:Ensemble T} {Bc:Bconst} {pf:Btype Bc = sig_set A}, transfer_dep (U:=fun T=>T->T->T) pf (Btimes Bc) = fun (x:sig_set A) (y:sig_set A) => transfer pf ((transfer_r pf x) * (transfer_r pf y)). intros T A Bc h1. destruct Bc. simpl in h1. subst. rewrite transfer_dep_eq_refl. simpl. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite transfer_eq_refl. do 2 rewrite transfer_r_eq_refl. reflexivity. Qed. Lemma transfer_dep_fun_comp : forall {T:Type} {A:Ensemble T} {Bc:Bconst} {pf:Btype Bc = sig_set A}, transfer_dep (U:=fun T=>T->T) pf (Bcomp Bc) = fun (x:sig_set A) => transfer pf (- (transfer_r pf x)). intros T A Bc h1. destruct Bc. simpl in h1. subst. rewrite transfer_dep_eq_refl. simpl. apply functional_extensionality. intro x. rewrite transfer_eq_refl, transfer_r_eq_refl. reflexivity. Qed. Lemma ba_sig_set_conv_assoc_sum : forall {T:Type} {A:Ensemble T} {B:Bool_Alg} {pf:bt B = sig_set A}, forall n m p : Btype_p T (bc_sig_set_conv _ _ pf) , n %+ (m %+ p) = n %+ m %+ p. intros T A B h1 n m p. simpl. rewrite transfer_dep_fun_plus. do 2 rewrite transfer_r_undoes_transfer. rewrite assoc_sum. reflexivity. Qed. Lemma ba_sig_set_conv_assoc_prod : forall {T:Type} {A:Ensemble T} {B:Bool_Alg} {pf:bt B = sig_set A}, forall n m p : Btype_p T (bc_sig_set_conv _ _ pf), n %* (m %* p) = n %* m %* p. intros T A B h1 n m p. simpl. rewrite transfer_dep_fun_times. do 2 rewrite transfer_r_undoes_transfer. rewrite assoc_prod. reflexivity. Qed. Lemma ba_sig_set_conv_comm_sum : forall {T:Type} {A:Ensemble T} {B:Bool_Alg} {pf:bt B = sig_set A}, forall n m : Btype_p T (bc_sig_set_conv _ _ pf), n %+ m = m %+ n. intros T A B h1 n m. simpl. rewrite transfer_dep_fun_plus. rewrite comm_sum. reflexivity. Qed. Lemma ba_sig_set_conv_comm_prod : forall {T:Type} {A:Ensemble T} {B:Bool_Alg} {pf:bt B = sig_set A}, forall n m : Btype_p T (bc_sig_set_conv _ _ pf), n %* m = m %* n. intros T A B h1 n m. simpl. rewrite transfer_dep_fun_times. rewrite comm_prod. reflexivity. Qed. Lemma ba_sig_set_conv_abs_sum : forall {T:Type} {A:Ensemble T} {B:Bool_Alg} {pf:bt B = sig_set A}, forall n m : Btype_p T (bc_sig_set_conv _ _ pf), n %+ n %* m = n. intros T A B h1 n m. simpl. rewrite transfer_dep_fun_plus, transfer_dep_fun_times. rewrite transfer_r_undoes_transfer. rewrite abs_sum. rewrite transfer_undoes_transfer_r. reflexivity. Qed. Lemma ba_sig_set_conv_abs_prod : forall {T:Type} {A:Ensemble T} {B:Bool_Alg} {pf:bt B = sig_set A}, forall n m : Btype_p T (bc_sig_set_conv _ _ pf), n %* (n %+ m) = n. intros T A B h1 n m. simpl. rewrite transfer_dep_fun_plus, transfer_dep_fun_times. rewrite transfer_r_undoes_transfer. rewrite abs_prod. rewrite transfer_undoes_transfer_r. reflexivity. Qed. Lemma ba_sig_set_conv_dist_sum : forall {T:Type} {A:Ensemble T} {B:Bool_Alg} {pf:bt B = sig_set A}, forall n m p : Btype_p T (bc_sig_set_conv _ _ pf), p %* (n %+ m) = p %* n %+ p %* m. intros T A B h1 n m p. simpl. rewrite transfer_dep_fun_plus, transfer_dep_fun_times. do 3 rewrite transfer_r_undoes_transfer. rewrite dist_sum. reflexivity. Qed. Lemma ba_sig_set_conv_dist_prod : forall {T:Type} {A:Ensemble T} {B:Bool_Alg} {pf:bt B = sig_set A}, forall n m p : Btype_p T (bc_sig_set_conv _ _ pf), p %+ n %* m = (p %+ n) %* (p %+ m). intros T A B h1 n m p. simpl. rewrite transfer_dep_fun_plus, transfer_dep_fun_times. do 3 rewrite transfer_r_undoes_transfer. rewrite dist_prod. reflexivity. Qed. Lemma ba_sig_set_conv_comp_sum : forall {T:Type} {A:Ensemble T} {B:Bool_Alg} {pf:bt B = sig_set A}, forall n : Btype_p T (bc_sig_set_conv _ _ pf), n %+ %-n = %1. intros T A B h1 n. simpl. rewrite transfer_dep_fun_plus, transfer_dep_fun_comp. rewrite transfer_r_undoes_transfer. rewrite comp_sum. reflexivity. Qed. Lemma ba_sig_set_conv_comp_prod : forall {T:Type} {A:Ensemble T} {B:Bool_Alg} {pf:bt B = sig_set A}, forall n : Btype_p T (bc_sig_set_conv _ _ pf), n %* %-n = %0. intros T A B h1 n. simpl. rewrite transfer_dep_fun_times, transfer_dep_fun_comp. rewrite transfer_r_undoes_transfer. rewrite comp_prod. reflexivity. Qed. Definition ba_sig_set_conv {T:Type} (A:Ensemble T) (B:Bool_Alg) (pf:bt B = sig_set A) : Bool_Alg_p T := Build_Bool_Alg_p T (bc_sig_set_conv _ _ pf) ba_sig_set_conv_und_set ba_sig_set_conv_assoc_sum ba_sig_set_conv_assoc_prod ba_sig_set_conv_comm_sum ba_sig_set_conv_comm_prod ba_sig_set_conv_abs_sum ba_sig_set_conv_abs_prod ba_sig_set_conv_dist_sum ba_sig_set_conv_dist_prod ba_sig_set_conv_comp_sum ba_sig_set_conv_comp_prod. Lemma ba_sig_set_conv_bc_p_compat : forall {T:Type} (A:Ensemble T) (B:Bool_Alg) (pf:bt B = sig_set A), Bc_p T (ba_sig_set_conv A B pf) = bc_sig_set_conv A (Bc B) pf. intros; auto. Qed. Section Basics. Variable B:Bool_Alg. Let Bt := bt B. (*We start out with some tactics for simplifying Boolean algebraic expressions, but it turns out that to use these tactics requires module functionality, which is very inconvenient, and using these tactics is like the equivalent of doing a quicksort on a small list. so I have included these for reference purposes only. I use them in a few places in this file only. Someone can make these commutativity/associativity tactics more ubiquitious if they want, by incoporating them into modules.*) Section assoc. (*************) (* All these functions/tactics in "Section assoc" and "Section commut" (with two operations) are adapted from the one-operation versions in Coq'Art Section 16.3.*) (* Binary tree with two types of nodes corresponding to two binary operations, plus and times, and a leaf which contains an integer rank corresponding to its position in a derived list. *) Inductive bin : Set := pnode : bin -> bin -> bin | tnode : bin -> bin -> bin | leaf : nat-> bin. Fixpoint flatten_aux_p (p fin:bin){struct p} : bin := match p with | pnode x1 x2 => flatten_aux_p x1 (flatten_aux_p x2 fin) | x => pnode x fin end. Fixpoint flatten_aux_t (t fin:bin){struct t} : bin := match t with | tnode x1 x2 => flatten_aux_t x1 (flatten_aux_t x2 fin) | x => tnode x fin end. Fixpoint flatten (t:bin) : bin := match t with | pnode (tnode x1 x2 as y1) (tnode x3 x4 as y2) => pnode (flatten y1) (flatten y2) | tnode (pnode x1 x2 as y1) (pnode x3 x4 as y2) => tnode (flatten y1) (flatten y2) | tnode x1 x2 => flatten_aux_t (flatten x1) (flatten x2) | pnode x1 x2 => flatten_aux_p (flatten x1) (flatten x2) | x => x end. Fixpoint bin_Bt (l: list Bt) (def:Bt) (t:bin) {struct t} : Bt := match t with | pnode t1 t2 => (bin_Bt l def t1) + (bin_Bt l def t2) | tnode t1 t2 => (bin_Bt l def t1) * (bin_Bt l def t2) | leaf n => nth n l def end. Theorem flatten_aux_valid_Bt_p : forall (l:list Bt)(def:Bt)(t t':bin), (bin_Bt l def t) + (bin_Bt l def t') = bin_Bt l def (flatten_aux_p t t'). intros l def t. elim t; simpl; auto. intros t1 iht1 t2 iht2 t'. rewrite <- iht1. rewrite <- iht2. rewrite (assoc_sum B). reflexivity. Qed. Theorem flatten_aux_valid_Bt_t : forall (l:list Bt)(def:Bt)(t t':bin), (bin_Bt l def t) * (bin_Bt l def t') = bin_Bt l def (flatten_aux_t t t'). intros l def t. elim t; simpl; auto. intros t1 iht1 t2 iht2 t'. rewrite <- iht1. rewrite <- iht2. rewrite (assoc_prod B). reflexivity. Qed. Theorem flatten_valid_Bt : forall (l:list Bt)(def:Bt)(t:bin), bin_Bt l def t = bin_Bt l def (flatten t). intros l def t. elim t; simpl; auto. intros t1 iht1 t2 iht2. destruct t1. destruct t2. rewrite iht1. rewrite iht2. rewrite flatten_aux_valid_Bt_p. reflexivity. rewrite iht1. rewrite iht2. rewrite flatten_aux_valid_Bt_p. reflexivity. rewrite iht1. rewrite iht2. rewrite flatten_aux_valid_Bt_p. reflexivity. destruct t2. rewrite iht1. rewrite iht2. rewrite flatten_aux_valid_Bt_p. reflexivity. rewrite iht1. rewrite iht2. unfold bin_Bt at 3. fold bin_Bt. reflexivity. rewrite iht1. rewrite iht2. rewrite flatten_aux_valid_Bt_p. reflexivity. rewrite iht1. rewrite iht2. rewrite flatten_aux_valid_Bt_p. reflexivity. intros t1 iht1 t2 iht2. destruct t1. destruct t2. unfold bin_Bt at 3. fold bin_Bt. rewrite <- iht1; rewrite <- iht2; reflexivity. rewrite iht1; rewrite iht2; rewrite flatten_aux_valid_Bt_t; reflexivity. rewrite iht1. rewrite iht2. rewrite flatten_aux_valid_Bt_t. reflexivity. rewrite iht1. rewrite iht2. rewrite flatten_aux_valid_Bt_t. reflexivity. rewrite iht1. rewrite iht2. rewrite flatten_aux_valid_Bt_t. reflexivity. Qed. Theorem flatten_valid_Bt_2 : forall (t t':bin)(l:list Bt)(def:Bt), bin_Bt l def (flatten t) = bin_Bt l def (flatten t') -> bin_Bt l def t = bin_Bt l def t'. intros t t' l def; rewrite (flatten_valid_Bt l def t); rewrite (flatten_valid_Bt l def t'). trivial. Qed. End assoc. Ltac term_list l v := match v with | (?X1 + ?X2) => let l1 := term_list l X2 in term_list l1 X1 | (?X1 * ?X2) => let l1 := term_list l X2 in term_list l1 X1 | ?X1 => constr:(cons X1 l) end. Ltac compute_rank l n v := match l with | (cons ?X1 ?X2) => let t1 := constr:X2 in match constr:(X1 = v) with | (?X1 = ?X1) => n | _ => compute_rank t1 (S n) v end end. Ltac model_aux l v := match v with | (?X1 + ?X2) => let r1 := model_aux l X1 with r2 := model_aux l X2 in constr:(pnode r1 r2) | (?X1 * ?X2) => let r1 := model_aux l X1 with r2 := model_aux l X2 in constr:(tnode r1 r2) | ?X1 => let n := compute_rank l 0%nat X1 in constr:(leaf n) | _ => constr:(leaf 0%nat) end. Ltac model_Bt def v := let l := term_list (nil (A:=Bt)) v in let t := model_aux l v in constr:(bin_Bt l def t). Ltac assoc_eq:= match goal with | [|- (?X1 = ?X2) ] => let term1 := model_Bt X1 X1 with term2 := model_Bt X1 X2 in (change (term1 = term2); apply flatten_valid_Bt_2; auto) end. (* Variables a b c d e f g h i j k l: Bt. Goal ((a+(b+c))*(d+(e+f)))*(g+(h+i)) = ((a+b)+c) * (((d+e)+f)*((g+h)+i)). assoc_eq. Qed. (*Checks*) Goal ((a*(b*c))+(d*(e*f)))+(g*(h*i)) = ((a*b)*c) + (((d*e)*f)+((g*h)*i)). assoc_eq. Qed. (*Checks*) Goal (a+(b*(c*d)))+(e*f*(g+h))+(i*j*(k+l)) = (a + ((b*c)*d)) + ((e*(f*(g+h))) + (i*(j*(k+l)))). assoc_eq. (*Checks*) Goal (a+b+c+d)*(e+f+g+h)*(i+j+k+l) = ((a+b)+(c+d))*(((e+f)+(g+h))*((i+j)+(k+l))). assoc_eq. (*Checks*) Goal (a*b*c*d)+(e*f*g*h)+(i*j*k*l) = ((a*b)*(c*d))+(((e*f)*(g*h))+((i*j)*(k*l))). assoc_eq. (*Checks*) *) Section commut. Fixpoint nat_eq_bool (m n:nat){struct n} : bool := match m, n with | 0, 0 => true | S m, 0 => false | 0, S n => false | S m, S n => nat_eq_bool m n end. Fixpoint nat_lt_bool (m n:nat){struct n} : bool := match m, n with | _, 0 => false | 0, S n => true | S m, S n => nat_lt_bool m n end. Fixpoint bin_eq_bool (t t':bin){struct t'} : bool := match t, t' with | (leaf m), (leaf n) => nat_eq_bool m n | (pnode p1 p2), (pnode p'1 p'2) => andb (bin_eq_bool p1 p'1) (bin_eq_bool p2 p'2) | (tnode t1 t2), (tnode t'1 t'2) => andb (bin_eq_bool t1 t'1) (bin_eq_bool t2 t'2) | _, _ => false end. (* The order is arbitrary. Leaves are less than pnodes are less than tnodes, all other cases depend on the constructors. *) Fixpoint bin_lt_bool (t t':bin){struct t'} : bool := match t, t' with | (leaf m), (leaf n) => nat_lt_bool m n | (leaf m), _=> true | (pnode p1 p2), (pnode p'1 p'2) => if (bin_eq_bool p1 p'1) then (bin_lt_bool p2 p'2) else (bin_lt_bool p1 p'1) | (pnode p1 p2), (leaf m) => false | (pnode _ _), (tnode _ _) => true | (tnode t1 t2), (tnode t'1 t'2) => if (bin_eq_bool t1 t'1) then (bin_lt_bool t2 t'2) else (bin_lt_bool t1 t'1) | (tnode t1 t2), _ => false end. Fixpoint insert_bin_p (t0:bin) (t:bin) {struct t} : bin := match t with | leaf m => match (bin_lt_bool t0 t) with | true => pnode t0 t | false => pnode t t0 end | tnode _ _ => match (bin_lt_bool t0 t) with | true => pnode t0 t | false => pnode t t0 end | pnode ((leaf m) as l) t' => match (bin_lt_bool t0 l) with | true => pnode t0 t | false => pnode l (insert_bin_p t0 t') end | pnode ((tnode _ _) as x) t' => match (bin_lt_bool t0 x) with | true => pnode t0 t | false => pnode x (insert_bin_p t0 t') end | t => pnode t0 t end. Fixpoint insert_bin_t (t0:bin)(t:bin){struct t} : bin := match t with | leaf m => match (bin_lt_bool t0 t) with | true => tnode t0 t | false => tnode t t0 end | pnode _ _ => match (bin_lt_bool t0 t) with | true => tnode t0 t | false => tnode t t0 end | tnode ((leaf m) as l) t' => match (bin_lt_bool t0 l) with | true => tnode t0 t | false => tnode l (insert_bin_t t0 t') end | tnode ((pnode _ _) as x) t' => match (bin_lt_bool t0 x) with | true => tnode t0 t | false => tnode x (insert_bin_t t0 t') end | t => tnode t0 t end. Fixpoint sort_bin (t:bin) : bin := match t with | tnode ((leaf n) as l) t' => insert_bin_t l (sort_bin t') | tnode (pnode _ _ as p) t' => insert_bin_t (sort_bin p) (sort_bin t') | pnode ((leaf n) as l) p' => insert_bin_p l (sort_bin p') | pnode (tnode _ _ as t') p' => insert_bin_p (sort_bin t') (sort_bin p') | t => t end. Theorem bin_Bt_ref_p : forall (l:list Bt)(def:Bt)(t1 t2:bin), bin_Bt l def (pnode t1 t2) = (bin_Bt l def t1) + (bin_Bt l def t2). intros l def t1 t2. unfold bin_Bt. fold bin_Bt. reflexivity. Qed. Theorem bin_Bt_ref_t : forall (l:list Bt)(def:Bt)(t1 t2:bin), bin_Bt l def (tnode t1 t2) = (bin_Bt l def t1) * (bin_Bt l def t2). intros l def t1 t2. unfold bin_Bt. fold bin_Bt. reflexivity. Qed. Theorem bin_Bt_leaf : forall (l:list Bt) (def:Bt) n, bin_Bt l def (leaf n) = nth n l def. intro l; induction l; auto. Qed. Theorem insert_eq_p : forall (l:list Bt)(def:Bt)(t0 t:bin), bin_Bt l def (insert_bin_p t0 t) = (bin_Bt l def t0) + (bin_Bt l def t). intros l def t0 t. induction t; unfold insert_bin_p. case t1. intros b b0. apply bin_Bt_ref_p. fold insert_bin_p. intros b b0. case (bin_lt_bool t0 (tnode b b0)). apply bin_Bt_ref_p. rewrite bin_Bt_ref_p. rewrite IHt2. unfold bin_Bt. fold bin_Bt. rewrite (assoc_sum B) with (n:= bin_Bt l def t0). rewrite (comm_sum B) with (n := bin_Bt l def t0) (m := bin_Bt l def b * bin_Bt l def b0). apply (assoc_sum B). fold insert_bin_p. intros n. case (bin_lt_bool t0 (leaf n)). rewrite bin_Bt_ref_p. reflexivity. rewrite bin_Bt_ref_p. rewrite IHt2. rewrite bin_Bt_ref_p. rewrite (assoc_sum B) with (n := bin_Bt l def t0). rewrite (comm_sum B) with (n := bin_Bt l def t0) (m := bin_Bt l def (leaf n)). rewrite (assoc_sum B). reflexivity. case (bin_lt_bool t0 (tnode t1 t2)). rewrite bin_Bt_ref_p. reflexivity. rewrite bin_Bt_ref_p. apply (comm_sum B). case (bin_lt_bool t0 (leaf n)). rewrite bin_Bt_ref_p. reflexivity. rewrite bin_Bt_ref_p. apply (comm_sum B). Qed. Theorem insert_eq_t : forall (l:list Bt)(def:Bt)(t0 t:bin), bin_Bt l def (insert_bin_t t0 t) = (bin_Bt l def t0) * (bin_Bt l def t). intros l def t0 t. induction t; unfold insert_bin_t. case (bin_lt_bool t0 (pnode t1 t2)). rewrite bin_Bt_ref_t. reflexivity. rewrite bin_Bt_ref_t. apply (comm_prod B). fold insert_bin_t. destruct t1. case (bin_lt_bool t0 (pnode t1_1 t1_2)). apply bin_Bt_ref_t. rewrite bin_Bt_ref_t. rewrite IHt2. unfold bin_Bt. fold bin_Bt. rewrite (assoc_prod B) with (n := bin_Bt l def t0). rewrite (comm_prod B) with (n := bin_Bt l def t0) (m := (bin_Bt l def t1_1 + bin_Bt l def t1_2)). apply (assoc_prod B). unfold bin_Bt. fold bin_Bt. reflexivity. case (bin_lt_bool t0 (leaf n)). apply bin_Bt_ref_t. rewrite bin_Bt_ref_t. rewrite IHt2. rewrite bin_Bt_ref_t. rewrite (assoc_prod B) with (n := bin_Bt l def t0). rewrite (comm_prod B) with (n := bin_Bt l def t0) (m := bin_Bt l def (leaf n)). apply (assoc_prod B). case (bin_lt_bool t0 (leaf n)). apply bin_Bt_ref_t. rewrite bin_Bt_ref_t. apply (comm_prod B). Qed. Theorem sort_eq : forall (l: list Bt)(def:Bt)(t:bin), bin_Bt l def (sort_bin t) = bin_Bt l def t. intros l def. induction t. unfold sort_bin. destruct t1. reflexivity. fold sort_bin. unfold sort_bin in IHt1. fold sort_bin in IHt1. rewrite insert_eq_p. rewrite IHt1. rewrite IHt2. rewrite <- bin_Bt_ref_p. reflexivity. fold sort_bin. rewrite insert_eq_p. rewrite IHt2. rewrite bin_Bt_ref_p. reflexivity. unfold sort_bin. destruct t1. fold sort_bin. unfold sort_bin in IHt1. fold sort_bin in IHt1. rewrite insert_eq_t. rewrite IHt1. rewrite IHt2. rewrite <- bin_Bt_ref_t. reflexivity. reflexivity. fold sort_bin. rewrite insert_eq_t. rewrite bin_Bt_ref_t. rewrite IHt2. reflexivity. unfold sort_bin. reflexivity. Qed. Theorem sort_eq_2 : forall (l:list Bt)(def:Bt)(t1 t2:bin), bin_Bt l def (sort_bin t1) = bin_Bt l def (sort_bin t2) -> bin_Bt l def t1 = bin_Bt l def t2. intros l def t1 t2. rewrite sort_eq. rewrite sort_eq. trivial. Qed. End commut. Ltac comm_eq := match goal with | [ |- (?X1 = ?X2) ] => let l := term_list (nil (A := Bt)) X1 in let term1 := model_aux l X1 with term2 := model_aux l X2 in (change (bin_Bt l X1 term1 = bin_Bt l X1 term2); apply flatten_valid_Bt_2; apply sort_eq_2; auto) end. (* (*All check!*) Variables a b c d e f g h i j k l : Bt. Goal (a*b)+(c*d) = (b*a) +(d*c). comm_eq. Goal (a*b)+(d*c)+(e*f*g) = (g*f*e) + (c*d) + (b*a). comm_eq. Goal (a*b*c+(d*e*f+(h*i*j)+k) = k + c*b*a + (f*d*e) + (j*i*h)). comm_eq. Goal a*(b+c)*d = (c+b)*d*a. comm_eq. Goal a+(b*(c+(d*e*f))*g) = ((f*e*d) + c)*b*g + a. comm_eq. *) (*Theorems *) Lemma in_bs : forall x:Bt, Inn (BS (Bc B)) x. intros; rewrite und_set; constructor. Qed. Lemma bc_inj : forall {A A':Bool_Alg}, Bc A = Bc A' -> A = A'. intros A A' h1. destruct A; destruct A'. simpl in h1. subst. f_equal; apply proof_irrelevance. Qed. Lemma dist_sum_r : forall (n m p: Bt), (n + m) * p = n*p + m*p. intros n m p. pose proof (comm_prod _ (n + m) p). pose proof (dist_sum B n m p). assert (h3: p*n + p*m = n*p + m*p). comm_eq. congruence. Qed. Hint Resolve dist_sum_r. Lemma dist_prod_r : forall (n m p: Bt), (n*m) + p = (n+p) * (m + p). intros n m p. pose proof (comm_sum _ (n * m) p). pose proof (dist_prod B n m p). assert (h3: (p+n) * (p+m) = (n+p) * (m+p)). comm_eq. congruence. Qed. Lemma idem_sum : forall b : Bt, b + b = b. intro b. pose proof (abs_sum _ b (b + b)) as h1. pose proof (abs_prod _ b b) as h2. rewrite h2 in h1; auto. Qed. Hint Resolve idem_sum. Lemma idem_prod: forall (b: Bt), b * b = b. intro b. pose proof (abs_prod _ b (b * b)) as h1. pose proof (abs_sum _ b b) as h2. rewrite h2 in h1; auto. Qed. Hint Resolve idem_prod. Lemma comp_prod_r : forall (b:Bt), (-b) * b = 0. intro; rewrite comm_prod; apply comp_prod. Qed. Lemma comp_sum_r : forall (b:Bt), (-b) + b = 1. intro; rewrite comm_sum; apply comp_sum. Qed. Lemma eq_ord : forall (x y: Bt), x + y = y <-> x * y = x. intros x y; split; intro h1. pose proof (abs_prod _ x y) as h2. rewrite h1 in h2; auto. pose proof (abs_sum _ y x) as h2. rewrite comm_prod in h2. rewrite h1, comm_sum in h2; auto. Qed. Hint Resolve eq_ord. Definition le (x y:Bt): Prop := x + y = y. Hint Unfold le. Lemma refl_le : forall (x: Bt), le x x. intro x; pose proof (idem_sum x); red; auto. Qed. Hint Resolve refl_le. Lemma anti_sym_le: forall (x y: Bt), le x y -> le y x -> x = y. intros x y h1 h2; red in h1, h2. rewrite comm_sum in h2. congruence. Qed. Hint Resolve anti_sym_le. Lemma trans_le: forall (x y z : Bt), le x y -> le y z -> le x z. intros x y z h1 h2; red in h1, h2; red. rewrite <- h2, assoc_sum, h1 at 1; auto. Qed. Hint Resolve trans_le. Lemma order_le : Order _ le. constructor. red; apply refl_le. red; apply trans_le. red; apply anti_sym_le. Qed. Lemma times_le : forall (x y : Bt), le (x * y) x. intros x y. red. rewrite eq_ord. rewrite <- assoc_prod. rewrite (comm_prod _ y x). rewrite assoc_prod. rewrite idem_prod. reflexivity. Qed. Lemma le_plus : forall (x y : Bt), le x (x + y). intros x y. red. rewrite assoc_sum. rewrite idem_sum. reflexivity. Qed. Definition lt (x y:Bt) := le x y /\ x <> y. Lemma lt_irrefl : forall (x:Bt), ~lt x x. intros x h1. red in h1. destruct h1 as [h1 h2]. contradict h2; auto. Qed. Lemma strict_order_lt : Strict_Order lt. constructor. red. intros x y z h1 h2. destruct h1 as [h1a h1b], h2 as [h2a h2b]. pose proof (trans_le _ _ _ h1a h2a) as h3. split; auto. intro h4. subst. pose proof (anti_sym_le _ _ h2a h1a) as h4. subst. contradict h1b. reflexivity. red. intro x. apply lt_irrefl. Qed. Lemma ba_le_dec : forall x y:Bt, {le x y} + {~le x y}. intros x y. unfold le. apply ba_eq_dec. Qed. Lemma ba_neq_lt_dec : forall x y:Bt, x <> y -> {lt x y} + {~lt x y}. intros x y h1. unfold lt. destruct (ba_le_dec x y) as [h2 | h2]. left; split; auto. right; contradict h2; destruct h2; auto. Qed. (* b is a lower bound of S in B*) Definition lb (S: Ensemble Bt) (b: Bt) : Prop := forall (s:Bt), Inn S s -> le b s. Hint Unfold lb. (* b is an upper bound of S in B*) Definition ub (S: Ensemble Bt) (b: Bt) : Prop := forall (s:Bt), Inn S s -> le s b. Hint Unfold ub. (* b is the infimum (greatest lower bound) of S in B *) Definition inf (S: Ensemble Bt) (b: Bt) : Prop := lb S b /\ forall b':Bt, lb S b' -> le b' b. Hint Unfold inf. (* b is the supremum (least upper bound) of S in B*) Definition sup (S: Ensemble Bt) (b: Bt) : Prop := ub S b /\ forall b':Bt, ub S b' -> le b b'. Hint Unfold sup. Lemma inf_le : forall (A:Ensemble Bt) (b:Bt), inf A b -> forall b':Bt, Inn A b' -> le b b'. intros A b h1 b' h2. red in h1. destruct h1 as [h1]. red in h1. apply h1; auto. Qed. Lemma le_sup : forall (A:Ensemble Bt) (b:Bt), sup A b -> forall b':Bt, Inn A b' -> le b' b. intros A b h1 b' h2. red in h1. destruct h1 as [h1]. red in h1. apply h1; auto. Qed. Lemma inf_singleton : forall (x:Bt), inf (Singleton x) x. intro x. red. split. (* lower bound *) red. intros y h1. inversion h1. subst. auto. (* greatest lower bound *) intros b' h1. red in h1. pose proof (Singleton_intro _ _ _ (eq_refl x)) as h2. apply h1. assumption. Qed. Lemma sup_singleton : forall (x:Bt), sup (Singleton x) x. intro x. red. split. (* lower bound *) red. intros y h1. inversion h1. subst. auto. (* greatest lower bound *) intros b' h1. red in h1. pose proof (Singleton_intro _ _ _ (eq_refl x)) as h2. apply h1. assumption. Qed. (* B is a lattice under partial order le (sum portion). *) Lemma lat_sum : forall (x y: Bt), sup (Couple x y) (x + y). intros; red; split; red; intros s h1. destruct h1. apply le_plus. rewrite comm_sum. apply le_plus. red in h1. pose proof (h1 x) as h2. hfirst h2. left; auto. apply h2 in hf. red in hf. pose proof (h1 y) as h3. hfirst h3. right; auto. apply h3 in hf0. red in hf0. rewrite <- assoc_sum, hf0, hf; auto. Qed. Hint Resolve lat_sum. (* B is a lattice under partial order le (product portion). *) Lemma lat_prod : forall (x y: Bt), inf (Couple x y) (x * y). intros; red; split; red; intros s h1. destruct h1. apply times_le. rewrite comm_prod. apply times_le. red in h1. pose proof (h1 x) as h2. hfirst h2. left; auto. apply h2 in hf. rewrite eq_ord. red in hf. rewrite eq_ord in hf. rewrite assoc_prod, hf. specialize (h1 y). hfirst h1. right; auto. apply h1 in hf0. red in hf0. rewrite eq_ord in hf0; auto. Qed. Hint Resolve lat_prod. Lemma inf_unq : forall (S:Ensemble Bt) (x y: Bt), inf S x -> inf S y -> x = y. intros S x y h1 h2; red in h1, h2. destruct h1 as [h1a h1b], h2 as [h2a h2b]; red in h1a, h2a. apply anti_sym_le; auto. Qed. Hint Resolve inf_unq. Lemma sup_unq : forall (S:Ensemble Bt) (x y: Bt), sup S x -> sup S y -> x = y. intros S x y h1 h2; red in h1, h2. destruct h1 as [h1a h1b], h2 as [h2a h2b]; red in h1a, h2a. apply anti_sym_le; auto. Qed. Hint Resolve sup_unq. Lemma ex_inf_unq : forall (A:Ensemble Bt), (exists p:Bt, inf A p) -> (exists! p:Bt, inf A p). intros A h1. destruct h1 as [p h1]. exists p. red. split; auto. intros q h2. apply inf_unq with A; auto. Qed. Lemma ex_sup_unq : forall (A:Ensemble Bt), (exists p:Bt, sup A p) -> (exists! p:Bt, sup A p). intros A h1. destruct h1 as [p h1]. exists p. red. split; auto. intros q h2. apply sup_unq with A; auto. Qed. (*0 <= x*) Lemma zero_min : forall (x : Bt), le 0 x. intro x; red. pose proof (lat_prod x (- x)) as h1. rewrite comp_prod in h1. red in h1. destruct h1 as [h1 h2]. red in h1. specialize (h1 x). hfirst h1. left; auto. apply h1 in hf. red in hf; auto. Qed. Hint Resolve zero_min. Lemma le_x_0 : forall (x:Bt), le x 0 -> x = 0. intros x h1. pose proof (zero_min x) as h2. apply anti_sym_le; assumption. Qed. (* x <= 1 *) Lemma one_max : forall (x : Bt), le x 1. intro x; red. pose proof (lat_sum x (- x)) as h1. rewrite comp_sum in h1. red in h1. destruct h1 as [h1 h2]. red in h1. specialize (h1 x). hfirst h1. left; auto. apply h1 in hf. red in hf; auto. Qed. Hint Resolve one_max. Lemma le_1_x : forall (x:Bt), le 1 x -> x = 1. intros x h1. pose proof (one_max x) as h2. apply anti_sym_le; assumption. Qed. Lemma zero_lt_or : forall x:Bt, x = 0 \/ lt 0 x. intro x. destruct (ba_eq_dec 0 x) as [h1 | h2]. left; auto. right. red. split. apply zero_min. assumption. Qed. Lemma lt_one_or : forall x:Bt, x = 1 \/ lt x 1. intro. destruct (ba_eq_dec x 1) as [h1 | h2]. left; auto. right. red. split. apply one_max. assumption. Qed. Lemma neq_zero_lt_iff : forall (x:Bt), x <> 0 <-> lt 0 x. intro. split. intro h1. apply neq_sym in h1. red. split; auto. intros h1 h2. subst. apply lt_irrefl in h1; auto. Qed. Lemma lb_couple0 : forall x, lb (Couple x 0) 0. intro; red; intros ? h1. destruct h1. apply zero_min. apply refl_le. Qed. Lemma ub_couple0 : forall x, ub (Couple x 0) x. intro; red; intros ? h1. destruct h1. apply refl_le. apply zero_min. Qed. Lemma lb_couple1 : forall x, lb (Couple x 1) x. intro; red; intros ? h1. destruct h1. apply refl_le. apply one_max. Qed. Lemma ub_couple1 : forall x, ub (Couple x 1) 1. intro; red; intros ? h1. destruct h1. apply one_max. apply refl_le. Qed. Lemma zero_sum : forall (x : Bt), x + 0 = x. intro x. pose proof (lat_sum x 0) as h1. destruct h1 as [h1 h2]. red in h1. specialize (h1 x). hfirst h1. left; auto. apply h1 in hf. specialize (h2 x). specialize (h2 (ub_couple0 _)). apply anti_sym_le; auto. Qed. Hint Resolve zero_sum. Lemma zero_prod : forall (x : Bt), x * 0 = 0. intro x. pose proof (lat_prod x 0) as h1. destruct h1 as [h1 h2]. red in h1. specialize (h1 0). hfirst h1. right; auto. apply h1 in hf. specialize (h2 0 (lb_couple0 _)). apply anti_sym_le; auto. Qed. Hint Resolve zero_prod. Lemma one_sum : forall (x : Bt), x + 1 = 1. intro x. pose proof (lat_sum x 1) as h1. destruct h1 as [h1 h2]. red in h1. specialize (h1 1). hfirst h1. right; auto. apply h1 in hf. specialize (h2 1 (ub_couple1 _)). apply anti_sym_le; auto. Qed. Hint Resolve one_sum. Lemma one_prod : forall x : Bt, x * 1 = x. intro x. pose proof (lat_prod x 1) as h1. destruct h1 as [h1 h2]. red in h1. specialize (h1 x). hfirst h1. left; auto. apply h1 in hf. specialize (h2 x (lb_couple1 _)). apply anti_sym_le; auto. Qed. Hint Resolve one_prod. Lemma zero_sum_l : forall (x : Bt), 0 + x = x. intro; rewrite comm_sum, zero_sum; auto. Qed. Lemma zero_prod_l : forall (x : Bt), 0 * x = 0. intro; rewrite comm_prod, zero_prod; auto. Qed. Lemma one_sum_l : forall (x : Bt), 1 + x = 1. intro; rewrite comm_sum , one_sum; auto. Qed. Lemma one_prod_l : forall x : Bt, 1 * x= x. intro; rewrite comm_prod, one_prod; auto. Qed. Lemma non_zero_prod_non_zero_l : forall x y:Bt, x * y <> 0 -> x <> 0. intros x y h1 h2. subst. rewrite comm_prod in h1. rewrite zero_prod in h1. contradict h1; auto. Qed. Lemma non_zero_prod_non_zero_r : forall x y:Bt, x * y <> 0 -> y <> 0. intros x y h1 h2. subst. rewrite zero_prod in h1. contradict h1; auto. Qed. Lemma non_one_sum_non_one_l : forall x y:Bt, x + y <> 1 -> x <> 1. intros x y h1 h2. subst. rewrite comm_sum in h1. rewrite one_sum in h1. contradict h1; auto. Qed. Lemma non_one_sum_non_one_r : forall x y:Bt, x + y <> 1 -> y <> 1. intros x y h1 h2. subst. rewrite one_sum in h1. contradict h1; auto. Qed. Definition degen := Bzero (Bc B) = Bone (Bc B). Lemma degen_eq : degen -> forall x y:Bt, x = y. intros h1 x y. rewrite <- one_prod. rewrite <- (one_prod x). rewrite <- h1. do 2 rewrite zero_prod. reflexivity. Qed. Lemma degen_dec : {degen} + {~degen}. destruct (ba_eq_dec (bzero B) (bone B)) as [h1 | h1]. left. red; auto. right; contradict h1; auto. Qed. Lemma eq_comp_iff : forall x:Bt, x = - x <-> degen. intro x; split; intro h1. red. rewrite <- (comp_sum _ x). rewrite <- h1. rewrite idem_sum. rewrite <- (comp_prod _ x). rewrite <- h1. apply idem_prod. apply degen_eq; auto. Qed. Lemma inf_empty : inf (Empty_set Bt) 1. red. split. red. intros. contradiction. auto. Qed. Lemma sup_empty : sup (Empty_set Bt) 0. red. split. red. intros. contradiction. auto. Qed. Lemma inf_subtract_one : forall (A:Ensemble Bt) (b:Bt), inf A b <-> inf (Subtract A 1) b. intros A b. split. intro h1. red in h1. destruct h1 as [h1 h2]. red in h1. red. split. red. intros s h3. destruct h3 as [h3 h4]. apply h1; auto. intros b' h3. assert (h4:lb A b'). red. intros s h4. destruct (ba_eq_dec s 1) as [h5 | h6]. subst. apply one_max. red in h3. assert (h7:Inn (Subtract A 1) s). constructor; auto. intro h7. destruct h7. contradict h6. reflexivity. apply h3; auto. apply h2; auto. intro h1. red in h1. destruct h1 as [h1 h2]. red in h1. red. split. red. intros s h3. destruct (ba_eq_dec s 1) as [h4 | h5]. subst. apply one_max. assert (h6:Inn (Subtract A 1) s). constructor; auto. intro h6. destruct h6. contradict h5. reflexivity. apply h1; auto. intros b' h3. red in h3. assert (h4:lb (Subtract A 1) b'). red. intros s h4. destruct h4 as [h4 h5]. apply h3; auto. apply h2; auto. Qed. Lemma sup_subtract_zero : forall (A:Ensemble Bt) (b:Bt), sup A b <-> sup (Subtract A 0) b. intros A b. split. intro h1. red in h1. destruct h1 as [h1 h2]. red in h1. red. split. red. intros s h3. destruct h3 as [h3 h4]. apply h1; auto. intros b' h3. assert (h4:ub A b'). red. intros s h4. destruct (ba_eq_dec s 0) as [h5 | h6]. subst. apply zero_min. red in h3. assert (h7:Inn (Subtract A 0) s). constructor; auto. intro h7. destruct h7. contradict h6. reflexivity. apply h3; auto. apply h2; auto. intro h1. red in h1. destruct h1 as [h1 h2]. red in h1. red. split. red. intros s h3. destruct (ba_eq_dec s 0) as [h4 | h5]. subst. apply zero_min. assert (h6:Inn (Subtract A 0) s). constructor; auto. intro h6. destruct h6. contradict h5. reflexivity. apply h1; auto. intros b' h3. red in h3. assert (h4:ub (Subtract A 0) b'). red. intros s h4. destruct h4 as [h4 h5]. apply h3; auto. apply h2; auto. Qed. Lemma lb_subtract_one : forall (A:Ensemble Bt) (b:Bt), lb A b <-> lb (Subtract A 1) b. intros A b. split. intro h1. red in h1. red. intros s h2. destruct h2 as [h2 h3]. apply h1; auto. intro h1. red in h1. red. intros s h2. destruct (ba_eq_dec s 1) as [h3 | h4]. subst. apply one_max. apply h1. constructor; auto. intro h5. destruct h5. contradict h4. reflexivity. Qed. Lemma ub_subtract_zero : forall (A:Ensemble Bt) (b:Bt), ub A b <-> ub (Subtract A 0) b. intros A b. split. intro h1. red in h1. red. intros s h2. destruct h2 as [h2 h3]. apply h1; auto. intro h1. red in h1. red. intros s h2. destruct (ba_eq_dec s 0) as [h3 | h4]. subst. apply zero_min. apply h1. constructor; auto. intro h5. destruct h5. contradict h4. reflexivity. Qed. (* The goal of this section is to develop a tactic that reduces equalities that contain any 0 or 1 in a disjunction/conjunction or any adjacent combination of x/-x, so that 0s and 1s and inverses get eliminated on both sides of the equation, i.e. reduced form. *) Section const_eq. Inductive bin2 : Set := pnode2 : bin2 -> bin2 -> bin2 | tnode2 : bin2 -> bin2 -> bin2 | leaf2 : nat-> bin2 | zero : bin2 | one : bin2 | popps : nat -> bool -> bin2 | topps : nat -> bool -> bin2. Fixpoint bin2_Bt (l: list Bt)(def:Bt)(t:bin2){struct t}: Bt := match t with | pnode2 t1 t2 => (bin2_Bt l def t1) + (bin2_Bt l def t2) | tnode2 t1 t2 => (bin2_Bt l def t1) * (bin2_Bt l def t2) | leaf2 n => nth n l def | zero => 0 | one => 1 | popps n b => let x := (nth n l def) in if b then (x + -x) else (-x + x) | topps n b => let x := (nth n l def) in if b then (x * -x) else (-x * x) end. Fixpoint elim_bin2_aux (t:bin2) : bin2 := match t with | pnode2 zero t' => elim_bin2_aux t' | pnode2 t' zero => elim_bin2_aux t' | pnode2 one t' => one | pnode2 t' one => one | tnode2 zero t' => zero | tnode2 t' zero => zero | tnode2 one t' => elim_bin2_aux t' | tnode2 t' one => elim_bin2_aux t' | pnode2 t1 t2 => pnode2 (elim_bin2_aux t1) (elim_bin2_aux t2) | tnode2 t1 t2 => tnode2 (elim_bin2_aux t1) (elim_bin2_aux t2) | topps _ _ => zero | popps _ _ => one | _ => t end. Fixpoint elim_bin2 (t: bin2) : bin2 := match t with | pnode2 t1 t2 => elim_bin2_aux (pnode2 (elim_bin2 t1) (elim_bin2 t2)) | tnode2 t1 t2 => elim_bin2_aux (tnode2 (elim_bin2 t1) (elim_bin2 t2)) | topps _ _ => zero | popps _ _ => one | _ => t end. Theorem bin_Bt2_ref_p : forall (l:list Bt)(def:Bt)(t1 t2:bin2), bin2_Bt l def (pnode2 t1 t2) = (bin2_Bt l def t1) + (bin2_Bt l def t2). intros l def t1 t2. unfold bin2_Bt. fold bin2_Bt. reflexivity. Qed. Theorem bin_Bt2_ref_t : forall (l:list Bt)(def:Bt)(t1 t2:bin2), bin2_Bt l def (tnode2 t1 t2) = (bin2_Bt l def t1) * (bin2_Bt l def t2). intros l def t1 t2. unfold bin2_Bt. fold bin2_Bt. reflexivity. Qed. Theorem elim_bin2_aux_p : forall (l:list Bt) (def : Bt) (t1 t2:bin2), bin2_Bt l def (elim_bin2_aux (pnode2 t1 t2)) = bin2_Bt l def (pnode2 (elim_bin2_aux t1) (elim_bin2_aux t2)). intros l def t1 t2. unfold elim_bin2_aux. fold elim_bin2_aux. destruct t1. destruct t2; auto. unfold elim_bin2_aux at 3. rewrite bin_Bt2_ref_p. unfold bin2_Bt at 3. rewrite zero_sum. reflexivity. unfold elim_bin2_aux at 2. rewrite bin_Bt2_ref_p. unfold bin2_Bt at 3. unfold bin2_Bt at 1. rewrite one_sum. reflexivity. destruct t2. reflexivity. reflexivity. reflexivity. unfold elim_bin2_aux at 3. rewrite bin_Bt2_ref_p. unfold bin2_Bt at 3. rewrite zero_sum. reflexivity. unfold elim_bin2_aux at 2. unfold bin2_Bt. fold bin2_Bt. rewrite one_sum. reflexivity. reflexivity. reflexivity. destruct t2. reflexivity. reflexivity. reflexivity. unfold elim_bin2_aux at 3. unfold bin2_Bt. fold bin2_Bt. rewrite zero_sum. reflexivity. unfold elim_bin2_aux at 2. unfold bin2_Bt. fold bin2_Bt. rewrite one_sum. reflexivity. reflexivity. reflexivity. unfold elim_bin2_aux at 2. unfold bin2_Bt. fold bin2_Bt. rewrite (comm_sum B). rewrite zero_sum. reflexivity. destruct t2; unfold elim_bin2_aux at 1; unfold bin2_Bt; fold bin2_Bt; rewrite (comm_sum B); rewrite one_sum; reflexivity. destruct t2; try reflexivity. unfold elim_bin2_aux at 3. unfold bin2_Bt. fold bin2_Bt. rewrite zero_sum. reflexivity. unfold elim_bin2_aux at 2. unfold bin2_Bt. fold bin2_Bt. rewrite one_sum. reflexivity. destruct t2; try reflexivity. unfold elim_bin2_aux. unfold bin2_Bt. fold bin2_Bt. rewrite zero_sum. reflexivity. unfold elim_bin2_aux at 2. unfold bin2_Bt. fold bin2_Bt. rewrite one_sum. reflexivity. Qed. Theorem elim_bin2_aux_t : forall (l:list Bt) (def : Bt) (t1 t2:bin2), bin2_Bt l def (elim_bin2_aux (tnode2 t1 t2)) = bin2_Bt l def (tnode2 (elim_bin2_aux t1) (elim_bin2_aux t2)). intros l def t1 t2. unfold elim_bin2_aux. fold elim_bin2_aux. destruct t1. destruct t2; try reflexivity. unfold elim_bin2_aux at 2. unfold bin2_Bt. fold bin2_Bt. rewrite zero_prod. reflexivity. unfold elim_bin2_aux at 3. unfold bin2_Bt. fold bin2_Bt. rewrite one_prod. reflexivity. destruct t2; try reflexivity. unfold elim_bin2_aux at 2. unfold bin2_Bt. fold bin2_Bt. rewrite zero_prod. reflexivity. unfold elim_bin2_aux at 3. unfold bin2_Bt. fold bin2_Bt. rewrite one_prod. reflexivity. destruct t2; try reflexivity. unfold elim_bin2_aux at 2. unfold bin2_Bt. fold bin2_Bt. rewrite zero_prod. reflexivity. unfold elim_bin2_aux at 3. unfold bin2_Bt. fold bin2_Bt. rewrite one_prod. reflexivity. unfold elim_bin2_aux at 1. unfold bin2_Bt. fold bin2_Bt. rewrite (comm_prod B). rewrite zero_prod. reflexivity. unfold elim_bin2_aux at 5. unfold bin2_Bt. fold bin2_Bt. rewrite (comm_prod B). rewrite (one_prod). destruct t2; try reflexivity. destruct t2; try reflexivity. unfold elim_bin2_aux at 2. unfold bin2_Bt. fold bin2_Bt. rewrite zero_prod. reflexivity. unfold elim_bin2_aux at 3. unfold bin2_Bt. fold bin2_Bt. rewrite one_prod. reflexivity. destruct t2; try reflexivity. unfold elim_bin2_aux at 2. unfold bin2_Bt. fold bin2_Bt. rewrite zero_prod. reflexivity. unfold elim_bin2_aux at 3. unfold bin2_Bt. fold bin2_Bt. rewrite one_prod. reflexivity. Qed. Theorem elim_bin2_aux_valid : forall (l:list Bt)(def:Bt)(t:bin2), bin2_Bt l def t = bin2_Bt l def (elim_bin2_aux t). intros l def t. induction t; unfold bin2_Bt; fold bin2_Bt. rewrite IHt1. rewrite IHt2. rewrite <- bin_Bt2_ref_p. rewrite elim_bin2_aux_p. reflexivity. rewrite IHt1. rewrite IHt2. rewrite <- bin_Bt2_ref_t. rewrite elim_bin2_aux_t. reflexivity. unfold elim_bin2_aux. unfold bin2_Bt. reflexivity. unfold elim_bin2_aux. unfold bin2_Bt. reflexivity. unfold elim_bin2_aux. unfold bin2_Bt. reflexivity. case b. unfold elim_bin2_aux. unfold bin2_Bt. apply (comp_sum B). unfold elim_bin2_aux. unfold bin2_Bt. rewrite (comm_sum B). apply (comp_sum B). case b. unfold elim_bin2_aux. unfold bin2_Bt. apply (comp_prod B). unfold elim_bin2_aux. unfold bin2_Bt. rewrite (comm_prod B). apply (comp_prod B). Qed. Theorem elim_bin2_valid : forall (l:list Bt)(def:Bt)(t:bin2), bin2_Bt l def t = bin2_Bt l def (elim_bin2 t). intros l def t. induction t; unfold elim_bin2; fold elim_bin2; try rewrite <- elim_bin2_aux_valid; unfold bin2_Bt; fold bin2_Bt; try (rewrite IHt1; rewrite IHt2; reflexivity) || reflexivity. case b. apply (comp_sum B). rewrite (comm_sum B). apply (comp_sum B). case b. apply (comp_prod B). rewrite (comm_prod B). apply (comp_prod B). Qed. Theorem elim_bin2_valid_2 : forall (t t':bin2)(l:list Bt)(def:Bt), bin2_Bt l def (elim_bin2 t) = bin2_Bt l def (elim_bin2 t') -> bin2_Bt l def t = bin2_Bt l def t'. intros t t' l def. rewrite <- elim_bin2_valid. rewrite <- elim_bin2_valid. tauto. Qed. End const_eq. (*Tactics to build the const_eq tactic*) Ltac term_list2 l v := match v with | (?X1 + (-?X1)) => constr:(cons X1 l) | ((-?X1) + ?X1) => constr:(cons X1 l) | (?X1 * (-?X1)) => constr:(cons X1 l) | ((-?X1) * ?X1) => constr:(cons X1 l) | (?X1 + ?X2) => let l1 := term_list2 l X2 in term_list2 l1 X1 | (?X1 * ?X2) => let l1 := term_list2 l X2 in term_list2 l1 X1 | ?X1 => constr:(cons X1 l) end. Ltac model_aux2 l v := match v with | (?X1 + (-?X1)) => let n := compute_rank l 0%nat X1 in constr:(popps n true) | ((-?X1) + ?X1) => let n := compute_rank l 0%nat X1 in constr:(popps n false) | (?X1 * (-?X1)) => let n := compute_rank l 0%nat X1 in constr:(topps n true) | ((-?X1) * ?X1) => let n := compute_rank l 0%nat X1 in constr:(topps n false) | (?X1 + ?X2) => let r1 := model_aux2 l X1 with r2 := model_aux2 l X2 in constr:(pnode2 r1 r2) | (?X1 * ?X2) => let r1 := model_aux2 l X1 with r2 := model_aux2 l X2 in constr:(tnode2 r1 r2) | 0 => constr:(zero) | 1 => constr:(one) | ?X1 => let n := compute_rank l 0%nat X1 in constr:(leaf2 n) | _ => constr:(leaf2 0%nat) end. Ltac const_eq := match goal with | [ |- (?X1 = ?X2) ] => let l1 := term_list2 (nil (A := Bt)) X1 in let l2 := term_list2 l1 X2 in let term1 := model_aux2 l2 X1 with term2 := model_aux2 l2 X2 in (change (bin2_Bt l2 X1 term1 = bin2_Bt l2 X1 term2); apply elim_bin2_valid_2; auto) end. (* (*All tests check *) Variables a b c d e f g h : Bt. Goal ((a + 0) = a). const_eq. Qed. Goal ((a + (b*0)) = a). const_eq. Qed. Goal ((a+b+(c*1)*(d+0)) = a + b +(c*d)). const_eq. Qed. Goal (0*(a+b) + c*1) = c. const_eq. Qed. Goal (a+(b*-b)) = a. const_eq. Qed. Goal (a+(c+d)*(-(c+d))) + (b*(c+d*((e+f) + (-(e+f))))) = a + b*(c+d). const_eq. Qed. *) (* More Theorems *) Lemma comp_unq : forall (x y : Bt), -x = -y -> x = y. intros x y h1. assert (h2: x = x*y). assert (h3: x*y = x*y + x * -x). const_eq. assert (h4: x*y + x*-x = x * (y + -x)). rewrite <- (dist_sum B); auto. rewrite h1 in h4. assert (h5: x * (y + -y) = x). const_eq. congruence. assert (h6: y = y *x). assert (h7 : y*x = y*x + y*-y). const_eq. assert (h8: y*x + y*-y = y*(x + -y)). rewrite <- (dist_sum B); auto. rewrite <- h1 in h8. assert (h9: y*(x + -x) = y). const_eq. congruence. rewrite h2. rewrite h6 at 2. comm_eq. Qed. Lemma inj_comp : Injective (Bcomp (Bc B)). red; intros; apply comp_unq; auto. Qed. Hint Resolve comp_unq. Lemma doub_neg : forall (x : Bt), x = --x. intro x. pose proof (one_prod (---x)) as h1. pose proof (comp_sum _ (-x)) as h2. rewrite <- h2, dist_sum, (comm_prod _ (---x) (--x)), (comp_prod _ (--x)), zero_sum, <-eq_ord in h1. pose proof (anti_sym_le (---x) (-x)) as h3. hfirst h3. red; auto. specialize (h3 hf). hfirst h3. pose proof (one_prod (-x)) as h4. rewrite <- (comp_sum _ (--x)), dist_sum, comp_prod, comm_sum, zero_sum in h4. red; rewrite eq_ord; auto. auto. Qed. Hint Resolve doub_neg. (* Zero is characterized by the property x + y = x for all y*) Lemma zero_char : forall (y : Bt), (forall (x : Bt), x + y = x) -> y = 0. intros y h1; specialize (h1 0); rewrite zero_sum_l in h1; auto. Qed. Hint Resolve zero_char. (* One is characterized by the property x * y = x for all y*) Lemma one_char : forall (y : Bt), (forall (x : Bt), x * y = x) -> y = 1. intros y h1; specialize (h1 1); rewrite one_prod_l in h1; auto. Qed. (* complement is characterized by x * y = 0, x + y =1 *) Lemma comp_char : forall (x y : Bt), x * y = 0 -> x + y = 1 -> y = -x. intros x y h1 h2. assert (h3: y = y * (x + -x)). const_eq. pose proof (dist_sum _ x (-x) y) as h4. assert (h5: (y * x) + (y * -x) = (x * y) + (y * -x)). comm_eq. rewrite h1 in h5. assert (h6 : 0 + y * -x = (-x * y) + 0). comm_eq. assert (h7: (-x * y) + 0 = -x * y + -x * x). const_eq. assert (h8: -x * y + -x * x = -x * (y + x)). rewrite dist_sum; auto. assert (h9: -x * (y + x) = -x * (x + y)). comm_eq. rewrite h2 in h9. assert (h10 : -x * 1 = -x). const_eq. congruence. Qed. Hint Resolve comp_char. (* -0 = 1 *) Lemma zero_comp: -0 = Bone (Bc B). pose proof (one_prod 0). pose proof (one_sum 0). pose proof (comp_char 0 1) as h1. rewrite h1; auto. Qed. Hint Resolve zero_comp. (* -1 = 0 *) Lemma one_comp: -(1) = Bzero (Bc B). pose proof (zero_sum 1). pose proof (zero_prod 1). pose proof (comp_char 1 0) as h1. rewrite h1; auto. Qed. Hint Resolve one_comp. Lemma comp_decomp_sum : forall (x y:Bt), x = x*y + x*-y. intros x y. rewrite <- dist_sum. rewrite comp_sum. rewrite one_prod; auto. Qed. Lemma comp_decomp_prod : forall (x y:Bt), x = (x+y) * (x+-y). intros x y. rewrite <- dist_prod. rewrite comp_prod. rewrite zero_sum; auto. Qed. Lemma zero_iff : forall x:Bt, x = 0 <-> -x = 1. intro x. split. intro h1. subst. rewrite zero_comp; auto. intro h1. pose proof (f_equal (@Bcomp (Bc B)) h1) as h2. rewrite <- doub_neg in h2. subst. rewrite one_comp; auto. Qed. Lemma one_iff : forall x:Bt, x = 1 <-> -x = 0. intro x. split. intro h1. subst. rewrite one_comp; auto. intro h1. pose proof (f_equal (@Bcomp (Bc B)) h1) as h2. rewrite <- doub_neg in h2. subst. rewrite zero_comp; auto. Qed. Lemma de_mor_sum : forall (x y : Bt), -(x + y) = -x * -y. intros x y. pose proof (dist_sum_r x y (-x * -y)) as h1. rewrite (assoc_prod _ x), comp_prod, (comm_prod _ 0 (-y)), (zero_prod (-y)), zero_sum_l, (comm_prod _ (-x)), (assoc_prod _ y (-y)), comp_prod, (comm_prod _ 0), zero_prod, (comm_prod _ (-y)) in h1. pose proof (dist_prod _ (-x) (-y) (x+y)) as h2. rewrite <- (assoc_sum _ x y (-y)), comp_sum, one_sum, one_prod, <- (assoc_sum _ x y (-x)), (comm_sum _ y (-x)), (assoc_sum _ x (-x) y), comp_sum, (comm_sum _ 1), one_sum in h2. symmetry. apply comp_char; auto. Qed. Hint Resolve de_mor_sum. Lemma de_mor_prod : forall (x y : Bt), -(x*y) = -x + -y. intros x y. pose proof (dist_prod_r x y (-x + -y)) as h1. rewrite (assoc_sum _ x), comp_sum, (comm_sum _ 1 (-y)), (one_sum (-y)), one_prod_l, (comm_sum _ (-x)), (assoc_sum _ y (-y)), comp_sum, (comm_sum _ 1), one_sum, (comm_sum _ (-y)) in h1. pose proof (dist_sum _ (-x) (-y) (x*y)) as h2. rewrite <- (assoc_prod _ x y (-y)), comp_prod, zero_prod, zero_sum, <- (assoc_prod _ x y (-x)), (comm_prod _ y (-x)), (assoc_prod _ x (-x) y), comp_prod, (comm_prod _ 0), zero_prod in h2. symmetry. apply comp_char; auto. Qed. Hint Resolve de_mor_prod. Lemma times_eq_plus : forall (x y: Bt), x*y = -((-x) + (-y)). intros x y. rewrite de_mor_sum with (x := -x) (y := -y). rewrite <- doub_neg. rewrite <- doub_neg. reflexivity. Qed. Lemma plus_eq_times : forall (x y:Bt), x+y = -((-x)*(-y)). intros x y. rewrite de_mor_prod with (x := -x) (y := -y). rewrite <- doub_neg. rewrite <- doub_neg. reflexivity. Qed. Lemma mono_sum : forall (x x' y y' : Bt), le x x' -> le y y' -> le (x+y) (x' + y'). intros x x' y y' h1 h2; red in h1, h2; red. assert (h3:(x + y) + (x' + y') = (x+x') + (y + y')). comm_eq. rewrite h3, h1, h2; auto. Qed. Lemma mono_prod : forall (x x' y y': Bt), le x x' -> le y y' -> le (x * y) (x' * y'). intros x x' y y' h1 h2; red in h1, h2; red; rewrite eq_ord in h1, h2; rewrite eq_ord. assert (h3:(x * y) * (x' * y') = (x * x') * (y * y')). comm_eq. rewrite h3, h1, h2; auto. Qed. Lemma plus_preserves_le : forall (p q r:Bt), le p r -> le q r -> le (p + q) r. intros p q r h1 h2. red in h1, h2. red. rewrite <- assoc_sum. rewrite h2. assumption. Qed. Lemma times_preserves_le : forall (p q r:Bt), le r p -> le r q -> le r (p * q). intros p q r h1 h2. red in h1, h2. red. rewrite eq_ord. rewrite eq_ord in h1, h2. rewrite assoc_prod. rewrite h1. rewrite h2. reflexivity. Qed. Lemma le_iff' : forall (a b:Bt), le a (-b) <-> a * b = 0. intros a b. split. intro h1. pose proof refl_le. pose proof (refl_le b) as h2. pose proof (mono_prod _ _ _ _ h1 h2) as h3. rewrite (comm_prod _ (-b) b) in h3. rewrite comp_prod in h3. apply le_x_0. assumption. intro h1. pose proof (zero_sum (a*-b)) as h2. rewrite <- h1 in h2. rewrite <- dist_sum in h2. rewrite comm_sum in h2. rewrite comp_sum in h2. rewrite one_prod in h2. red. rewrite eq_ord. rewrite h2 at 2. reflexivity. Qed. Lemma le_iff : forall (a b:Bt), le a b <-> a * -b = 0. intros a b. rewrite (doub_neg b) at 1. rewrite le_iff'. tauto. Qed. Lemma mono_comp : forall (x x' : bt B), le x x' -> le (-x') (-x). intros x x' h1; red in h1; red; rewrite eq_ord in h1. rewrite comm_sum, <- (de_mor_prod x x'), h1; auto. Qed. Lemma mono_comp_iff : forall (x x' : bt B), le x x' <-> le (-x') (-x). intros x x'. split. apply mono_comp. intro h1. apply mono_comp in h1. do 2 rewrite <- doub_neg in h1. assumption. Qed. Lemma mono_comp_iff' : forall (x x' : bt B), le x (- x') <-> le (x') (-x). intros x x'. split. intro h1. red in h1. red. rewrite eq_ord in h1. rewrite eq_ord. rewrite <- h1. rewrite de_mor_prod. rewrite <- doub_neg. rewrite dist_sum. rewrite idem_prod. rewrite comm_sum. apply abs_sum. intro h1. red in h1. red. rewrite eq_ord in h1. rewrite eq_ord. rewrite <- h1. rewrite de_mor_prod. rewrite <- doub_neg. rewrite dist_sum. rewrite idem_prod. rewrite comm_sum. apply abs_sum. Qed. Definition comp_set {B:Bool_Alg} (A:Ensemble (bt B)) := Im A (Bcomp (Bc B)). Lemma comp_set_comp_set : forall (A:Ensemble Bt), comp_set (comp_set A) = A. intro A. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]. subst. destruct h1 as [x h1]. subst. rewrite <- doub_neg. assumption. red. apply Im_intro with (-x). apply Im_intro with x; auto. apply doub_neg. Qed. Lemma lb_ub_compat_iff : forall (A:Ensemble Bt) (p:Bt), lb A p <-> ub (comp_set A) (-p). intros A p. split. intros h1. red in h1. red. intros s h2. destruct h2 as [s h2]. subst. specialize (h1 _ h2). apply mono_comp; auto. intro h1. red in h1. red. intros s h2. specialize (h1 (-s)). assert (h3:Inn (comp_set A) (-s)). apply Im_intro with s; auto. specialize (h1 h3). pose proof (mono_comp _ _ h1) as h4. rewrite <- doub_neg in h4. rewrite <- doub_neg in h4. assumption. Qed. Lemma ub_lb_compat_iff : forall (A:Ensemble Bt) (p:Bt), ub A p <-> lb (comp_set A) (-p). intros A p. split. intros h1. red in h1. red. intros s h2. destruct h2 as [s h2]. subst. specialize (h1 _ h2). apply mono_comp; auto. intro h1. red in h1. red. intros s h2. specialize (h1 (-s)). assert (h3:Inn (comp_set A) (-s)). apply Im_intro with s; auto. specialize (h1 h3). pose proof (mono_comp _ _ h1) as h4. rewrite <- doub_neg in h4. rewrite <- doub_neg in h4. assumption. Qed. Lemma sup_inf_compat_iff : forall (A:Ensemble Bt) (p:Bt), sup A p <-> inf (comp_set A) (-p). intros A p. split. intro h1. red in h1. red. destruct h1 as [h1l h1r]. split. red. intros s h2. destruct h2 as [s h2]. subst. apply h1l in h2. apply mono_comp; auto. intros b' h2. rewrite lb_ub_compat_iff in h2. rewrite comp_set_comp_set in h2. pose proof (h1r _ h2) as h3. rewrite (doub_neg b'). apply mono_comp; auto. intro h1. red in h1. red. destruct h1 as [h1l h1r]. split. red. intros s h2. assert (h3:Inn (comp_set A) (-s)). apply Im_intro with s; auto. apply h1l in h3. rewrite doub_neg. rewrite (doub_neg s). apply mono_comp; auto. intros b' h2. rewrite ub_lb_compat_iff in h2. specialize (h1r _ h2). rewrite doub_neg. rewrite (doub_neg p). apply mono_comp; auto. Qed. Lemma inf_sup_compat_iff : forall (A:Ensemble Bt) (p:Bt), inf A p <-> sup (comp_set A) (-p). intros A p. split. intro h1. red in h1. red. destruct h1 as [h1l h1r]. split. red. intros s h2. destruct h2 as [s h2]. subst. apply h1l in h2. apply mono_comp; auto. intros b' h2. rewrite ub_lb_compat_iff in h2. rewrite comp_set_comp_set in h2. pose proof (h1r _ h2) as h3. rewrite (doub_neg b'). apply mono_comp; auto. intro h1. red in h1. red. destruct h1 as [h1l h1r]. split. red. intros s h2. assert (h3:Inn (comp_set A) (-s)). apply Im_intro with s; auto. apply h1l in h3. rewrite doub_neg. rewrite (doub_neg p). apply mono_comp; auto. intros b' h2. rewrite lb_ub_compat_iff in h2. specialize (h1r _ h2). rewrite doub_neg. rewrite (doub_neg b'). apply mono_comp; auto. Qed. Definition atom (a:Bt) : Prop := a <> 0 /\ forall b:Bt, le b a -> a = b \/ b = 0. Lemma atom_iff : forall (a:Bt), atom a <-> (forall b:Bt, (le a b \/ a * b = 0) /\ ~ (le a b /\ a*b = 0)). intro a. split. (* -> *) intros h1 b. split. red in h1. destruct h1 as [h1l h1r]. pose proof (times_le a b) as h2. specialize (h1r _ h2). destruct h1r. left. red. rewrite eq_ord. symmetry. assumption. right. assumption. intro h2. destruct h2 as [h2l h2r]. rewrite <- le_iff' in h2r. pose proof (mono_prod _ _ _ _ h2l h2r) as h3. rewrite idem_prod in h3. rewrite comp_prod in h3. pose proof (le_x_0 _ h3). subst. red in h1. destruct h1 as [h1]. contradict h1. apply reflexivity. (* <- *) intro h1. red. split. intro h2. subst. specialize (h1 0). destruct h1 as [h1l h1r]. contradict h1r. split. apply zero_min. apply zero_prod. intros b h2. specialize (h1 b). destruct h1 as [h1l h1r]. destruct h1l as [h1a | h1b]. pose proof (anti_sym_le _ _ h1a h2) as h3. subst. left. reflexivity. rewrite comm_prod in h1b. rewrite <- le_iff' in h1b. pose proof (mono_prod _ _ _ _ h1b h2) as h3. rewrite idem_prod in h3. rewrite comm_prod in h3. rewrite comp_prod in h3. right. apply le_x_0. assumption. Qed. Lemma atom_impl : forall (a:Bt), atom a -> a <> 0 /\ (forall b:Bt, a * b = a \/ a * b = 0). intros a h1. red in h1. destruct h1 as [h1l h1r]. split; auto. intro b. pose proof (times_le a b) as h2. specialize (h1r _ h2). destruct h1r. left. symmetry. assumption. right. auto. Qed. Lemma atoms_disjoint : forall (a b:Bt), a <> b -> atom a -> atom b -> a * b = 0. intros a b h1 h2 h3. pose proof (atom_impl a h2) as h4. destruct h4 as [h4 h5]. specialize (h5 b). destruct h5 as [h5 |]; auto. red in h3. destruct h3 as [h3 h6]. rewrite <- eq_ord in h5. specialize (h6 _ h5). destruct h6 as [h6 | h8]. clear h5. subst. contradict h1. reflexivity. contradiction. Qed. Definition atom_set := [x:bt B | atom x]. End Basics. (*Theorems*) Arguments flatten_aux_valid_Bt_p [B] l def t t'. Arguments flatten_aux_valid_Bt_t [B] l def t t'. Arguments flatten_valid_Bt [B] l def t. Arguments flatten_valid_Bt_2 [B] t t' l def _. Arguments bin_Bt_ref_p [B] l def t1 t2. Arguments bin_Bt_ref_t [B] l def t1 t2. Arguments insert_eq_p [B] l def t0 t. Arguments insert_eq_t [B] l def t0 t. Arguments sort_eq [B] l def t. Arguments sort_eq_2 [B] l def t1 t2 _. Arguments dist_sum_r [B] n m p. Arguments dist_prod_r [B] n m p. Arguments idem_sum [B] b. Arguments idem_prod [B] b. Arguments eq_ord [B] x y. Arguments refl_le [B] x. Arguments anti_sym_le [B] x y _ _. Arguments trans_le [B] x y z _ _. Arguments lat_sum [B] x y. Arguments lat_prod [B] x y. Arguments inf_unq [B] S x y _ _. Arguments sup_unq [B] S x y _ _. Arguments zero_min [B] x. Arguments one_max [B] x. Arguments zero_sum [B] x. Arguments zero_prod [B] x. Arguments one_sum [B] x. Arguments one_prod [B] x. Arguments comp_unq [B] x y _. Arguments doub_neg [B] x. Arguments zero_char [B] y _. Arguments one_char [B] y _. Arguments comp_char [B] x y _ _. Arguments comp_unq [B] x y _. Arguments zero_comp [B]. Arguments one_comp [B]. Arguments de_mor_sum [B] x y. Arguments de_mor_prod [B] x y. Arguments mono_sum [B] x x' y y' _ _. Arguments mono_prod [B] x x' y y' _ _. Arguments mono_comp [B] x x' _. Arguments le_plus [B] _ _. Arguments le_x_0 [B] _ _. Arguments le_1_x [B] _ _. Arguments times_le [B] _ _. Arguments atom_iff [B] _. Arguments in_bs [B] _. (*Definitions*) Arguments le [B] x y. Arguments lt [B] x y. Arguments lb [B] S b. Arguments ub [B] S b. Arguments inf [B] S b. Arguments sup [B] S b. Arguments atom [B] _. Definition sym_diff {B:Bool_Alg} (x y:(bt B)) := x * -y + y* -x. Notation "x /_\ y" := (sym_diff x y) (at level 50, left associativity) : ba_scope. Section SymDiff. Variable B:Bool_Alg. Let Bt := bt B. Lemma comm_sym_diff : forall (x y:Bt), x /_\ y = y /_\ x. intros. unfold sym_diff. apply comm_sum. Qed. Lemma sym_diff_ref : forall x:Bt, x /_\ x = 0. intros. unfold sym_diff. rewrite comp_prod. apply zero_sum. Qed. Lemma zero_sum_each_zero: forall (x y:Bt), x + y = 0 <-> x = 0 /\ y = 0. intros x y. split. intro h1. pose proof (le_plus x y) as h2. rewrite h1 in h2. pose proof (le_x_0 _ h2). subst. rewrite comm_sum in h1. rewrite zero_sum in h1. subst. split; reflexivity. intro h2. destruct h2; subst. rewrite zero_sum. reflexivity. Qed. Lemma one_prod_each_one: forall (x y:Bt), x * y = 1 <-> x = 1 /\ y = 1. intros x y. split. intro h1. pose proof (times_le x y) as h2. rewrite h1 in h2. pose proof (le_1_x _ h2). subst. rewrite comm_prod in h1. rewrite one_prod in h1. subst. split; reflexivity. intro h2. destruct h2; subst. rewrite one_prod. reflexivity. Qed. Lemma sym_diff_iff : forall (x y:Bt), x /_\ y = 0 <-> x = y. intros x y. split. intro h1. unfold sym_diff in h1. rewrite zero_sum_each_zero in h1. destruct h1 as [h1l h1r]. unfold Bt, bt in h1l. rewrite <- le_iff in h1l. unfold Bt, bt in h1r. rewrite <- le_iff in h1r. apply anti_sym_le; assumption. intros; subst. apply sym_diff_ref. Qed. End SymDiff. Lemma ba_subst_pred : forall (A B:Bool_Alg), A = B -> forall (P:Bool_Alg->Prop), P A <-> P B. intros; subst; tauto. Qed. Section ParametricAnalogues. Variable T:Type. Variable Bp:Bool_Alg_p T. Let Btp := btp Bp. Lemma sig_set_ba_p_ens_eq : forall {T:Type} (Bp:Bool_Alg_p T), sig_set (ba_p_ens Bp) = btp Bp. unfold btp; simpl; auto. Qed. Lemma ba_p_in_add_ens : forall (S:Ensemble T) (pf:Included S (ba_p_ens Bp)) (a:Btp), Included (Add S (proj1_sig a)) (ba_p_ens Bp). intros S h0 a. red; intros x h1. destruct h1 as [x h1 | x h1]; auto. destruct h1. apply proj2_sig. Qed. Lemma in_ba_p_ens_times : forall (x y:T) (pfx:Inn (ba_p_ens Bp) x) (pfy:Inn (ba_p_ens Bp) y), Inn (ba_p_ens Bp) (proj1_sig ((exist _ _ pfx) %* (exist _ _ pfy))). intros. apply proj2_sig. Qed. Lemma in_ba_p_ens_plus : forall (x y:T) (pfx:Inn (ba_p_ens Bp) x) (pfy:Inn (ba_p_ens Bp) y), Inn (ba_p_ens Bp) (proj1_sig ((exist _ _ pfx) %+ (exist _ _ pfy))). intros. apply proj2_sig. Qed. Lemma in_ba_p_ens_comp : forall (x:T) (pfx:Inn (ba_p_ens Bp) x), Inn (ba_p_ens Bp) (proj1_sig (%-(exist _ _ pfx))). intros. apply proj2_sig. Qed. Lemma in_ba_p_ens_one : Inn (ba_p_ens Bp) (proj1_sig (Bone_p T (Bc_p T Bp))). apply proj2_sig. Qed. Lemma in_ba_p_ens_zero : Inn (ba_p_ens Bp) (proj1_sig (Bzero_p T (Bc_p T Bp))). apply proj2_sig. Qed. Lemma in_ba_p_ens_plus' : forall (x y:Btp), Inn (ba_p_ens Bp) (proj1_sig (x %+ y)). destruct x, y. apply in_ba_p_ens_plus. Qed. Lemma in_ba_p_ens_times' : forall (x y:Btp), Inn (ba_p_ens Bp) (proj1_sig (x %* y)). destruct x, y. apply in_ba_p_ens_times. Qed. Lemma in_ba_p_ens_comp' : forall x:Btp, Inn (ba_p_ens Bp) (proj1_sig (%- x)). destruct x. apply in_ba_p_ens_comp. Qed. Lemma in_bs_p : forall x:Btp, Inn (BS_p T (Bc_p T Bp)) x. intros; rewrite und_set_p; constructor. Qed. Lemma bc_inj_p : forall {A A':Bool_Alg_p T}, Bc_p T A = Bc_p T A' -> A = A'. intros A A' h1. destruct A; destruct A'. simpl in h1. subst. f_equal; apply proof_irrelevance. Qed. Lemma dist_sum_r_p : forall (n m p: Btp), (n %+ m) %* p = n%*p %+ m%*p. intros n m p. do 2 rewrite ba_conv_plus. do 3 rewrite ba_conv_times. repeat rewrite <- ba_conv_elt_eq. apply (@dist_sum_r (ba_conv Bp)). Qed. Lemma dist_prod_r_p : forall (n m p: Btp), (n %* m) %+ p = (n%+p) %* (m%+p). intros n m p. do 3 rewrite ba_conv_plus. do 2 rewrite ba_conv_times. repeat rewrite <- ba_conv_elt_eq. apply (@dist_prod_r (ba_conv Bp)). Qed. Lemma idem_sum_p : forall b : Btp, b %+ b = b. intro b. rewrite ba_conv_plus. repeat rewrite <- ba_conv_elt_eq. apply (@idem_sum (ba_conv Bp)). Qed. Lemma idem_prod_p: forall (b: Btp), b %* b = b. intro b. rewrite ba_conv_times. repeat rewrite <- ba_conv_elt_eq. apply (@idem_prod (ba_conv Bp)). Qed. Lemma comp_prod_r_p : forall (b:Btp), (%-b) %* b = %0. intro; rewrite comm_prod_p; apply comp_prod_p. Qed. Lemma comp_sum_r_p : forall (b:Btp), (%-b) %+ b = %1. intro; rewrite comm_sum_p; apply comp_sum_p. Qed. Lemma eq_ord_p: forall (x y: Btp), x %+ y = y <-> x %* y = x. intros x y. rewrite ba_conv_times. rewrite ba_conv_plus. repeat rewrite <- ba_conv_elt_eq. apply (@eq_ord (ba_conv Bp)). Qed. Definition le_p (x y:Btp): Prop := x %+ y = y. Lemma le_p_iff : forall (x y:Btp), le_p x y <-> le (ba_conv_elt x) (ba_conv_elt y). intros x y. unfold le_p, le. repeat rewrite <- ba_conv_elt_eq. simpl. tauto. Qed. Lemma refl_le_p : forall (x: Btp), le_p x x. intro x. rewrite le_p_iff. repeat rewrite <- ba_conv_elt_eq. apply refl_le. Qed. Lemma anti_sym_le_p: forall (x y: Btp), le_p x y -> le_p y x -> x = y. intros x y. rewrite le_p_iff. repeat rewrite <- ba_conv_elt_eq. apply anti_sym_le. Qed. Lemma trans_le_p: forall (x y z : Btp), le_p x y -> le_p y z -> le_p x z. intros x y z. rewrite le_p_iff. repeat rewrite <- ba_conv_elt_eq. apply trans_le. Qed. Lemma times_le_p : forall (x y : Btp), le_p (x %* y) x. intros x y. rewrite le_p_iff. rewrite ba_conv_times. repeat rewrite <- ba_conv_elt_eq. apply times_le. Qed. Lemma le_plus_p : forall (x y : Btp), le_p x (x %+ y). intros x y. rewrite le_p_iff. rewrite ba_conv_plus. repeat rewrite <- ba_conv_elt_eq. apply le_plus. Qed. Definition lt_p (x y:Btp) := le_p x y /\ x <> y. Lemma lt_p_iff : forall (x y:Btp), lt_p x y <-> lt (transfer (ba_conv_type Bp) x) (transfer (ba_conv_type Bp) y). intros x y. unfold lt_p. rewrite le_p_iff. unfold lt. tauto. Qed. (* b is a lower bound of S in B*) Definition lb_p (S: Ensemble Btp) (b: Btp) : Prop := forall (s:Btp), Inn S s -> le_p b s. Lemma lb_p_iff : forall (S:Ensemble Btp) (b:Btp), lb_p S b <-> lb (ba_conv_set S) (ba_conv_elt b). intros S b. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. unfold transfer. unfold eq_rect_r. simpl. unfold lb_p. unfold lb. simpl. split. intros h1 s h2. specialize (h1 _ h2). rewrite le_p_iff in h1. repeat rewrite <- ba_conv_elt_eq in h1. assumption. intros h1 s h2. rewrite le_p_iff. repeat rewrite <- ba_conv_elt_eq. apply h1; auto. Qed. (* b is an upper bound of S in B.*) Definition ub_p (S: Ensemble Btp) (b: Btp) : Prop := forall (s:Btp), Inn S s -> le_p s b. Lemma ub_p_iff : forall (S:Ensemble Btp) (b:Btp), ub_p S b <-> ub (ba_conv_set S) (ba_conv_elt b). intros S b. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. unfold transfer. unfold eq_rect_r. simpl. unfold lb_p. unfold lb. simpl. split. intros h1 s h2. specialize (h1 _ h2). rewrite le_p_iff in h1. repeat rewrite <- ba_conv_elt_eq in h1. assumption. intros h1 s h2. rewrite le_p_iff. repeat rewrite <- ba_conv_elt_eq. apply h1; auto. Qed. (* b is the infimum (greatest lower bound) of S in B.*) Definition inf_p (S: Ensemble Btp) (b: Btp) : Prop := lb_p S b /\ forall b':Btp, lb_p S b' -> le_p b' b. Lemma inf_p_iff : forall (S:Ensemble Btp) (b:Btp), inf_p S b <-> inf (ba_conv_set S) (ba_conv_elt b). intros S b. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. unfold transfer. unfold eq_rect_r. simpl. unfold inf_p. unfold inf. simpl. split. intros h1. destruct h1 as [h1l h1r]. split. rewrite lb_p_iff in h1l. rewrite <- ba_conv_elt_eq in h1l. unfold ba_conv_set in h1l. unfold ba_conv_type in h1l. rewrite transfer_dep_eq_refl in h1l. assumption. intros b' h2. pose proof (lb_p_iff S b') as h3. rewrite <- ba_conv_elt_eq in h3. unfold ba_conv_set in h3. unfold ba_conv_type in h3. rewrite transfer_dep_eq_refl in h3. rewrite <- h3 in h2. pose proof (h1r _ h2) as h4. rewrite le_p_iff in h4. repeat rewrite <- ba_conv_elt_eq in h4. assumption. intro h1. destruct h1 as [h1l h1r]. split. rewrite lb_p_iff. repeat rewrite <- ba_conv_elt_eq. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. assumption. intros b' h2. rewrite lb_p_iff in h2. repeat rewrite <- ba_conv_elt_eq in h2. unfold ba_conv_set in h2. unfold ba_conv_type in h2. rewrite transfer_dep_eq_refl in h2. specialize (h1r _ h2). rewrite le_p_iff. do 2 rewrite <- ba_conv_elt_eq. assumption. Qed. (* b is the supremum (least upper bound) of S in B.*) Definition sup_p (S: Ensemble Btp) (b: Btp) : Prop := ub_p S b /\ forall b':Btp, ub_p S b' -> le_p b b'. Lemma sup_p_iff : forall (S:Ensemble Btp) (b:Btp), sup_p S b <-> sup (ba_conv_set S) (ba_conv_elt b). intros S b. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. unfold transfer. unfold eq_rect_r. simpl. unfold sup_p. unfold sup. simpl. split. intros h1. destruct h1 as [h1l h1r]. split. rewrite ub_p_iff in h1l. rewrite <- ba_conv_elt_eq in h1l. unfold ba_conv_set in h1l. unfold ba_conv_type in h1l. rewrite transfer_dep_eq_refl in h1l. assumption. intros b' h2. pose proof (ub_p_iff S b') as h3. rewrite <- ba_conv_elt_eq in h3. unfold ba_conv_set in h3. unfold ba_conv_type in h3. rewrite transfer_dep_eq_refl in h3. rewrite <- h3 in h2. pose proof (h1r _ h2) as h4. rewrite le_p_iff in h4. repeat rewrite <- ba_conv_elt_eq in h4. assumption. intro h1. destruct h1 as [h1l h1r]. split. rewrite ub_p_iff. repeat rewrite <- ba_conv_elt_eq. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. assumption. intros b' h2. rewrite ub_p_iff in h2. repeat rewrite <- ba_conv_elt_eq in h2. unfold ba_conv_set in h2. unfold ba_conv_type in h2. rewrite transfer_dep_eq_refl in h2. specialize (h1r _ h2). rewrite le_p_iff. do 2 rewrite <- ba_conv_elt_eq. assumption. Qed. Lemma inf_le_p : forall (A:Ensemble Btp) (b:Btp), inf_p A b -> forall b':Btp, Inn A b' -> le_p b b'. intros A b h1 b' h2. rewrite le_p_iff. do 2 rewrite <- ba_conv_elt_eq. rewrite inf_p_iff in h1. rewrite <- ba_conv_elt_eq in h1. unfold ba_conv_set in h1. unfold ba_conv_type in h1. rewrite transfer_dep_eq_refl in h1. apply inf_le with A; auto. Qed. Lemma le_sup_p : forall (A:Ensemble Btp) (b:Btp), sup_p A b -> forall b':Btp, Inn A b' -> le_p b' b. intros A b h1 b' h2. rewrite le_p_iff. do 2 rewrite <- ba_conv_elt_eq. rewrite sup_p_iff in h1. rewrite <- ba_conv_elt_eq in h1. unfold ba_conv_set in h1. unfold ba_conv_type in h1. rewrite transfer_dep_eq_refl in h1. apply le_sup with A; auto. Qed. Lemma inf_singleton_p : forall (x:Btp), inf_p (Singleton x) x. intro x. rewrite inf_p_iff. rewrite <- ba_conv_elt_eq. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. apply inf_singleton. Qed. Lemma sup_singleton_p : forall (x:Btp), sup_p (Singleton x) x. intro x. rewrite sup_p_iff. rewrite <- ba_conv_elt_eq. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. apply sup_singleton. Qed. Lemma lat_sum_p : forall (x y: Btp), sup_p (Couple x y) (x %+ y). intros x y. rewrite sup_p_iff. rewrite ba_conv_plus. repeat rewrite <- ba_conv_elt_eq. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. apply lat_sum; auto. Qed. Lemma lat_prod_p : forall (x y: Btp), inf_p (Couple x y) (x %* y). intros x y. rewrite inf_p_iff. rewrite ba_conv_times. repeat rewrite <- ba_conv_elt_eq. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. apply lat_prod; auto. Qed. Lemma inf_unq_p : forall (S:Ensemble Btp) (x y: Btp), inf_p S x -> inf_p S y -> x = y. intros S x y h1 h2. rewrite inf_p_iff in h1. rewrite inf_p_iff in h2. rewrite <- ba_conv_elt_eq in h1. rewrite <- ba_conv_elt_eq in h2. unfold ba_conv_set in h1. unfold ba_conv_type in h1. unfold ba_conv_set in h2. unfold ba_conv_type in h2. rewrite transfer_dep_eq_refl in h1. rewrite transfer_dep_eq_refl in h2. apply (@inf_unq (ba_conv Bp) S); auto. Qed. Lemma sup_unq_p : forall (S:Ensemble Btp) (x y: Btp), sup_p S x -> sup_p S y -> x = y. intros S x y h1 h2. rewrite sup_p_iff in h1. rewrite sup_p_iff in h2. rewrite <- ba_conv_elt_eq in h1. rewrite <- ba_conv_elt_eq in h2. unfold ba_conv_set in h1. unfold ba_conv_type in h1. unfold ba_conv_set in h2. unfold ba_conv_type in h2. rewrite transfer_dep_eq_refl in h1. rewrite transfer_dep_eq_refl in h2. apply (@sup_unq (ba_conv Bp) S); auto. Qed. Lemma ex_inf_unq_p : forall (A:Ensemble Btp), (exists p:Btp, inf_p A p) -> (exists! p:Btp, inf_p A p). intros A h1. destruct h1 as [p h1]. rewrite inf_p_iff in h1. rewrite <- ba_conv_elt_eq in h1. unfold ba_conv_set in h1. unfold ba_conv_type in h1. rewrite transfer_dep_eq_refl in h1. pose proof (ex_inf_unq _ (transfer_dep (ba_conv_type Bp) A)) as h2. simpl in h2. specialize (h2 (ex_intro _ p h1)). destruct h2 as [b h2]. exists b. red in h2. red. destruct h2 as [h2 h3]. split. rewrite inf_p_iff. rewrite <- ba_conv_elt_eq. assumption. intros x h4. rewrite inf_p_iff in h4. rewrite <- ba_conv_elt_eq in h4. apply h3; auto. Qed. Lemma ex_sup_unq_p : forall (A:Ensemble Btp), (exists p:Btp, sup_p A p) -> (exists! p:Btp, sup_p A p). intros A h1. destruct h1 as [p h1]. rewrite sup_p_iff in h1. rewrite <- ba_conv_elt_eq in h1. unfold ba_conv_set in h1. unfold ba_conv_type in h1. rewrite transfer_dep_eq_refl in h1. pose proof (ex_sup_unq _ (transfer_dep (ba_conv_type Bp) A)) as h2. simpl in h2. specialize (h2 (ex_intro _ p h1)). destruct h2 as [b h2]. exists b. red in h2. red. destruct h2 as [h2 h3]. split. rewrite sup_p_iff. rewrite <- ba_conv_elt_eq. assumption. intros x h4. rewrite sup_p_iff in h4. rewrite <- ba_conv_elt_eq in h4. apply h3; auto. Qed. Lemma inj_comp_p : Injective (Bcomp_p T (Bc_p T Bp)). red; intros x y h1. rewrite (ba_conv_comp _ x), (ba_conv_comp _ y) in h1. pose proof (inj_comp _ (ba_conv_elt x) (ba_conv_elt y)) as h2. apply h2 in h1; auto. Qed. Lemma zero_min_p : forall (x : Btp), le_p %0 x. intro x. rewrite le_p_iff. rewrite ba_conv_zero. do 2 rewrite <- ba_conv_elt_eq. apply zero_min. Qed. Lemma le_x_0_p : forall (x:Btp), le_p x %0 -> x = %0. intros x h1. rewrite le_p_iff in h1. rewrite ba_conv_zero in h1. do 2 rewrite <- ba_conv_elt_eq in h1. rewrite ba_conv_zero. apply (@le_x_0 (ba_conv Bp)); auto. Qed. Lemma one_max_p : forall (x : Btp), le_p x %1. intro x. rewrite le_p_iff. rewrite ba_conv_one. do 2 rewrite <- ba_conv_elt_eq. apply one_max. Qed. Lemma le_1_x_p : forall (x:Btp), le_p %1 x -> x = %1. intros x h1. rewrite le_p_iff in h1. rewrite ba_conv_one in h1. do 2 rewrite <- ba_conv_elt_eq in h1. rewrite ba_conv_one. apply (@le_1_x (ba_conv Bp)); auto. Qed. Lemma zero_sum_p : forall (x : Btp), x %+ %0 = x. intro x. rewrite ba_conv_plus. rewrite ba_conv_zero. repeat rewrite <- ba_conv_elt_eq. apply (@zero_sum (ba_conv Bp)). Qed. Lemma zero_prod_p : forall (x : Btp), x %* %0 = %0. intro x. rewrite ba_conv_times. rewrite ba_conv_zero. repeat rewrite <- ba_conv_elt_eq. apply (@zero_prod (ba_conv Bp)). Qed. Lemma one_sum_p : forall (x : Btp), x %+ %1 = %1. intro x. rewrite ba_conv_plus. rewrite ba_conv_one. repeat rewrite <- ba_conv_elt_eq. apply (@one_sum (ba_conv Bp)). Qed. Lemma one_prod_p : forall x : Btp, x %* %1 = x. intros x. rewrite ba_conv_times. rewrite ba_conv_one. repeat rewrite <- ba_conv_elt_eq. apply (@one_prod (ba_conv Bp)). Qed. Lemma non_zero_prod_non_zero_l_p : forall x y:Btp, x %* y <> %0 -> x <> %0. intros x y h1 h2. subst. rewrite comm_prod_p in h1. rewrite zero_prod_p in h1. contradict h1; auto. Qed. Lemma non_zero_prod_non_zero_r_p : forall x y:Btp, x %* y <> %0 -> y <> %0. intros x y h1 h2. subst. rewrite zero_prod_p in h1. contradict h1; auto. Qed. Lemma non_one_sum_non_one_l_p : forall x y:Btp, x %+ y <> %1 -> x <> %1. intros x y h1 h2. subst. rewrite comm_sum_p in h1. rewrite one_sum_p in h1. contradict h1; auto. Qed. Lemma non_one_sum_non_one_r_p : forall x y:Btp, x %+ y <> %1 -> y <> %1. intros x y h1 h2. subst. rewrite one_sum_p in h1. contradict h1; auto. Qed. Lemma inf_empty_p : inf_p (Empty_set Btp) %1. rewrite inf_p_iff. rewrite ba_conv_one. repeat rewrite <- ba_conv_elt_eq. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. apply inf_empty. Qed. Lemma sup_empty_p : sup_p (Empty_set Btp) %0. rewrite sup_p_iff. rewrite ba_conv_zero. repeat rewrite <- ba_conv_elt_eq. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. apply sup_empty. Qed. Lemma inf_subtract_one_p : forall (A:Ensemble Btp) (b:Btp), inf_p A b <-> inf_p (Subtract A %1) b. intros A b. rewrite inf_p_iff. apply inf_subtract_one. Qed. Lemma sup_subtract_zero_p : forall (A:Ensemble Btp) (b:Btp), sup_p A b <-> sup_p (Subtract A %0) b. intros A b. rewrite sup_p_iff. apply sup_subtract_zero. Qed. Lemma lb_subtract_one_p : forall (A:Ensemble Btp) (b:Btp), lb_p A b <-> lb_p (Subtract A %1) b. intros A b. rewrite lb_p_iff. apply lb_subtract_one. Qed. Lemma ub_subtract_zero_p : forall (A:Ensemble Btp) (b:Btp), ub_p A b <-> ub_p (Subtract A %0) b. intros A b. rewrite ub_p_iff. apply ub_subtract_zero. Qed. Lemma comp_unq_p : forall (x y : Btp), %-x = %-y -> x = y. intros x y. rewrite ba_conv_comp. rewrite ba_conv_comp. apply (@comp_unq (ba_conv Bp)). Qed. Lemma doub_neg_p : forall (x : Btp), x = %- %-x. intro x. do 2 rewrite ba_conv_comp. apply (@doub_neg (ba_conv Bp)). Qed. Lemma zero_char_p : forall (y : Btp), (forall (x : Btp), x %+ y = x) -> y = %0. intros y h1. assert (h2:forall x:(Btype (Bc (ba_conv Bp))), x + y = x). intro x. apply h1. rewrite ba_conv_zero. apply (@zero_char (ba_conv Bp)); auto. Qed. Lemma one_char_p : forall (y : Btp), (forall (x : Btp), x %* y = x) -> y = %1. intros y h1. assert (h2:forall x:(Btype (Bc (ba_conv Bp))), x * y = x). intro x. apply h1. rewrite ba_conv_one. apply (@one_char (ba_conv Bp)); auto. Qed. Lemma comp_char_p : forall (x y : Btp), x %* y = %0 -> x %+ y = %1 -> y = %-x. intros x y. rewrite ba_conv_zero. rewrite ba_conv_one. rewrite ba_conv_times. rewrite ba_conv_comp. rewrite ba_conv_plus. apply (@comp_char (ba_conv Bp)). Qed. Lemma zero_comp_p: %-%0 = Bone_p T (Bc_p T Bp). rewrite ba_conv_zero. rewrite ba_conv_comp. rewrite ba_conv_one. apply (@zero_comp (ba_conv Bp)). Qed. Lemma one_comp_p: %-(%1) = Bzero_p T (Bc_p T Bp). rewrite ba_conv_zero. rewrite ba_conv_comp. rewrite ba_conv_one. apply (@one_comp (ba_conv Bp)). Qed. Lemma zero_iff_p : forall x:Btp, x = %0 <-> %- x = %1. intro x. rewrite ba_conv_zero, ba_conv_one. pose proof (zero_iff _ (ba_conv_elt x)). assumption. Qed. Lemma one_iff_p : forall x:Btp, x = %1 <-> %- x = %0. intro x. rewrite ba_conv_zero, ba_conv_one. pose proof (one_iff _ (ba_conv_elt x)). assumption. Qed. Lemma de_mor_sum_p : forall (x y : Btp), %-(x %+ y) = %-x %* %-y. intros x y. do 3 rewrite ba_conv_comp. rewrite ba_conv_plus. rewrite ba_conv_times. apply (@de_mor_sum (ba_conv Bp)). Qed. Lemma de_mor_prod_p : forall (x y : Btp), %-(x%*y) = %-x %+ %-y. intros x y. do 3 rewrite ba_conv_comp. rewrite ba_conv_plus. rewrite ba_conv_times. apply (@de_mor_prod (ba_conv Bp)). Qed. Lemma times_eq_plus_p : forall (x y: Btp), x%*y = %-((%-x) %+ (%-y)). intros x y. rewrite ba_conv_times. do 3 rewrite ba_conv_comp. rewrite ba_conv_plus. apply (@times_eq_plus (ba_conv Bp)). Qed. Lemma plus_eq_times_p : forall (x y:Btp), x%+y = %-((%-x)%*(%-y)). intros x y. rewrite ba_conv_times. do 3 rewrite ba_conv_comp. rewrite ba_conv_plus. apply (@plus_eq_times (ba_conv Bp)). Qed. Lemma mono_sum_p : forall (x x' y y' : Btp), le_p x x' -> le_p y y' -> le_p (x%+y) (x' %+ y'). intros x x' y y'. rewrite le_p_iff. do 2 rewrite ba_conv_plus. apply (@mono_sum (ba_conv Bp)). Qed. Lemma mono_prod_p : forall (x x' y y': Btp), le_p x x' -> le_p y y' -> le_p (x %* y) (x' %* y'). intros x x' y y'. rewrite le_p_iff. do 2 rewrite ba_conv_times. apply (@mono_prod (ba_conv Bp)). Qed. Lemma le_iff_p' : forall (a b:Btp), le_p a (%-b) <-> a %* b = %0. intros a b. rewrite le_p_iff. rewrite ba_conv_times. rewrite ba_conv_comp. rewrite ba_conv_zero. apply (@le_iff' (ba_conv Bp)). Qed. Lemma le_iff_p : forall (a b:Btp), le_p a b <-> a %* %-b = %0. intros a b. rewrite le_p_iff. rewrite ba_conv_times. rewrite ba_conv_comp. rewrite ba_conv_zero. apply (@le_iff (ba_conv Bp)). Qed. Lemma mono_comp_p : forall (x x' : Btp), le_p x x' -> le_p (%-x') (%-x). intros x x'. rewrite le_p_iff. do 2 rewrite ba_conv_comp. apply (@mono_comp (ba_conv Bp)). Qed. Definition comp_set_p (A:Ensemble Btp) := Im A (Bcomp_p T (Bc_p T Bp)). Lemma comp_set_p_eq : forall (A:Ensemble Btp), comp_set_p A = comp_set (ba_conv_set A). intros A. unfold comp_set_p. unfold comp_set. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. f_equal. Qed. Lemma comp_set_p_comp_set_p : forall (A:Ensemble Btp), comp_set_p (comp_set_p A) = A. intro A. do 2 rewrite comp_set_p_eq. unfold ba_conv_set. unfold ba_conv_type. do 2 rewrite transfer_dep_eq_refl. apply (@comp_set_comp_set (ba_conv Bp)). Qed. Lemma lb_ub_compat_iff_p : forall (A:Ensemble Btp) (p:Btp), lb_p A p <-> ub_p (comp_set_p A) (%-p). intros A p. rewrite lb_p_iff. rewrite ub_p_iff. rewrite comp_set_p_eq. rewrite ba_conv_comp. apply (@lb_ub_compat_iff (ba_conv Bp)). Qed. Lemma ub_lb_compat_iff_p : forall (A:Ensemble Btp) (p:Btp), ub_p A p <-> lb_p (comp_set_p A) (%-p). intros A p. rewrite lb_p_iff. rewrite ub_p_iff. rewrite comp_set_p_eq. rewrite ba_conv_comp. apply (@ub_lb_compat_iff (ba_conv Bp)). Qed. Lemma sup_inf_compat_iff_p : forall (A:Ensemble Btp) (p:Btp), sup_p A p <-> inf_p (comp_set_p A) (%-p). intros A p. rewrite sup_p_iff. rewrite inf_p_iff. rewrite comp_set_p_eq. rewrite ba_conv_comp. apply (@sup_inf_compat_iff (ba_conv Bp)). Qed. Lemma inf_sup_compat_iff_p : forall (A:Ensemble Btp) (p:Btp), inf_p A p <-> sup_p (comp_set_p A) (%-p). intros A p. rewrite sup_p_iff. rewrite inf_p_iff. rewrite comp_set_p_eq. rewrite ba_conv_comp. apply (@inf_sup_compat_iff (ba_conv Bp)). Qed. Definition atom_p (a:Btp) : Prop := a <> %0 /\ forall b:Btp, le_p b a -> a = b \/ b = %0. Lemma atom_p_iff : forall (a:Btp), atom_p a <-> atom (ba_conv_elt a). intro a. unfold atom_p, atom. rewrite ba_conv_zero. rewrite <- ba_conv_elt_eq. tauto. Qed. Lemma atom_iff_p : forall (a:Btp), atom_p a <-> (forall b:Btp, (le_p a b \/ a %* b = %0) /\ ~ (le_p a b /\ a%*b = %0)). intro a. rewrite atom_p_iff. rewrite <- ba_conv_elt_eq. split. intros h1 b. rewrite atom_iff in h1. specialize (h1 b). rewrite le_p_iff. assumption. intros h1. rewrite atom_iff. assumption. Qed. Lemma atom_impl_p : forall (a:Btp), atom_p a -> a <> %0 /\ (forall b:Btp, a %* b = a \/ a %* b = %0). intro a. rewrite atom_p_iff. rewrite <- ba_conv_elt_eq. intro h1. apply atom_impl in h1. assumption. Qed. Definition sym_diff_p (x y:Btp) := x %* %-y %+ y %* %-x. Lemma incl_ba_conv_set_iff : forall (E F:Ensemble (btp Bp)), Included E F <-> Included (ba_conv_set E) (ba_conv_set F). intros E F. split. intro h1. red. intros x h2. unfold ba_conv_set in h2. unfold ba_conv_type in h2. rewrite transfer_dep_eq_refl in h2. apply h1; auto. intros h1. red. intros x h2. apply h1; auto. Qed. Lemma eq_ba_conv_set_iff : forall (E F:Ensemble (btp Bp)), E = F <-> ba_conv_set E = ba_conv_set F. intros; subst; tauto. Qed. Lemma eq_ba_conv_set_iff' : forall (E F:Ensemble (btp Bp)), E = F <-> ba_conv_set' E = ba_conv_set' F. intros; subst; tauto. Qed. Lemma list_to_set_ba_conv_list : forall (l:list (btp Bp)), ListUtilities.list_to_set (ba_conv_list l) = ba_conv_set (ListUtilities.list_to_set l). intro l. apply Extensionality_Ensembles. red. split. red. intros x h1. unfold ba_conv_set. unfold ba_conv_type. rewrite transfer_dep_eq_refl. unfold ba_conv_list in h1. unfold ba_conv_type in h1. rewrite transfer_dep_eq_refl in h1. assumption. red. intros x h1. unfold ba_conv_set in h1. unfold ba_conv_type in h1. rewrite transfer_dep_eq_refl in h1. unfold ba_conv_list. unfold ba_conv_type. rewrite transfer_dep_eq_refl. assumption. Qed. Definition degen_p : Prop := Bzero_p T (Bc_p T Bp) = Bone_p T (Bc_p T Bp). Lemma degen_p_iff : degen_p <-> degen (ba_conv Bp). split; intro h1; red in h1; red. rewrite ba_conv_zero, ba_conv_one in h1. assumption. rewrite ba_conv_zero, ba_conv_one. assumption. Qed. Lemma degen_eq_p : degen_p -> forall x y:Btp, x = y. intros h1 x y. rewrite degen_p_iff in h1. rewrite (ba_conv_elt_eq x), (ba_conv_elt_eq y). apply (degen_eq _ h1 (ba_conv_elt x) (ba_conv_elt y)). Qed. Lemma degen_p_dec : {degen_p} + {~degen_p}. destruct (ba_eq_dec_p (bzero_p Bp) (bone_p Bp)) as [h1 | h1]. left. red; auto. right; contradict h1; auto. Qed. Lemma eq_comp_iff_p : forall x:Btp, x = %- x <-> degen_p. intro x. rewrite ba_conv_comp. rewrite (ba_conv_elt_eq x) at 1. pose proof (eq_comp_iff _ (ba_conv_elt x)); auto. Qed. End ParametricAnalogues. (*Definitions*) Arguments le_p [T] [Bp] x y. Arguments lt_p [T] [Bp] x y. Arguments lb_p [T] [Bp] S b. Arguments ub_p [T] [Bp] S b. Arguments inf_p [T] [Bp] S b. Arguments sup_p [T] [Bp] S b. Arguments atom_p [T] [Bp] _. Arguments sym_diff_p [T] [Bp] _ _. Arguments comp_set_p [T] [Bp] _ _. Arguments degen_p [T] _. Notation "x /%\ y" := (sym_diff_p x y) (at level 50, left associativity) : ba_p_scope. (*broken up for top-level Notation*) Section ParametricAnalogues2. Variable T:Type. Variable Bp:Bool_Alg_p T. Let Btp := btp Bp. Lemma sym_diff_p_eq : forall (x y:Btp), x /%\ y = (ba_conv_elt x) /_\ (ba_conv_elt y). intros x y. unfold sym_diff_p, sym_diff. do 2 rewrite ba_conv_times. rewrite ba_conv_plus. do 2 rewrite ba_conv_comp. repeat rewrite <- ba_conv_elt_eq. reflexivity. Qed. Lemma comm_sym_diff_p : forall (x y:Btp), x /%\ y = y /%\ x. intros x y. rewrite sym_diff_p_eq. rewrite (sym_diff_p_eq y x). apply (@comm_sym_diff (ba_conv Bp)). Qed. Lemma sym_diff_ref_p : forall x:Btp, x /%\ x = %0. intros. rewrite sym_diff_p_eq. rewrite ba_conv_zero. apply (@sym_diff_ref (ba_conv Bp)). Qed. Lemma zero_sum_each_zero_p: forall (x y:Btp), x %+ y = %0 <-> x = %0 /\ y = %0. intros x y. rewrite ba_conv_plus. rewrite ba_conv_zero. apply (@zero_sum_each_zero (ba_conv Bp)). Qed. Lemma one_prod_each_one_p: forall (x y:Btp), x %* y = %1 <-> x = %1 /\ y = %1. intros x y. rewrite ba_conv_times. rewrite ba_conv_one. apply (@one_prod_each_one (ba_conv Bp)). Qed. Lemma sym_diff_iff_p : forall (x y:Btp), x /%\ y = %0 <-> x = y. intros x y. rewrite sym_diff_p_eq. rewrite ba_conv_zero. apply (@sym_diff_iff (ba_conv Bp)). Qed. End ParametricAnalogues2. (*Theorems*) Arguments dist_sum_r_p [T] [Bp] n m p. Arguments dist_prod_r_p [T] [Bp] n m p. Arguments idem_sum_p [T] [Bp] b. Arguments idem_prod_p [T] [Bp] b. Arguments eq_ord_p [T] [Bp] x y. Arguments refl_le_p [T] [Bp] x. Arguments anti_sym_le_p [T] [Bp] x y _ _. Arguments trans_le_p [T] [Bp] x y z _ _. Arguments lat_sum_p [T] [Bp] x y. Arguments lat_prod_p [T] [Bp] x y. Arguments inf_unq_p [T] [Bp] S x y _ _. Arguments sup_unq_p [T] [Bp] S x y _ _. Arguments zero_min_p [T] [Bp] x. Arguments one_max_p [T] [Bp] x. Arguments zero_sum_p [T] [Bp] x. Arguments zero_prod_p [T] [Bp] x. Arguments one_sum_p [T] [Bp] x. Arguments one_prod_p [T] [Bp] x. Arguments comp_unq_p [T] [Bp] x y _. Arguments doub_neg_p [T] [Bp] x. Arguments zero_char_p [T] [Bp] y _. Arguments one_char_p [T] [Bp] y _. Arguments comp_char_p [T] [Bp] x y _ _. Arguments comp_unq_p [T] [Bp] x y _. Arguments zero_comp_p [T] [Bp]. Arguments one_comp_p [T] [Bp]. Arguments de_mor_sum_p [T] [Bp] x y. Arguments de_mor_prod_p [T] [Bp] x y. Arguments mono_sum_p [T] [Bp] x x' y y' _ _. Arguments mono_prod_p [T] [Bp] x x' y y' _ _. Arguments mono_comp_p [T] [Bp] x x' _. Arguments le_plus_p [T] [Bp] _ _. Arguments le_x_0_p [T] [Bp] _ _. Arguments le_1_x_p [T] [Bp] _ _. Arguments times_le_p [T] [Bp] _ _. Arguments atom_iff_p [T] [Bp] _. Arguments in_bs_p [T] [Bp] _. Arguments comm_sym_diff_p [T] [Bp] _ _. Arguments sym_diff_ref_p [T] [Bp] _. Arguments zero_sum_each_zero_p [T] [Bp] _ _. Arguments one_prod_each_one_p [T] [Bp] _ _. Arguments sym_diff_iff_p [T] [Bp] _ _. Arguments comp_set_p_eq [T] [Bp] _. Arguments ba_p_ens [T] _ _. Arguments btp [T] _. Section MoreParametricBAs. Lemma in_ba_p_ens_eq : forall {T:Type} (x:T) (C D:Bool_Alg_p T), C = D -> (Inn (ba_p_ens C) x <-> Inn (ba_p_ens D) x). intros; subst; tauto. Qed. Lemma ba_p_subst_times : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y), proj1_sig (exist _ _ pfxa %* exist _ _ pfya) = proj1_sig (exist _ _ pfxb %* exist _ _ pfyb). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst; auto. Qed. Lemma ba_p_subst_times3 : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y z:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfza:Inn (ba_p_ens Ap) z) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y) (pfzb:Inn (ba_p_ens Bp) z), proj1_sig (exist _ _ pfxa %* exist _ _ pfya %* exist _ _ pfza) = proj1_sig (exist _ _ pfxb %* exist _ _ pfyb %* exist _ _ pfzb). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst. pose proof (proof_irrelevance _ pfza pfzb); subst; auto. Qed. Lemma ba_p_subst_times3' : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y z:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfza:Inn (ba_p_ens Ap) z) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y) (pfzb:Inn (ba_p_ens Bp) z), proj1_sig (exist _ _ pfxa %* (exist _ _ pfya %* exist _ _ pfza)) = proj1_sig (exist _ _ pfxb %* (exist _ _ pfyb %* exist _ _ pfzb)). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst. pose proof (proof_irrelevance _ pfza pfzb); subst; auto. Qed. Lemma ba_p_subst_plus : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y), proj1_sig (exist _ _ pfxa %+ exist _ _ pfya) = proj1_sig (exist _ _ pfxb %+ exist _ _ pfyb). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst; auto. Qed. Lemma ba_p_subst_plus3 : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y z:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfza:Inn (ba_p_ens Ap) z) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y) (pfzb:Inn (ba_p_ens Bp) z), proj1_sig (exist _ _ pfxa %+ exist _ _ pfya %+ exist _ _ pfza) = proj1_sig (exist _ _ pfxb %+ exist _ _ pfyb %+ exist _ _ pfzb). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst. pose proof (proof_irrelevance _ pfza pfzb); subst; auto. Qed. Lemma ba_p_subst_plus3' : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y z:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfza:Inn (ba_p_ens Ap) z) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y) (pfzb:Inn (ba_p_ens Bp) z), proj1_sig (exist _ _ pfxa %+ (exist _ _ pfya %+ exist _ _ pfza)) = proj1_sig (exist _ _ pfxb %+ (exist _ _ pfyb %+ exist _ _ pfzb)). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst. pose proof (proof_irrelevance _ pfza pfzb); subst; auto. Qed. Lemma ba_p_subst_times_plus : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y z:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfza:Inn (ba_p_ens Ap) z) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y) (pfzb:Inn (ba_p_ens Bp) z), proj1_sig (exist _ _ pfxa %* exist _ _ pfya %+ exist _ _ pfza) = proj1_sig (exist _ _ pfxb %* exist _ _ pfyb %+ exist _ _ pfzb). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst. pose proof (proof_irrelevance _ pfza pfzb); subst; auto. Qed. Lemma ba_p_subst_times_plus' : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y z:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfza:Inn (ba_p_ens Ap) z) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y) (pfzb:Inn (ba_p_ens Bp) z), proj1_sig (exist _ _ pfxa %* (exist _ _ pfya %+ exist _ _ pfza)) = proj1_sig (exist _ _ pfxb %* (exist _ _ pfyb %+ exist _ _ pfzb)). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst. pose proof (proof_irrelevance _ pfza pfzb); subst; auto. Qed. Lemma ba_p_subst_plus_times : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y z:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfza:Inn (ba_p_ens Ap) z) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y) (pfzb:Inn (ba_p_ens Bp) z), proj1_sig ((exist _ _ pfxa %+ exist _ _ pfya) %* exist _ _ pfza) = proj1_sig ((exist _ _ pfxb %+ exist _ _ pfyb) %* exist _ _ pfzb). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst. pose proof (proof_irrelevance _ pfza pfzb); subst; auto. Qed. Lemma ba_p_subst_plus_times' : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y z:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfza:Inn (ba_p_ens Ap) z) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y) (pfzb:Inn (ba_p_ens Bp) z), proj1_sig (exist _ _ pfxa %+ exist _ _ pfya %* exist _ _ pfza) = proj1_sig (exist _ _ pfxb %+ exist _ _ pfyb %* exist _ _ pfzb). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst. pose proof (proof_irrelevance _ pfza pfzb); subst; auto. Qed. Lemma ba_p_subst_times_plus_times : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y z w:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfza:Inn (ba_p_ens Ap) z) (pfwa:Inn (ba_p_ens Ap) w) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y) (pfzb:Inn (ba_p_ens Bp) z) (pfwb:Inn (ba_p_ens Bp) w), proj1_sig (exist _ _ pfxa %* exist _ _ pfya %+ exist _ _ pfza %* exist _ _ pfwa) = proj1_sig (exist _ _ pfxb %* exist _ _ pfyb %+ exist _ _ pfzb %* exist _ _ pfwb). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst. pose proof (proof_irrelevance _ pfza pfzb); subst; auto. pose proof (proof_irrelevance _ pfwa pfwb); subst; auto. Qed. Lemma ba_p_subst_plus_times_plus : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x y z w:T) (pfxa:Inn (ba_p_ens Ap) x) (pfya:Inn (ba_p_ens Ap) y) (pfza:Inn (ba_p_ens Ap) z) (pfwa:Inn (ba_p_ens Ap) w) (pfxb:Inn (ba_p_ens Bp) x) (pfyb:Inn (ba_p_ens Bp) y) (pfzb:Inn (ba_p_ens Bp) z) (pfwb:Inn (ba_p_ens Bp) w), proj1_sig ((exist _ _ pfxa %+ exist _ _ pfya) %* (exist _ _ pfza %+ exist _ _ pfwa)) = proj1_sig ((exist _ _ pfxb %+ exist _ _ pfyb) %* (exist _ _ pfzb %+ exist _ _ pfwb)). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst. pose proof (proof_irrelevance _ pfya pfyb); subst. pose proof (proof_irrelevance _ pfza pfzb); subst; auto. pose proof (proof_irrelevance _ pfwa pfwb); subst; auto. Qed. Lemma ba_p_subst_comp : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> forall (x:T) (pfxa:Inn (ba_p_ens Ap) x) (pfxb:Inn (ba_p_ens Bp) x), proj1_sig (%- exist _ _ pfxa) = proj1_sig (%- exist _ _ pfxb). intros; subst. pose proof (proof_irrelevance _ pfxa pfxb); subst; auto. Qed. Lemma ba_p_subst_one : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> proj1_sig (Bone_p T (Bc_p T Ap)) = proj1_sig (Bone_p T (Bc_p T Bp)). intros; subst; auto. Qed. Lemma ba_p_subst_zero : forall {T:Type} (Ap Bp:Bool_Alg_p T), Ap = Bp -> proj1_sig (Bzero_p T (Bc_p T Ap)) = proj1_sig (Bzero_p T (Bc_p T Bp)). intros; subst; auto. Qed. End MoreParametricBAs. Section TwoBoolBa. Definition two_bool_bc := Build_Bconst bool (Full_set bool) orb andb true false negb. Definition two_bool_ex : exists! (B:Bool_Alg), Bc B = two_bool_bc. assert (h1: BS (two_bool_bc) = Full_set bool). reflexivity. exists (Build_Bool_Alg two_bool_bc h1 orb_assoc andb_assoc orb_comm andb_comm absoption_orb absoption_andb (fun a b c => andb_orb_distrib_r c a b) (fun a b c => orb_andb_distrib_r c a b) orb_negb_r andb_negb_r). red. split. simpl. reflexivity. intros B h2. apply bc_inj. simpl. auto. Qed. Definition two_bool_ba := proj1_sig (constructive_definite_description _ two_bool_ex). Lemma two_bool_ba_compat : Bc two_bool_ba = two_bool_bc. unfold two_bool_ba. destruct constructive_definite_description. simpl. assumption. Qed. End TwoBoolBa. Section ExpBoolBa. Lemma exp_bool_times_ex : forall {T:Type} {A:Ensemble T} (F F':Fin_map A (Full_set bool) false), exists! (G:Fin_map A (Full_set bool) false), forall x:T, Inn A x -> G |-> x = F |-> x && F' |-> x. intros T A F F'. pose (fun_to_fin_map (fin_map_type_eq_dec F) _ false (fin_map_fin_dom F) (fun x => F |-> x && F' |-> x)) as G'. assert (h1:Included (Im A (fun x:T=>F|->x && F'|->x)) (Full_set bool)). red. intros x h1. constructor. pose (fin_map_new_ran G' finite_bool h1) as G. exists G. red. split. intros x ha. unfold G. rewrite <- fin_map_new_ran_compat. unfold G'. rewrite fun_to_fin_map_compat; auto. assumption. intros G'' h2. apply fin_map_ext. intros x h3. specialize (h2 _ h3). rewrite h2. unfold G. rewrite <- fin_map_new_ran_compat; auto. unfold G'. rewrite fun_to_fin_map_compat; auto. Qed. Definition exp_bool_times {T:Type} {A:Ensemble T} (F F':Fin_map A (Full_set bool) false) := proj1_sig (constructive_definite_description _ (exp_bool_times_ex F F')). Lemma exp_bool_times_compat : forall {T:Type} (A:Ensemble T) (F F':Fin_map A (Full_set bool) false), let G := exp_bool_times F F' in forall x:T, Inn A x -> G |-> x = F |-> x && F' |-> x. unfold exp_bool_times; intros; destruct constructive_definite_description; auto. Qed. Definition exp_bool_plus_ex : forall {T:Type} {A:Ensemble T} (F F':Fin_map A (Full_set bool) false), exists! (G:Fin_map A (Full_set bool) false), forall x:T, Inn A x -> G |-> x = F |-> x || F' |-> x. intros T A F F'. pose (fun_to_fin_map (fin_map_type_eq_dec F) _ false (fin_map_fin_dom F) (fun x => F |-> x || F' |-> x)) as G'. assert (h1:Included (Im A (fun x:T=>F|->x || F'|->x)) (Full_set bool)). red. intros x h1. constructor. pose (fin_map_new_ran G' finite_bool h1) as G. exists G. red. split. intros x ha. unfold G. rewrite <- fin_map_new_ran_compat. unfold G'. rewrite fun_to_fin_map_compat; auto. assumption. intros G'' h2. apply fin_map_ext. intros x h3. rewrite h2; auto. unfold G. rewrite <- fin_map_new_ran_compat; auto. unfold G'. rewrite fun_to_fin_map_compat; auto. Qed. Definition exp_bool_plus {T:Type} {A:Ensemble T} (F F':Fin_map A (Full_set bool) false) := proj1_sig (constructive_definite_description _ (exp_bool_plus_ex F F')). Lemma exp_bool_plus_compat : forall {T:Type} (A:Ensemble T) (F F':Fin_map A (Full_set bool) false), let G := exp_bool_plus F F' in forall x:T, Inn A x -> G |-> x = F |-> x || F' |-> x. unfold exp_bool_plus; intros; destruct constructive_definite_description; auto. Qed. Definition exp_bool_comp_ex : forall {T:Type} {A:Ensemble T} (F:Fin_map A (Full_set bool) false), exists! (G:Fin_map A (Full_set bool) false), forall x:T, Inn A x -> G |-> x = negb (F |-> x). intros T A F. pose (fun_to_fin_map (fin_map_type_eq_dec F) _ false (fin_map_fin_dom F) (fun x => negb( F |-> x))) as G'. assert (h1:Included (Im A (fun x:T=>negb (F|->x))) (Full_set bool)). red. intros x h1. constructor. pose (fin_map_new_ran G' finite_bool h1) as G. exists G. red. split. intros x ha. unfold G. rewrite <- fin_map_new_ran_compat. unfold G'. rewrite fun_to_fin_map_compat; auto. assumption. intros G'' h2. apply fin_map_ext. intros x h3. rewrite h2; auto. unfold G. rewrite <- fin_map_new_ran_compat. unfold G'. rewrite fun_to_fin_map_compat; auto. assumption. Qed. Definition exp_bool_comp {T:Type} {A:Ensemble T} (F:Fin_map A (Full_set bool) false) := proj1_sig (constructive_definite_description _ (exp_bool_comp_ex F)). Lemma exp_bool_comp_compat : forall {T:Type} (A:Ensemble T) (F:Fin_map A (Full_set bool) false), let G := exp_bool_comp F in forall x:T, Inn A x -> G |-> x = negb (F |-> x). intros T A F G x. unfold G. unfold exp_bool_comp. destruct constructive_definite_description as [G' h2]. simpl. apply h2; auto. Qed. Definition exp_bool_bc {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (pf:Finite A) := Build_Bconst (Fin_map A (Full_set bool) false) (Full_set (Fin_map A (Full_set bool) false)) exp_bool_plus exp_bool_times (fin_map_sing_ran pfdt A (Full_set bool) pf finite_bool false true (Full_intro _ true)) (fin_map_sing_ran pfdt A (Full_set bool) pf finite_bool false false (Full_intro _ false)) exp_bool_comp. Lemma exp_bool_assoc_sum : forall {T:Type} {A:Ensemble T} {pf:Finite A} (pfdt:type_eq_dec T) (n m p : Btype (exp_bool_bc pfdt A pf)), n + (m + p) = n + m + p. intros T A h1 hdt f g k. simpl. apply fin_map_ext. intros x h2. do 4 (rewrite exp_bool_plus_compat; auto). rewrite orb_assoc; auto. Qed. Lemma exp_bool_assoc_prod : forall {T:Type} {A:Ensemble T} {pf:Finite A} (pfdt:type_eq_dec T) (n m p : Btype (exp_bool_bc pfdt A pf)), n * (m * p) = n * m * p. intros T A h1 hdt f g k. simpl. apply fin_map_ext. intros x h2. do 4 (rewrite exp_bool_times_compat; auto). apply andb_assoc. Qed. Lemma exp_bool_comm_sum : forall {T:Type} {A:Ensemble T} {pf:Finite A} (pfdt:type_eq_dec T) (n m : Btype (exp_bool_bc pfdt A pf)), n + m = m + n. intros T A h1 hdt f g. simpl. apply fin_map_ext. intros x h2. do 2 (rewrite exp_bool_plus_compat; auto). apply orb_comm. Qed. Lemma exp_bool_comm_prod : forall {T:Type} {A:Ensemble T} {pf:Finite A} (pfdt:type_eq_dec T) (n m : Btype (exp_bool_bc pfdt A pf)), n * m = m * n. intros T A h1 hdt f g. simpl. apply fin_map_ext. intros x h2. do 2 (rewrite exp_bool_times_compat; auto). apply andb_comm. Qed. Lemma exp_bool_abs_sum : forall {T:Type} {A:Ensemble T} {pf:Finite A} (pfdt:type_eq_dec T) (n m : Btype (exp_bool_bc pfdt A pf)), n + n * m = n. intros T A h1 hdt n m. apply fin_map_ext. intros x h2. simpl. rewrite exp_bool_plus_compat; auto. rewrite exp_bool_times_compat; auto. apply absoption_orb. Qed. Lemma exp_bool_abs_prod : forall {T:Type} {A:Ensemble T} {pf:Finite A} (pfdt:type_eq_dec T) (n m : Btype (exp_bool_bc pfdt A pf)), n * (n + m) = n. intros T A h1 hdt n m. apply fin_map_ext. intros x h2. simpl. rewrite exp_bool_times_compat; auto. rewrite exp_bool_plus_compat; auto. apply absoption_andb. Qed. Lemma exp_bool_dist_sum : forall {T:Type} {A:Ensemble T} {pf:Finite A} (pfdt:type_eq_dec T) (n m p : Btype (exp_bool_bc pfdt A pf)), p * (n + m) = p * n + p * m. intros T A h1 hdt f g k. apply fin_map_ext. intros x h2. simpl. rewrite exp_bool_times_compat; auto. do 2 (rewrite exp_bool_plus_compat; auto). do 2 (rewrite exp_bool_times_compat; auto). apply (fun a b c => andb_orb_distrib_r c a b). Qed. Lemma exp_bool_dist_prod : forall {T:Type} {A:Ensemble T} {pf:Finite A} (pfdt:type_eq_dec T) (n m p : Btype (exp_bool_bc pfdt A pf)), p + n * m = (p + n) * (p + m). intros T A h1 hdt f g k. apply fin_map_ext. intros x h2. simpl. rewrite exp_bool_plus_compat; auto. do 2 (rewrite exp_bool_times_compat; auto). do 2 (rewrite exp_bool_plus_compat; auto). apply (fun a b c => orb_andb_distrib_r c a b). Qed. Lemma exp_bool_comp_sum : forall {T:Type} {A:Ensemble T} {pf:Finite A} (pfdt:type_eq_dec T) (n : Btype (exp_bool_bc pfdt A pf)) , n + - n = 1. intros T A h1 hdt f. apply fin_map_ext. intros x h2. simpl. rewrite fin_map_sing_ran_compat; auto. rewrite exp_bool_plus_compat; auto. rewrite exp_bool_comp_compat; auto. apply orb_negb_r. Qed. Lemma exp_bool_comp_prod : forall {T:Type} {A:Ensemble T} {pf:Finite A} (pfdt:type_eq_dec T) (n : Btype (exp_bool_bc pfdt A pf)) , n * - n = 0. intros T A h1 hdt f. apply fin_map_ext. intros x h2. simpl. rewrite fin_map_sing_ran_compat; auto. rewrite exp_bool_times_compat; auto. rewrite exp_bool_comp_compat; auto. apply andb_negb_r. Qed. Definition exp_bool_ba {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (pf:Finite A) := Build_Bool_Alg (exp_bool_bc pfdt A pf) (eq_refl _) (exp_bool_assoc_sum pfdt) (exp_bool_assoc_prod pfdt) (exp_bool_comm_sum pfdt) (exp_bool_comm_prod pfdt) (exp_bool_abs_sum pfdt) (exp_bool_abs_prod pfdt) (exp_bool_dist_sum pfdt) (exp_bool_dist_prod pfdt) (exp_bool_comp_sum pfdt) (exp_bool_comp_prod pfdt). End ExpBoolBa. Section FamBa. (*Before I had employed parametric Bool_Alg records, I couldn't define [Ensemble Bool_Alg] (univesere inconsistency), so I had to dance around a universe inconsistency with a clumsy definition. Now that I have defined parametric analogues, the below defintion [fam_ba_p] works great, I can easily define but I decided to leave the old [fam_ba], just in case for some perverse reason I need to have an ensemble of Bool_Alg records without a parametric underlying type. *) Definition fam_ba_p (T:Type) := (Ensemble (Bool_Alg_p T)). (*the family of underlying sets in a family of parametric Boolean algebras*) Definition fam_ba_p_family {T:Type} (E:fam_ba_p T) := Im E (fun A=>(A_p T (Bc_p T A))). (*probably obsolete*) Inductive fam_ba : Type := ens_ba_intro :forall (A:Ensemble Type), (forall T:Type, Inn A T -> exists B:Bool_Alg, bt B = T) -> fam_ba. (*probably obsolete*) Definition fam_ba_types (S:fam_ba) : Ensemble Type. destruct S. refine A. Defined. (*another fumbled syntax required to avoid universe inconsistencies*) Definition in_fam_ba (S:fam_ba) (B:Bool_Alg) : Prop := exists T:Type, Inn (fam_ba_types S) T /\ T = bt B. End FamBa. Section TransferBA. Definition transfer_ba_elt {A B:Bool_Alg} (pf:A = B) (x:bt A) : bt B. subst; refine x. Defined. Definition transfer_ba_elt_r {A B:Bool_Alg} (pf:A = B) (x:bt B) : bt A. subst; refine x. Defined. Lemma transfer_ba_elt_undoes_transfer_ba_elt_r : forall {A B:Bool_Alg} (pf:A = B) (x:bt B), transfer_ba_elt pf (transfer_ba_elt_r pf x) = x. intros; subst; auto. Qed. Lemma transfer_ba_elt_r_undoes_transfer_ba_elt : forall {A B:Bool_Alg} (pf:A = B) (x:bt A), transfer_ba_elt_r pf (transfer_ba_elt pf x) = x. intros; subst; auto. Qed. Lemma transfer_ba_elt_inj : forall {A B:Bool_Alg} (pf:A = B) (x y:bt A), transfer_ba_elt pf x = transfer_ba_elt pf y -> x = y. intros A B h1 x y h2. subst. unfold transfer_ba_elt in h2. unfold eq_rect_r in h2. simpl in h2. assumption. Qed. Lemma transfer_ba_elt_r_inj : forall {A B:Bool_Alg} (pf:A = B) (x y:bt B), transfer_ba_elt_r pf x = transfer_ba_elt_r pf y -> x = y. intros A B h1 x y h2. subst. unfold transfer_ba_elt_r in h2. unfold eq_rect_r in h2. simpl in h2. assumption. Qed. Lemma transfer_dep_ba_p_eq : forall {T:Type} (Ap Bp:Bool_Alg_p T) (pf:Ap = Bp) (x:btp Ap), (@transfer_dep _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf x) = exist _ _ (iff1 (in_ba_p_ens_eq _ _ _ pf) (proj2_sig x)). intros; subst. rewrite transfer_dep_eq_refl. destruct x. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma transfer_dep_r_ba_p_eq : forall {T:Type} (Ap Bp:Bool_Alg_p T) (pf:Ap = Bp) (x:btp Bp), (@transfer_dep_r _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf x) = exist _ _ (iff2 (in_ba_p_ens_eq _ _ _ pf) (proj2_sig x)). intros; subst. rewrite transfer_dep_r_eq_refl. destruct x. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma transfer_dep_times : forall {T:Type} (Ap Bp:Bool_Alg_p T) (pf:Ap = Bp) (x y:btp Ap), @transfer_dep _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf (x %* y) = (@transfer_dep _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf x) %* (@transfer_dep _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf y). intros. subst. do 3 rewrite transfer_dep_eq_refl. reflexivity. Qed. Lemma transfer_dep_r_times : forall {T:Type} (Ap Bp:Bool_Alg_p T) (pf:Ap = Bp) (x y:btp Bp), @transfer_dep_r _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf (x %* y) = (@transfer_dep_r _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf x) %* (@transfer_dep_r _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf y). intros. subst. do 3 rewrite transfer_dep_r_eq_refl. reflexivity. Qed. Lemma transfer_dep_plus : forall {T:Type} (Ap Bp:Bool_Alg_p T) (pf:Ap = Bp) (x y:btp Ap), @transfer_dep _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf (x %+ y) = (@transfer_dep _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf x) %+ (@transfer_dep _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf y). intros. subst. do 3 rewrite transfer_dep_eq_refl. reflexivity. Qed. Lemma transfer_dep_r_plus : forall {T:Type} (Ap Bp:Bool_Alg_p T) (pf:Ap = Bp) (x y:btp Bp), @transfer_dep_r _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf (x %+ y) = (@transfer_dep_r _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf x) %+ (@transfer_dep_r _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf y). intros. subst. do 3 rewrite transfer_dep_r_eq_refl. reflexivity. Qed. Lemma transfer_dep_comp : forall {T:Type} (Ap Bp:Bool_Alg_p T) (pf:Ap = Bp) (x:btp Ap), @transfer_dep _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf (%- x) = %- (@transfer_dep _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf x). intros. subst. do 2 rewrite transfer_dep_eq_refl. reflexivity. Qed. Lemma transfer_dep_r_comp : forall {T:Type} (Ap Bp:Bool_Alg_p T) (pf:Ap = Bp) (x:btp Bp), @transfer_dep_r _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf (%- x) = %- (@transfer_dep_r _ (fun B=>{a:T | Inn (ba_p_ens B) a}) _ _ pf x). intros. subst. do 2 rewrite transfer_dep_r_eq_refl. reflexivity. Qed. End TransferBA. Section RandomBAFacts. Definition finite_ba (B:Bool_Alg) : Prop := Finite (ba_ens B). Definition finite_ba_p {T:Type} (Bp:Bool_Alg_p T) : Prop := Finite (ba_p_ens Bp). Definition countable_ba (B:Bool_Alg) : Prop := CountableT (bt B). Definition countable_ba_p {T:Type} (Bp:Bool_Alg_p T) : Prop := CountableT (btp Bp). Definition inf_countable_ba (B:Bool_Alg) : Prop := ~finite_ba B /\ countable_ba B. Definition inf_countable_ba_p {T:Type} (Bp:Bool_Alg_p T) : Prop := ~finite_ba_p Bp /\ countable_ba_p Bp. Lemma finite_ba_p_iff : forall {T:Type} (Bp:Bool_Alg_p T), finite_ba_p Bp <-> finite_ba (ba_conv Bp). intros T Bp. unfold finite_ba_p, finite_ba. rewrite ba_p_ens_eq. split. intro h1. rewrite finite_im_proj1_sig_iff in h1. assumption. intro h1. rewrite finite_im_proj1_sig_iff. assumption. Qed. Lemma countable_ba_p_iff : forall {T:Type} (Bp:Bool_Alg_p T), countable_ba_p Bp <-> countable_ba (ba_conv Bp). intros T Bp. unfold countable_ba_p, countable_ba. tauto. Qed. Lemma inf_countable_ba_p_iff : forall {T:Type} (Bp:Bool_Alg_p T), inf_countable_ba_p Bp <-> inf_countable_ba (ba_conv Bp). intros T Bp. unfold inf_countable_ba_p, inf_countable_ba. split. intro h1. destruct h1 as [h1 h2]. rewrite finite_ba_p_iff in h1. rewrite countable_ba_p_iff in h2. split; auto. intro h1. destruct h1 as [h1 h2]. split. rewrite finite_ba_p_iff; auto. rewrite countable_ba_p_iff; auto. Qed. Lemma finite_ba_bt_iff : forall (B:Bool_Alg), finite_ba B <-> FiniteT (bt B). intro B. split; intro h1. red in h1. rewrite <- Finite_FiniteT_iff. unfold bt. assumption. red. rewrite Finite_FiniteT_iff. assumption. Qed. Lemma finite_ba_btp_iff : forall {T:Type} (Bp:Bool_Alg_p T), finite_ba_p Bp <-> FiniteT (btp Bp). intros T B. split; intro h1. red in h1. unfold btp. unfold Btype_p. apply Finite_ens_type; auto. red. apply FiniteT_sig_Finite in h1. assumption. Qed. Lemma countable_ba_ba_ens_iff : forall (B:Bool_Alg), countable_ba B <-> Countable (ba_ens B). intro B. split. intro h1. apply countable_type_ensemble; auto. intro h1. red. rewrite countable_full_iff. assumption. Qed. Lemma countable_ba_ba_ens_iff_p : forall {T:Type} (Bp:Bool_Alg_p T), countable_ba_p Bp <-> Countable (ba_p_ens Bp). intros; split; auto. Qed. Lemma inf_countable_ba_non_trivial : forall (B:Bool_Alg), inf_countable_ba B -> exists b:bt B, b <> 0 /\ b <> 1. intros B h1. destruct h1 as [h1a h1b]. apply NNPP. intro h2. pose proof (not_ex_all_not _ (fun x:bt B => x <> 0 /\ x <> 1) h2) as h3. clear h2. simpl in h3. assert (h4:ba_ens B = Couple 0 1). apply Extensionality_Ensembles; red; split; red; intros x h4. specialize (h3 x). apply not_and_or in h3. destruct h3 as [h3 | h3]; apply NNPP in h3; subst. left. right. constructor. unfold finite_ba in h1a. rewrite h4 in h1a. contradict h1a. apply finite_couple. Qed. Lemma inf_countable_ba_p_index_iff : forall {T:Type} (Bp:Bool_Alg_p T), inf_countable_ba_p Bp <-> exists f:nat->btp Bp, bijective f. intros T Bp. split. intro h1. destruct h1 as [h1 h2]. apply inf_countable_index; auto. rewrite finite_ba_btp_iff in h1. assumption. intro h1. destruct h1 as [f h1]. red. split. pose proof infinite_nat as h2. intro h3. red in h3. assert (h5:Finite (Im (Full_set nat) f)). eapply Finite_downward_closed. rewrite finite_full_sig_iff in h3. apply h3. red; intros; constructor. pose proof (Pigeonhole_bis _ _ _ f h2 h5) as h4. destruct h1 as [h1 h1']. rewrite inj_image_fun_prop_compat in h1. contradiction. red. eapply surj_countable. apply countable_nat. destruct h1 as [? h1]. apply h1. Qed. Lemma inf_countable_ba_p_non_trivial : forall {T:Type} (Bp:Bool_Alg_p T), inf_countable_ba_p Bp -> exists b:btp Bp, b <> %0 /\ b <> %1. intros T B h1. rewrite inf_countable_ba_p_iff in h1. apply (inf_countable_ba_non_trivial _ h1). Qed. Lemma finite_ba_finite_subset : forall B:Bool_Alg, finite_ba B -> forall S:Ensemble (bt B), Finite S. intros B h1 S. eapply Finite_downward_closed. apply h1. red; intros; constructor. Qed. Definition ba_pos (B:Bool_Alg) := [x:bt B | lt 0 x]. Definition bt_pos (B:Bool_Alg) := {x:bt B | lt 0 x}. Lemma finite_ba_pos : forall (B:Bool_Alg), finite_ba B -> Finite (ba_pos B). intros; apply finite_ba_finite_subset; auto. Qed. Lemma finite_bt_pos : forall (B:Bool_Alg), finite_ba B -> FiniteT (bt_pos B). intros B h1. rewrite finite_ba_bt_iff in h1. eapply finite_subtype; auto. intros; apply classic. Qed. Lemma ba_pos_subtract : forall (B:Bool_Alg), ba_pos B = Subtract (ba_ens B) 0. intro B. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1]. constructor. constructor. intro h2. destruct h2. apply lt_irrefl in h1; auto. destruct h1 as [h1 h2]. constructor. rewrite <- neq_zero_lt_iff. intro h3. subst. contradict h2. constructor; auto. Qed. Lemma card_bt_pos : forall (B:Bool_Alg), finite_ba B -> card_fun1 (ba_ens B) = S (card_fun1 (ba_pos B)). intros B h1. rewrite ba_pos_subtract. apply card_sub_in. red in h1. unfold ba_ens. assumption. constructor. Qed. Lemma ba_pos_eq : forall (B:Bool_Alg), ba_pos B = im_proj1_sig (Full_set (bt_pos B)). intro B. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1]. apply Im_intro with (exist _ _ h1). constructor. simpl; auto. destruct h1 as [x h1]. subst. constructor. apply proj2_sig. Qed. Lemma card_bt_pos_ba_pos_eq : forall (B:Bool_Alg), finite_ba B -> card_fun1 (ba_pos B) = card_fun1 (Full_set (bt_pos B)). intros B h1. pose proof (ba_pos_eq B) as h2. rewrite h2. assert (hinj: Image.injective (bt_pos B) (bt B) (proj1_sig (P:=fun x : bt B => lt 0 x))). red. apply proj1_sig_injective. pose proof (card_fun1_compat (im_proj1_sig (Full_set (bt_pos B)))) as h3. pose proof (card_fun1_compat (Full_set (bt_pos B))) as h4. simpl in h3, h4. destruct h3 as [h3 h3'], h4 as [h4 h4']. apply finite_bt_pos in h1. rewrite <- Finite_FiniteT_iff in h1. assert (h5:Finite (im_proj1_sig (Full_set (bt_pos B)))). apply finite_image; auto. specialize (h3 h5). specialize (h4 h1). clear h5 h3' h4'. eapply injective_preserves_cardinal. apply hinj. apply h4. apply h3. Qed. Lemma trivial_ba_p_dec : forall {T:Type} (Bp:Bool_Alg_p T), ba_p_ens Bp = im_proj1_sig (Couple (Bzero_p T (Bc_p T Bp)) (Bone_p T (Bc_p T Bp))) -> forall x:btp Bp, {x = %0} + {x = %1}. intros T Bp h4 x. destruct x as [x h6]. unfold ba_p_ens in h4. pose proof h6 as h7. rewrite h4 in h7. assert (h8:x = proj1_sig (exist _ _ h6)). auto. rewrite h8 in h7. rewrite in_im_proj1_sig_iff in h7 at 1. apply couple_inv_dec in h7. destruct h7 as [h7 | h7]. left. apply proj1_sig_injective. simpl. pose proof (f_equal (@proj1_sig _ _) h7) as h9. simpl in h9. assumption. right. apply proj1_sig_injective. simpl. pose proof (f_equal (@proj1_sig _ _) h7) as h9. simpl in h9. assumption. Qed. End RandomBAFacts. bool2-v0-3/Ordinals.v0000644000175000017500000002570513575574055015315 0ustar dwyckoff76dwyckoff76(*Copyright (C) 2016, Daniel Wyckoff, except for the functions and theorems labeled "Schepler" which were written by Daniel Schepler in his "Zorn's Lemma" package.*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Import Classical. Require Import WellOrders. Require Import SetUtilities. Require Import Eqdep. Require Import TypeUtilities. Require Import Equality. Require Import Description. Require Import DecidableDec. Require Import FunctionProperties. Require Import WellOrders. Require Import FunctionalExtensionality. Require Import LogicUtilities. (*Schepler*) Inductive Ordinal : Type := | ordS : Ordinal -> Ordinal | ord_sup: forall {I:Type}, (I->Ordinal) -> Ordinal. (*Schepler*) Inductive ord_le : Ordinal -> Ordinal -> Prop := | ord_le_respects_succ: forall alpha beta:Ordinal, ord_le alpha beta -> ord_le (ordS alpha) (ordS beta) | ord_le_S_sup: forall (alpha:Ordinal) (J:Type) (beta:J->Ordinal) (j:J), ord_le (ordS alpha) (beta j) -> ord_le (ordS alpha) (ord_sup beta) | ord_sup_minimal: forall (I:Type) (alpha:I->Ordinal) (beta:Ordinal), (forall i:I, ord_le (alpha i) beta) -> ord_le (ord_sup alpha) beta. (*Schepler*) Definition ord_lt (alpha beta:Ordinal) := ord_le (ordS alpha) beta. (*Schepler*) Definition ord_eq (alpha beta:Ordinal) := ord_le alpha beta /\ ord_le beta alpha. (*Schepler*) Definition ord_ge (alpha beta:Ordinal) := ord_le beta alpha. (*Schepler*) Definition ord_gt (alpha beta:Ordinal) := ord_lt beta alpha. Delimit Scope ordinal_scope with Ordinal. Open Scope ordinal_scope. Notation "alpha < beta" := (ord_lt alpha beta) : ordinal_scope. Notation "alpha <= beta" := (ord_le alpha beta) : ordinal_scope. Notation "alpha == beta" := (ord_eq alpha beta) (at level 70) : ordinal_scope. Notation "alpha > beta" := (ord_gt alpha beta) : ordinal_scope. Notation "alpha >= beta" := (ord_ge alpha beta) : ordinal_scope. (*Schepler*) Lemma ord_le_respects_succ_converse: forall alpha beta:Ordinal, ordS alpha <= ordS beta -> alpha <= beta. Proof. intros. inversion_clear H. assumption. Qed. (*Schepler*) Lemma ord_le_S_sup_converse: forall (alpha:Ordinal) (J:Type) (beta:J->Ordinal), ordS alpha <= ord_sup beta -> exists j:J, ordS alpha <= beta j. Proof. intros. inversion H. exists j. assumption. Qed. (*Schepler*) Lemma ord_sup_minimal_converse: forall (I:Type) (alpha:I->Ordinal) (beta:Ordinal), ord_sup alpha <= beta -> forall i:I, alpha i <= beta. Proof. intros. inversion H. apply inj_pair2 in H2. destruct H2. apply H3. Qed. (*Schepler*) Lemma ord_le_trans: forall alpha beta gamma:Ordinal, alpha <= beta -> beta <= gamma -> alpha <= gamma. Proof. induction alpha. induction beta. induction gamma. intros. apply ord_le_respects_succ. apply IHalpha with beta. apply ord_le_respects_succ_converse; trivial. apply ord_le_respects_succ_converse; trivial. intros. apply ord_le_S_sup_converse in H1. destruct H1 as [i]. apply ord_le_S_sup with i. apply H; trivial. intros. pose proof (ord_sup_minimal_converse _ _ _ H1). apply ord_le_S_sup_converse in H0. destruct H0 as [i]. apply H with i; trivial. intros. pose proof (ord_sup_minimal_converse _ _ _ H0). constructor. intro. apply H with beta; trivial. Qed. (*Schepler*) Lemma ord_le_sup: forall (I:Type) (alpha:I->Ordinal) (i:I), alpha i <= ord_sup alpha. Proof. assert (forall beta:Ordinal, beta <= beta /\ forall (I:Type) (alpha:I->Ordinal) (i:I), beta <= alpha i -> beta <= ord_sup alpha). induction beta. destruct IHbeta. split. apply ord_le_respects_succ; trivial. intros. apply ord_le_S_sup with i. trivial. split. apply ord_sup_minimal. intro. destruct (H i). apply H1 with i; trivial. intros J alpha j ?. apply ord_sup_minimal. intro. destruct (H i). apply H2 with j. apply ord_le_trans with (ord_sup o). apply H2 with i; trivial. trivial. intros. destruct (H (alpha i)). apply H1 with i; trivial. Qed. (*Schepler*) Lemma ord_le_refl: forall alpha:Ordinal, alpha <= alpha. Proof. induction alpha. apply ord_le_respects_succ; trivial. apply ord_sup_minimal. apply ord_le_sup. Qed. (*Schepler*) Lemma ord_le_S: forall alpha:Ordinal, alpha <= ordS alpha. Proof. induction alpha. apply ord_le_respects_succ; trivial. apply ord_sup_minimal. intro. apply ord_le_trans with (ordS (o i)). apply H. apply ord_le_respects_succ. apply ord_le_sup. Qed. (*Wyckoff*) Lemma ord_lt_S: forall alpha:Ordinal, alpha < ordS alpha. intro alpha. red. constructor. apply ord_le_refl. Qed. (*Schepler*) Lemma ord_lt_le: forall alpha beta:Ordinal, alpha < beta -> alpha <= beta. Proof. intros. apply ord_le_trans with (ordS alpha); trivial. apply ord_le_S. Qed. (*Schepler*) Lemma ord_lt_le_trans: forall alpha beta gamma:Ordinal, alpha < beta -> beta <= gamma -> alpha < gamma. Proof. intros. apply ord_le_trans with beta; trivial. Qed. (*Schepler*) Lemma ord_le_lt_trans: forall alpha beta gamma:Ordinal, alpha <= beta -> beta < gamma -> alpha < gamma. Proof. intros. apply ord_le_trans with (ordS beta); trivial. apply ord_le_respects_succ; trivial. Qed. (*Schepler*) Lemma ord_lt_trans: forall alpha beta gamma:Ordinal, alpha < beta -> beta < gamma -> alpha < gamma. Proof. intros. apply ord_lt_le_trans with beta; trivial; apply ord_lt_le; trivial. Qed. (*Schepler*) Lemma ord_lt_respects_succ: forall alpha beta:Ordinal, alpha < beta -> ordS alpha < ordS beta. Proof. intros. apply ord_le_respects_succ; trivial. Qed. (*Schepler*) Lemma ord_total_order: forall alpha beta:Ordinal, alpha < beta \/ alpha == beta \/ alpha > beta. Proof. induction alpha. induction beta. destruct (IHalpha beta) as [|[|]]. left; apply ord_lt_respects_succ; trivial. right; left. split. apply ord_le_respects_succ; apply H. apply ord_le_respects_succ; apply H. right; right. apply ord_lt_respects_succ; trivial. destruct (classic (exists i:I, ordS alpha < o i)). destruct H0 as [i]. left. apply ord_lt_le_trans with (o i); trivial. apply ord_le_sup. destruct (classic (exists i:I, ordS alpha == o i)). destruct H1 as [i]. right; left. split. apply ord_le_trans with (o i). apply H1. apply ord_le_sup. apply ord_sup_minimal. intro. destruct (H i0) as [|[|]]. contradiction H0; exists i0; trivial. apply H2. apply ord_lt_le; trivial. assert (forall i:I, ordS alpha > o i). intros. destruct (H i) as [|[|]]. contradiction H0; exists i; trivial. contradiction H1; exists i; trivial. trivial. right; right. apply ord_le_lt_trans with alpha. apply ord_sup_minimal. intro. apply ord_le_respects_succ_converse. apply H2. apply ord_le_refl. induction beta. case (classic (exists i:I, o i > ordS beta)); intro. destruct H0 as [i]. right; right. apply ord_lt_le_trans with (o i); trivial. apply ord_le_sup. case (classic (exists i:I, o i == ordS beta)); intro. right; left. destruct H1 as [i]. split. apply ord_sup_minimal. intro j. destruct (H j (ordS beta)) as [|[|]]. apply ord_lt_le; trivial. apply H2. contradiction H0; exists j; trivial. apply ord_le_trans with (o i). apply H1. apply ord_le_sup. left. apply ord_le_respects_succ. apply ord_sup_minimal. intro. destruct (H i (ordS beta)) as [|[|]]. apply ord_le_respects_succ_converse; trivial. contradiction H1; exists i; trivial. contradiction H0; exists i; trivial. case (classic (exists j:I0, ord_sup o < o0 j)); intro. left. destruct H1 as [j]. apply ord_lt_le_trans with (o0 j); trivial. apply ord_le_sup. case (classic (exists i:I, o i > ord_sup o0)); intro. destruct H2 as [i]. right; right. apply ord_lt_le_trans with (o i); trivial. apply ord_le_sup. right; left. split. apply ord_sup_minimal; intro. destruct (H i (ord_sup o0)) as [|[|]]. apply ord_lt_le; trivial. apply H3. contradiction H2; exists i; trivial. apply ord_sup_minimal; intro j. destruct (H0 j) as [|[|]]. contradiction H1; exists j; trivial. apply H3. apply ord_lt_le; trivial. Qed. (*Schepler*) Lemma ordinals_well_founded: well_founded ord_lt. Proof. red; intro alpha. induction alpha. constructor. intros beta ?. apply ord_le_respects_succ_converse in H. constructor; intros gamma ?. destruct IHalpha. apply H1. apply ord_lt_le_trans with beta; trivial. constructor; intros alpha ?. apply ord_le_S_sup_converse in H0. destruct H0 as [j]. destruct (H j). apply H1; trivial. Qed. (*Schepler*) Lemma ord_lt_irrefl: forall alpha:Ordinal, ~(alpha < alpha). Proof. intro; red; intro. assert (forall beta:Ordinal, beta <> alpha). intro. pose proof (ordinals_well_founded beta). induction H0. red; intro. symmetry in H2; destruct H2. contradiction (H1 alpha H); trivial. contradiction (H0 alpha); trivial. Qed. (*Schepler*) Inductive successor_ordinal : Ordinal->Prop := | intro_succ_ord: forall alpha:Ordinal, successor_ordinal (ordS alpha) | succ_ord_wd: forall alpha beta:Ordinal, successor_ordinal alpha -> alpha == beta -> successor_ordinal beta. (*Schepler*) Inductive limit_ordinal : Ordinal->Prop := | intro_limit_ord: forall {I:Type} (alpha:I->Ordinal), (forall i:I, exists j:I, alpha i < alpha j) -> limit_ordinal (ord_sup alpha) | limit_ord_wd: forall alpha beta:Ordinal, limit_ordinal alpha -> alpha == beta -> limit_ordinal beta. (*Schepler*) Lemma ord_successor_or_limit: forall alpha:Ordinal, successor_ordinal alpha \/ limit_ordinal alpha. Proof. induction alpha. left; constructor. destruct (classic (forall i:I, exists j:I, o i < o j)). right; constructor; trivial. destruct (not_all_ex_not _ _ H0) as [i]. assert (forall j:I, o j <= o i). intro. destruct (ord_total_order (o i) (o j)) as [|[|]]. contradiction H1; exists j; trivial. apply H2. apply ord_lt_le; trivial. assert (ord_sup o == o i). split. apply ord_sup_minimal; trivial. apply ord_le_sup. case (H i); intro. left; apply succ_ord_wd with (o i); trivial. split; apply H3. right. apply limit_ord_wd with (o i); trivial. split; apply H3. Qed. (*Schepler*) Lemma successor_ordinal_not_limit: forall alpha:Ordinal, successor_ordinal alpha -> ~ limit_ordinal alpha. Proof. intros; red; intro. induction H. inversion_clear H0. induction H as [I beta|]. assert (ord_sup beta <= alpha). apply ord_sup_minimal. intro. apply ord_le_respects_succ_converse. destruct (H i) as [j]. apply ord_le_trans with (beta j); trivial. apply ord_le_trans with (ord_sup beta). apply ord_le_sup. apply H1. contradiction (ord_lt_irrefl alpha). apply ord_le_trans with (ord_sup beta); trivial. apply H1. apply IHlimit_ordinal. split; apply ord_le_trans with beta; (apply H0 || apply H1). contradiction IHsuccessor_ordinal. apply limit_ord_wd with beta; trivial. split; apply H1. Qed. bool2-v0-3/Makefile.bak0000644000175000017500000001516413575574030015536 0ustar dwyckoff76dwyckoff76############################################################################# ## v # The Coq Proof Assistant ## ## "$@" || ( RV=$$?; rm -f "$@"; exit $${RV} ) %.v.beautified: $(COQC) $(COQDEBUG) $(COQFLAGS) -beautify $* # WARNING # # This Makefile has been automagically generated # Edit at your own risks ! # # END OF WARNING bool2-v0-3/Relation_Definitions_Implicit.v0000644000175000017500000000337313575574055021501 0ustar dwyckoff76dwyckoff76(*Copyright (C) 2016, Daniel Wyckoff.*) (*This file is an upgrade of a complete copy and paste from Daniel Schepler's "Zorn's Lemma." *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Export Relation_Definitions. Arguments reflexive [A] _. Arguments transitive [A] _. Arguments symmetric [A] _. Arguments antisymmetric [A] _. Arguments equiv [A] _. Arguments preorder [A] _. Arguments Build_preorder [A] _ _ _. Arguments preord_refl [A] [R] _ _. Arguments preord_trans [A] [R] _ _ _ _ _ _. Arguments order [A] _. Arguments Build_order [A] _ _ _ _. Arguments ord_refl [A] [R] _ _. Arguments ord_trans [A] [R] _ _ _ _ _ _. Arguments ord_antisym [A] [R] _ _ _ _ _. Arguments equivalence [A] _. Arguments Build_equivalence [A] _ _ _ _. Arguments equiv_refl [A] [R] _ _. Arguments equiv_trans [A] [R] _ _ _ _ _ _. Arguments equiv_sym [A] [R] _ _ _ _. Arguments PER [A] _. Arguments Build_PER [A] _ _ _. Arguments per_sym [A] [R] _ _ _ _. Arguments per_trans [A] [R] _ _ _ _ _ _. Arguments inclusion [A] _ _. Arguments same_relation [A] _ _. Arguments commut [A] _ _.bool2-v0-3/TypeUtilities.v0000644000175000017500000065651113575574055016364 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-2019, Daniel Wyckoff, except for the portions so labeleled which I got from Daniel Schepler*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Import ProofIrrelevance. Require Import FunctionalExtensionality. Require Export Ensembles. Require Import NaryFunctions. Require Import Image. Require Import EqdepFacts. Require Import ClassicalChoice. Require Import Arith. Require Import Basics. Require Import Omega. Require Import Description. Require QArith. Require ZArith. Require Import Equality. Require Import Setoid. Require Export EnsemblesImplicit. Require Import FunctionProperties. Require Export Finite_sets. Require Import ImageImplicit. Require Import DecidableDec. Require Import SetUtilities. Require Import CSB. Require Import LogicUtilities. Require Import ArithUtilities. Require Import EqDec. (*Wyckoff*) Lemma type_eq_dec_false : type_eq_dec False. red; intro; contradiction. Qed. (*Wyckoff*) Lemma unit_eq : forall (u u':unit), u = u'. intros u u'; destruct u, u'; auto. Qed. (*Wyckoff*) (*Equinumerous relation on arbitrary types.*) Inductive equi (T U:Type) : Prop := | equi_intro : forall f:T->U, invertible f -> equi T U. (*Wyckoff*) Lemma equi_refl : forall T, equi T T. intro T. econstructor. apply invertible_id. Qed. (*Wyckoff*) Lemma equi_sym : forall T U, equi T U -> equi U T. intros ? ? h1. destruct h1. econstructor. apply invertible_impl_inv_invertible. Grab Existential Variables. apply H. Qed. (*Wyckoff*) Lemma equi_trans : forall T U V, equi T U -> equi U V -> equi T V. intros ? ? ? h1 h2; destruct h1, h2; apply equi_intro with (compose f0 f); apply invertible_compose; auto. Qed. (*Wyckoff*) Lemma some_neq_none : forall {T} (x:T), Some x <> None. intro; discriminate. Qed. (*Wyckoff*) Lemma none_dec : forall {T} (o:option T), {o = None} + {o <> None}. intros T o. destruct o. right; intro; discriminate. left; auto. Qed. (*Wyckoff*) Definition neq_none_inv {T} (o:option T) : o <> None -> T := match o with | Some x => fun _ => x | None => fun pf => False_rect _ (pf (eq_refl _)) end. (*Wyckoff*) Lemma inj_neq_none_inv : forall {T} (o o':option T) (pfo: o <> None) (pfo':o' <> None), neq_none_inv o pfo = neq_none_inv o' pfo' -> o = o'. unfold neq_none_inv. intros ? o o'; destruct o, o'; intros ho ho' ?; subst; auto; [contradict ho'; auto | contradict ho; auto]. Qed. (*Wyckoff*) Lemma some_neq_none_inv : forall {T} (o:option T) (pf:o <> None), Some (neq_none_inv o pf) = o. intros T o; destruct o; auto. intro h1; contradict h1; auto. Qed. (*Wyckoff*) Lemma simpl_sig : forall {T:Type} (P:T->Prop) (a:{x:T | P x}) (pf:P (proj1_sig a)), exist _ _ pf = a. intros T P a h1. destruct a as [a h2]. simpl. apply proj1_sig_injective. simpl. reflexivity. Qed. (*Many of these functions are copied and pasted from Daniel Schepler's "Zorn's Lemma" user contribution. Those which are are commented "Schepler." Those which aren't are commented "Wyckoff"*) Section FiniteTypes. (*Schepler*) Inductive FiniteT : Type -> Prop := | empty_finite: FiniteT False | add_finite: forall T:Type, FiniteT T -> FiniteT (option T) | bij_finite: forall (X Y:Type) (f:X->Y), FiniteT X -> invertible f -> FiniteT Y. (*Wyckoff*) Lemma FiniteT_functional : forall {T U}, T = U -> FiniteT T -> FiniteT U. intros; subst; auto. Qed. (*Schepler*) Lemma True_finite: FiniteT True. Proof. apply bij_finite with (option False) (fun _ => I). constructor; constructor. exists (True_rect None). destruct x as [[]|]. remember (True_rect (@None False) I) as LHS. destruct LHS as [[]|]. reflexivity. exact (fun y:True => match y with | I => refl_equal I end). Qed. (*Schepler*) Lemma finite_dec_exists: forall (X:Type) (P:X->Prop), FiniteT X -> (forall x:X, {P x} + {~ P x}) -> { exists x:X, P x } + { forall x:X, ~ P x }. Proof. intros. apply exclusive_dec. red; intro. destruct H0. destruct H0. contradiction (H1 x). revert P X0. induction H. right. destruct x. intros. case (IHFiniteT (fun x:T => P (Some x)) (fun x:T => X0 (Some x))). left. destruct H0. exists (Some x). assumption. intro. case (X0 None). left. exists None. assumption. right. destruct x. apply H0. assumption. destruct H0. intros. case (IHFiniteT (fun x:X => P (f x)) (fun x:X => X0 (f x))). left. destruct H2. exists (f x). assumption. right. intro. rewrite <- H1 with x. apply H2. Qed. (*Schepler*) Lemma finite_dec_forall: forall (X:Type) (P:X->Prop), FiniteT X -> (forall x:X, { P x } + { ~ P x }) -> { forall x:X, P x } + { exists x:X, ~ P x }. Proof. intros. apply exclusive_dec. intuition. destruct H2. contradiction (H1 x). revert P X0. induction H. left. destruct x. intros. case (IHFiniteT (fun x:T => P (Some x)) (fun x:T => X0 (Some x))). intro. case (X0 None). left. destruct x. apply H0. assumption. right. exists None. assumption. right. destruct H0. exists (Some x). assumption. intros. destruct H0. case (IHFiniteT (fun x:X => P (f x)) (fun x:X => X0 (f x))). left. intro y. rewrite <- H1. apply H2. right. destruct H2. exists (f x). assumption. Qed. (*Schepler*) Lemma FiniteT_img: forall (X Y:Type) (f:X->Y), FiniteT X -> (forall y1 y2:Y, y1=y2 \/ y1<>y2) -> Finite (Im (Full_set _) f). Proof. intros. induction H. assert (Im (Full_set _) f = Empty_set _). apply Extensionality_Ensembles. red; split. red; intros. destruct H. destruct x. auto with sets. rewrite H. constructor. assert ({exists x:T, f (Some x) = f None} + {forall x:T, f (Some x) <> f None}). apply finite_dec_exists. assumption. intro. apply decidable_dec. apply H0. case H1. intro. pose (g := fun (x:T) => f (Some x)). assert (Im (Full_set _) f = Im (Full_set _) g). apply Extensionality_Ensembles. red; split. red; intros. destruct H2. destruct x. exists t. constructor. assumption. destruct e. exists x. constructor. transitivity (f None). assumption. symmetry; assumption. red; intros. destruct H2. exists (Some x). constructor. assumption. rewrite H2. apply IHFiniteT. intros. pose (g := fun x:T => f (Some x)). assert (Im (Full_set _) f = Add (Im (Full_set _) g) (f None)). apply Extensionality_Ensembles. red; split. red; intros. destruct H2. destruct x. left. exists t. constructor. assumption. right. auto with sets. red; intros. destruct H2. destruct H2. exists (Some x). constructor. assumption. destruct H2. exists None. constructor. reflexivity. rewrite H2. constructor. apply IHFiniteT. red; intro. destruct H3. contradiction (n x). symmetry; assumption. pose (g := fun (x:X) => f (f0 x)). assert (Im (Full_set _) f = Im (Full_set _) g). apply Extensionality_Ensembles. red; split. red; intros. destruct H2. destruct H1. rewrite H3. rewrite <- H4 with x. exists (g0 x). constructor. unfold g. reflexivity. red; intros. destruct H2. exists (f0 x). constructor. assumption. rewrite H2. apply IHFiniteT. Qed. (*Schepler*) Lemma finite_subtype: forall (X:Type) (P:X->Prop), FiniteT X -> (forall x:X, P x \/ ~ P x) -> FiniteT {x:X | P x}. Proof. intros. induction H. apply bij_finite with False (False_rect _). constructor. exists (@proj1_sig _ _). destruct x. intro s; destruct s; destruct x. destruct (H0 None). pose (g := fun (x:option {x:T | P (Some x)}) => match x return {x:option T | P x} with | Some (exist x0 i) => exist (fun x:option T => P x) (Some x0) i | None => exist (fun x:option T => P x) None H1 end). apply bij_finite with _ g. apply add_finite. apply IHFiniteT. intro; apply H0. pose (ginv := fun (s:{x0:option T | P x0}) => match s return option {x:T | P (Some x)} with | exist (Some x0) i => Some (exist (fun y:T => P (Some y)) x0 i) | exist None _ => None end). exists ginv. destruct x as [[x0]|]. simpl. reflexivity. simpl. reflexivity. destruct y as [[x0|]]. simpl. reflexivity. simpl. destruct (proof_irrelevance _ H1 p). reflexivity. pose (g := fun (x:{x:T | P (Some x)}) => match x return {x:option T | P x} with | exist x0 i => exist (fun x:option T => P x) (Some x0) i end). apply bij_finite with _ g. apply IHFiniteT. intro; apply H0. pose (ginv := fun s:{x0:option T | P x0} => match s return {x:T | P (Some x)} with | exist (Some x0) i => exist (fun x:T => P (Some x)) x0 i | exist None i => False_rect _ (H1 i) end). exists ginv. destruct x; simpl. reflexivity. destruct y as [[x0|]]. simpl. reflexivity. contradiction H1. pose (g := fun (x:{x:X | P (f x)}) => match x with | exist x0 i => exist (fun x:Y => P x) (f x0) i end). apply bij_finite with _ g. apply IHFiniteT. intro; apply H0. destruct H1. assert (forall y:Y, P y -> P (f (g0 y))). intros; rewrite H2; assumption. pose (ginv := fun (y:{y:Y | P y}) => match y with | exist y0 i => exist (fun x:X => P (f x)) (g0 y0) (H3 y0 i) end). exists ginv. destruct x; simpl. generalize (H3 (f x) p). rewrite H1. intro; destruct (proof_irrelevance _ p p0). reflexivity. destruct y; simpl. generalize (H3 x p). rewrite H2. intro; destruct (proof_irrelevance _ p p0). reflexivity. Qed. (*Wyckoff*) Lemma FiniteT_Finite : forall (T:Type), FiniteT T -> (forall (S:Ensemble T), Finite S). intros T h1 S. assert (h2: forall x:T, In S x \/ ~In S x). intro. tauto. pose proof (finite_subtype T _ h1 h2) as h3. assert (h4: forall y1 y2 : T, y1 = y2 \/ y1 <> y2). intros. tauto. pose proof (FiniteT_img (sig_set S) T (@proj1_sig T _) h3 h4) as h5. assert (h6: (Im (full_sig S) (proj1_sig (P:=fun x : T => In S x))) = S). apply Extensionality_Ensembles. red. unfold Included. split. (*left*) intros x h7. inversion h7 as [y h8 z h9]. pose proof (proj2_sig y) as h10. rewrite <- h9 in h10. assumption. (*right*) intros x h11. apply Im_intro with (exist _ x h11). apply Full_intro. simpl. reflexivity. unfold full_sig in h6. rewrite h6 in h5. assumption. Qed. (*Wyckoff*) Lemma FiniteT_sig_Finite : forall {T:Type} (S:Ensemble T), FiniteT (sig_set S) -> Finite S. intros T S h1. pose proof (FiniteT_img _ _ (@proj1_sig _ _) h1) as h2. assert (h3: forall y1 y2 : T, y1 = y2 \/ y1 <> y2). intros. apply classic. apply h2 in h3. assert (h4:(Im (full_sig S) (proj1_sig (P:=fun x : T => In S x))) = S). apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h4. inversion h4 as [? ? ? h5]. rewrite h5. apply proj2_sig. (* >= *) red. intros x h4. eapply Im_intro. apply (Full_intro _ (exist _ x h4)). reflexivity. unfold full_sig in h4. unfold sig_set in h4. rewrite h4 in h3; assumption. Qed. (*Schepler*) Lemma Finite_ens_type: forall {X:Type} (S:Ensemble X), Finite S -> FiniteT (sig_set S). Proof. intros. induction H. apply bij_finite with False (False_rect _). constructor. assert (g:{x:X | In (Empty_set _) x}->False). intro. destruct X0. destruct i. exists g. destruct x. destruct y. destruct g. assert (Included A (Add A x)). auto with sets. assert (In (Add A x) x). auto with sets. pose (g := fun (y: option {x:X | In A x}) => match y return {x0:X | In (Add A x) x0} with | Some (exist y0 i) => exist (fun x2:X => In (Add A x) x2) y0 (H1 y0 i) | None => exist (fun x2:X => In (Add A x) x2) x H2 end). apply bij_finite with _ g. apply add_finite. assumption. assert (h:forall x0:X, In (Add A x) x0 -> { In A x0 } + { x0 = x }). intros; apply exclusive_dec. intuition. destruct H6; auto. destruct H3. left; assumption. right; destruct H3; reflexivity. pose (ginv := fun s:{x0:X | In (Add A x) x0} => match s return option {x:X | In A x} with | exist x0 i => match (h x0 i) with | left iA => Some (exist _ x0 iA) | right _ => None end end). exists ginv. intro; destruct x0. destruct s. simpl. remember (h x0 (H1 x0 i)) as sum; destruct sum. destruct (proof_irrelevance _ i i0). reflexivity. contradiction H0. rewrite <- e; assumption. simpl. remember (h x H2) as sum; destruct sum. contradiction H0. reflexivity. intro. unfold ginv. destruct y. destruct (h x0 i). simpl. generalize (H1 x0 i0); intro. destruct (proof_irrelevance _ i i1). reflexivity. simpl. destruct e. destruct (proof_irrelevance _ H2 i). reflexivity. Qed. (*Wyckoff*) Lemma FiniteT_Full : forall T:Type, FiniteT T <-> FiniteT {x:T | In (Full_set T) x}. intros T. pose (fun (a:T) => exist _ a (Full_intro _ a)) as f. assert (h2:invertible f). apply bijective_impl_invertible. red. split. (*inj*) red. intros x1 x2 h3. unfold f in h3. assert (h4: x1 = proj1_sig (exist (In (Full_set T)) x1 (Full_intro T x1))). simpl. reflexivity. assert (h5: x2 = proj1_sig (exist (In (Full_set T)) x2 (Full_intro T x2))). simpl. reflexivity. congruence. (*surj*) red. intro y. exists (proj1_sig y). destruct y as [z h1]. unfold f. simpl. pose proof (proof_irrelevance _ h1 (Full_intro T z)) as h2. rewrite h2. reflexivity. pose (fun (a:{x:T | In (Full_set T) x}) => proj1_sig a) as f'. assert (h3:invertible f'). apply bijective_impl_invertible. red. split. (*inj*) red. intros x1 x2 h3. unfold f' in h3. apply proj1_sig_injective; trivial. (*surj*) red. intro y. exists (exist _ y (Full_intro _ y)). unfold f'. simpl. reflexivity. split. intro h1. apply (bij_finite _ _ f); trivial. intro h1. apply (bij_finite _ _ f'); trivial. Qed. (*Wyckoff*) Lemma Finite_FiniteT : forall (T:Type), Finite (Full_set T) -> FiniteT T. intros T h1. rewrite FiniteT_Full. apply Finite_ens_type; trivial. Qed. (* Wyckoff *) Lemma Finite_FiniteT_iff : forall (T:Type), Finite (Full_set T) <-> FiniteT T. intro T. split. apply Finite_FiniteT. intro h1. apply FiniteT_Finite. assumption. Qed. (* Wyckoff *) Lemma finite_full_sig_iff : forall {T:Type} (A:Ensemble T), Finite A <-> Finite (full_sig A). intros T A. split. intro h1. pose proof (Finite_ens_type _ h1) as h2. rewrite <- Finite_FiniteT_iff in h2. assumption. intro h1. rewrite Finite_FiniteT_iff in h1. apply FiniteT_sig_Finite. assumption. Qed. (*Wyckoff*) Lemma finite_subset_sig_iff : forall {T:Type} (A B:Ensemble T) (pf:Included A B), Finite A <-> Finite (subset_sig _ _ pf). intros T A B h1. split. intro h2. red in h1. unfold subset_sig. apply finite_image. rewrite <- finite_full_sig_iff. assumption. intro h2. unfold subset_sig in h2. assert (h3:Injective (fun x : {x : T | In A x} => exist (In B) (proj1_sig x) (h1 (proj1_sig x) (proj2_sig x)))). red. intros x1 x2 h3. pose proof (exist_injective _ _ _ _ _ h3) as h4. apply proj1_sig_injective. assumption. pose proof (finite_image_rev_inj _ _ h3 h2) as h4. rewrite finite_full_sig_iff. assumption. Qed. (*Wyckoff*) Lemma finite_incl_sig_iff : forall {T:Type} (A B:Ensemble T) (pf:Included A B), Finite (incl_sig A B pf) <-> Finite A. intros T A B h1. rewrite incl_sig_eq_subset_sig. symmetry. apply finite_subset_sig_iff. Qed. (*Wyckoff*) Lemma card_fun_full_sig_eq : forall {T:Type} (A:Ensemble T) (pf:Finite A), card_fun A pf = card_fun (full_sig A) (iff1 (finite_full_sig_iff A) pf). intros T A h1. pose proof (card_fun_compat _ h1) as h2. pose proof (card_fun_compat _ (iff1 (finite_full_sig_iff A) h1)) as h3. pose proof (proj1_sig_injective (fun x=>In A x)) as h4. eapply injective_preserves_cardinal. red. apply h4. apply h3. pose proof (im_full_sig_proj1_sig A) as h5. rewrite <- h5 at 1. assumption. Qed. (*Schepler*) Lemma finite_eq_dec: forall X:Type, FiniteT X -> forall x y:X, {x=y} + {x<>y}. Proof. intros. apply decidable_dec. induction H. destruct x. decide equality. destruct H0. case (IHFiniteT (g x) (g y)). left. rewrite <- H1. rewrite <- H1 with x. rewrite H2. reflexivity. right. contradict H2. rewrite H2. reflexivity. Qed. (*Schepler*) Lemma finite_inj_surj: forall (X:Type) (f:X->X), FiniteT X -> Injective f -> surjective f. Proof. intros. induction H. red. destruct y. remember (f None) as f0; destruct f0 as [a|]. assert (forall x:T, f (Some x) <> Some a). unfold not; intros. assert (Some x = None). apply H0. congruence. discriminate H2. pose (g := fun x:T => match f (Some x) with | Some y => y | None => a end). assert (surjective g). apply IHFiniteT. red; intros. remember (f (Some x1)) as fx1; destruct fx1; remember (f (Some x2)) as fx2; destruct fx2. unfold g in H2. rewrite <- Heqfx1 in H2; rewrite <- Heqfx2 in H2. destruct H2; assert (f (Some x1) = f (Some x2)). congruence. apply H0 in H2. injection H2; trivial. unfold g in H2; rewrite <- Heqfx1 in H2; rewrite <- Heqfx2 in H2. destruct H2. contradiction (H1 x1). symmetry; assumption. unfold g in H2; rewrite <- Heqfx1 in H2; rewrite <- Heqfx2 in H2. destruct H2. contradiction (H1 x2). symmetry; assumption. assert (Some x1 = Some x2). apply H0. congruence. injection H3; trivial. red; intro. destruct y. case (finite_eq_dec _ H t a). exists None. congruence. destruct (H2 t). exists (Some x). unfold g in H3. destruct (f (Some x)). congruence. contradiction n. symmetry; assumption. destruct (H2 a). exists (Some x). unfold g in H3. remember (f (Some x)) as fx; destruct fx. destruct H3. contradiction (H1 x). symmetry; assumption. reflexivity. assert (forall x:T, { y:T | f (Some x) = Some y }). intros. remember (f (Some x)) as fx; destruct fx. exists t; reflexivity. assert (Some x = None). apply H0. congruence. discriminate H1. pose (g := fun x:T => proj1_sig (X x)). assert (surjective g). apply IHFiniteT. red; intros. unfold g in H1. repeat destruct X in H1. simpl in H1. assert (Some x1 = Some x2). apply H0. congruence. injection H2; trivial. red; intro. destruct y. destruct (H1 t). unfold g in H2; destruct X in H2. simpl in H2. exists (Some x). congruence. exists None. symmetry; assumption. destruct H1. pose (f' := fun (x:X) => g (f (f0 x))). assert (surjective f'). apply IHFiniteT. red; intros. unfold f' in H3. assert (f (f0 x1) = f (f0 x2)). congruence. apply H0 in H4. congruence. red; intro. destruct (H3 (g y)). unfold f' in H4. exists (f0 x). congruence. Qed. (*Schepler *) Lemma surj_finite: forall (X Y:Type) (f:X->Y), FiniteT X -> surjective f -> (forall y1 y2:Y, y1=y2 \/ y1<>y2) -> FiniteT Y. Proof. intros. apply bij_finite with {y:Y | In (Im (Full_set _) f) y} (@proj1_sig _ (fun y:Y => In (Im (Full_set _) f) y)). apply Finite_ens_type. apply FiniteT_img. assumption. assumption. assert (forall y:Y, In (Im (Full_set _) f) y). intro. destruct (H0 y). exists x; auto with sets. constructor. pose (proj1_sig_inv := fun y:Y => exist (fun y0:Y => In (Im (Full_set _) f) y0) y (H2 y)). exists proj1_sig_inv. destruct x. simpl. unfold proj1_sig_inv. destruct (proof_irrelevance _ (H2 x) i); trivial. intros; simpl; reflexivity. Qed. (*Schepler*) Lemma inj_finite: forall (X Y:Type) (f:X->Y), FiniteT Y -> Injective f -> (forall y:Y, (exists x:X, f x = y) \/ (~ exists x:X, f x = y)) -> FiniteT X. Proof. intros. assert (forall y:{y:Y | exists x:X, f x = y}, {x:X | f x = proj1_sig y}). intro. destruct y. simpl. apply constructive_definite_description. destruct e. exists x0. red; split. assumption. intros. apply H0. transitivity x. assumption. symmetry; assumption. pose (g := fun y:{y:Y | exists x:X, f x = y} => proj1_sig (X0 y)). apply bij_finite with _ g. apply finite_subtype. assumption. assumption. pose (ginv := fun (x:X) => exist (fun y:Y => exists x:X, f x = y) (f x) (ex_intro _ x (refl_equal _))). exists ginv. destruct x as [y [x e]]. unfold g; simpl. match goal with |- context [X0 ?arg] => destruct (X0 arg) end. simpl. unfold ginv; simpl. simpl in e0. repeat match goal with |- context [ex_intro ?f ?x ?e] => generalize (ex_intro f x e) end. rewrite <- e0. intros; destruct (proof_irrelevance _ e1 e2). reflexivity. intro; unfold ginv. unfold g; simpl. match goal with |- context [X0 ?arg] => destruct (X0 arg) end. simpl. simpl in e. auto. Qed. (*Wyckoff*) Lemma finite_sig_sat_compat : forall {T:Type} (P:T->Prop), FiniteT {x:T | P x} <-> Finite [x:T | P x]. intros T P. split. intro h1. rewrite finite_full_sig_iff. unfold full_sig. unfold sig_set. rewrite Finite_FiniteT_iff. pose (fun a:{x:T| In [x0:T | P x0] x} => (exist _ (proj1_sig a) (match (proj2_sig a) with |intro_characteristic_sat pf => pf end))) as f. assert (h2:Injective f). red. intros a b h2. destruct a; destruct b. destruct i; destruct i0. unfold f in h2. simpl in h2. apply proj1_sig_injective. simpl. apply exist_injective in h2. assumption. eapply inj_finite; auto. apply h1. apply h2. intros; apply classic. intro h2. rewrite finite_full_sig_iff in h2. unfold full_sig in h2. unfold sig_set in h2. rewrite Finite_FiniteT_iff in h2. pose (fun a:{x:T|P x} => exist _ (proj1_sig a) (intro_characteristic_sat _ _ (proj2_sig a))) as f. assert (h1:Injective f). red. intros a b h1. destruct a; destruct b. unfold f in h1. apply exist_injective in h1. simpl in h1. apply proj1_sig_injective. simpl. assumption. eapply inj_finite; auto. apply h2. apply h1. intros; apply classic. Qed. (* Wyckoff *) Lemma cart_prod_fin_comm : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), Finite (cart_prod A B) -> Finite (cart_prod B A). intros T U A B h1. apply Finite_ens_type in h1. pose (fun pr:{p:U*T | In (cart_prod B A) p} => exist _ (snd (proj1_sig pr), fst (proj1_sig pr)) (in_cart_prod_comm B A (proj1_sig pr) (proj2_sig pr))) as f. assert (h2:Injective f). red. unfold f. intros pr1 pr2 h2. destruct pr1 as [p1 h3]. destruct pr2 as [p2 h4]. simpl in h2. apply existTexist. apply subsetT_eq_compat. inversion h2 as [h5]. rewrite surjective_pairing. rewrite <- h5. rewrite <- H. apply surjective_pairing. apply FiniteT_sig_Finite. apply (inj_finite _ _ f h1 h2). intro pr. destruct pr as [p h4]. unfold f. left. exists (exist (fun p:U*T => In (cart_prod B A) p) (snd p, fst p) (in_cart_prod_comm A B p h4)). simpl. apply existTexist. apply subsetT_eq_compat. rewrite surjective_pairing. reflexivity. Qed. (*Wyckoff*) Lemma FiniteT_Finite_Union : forall {It T:Type} (F:IndexedFamily It T), (forall i:It, Finite (F i)) -> FiniteT It -> Finite (IndexedUnion F). Proof. intros It T F h1 h2. induction h2. (*Empty*) assert (h3:(IndexedUnion F) = Empty_set _ ). apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h4. inversion h4. contradiction. (* >= *) red. intros; contradiction. rewrite h3. constructor. (*Add*) rename T0 into It. pose (fun (i:It) => F (Some i)) as f. assert (h4:(IndexedUnion F) = Union (F None) (IndexedUnion f)). apply Extensionality_Ensembles. red. split. (*<=*) red. intros x h5. inversion h5 as [i ? h6]. destruct i. (*Some*) right. unfold f. apply indexed_union_intro with i. assumption. (*None*) constructor. assumption. (*>=*) red. intros x h5. pose proof (Union_inv _ _ _ x h5) as h6. case h6 as [h7 | h8]. apply indexed_union_intro with None; assumption. inversion h8. apply indexed_union_intro with (Some i); assumption. assert (h5: forall i:It, Finite (f i)). intro i. apply (h1 (Some i)). pose proof (IHh2 f h5) as h9. assert (h10: Finite (F None)). apply (h1 None). rewrite h4. apply Union_preserves_Finite; assumption. (*Bij*) rename Y into It. rename H into h3. pose (fun (x:X) => F (f x)) as Ff. assert (h4: forall x:X, Finite (Ff x)). intro x. apply (h1 (f x)). pose proof (IHh2 Ff h4) as h5. assert (h6:IndexedUnion Ff = IndexedUnion F). apply Extensionality_Ensembles. red. split. (*<=*) red. intros x h7. inversion h7 as [? ? h8]. unfold Ff in h8. apply indexed_union_intro with (f i); assumption. (*>=*) pose proof (invertible_impl_bijective _ h3) as h6. red. intros t h7. inversion h7 as [i y h8]. unfold bijective in h6. destruct h6 as [h9 h10]. red in h10. pose proof (h10 i) as h11. elim h11. intros x h12. rewrite <- h12 in h8. unfold Ff. apply indexed_union_intro with x; assumption. rewrite <- h6. assumption. Qed. (*Wyckoff*) Lemma bool_finite : FiniteT bool. Proof. pose (fun x:(option (option False)) => match x with | None => false | Some y => match y with | None => true | Some z => (False_rec bool z) end end) as f. apply (bij_finite _ _ f); repeat constructor. apply bijective_impl_invertible. red. split. (*injective*) red. intros a b h1. unfold f in h1. destruct a as [a2 | a3]; destruct b as [b2 | b3]. destruct a2 as [a2a | a2b]; destruct b2 as [b2a | b2b]; try contradiction; try reflexivity. destruct a2 as [a2a | a2b]; try contradiction; try discriminate. destruct b2 as [b2a | b2b]; try contradiction; try discriminate. reflexivity. (*surjective*) red. intro x. unfold f. destruct x. exists (Some None). reflexivity. exists None. reflexivity. Qed. (*Wyckoff*) Lemma bool_eq_true_iff : forall v, v = true <-> v <> false. intro v; destruct v; intuition. Qed. (*Wyckoff*) Lemma bool_eq_false_iff : forall v, v = false <-> v <> true. intro v; destruct v; intuition. Qed. (*Wyckoff*) Inductive bool_lt : bool -> bool -> Prop := | bool_lt_intro : bool_lt false true. (*Wyckoff*) Lemma tso_dec_bool_lt : tso_dec bool_lt. red; intros x y; destruct x, y. left; right; auto. right; constructor. left; left; constructor. left; right; auto. Qed. (*Wyckoff*) Lemma total_strict_order_bool_lt : total_strict_order bool_lt. red; intros x y; destruct x, y. right; left; auto. right; right; constructor. left; constructor. right; left; auto. Qed. (*Wyckoff*) Lemma transitive_bool_lt : transitive _ bool_lt. red; intros x y z h1 h2; destruct x, y, z; try inversion h1; try inversion h2. Qed. (*Wyckoff*) Lemma irrefl_bool_lt : irrefl bool_lt. red; intro b; destruct b; intro h1; inversion h1. Qed. (*Wyckoff*) Lemma sortable_dec_bool_lt : sort_dec bool_lt. constructor. apply tso_dec_bool_lt. apply transitive_bool_lt. apply irrefl_bool_lt. Qed. (*Wyckoff*) Lemma sortable_bool_lt : sortable bool_lt. constructor. apply total_strict_order_bool_lt. apply transitive_bool_lt. apply irrefl_bool_lt. Qed. (*Wyckoff*) Fixpoint bool_le (p q:bool) : Prop := if p then if q then True else False else True. (*Wyckoff*) Lemma false_bool_le : forall q, bool_le false q. intro q; destruct q; auto; simpl; auto. Qed. (*Wyckoff*) Lemma bool_le_true : forall p, bool_le p true. intro p; destruct p; auto; simpl; auto. Qed. (*Wyckoff*) Lemma bool_le_dec : forall v w, {bool_le v w} + {bool_le w v}. intros v w; destruct v, w; simpl. left; auto. right; auto. left; auto. right; auto. Qed. (*Wyckoff*) Lemma total_strict_order_bool_le : total_strict_order bool_le. red. intros x y; destruct x, y; simpl. left; auto. right; auto. left; auto. left; auto. Qed. (*Wyckoff*) Lemma transitive_bool_le : transitive _ bool_le. red; intros x y z; destruct x, y, z; simpl; auto. Qed. (*Wyckoff*) Lemma eq_true_iff : forall b, b = true <-> b <> false. intro b; destruct b; split; intro h1; try intro h2; try discriminate; auto. Qed. (*Wyckoff*) Lemma eq_false_iff : forall b, b = false <-> b <> true. intro b; destruct b; split; intro h1; try intro h2; try discriminate; auto. contradict h1; auto. Qed. (*Schepler*) Lemma finite_dep_choice: forall (A:Type) (B:forall x:A, Type) (R:forall x:A, B x->Prop), FiniteT A -> (forall x:A, exists y:B x, R x y) -> exists f:(forall x:A, B x), forall x:A, R x (f x). Proof. intros. revert B R H0. induction H. intros. exists (fun x:False => False_rect (B x) x). destruct x. intros. pose proof (IHFiniteT (fun x:T => B (Some x)) (fun x:T => R (Some x)) (fun x:T => H0 (Some x))). destruct H1. pose proof (H0 None). destruct H2. exists (fun y:option T => match y return (B y) with | Some y0 => x y0 | None => x0 end). destruct x1. apply H1. assumption. intros. destruct H0. pose proof (IHFiniteT (fun x:X => B (f x)) (fun x:X => R (f x)) (fun x:X => H1 (f x))). destruct H3. pose (f0 := fun y:Y => x (g y)). pose (conv := fun (y:Y) (a:B (f (g y))) => eq_rect (f (g y)) B a y (H2 y)). exists (fun y:Y => conv y (x (g y))). intro. unfold conv; simpl. generalize (H2 x0). pattern x0 at 2 3 6. rewrite <- H2. intro. rewrite <- eq_rect_eq. apply H3. Qed. (*Schepler*) Lemma finite_choice : forall (A B:Type) (R:A->B->Prop), FiniteT A -> (forall x:A, exists y:B, R x y) -> exists f:A->B, forall x:A, R x (f x). Proof. intros. apply finite_dep_choice. assumption. assumption. Qed. (*Schepler*) Lemma finite_sum: forall X Y:Type, FiniteT X -> FiniteT Y -> FiniteT (X+Y). Proof. intros. induction H0. apply bij_finite with _ inl. assumption. pose (g := fun (x:X+False) => match x with | inl x => x | inr f => False_rect X f end). exists g. intro; simpl. reflexivity. destruct y. simpl. reflexivity. destruct f. pose (g := fun (x:option (X+T)) => match x with | Some (inl x) => inl _ x | Some (inr t) => inr _ (Some t) | None => inr _ None end). apply bij_finite with _ g. apply add_finite. assumption. pose (ginv := fun (x:X + option T) => match x with | inl x => Some (inl _ x) | inr (Some t) => Some (inr _ t) | inr None => None end). exists ginv. destruct x as [[x|t]|]; trivial. destruct y as [x|[t|]]; trivial. pose (g := fun (x:X+X0) => match x with | inl x0 => inl _ x0 | inr x0 => inr _ (f x0) end). destruct H1. pose (ginv := fun (x:X+Y) => match x with | inl x0 => inl _ x0 | inr y0 => inr _ (g0 y0) end). apply bij_finite with _ g. assumption. exists ginv. destruct x as [x0|x0]; trivial. simpl. rewrite H1; reflexivity. destruct y as [x|y0]; trivial. simpl. rewrite H2; reflexivity. Qed. (*Schepler*) Lemma finite_prod: forall (X Y:Type), FiniteT X -> FiniteT Y -> FiniteT (X*Y). Proof. intros. induction H0. apply bij_finite with _ (False_rect _). constructor. exists (@snd X False). destruct x. destruct y. destruct f. pose (g := fun (x:X*T + X) => match x with | inl (pair x0 t) => pair x0 (Some t) | inr x0 => pair x0 None end). pose (ginv := fun (x:X * option T) => match x with | (x0, Some t) => inl _ (x0, t) | (x0, None) => inr _ x0 end). apply bij_finite with _ g. apply finite_sum. assumption. assumption. exists ginv. destruct x as [[x0 t]|x0]; trivial. destruct y as [x0 [t|]]; trivial. pose (g := fun (y:X*X0) => match y with | pair x x0 => pair x (f x0) end). destruct H1. pose (ginv := fun (y:X*Y) => let (x,y0) := y in (x, g0 y0)). apply bij_finite with _ g. assumption. exists ginv. destruct x as [x x0]; unfold ginv, g; try rewrite H1; trivial. destruct y as [x y]; unfold ginv, g; try rewrite H2; trivial. Qed. (*Schepler*) Lemma finite_exp: forall X Y:Type, FiniteT X -> FiniteT Y -> FiniteT (X->Y). Proof. intros. induction H. pose (g := fun (x:True) (f:False) => False_rect Y f). pose (ginv := fun (_:False->Y) => I). apply bij_finite with _ g. apply True_finite. exists ginv. destruct x as []. trivial. intro y. apply functional_extensionality. intro; contradiction. pose (g := fun (p:(T->Y)*Y) (x:option T) => let (f,y0) := p in match x with | Some x0 => f x0 | None => y0 end). pose (ginv := fun (f:option T->Y) => (fun x:T => f (Some x), f None)). apply bij_finite with _ g. apply finite_prod. assumption. assumption. exists ginv. destruct x as [f y0]; try extensionality t; try destruct t as [t0|]; trivial. intro. extensionality t; destruct t as [t0|]; trivial. destruct H1. pose (g0 := fun (h:X->Y) (y:Y0) => h (g y)). apply bij_finite with _ g0. assumption. pose (g0inv := fun (h:Y0->Y) (x:X) => h (f x)). exists g0inv. intro. extensionality x0; unfold g0; unfold g0inv; simpl. rewrite H1; reflexivity. intro. extensionality y0; unfold g0; unfold g0inv; simpl. rewrite H2; reflexivity. Qed. (*Wyckoff*) (*Uses excluded middle*) Lemma power_set_finitet : forall {T:Type} (A:Ensemble T), Finite A -> FiniteT {S:Ensemble T | Included S A}. intros T A h1. assert (h2: (forall (S:Ensemble T) (x:T), In A x -> {In S x} + {~In S x})). intros; apply classic_dec. pose (fun S:{S':(Ensemble T) | Included S' A} => (fun x:{t:T | In A t} => if (h2 (proj1_sig S) (proj1_sig x) (proj2_sig x)) then true else false)) as f. pose proof (Finite_ens_type _ h1) as h3. pose proof bool_finite as h4. pose proof (finite_exp _ _ h3 h4) as h5. apply (inj_finite _ _ f). assumption. red. intros S1 S2 h6. apply proj1_sig_injective. apply Extensionality_Ensembles. red. split. (* -> *) red. intros x h7. pose proof (proj2_sig S1) as h11. simpl in h11. assert (h12:In A x). auto with sets. pose (exist _ _ h12) as x'. assert (h13: f S1 x' = f S2 x'). rewrite h6. reflexivity. unfold f in h13. pose proof (conditional_correspondence _ _ _ _ h13) as h14. simpl in h14. tauto. (* <- *) red. intros x h7. pose proof (proj2_sig S2) as h11. simpl in h11. assert (h12:In A x). auto with sets. pose (exist _ _ h12) as x'. assert (h13: f S1 x' = f S2 x'). rewrite h6. reflexivity. unfold f in h13. pose proof (conditional_correspondence _ _ _ _ h13) as h14. simpl in h14. tauto. intros; tauto. Qed. (*Wyckoff*) Lemma power_set_finite : forall {T:Type} (A:Ensemble T), Finite A -> Finite (power_set A). intros T A h1. pose proof (power_set_finitet A h1) as h2. unfold power_set. rewrite <- finite_sig_sat_compat. assumption. Qed. (*Use is made of Full_set rather than T to exploit previous theorems easier.*) (*Wyckoff*) Lemma bij_seg_finite_full_iff : forall {T:Type} (f:sig_set (seg (card_fun1 (Full_set T))) -> sig_set (Full_set _)), bijective f -> forall (P:T->Prop), (Finite [x:T | P x] <-> Finite [n : nat | exists pf : In (seg (card_fun1 (Full_set T))) n, P (proj1_sig (f (exist _ _ pf)))]). intros T f h1 P. match goal with |- _ <-> Finite ?S' => pose S' as S end. fold S. assert (h2:forall n:nat, In S n -> In (seg (card_fun1 (Full_set T))) n). intros n h2. destruct h2 as [h2]. destruct h2; auto. assert (h3:forall (n:nat) (pf:In S n), P (proj1_sig (f (exist (In (seg (card_fun1 (Full_set T)))) n (h2 n pf))))). intros n h3. destruct h3 as [h3]. destruct h3 as [h3 h4]. gterm0. redterm0 c. redterm0 c0. redterm0 c1. clear c1 c0 c. pose proof (proof_irrelevance _ h3 c2) as h5. fold c2. rewrite <- h5 at 1. assumption. pose (fun a:sig_set S => exist _ _ (intro_characteristic_sat P _ (h3 _ (proj2_sig a)))) as ff. apply bijective_impl_invertible in h1. pose proof (proj2_sig (function_inverse f h1)) as h7. destruct h7 as [h7 h8]. pose (proj1_sig (function_inverse _ h1)) as f'. assert (h6:forall a:sig_set [x:T | P x], P (proj1_sig (f (exist _ _ (proj2_sig (f' (exist _ _ (Full_intro _ (proj1_sig a))))))))). intro a. destruct a as [a h0]. unfold f'. simpl. rewrite <- unfold_sig. rewrite h8. simpl. destruct h0. assumption. assert (h9:forall a:sig_set [x:T | P x], In S (proj1_sig (f' (exist _ _ (Full_intro _ (proj1_sig a)))))). intro a. constructor. exists (proj2_sig _). apply h6. pose (fun a:sig_set [x:T | P x] => (exist _ _ (h9 a))) as ff'. assert (h4:Im (full_sig S) ff = full_sig [x:T | P x]). apply Extensionality_Ensembles. red; split; red; intros x h5. unfold full_sig. constructor. apply Im_intro with (ff' x). constructor. unfold ff, ff'. simpl. apply proj1_sig_injective. simpl. unfold f'. gtermr. redterm0 y. redterm0 c. redterm0 c0. clear c0 c y. fold c1. pose proof (proof_irrelevance _ c1 (proj2_sig _)) as h10. rewrite h10 at 1. rewrite <- unfold_sig. rewrite h8. simpl. reflexivity. assert (h5:Im (full_sig [x:T | P x]) ff' = full_sig S). apply Extensionality_Ensembles. red; split; red; intros x h5. unfold full_sig. constructor. apply Im_intro with (ff x). constructor. unfold ff, ff'. simpl. apply proj1_sig_injective. simpl. unfold f'. gtermr. redterm0 y. redterm0 c. redterm0 c0. clear c0 c y. pose proof (proof_irrelevance _ c1 (proj2_sig _)) as h10. fold c1. rewrite h10 at 1. rewrite <- unfold_sig. rewrite h7. simpl. reflexivity. pose proof (invertible_impl_bijective _ h1) as h13. destruct h13 as [h13 h14]. assert (h11:Injective ff). red. intros x y h12. destruct x as [x h15], y as [y h16]. apply proj1_sig_injective. simpl. unfold ff in h12. simpl in h12. apply exist_injective in h12. apply proj1_sig_injective in h12. apply h13 in h12. apply exist_injective in h12. assumption. pose proof (invertible_impl_inv_invertible _ h1) as h16. apply invertible_impl_bijective in h16. destruct h16 as [h16 h17]. assert (h12:Injective ff'). red. intros x y h12. destruct x as [x h15], y as [y h18]. apply proj1_sig_injective. simpl. unfold ff' in h12. simpl in h12. apply exist_injective in h12. apply proj1_sig_injective in h12. apply h16 in h12. apply exist_injective in h12. assumption. split. intro h10. apply FiniteT_sig_Finite. apply Finite_ens_type in h10. apply (inj_finite _ _ ff h10 h11). intros. apply classic. intro h10. apply FiniteT_sig_Finite. apply Finite_ens_type in h10. apply (inj_finite _ _ ff' h10 h12). intros; apply classic. Qed. Lemma finite_constructive_definite_description_full : forall (T:Type), FiniteT T -> forall f:sig_set (seg (card_fun1 (Full_set T))) -> sig_set (Full_set T), bijective f -> forall (P:T->Prop), (exists x:T, P x) -> {x:T | P x}. intros T h1 f h2 P h3. pose proof (ex_inh _ h3) as h4. pose proof (bij_seg_inh_full_iff f h2 P) as h5. rewrite h5 in h4. pose [n:nat | exists pf:In (seg (card_fun1 (Full_set T))) n, P (proj1_sig (f (exist _ _ pf)))] as S. pose proof (FiniteT_Finite _ h1 [x:T | P x]) as h6. rewrite bij_seg_finite_full_iff in h6. pose (max_set_nat _ h6 h4). pose proof (max_set_nat_compat _ h6 h4) as h10. simpl in h10. destruct h10 as [h10 h11]. destruct h10 as [h10]. pose proof (destruct_exists_pf_set _ _ _ h10) as h12. simpl in h12. refine (exist _ _ h12). assumption. Qed. Section SubtractT. Definition subtractT {U} T (f:option U->T) : Type := {x:T | x <> f None}. Lemma FiniteT_subtract : forall {T U} (pft:FiniteT T) (f:option U->T), FiniteT (subtractT T f). unfold subtractT; intros; apply finite_subtype; auto. intro x. destruct (finite_eq_dec _ pft x (f None)); subst. right; intro h1; contradict h1; auto. left; auto. Qed. Definition de_option {T} (o:option T) : o <> None -> T := match o with | None => fun pf => False_rect _ (pf (eq_refl None)) | Some o' => fun _ => o' end. Lemma inj_de_option : forall {T} (o o':option T) (pfo: o <> None) (pfo':o' <> None), de_option o pfo = de_option o' pfo' -> o = o'. intros T o o'. destruct o as [x|], o' as [x'|]; simpl; intros; subst; auto. contradict pfo'. contradict pfo. Qed. Lemma subtract_inv_neq_none : forall {T U} (f:T->option U) (pfinv:invertible f) (x:subtractT T (proj1_sig (function_inverse f pfinv))), f (proj1_sig x) <> None. intros T U f h1 x. destruct x as [x h2]; simpl. pose proof (proj2_sig (function_inverse f h1)) as h3; simpl in h3. destruct h3 as [h3 h4]. specialize (h3 x). contradict h2; rewrite <- h3, h2; auto. Qed. Lemma f_some_neq_none : forall {T U} (f:option T -> U), injective f -> forall x, f (Some x) <> f None. intros T U f h1 x h2; apply h1 in h2; discriminate. Qed. Lemma f_some_neq_none' : forall {T U} (f:option T -> U), invertible f -> forall x, f (Some x) <> f None. intros T U f h1 x. apply invertible_impl_bijective in h1. destruct h1. apply f_some_neq_none; auto. Qed. Definition f_from_option1 {T U} (f:option T-> U) (pfinv:invertible f) : T -> subtractT U f := fun x => exist _ (f (Some x)) (f_some_neq_none' f pfinv x). Definition f_from_option2 {T U} (f:T->option U) (pfinv:invertible f) : subtractT T (proj1_sig (function_inverse f pfinv)) -> U := fun x => de_option (f (proj1_sig x)) (subtract_inv_neq_none f pfinv x). Lemma inv_f_from_option1 : forall {T U} (f:option T-> U) (pfinv:invertible f), invertible (f_from_option1 f pfinv). intros T U f h1. pose proof (invertible_impl_bijective _ h1) as h2. apply bijective_impl_invertible. unfold f_from_option1. destruct h2 as [h2 h3]. red. split. red in h2; red. intros x y h4. apply exist_injective in h4. apply h2 in h4. inversion h4; auto. red in h3; red. intro y. destruct y as [y h4]. specialize (h3 y). destruct h3 as [x h3]; subst. destruct x as [x|]. exists x. apply proj1_sig_injective; simpl; auto. contradict h4; auto. Qed. Lemma inv_f_from_option2 : forall {T U} (f:T-> option U) (pfinv:invertible f), invertible (f_from_option2 f pfinv). intros T U f h1. pose proof (invertible_impl_bijective _ h1) as h2. apply bijective_impl_invertible. unfold f_from_option2. destruct h2 as [h2 h3]. red. split. red in h2; red. intros x y h4. apply inj_de_option in h4. apply h2 in h4. apply proj1_sig_injective in h4; auto. pose proof h3 as h3'. red in h3; red. intro y. specialize (h3 (Some y)). destruct h3 as [x h3]; subst. assert (h4:x <> (proj1_sig (function_inverse f h1)) None). intro; subst. pose proof (proj2_sig (function_inverse f h1)) as h4. simpl in h4. destruct h4 as [h4 h5]. rewrite h5 in h3; discriminate. exists (exist _ x h4). simpl. gen0. simpl. rewrite h3. intro. simpl; auto. Qed. Lemma equi_subtract : forall {T U} (f:option T -> U), invertible f -> equi T (subtractT U f). intros T U f h1. pose proof (inv_f_from_option1 f h1) as h2. econstructor. apply h2. Qed. Definition f_option_subtractT {T U} (f:option T->U) : option (subtractT U f) -> U := fun x => match x with | None => f None | Some x' => proj1_sig x' end. Lemma inv_f_option_subtractT : forall {T U} (f:option T->U), invertible f -> invertible (f_option_subtractT f). intros T U f hi. pose proof (invertible_impl_bijective _ hi) as h1. destruct h1 as [hinj hsurj]. apply bijective_impl_invertible. unfold f_option_subtractT. split. intros x y h3. destruct x as [x|], y as [y|]. apply proj1_sig_injective in h3; subst; auto. pose proof (proj2_sig x) as h9; simpl in h9. rewrite h3 in h9; contradict h9; auto. pose proof (proj2_sig y) as h10; simpl in h10. rewrite h3 in h10 at 1. contradict h10; auto. reflexivity. red. intro y. red in hsurj. specialize (hsurj y). destruct hsurj; subst. destruct x as [x|]. pose proof (f_some_neq_none f hinj x) as h9. exists (Some (exist (fun c => c <> f None) (f (Some x)) h9)); simpl; auto. exists None; auto. Qed. (*Think of this, as the bijective opposite of [map option] -- in the most "natural" way -- the nicest way to turn [f] into the proper signature!*) Definition f_from_option {T U} (f:option T->option U) (pf:invertible f) : T -> U := match (none_dec (f None)) with | left pfe => fun x => let pff := f_some_neq_none' f pf x in match (none_dec (f (Some x))) with | left pfe' => False_rect U (pff (eq_trans pfe' (eq_sym pfe))) | right pfn' => neq_none_inv _ pfn' end | right pfn => fun x => let pff := f_some_neq_none' f pf x in match (none_dec (f (Some x))) with | left pfe' => neq_none_inv _ pfn | right pfn' => neq_none_inv _ pfn' end end. Lemma f_from_option_none : forall {T U} (f:option T -> option U) (pf:invertible f) (pfe:f None = None) x, let pfn := neq_eq_trans (f_some_neq_none' f pf x) pfe in f_from_option f pf x = de_option (f (Some x)) pfn. unfold f_from_option; intros; do 2 lr_match_goal'; unfold neq_none_inv, de_option; try contradiction. gterml. redterm0 x0; contradiction. f_equal; apply proof_irrelevance. Qed. Lemma f_from_option_some_eq : forall {T U} (f:option T -> option U) (pf:invertible f) (pfn:f None <> None) x, f (Some x) = None -> f_from_option f pf x = de_option (f None) pfn. unfold f_from_option, de_option; intros; do 2 lr_match_goal'; unfold neq_none_inv; try contradiction; f_equal; apply proof_irrelevance. Qed. Lemma f_from_option_some_neq : forall {T U} (f:option T -> option U) (pf:invertible f), f None <> None -> forall x (pfn:f (Some x) <> None), f_from_option f pf x = de_option (f (Some x)) pfn. unfold f_from_option, de_option; intros; do 2 lr_match_goal'; unfold neq_none_inv; try contradiction; f_equal; apply proof_irrelevance. Qed. Lemma inv_f_from_option : forall {T U} (f:option T -> option U) (pf:invertible f), invertible (f_from_option f pf). unfold f_from_option; intros T U f h1. pose proof (invertible_impl_bijective _ h1) as h2. apply bijective_impl_invertible. destruct h2 as [h2 h3]; split. red in h2; red. intros x y h4. lr_match' h4. lr_match' h4. terml h4. redterm0 x0. contradiction. lr_match' h4. termr h4. redterm0 y0. contradiction. apply inj_neq_none_inv in h4. apply h2 in h4. inversion h4; auto. lr_match' h4. lr_match' h4. rewrite <- hl in hl0. apply h2 in hl0. inversion hl0; auto. apply inj_neq_none_inv in h4. apply h2 in h4; discriminate. lr_match' h4. apply inj_neq_none_inv in h4. apply h2 in h4; discriminate. apply inj_neq_none_inv in h4. apply h2 in h4. inversion h4; auto. red in h3; red. intro y. specialize (h3 (Some y)). destruct h3 as [x h3]. lr_match_goal'. destruct x as [x|]. exists x. lr_match_goal'. rewrite hl0 in h3; discriminate. unfold neq_none_inv. destruct (f (Some x)). inversion h3; auto. contradict hr; auto. rewrite hl in h3; discriminate. destruct x as [x|]. exists x. lr_match_goal'. unfold neq_none_inv. rewrite h3 in hl; discriminate. unfold neq_none_inv. pose proof (some_neq_none y) as h4. pose proof h4 as h4'. rewrite <- h3 in h4'. pose proof (subsetT_eq_compat _ (fun c => c <> None) _ _ hr0 h4 h3) as h5. dependent rewrite h5; auto. unfold neq_none_inv. pose proof hr as h4. rewrite h3 in h4. pose proof (subsetT_eq_compat _ (fun c => c <> None) _ _ hr h4 h3) as h5. dependent rewrite h5. destruct (none_dec (proj1_sig (function_inverse _ h1) None)) as [h6 | h6]. pose proof (proj2_sig (function_inverse _ h1)) as h7; simpl in h7; destruct h7 as [h7 h8]. specialize (h8 None). rewrite h6 in h8. contradiction. exists (neq_none_inv _ h6). lr_match_goal'; auto. pose proof hr0 as h7. rewrite some_neq_none_inv in h7. pose proof (proj2_sig (function_inverse _ h1)) as h8; simpl in h8; destruct h8 as [h8 h9]. specialize (h9 None). contradiction. Qed. End SubtractT. Lemma equi_option : forall U V, equi U V -> equi (option U) (option V). intros U V h1. destruct h1. econstructor. apply invertible_option_map. apply H. Qed. Lemma inj_option : forall T U, equi (option T) (option U) -> equi T U. intros T U h1. inversion h1. apply equi_intro with (f_from_option f H). apply inv_f_from_option. Qed. Lemma equi_sig_false : forall T, equi False {x:T | False}. intro T. pose (fun x:False => exist (fun c:T => False) (False_rect _ x) x) as f. assert (h1:invertible f). apply bijective_impl_invertible. red; split; red. intros; contradiction. intro y; destruct y; contradiction. pose proof (equi_intro _ _ _ h1); auto. Qed. Lemma false_option : forall T, ~ (eq False (option T) (A:=Type)). intros T h1. pose proof (@None T). rewrite <- h1 in X; auto. Qed. Lemma finite_segT : forall n, FiniteT (n @). intro n; induction n as [|n ih]. apply (bij_finite _ _ (fun (pff:False) => False_rect (0 @) pff)). apply empty_finite. apply bijective_impl_invertible; red; split; red; [intros; contradiction|intro y; destruct y; omega]. apply (bij_finite _ _ (fun x:option (n @) => match x with | Some e => let (a, pfa) := e in exist (fun c => c < S n) (S a) (lt_n_S _ _ pfa) | None => exist (fun c => c < S n) 0 (lt_0_Sn _) end)). apply add_finite; auto. apply bijective_impl_invertible; red; split; red. intros x y h1; destruct x, y. try destruct s, s0. apply exist_injective in h1. apply S_inj in h1; subst. f_equal; apply proj1_sig_injective; simpl; auto. destruct s. apply exist_injective in h1. discriminate. destruct s. apply exist_injective in h1. discriminate. reflexivity. intro y. destruct y as [y hy]. destruct y as [|y]. exists None. f_equal; apply proof_irrelevance. exists (Some (exist (fun c => c < n) y (lt_S_n _ _ hy))). f_equal; apply proof_irrelevance. Qed. End FiniteTypes. (*Whole section copied and pasted from file of same name in Schepler's Zorn's Lemma.*) Section InfiniteTypes. Open Scope nat. (*Schepler*) Lemma finite_nat_initial_segment: forall n:nat, FiniteT { m:nat | m < n }. Proof. intros. apply Finite_ens_type. rewrite <- characteristic_function_to_ensemble_is_identity. induction n. assert ([x:nat | x < 0] = (Empty_set _)). apply Extensionality_Ensembles; split; auto with sets. red; intros. destruct H. contradict H. auto with arith. rewrite H; constructor. assert ([x:nat | S x <= S n] = Add [x:nat | x < n] n). apply Extensionality_Ensembles; split. red; intros. destruct H. assert (x <= n); auto with arith. apply le_lt_or_eq in H0. case H0. left; constructor; trivial. right; auto with sets. red; intros. case H. intros. destruct H0; constructor. auto with arith. intros. destruct H0. constructor. auto with arith. assert (h1:forall x, S x <= S n <-> x < S n). intros; tauto; trivial with arith. rewrite (sat_iff _ _ h1) in H. rewrite H; constructor; trivial. red; intro. destruct H0. contradict H0. auto with arith. Qed. (*Schepler*) Lemma nat_infinite: ~ FiniteT nat. Proof. red; intro. assert (surjective S). apply finite_inj_surj; trivial. red; intros. injection H0; trivial. destruct (H0 0). discriminate H1. Qed. (*Wyckoff*) Lemma Z_infinite : ~FiniteT Z. intro h1. assert (h2:forall y : Z, (exists x : nat, Z.of_nat x = y) \/ ~ (exists x : nat, Z.of_nat x = y)). intros; apply classic. pose proof (inj_finite _ _ Z.of_nat h1 Nat2Z.inj h2) as h3. pose proof nat_infinite. contradiction. Qed. (*Wyckoff*) Lemma Q_infinite : ~FiniteT QArith_base.Q. intro h1. assert (h2: Injective QArith_base.inject_Z). red. intros ? ? h3. assert (h3': QArith_base.Qeq(QArith_base.inject_Z x1) (QArith_base.inject_Z x2)). rewrite h3. constructor. rewrite QArith_base.inject_Z_injective in h3'. assumption. assert (h3: forall y : QArith_base.Q, (exists x : Z, QArith_base.inject_Z x = y) \/ ~ (exists x : Z, QArith_base.inject_Z x = y)). intros; apply classic. pose proof (inj_finite _ _ QArith_base.inject_Z h1 h2 h3) as h4. pose proof Z_infinite. contradiction. Qed. (*Schepler*) Lemma infinite_nat_inj: forall X:Type, ~ FiniteT X -> exists f:nat->X, Injective f. Proof. intros. assert (inhabited (forall S:Ensemble X, Finite S -> { x:X | ~ In S x})). pose proof (choice (fun (x:{S:Ensemble X | Finite S}) (y:X) => ~ In (proj1_sig x) y)). simpl in H0. match type of H0 with | ?A -> ?B => assert B end. apply H0. intros. apply NNPP. red; intro. pose proof (not_ex_not_all _ _ H1); clear H1. destruct x. assert (x = (Full_set _)). apply Extensionality_Ensembles; red; split; auto with sets. intro x0; constructor. symmetry in H1; destruct H1. contradiction H. clear H2. apply bij_finite with (f:=@proj1_sig _ (fun x:X => In (Full_set _) x)). apply Finite_ens_type; assumption. exists (fun x:X => exist _ x (Full_intro _ x)). destruct x; simpl. generalize (Full_intro X x). intro i0; destruct (proof_irrelevance _ i i0); trivial. trivial. clear H0. destruct H1. exists. intros. exists (x (exist _ S H1)). exact (H0 (exist _ S H1)). destruct H0. assert (forall (n:nat) (g:forall m:nat, m X), { x:X | forall (m:nat) (Hlt:m x }). intros. assert (Finite (fun x:X => exists m:nat, exists Hlt:m g (proj1_sig x) (proj2_sig x)). match goal with |- Finite ?S => assert (S = Im (Full_set _) h) end. apply Extensionality_Ensembles; red; split; red; intros. destruct H0. destruct H0. exists (exist (fun m:nat => m < n) x0 x1). constructor. unfold h; simpl. symmetry; assumption. destruct H0. destruct x. unfold h in H1; simpl in H1. exists x; exists l; symmetry; assumption. rewrite H0; apply FiniteT_img. apply finite_nat_initial_segment. intros; apply classic. destruct (X0 _ H0). unfold In in n0. exists x. intros; red; intro. contradiction n0; exists m; exists Hlt; exact H1. pose (f := Fix lt_wf (fun n:nat => X) (fun (n:nat) (g:forall m:nat, mX) => proj1_sig (X1 n g))). simpl in f. assert (forall n m:nat, m f m <> f n). pose proof (Fix_eq lt_wf (fun n:nat => X) (fun (n:nat) (g:forall m:nat, mX) => proj1_sig (X1 n g))). fold f in H0. simpl in H0. match type of H0 with | ?A -> ?B => assert (B) end. apply H0. intros. assert (f0 = g). Require Import FunctionalExtensionality. extensionality y; extensionality p; apply H1. destruct H2; trivial. intros. pose proof (H1 n). destruct X1 in H3. simpl in H3. destruct H3. auto. exists f. red; intros m n ?. destruct (lt_eq_lt_dec m n) as [[Hlt|Heq]|Hlt]; trivial. contradiction (H0 n m). contradiction (H0 m n). symmetry; assumption. Qed. (*Schepler*) Lemma inf_comp_fin : forall {T:Type}, ~FiniteT T -> forall U:Ensemble T, Finite (Comp U) -> ~ Finite U. intros T h0 U h1. case (classic (Finite U)); trivial. intro h2. pose proof (Union_preserves_Finite T U (Comp U) h2 h1) as h3. rewrite excl_middle_full in h3. pose proof (Finite_ens_type (Full_set T) h3) as h4. pose (fun (a:{x:T | In (Full_set T) x}) => proj1_sig a) as f. assert (h5:invertible f). apply bijective_impl_invertible. red. split. (*left*) red. intros x1 x2 h6. unfold f in h6. apply proj1_sig_injective; trivial. (*right*) red. intro y. exists (exist _ y (Full_intro _ y)). unfold f. simpl. reflexivity. pose proof (bij_finite _ _ f h4 h5). contradiction. Qed. End InfiniteTypes. Section CountableT. (*Schepler*) Inductive CountableT (X:Type) : Prop := | intro_nat_injection: forall f:X->nat, Injective f -> CountableT X. (*Wyckoff*) Lemma countable_nat : CountableT nat. apply intro_nat_injection with id. apply inj_id. Qed. (*Schepler*) Lemma CountableT_is_FiniteT_or_countably_infinite: forall X:Type, CountableT X -> {FiniteT X} + {exists f:X->nat, bijective f}. Proof. intros. apply exclusive_dec. red; intro. destruct H0 as [? [f ?]]. contradiction nat_infinite. apply bij_finite with _ f; trivial. apply bijective_impl_invertible; trivial. case (classic (FiniteT X)). left; trivial. right. apply infinite_nat_inj in H0. destruct H. destruct H0 as [g]. apply CSB with f g; trivial. Qed. (*Wyckoff*) Lemma inf_countable_index : forall T:Type, CountableT T -> ~FiniteT T -> exists f:nat->T, bijective f. intros T h1 h2. pose proof (CountableT_is_FiniteT_or_countably_infinite T h1) as h3. destruct h3 as [h3 | h3]; try contradiction. destruct h3 as [f h4]. apply bijective_impl_invertible in h4. pose (proj1_sig (function_inverse f h4)) as f'. exists f'. pose proof (invertible_impl_inv_invertible f h4) as h5. apply invertible_impl_bijective in h5. assumption. Qed. (*Schepler*) Lemma countable_nat_product: CountableT (nat*nat). Proof. pose (sum_1_to_n := fix sum_1_to_n n:nat := match n with | O => O | S m => (sum_1_to_n m) + n end). exists (fun p:nat*nat => let (m,n):=p in (sum_1_to_n (m+n)) + n). assert (forall m n:nat, m sum_1_to_n m + m < sum_1_to_n n). intros. induction H. simpl. auto with arith. apply lt_trans with (sum_1_to_n m0). assumption. simpl. assert (0 < S m0); auto with arith. assert (sum_1_to_n m0 + 0 < sum_1_to_n m0 + S m0); auto with arith. assert (sum_1_to_n m0 + 0 = sum_1_to_n m0); auto with arith. rewrite H2 in H1; assumption. red; intros. destruct x1 as [x1 y1]. destruct x2 as [x2 y2]. Require Import Compare_dec. case (lt_eq_lt_dec (x1+y1) (x2+y2)); intro. case s; intro. assert (sum_1_to_n (x1+y1) + y1 < sum_1_to_n (x2+y2) + y2). apply le_lt_trans with (sum_1_to_n (x1+y1) + (x1+y1)). assert (sum_1_to_n (x1+y1) + (x1+y1) = (sum_1_to_n (x1+y1) + y1) + x1). Require Import ArithRing. ring. auto with arith. apply lt_le_trans with (sum_1_to_n (x2+y2)). apply H; trivial. auto with arith. rewrite H0 in H1. contradict H1. auto with arith. assert (y1=y2). rewrite e in H0. Require Import Arith. apply plus_reg_l in H0. assumption. f_equal; trivial. rewrite H1 in e. rewrite plus_comm in e. rewrite (plus_comm x2 y2) in e. apply plus_reg_l in e. assumption. assert (sum_1_to_n (x2+y2) + y2 < sum_1_to_n (x1+y1) + y1). apply le_lt_trans with (sum_1_to_n (x2+y2) + (x2+y2)). auto with arith. apply lt_le_trans with (sum_1_to_n (x1+y1)); auto with arith. rewrite H0 in H1. contradict H1. auto with arith. Qed. (*Schepler*) Lemma countable_sum: forall X Y:Type, CountableT X -> CountableT Y -> CountableT (X+Y). Proof. intros. destruct H as [f]. destruct H0 as [g]. destruct countable_nat_product as [h]. exists (fun s:X+Y => match s with | inl x => h (0, f x) | inr y => h (1, g y) end). red; intros s1 s2 ?. destruct s1 as [x1|y1]; destruct s2 as [x2|y2]; apply H1 in H2; try discriminate H2; intros; f_equal; (apply H || apply H0); injection H2; trivial. Qed. (*Schepler*) Lemma countable_product: forall X Y:Type, CountableT X -> CountableT Y -> CountableT (X*Y). Proof. intros. destruct H as [f]. destruct H0 as [g]. pose (fg := fun (p:X*Y) => let (x,y):=p in (f x, g y)). destruct countable_nat_product as [h]. exists (fun p:X*Y => h (fg p)). red; intros. apply H1 in H2. destruct x1 as [x1 y1]. destruct x2 as [x2 y2]. unfold fg in H2. injection H2; intros. apply H0 in H3. apply H in H4. f_equal; trivial. Qed. Require Import FunctionalExtensionality. (*Schepler*) Lemma countable_exp: forall X Y:Type, FiniteT X -> CountableT Y -> CountableT (X->Y). Proof. intros. induction H. exists (fun _ => 0). red; intros. extensionality f. destruct f. destruct (countable_product (T->Y) Y); trivial. exists (fun (g:option T->Y) => f (fun x:T => g (Some x), g None)). red; intros g1 g2 ?. apply H1 in H2. extensionality o. destruct o. injection H2; intros. pose proof (equal_f H4). simpl in H5. apply H5. injection H2; trivial. destruct H1. destruct IHFiniteT. exists (fun (h:Y0->Y) => f0 (fun x:X => h (f x))). red; intros h1 h2 ?. apply H3 in H4. pose proof (equal_f H4). simpl in H5. extensionality y. rewrite <- (H2 y). apply H5. Qed. (*Schepler*) Definition Countable {X:Type} (S:Ensemble X) : Prop := CountableT {x:X | In S x}. (*Schepler*) Lemma inj_countable: forall {X Y:Type} (f:X->Y), CountableT Y -> Injective f -> CountableT X. Proof. intros. destruct H as [g]. exists (fun x:X => g (f x)). red; intros; auto. Qed. (*Schepler*) Lemma surj_countable: forall {X Y:Type} (f:X->Y), CountableT X -> surjective f -> CountableT Y. Proof. intros. Require Import ClassicalChoice. pose proof (choice (fun (y:Y) (x:X) => f x = y)). destruct H1 as [finv]. exact H0. apply inj_countable with finv. assumption. red; intros. congruence. Qed. (*Schepler*) Lemma countable_downward_closed: forall {X:Type} (S T:Ensemble X), Countable T -> Included S T -> Countable S. Proof. intros. destruct H. exists (fun x:{x:X | In S x} => match x with | exist x0 i => f (exist _ x0 (H0 _ i)) end). red; intros. destruct x1 as [x1]. destruct x2 as [x2]. apply H in H1. injection H1; intros. destruct H2. destruct (proof_irrelevance _ i i0). trivial. Qed. (*Schepler*) Lemma countable_img: forall {X Y:Type} (f:X->Y) (S:Ensemble X), Countable S -> Countable (Im S f). Proof. intros. assert (forall x:X, In S x -> In (Im S f) (f x)). auto with sets. pose (fS := fun x:{x:X | In S x} => match x return {y:Y | In (Im S f) y} with | exist x0 i => exist _ (f x0) (H0 x0 i) end). apply surj_countable with fS; trivial. red; intros. destruct y. destruct i. exists (exist _ x i). simpl. generalize (H0 x i); intro. generalize (Im_intro X Y S f x i y e); intro. destruct e. destruct (proof_irrelevance _ i0 i1). trivial. Qed. (*slightly modified from Schepler*) Lemma countable_type_ensemble: forall {X:Type} (S:Ensemble X), CountableT X -> Countable S. Proof. intros. red. apply inj_countable with (@proj1_sig _ (fun x:X => In S x)). assumption. red. intros a1 a2 h1. destruct a1. destruct a2. apply existTexist. apply subsetT_eq_compat. simpl in h1. assumption. Qed. (*Wyckoff*) Lemma inf_countable_sig_index : forall T:Type, CountableT T -> ~FiniteT T -> forall S:Ensemble nat, ~Finite S -> exists f:sig_set S->T, bijective f. intros T h1 h2 S h3. pose proof countable_nat as h4. pose proof (countable_type_ensemble S h4) as h5. pose proof (CountableT_is_FiniteT_or_countably_infinite _ h5) as h6. destruct h6 as [h6 | h6]. apply FiniteT_sig_Finite in h6. contradiction. destruct h6 as [f h6]. pose proof (inf_countable_index _ h1 h2) as h7. destruct h7 as [g h7]. exists (compose g f). apply bij_compose; auto. Qed. Require Import Infinite_sets. (*Wyckoff*) Lemma inf_countable_index_iff_inf_countable : forall (T:Type), (exists S:Ensemble nat, ~Finite S /\ exists f:sig_set S->T, bijective f) <-> CountableT T /\ ~FiniteT T. intro T. split. intro h1. destruct h1 as [S h1]. destruct h1 as [h1 h2]. destruct h2 as [f h2]. split. eapply surj_countable. assert (h3:CountableT (sig_set S)). apply countable_type_ensemble. apply countable_nat. apply h3. red in h2. destruct h2 as [h2a h2b]. apply h2b. intro h3. rewrite finite_full_sig_iff in h1. assert (h4: Finite (Im (full_sig S) f)). eapply Finite_downward_closed. rewrite <- Finite_FiniteT_iff in h3. apply h3. red; intros; constructor. pose proof (Pigeonhole_bis _ _ _ f h1 h4) as h5. contradict h5. destruct h2; auto. intro h1. destruct h1 as [h1 h2]. exists (Full_set nat). split. rewrite Finite_FiniteT_iff. apply nat_infinite. pose proof (inf_countable_index _ h1 h2) as h3. destruct h3 as [f h3]. pose (fun n:sig_set (Full_set nat) => (f (proj1_sig n))) as f'. exists f'. destruct h3 as [h3 h4]. red. split. red; intros x y h5. destruct x as [x h6], y as [y h7]. unfold f' in h5. simpl in h5. apply proj1_sig_injective; simpl. apply h3; auto. red in h4. red. intro y. specialize (h4 y). destruct h4 as [n h4]. exists (exist _ _ (Full_intro _ n)). unfold f'. simpl. assumption. Qed. (*Schepler*) Lemma FiniteT_impl_CountableT: forall X:Type, FiniteT X -> CountableT X. Proof. intros. induction H. exists (False_rect nat). red; intros. destruct x1. destruct IHFiniteT. exists (fun x:option T => match x with | Some x0 => S (f x0) | None => 0 end). red; intros. destruct x1; destruct x2; try (injection H1 || discriminate H1); trivial. intro. apply H0 in H2. destruct H2; trivial. destruct IHFiniteT as [g]. destruct H0 as [finv]. exists (fun y:Y => g (finv y)). red; intros y1 y2 ?. apply H1 in H3. congruence. Qed. (*Schepler*) Lemma Finite_impl_Countable: forall {X:Type} (S:Ensemble X), Finite S -> Countable S. Proof. intros. apply FiniteT_impl_CountableT. apply Finite_ens_type; trivial. Qed. (* Schepler *) Lemma countable_union: forall {X A:Type} (F:IndexedFamily A X), CountableT A -> (forall a:A, Countable (F a)) -> Countable (IndexedUnion F). Proof. intros. destruct (choice_on_dependent_type (fun (a:A) (f:{x:X | In (F a) x} -> nat) => Injective f)) as [choice_fun_inj]. intro. destruct (H0 a). exists f; trivial. destruct (choice (fun (x:{x:X | In (IndexedUnion F) x}) (a:A) => In (F a) (proj1_sig x))) as [choice_fun_a]. destruct x as [x [a]]. exists a. assumption. destruct countable_nat_product as [g]. destruct H as [h]. exists (fun x:{x:X | In (IndexedUnion F) x} => g (h (choice_fun_a x), choice_fun_inj (choice_fun_a x) (exist _ (proj1_sig x) (H2 x)))). red; intros. apply H3 in H4. injection H4; intros. apply H in H6. revert H5. generalize (H2 x1). generalize (H2 x2). rewrite H6. intros. apply H1 in H5. injection H5. apply proj1_sig_injective. Qed. (*Wyckoff*) Lemma countable_union2 : forall (T:Type) (S1 S2:Ensemble T), Countable S1 -> Countable S2 -> Countable (Union S1 S2). intros T S1 S2 h1 h2. pose (option (option False)) as twoset. pose (fun (t:twoset) => match t with | None => S1 | Some _ => S2 end) as F. assert (h3:FiniteT twoset). apply add_finite. apply add_finite. apply empty_finite. pose proof (FiniteT_impl_CountableT _ h3) as h4. assert (h5: forall t:twoset, Countable (F t)). destruct t; simpl; assumption. pose proof (countable_union F h4 h5) as h6. assert (h7:IndexedUnion F = Union S1 S2). apply Extensionality_Ensembles. red. split. red. intros x h7. (* <= *) inversion h7 as [i t h8]. destruct i. simpl in h8. apply Union_intror; assumption. simpl in h8. apply Union_introl; assumption. (* >= *) red. intros x h7. pose proof (Union_inv _ _ _ _ h7) as h8. case h8. unfold F. intro h9. apply indexed_union_intro with None; assumption. unfold F. intro h9. apply indexed_union_intro with (Some None); assumption. rewrite h7 in h6. assumption. Qed. (*Wyckoff*) Lemma countable_intersection: forall (T:Type) (S1 S2:Ensemble T), Countable S1 -> Countable S2 -> Countable (Intersection S1 S2). intros T S1 S2 h1 h2. apply (countable_downward_closed (Intersection S1 S2) S1). assumption. auto with sets. Qed. (*Schepler*) Lemma Countable_Full : forall {T:Type}, CountableT T->Countable (Full_set T). intros T h1. red. pose (fun (x:T) => (exist _ x (Full_intro _ x))) as f. apply surj_countable with f. assumption. red. intro y. exists (proj1_sig y). unfold f. destruct y. simpl. pose proof (proof_irrelevance _ i (Full_intro T x)) as h2. rewrite h2. reflexivity. Qed. (*Wyckoff*) Lemma countable_full_iff : forall {T:Type}, CountableT T <-> Countable (Full_set T). intro T. split; try apply Countable_Full. intro h1. red in h1. pose (fun x:T => (exist _ _ (Full_intro _ x))) as f. apply (inj_countable f); auto. unfold f. red; intros ? ? h2. apply exist_injective in h2. assumption. Qed. (*Schepler*) Lemma Countable_Sig : forall {T:Type} (P:T->Prop), CountableT T -> CountableT {t:T|P t}. intros T P h1. pose [t:T|P t] as S. pose proof (countable_downward_closed S (Full_set _)) as h2. pose proof (Countable_Full h1) as h3. assert (h4:Included S (Full_set _)). red. intros. constructor. pose proof (h2 h3 h4) as h5. red in h5. assert (h6:forall x:T, In S x <-> P x). intro x. split. (*->*) intro h7. inversion h7. assumption. (*<-*) intro h8. constructor. assumption. assert (h7:forall x:T, In S x -> P x). intros x h8. rewrite h6 in h8. assumption. pose (fun a:{x:T | In S x} => exist _ (proj1_sig a) (h7 (proj1_sig a) (proj2_sig a))) as f. apply surj_countable with f. assumption. red. intros y. destruct y as [t h8]. assert (h9:In S t). rewrite h6. assumption. exists (exist _ t (h9)). unfold f. simpl. pose proof (proof_irrelevance _ (h7 t h9) h8) as h10. rewrite h10. reflexivity. Qed. (*Wyckoff*) Corollary CountableT_Countable : forall (T:Type), CountableT T -> (forall (S:Ensemble T), Countable S). intros T h1 S. red. apply Countable_Sig; assumption. Qed. (*Schepler*) Lemma positive_countable: CountableT BinPos.positive. Proof. exists nat_of_P. red; intros. apply nat_of_P_inj; trivial. Qed. (*Schepler*) Lemma Z_countable: CountableT Z. Proof. destruct (countable_nat_product) as [f]. destruct positive_countable as [g]. exists (fun n:Z => match n with | Z0 => f (0, 0) | Zpos p => f (1, g p) | Zneg p => f (2, g p) end). red; intros n1 n2 ?. destruct n1 as [|p1|p1]; destruct n2 as [|p2|p2]; apply H in H1; try discriminate H1. trivial. injection H1; intro; f_equal; auto. injection H1; intro; f_equal; auto. Qed. (*Schepler*) Lemma Q_countable: CountableT QArith_base.Q. Proof. destruct countable_nat_product as [f]. destruct positive_countable as [g]. destruct Z_countable as [h]. exists (fun q:QArith_base.Q => match q with (QArith_base.Qmake n d) => f (h n, g d) end). red; intros q1 q2 ?. destruct q1 as [n1 d1]; destruct q2 as [n2 d2]. apply H in H2. injection H2; intros. f_equal; auto. Qed. End CountableT. (*Wyckoff*) Inductive point_true {T} (P:T->Prop) : Prop := | point_true_intro : forall x, P x -> point_true P. Section FiniteT_card. (*Schepler*) Lemma FiniteT_has_nat_cardinal: forall X:Type, FiniteT X -> exists! n:nat, cardinal _ (@Full_set X) n. Proof. intros. apply -> unique_existence; split. apply finite_cardinal. pose (idX := fun x:X => x). assert (Im (Full_set _) idX = Full_set _). apply Extensionality_Ensembles. red; split. red; intros; constructor. red; intros. exists x. constructor. trivial. rewrite <- H0. apply FiniteT_img with (f:=fun x:X => x). assumption. intros. case (finite_eq_dec X H y1 y2); tauto. red; intros. apply cardinal_unicity with X (Full_set _); trivial. Qed. (*Schepler*) Definition FiniteT_card (X:Type) (H:FiniteT X) : nat := proj1_sig (constructive_definite_description _ (FiniteT_has_nat_cardinal X H)). (*Schepler*) Lemma FiniteT_card_def: forall (X:Type) (H:FiniteT X), cardinal _ (@Full_set X) (FiniteT_card X H). Proof. intros; unfold FiniteT_card. destruct constructive_definite_description. assumption. Qed. (*Schepler*) Lemma FiniteT_card_cond: forall (X:Type) (H:FiniteT X) (n:nat), cardinal _ (@Full_set X) n -> FiniteT_card X H = n. Proof. intros. pose proof (FiniteT_has_nat_cardinal X H). destruct H1. red in H1. destruct H1. transitivity x. symmetry; apply H2. apply FiniteT_card_def. apply H2; trivial. Qed. (*Schepler*) Lemma FiniteT_card_False: FiniteT_card False empty_finite = 0. Proof. apply FiniteT_card_cond. assert (@Full_set False = @Empty_set False). apply Extensionality_Ensembles; red; split; auto with sets. red; intros. destruct x. rewrite H. constructor. Qed. (*Schepler*) Lemma injection_preserves_cardinal: forall (X Y:Type) (f:X->Y) (n:nat) (S:Ensemble X), cardinal _ S n -> Injective f -> cardinal _ (Im S f) n. Proof. intros. induction H. assert (Im (Empty_set _) f = (Empty_set _)). apply Extensionality_Ensembles; split; auto with sets. red; intros. destruct H. destruct H. rewrite H; constructor. assert (Im (Add A x) f = Add (Im A f) (f x)). apply Extensionality_Ensembles; split. red; intros. destruct H2. symmetry in H3; destruct H3. destruct H2. left; exists x0; auto with sets. destruct H2; right; auto with sets. red; intros. destruct H2. destruct H2. exists x0. left; auto with sets. assumption. destruct H2. exists x; trivial; right; auto with sets. rewrite H2. constructor; trivial. red; intro H3; inversion H3. apply H0 in H5; destruct H5. contradiction H1. Qed. (*Schepler*) Lemma FiniteT_card_option: forall (X:Type) (H:FiniteT X), FiniteT_card (option X) (add_finite X H) = S (FiniteT_card X H). Proof. intros. apply FiniteT_card_cond. assert ((Full_set _) = Add (Im (Full_set _) (@Some X)) None). apply Extensionality_Ensembles; split. red; intros. destruct x. left; exists x; constructor. right; constructor. red; intros; constructor. rewrite H0. constructor. apply injection_preserves_cardinal. apply FiniteT_card_def. red; intros x1 x2 Heq; injection Heq; trivial. red; intro. inversion H1. discriminate H3. Qed. (*Schepler*) Lemma FiniteT_card_bijection: forall (X Y:Type) (H:FiniteT X) (g:X->Y) (Hinv:invertible g), FiniteT_card Y (bij_finite X Y g H Hinv) = FiniteT_card X H. Proof. intros. apply FiniteT_card_cond. apply invertible_impl_bijective in Hinv. destruct Hinv as [g_inj g_surj]. assert ((Full_set _)= Im (Full_set _) g). apply Extensionality_Ensembles; split; red; intros; try constructor. destruct (g_surj x). exists x0; try constructor; auto. rewrite H0; apply injection_preserves_cardinal; trivial. apply FiniteT_card_def. Qed. (*Wyckoff*) Lemma unit_finite : FiniteT unit. pose (fun x:(option False) => tt) as f. assert (h1:invertible f). apply bijective_impl_invertible. red; split. red. intros x1 x2. destruct x1; destruct x2; try contradiction; auto. red. intros x1. exists None. unfold f. destruct x1. reflexivity. eapply bij_finite; auto. eapply add_finite. apply empty_finite. apply h1. Qed. (*Wyckoff*) Lemma FiniteT_card_unit : FiniteT_card _ unit_finite = 1. assert (h2:FiniteT_card _ (add_finite _ empty_finite) = 1). pose proof (FiniteT_card_False) as h2. rewrite <- h2. apply FiniteT_card_option. rewrite <- h2. pose (fun x:(option False) => tt) as f. assert (h3:invertible f). apply bijective_impl_invertible. red; split. red. intros x1 x2. destruct x1; destruct x2; try contradiction; auto. red. intros x1. exists None. unfold f. destruct x1. reflexivity. generalize unit_finite. intro h4. pose proof (proof_irrelevance _ h4 (bij_finite _ _ f (add_finite _ empty_finite) h3)); subst. eapply FiniteT_card_bijection. Qed. (*Wyckoff*) Lemma FiniteT_card_fun1_compat : forall {T:Type} (pf:FiniteT T), FiniteT_card _ pf = card_fun1 (Full_set T). intros T h1. pose proof (card_fun1_compat (Full_set T)) as h2. destruct h2 as [h2l h2r]. pose proof (Finite_FiniteT_iff T) as h3. pose proof h1 as h1'. rewrite <- h3 in h1'. specialize (h2l h1'). pose proof (FiniteT_card_cond T h1 _ h2l) as h4. assumption. Qed. (*Wyckoff*) Lemma FiniteT_card_add_empty : FiniteT_card _ (add_finite _ empty_finite) = 1. rewrite (FiniteT_card_option _ empty_finite). f_equal; apply FiniteT_card_False. Qed. (*Wyckoff*) Lemma FiniteT_card_eq0 : forall T (pf:FiniteT T), FiniteT_card T pf = 0 -> equi T False. intros T h1. pose proof card_empty T as h2. intros h3. pose proof (FiniteT_card_def T h1) as h4. rewrite h3 in h4. clear h3 h2. revert h4. induction h1; auto. intros h2. exists id. apply invertible_id. intro h2. inversion h2. pose proof (Full_intro (option T) None) as h3. rewrite <- H0 in h3. inversion h3. intro h2; auto. inversion h2. pose (fun y:Y => (Full_intro _ y)) as g. rewrite <- H1 in g. pose proof (equi_intro _ False (fun y => match (g y) with | end)) as h3. apply h3. econstructor. intro x. pose proof (g x) as h4. contradiction. intro; contradiction. Grab Existential Variables. intro; contradiction. Qed. End FiniteT_card. (*Wyckoff*) Lemma FiniteT_iff : forall T, FiniteT T <-> (equi T False \/ (exists U, FiniteT U /\ equi T (option U)) \/ (exists U (f:U->T), invertible f /\ FiniteT U)). intros T; split; intro h1. induction h1. left; auto. apply equi_refl. destruct IHh1 as [|h2]; subst. right. left. exists False. split; auto. apply empty_finite. inversion H. apply equi_intro with (option_map f). apply invertible_option_map; auto. destruct h2 as [h2 | h2]. destruct h2 as [U [h2]]. inversion H. right. left. exists (option U). split. constructor; auto. apply equi_intro with (option_map f). apply invertible_option_map; auto. destruct h2 as [U [f h2]]. right; right. destruct h2 as [h2 h3]. exists (option U), (option_map f). split. apply invertible_option_map; auto. constructor; auto. destruct IHh1 as [h2|h2]. left. apply (equi_trans _ _ _ (equi_intro _ _ _ (invertible_impl_inv_invertible _ H)) h2). right. destruct h2 as [h2|h2]. destruct h2 as [U [h2 h3]]. left. exists U. split; auto. apply (equi_trans _ _ _ (equi_intro _ _ _ (invertible_impl_inv_invertible _ H)) h3). right. destruct h2 as [U [g [h3 h4]]]. exists U. exists (compose f g). split; auto. apply invertible_compose; auto. destruct h1 as [h1 | [h1 | h1]]. inversion h1. apply (bij_finite _ _ _ empty_finite (invertible_impl_inv_invertible _ H)). destruct h1 as [U [h1 h2]]. inversion h2. apply (bij_finite _ _ _ (add_finite _ h1) (invertible_impl_inv_invertible _ H)). destruct h1 as [U [f [h1 h2]]]. apply (bij_finite _ _ _ h2 h1). Qed. (*Wyckoff*) Inductive FiniteT_prop (P:forall T, FiniteT T -> Prop) : Prop := | empty_finite_prop : P False empty_finite -> FiniteT_prop P | add_finite_prop : (forall U (pfu:FiniteT U), P U pfu -> P (option U) (add_finite U pfu)) -> FiniteT_prop P | bij_finite_prop : (forall U V (pfu:FiniteT U) (f:U->V) (pfinv:invertible f), P U pfu -> P V (bij_finite _ _ f pfu pfinv)) -> FiniteT_prop P. (*Wyckoff*) Inductive FiniteT_prop_dep (P:forall T, FiniteT T -> Prop) : forall T, FiniteT T -> Prop := | empty_finite_prop_dep : P False empty_finite -> FiniteT_prop_dep P False empty_finite | add_finite_prop_dep : (forall U (pfu:FiniteT U), P U pfu -> P (option U) (add_finite U pfu)) -> forall T (pft:FiniteT T), P T pft -> FiniteT_prop_dep P (option T) (add_finite T pft) | bij_finite_prop_dep : (forall U V (pfu:FiniteT U) (f:U->V) (pfinv:invertible f), P U pfu -> P V (bij_finite _ _ f pfu pfinv)) -> forall T U (pft:FiniteT T) (f:T->U) (pfinv:invertible f), P T pft -> FiniteT_prop_dep P U (bij_finite _ _ f pft pfinv). (*Wyckoff*) Lemma FiniteT_ind' : forall {T} (pft:FiniteT T) (P:forall T, FiniteT T->Prop), FiniteT_prop_dep P T pft -> P T pft. intros ? ? ? h1; induction h1; auto. Qed. (*Wyckoff*) (*This section is included as a curiosity only; its use of providing a canonical index on elements in a finite type is better served by [FinT] in [TypeUtilities2].*) Section FiniteTIndex. (*option (option ( ... False*) Fixpoint option_false n : Type := match n with | 0 => False | S n' => option (option_false n') end. Inductive FiniteT_index : forall {T} (pff:FiniteT T), (T -> (FiniteT_card _ pff) @) -> Prop := | FiniteT_ind_empty : FiniteT_index empty_finite (fun x => False_rect _ x) | FiniteT_ind_add : forall {T} (pff:FiniteT T) (f:T->(FiniteT_card _ pff) @), FiniteT_index pff f -> FiniteT_index (add_finite _ pff) (fun x => let pf0 := lt_eq_trans _ _ _ (lt_0_Sn (FiniteT_card _ pff)) (eq_sym (FiniteT_card_option _ pff)) in match x with | Some x => let z := f x in let pflt := lt_eq_trans _ _ _ (lt_n_S _ _ (proj2_sig z)) (eq_sym (FiniteT_card_option _ pff)) in exist (fun i => i < FiniteT_card _ (add_finite T pff)) (S (proj1_sig z)) pflt | None => exist _ 0 pf0 end) | FiniteT_ind_bij : forall {T U} (g:T->U) (pff:FiniteT T) (pfg:invertible g) (f:T->(FiniteT_card _ pff) @), FiniteT_index pff f -> FiniteT_index (bij_finite T U g pff pfg) (fun y => let z := f (proj1_sig (function_inverse _ pfg) y) in let pfz := lt_eq_trans _ _ _ (proj2_sig z) (eq_sym (FiniteT_card_bijection T U pff g pfg)) in exist (fun i => i < FiniteT_card _ (bij_finite T U g pff pfg)) (proj1_sig z) pfz). Lemma injective_FiniteT_index : forall {T} (pff:FiniteT T) (f:T->(FiniteT_card _ pff) @), FiniteT_index pff f -> injective f. intros T hff f h1. induction h1. apply inj_false. red in IHh1; red. intros x y; destruct x as [x|], y as [y|]; intro h2; apply exist_injective in h2; f_equal; try discriminate; auto. apply S_inj in h2. apply proj1_sig_injective in h2; auto. red; intros x y h2. apply exist_injective in h2. apply proj1_sig_injective in h2. apply IHh1 in h2. pose proof (invertible_impl_inv_invertible _ pfg) as h3. pose proof (invertible_impl_bijective _ h3) as h4. destruct h4 as [h4 h5]. apply h4 in h2; auto. Qed. Lemma surjective_FiniteT_index : forall {T} (pff:FiniteT T) (f:T->(FiniteT_card _ pff) @), FiniteT_index pff f -> surjective f. intros T hf f h1. induction h1. red. intro y. destruct y as [? hy]. pose proof hy as hy'. rewrite FiniteT_card_False in hy'; omega. red. intro y. destruct y as [y hy]. destruct y as [|y]. exists None. apply proj1_sig_injective; auto. pose proof hy as hy'. rewrite FiniteT_card_option in hy'. apply lt_S_n in hy'. red in IHh1. specialize (IHh1 (exist (fun c => c < FiniteT_card T pff) _ hy')). destruct IHh1 as [x h3]. pose proof (injective_FiniteT_index pff _ h1) as h4. pose proof (f_equal (@proj1_sig _ _) h3) as h5. simpl in h5; subst. pose proof (in_im x f) as h6. exists (Some x). f_equal; apply proof_irrelevance. red in IHh1. red. intro y; destruct y as [y hy]. pose proof hy as hy'. rewrite FiniteT_card_bijection in hy'. specialize (IHh1 (exist (fun c => c < FiniteT_card _ pff) y hy')). destruct IHh1 as [x hx]. pose proof (f_equal (@proj1_sig _ _) hx) as h2. simpl in h2; subst. exists (g x). apply proj1_sig_injective; simpl. pose proof (proj2_sig (function_inverse g pfg)) as h2. simpl in h2. destruct h2 as [h2 h3]. rewrite h2; auto. Qed. Lemma bijective_FiniteT_index : forall {T} (pff:FiniteT T) (f:T->(FiniteT_card _ pff) @), FiniteT_index pff f -> bijective f. intros; red; split; [apply injective_FiniteT_index | apply surjective_FiniteT_index]; auto. Qed. End FiniteTIndex. (*Wyckoff*) (*[Set]-version of [sumor], which is usually not necessary.*) Inductive sumor' (A:Set) (B:Prop) : Set := | inleft' : A -> sumor' A B | inright' : B -> sumor' A B. (*Wyckoff*) Fixpoint nsumorl (l:list Prop) : Set := match l return Set with | nil => False | cons P l' => sumor (nsumorl l') P end. Section NProd. (*Wyckoff*) Lemma nprod0_eq : forall {T} (x y:T^0), x = y. intros T x y; destruct x, y; auto. Qed. (*Wyckoff*) Fixpoint in_nprod {T n} y : T^n -> Prop := match n with | 0 => fun _ => False | S n' => fun x => y = fst x \/ in_nprod y (snd x) end. (*Wyckoff*) Fixpoint hd_nprod {T n} : n > 0 -> forall (x:T^n), T := match n with | 0 => fun pf => False_rect _ (lt_irrefl _ pf) | S _ => fun _ x => fst x end. (*Wyckoff*) Fixpoint tl_nprod {T n} : n > 0 -> forall (x:T^n), T^pred n := match n with | 0 => fun pf => False_rect _ (lt_irrefl _ pf) | S _ => fun _ x => snd x end. (*Wyckoff*) Lemma in_nprod_fst : forall {T n} (x:T^S n), in_nprod (fst x) x. intros; simpl; left; auto. Qed. (*Wyckoff*) Definition all_tuples (T:Type) := {n:nat & T^n}. (*Wyckoff*) Lemma full_set_all_tuples_eq : forall T:Type, Full_set (all_tuples T) = IndexedUnion (fun n:nat => [x:all_tuples T | projT1 x = n]). intro T. apply Extensionality_Ensembles; red; split; red; intros x h1. clear h1. destruct x as [n x]. apply indexed_union_intro with n. constructor. simpl. reflexivity. constructor. Qed. (*Wyckoff*) Definition match_all_tuples_type T n : Type := match n with | 0 => unit | S _ => (T * all_tuples T)%type end. (*Wyckoff*) Definition match_all_tuples {T} (tpl:all_tuples T) : match_all_tuples_type T (projT1 tpl) := let (n, x) as pr return (match_all_tuples_type T (projT1 pr)) := tpl in match n as m return (T ^ m -> match_all_tuples_type T m) with | 0 => fun _ => tt | S m' => fun tpl': T ^ S m' => (fst tpl', existT _ m' (snd tpl')) end x. (*Wyckoff*) Lemma finite_nprod : forall T n, FiniteT T -> FiniteT (T^n). intros T n. induction n as [|n ih]; simpl. intros. apply unit_finite. intro h1. specialize (ih h1). apply finite_prod; auto. Qed. (*Wyckoff*) Lemma countable_nprod : forall (T:Type) (n:nat), CountableT T -> CountableT (T^n). intros T n. induction n as [|n h1]; simpl; auto. intros. apply FiniteT_impl_CountableT. apply unit_finite. intro h2. specialize (h1 h2). apply countable_product; auto. Qed. (*Wyckoff*) Lemma countable_all_tuples_of_countable_type : forall T:Type, CountableT T -> CountableT (all_tuples T). intros T h1. rewrite countable_full_iff. rewrite full_set_all_tuples_eq. apply countable_union. apply countable_nat. intro n. pose (fun x:T^n => existT _ n x) as f. assert (h3: [x : all_tuples T | projT1 x = n] = Im (Full_set (T ^ n)) f). apply Extensionality_Ensembles; red; split; red; intros x h3. destruct h3 as [h3]. destruct x as [n' x]. simpl in h3. subst. apply Im_intro with x. constructor. unfold f. f_equal. destruct h3 as [x h3]. subst. clear h3. constructor. unfold f. simpl. reflexivity. rewrite h3. apply countable_img. rewrite <- countable_full_iff. apply countable_nprod; auto. Qed. (*Wyckoff*) Lemma nprod_to_list_eq_nil_iff : forall {T n} (x:T^n), nprod_to_list _ _ x = nil <-> n = 0. intros T n. destruct n as [|n]; simpl. tauto. intro x; destruct x as [fx sx]. split; intro h1; discriminate. Qed. (*Wyckoff*) Lemma length_nprod_to_list : forall {T n} (x:T^n), length (nprod_to_list _ _ x) = n. intros ? n. induction n as [|n ih]; simpl; auto. intro x. destruct x as [x y]. simpl; f_equal. apply ih. Qed. (*Wyckoff*) Fixpoint nprod_deg {n T} : T^n -> nat := match n with | 0 => fun _ => 0 | S n' => fun x => S (nprod_deg (snd x)) end. (*Wyckoff*) Lemma length_nprod_of_list : forall {T} (l:list T), nprod_deg (nprod_of_list _ l) = length l. intros ? l; induction l; simpl; auto. Qed. (*Wyckoff*) Definition hd_prod {T n} : T^S n -> T := fun x => fst x. (*Wyckoff*) Definition tl_prod {T n} : T^S n->T ^ n := fun x => snd x. (*Wyckoff*) (*A tuple of [n] identical copies of [x]*) Fixpoint nprod_sing {T} (x:T) n : T^n := match n with | 0 => tt | S n' => (x, nprod_sing x n') end. (*Wyckoff*) Definition None_n {T:Type} n : (option T)^n := nprod_sing None n. (*Wyckoff*) Fixpoint removelast_prod {T:Type} {n:nat} : T ^ S n -> T^n := match n return (T ^ S n -> T ^ n) with | O => fun y => tt | S n' => fun y => (fst y, removelast_prod (snd y)) end. (*Wyckoff*) Fixpoint nth_prod {T:Type} {n:nat} i : i < n -> T^n -> T := match n return (i < n -> T^n -> T) with | 0 => fun pf _ => False_rect _ (lt_n_0 _ pf) | S n' => match i with | 0 => fun _ x => fst x | S i' => fun pflt x => let pflt' := lt_S_n _ _ pflt in @nth_prod T _ _ pflt' (snd x) end end. (*Wyckoff*) Lemma nth_prod0 : forall {T:Type} {n} (pf:0 < S n) (x:T^S n), nth_prod 0 pf x = fst x. intros ? n; destruct n; simpl; auto. Qed. (*Wyckoff*) Lemma hd_nprod_eq : forall {T n} (pf:n > 0) (x:T^n), hd_nprod pf x = nth_prod 0 pf x. intros T n; destruct n; auto. intros; omega. Qed. (*Wyckoff*) Lemma nth_prod_S : forall {T n} i (pfsi:S i < S n) (x:T^S n), nth_prod (S i) pfsi x = nth_prod i (lt_S_n _ _ pfsi) (snd x). auto. Qed. (*Wyckoff*) Lemma in_nth_prod : forall {T n} i (pfi:i < n) (x:T^n), in_nprod (nth_prod i pfi x) x. intros T n. induction n as [|n ih]; simpl. intros; omega. intro i. destruct i as [|i]. intros; left; auto. intros hi x. specialize (ih i (lt_S_n _ _ hi) (snd x)); auto. Qed. (*Wyckoff*) Lemma in_nprod_iff : forall {T n} y (x:T^n), in_nprod y x <-> exists i (pfi:i < n), nth_prod i pfi x = y. intros T n. induction n as [|n ih]; simpl. intros; split; intro h1. contradiction. destruct h1 as [? [? ?]]; omega. intros y x; split; intro h1. destruct h1 as [|h1]; subst. exists 0, (lt_0_Sn _); auto. rewrite ih in h1. destruct h1 as [i [hi ?]]; subst. exists (S i), (lt_n_S _ _ hi). f_equal; apply proof_irrelevance. destruct h1 as [i [hi h2]]. destruct i as [|i]; subst. left; auto. right. rewrite ih. exists i, (lt_S_n _ _ hi); auto. Qed. (*Wyckoff*) Lemma nth_prod_seg_fun_pf_eq : forall {T n} (x:T^n) (lt_fun : forall i, i < n -> i < n), (fun i (pfi:i < n) => (nth_prod i pfi x)) = (fun i (pfi:i < n) => (nth_prod i (lt_fun i pfi) x)). intros; apply functional_extensionality_dep; intro; apply functional_extensionality_dep; intros; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma nth_prod_seg_fun_pf2_eq : forall {T n} (x:T^n) (lt_fun1 : forall i, i < n -> i < n) (lt_fun2 : forall i, i < n -> i < n), (fun i (pfi:i < n) => (nth_prod i (lt_fun1 i pfi) x)) = (fun i (pfi:i < n) => (nth_prod i (lt_fun2 i pfi) x)). intros; apply functional_extensionality_dep; intro; apply functional_extensionality_dep; intros; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Fixpoint plus_nprod {n} : forall (x:nat^n), nat := match n with | 0 => fun _ => 0 | S n' => fun x => (fst x) + (plus_nprod (snd x)) end. (*Wyckoff*) Fixpoint times_nprod {n} : forall (x:nat^n), nat := match n with | 0 => fun _ => 1 | S n' => fun x => (fst x) * (times_nprod (snd x)) end. (*Wyckoff*) Definition all_sing_nprod {T n} (x:T^n) (a:T) := forall b, in_nprod b x -> b = a. (*Wyckoff*) Definition in_pre_w {X S n} (s:S) (v:X->S) (w:S^n) : Prop := forall i (pfi:i < n), in_pre (nth_prod i pfi w) v. (*Wyckoff*) Definition last_prod {T n} : T^S n -> T := fun x => nth_prod n (lt_n_Sn _) x. (*Wyckoff*) Fixpoint im_nprod {T U n} (f:T->U) : T^n -> U^n := match n with | 0 => fun _ => tt | S n' => fun x => ((f (fst x)), im_nprod f (snd x)) end. (*Wyckoff*) Lemma im_nprod_im_nprod : forall {T U V:Type} {n:nat} (t:T^n) (f:T->U) (g:U->V), im_nprod g (im_nprod f t) = im_nprod (fun x=> g (f x)) t. intros T U V n. induction n as [|n h1]. intros t f g; simpl. reflexivity. intros t f g; simpl. apply injective_projections; simpl. reflexivity. destruct t as [x t]. simpl. apply h1. Qed. (*Wyckoff*) Lemma im_nprod_inj : forall {T U:Type} {n:nat} (t1 t2:T^n) (f:T->U), Injective f -> im_nprod f t1 = im_nprod f t2 -> t1 = t2. intros T U. induction n as [|n h1]; intros t1 t2 f h2 h3; simpl. simpl in h3. destruct t1. destruct t2. reflexivity. destruct t1 as [ft1 st1]. destruct t2 as [ft2 st2]. simpl in h3. inversion h3 as [h4]. specialize (h1 _ _ _ h2 H). red in h2. specialize (h2 _ _ h4). apply injective_projections; auto. Qed. (*Wyckoff*) Fixpoint map_nprod {U n} : seg_fun n U -> U^n := match n with | 0 => fun _ => tt | S n' => fun f => let f' := seg_fun_to_S f in (f 0 (lt_0_Sn _), map_nprod f') end. (*Wyckoff*) Lemma nth_map_nprod : forall {U n} (f:seg_fun n U) i (pfi:i < n), nth_prod i pfi (map_nprod f) = f i pfi. intros U n. revert U. induction n as [|n ih]. intros; omega. intros U f i. destruct i as [|i]. simpl. intros; f_equal; apply proof_irrelevance. intro h1. specialize (ih U (seg_fun_to_S f) i (lt_S_n _ _ h1)). simpl. rewrite ih. unfold seg_fun_to_S; f_equal. apply proof_irrelevance. Qed. Definition map_nprod_functional : forall {U n} (f g:seg_fun n U), f = g -> map_nprod f = map_nprod g. intros; subst; auto. Qed. Lemma map_nprod_eq_iff : forall {U n} (f:seg_fun n U) (y:U^n), map_nprod f = y <-> forall i (pfi:i < n), f i pfi = nth_prod i pfi y. intros U n; induction n as [|n ih]; simpl. intros; split; subst. intros; omega. intros; apply unit_eq. unfold seg_fun_to_S. intros f y; destruct y as [y y']; split; intro h1. inversion h1; subst. intro i; destruct i as [|i]; simpl; intros; repeat f_equal. apply proof_irrelevance. rewrite nth_map_nprod; f_equal; apply proof_irrelevance. pose proof (h1 0 (lt_0_Sn _)) as h2; simpl in h2; subst. f_equal. rewrite ih. intros i hi. rewrite (h1 (S i) (lt_n_S _ _ hi)); f_equal; apply proof_irrelevance. Qed. (*This transfer is done pointwise, rather than as a whole, to avoid universe inconsistencies, which in general arise when you're assuming too much type-theoretical information, which in this case is that, type-substitutions would destruct correspdondingly.*) (*Wyckoff*) Definition transfer_nprod {T n n'} (pfe:n = n') (x:T^n) : T^n' := map_nprod (fun i (pfi:i < n') => nth_prod i (lt_eq_trans _ _ _ pfi (eq_sym pfe)) x). (*Wyckoff*) Definition transfer_nprod_r {T n n'} (pfe:n = n') (x:T^n') : T^n := map_nprod (fun i (pfi:i < n) => nth_prod i (lt_eq_trans _ _ _ pfi pfe) x). (*Wyckoff*) Lemma transfer_nprod_eq_refl : forall {T n} (pfe:n = n) (x:T^n), transfer_nprod pfe x = x. unfold transfer_nprod; intros T n. revert T. induction n as [|n ih]; simpl. intros; apply unit_eq. intros T he x; simpl. destruct x as [x w]; simpl. f_equal. specialize (ih T (S_inj _ _ he) w). rewrite <- ih. f_equal. apply functional_extensionality_dep. intro. unfold seg_fun_to_S. apply functional_extensionality_dep. intro. f_equal. apply proof_irrelevance. simpl; auto. Qed. (*Wyckoff*) Lemma transfer_nprod_r_eq_refl : forall {T n} (pfe:n = n) (x:T^n), transfer_nprod_r pfe x = x. unfold transfer_nprod_r; intros T n. revert T. induction n as [|n ih]; simpl. intros; apply unit_eq. intros T he x; simpl. destruct x as [x w]; simpl. f_equal. specialize (ih T (S_inj _ _ he) w). rewrite <- ih. f_equal. apply functional_extensionality_dep. intro. unfold seg_fun_to_S. apply functional_extensionality_dep. intro. f_equal. apply proof_irrelevance. simpl; auto. Qed. (*Wyckoff*) Lemma transfer_nprod_eq : forall {T m n} (pfe:m = n) (x:T^m), transfer_nprod pfe x = transfer_nprod_r (eq_sym pfe) x. intros; subst; rewrite transfer_nprod_eq_refl, transfer_nprod_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_nprod_r_eq : forall {T m n} (pfe:m = n) (x:T^n), transfer_nprod_r pfe x = transfer_nprod (eq_sym pfe) x. intros; subst; rewrite transfer_nprod_eq_refl, transfer_nprod_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_nprod_r_undoes_transfer_nprod : forall {T m n} (pfe pfe':m = n) (x:T^m), transfer_nprod_r pfe' (transfer_nprod pfe x) = x. intros; subst; rewrite transfer_nprod_eq_refl, transfer_nprod_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_nprod_undoes_transfer_nprod_r : forall {T m n} (pfe pfe':m = n) (x:T^n), transfer_nprod pfe' (transfer_nprod_r pfe x) = x. intros; subst; rewrite transfer_nprod_eq_refl, transfer_nprod_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_nprod_S : forall {T n m} (pfe:S n = S m) (x:T^S n), transfer_nprod pfe x = (fst x, transfer_nprod (S_inj _ _ pfe) (snd x)). intros ? ? ? h1 x; pose proof (S_inj _ _ h1); subst; intros; do 2 rewrite transfer_nprod_eq_refl; destruct x; auto. Qed. (*Wyckoff*) Lemma nprod_transfer_plus_pr : forall {T m n} (pf:S m = S m + n) (x:T^m) (a:T), let pf' := S_eq_plus _ _ pf in transfer_nprod pf (a, x) = (a, transfer_nprod pf' x). intros T m n h1. pose proof (eq_plus_l _ _ h1); subst. intros x a. simpl. generalize (S_eq_plus m 0 h1). revert h1. simpl. rewrite <- plus_n_O. intros. do 2 rewrite transfer_nprod_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_nprod_r_S : forall {T n m} (pfe:S n = S m) (x:T^S m), transfer_nprod_r pfe x = (fst x, transfer_nprod_r (S_inj _ _ pfe) (snd x)). intros ? ? ? h1 x; pose proof (S_inj _ _ h1); subst; intros; do 2 rewrite transfer_nprod_r_eq_refl; destruct x; auto. Qed. (*Wyckoff*) Lemma nth_prod_transfer_prod : forall {T m n} {U:T->Type} (pfe:m = n) (w:T^m) i (pfi:i < n), let pfi' := lt_eq_trans _ _ _ pfi (eq_sym pfe) in nth_prod i pfi (transfer_nprod pfe w) = nth_prod i pfi' w. simpl; intros; subst; rewrite transfer_nprod_eq_refl; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma nth_prod_transfer_prod_r : forall {T m n} {U:T->Type} (pfe:m = n) (w:T^n) i (pfi:i < m), let pfi' := lt_eq_trans _ _ _ pfi pfe in nth_prod i pfi (transfer_nprod_r pfe w) = nth_prod i pfi' w. simpl; intros; subst; rewrite transfer_nprod_r_eq_refl; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma nth_prod_functional0 : forall {T n n'} i (pfi:i < n) (pfi':i < n') (w:T^n) (pfe:n = n'), @nth_prod _ n i pfi w = @nth_prod _ n' i pfi' (transfer_nprod pfe w). intros; subst; rewrite transfer_nprod_eq_refl; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma nth_prod_functional1 : forall {T n} i i' (pfi:i < n) (pfi':i' < n) (w:T^n) (pfe:i = i'), nth_prod i pfi w = nth_prod i' pfi' w. intros; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma nth_prod_functional2 : forall {T n} i (pfi pfi':i < n) (w:T^n), nth_prod i pfi w = nth_prod i pfi' w. intros; f_equal; apply proof_irrelevance. Qed. (**Wyckoff*) Lemma nth_prod_functional3 : forall {T n} i (pfi pfi':i < n) (w w':T^n), w = w' -> nth_prod i pfi w = nth_prod i pfi' w'. intros; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma nprod_ext : forall {T n} (x y:T^n), (forall i (pfi:i < n), nth_prod i pfi x = nth_prod i pfi y) -> x = y. intros T n; induction n as [|n ih]; simpl. intros; apply unit_eq. intros x y. destruct x as [t x], y as [u y]; simpl. intro h1. pose proof (h1 0 (lt_0_Sn _)) as h2; simpl in h2; subst; f_equal. apply ih. intros i hi. specialize (h1 (S i) (lt_n_S _ _ hi)); simpl in h1. erewrite nth_prod_functional2. rewrite h1 at 1; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) (*This is just the first [n] entries, as opposed to the last [n], as in [tl_prod].*) Fixpoint fst_tl_prod {T n} : T^S n->T ^ n := fun x => map_nprod (fun i (pf:i < n) => let pf' := lt_trans _ _ _ pf (lt_n_Sn _) in nth_prod i pf' x). (*Wyckoff*) (*This function allows you to specify a (presumably surjective) mapping from [seg n] to [seg r] that corresponds indices of the returned tuple in [T^n] to indices of the given tuple, [x:T^r].*) Definition nprod_identify {T n r} {t:seg_fun n nat} : seg_fun_endo t r -> T^r -> T^n := fun pfs x => map_nprod (fun i (pfi:i < n) => let pflt := lt_seg_fun_endo t pfs i pfi in nth_prod (t i pfi) pflt x). (*This is the analogue of [nprod] in [Set]. Usaully unnecessary.*) Fixpoint setprod (U:Set) (n:nat) : Set := match n with | 0 => unit | S n' => (U * setprod U n')%type end. (*Wyckoff*) Definition full_segT_ind (n:nat) : (n @)^n := map_nprod (fun i (pfi:i < n) => exist (fun c => c < n) i pfi). (*Wyckoff*) Lemma nth_full_segT_ind : forall {n} i (pfi:i < n), proj1_sig (nth_prod i pfi (full_segT_ind n)) = i. intros n i hi. unfold full_segT_ind. rewrite nth_map_nprod; simpl; auto. Qed. (*Wyckoff*) Lemma nth_full_segT_ind_S : forall {n} i (pfi:i < n), proj1_sig (nth_prod i pfi (snd (full_segT_ind (S n)))) = S i. intro n. induction n as [|n ih]; simpl. intros; omega. intros i hi. destruct i as [|i]; simpl; auto. specialize (ih i (lt_S_n _ _ hi)). rewrite nth_map_nprod. unfold seg_fun_to_S; simpl; auto. Qed. (*Wyckoff*) Definition seg_fun_type_to_ind {n} (f:seg_fun n Type) : n @ -> Type := fun i => f (proj1_sig i) (proj2_sig i). (*Wyckoff*) Lemma seg_fun_type_to_ind0 : forall {n} (f:seg_fun (S n) Type), seg_fun_type_to_ind f (fst (full_segT_ind (S n))) = f 0 (lt_0_Sn n). unfold seg_fun_type_to_ind; auto. Qed. (*Wyckoff*) Definition pointwise_ind_eq {T U V n} (P:T->V) (Q:U->V) (w:T^n) (w':U^n) : Prop := forall i (pfi:i < n), P (nth_prod i pfi w) = Q (nth_prod i pfi w'). (*Wyckoff*) Lemma eq_pointwise_ind_eq : forall {T U n} {w x:T^n}, w = x -> forall (P:T->U), pointwise_ind_eq P P w x. intros; subst; red; intros; auto. Qed. (*Wyckoff*) Definition pointwise_ind_eq_to_S {T U V n} {P:T->V} {Q:U->V} {w:T^(S n)} {w':U^(S n)} (pfp:pointwise_ind_eq P Q w w') : pointwise_ind_eq P Q (snd w) (snd w'). red in pfp; red; intros i hi; specialize (pfp (S i) (lt_n_S _ _ hi)); simpl in pfp; erewrite nth_prod_functional2. rewrite pfp at 1. repeat f_equal; apply proof_irrelevance. Defined. (*Wyckoff*) Lemma pointwise_ind_eq_seg_fun_type_to_ind : forall {n} (f:seg_fun (S n) Type), pointwise_ind_eq (seg_fun_type_to_ind (seg_fun_to_S f)) (seg_fun_type_to_ind f) (full_segT_ind n) (snd (full_segT_ind (S n))). intros n f; red; intros i hi. unfold seg_fun_type_to_ind, seg_fun_to_S; f_equal. gen0. intro h1. symmetry. gen0. revert h1. rewrite nth_full_segT_ind. rewrite nth_full_segT_ind_S. intros; f_equal. apply proof_irrelevance. Qed. (*Wyckoff*) (*A binary relation establishing corresponding types of tuples.*) Definition pointwise_ind_eq_rel {T V} (P:T->V) n : Relation (T ^ n) := fun w w' => forall i (pfi:i < n), P (nth_prod i pfi w) = P (nth_prod i pfi w'). (*Wyckoff*) Lemma refl_pointwise_ind_eq_rel : forall {T V} (P:T->V) n, Reflexive (pointwise_ind_eq_rel P n). intros; red; intro; red; intros; auto. Qed. (*Wyckoff*) Lemma symm_pointwise_ind_eq_rel : forall {T V} (P:T->V) n, Symmetric (pointwise_ind_eq_rel P n). intros; red; intro; red; intros; auto. Qed. (*Wyckoff*) Lemma trans_pointwise_ind_eq_rel : forall {T V} (P:T->V) n, Transitive (pointwise_ind_eq_rel P n). intros; red; intro; red; intros; auto. rename H into h1, H0 into h2. red in h1, h2. specialize (h1 i pfi). specialize (h2 i pfi). congruence. Qed. (*Wyckoff*) Add Parametric Relation {T V n} (P:T->V) : (T ^ n) (pointwise_ind_eq_rel P n) reflexivity proved by (refl_pointwise_ind_eq_rel P n) symmetry proved by (symm_pointwise_ind_eq_rel P n) transitivity proved by (trans_pointwise_ind_eq_rel P n) as pointwise_ind_eq_rel_equiv. Section TupleProd. (*For "many sorts"*) Fixpoint tuple_prod {T n} (A:T->Type) : T^n -> Type := match n with | 0 => fun _ => unit | S n' => fun x:T^S n' => ((A (fst x)) * (tuple_prod A (snd x)))%type end. Fixpoint nth_tuple_prod {T n} {A:T->Type} : forall {t:T^n}, tuple_prod A t -> forall i (pf:i < n), A (nth_prod i pf t) := match n with | 0 => fun _ _ _ pf0 => False_rect _ (lt_n_0 _ pf0) | S n' => fun (t:T^(S n')) (y:tuple_prod A t) i => match i with | 0 => fun _ => (fst y) | S i' => fun pfi => let pfi' := lt_S_n _ _ pfi in @nth_tuple_prod T n' A (snd t) (snd y) i' pfi' end end. Lemma tuple_prod_const : forall {n} T U (x:T^n), tuple_prod (fun _ => U) x = (U^n)%type. intros n; induction n as [|n ih]; auto. intros ? ? x. destruct x as [fx sx]. specialize (ih T U sx). simpl. f_equal; auto. Defined. Lemma tuple_prod0 : forall {S} {A:S->Type} (x:S^0), tuple_prod A x = unit. auto. Qed. Lemma tuple_prod_S : forall {T n} (A:T->Type) (x:T^S n), tuple_prod A x = (A (fst x) * tuple_prod A (snd x))%type. auto. Qed. Definition transfer_tuple_prod1 {T} {A:T->Type} {x:T^1} (p:tuple_prod A x) : A (fst x) := fst p. Definition transfer_to_tuple_prod1 {T} {A:T->Type} {x:T} (a:A x) : tuple_prod A (x, tt) (n := 1) := (a, tt). Definition tuple_prod_all {S} (A:S->Type) (x:all_tuples S) : Type := tuple_prod A (projT2 x). Lemma tuple_prod_all0 : forall {S} (A:S->Type) (x:S^0), tuple_prod_all A (existT _ _ x) = unit. unfold tuple_prod_all; intros; simpl; auto. Qed. Lemma tuple_prod_all_tuple : forall {S} (A:S->Type) (x:all_tuples S), tuple_prod_all A x = tuple_prod A (projT2 x). unfold tuple_prod_all, tuple_prod; intros ? ? x; destruct x as [n x]. revert x A. revert S. induction n as [|n ih]; simpl; intros; simpl; try simpl in ih; auto. Qed. Lemma tuple_prod_functional : forall {T n} {A B:T->Type} {w w':T^n}, pointwise_ind_eq A B w w' -> tuple_prod A w = tuple_prod B w'. intros T n; revert T. induction n as [|n ih]; simpl; auto. intros T A B w w' h1. pose proof (h1 0 (lt_0_Sn _)) as h2; simpl in h2. f_equal; auto. apply ih. intros i hi. specialize (h1 (S i) (lt_n_S _ _ hi)). simpl in h1. erewrite f_equal_dep. rewrite h1 at 1. reflexivity. reflexivity. Qed. Lemma tuple_prod_seg_functional : forall {n r s} (A:r @->Type) (B:s @ ->Type) (w: (r @)^n) (w':(s @)^n), pointwise_ind_eq A B w w' -> tuple_prod A w = tuple_prod B w'. intro n. induction n as [|n ih]; simpl; auto; intros r s A B w w' h1. f_equal. specialize (h1 0 (lt_0_Sn _)); simpl in h1; auto. specialize (ih r s A B (snd w) (snd w')). apply ih. intros i hi. specialize (h1 (S i) (lt_n_S _ _ hi)); simpl in h1. erewrite f_equal_dep. rewrite h1 at 1. repeat f_equal. apply proof_irrelevance. reflexivity. Qed. Lemma tuple_prod_ext : forall {T n} {A:T->Type} {w:T^n} (p q:tuple_prod A w), (forall i (pfi:i < n), nth_tuple_prod p i pfi = nth_tuple_prod q i pfi) -> p = q. intros T n. induction n as [|n ih]; simpl. intros; apply unit_eq. intros A w p q. destruct w as [w w'], p as [p p'], q as [q q']; simpl; simpl in w', p, q, p', q'. intro h1. pose proof (h1 0 (lt_0_Sn _)) as h2; simpl in h2; subst; f_equal. apply ih. intros i hi. specialize (h1 _ (lt_n_S _ _ hi)); simpl in h1. rewrite (proof_irrelevance _ hi (lt_S_n _ _ (lt_n_S _ _ hi))); auto. Qed. Definition Sfun {S} (A B:S -> Type) : Type := forall s, A s -> B s. Definition Sfun' {S U} (A B:S -> Ensemble U) : Type := forall s, fun_in_ens (A s) (sig_set (B s)). Definition Sfun_dep {S} {U V:S->Type} (A:forall s:S, Ensemble (U s)) (B:forall s:S, Ensemble (V s)) : Type := forall s, fun_in_ens (A s) (sig_set (B s)). End TupleProd. Fixpoint snth_prod_aux {U n} : setprod U (S n) -> nat -> U := match n with | 0 => fun x _ => fst x | n' => fun x i => match i with | 0 => fst x | S i' => snth_prod_aux (snd x) i' end end. Definition snth_prod {U n} : setprod U n -> (match n with | 0 => unit | S _ => nat -> U end) := match n with | 0 => fun _ => tt | S n' => fun x => snth_prod_aux x end. Fixpoint setprod_to_nprod {U n} : setprod U n -> U^n := match n with | 0 => fun _ => tt | S n' => fun x => let f := fst x in let s := snd x in let stn := setprod_to_nprod s in (f, stn) end. Lemma snth_prod_compat : forall {U n} (x:setprod U (S n)) i (pfi:i < S n), snth_prod x i = nth_prod i pfi (setprod_to_nprod x). intros U n. induction n as [|n ih]. intros x i h1. pose proof (lt1 _ h1); subst; simpl; auto. intros x i h1. destruct i as [|i]. simpl; auto. destruct i as [|i]. specialize (ih (snd x) 0 (lt_0_Sn _)). simpl in ih; simpl; auto. specialize (ih (snd x) (S i) (lt_S_n _ _ h1)). simpl in ih; simpl; auto. Qed. Lemma snth_prod_compat' : forall {U n} (x:setprod U (S n)) i (pfi:i >= S n), snth_prod x i = nth_prod n (lt_n_Sn _) (setprod_to_nprod x). intros U n. induction n as [|n ih]. intros; auto. intros x i h1. destruct i as [|i]. destruct n; omega. specialize (ih (snd x) i (le_S_n _ _ h1)). simpl in ih; simpl; auto. rewrite ih. f_equal. apply proof_irrelevance. Qed. End NProd. Definition ind_prod {T n} (x:T^n) (i:nat) : option T := match (lt_le_dec i n) with | left pflt => Some (nth_prod i pflt x) | right _ => None end. (*Chosen as [right associativity] as a dual to "exponentation."*) Notation "x '__' i" := (ind_prod x i) (at level 30, right associativity). Section RandomFacts. (*Wyckoff*) Lemma finite_seg_one_t : forall n:nat, FiniteT (seg_one_t n). intro n. pose proof (finite_seg_one n) as h1. apply Finite_ens_type in h1. assert (h2:forall m:seg_one_t n, In (seg_one n) (proj1_sig m)). intro m. destruct m as [m h3]. constructor. simpl. assumption. pose (fun i => (exist (fun m => In (seg_one n) m) _ (h2 i))) as f. assert (h3:Injective f). red. intros x y h4. destruct x as [x h5], y as [y h6]. unfold f in h4. simpl in h4. apply exist_injective in h4. apply proj1_sig_injective; auto. eapply inj_finite. apply h1. apply h3. intros; apply classic. Qed. (*Wyckoff*) Lemma iff_pred_sig_eq : forall {T:Type} (P Q:T->Prop), iff_pred P Q -> {x|P x}={x|Q x}. intros T P Q h1; destruct h1 as [h1 h2]; red in h1; do 2 red in h2. f_equal. apply functional_extensionality. intro x. specialize (h1 x); specialize (h2 x). apply prop_ext; tauto. Qed. End RandomFacts. (*Reserach standard library [eq_dep] to interface with [transfer]!*) Section Transfer. (* Wyckoff *) Definition transfer {T U:Type} (pf:T=U) (x:T) : U. revert x. rewrite pf. intro u. refine u. Defined. (*Wyckoff*) Definition transfer_r {T U:Type} (pf:T = U) (y:U) : T. revert y. rewrite pf. intro u. refine u. Defined. (*Wyckoff*) Lemma transfer_eq : forall {T U:Type} (pf:T=U) (x:T), transfer pf x = eq_rect _ id x _ pf. intros; subst; simpl; auto. Qed. (*Wyckoff*) Lemma transfer_r_eq : forall {T U:Type} (pf:T=U) (x:U), transfer_r pf x = eq_rect_r id x pf. intros; subst; simpl; auto. Qed. (*Wyckoff*) Lemma transfer_eq_refl : forall {T:Type} (x:T) (pfe:T = T), transfer pfe x = x. intros T x h1; pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. unfold transfer, eq_rec_r, eq_rec, eq_rect; simpl; auto. Defined. (*Wyckoff*) Lemma transfer_r_eq_refl : forall {T:Type} (x:T) (pfe:T = T), transfer_r pfe x = x. intros T x h1; pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. unfold transfer_r, eq_rec_r, eq_rec, eq_rect; simpl; auto. Defined. (*Wyckoff*) Lemma match_eq_refl_eq : forall {T} {U:T->Type} {x y} (pfxy:x = y) (pfeu:U x = U y) (ux:U x), match pfeu in (_ = y) return y with | eq_refl => ux end = transfer pfeu ux. intros; subst; rewrite transfer_eq_refl; pose proof (proof_irrelevance _ pfeu (eq_refl _)); subst; auto. Qed. (*Wyckoff*) Lemma transfer_transfer_r_compat : forall {T U:Type} (pf:T=U) (x:T), transfer pf x = transfer_r (eq_sym pf) x. intros; subst; rewrite transfer_eq_refl, transfer_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_r_transfer_compat : forall {T U:Type} (pf:T=U) (y:U), transfer_r pf y = transfer (eq_sym pf) y. intros; subst; rewrite transfer_eq_refl, transfer_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_undoes_transfer_r : forall {T U:Type} (pf:T=U) (y:U), transfer pf (transfer_r pf y) = y. intros; subst; rewrite transfer_eq_refl, transfer_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_r_undoes_transfer : forall {T U:Type} (pf:T=U) (x:T), transfer_r pf (transfer pf x) = x. intros; subst; rewrite transfer_eq_refl, transfer_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_undoes_transfer_r' : forall {T U:Type} (pf pf':T=U) (y:U), transfer pf (transfer_r pf' y) = y. intros T U h1 h2 y; pose proof (proof_irrelevance _ h1 h2); subst. rewrite transfer_eq_refl, transfer_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_r_undoes_transfer' : forall {T U:Type} (pf pf':T=U) (x:T), transfer_r pf (transfer pf' x) = x. intros T U h1 h2 y; pose proof (proof_irrelevance _ h1 h2); subst. rewrite transfer_eq_refl, transfer_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_inj : forall {T U:Type} (pf:T=U) (x y:T), transfer pf x = transfer pf y -> x = y. intros T U pf x y h1. subst. unfold transfer in h1. unfold eq_rect_r in h1. rewrite <- eq_rect_eq in h1. assumption. Qed. (*Wyckoff*) Lemma transfer_hetero_inj : forall {T U V:Type} (pft:T=V) (pfu:U=V) (x:T) (y:U), transfer pft x = transfer pfu y -> existT id _ x = existT id _ y. intros T U V h1 h2 x y. subst. intro h1. apply transfer_inj in h1. subst. reflexivity. Qed. (*Wyckoff*) Lemma transfer_r_inj : forall {T U:Type} (pf:T=U) (x y:U), transfer_r pf x = transfer_r pf y -> x = y. intros T U pf x y h1. subst. unfold transfer_r in h1. unfold eq_rect_r in h1. rewrite <- eq_rect_eq in h1. assumption. Qed. (*Wyckoff*) Lemma transfer_r_hetero_inj : forall {T U V:Type} (pft:V=T) (pfu:V=U) (x:T) (y:U), transfer_r pft x = transfer_r pfu y -> existT id _ x = existT id _ y. intros T U V h1 h2 x y. subst. subst. intro h1. apply transfer_r_inj in h1. subst. reflexivity. Qed. (*Wyckoff*) Lemma transfer_existT_iff : forall {T U:Type} (pf:T = U) (x:T) (y:U), transfer pf x = y <-> existT id _ x = existT id _ y. intros T U h1 x y; subst. rewrite transfer_eq_refl. split. intros; subst; auto. intro h1. apply inj_pair2 in h1; auto. Qed. (*Wyckoff*) Lemma transfer_r_existT_iff : forall {T U:Type} (pf:T = U) (x:T) (y:U), transfer_r pf y = x <-> existT id _ y = existT id _ x. intros T U h1 x y; subst. rewrite transfer_r_eq_refl. split. intros; subst; auto. intro h1. apply inj_pair2 in h1; auto. Qed. (*Wyckoff*) Lemma transfer_transfer_hetero : forall {T U V:Type} (pf:T = U) (pf':U = V) (a:T), transfer pf' (transfer pf a) = transfer (eq_trans pf pf') a. intros T U V h1 h2 a. subst. subst. unfold transfer. unfold transfer_r. unfold eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_transfer_r_hetero : forall {T U V:Type} (pf:T = U) (pf':T = V) (a:U), transfer pf' (transfer_r pf a) = transfer (eq2 pf pf') a. intros T U V h1 h2 a. subst. subst. unfold transfer. unfold transfer_r. unfold eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_r_transfer_hetero : forall {T U V:Type} (pf:T = U) (pf':V = U) (a:T), transfer_r pf' (transfer pf a) = transfer (eq2' pf pf') a. intros T U V h1 h2 a. subst. subst. unfold transfer. unfold transfer_r. unfold eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_r_transfer_r_hetero : forall {T U V:Type} (pf:T = U) (pf':V = T) (a:U), transfer_r pf' (transfer_r pf a) = transfer_r (eq_trans' pf pf') a. intros T U V h1 h2 a. subst. subst. unfold transfer. unfold transfer_r. unfold eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Definition transfer_pred {T} {P Q:T->Prop} (pfe:pred_eq P Q) (x:{t:T | P t}) : {t:T | Q t} := exist _ _ (iff1 (pfe (proj1_sig x)) (proj2_sig x)). (*Wyckoff*) Definition transfer_pred_r {T} {P Q:T->Prop} (pfe:pred_eq P Q) (x:{t:T | Q t}) : {t:T | P t} := exist _ _ (iff2 (pfe (proj1_sig x)) (proj2_sig x)). (*Wyckoff*) Lemma transfer_pred_undoes_transfer_pred_r : forall {T} {P Q:T->Prop} (pfe:pred_eq P Q) (x:{t:T | Q t}), transfer_pred pfe (transfer_pred_r pfe x) = x. unfold transfer_pred_r, transfer_pred; simpl; intros; apply proj1_sig_injective; simpl; auto. Qed. (*Wyckoff*) Lemma transfer_pred_r_undoes_transfer_pred : forall {T} {P Q:T->Prop} (pfe:pred_eq P Q) (x:{t:T | P t}), transfer_pred_r pfe (transfer_pred pfe x) = x. unfold transfer_pred_r, transfer_pred; simpl; intros; apply proj1_sig_injective; simpl; auto. Qed. (*Wyckoff*) Lemma transfer_pred_refl : forall {T} {P:T->Prop} (x:{t:T | P t}) (pf:pred_eq P P), transfer_pred pf x = x. unfold transfer_pred; intros; apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma transfer_pred_refl_r : forall {T} {P:T->Prop} (x:{t:T | P t}) (pf:pred_eq P P), transfer_pred_r pf x = x. unfold transfer_pred_r; intros; apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma transfer_pred_sigT : forall {T T':Type} (pf:T = T') (P:{U:Type & U} -> Prop), forall x:T, P (existT id T x) <-> P (existT id T' (transfer pf x)). intros T T' h1 P x. subst. unfold transfer. unfold eq_rect_r. simpl. tauto. Qed. (*Wyckoff*) Lemma transfer_fun_pred : forall {T T' U U':Type} (pft:T = T') (pfu:U = U') (f:T->U) (P:{pT:Type*Type & ((fst pT)->(snd pT))}->Prop), P (existT _ (T,U) f) <-> P (existT _ (T', U') (fun x:(fst (T', U')) => transfer (eq_refl (snd (T', U'))) (transfer pfu (f (transfer_r pft x))))). simpl. intros T T' U U' h1 h2 f P. subst. unfold transfer, transfer_r, eq_rect_r. simpl. tauto. Qed. (*Wyckoff*) Lemma transfer_fun_pred' : forall {T T' U U':Type} (pft:T' = T) (pfu:U = U') (f:T->U) (P:{pT:Type*Type & ((fst pT)->(snd pT))}->Prop), P (existT _ (T,U) f) <-> P (existT _ (T', U') (fun x:(fst (T', U')) => transfer (eq_refl (snd (T', U'))) (transfer pfu (f (transfer pft x))))). simpl. intros T T' U U' h1 h2 f P. subst. unfold transfer, eq_rect_r. simpl. tauto. Qed. (*Wyckoff*) Definition transfer_dep {T:Type} {U:T->Type} {x y:T} (pf:x = y) (a:U x):U y. subst. refine a. Defined. (*Wyckoff*) Definition transfer_dep_r {T:Type} {U:T->Type} {x y:T} (pf:y = x) (a:U x):U y. subst. refine a. Defined. (*Wyckoff*) Definition transfer_dep2 {T U:Type} {V:T->U->Type} {x x':T} {y y':U} (pfx:x = x') (pfy:y = y') (a:V x y) : V x' y'. subst. refine a. Defined. (*Wyckoff*) Definition transfer_dep2_r {T U:Type} {V:T->U->Type} {x x':T} {y y':U} (pfx:x = x') (pfy:y = y') (a:V x' y') : V x y. subst. refine a. Defined. (*Wyckoff*) Lemma transfer_dep_eq_refl : forall {T:Type} (U:T->Type) (x:T) (pfe:x = x) (p:U x), transfer_dep pfe p = p. intros ? ? ? h1 ?. pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. unfold transfer_dep, eq_rect_r, eq_rect; simpl; auto. Defined. (*Wyckoff*) Lemma transfer_dep_r_eq_refl : forall {T:Type} (U:T->Type) (x:T) (pfe:x = x) (p:U x), transfer_dep_r pfe p = p. intros ? ? ? h1 ?. pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. unfold transfer_dep_r, eq_rect_r, eq_rect; simpl; auto. Defined. (*Wyckoff*) Lemma transfer_dep2_eq_refl : forall {T U:Type} {V:T->U->Type} {x:T} {y:U} (pfx:x = x) (pfy:y = y) (a:V x y), transfer_dep2 pfx pfy a = a. unfold transfer_dep2, eq_rect_r, eq_rect. intros T U V x y hx hy a. pose proof (proof_irrelevance _ (eq_sym hy) (eq_refl _)) as h1. rewrite h1. pose proof (proof_irrelevance _ (eq_sym hx) (eq_refl _)) as h2. rewrite h2; auto. Qed. (*Wyckoff*) Lemma transfer_dep2_r_eq_refl : forall {T U:Type} {V:T->U->Type} {x:T} {y:U} (pfx:x = x) (pfy:y = y) (a:V x y), transfer_dep2_r pfx pfy a = a. unfold transfer_dep2_r, eq_rect_r, eq_rect. intros T U V x y hx hy a. pose proof (proof_irrelevance _ (eq_sym hy) (eq_refl _)) as h1. rewrite h1. pose proof (proof_irrelevance _ (eq_sym hx) (eq_refl _)) as h2. rewrite h2; auto. Qed. Lemma transfer_dep2_transfer_dep2_r_compat : forall {T U} {V:T->U->Type} {x x':T} {y y':U} (pfex:x = x') (pfey:y = y') (pfex':x' = x) (pfey':y' = y) (a:V x y), transfer_dep2 pfex pfey a = transfer_dep2_r pfex' pfey' a. intros T U V x x' y y' hx hy hx' hy' a. pose proof (proof_irrelevance _ hx' (eq_sym hx)) as h1. rewrite h1. pose proof (proof_irrelevance _ hy' (eq_sym hy)) as h2. rewrite h2. clear hx' hy' h1 h2. subst. rewrite transfer_dep2_eq_refl, transfer_dep2_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep_transfer_dep_r_compat : forall {T:Type} {U:T->Type} {x y:T} (pf:x = y) (pf':y = x) (a:U x), transfer_dep pf a = transfer_dep_r pf' a. intros T U x y h1 h2 a. unfold transfer_dep, transfer_dep_r. destruct h2. unfold eq_rect_r. simpl. pose proof (proof_irrelevance _ (eq_sym h1) (eq_refl _)) as h2. rewrite h2. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep_r_undoes_transfer_dep : forall {T:Type} {U:T->Type} {x y:T} (pf:x = y) (a:U x), transfer_dep_r pf (transfer_dep pf a) = a. intros; subst; rewrite transfer_dep_eq_refl, transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep_undoes_transfer_dep_r : forall {T:Type} {U:T->Type} {x y:T} (pf:x = y) (a:U y), transfer_dep pf (transfer_dep_r pf a) = a. intros; subst; rewrite transfer_dep_eq_refl, transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep_r_undoes_transfer_dep' : forall {T:Type} {U:T->Type} {x y:T} (pf pf':x = y) (a:U x), transfer_dep_r pf (transfer_dep pf' a) = a. intros T U x y h1 h2 a; pose proof (proof_irrelevance _ h1 h2); subst. rewrite transfer_dep_eq_refl, transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep_undoes_transfer_dep_r' : forall {T:Type} {U:T->Type} {x y:T} (pf pf':x = y) (a:U y), transfer_dep pf (transfer_dep_r pf' a) = a. intros T U x y h1 h2 a; pose proof (proof_irrelevance _ h1 h2); subst. rewrite transfer_dep_eq_refl, transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep2_undoes_transfer_dep2_r : forall {T U} {V:T->U->Type} {x x':T} {y y':U} (pfex pfex':x' = x) (pfey pfey':y' = y) (a:V x y), transfer_dep2 pfex pfey (transfer_dep2_r pfex' pfey' a) = a. intros; subst; rewrite transfer_dep2_eq_refl, transfer_dep2_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep2_r_undoes_transfer_dep2 : forall {T U} {V:T->U->Type} {x x':T} {y y':U} (pfex pfex':x' = x) (pfey pfey':y' = y) (a:V x y), transfer_dep2 pfex pfey (transfer_dep2_r pfex' pfey' a) = a. intros; subst; rewrite transfer_dep2_eq_refl, transfer_dep2_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep_trivial : forall {T U} {x y:T} (pfe:x = y) (u:U), transfer_dep (U := fun _ => U) pfe u = u. intros; subst; rewrite transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep_r_trivial : forall {T U} {x y:T} (pfe:x = y) (u:U), transfer_dep_r (U := fun _ => U) pfe u = u. intros; subst; rewrite transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep2_trivial1 : forall {T U} {V:U->Type} {x x' x0 x0':T} {y y':U} (pfex:x = x') (pfex0:x0 = x0') (pfey:y = y') (a:V y), transfer_dep2 (V := fun _ q => V q) pfex pfey a = transfer_dep2 (V := fun _ q => V q) pfex0 pfey a. intros; subst; auto. Qed. (*Wyckoff*) Lemma transfer_dep2_trivial2 : forall {T U} {V:T->Type} {x x':T} {y y' y0 y0':U} (pfex:x = x') (pfey:y = y') (pfey0:y0 = y0') (a:V x), transfer_dep2 (V := fun q _ => V q) pfex pfey a = transfer_dep2 (V := fun q _ => V q) pfex pfey0 a. intros; subst; auto. Qed. (*Wyckoff*) Lemma transfer_dep2_r_trivial1 : forall {T U} {V:U->Type} {x x' x0 x0':T} {y y':U} (pfex:x = x') (pfex0:x0 = x0') (pfey:y = y') (a:V y'), transfer_dep2_r (V := fun _ q => V q) pfex pfey a = transfer_dep2_r (V := fun _ q => V q) pfex0 pfey a. intros; subst; auto. Qed. (*Wyckoff*) Lemma transfer_dep2_r_trivial2 : forall {T U} {V:T->Type} {x x':T} {y y' y0 y0':U} (pfex:x = x') (pfey:y = y') (pfey0:y0 = y0') (a:V x'), transfer_dep2_r (V := fun q _ => V q) pfex pfey a = transfer_dep2_r (V := fun q _ => V q) pfex pfey0 a. intros; subst; auto. Qed. (*Wyckoff*) Lemma transfer_dep_inj : forall {T:Type} {U:T->Type} {x y:T} (pf:x = y) (a b:U x), transfer_dep pf a = transfer_dep pf b -> a = b. intros T U x y h1 a b h2. subst. unfold transfer_dep in h2. unfold eq_rect_r in h2. simpl in h2. assumption. Qed. (*Wyckoff*) Lemma transfer_dep_r_inj : forall {T:Type} {U:T->Type} {x y:T} (pf:x = y) (a b:U y), transfer_dep_r pf a = transfer_dep_r pf b -> a = b. intros T U x y h1 a b h2. subst. unfold transfer_dep_r, eq_rect_r in h2. simpl in h2. assumption. Qed. (*Wyckoff*) Lemma transfer_dep_existT : forall {T:Type} {U:T->Type} (x y:T) (pf:x = y) (a:U x), existT _ _ a = existT _ _ (transfer_dep pf a). intros T U x y h1 a. subst. f_equal. Qed. (*Wyckoff*) Lemma projT_injective : forall {T:Type} (A:T->Type) (x y:{a:T & (A a)}) (pf:projT1 x = projT1 y), (transfer_dep pf (projT2 x)) = projT2 y-> x = y. intros T A x y h1 h2. destruct x; destruct y. simpl in h1. simpl in h2. subst. apply transfer_dep_existT. Qed. (*Wyckoff*) Definition transfer_pred_dep {T} {P Q:T->Prop} {U:forall R:T->Prop, {x:T | R x} -> Type} {p:{x:T | P x}} (pfeq:pred_eq P Q) (u:U P p) : U Q (transfer_pred pfeq p). pose proof (pred_ext _ _ pfeq); subst. rewrite transfer_pred_refl; auto. Defined. (*Wyckoff*) Definition transfer_pred_dep_r {T} {P Q:T->Prop} {U:forall R:T->Prop, {x:T | R x} -> Type} {q:{x:T | Q x}} (pfeq:pred_eq P Q) (u:U Q q) : U P (transfer_pred_r pfeq q). pose proof (pred_ext _ _ pfeq); subst. rewrite transfer_pred_refl_r; auto. Defined. (*Wyckoff*) Lemma existT_injective1 : forall {T:Type} (U:T->Type) (x y:T) (p:U x) (q:U y), existT _ x p = existT _ y q -> x = y. intros T U x y p q h1. pose proof (f_equal (@projT1 _ _) h1) as h2. simpl in h2. assumption. Qed. (*Wyckoff*) Lemma existT_injective2 : forall {T:Type} (U:T->Type) (x:T) (p q:U x), existT _ x p = existT _ x q -> p = q. intros; apply inj_pair2; auto. Qed. (*Wyckoff*) Lemma existT_injective2' : forall {T:Type} (U:T->Type) (x x':T) (p:U x) (q:U x') (pfx:x = x'), existT _ x p = existT _ x' q -> transfer_dep pfx p = q. intros; subst; rewrite transfer_dep_eq_refl; apply existT_injective2 in H; auto. Qed. (*Wyckoff*) Lemma transfer_dep_type : forall {T:Type} {U:Type->Type} (y y':U T) (P:T->U T->Prop) (f:forall t:T, P t y -> Prop) (pfe:y = y') (x:T) (pfxy:P x y'), @transfer_dep (U T) (fun l:U T => forall t:T, P t l -> Prop) y y' pfe f x pfxy = f x (@transfer_dep (U T) (fun l:U T => P x l) y' y (eq_sym pfe) pfxy). intros; subst; rewrite transfer_dep_eq_refl; reflexivity. Qed. Definition transfer_dep_pf {T} {P:T->Prop} {U:forall x:T, P x -> Type} {x} (pf pf':P x) (u:U x pf) : U x pf' := transfer_dep (proof_irrelevance (P x) pf pf') u. Definition transfer_dep_pf_r {T} {P:T->Prop} {U:forall x:T, P x -> Type} {x} (pf pf':P x) (u:U x pf') : U x pf := transfer_dep_r (proof_irrelevance (P x) pf pf') u. Lemma transfer_dep_pf_eq : forall {T U} {P:T->Prop} (f:prop_dep_fun P U) x x' (pfe:x = x') (pfpx:P x) (pfpx':P x') (pfe':pfpx = transfer_dep_r pfe pfpx'), transfer_dep pfe' (f x pfpx) (U := fun pf:P x => U) = f x' (transfer_dep pfe pfpx). intros; subst; rewrite transfer_dep_eq_refl; auto. Qed. Lemma transfer_dep_r_pf_eq : forall {T U} {P:T->Prop} (f:prop_dep_fun P U) x x' (pfe:x = x') (pfpx:P x) (pfpx':P x') (pfe':transfer_dep pfe pfpx = pfpx'), transfer_dep_r pfe' (f x' pfpx') (U := fun pf:P x' => U) = f x (transfer_dep_r pfe pfpx'). intros; subst; rewrite transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep_prop : forall {T:Type} {U:T->Type} (x y:T) (pf:x = y) (a:U x) (P:{t:T & (U t)}->Prop), P (existT _ _ a) <-> P (existT _ _ (transfer_dep pf a)). intros T U x y h1 a P. rewrite (transfer_dep_existT _ _ h1 a). tauto. Qed. (*Wyckoff*) Definition transfer_dep_prop1 {T V:Type} {P:T->Prop} {U: forall x:T, P x -> Type} {x y:T} (pfe:x = y) (z:{pfpx:P x & U x pfpx}) : {pfpy:P y & U y pfpy} := let z' := transfer_dep pfe (projT1 z) in let u := transfer (f_equal_dep U _ _ (projT1 z) z' pfe) (projT2 z) in existT (fun pfpy:P y => U y pfpy) z' u. (*Wyckoff*) Lemma transfer_dep_eq_iff : forall {T:Type} {U:T->Type} (x y:T) (pf:x = y) (a:U x) (b:U y), existT _ _ a = existT _ _ b <-> transfer_dep pf a = b. intros T U x y h1 a b. subst. split. intro h1. apply inj_pair2 in h1. unfold transfer_dep. unfold eq_rect_r. rewrite <- eq_rect_eq. assumption. intro h1. unfold transfer_dep in h1. unfold eq_rect_r in h1. rewrite <- eq_rect_eq in h1. subst. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep_r_eq_iff : forall {T:Type} {U:T->Type} (x y:T) (pf:y = x) (a:U x) (b:U y), existT _ _ a = existT _ _ b <-> transfer_dep_r pf a = b. intros T U x y h1 a b. subst. split. intro h1. apply inj_pair2 in h1. unfold transfer_dep_r. unfold eq_rect_r. rewrite <- eq_rect_eq. assumption. intro h1. unfold transfer_dep_r in h1. unfold eq_rect_r in h1. rewrite <- eq_rect_eq in h1. subst. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep2_eq_iff : forall {T U} {V:T->U->Type} {x x':T} {y y':U} (pfe:x = x') (pfe':y = y') (a:V x y) (b:V x' y'), existT _ _ (existT _ _ a) = existT _ _ (existT _ _ b)<-> transfer_dep2 pfe pfe' a = b. intros; subst; rewrite transfer_dep2_eq_refl. split; intro h1. apply inj_pair2 in h1. apply inj_pair2 in h1; auto. subst; auto. Qed. (*Wyckoff*) Lemma transfer_dep2_r_eq_iff : forall {T U} {V:T->U->Type} {x x':T} {y y':U} (pfe:x' = x) (pfe':y' = y) (a:V x y) (b:V x' y'), existT _ _ (existT _ _ a) = existT _ _ (existT _ _ b)<-> transfer_dep2_r pfe pfe' a = b. intros; subst; rewrite transfer_dep2_r_eq_refl. split; intro h1. apply inj_pair2 in h1. apply inj_pair2 in h1; auto. subst; auto. Qed. (*Wyckoff*) Definition transfer_prop_app2 {T} {P:T->Prop} {f:forall x:T, P x -> Type} {x x'} (pfe:x = x') (pfx:P x) (y:f x pfx) : f x' (transfer_dep pfe pfx). subst; rewrite transfer_dep_eq_refl; auto. Defined. (*Wyckoff*) Lemma existT_inv : forall {T:Type} {U:T->Type} (x y:T) (pf:x = y) (f:U x) (g:U y), @existT T U x f = @existT T U y g -> f = transfer_dep_r pf g. intros T U x y h1 f g h2. symmetry in h2. rewrite transfer_dep_r_eq_iff in h2. symmetry. apply h2. Qed. (*Wyckoff*) Lemma existT_s_inv : forall {T:Set} {V:Type} {U:T->Type} (x y:T) (pf:x = y) (f:U x) (g:U y), @existT T U x f = @existT T U y g -> f = transfer_dep_r pf g. intros T V U x y h1 f g h2. symmetry in h2. rewrite transfer_dep_r_eq_iff in h2. symmetry. apply h2. Qed. (*Wyckoff*) Lemma transfer_dep_fun1_eq : forall {T U:Type} (pf:T = U) (f:T->T) (x:U), (transfer_dep (U:=fun x=>x->x) pf f) x = (transfer pf (f (transfer_r pf x))). intros T U h1 f x. subst. rewrite transfer_dep_eq_refl. unfold transfer, transfer_r, eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep_fun2_eq : forall {T U:Type} (pf:T = U) (f:T->T->T) (x y:U), (transfer_dep (U:=fun x=>x->x->x) pf f) x y = (transfer pf (f (transfer_r pf x) (transfer_r pf y))). intros T U h1 f x y. subst. rewrite transfer_dep_eq_refl. unfold transfer, transfer_r, eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep_hetero : forall {T:Type} {U:T->Type} (x x' y:T) (pf:x = y) (pf':x' = y) (a:U x) (a':U x'), transfer_dep pf a = transfer_dep pf' a' <-> a = transfer_dep (eq_trans pf' (eq_sym pf)) a'. intros; subst; do 3 rewrite transfer_dep_eq_refl; tauto. Qed. (*Wyckoff*) Lemma transfer_dep_r_fun1_eq : forall {T U:Type} (pf:T = U) (f:U->U) (x:T), (transfer_dep_r (U:=fun x=>x->x) pf f) x = (transfer_r pf (f (transfer pf x))). intros T U h1 f x. subst. rewrite transfer_dep_r_eq_refl. unfold transfer, transfer_r, eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep_r_fun2_eq : forall {T U:Type} (pf:T = U) (f:U->U->U) (x y:T), (transfer_dep_r (U:=fun x=>x->x->x) pf f) x y = (transfer_r pf (f (transfer pf x) (transfer pf y))). intros T U h1 f x y. subst. rewrite transfer_dep_r_eq_refl. unfold transfer, transfer_r, eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep_r_hetero : forall {T:Type} {U:T->Type} (x y y':T) (pf:x = y) (pf':x = y') (a:U y) (a':U y'), transfer_dep_r pf a = transfer_dep_r pf' a' <-> a = transfer_dep_r (eq2 pf pf') a'. intros T U x y y' h1 h2 a a'. subst. subst. do 2 rewrite transfer_dep_r_eq_refl. unfold eq2. unfold eq_ind_r. unfold eq_ind. simpl. rewrite transfer_dep_r_eq_refl. tauto. Qed. (*Wyckoff*) Definition transfer_fun {T U V:Type} (pf:T=U) (f:T->V) : U->V. revert f. rewrite pf. intro f. refine f. Defined. (*Wyckoff*) Definition transfer_fun' {T U:Type} (pf:T=U) (f:T->T) : U->U. revert f. rewrite pf. intro f. refine f. Defined. (*Wyckoff*) Definition transfer_fun2 {T U V:Type} (pf:T=U) (f:T->T->V) : U->U->V. revert f. rewrite pf. intro f. refine f. Defined. (*Wyckoff*) Definition transfer_fun2' {T U:Type} (pf:T=U) (f:T->T->T) : U->U->U. revert f. rewrite pf. intro f. refine f. Defined. (*Wyckoff*) Definition transfer_fun_r {T U V:Type} (pf:T=U) (f:U->V) : T->V. revert f. rewrite pf. intro f. refine f. Defined. (*Wyckoff*) Definition transfer_fun_r' {T U:Type} (pf:T=U) (f:U->U) : T->T. revert f. rewrite pf. intro f. refine f. Defined. (*Wyckoff*) Definition transfer_fun2_r {T U V:Type} (pf:T=U) (f:U->U->V) : T->T->V. revert f. rewrite pf. intro f. refine f. Defined. (*Wyckoff*) Definition transfer_fun2_r' {T U:Type} (pf:T=U) (f:U->U->U) : T->T->T. revert f. rewrite pf. intro f. refine f. Defined. (*Wyckoff*) Lemma transfer_fun_eq_refl : forall {T U:Type} (f:T->U) (pfe:T = T), transfer_fun pfe f = f. intros. apply functional_extensionality. intro. unfold transfer_fun, eq_rect_r, eq_rect; simpl. pose proof (proof_irrelevance _ pfe (eq_refl _)) as h1. rewrite h1; simpl; auto. Defined. (*Wyckoff*) Lemma transfer_fun2_eq_refl : forall {T U:Type} (f:T->T->U) (pfe:T = T), transfer_fun2 pfe f = f. intros. apply functional_extensionality. intro. apply functional_extensionality. intro. unfold transfer_fun2, eq_rect_r, eq_rec, eq_rect. simpl. pose proof (proof_irrelevance _ pfe (eq_refl _)) as h1. rewrite h1; simpl; auto. Defined. (*Wyckoff*) Lemma transfer_fun_eq_refl' : forall {T:Type} (f:T->T), transfer_fun' eq_refl f = f. intros. apply functional_extensionality. intro. unfold transfer_fun'. unfold eq_rect_r. simpl. reflexivity. Defined. (*Wyckoff*) Lemma transfer_fun2_eq_refl' : forall {T:Type} (f:T->T->T), transfer_fun2' eq_refl f = f. intros. apply functional_extensionality. intro. apply functional_extensionality. intro. unfold transfer_fun2'. unfold eq_rect_r. simpl. reflexivity. Defined. (*Wyckoff*) Lemma transfer_fun_r_eq_refl : forall {T U:Type} (f:T->U), transfer_fun_r eq_refl f = f. intros. apply functional_extensionality. intro. unfold transfer_fun. unfold eq_rect_r. simpl. reflexivity. Defined. (*Wyckoff*) Lemma transfer_fun2_r_eq_refl : forall {T U:Type} (f:T->T->U), transfer_fun2_r eq_refl f = f. intros. apply functional_extensionality. intro. apply functional_extensionality. intro. unfold transfer_fun2. unfold eq_rect_r. simpl. reflexivity. Defined. (*Wyckoff*) Lemma transfer_fun_r_eq_refl' : forall {T:Type} (f:T->T), transfer_fun_r' eq_refl f = f. intros. apply functional_extensionality. intro. unfold transfer_fun_r'. unfold eq_rect_r. simpl. reflexivity. Defined. (*Wyckoff*) Lemma transfer_fun2_r_eq_refl' : forall {T:Type} (f:T->T->T), transfer_fun2_r' eq_refl f = f. intros. apply functional_extensionality. intro. apply functional_extensionality. intro. unfold transfer_fun2_r'. unfold eq_rect_r. simpl. reflexivity. Defined. (*Wyckoff*) Lemma transfer_fun_eq : forall {T U V:Type} (pf:T = U) (f:T->V), transfer_fun pf f = (fun x:U => f (transfer_r pf x)). intros T U V h1 f. subst. rewrite transfer_fun_eq_refl. apply functional_extensionality. intro x. rewrite transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_eq' : forall {T U:Type} (pf:T=U) (f:T->T), transfer_fun' pf f = (fun x:U => (transfer pf (f (transfer_r pf x)))). intros T U h1 f. subst. rewrite transfer_fun_eq_refl'. apply functional_extensionality. intro x. rewrite transfer_eq_refl. rewrite transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_eq : forall {T U V:Type} (pf:T = U) (f:T->T->V), transfer_fun2 pf f = (fun (x:U) (y:U) => f (transfer_r pf x) (transfer_r pf y)). intros T U V h1 f. subst. rewrite transfer_fun2_eq_refl. apply functional_extensionality. intro x. rewrite transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_eq' : forall {T U:Type} (pf:T=U) (f:T->T->T), transfer_fun2' pf f = (fun (x:U) (y:U) => (transfer pf (f (transfer_r pf x) (transfer_r pf y)))). intros T U h1 f. subst. rewrite transfer_fun2_eq_refl'. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite transfer_eq_refl. do 2 rewrite transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_r_eq : forall {T U V:Type} (pf:T = U) (f:U->V), transfer_fun_r pf f = (fun x:T => f (transfer pf x)). intros T U V h1 f. subst. rewrite transfer_fun_r_eq_refl. apply functional_extensionality. intro x. rewrite transfer_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_r_eq' : forall {T U:Type} (pf:T=U) (f:U->U), transfer_fun_r' pf f = (fun x:T => (transfer_r pf (f (transfer pf x)))). intros T U h1 f. subst. rewrite transfer_fun_r_eq_refl'. apply functional_extensionality. intro x. rewrite transfer_eq_refl. rewrite transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_r_eq : forall {T U V:Type} (pf:T = U) (f:U->U->V), transfer_fun2_r pf f = (fun (x:T) (y:T) => f (transfer pf x) (transfer pf y)). intros T U V h1 f. subst. rewrite transfer_fun2_r_eq_refl. apply functional_extensionality. intro x. rewrite transfer_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_r_eq' : forall {T U:Type} (pf:T=U) (f:U->U->U), transfer_fun2_r' pf f = (fun (x:T) (y:T) => (transfer_r pf (f (transfer pf x) (transfer pf y)))). intros T U h1 f. subst. rewrite transfer_fun2_r_eq_refl'. apply functional_extensionality. intro x. apply functional_extensionality. intro y. do 2 rewrite transfer_eq_refl. rewrite transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_compat : forall {T U V:Type} (pf:T = U) (f:T->V) (x:T), (transfer_fun pf f) (transfer pf x) = f x. intros T U V h1 f x. unfold transfer_fun. unfold transfer. destruct h1. compute. reflexivity. Qed. Lemma transfer_fun_inj_iff : forall {T U V:Type} (pf: T = U) (f:T->V), Injective f <-> Injective (transfer_fun pf f). intros T U V h1 f. subst. rewrite transfer_fun_eq_refl. tauto. Qed. (*Wyckoff*) Lemma transfer_fun2_compat : forall {T U V:Type} (pf:T = U) (f:T->T->V) (x y:T), (transfer_fun2 pf f) (transfer pf x) (transfer pf y) = f x y. intros T U V h1 f x y. unfold transfer_fun2. unfold transfer, eq_rect_r. subst. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_compat' : forall {T U:Type} (pf:T = U) (f:T->T) (x:T), (transfer_fun' pf f) (transfer pf x) = transfer pf (f x). intros T U h1 f x. subst. unfold transfer_fun'. unfold transfer. unfold eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_compat' : forall {T U:Type} (pf:T = U) (f:T->T->T) (x y:T), (transfer_fun2' pf f) (transfer pf x) (transfer pf y) = transfer pf (f x y). intros T U h1 f x y. unfold transfer_fun2'. unfold transfer, eq_rect_r. subst. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_r_compat : forall {T U V:Type} (pf:T = U) (f:U->V) (x:U), f x = (transfer_fun_r pf f) (transfer_r pf x). intros T U V h1 f x. unfold transfer_fun_r. unfold transfer_r. destruct h1. compute. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_r_compat : forall {T U V:Type} (pf:T = U) (f:U->U->V) (x y:U), f x y = (transfer_fun2_r pf f) (transfer_r pf x) (transfer_r pf y). intros T U V h1 f x y. unfold transfer_fun2_r. unfold transfer, eq_rect_r. subst. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_r_compat' : forall {T U:Type} (pf:T = U) (f:U->U) (x:U), (transfer_fun_r' pf f) (transfer_r pf x) = transfer_r pf (f x). intros T U h1 f x. unfold transfer_fun_r'. unfold transfer_r. destruct h1. compute. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_r_compat' : forall {T U:Type} (pf:T = U) (f:U->U->U) (x y:U), (transfer_fun2_r' pf f) (transfer_r pf x) (transfer_r pf y) = transfer_r pf (f x y). intros T U h1 f x y. unfold transfer_fun2_r'. unfold transfer_r, eq_rect_r. subst. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_transfer_dep_compat : forall {T U V:Type} (pf:T = U) (f:T->V), transfer_fun pf f = transfer_dep pf (U:=fun A:Type=>A->V) f. intros T U V h1 f. subst. simpl. rewrite transfer_fun_eq_refl, transfer_dep_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_transfer_dep_compat : forall {T U V:Type} (pf:T = U) (f:T->T->V), transfer_fun2 pf f = transfer_dep pf (U:=fun A:Type=>A->A->V) f. intros T U V h1 f. subst. simpl. rewrite transfer_fun2_eq_refl, transfer_dep_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_r_transfer_dep_r_compat : forall {T U V:Type} (pf:T = U) (f:U->V), transfer_fun_r pf f = transfer_dep_r pf (U:=fun A:Type=>A->V) f. intros T U V h1 f. subst. simpl. rewrite transfer_fun_r_eq_refl, transfer_dep_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_r_transfer_dep_r_compat : forall {T U V:Type} (pf:T = U) (f:U->U->V), transfer_fun2_r pf f = transfer_dep_r pf (U:=fun A:Type=>A->A->V) f. intros T U V h1 f. subst. simpl. rewrite transfer_fun2_r_eq_refl, transfer_dep_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_transfer_dep_compat' : forall {T U:Type} (pf:T = U) (f:T->T), transfer_fun' pf f = transfer_dep pf (U:=fun A:Type=>A->A) f. intros T U h1 f. subst. simpl. rewrite transfer_fun_eq_refl', transfer_dep_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_transfer_dep_compat' : forall {T U:Type} (pf:T = U) (f:T->T->T), transfer_fun2' pf f = transfer_dep pf (U:=fun A:Type=>A->A->A) f. intros T U h1 f. subst. simpl. rewrite transfer_fun2_eq_refl', transfer_dep_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_r_transfer_dep_r_compat' : forall {T U:Type} (pf:T = U) (f:U->U), transfer_fun_r' pf f = transfer_dep_r pf (U:=fun A:Type=>A->A) f. intros T U h1 f. subst. simpl. rewrite transfer_fun_r_eq_refl', transfer_dep_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_r_transfer_dep_r_compat' : forall {T U:Type} (pf:T = U) (f:U->U->U), transfer_fun2_r' pf f = transfer_dep_r pf (U:=fun A:Type=>A->A->A) f. intros T U h1 f. subst. simpl. rewrite transfer_fun2_r_eq_refl', transfer_dep_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_transfer_compat : forall {T U V:Type} (pf:T = U) (f:T->V), transfer_fun pf f = transfer (f_equal (fun A=>A->V) pf) f. intros T U V h1 f. subst. rewrite transfer_fun_eq_refl. simpl. rewrite transfer_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_transfer_compat' : forall {T U:Type} (pf:T = U) (f:T->T), transfer_fun' pf f = transfer (f_equal (fun A=>A->A) pf) f. intros T U h1 f. subst. rewrite transfer_fun_eq_refl'. simpl. rewrite transfer_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_r_transfer_r_compat : forall {T U V:Type} (pf:T = U) (f:U->V), transfer_fun_r pf f = transfer_r (f_equal (fun A=>A->V) pf) f. intros T U V h1 f. subst. rewrite transfer_fun_r_eq_refl. simpl. rewrite transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun_r_transfer_r_compat' : forall {T U:Type} (pf:T = U) (f:U->U), transfer_fun_r' pf f = transfer_r (f_equal (fun A=>A->A) pf) f. intros T U h1 f. subst. rewrite transfer_fun_r_eq_refl'. simpl. rewrite transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_transfer_compat : forall {T U V:Type} (pf:T = U) (f:T->T->V), transfer_fun2 pf f = transfer (f_equal (fun A=>A->A->V) pf) f. intros T U V h1 f. subst. rewrite transfer_fun2_eq_refl. simpl. rewrite transfer_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_transfer_compat' : forall {T U:Type} (pf:T = U) (f:T->T->T), transfer_fun2' pf f = transfer (f_equal (fun A=>A->A->A) pf) f. intros T U h1 f. subst. rewrite transfer_fun2_eq_refl'. simpl. rewrite transfer_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_r_transfer_r_compat : forall {T U V:Type} (pf:T = U) (f:U->U->V), transfer_fun2_r pf f = transfer_r (f_equal (fun A=>A->A->V) pf) f. intros T U V h1 f. subst. rewrite transfer_fun2_r_eq_refl. simpl. rewrite transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_fun2_r_transfer_r_compat' : forall {T U:Type} (pf:T = U) (f:U->U->U), transfer_fun2_r' pf f = transfer_r (f_equal (fun A=>A->A->A) pf) f. intros T U h1 f. subst. rewrite transfer_fun2_r_eq_refl'. simpl. rewrite transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) (*[transfer_fun_sig], is useful when the type equality in transfer_fun (applied to a sigma function) is too coarse, in that you can't extract more fundamental information from it, like equality of underlying sets in the sigma type.*) Definition transfer_fun_sig {T U:Type} {A B:Ensemble T} (pf:A = B) (f:sig_set A->U) : sig_set B->U. subst. refine f. Defined. (*Wyckoff*) Lemma transfer_fun_sig_eq_refl : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), transfer_fun_sig eq_refl f = f. intros T U A f. apply functional_extensionality. intro x. unfold transfer_fun_sig, eq_rect_r, eq_rect. simpl. reflexivity. Defined. (*Wyckoff*) Lemma transfer_fun_sig_compat : forall {T U:Type} {A B:Ensemble T} (pf : A = B) (f:sig_set A->U) (x:sig_set B), transfer_fun_sig pf f x = f (transfer_dep_r pf x). intros T U A B h1 f x. subst. rewrite transfer_dep_r_eq_refl, transfer_fun_sig_eq_refl. reflexivity. Qed. (*Wyckoff*) Definition transfer_fun_sig_r {T U:Type} {A B:Ensemble T} (pf : A = B) (f:sig_set B->U) : sig_set A->U. subst. refine f. Defined. (*Wyckoff*) Lemma transfer_fun_sig_r_eq_refl : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), transfer_fun_sig_r eq_refl f = f. intros T U A f. apply functional_extensionality. intro x. unfold transfer_fun_sig_r, eq_rect_r, eq_rect. simpl. reflexivity. Defined. (*Wyckoff*) Lemma transfer_fun_sig_r_compat : forall {T U:Type} {A B:Ensemble T} (pf : A = B) (f:sig_set B->U) (x:sig_set A), transfer_fun_sig_r pf f x = f (transfer_dep pf x). intros T U A B h1 f x. subst. rewrite transfer_dep_eq_refl, transfer_fun_sig_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_prop : forall {T U:Type} (pf:T=U) (f:Type->Prop), f T <-> f U. intros T U h1. subst. tauto. Qed. (*Wyckoff*) Lemma transfer_prop_dependent : forall {T:Type} {PT:T->Type} (x y:T) (P:PT x->Prop) (pf:PT x = PT y) (a:PT x), P a <-> (transfer_fun pf P) (transfer pf a). intros T PT x y P h1 a. split. intro h2. rewrite <- (transfer_fun_compat h1 P a) in h2. assumption. intro h2. rewrite <- (transfer_fun_compat h1 P a). assumption. Qed. (*Wyckoff*) Lemma transfer_prop' {T U:Type} (pf:T= U) (P:T->Prop) : U->Prop. subst. refine P. Defined. (*Wyckoff*) Definition transfer_sig_set {T:Type} (A B:Ensemble T) (pf:sig_set A = sig_set B) (x:sig_set A) : sig_set B. rewrite <- pf. refine x. Defined. (*Wyckoff*) Definition transfer_sig_set_r {T:Type} (A B:Ensemble T) (pf:sig_set A = sig_set B) (x:sig_set B) : sig_set A. rewrite pf. refine x. Defined. (*Wyckoff*) Lemma transfer_dep_sig_set_eq : forall {T:Type} (A B:Ensemble T) (pfeq:A = B) (x:T) (pfin:Inn A x), transfer_dep pfeq (exist _ _ pfin) = exist _ _ (in_eq_set _ _ pfeq x pfin). intros T A B h1 x h2. subst. rewrite transfer_dep_eq_refl. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma transfer_dep_r_sig_set_eq : forall {T:Type} (A B:Ensemble T) (pfeq:A = B) (x:T) (pfin:Inn B x), transfer_dep_r pfeq (exist _ _ pfin) = exist _ _ (in_eq_set _ _ (eq_sym pfeq) x pfin). intros T A B h1 x h2. subst. rewrite transfer_dep_r_eq_refl. apply proj1_sig_injective; auto. Qed. (*Wyckoff*) Lemma f_equal_nprod_fst_fun : forall {T n} {U:T->Type} {w:T} {w':T^n} {x:T^(S n)} (pfe:x = (w, w')) (p:U (fst x)), transfer_dep (U := fun c => U (fst c)) pfe p = transfer_dep (U := U) (f_equal (@fst _ _) pfe) p. intros; subst; rewrite transfer_dep_eq_refl; simpl; rewrite transfer_dep_eq_refl; auto. Qed. (*These are [Set]-analogues of much of the above.*) Section SetSortTransfer. (*Wyckoff*) Definition s_transfer {T U:Set} (pf:T=U) (x:T) : U. revert x. rewrite pf. intro u. refine u. Defined. (*Wyckoff*) Definition s_transfer_r {T U:Set} (pf:T = U) (y:U) : T. revert y. rewrite pf. intro u. refine u. Defined. (*Wyckoff*) Lemma s_transfer_eq : forall {T U:Set} (pf:T=U) (x:T), s_transfer pf x = eq_rec _ id x _ pf. intros; subst; simpl; auto. Qed. (*Wyckoff*) Lemma s_transfer_eq_refl : forall {T:Set} (x:T) (pfe:T = T), s_transfer pfe x = x. intros T x h1; pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. unfold s_transfer, eq_rec_r, eq_rec, eq_rect; simpl; auto. Defined. (*Wyckoff*) Lemma s_transfer_r_eq_refl : forall {T:Set} (x:T) (pfe:T = T), s_transfer_r pfe x = x. intros T x h1; pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. unfold s_transfer_r, eq_rec_r, eq_rec, eq_rect; simpl; auto. Defined. (*Wyckoff*) Lemma s_transfer_r_eq : forall {T U:Set} (pf:T=U) (x:U), s_transfer_r pf x = eq_rec_r id x pf. intros; subst; simpl; auto. Qed. (*Wyckoff*) Lemma s_transfer_transfer_r_compat : forall {T U:Set} (pf:T=U) (pf':U=T) (x:T), s_transfer pf x = s_transfer_r pf' x. intros T U ? ? x. unfold s_transfer. unfold s_transfer_r. destruct pf. pose proof (proof_irrelevance _ pf' (eq_refl _)); subst. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_undoes_transfer_r : forall {T U:Set} (pf:T=U) (y:U), s_transfer pf (s_transfer_r pf y) = y. intros T U h1 y. subst. unfold s_transfer, s_transfer_r. simpl. unfold eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_r_undoes_transfer : forall {T U:Set} (pf:T=U) (x:T), s_transfer_r pf (s_transfer pf x) = x. intros T U h1 y. subst. unfold s_transfer, s_transfer_r. simpl. unfold eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_inj : forall {T U:Set} (pf:T=U) (x y:T), s_transfer pf x = s_transfer pf y -> x = y. intros T U pf x y h1. subst. unfold s_transfer in h1. unfold eq_rec_r in h1. rewrite <- eq_rec_eq in h1. assumption. Qed. (*Wyckoff*) Lemma s_transfer_hetero_inj : forall {T U V:Set} (pft:T=V) (pfu:U=V) (x:T) (y:U), s_transfer pft x = s_transfer pfu y -> existT id _ x = existT id _ y. intros T U V h1 h2 x y. subst. intro h1. apply s_transfer_inj in h1. subst. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_r_inj : forall {T U:Set} (pf:T=U) (x y:U), s_transfer_r pf x = s_transfer_r pf y -> x = y. intros T U pf x y h1. subst. unfold s_transfer_r in h1. unfold eq_rec_r in h1. rewrite <- eq_rec_eq in h1. assumption. Qed. (*Wyckoff*) Lemma s_transfer_r_hetero_inj : forall {T U V:Set} (pft:V=T) (pfu:V=U) (x:T) (y:U), s_transfer_r pft x = s_transfer_r pfu y -> existT id _ x = existT id _ y. intros T U V h1 h2 x y. subst. subst. intro h1. apply s_transfer_r_inj in h1. subst. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_transfer_hetero : forall {T U V:Set} (pf:T = U) (pf':U = V) (a:T), s_transfer pf' (s_transfer pf a) = s_transfer (eq_trans pf pf') a. intros T U V h1 h2 a. subst. subst. unfold s_transfer. unfold s_transfer_r. unfold eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_transfer_r_hetero : forall {T U V:Set} (pf:T = U) (pf':T = V) (a:U), s_transfer pf' (s_transfer_r pf a) = s_transfer (eq2 pf pf') a. intros T U V h1 h2 a. subst. subst. unfold s_transfer. unfold s_transfer_r. unfold eq_rec_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_r_transfer_hetero : forall {T U V:Set} (pf:T = U) (pf':V = U) (a:T), s_transfer_r pf' (s_transfer pf a) = s_transfer (eq2' pf pf') a. intros T U V h1 h2 a. subst. subst. unfold s_transfer. unfold s_transfer_r. unfold eq_rec_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_r_transfer_r_hetero : forall {T U V:Set} (pf:T = U) (pf':V = T) (a:U), s_transfer_r pf' (s_transfer_r pf a) = s_transfer_r (eq_trans' pf pf') a. intros T U V h1 h2 a. subst. subst. unfold s_transfer. unfold s_transfer_r. unfold eq_rec_r. simpl. reflexivity. Qed. (*Wyckoff*) Definition s_transfer_dep {T:Set} {U:T->Set} {x y:T} (pf:x = y) (a:U x):U y. subst. refine a. Defined. (*Wyckoff*) Definition s_transfer_dep_r {T:Set} {U:T->Set} {x y:T} (pf:y = x) (a:U x):U y. subst. refine a. Defined. (*Wyckoff*) Lemma s_transfer_dep_transfer_dep_r_compat : forall {T:Set} {U:T->Set} {x y:T} (pf:x = y) (pf':y = x) (a:U x), s_transfer_dep pf a = s_transfer_dep_r pf' a. intros T U x y h1 h2 a. unfold s_transfer_dep, s_transfer_dep_r. destruct h2. unfold eq_rec_r. simpl. pose proof (proof_irrelevance _ (eq_sym h1) (eq_refl _)) as h2. rewrite h2. simpl. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_dep_r_undoes_transfer_dep : forall {T:Set} {U:T->Set} {x y:T} (pf:x = y) (a:U x), s_transfer_dep_r pf (s_transfer_dep pf a) = a. intros T U x y h1 a. subst. unfold transfer_dep_r, transfer_dep, eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_dep_undoes_transfer_dep_r : forall {T:Set} {U:T->Set} {x y:T} (pf:x = y) (a:U y), s_transfer_dep pf (s_transfer_dep_r pf a) = a. intros T U x y h1 a. subst. unfold s_transfer_dep_r, s_transfer_dep, eq_rect_r. simpl. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_dep_inj : forall {T:Set} {U:T->Set} {x y:T} (pf:x = y) (a b:U x), s_transfer_dep pf a = s_transfer_dep pf b -> a = b. intros T U x y h1 a b h2. subst. unfold transfer_dep in h2. unfold eq_rect_r in h2. simpl in h2. assumption. Qed. (*Wyckoff*) Lemma s_transfer_dep_r_inj : forall {T:Set} {U:T->Set} {x y:T} (pf:x = y) (a b:U y), s_transfer_dep_r pf a = s_transfer_dep_r pf b -> a = b. intros T U x y h1 a b h2. subst. unfold s_transfer_dep_r, eq_rect_r in h2. simpl in h2. assumption. Qed. (*Wyckoff*) Lemma s_transfer_dep_existT : forall {T:Set} {U:T->Set} (x y:T) (pf:x = y) (a:U x), existT _ _ a = existT _ _ (transfer_dep pf a). intros T U x y h1 a. subst. f_equal. Qed. (*Wyckoff*) Lemma s_projT_injective : forall {T:Set} (A:T->Set) (x y:{a:T & (A a)}) (pf:projT1 x = projT1 y), s_transfer_dep pf (projT2 x) = projT2 y -> x = y. intros T A x y h1 h2. destruct x; destruct y. simpl in h1. simpl in h2. subst. apply s_transfer_dep_existT. Qed. (*Wyckoff*) Lemma s_transfer_dep_eq_refl : forall {T:Set} (U:T->Set) (x:T) (p:U x) (pfe:x = x), transfer_dep pfe p = p. intros ? ? ? ? h1. pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. unfold s_transfer_dep, eq_rect_r, eq_rect; simpl; auto. Defined. (*Wyckoff*) Lemma s_transfer_dep_r_eq_refl : forall {T:Set} (U:T->Set) (x:T) (p:U x) (pfe:x = x), s_transfer_dep_r pfe p = p. intros ? ? ? ? h1. pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. unfold s_transfer_dep_r, eq_rect_r, eq_rect; simpl; auto. Defined. (*Wyckoff*) Lemma s_existT_injective1 : forall {T:Set} (U:T->Set) (x y:T) (p:U x) (q:U y), existT _ x p = existT _ y q -> x = y. intros T U x y p q h1. pose proof (f_equal (@projT1 _ _) h1) as h2. simpl in h2. assumption. Qed. (*Wyckoff*) Lemma s_transfer_dep_prop : forall {T:Set} {U:T->Set} (x y:T) (pf:x = y) (a:U x) (P:{t:T & (U t)}->Prop), P (existT _ _ a) <-> P (existT _ _ (transfer_dep pf a)). intros T U x y h1 a P. rewrite (s_transfer_dep_existT _ _ h1 a). tauto. Qed. (*Wyckoff*) Lemma s_transfer_dep_eq_iff : forall {T:Set} {U:T->Set} (x y:T) (pf:x = y) (a:U x) (b:U y), existT _ _ a = existT _ _ b <-> s_transfer_dep pf a = b. intros T U x y h1 a b. subst. split. intro h1. apply inj_pair2 in h1. unfold s_transfer_dep. unfold eq_rec_r. rewrite <- eq_rect_eq. assumption. intro h1. unfold s_transfer_dep in h1. unfold eq_rec_r in h1. rewrite <- eq_rec_eq in h1. subst. reflexivity. Qed. (*Wyckoff*) Lemma s_transfer_dep_r_eq_iff : forall {T:Set} {U:T->Set} (x y:T) (pf:y = x) (a:U x) (b:U y), existT _ _ a = existT _ _ b <-> s_transfer_dep_r pf a = b. intros T U x y h1 a b. subst. split. intro h1. apply inj_pair2 in h1. unfold s_transfer_dep_r. unfold eq_rec_r. rewrite <- eq_rec_eq. assumption. intro h1. unfold s_transfer_dep_r in h1. unfold eq_rec_r in h1. rewrite <- eq_rec_eq in h1. subst. reflexivity. Qed. End SetSortTransfer. (*Wyckoff*) Lemma transfer_in : forall {T T':Type} (pf:T = T') (A:Ensemble T) (x:T), In A x <-> In (transfer_dep pf A) (transfer pf x). intros T T' h1 A x. subst. unfold transfer_dep. unfold eq_rect_r. simpl. unfold transfer. unfold eq_rect_r. simpl. tauto. Qed. (*Wyckoff*) Lemma transfer_in' : forall {T T':Type} (pf:T = T') (A:Ensemble T') (x:T), In A (transfer pf x) <-> In (transfer_dep_r pf A) x. intros T T' h1 A x. subst. rewrite transfer_eq_refl, transfer_dep_r_eq_refl. tauto. Qed. (*Wyckoff*) Lemma transfer_sig_set_eq_in : forall {T:Type} (A B:Ensemble T) (pfa:A = B) (D:Ensemble (sig_set B)) (x:sig_set A), In D (transfer (sig_set_eq _ _ pfa) x) <-> Inn (transfer_dep_r (U:=(fun C : Ensemble T => Ensemble (sig_set C))) pfa D) x. intros T A B h1 D x. subst. pose proof (proof_irrelevance _ (sig_set_eq B B eq_refl) eq_refl) as h1. rewrite h1. rewrite transfer_eq_refl, transfer_dep_r_eq_refl. tauto. Qed. (*Wyckoff*) Lemma transfer_in_r : forall {T T':Type} (pf:T' = T) (A:Ensemble T) (x:T), In A x <-> In (transfer_dep_r pf A) (transfer_r pf x). intros T T' h1 A x. subst. tauto. Qed. (*Wyckoff*) Lemma transfer_in_r' : forall {T T':Type} (pf:T = T') (A:Ensemble T) (x:T'), In A (transfer_r pf x) <-> In (transfer_dep pf A) x. intros T T' h1 A x. subst. rewrite transfer_r_eq_refl, transfer_dep_eq_refl. tauto. Qed. (*Wyckoff*) Lemma transfer_r_sig_set_eq_in : forall {T:Type} (A B:Ensemble T) (pfa:A = B) (D:Ensemble (sig_set A)) (x:sig_set B), In D (transfer_r (sig_set_eq _ _ pfa) x) <-> Inn (transfer_dep (U:=(fun C : Ensemble T => Ensemble (sig_set C))) pfa D) x. intros T A B h1 D x. subst. pose proof (proof_irrelevance _ (sig_set_eq B B eq_refl) eq_refl) as h1. rewrite h1, transfer_r_eq_refl, transfer_dep_eq_refl. tauto. Qed. (*Wyckoff*) Lemma transfer_sig_set_eq : forall {T:Type} (A A':Ensemble T) (pf:A = A') (pfs:sig_set A = sig_set A') (x:sig_set A), transfer pfs x = exist (fun x=>In A' x) (proj1_sig x) (@transfer_dep _ (fun S=>In S (proj1_sig x)) _ _ pf (proj2_sig x)). intros T A A' h1 h2 x. subst. pose proof (proof_irrelevance _ h2 (eq_refl _)); subst. simpl. destruct x as [x h1]. apply proj1_sig_injective. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_r_sig_set_eq : forall {T:Type} (A A':Ensemble T) (pf:A = A') (pfs:sig_set A = sig_set A') (x:sig_set A'), transfer_r pfs x = exist (fun x=>In A x) (proj1_sig x) (@transfer_dep_r _ (fun S=>In S (proj1_sig x)) _ _ pf (proj2_sig x)). intros T A A' h1 h2 x. subst. pose proof (proof_irrelevance _ h2 (eq_refl _)); subst. simpl. destruct x as [x h1]. apply proj1_sig_injective. simpl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep_transfer_dep_sig_set_eq_compat : forall {T:Type} (A B:Ensemble T) (pf:A = B) (U: Type->Type) (f:U (sig_set A)), transfer_dep pf f (U:=fun C:Ensemble T=> U (sig_set C)) = transfer_dep (sig_set_eq _ _ pf) f. intros T A B h1 U f. subst. rewrite transfer_dep_eq_refl. pose proof (proof_irrelevance _ (sig_set_eq B B eq_refl) eq_refl) as h1. rewrite h1. rewrite transfer_dep_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep_r_transfer_dep_r_sig_set_eq_compat : forall {T:Type} (A B:Ensemble T) (pf:A = B) (U: Type->Type) (f:U (sig_set B)), transfer_dep_r pf f (U:=fun C:Ensemble T=> U (sig_set C)) = transfer_dep_r (sig_set_eq _ _ pf) f. intros T A B h1 U f. subst. rewrite transfer_dep_r_eq_refl. pose proof (proof_irrelevance _ (sig_set_eq B B eq_refl) eq_refl) as h1. rewrite h1, transfer_dep_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep_transfer_compat : forall {T:Type} (A B:T) (pf:A = B) (U: T->Type) (f:U A), transfer_dep pf f (U:=fun C:T=> U C) = transfer (f_equal U pf) f. intros T A B h1 U f. subst. simpl. rewrite transfer_dep_eq_refl. rewrite transfer_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep_r_transfer_r_compat : forall {T:Type} (A B:T) (pf:A = B) (U: T->Type) (f:U B), transfer_dep_r pf f (U:=fun C:T=> U C) = transfer_r (f_equal U pf) f. intros T A B h1 U f. subst. simpl. rewrite transfer_dep_r_eq_refl. rewrite transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_transfer_dep_compat : forall {T:Type} {U: T->Type} {A B:T} (pfe:A = B) (pfeu:U A = U B) (f:U A), transfer pfeu f = transfer_dep pfe f (U:=fun C:T=> U C). intros; subst; rewrite transfer_eq_refl, transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_transfer_dep_r_compat : forall {T:Type} {U: T->Type} {A B:T} (pfe:A = B) (pfeu:U A = U B) (f:U B), transfer_r pfeu f = transfer_dep_r pfe f (U:=fun C:T=> U C). intros; subst; rewrite transfer_r_eq_refl, transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep_id_transfer_compat : forall (T U:Type) (pf:T = U) (a:T), transfer_dep pf a (U:=@id Type) = transfer pf a. intros T U h1 a. subst. rewrite transfer_dep_eq_refl, transfer_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma transfer_dep_r_id_transfer_r_compat : forall (T U:Type) (pf:T = U) (a:U), transfer_dep_r pf a (U:=@id Type) = transfer_r pf a. intros T U h1 a. subst. rewrite transfer_dep_r_eq_refl, transfer_r_eq_refl. reflexivity. Qed. (*Wyckoff*) Lemma extends_prim_impl_extends_sig : forall {T T' U:Type} (f:T->U) (g:T'->U), extends_prim f g -> exists (A:Ensemble T) (pf:T' = sig_set A), extends_sig1 f (transfer_dep (U:=fun V=>V->U) pf g). intros T T' U f g h1. red in h1. destruct h1 as [A h1]. destruct h1 as [h1 h2]. subst. exists A. exists eq_refl. red. intro x. rewrite transfer_dep_eq_refl. specialize (h2 (proj1_sig x) x (proj2_sig x)). rewrite <- unfold_sig in h2. specialize (h2 eq_refl). rewrite h2. reflexivity. Qed. (*Wyckoff*) Lemma proj1_sig_transfer_dep : forall {T:Type} {U:Type->Type} {P:T->(U T)->Prop} (l l':U T) (pfx:l = l') (x:{c:T | P c l}), proj1_sig (transfer_dep (U:=(fun xl => {t:T | P t xl})) pfx x) = proj1_sig x. intros; subst; rewrite transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma proj1_sig_transfer_dep_r : forall {T:Type} {U:Type->Type} {P:T->(U T)->Prop} (l l':U T) (pfx:l = l') (x:{c:T | P c l'}), proj1_sig (transfer_dep_r (U:=(fun xl => {t:T | P t xl})) pfx x) = proj1_sig x. intros; subst; rewrite transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma proj1_sig_transfer_dep_s : forall {T:Type} {U:Type->Set} {P:T->(U T)->Prop} (l l':U T) (pfx:l = l') (x:{c:T | P c l}), proj1_sig (transfer_dep (U:=(fun xl => {t:T | P t xl})) pfx x) = proj1_sig x. intros; subst; rewrite transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma proj1_sig_transfer_dep_r_s : forall {T:Type} {U:Type->Set} {P:T->(U T)->Prop} (l l':U T) (pfx:l = l') (x:{c:T | P c l'}), proj1_sig (transfer_dep_r (U:=(fun xl => {t:T | P t xl})) pfx x) = proj1_sig x. intros; subst; rewrite transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep_fun_compat : forall {T V:Type} {U:T->Type} {a b:T} (pf:a = b) (f:U a->V) (x:U b), transfer_dep (U:=(fun x:T => U x -> V)) pf f x = f (transfer_dep_r pf x). intros; subst. rewrite transfer_dep_eq_refl, transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep_r_fun_compat : forall {T V:Type} {U:T->Type} {a b:T} (pf:a = b) (f:U b->V) (x:U a), transfer_dep_r (U:=(fun x:T => U x -> V)) pf f x = f (transfer_dep pf x). intros; subst. rewrite transfer_dep_eq_refl, transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Definition transfer_seg_fun {n n':nat} {U:Type} (pf:n = n') (f:seg_fun n U) : seg_fun n' U := transfer_dep (U:=fun a => seg_fun a U) pf f. (*Wyckoff*) Lemma transfer_seg_fun_eq_refl : forall {n:nat} {U:Type} (f:seg_fun n U), transfer_seg_fun eq_refl f = f. intros; unfold transfer_seg_fun; rewrite transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_seg_fun_compat : forall {n n':nat} {U:Type} (pfeq:n = n') (f:seg_fun n U) (m:nat) (pflt:m < n'), transfer_seg_fun pfeq f m pflt = f m (transfer_dep_r (U:= fun c => m < c) pfeq pflt). intros n n' U h1 f m h2. subst. rewrite transfer_seg_fun_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_seg_fun_app : forall {n} {U:Type} {i i'} (pfe:i = i') (f:seg_fun n U) (pfi:i < n), f i pfi = f i' (eq_lt_trans _ _ _ (eq_sym pfe) pfi). intros; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma transfer_seg_fun_app_r : forall {n} {U:Type} {i i'} (pfe:i = i') (f:seg_fun n U) (pfi':i' < n), f i' pfi' = f i (eq_lt_trans _ _ _ pfe pfi'). intros; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma transfer_seg_fun_app' : forall {n} {U:Type} {i i'} (pfe:i = i') (f:seg_fun n U) (pfi: i < n) (pfi':i' < n), f i pfi = f i' pfi'. intros; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma transfer_seg_fun_app_r' : forall {n} {U:Type} {i i'} (pfe:i = i') (f:seg_fun n U) (pfi:i < n) (pfi':i' < n), f i' pfi' = f i pfi. intros; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Definition transfer_seg_fun_depT2 {T n} {U:T->Type} {f:seg_fun n T} {i i'} {pfi: i < n} (pfe:i = i') (pfi':i' < n) (u:U (f i pfi)) : U (f i' pfi'). subst; assert (pfi = pfi'); try apply proof_irrelevance; subst; auto. Defined. (*Wyckoff*) Definition transfer_seg_fun_depT2_r {T n} {U:T->Type} {f:seg_fun n T} {i i'} {pfi': i' < n} (pfe:i = i') (pfi:i < n) (u:U (f i' pfi')) : U (f i pfi). subst; assert (pfi = pfi'); try apply proof_irrelevance; subst; auto. Defined. (*Wyckoff*) Lemma transfer_seg_fun_depT2_eq_refl : forall {T n} {U:T->Type} {f:seg_fun n T} {i} (pfi:i < n) (pfe: i = i) (u:U (f i pfi)), transfer_seg_fun_depT2 pfe pfi u = u. unfold transfer_seg_fun_depT2, eq_rect_r, eq_rect; intros. pose proof (proof_irrelevance _ (eq_sym pfe) (eq_refl _)) as h1. rewrite h1. pose proof (proof_irrelevance _ (eq_sym (proof_irrelevance (i < n) pfi pfi)) (eq_refl _)) as h2. rewrite h2; auto. Qed. (*Wyckoff*) Lemma transfer_seg_fun_depT2_r_eq_refl : forall {T n} {U:T->Type} {f:seg_fun n T} {i} (pfi:i < n) (pfe: i = i) (u:U (f i pfi)), transfer_seg_fun_depT2_r pfe pfi u = u. unfold transfer_seg_fun_depT2_r, eq_rect_r, eq_rect; intros. pose proof (proof_irrelevance _ (eq_sym pfe) (eq_refl _)) as h1. pose proof (proof_irrelevance _ (eq_sym (proof_irrelevance (i < n) pfi pfi)) (eq_refl _)) as h2. rewrite h1, h2; auto. Qed. (*Wyckoff*) Lemma transfer_seg_fun_depT2_eq_iff : forall {T n} {U:T->Type} {f:seg_fun n T} {i i'} {pfi: i < n} (pfe:i = i') (pfi':i' < n) (u:U (f i pfi)) (v:U (f i' pfi')), transfer_seg_fun_depT2 pfe pfi' u = v <-> existT id _ u = existT id _ v. intros; subst; pose proof (proof_irrelevance _ pfi pfi'); subst; rewrite transfer_seg_fun_depT2_eq_refl; split; intro h1; subst; auto; apply existT_injective2 in h1; auto. Qed. (*Wyckoff*) Lemma transfer_seg_fun_depT2_r_eq_iff : forall {T n} {U:T->Type} {f:seg_fun n T} {i i'} {pfi':i' < n} (pfe:i = i') (pfi:i < n) (u:U (f i' pfi')) (v:U (f i pfi)), transfer_seg_fun_depT2_r pfe pfi u = v <-> existT id _ u = existT id _ v. intros; subst; pose proof (proof_irrelevance _ pfi pfi'); subst; rewrite transfer_seg_fun_depT2_r_eq_refl; split; intro h1; subst; auto; apply existT_injective2 in h1; auto. Qed. (*Wyckoff*) Lemma transfer_seg_fun_depT2_transfer_dep : forall {T n} {U:T->Type} {f:seg_fun n T} {i i' t} {pfi: i < n} (pfi':i' < n) (pfe: i = i') (u:U t) (pfe':t = f i pfi), transfer_seg_fun_depT2 pfe pfi' (transfer_dep pfe' u) = transfer_dep (eq_trans pfe' (transfer_seg_fun_app' pfe f pfi pfi')) u. intros; subst; rewrite transfer_dep_eq_refl; pose proof (proof_irrelevance _ pfi pfi'); subst; rewrite transfer_seg_fun_depT2_eq_refl, transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_seg_fun_depT2_r_transfer_dep : forall {T n} {U:T->Type} {f:seg_fun n T} {i i' t} {pfi: i < n} (pfi':i' < n) (pfe: i = i') (u:U t) (pfe':t = f i' pfi'), transfer_seg_fun_depT2_r pfe pfi (transfer_dep pfe' u) = transfer_dep (eq_trans pfe' (transfer_seg_fun_app_r' pfe f pfi pfi')) u. intros; subst; rewrite transfer_dep_eq_refl; pose proof (proof_irrelevance _ pfi pfi'); subst; rewrite transfer_seg_fun_depT2_r_eq_refl, transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Definition transfer_seg_fun_depT_pf {T n} {U:T->Type} {f:seg_fun n T} {i} {pfi:i < n} (pfi':i < n) (x:U (f i pfi)) : U (f i pfi') := transfer_dep (proof_irrelevance _ pfi pfi') x (U := fun pfc => U (f i pfc)). (*Wyckoff*) Definition transfer_sig {T} {P Q:T->Prop} (pfi:pred_impl P Q) (x:{t:T | P t}) : {x:T | Q x} := exist Q (proj1_sig x) (pfi _ (proj2_sig x)). (*Wyckoff*) Definition transfer_sig' {T} {P Q:T->Prop} (pfi:iff_pred P Q) (x:{t:T | P t}) : {x:T | Q x} := transfer_sig (iff_pred1 _ _ pfi) x. (*Wyckoff*) Lemma in_transfer_dep : forall {T} {U:T->Type} {x y:T} {A:forall t, Ensemble (U t)} (pfe:x = y) (z:U x), Inn (A x) z -> Inn (A y) (transfer_dep pfe z). intros; subst; rewrite transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma in_transfer_dep_r : forall {T} {U:T->Type} {x y:T} {A:forall t, Ensemble (U t)} (pfe:x = y) (z:U y), Inn (A y) z -> Inn (A x) (transfer_dep_r pfe z). intros; subst; rewrite transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma in_transfer_dep_iff : forall {T U} (pfe:T = U) (A:Ensemble T) (y:U), Inn (transfer_dep pfe A) y <-> Inn A (transfer_r pfe y). intros; subst; tauto. Qed. (*Wyckoff*) Lemma in_transfer_dep_r_iff : forall {T U} (pfe:T = U) (A:Ensemble U) (x:T), Inn (transfer_dep_r pfe A) x <-> Inn A (transfer pfe x). intros; subst; tauto. Qed. (*Wyckoff*) Lemma in_transfer_ens : forall {T U} {A:T->Ensemble U} {x y} {u} (pfe:x = y), Inn (A x) u -> Inn (A y) u. intros; subst; auto. Qed. (*Wyckoff*) Lemma in_transfer_ens_r : forall {T U} {A:T->Ensemble U} {x y} {u} (pfe:x = y), Inn (A y) u -> Inn (A x) u. intros; subst; auto. Qed. (*Wyckoff*) Lemma in_transfer_dep_ens : forall {T} {U:T->Type} {A:forall x:T, Ensemble (U x)} {x y} {u:U x} (pfe:x = y), Inn (A x) u -> Inn (A y) (transfer_dep pfe u). intros; subst; auto. Qed. (*Wyckoff*) Lemma in_transfer_dep_ens_r : forall {T} {U:T->Type} {A:forall t:T, Ensemble (U t)} {x y} {u:U y} (pfe:x = y), Inn (A y) u -> Inn (A x) (transfer_dep_r pfe u). intros; subst; auto. Qed. (*Wyckoff*) Definition transfer_dep_ens' {T} {U:T->Type} {x y:T} (pfe:x = y) (A:Ensemble (U x)) : Ensemble (U y). subst; auto. Defined. (*Wyckoff*) Definition transfer_dep_ens_r' {T} {U:T->Type} {x y:T} (pfe:x = y) (A:Ensemble (U y)) : Ensemble (U x). subst; auto. Defined. (*Wyckoff*) Lemma in_transfer_dep_ens' : forall {T} {U:T->Type} {x y:T} {A:Ensemble (U x)} {z:U x} (pfe:x = y), Inn A z -> Inn (transfer_dep_ens' pfe A) (transfer_dep pfe z). intros; subst; auto. Qed. (*Wyckoff*) Lemma in_transfer_dep_ens_r' : forall {T} {U:T->Type} {x y:T} {A:Ensemble (U y)} {z:U y} (pfe:x = y), Inn A z -> Inn (transfer_dep_ens_r' pfe A) (transfer_dep_r pfe z). intros; subst; auto. Qed. (*Wyckoff*) Definition transfer_prod {U U' V V'} (pfu:U = U') (pfv:V = V') (pr:U*V) : U' * V' := (transfer pfu (fst pr), transfer pfv (snd pr)). (*Wyckoff*) Fixpoint transfer_tuple_prod {T V n} {U:T->Type} {W:V->Type} : forall {w:T^n} {v:V^n} (pfe:pointwise_ind_eq U W w v) (t:tuple_prod U w), tuple_prod W v := match n with | 0 => fun _ _ _ _ => tt | S n' => fun w v pfe t => let pfe' := pointwise_ind_eq_to_S pfe in (transfer (pfe 0 (lt_0_Sn _)) (fst t), @transfer_tuple_prod _ _ n' _ _ (snd w) (snd v) pfe' (snd t)) end. (*Wyckoff*) Fixpoint transfer_tuple_prod_r {T V n} {U:T->Type} {W:V->Type} : forall {w:T^n} {v:V^n} (pfe:pointwise_ind_eq U W w v) (t:tuple_prod W v), tuple_prod U w := match n with | 0 => fun _ _ _ _ => tt | S n' => fun w v pfe t => let pfe' := pointwise_ind_eq_to_S pfe in (transfer_r (pfe 0 (lt_0_Sn _)) (fst t), @transfer_tuple_prod_r _ _ n' _ _ (snd w) (snd v) pfe' (snd t)) end. (*Wyckoff*) Fixpoint transfer_tuple_prod_const {T U n} : forall {x:T^n} (p:tuple_prod (fun _:T => U) x), U ^ n := match n with | 0 => fun _ _ => tt | S n' => fun x p => let (fx, sx) := x in let (fp, sp) := p in let t := transfer_tuple_prod_const sp in (fp, t) end. (*Wyckoff*) Fixpoint transfer_to_tuple_prod_const {T U n} : forall (x:T^n) (u:U^n), tuple_prod (fun _:T => U) x := match n with | 0 => fun _ _ => tt | S n' => fun x u => let (fu, su) := u in let t := transfer_to_tuple_prod_const _ su in (fu, t) end. (*Wyckoff*) Lemma transfer_tuple_prod_const_to_compat : forall {T U n} (x:T^n) (u:U^n), transfer_tuple_prod_const (transfer_to_tuple_prod_const x u) = u. intros T U n; induction n; simpl. intros; apply unit_eq. intros x u; destruct x, u; f_equal; auto. Qed. (*Wyckoff*) Lemma transfer_to_tuple_prod_const_compat : forall {T U n} {x:T^n} (p:tuple_prod (fun _:T => U) x), transfer_to_tuple_prod_const x (transfer_tuple_prod_const p) = p. intros T U n; induction n; simpl. intros; apply unit_eq. intros x u; destruct x, u; f_equal; auto. Qed. (*Wyckoff*) Lemma inj_transfer_tuple_prod_const : forall {T U n} (x:T^n) (p p':tuple_prod (fun _:T => U) x), transfer_tuple_prod_const p = transfer_tuple_prod_const p' -> p = p'. intros T U n. induction n as [|n ih]; simpl. intros; apply unit_eq. intros x p p' h1; simpl in h1. destruct x as [fx sx], p as [fp sp], p' as [fp' sp']. inversion h1; subst; f_equal; auto. Qed. (*Wyckoff*) Lemma inj_transfer_to_tuple_prod_const : forall {T U n} (x:T^n) (u u':U^n), transfer_to_tuple_prod_const x u = transfer_to_tuple_prod_const x u' -> u = u'. intros T U n. induction n as [|n ih]; simpl. intros; apply unit_eq. intros x u u' h1; simpl in h1. destruct x as [fx sx], u as [fu su], u' as [fu' su']. inversion h1; subst; f_equal. apply ih in H1; auto. Qed. (*Wyckoff*) Lemma nth_tuple_prod_transfer_to_tuple_prod_const : forall {T U n} (x:T^n) (y:U^n) i (pfi:i < n), nth_tuple_prod (transfer_to_tuple_prod_const x y) i pfi = nth_prod i pfi y. intros ? ? n; induction n; simpl. intros; omega. intros x y i; destruct x, y, i; simpl; auto. Qed. (*Wyckoff*) Lemma nth_prod_transfer_tuple_prod_const : forall {T U n} {x:T^n} (y:tuple_prod (fun _:T => U) x) i (pfi:i < n), nth_prod i pfi (transfer_tuple_prod_const y) = nth_tuple_prod y i pfi. intros ? ? n; induction n; simpl. intros; omega. intros x y i; destruct x, y, i; simpl; auto. Qed. (*Wyckoff*) Lemma tuple0_eq : forall {T n} (x:T^n) (pfe:n = 0), x = transfer_dep_r pfe tt. intros; subst; rewrite transfer_dep_r_eq_refl; apply unit_eq. Qed. (*Wyckoff*) Definition transfer_tt0 T : T^0 := tt. (*Wyckoff*) Definition transfer_tt {n} (pfeq:n = 0) T : T^n := transfer_dep_r pfeq (transfer_tt0 T). (*Wyckoff*) Lemma transfer_tt_eq_refl : forall T (pf0:0 = 0), transfer_tt pf0 T = tt. unfold transfer_tt; intros. rewrite transfer_dep_r_eq_refl. apply nprod0_eq. Defined. (*Wyckoff*) Definition transfer_tt0_tuple_prod {T} (A:T->Type) : tuple_prod A (transfer_tt0 T) := tt. (*Wyckoff*) Definition transfer_tt0_tuple_prod' {T} (A:T->Type) (x:T^0) : tuple_prod A x := tt. (*Wyckoff*) Definition transfer_tt_tuple_prod {T n} (A:T->Type) (pfeq:n = 0) (x:T^n) : tuple_prod A x. intros; subst; destruct x; refine tt. Defined. (*Wyckoff*) Lemma transfer_tt_tuple_prod_eq_refl : forall {T} (A:T->Type) (pfeq:0 = 0) (x:T^0), transfer_tt_tuple_prod A pfeq x = tt. unfold transfer_tt_tuple_prod, eq_rect_r, eq_rect; intros. pose proof (proof_irrelevance _ pfeq (eq_refl _)); subst. simpl. destruct x; auto. Defined. (*Wyckoff*) Definition transfer_to_tuple_prod0 {T} (A:T->Type) (x:T^0) (p:tuple_prod A x) : tuple_prod A (transfer_tt0 T) := p. (*Wyckoff*) Definition transfer_tuple_prod0 {T} (A:T->Type) (x:T^0) (p:tuple_prod A (transfer_tt0 T)) : tuple_prod A x := p. (*Wyckoff*) Lemma tuple_prod0_eq : forall {T} (A:T->Type) (x:T^0) (p q:tuple_prod A x), p = q. simpl; intros; apply unit_eq. Qed. (*Wyckoff*) Lemma tuple_prod0_eq' : forall {T} (A:T->Type) (x y:T^0) (p:tuple_prod A x) (q:tuple_prod A y) (pfe:x = y), p = transfer_dep_r pfe q. simpl; intros; apply unit_eq. Qed. (*Wyckoff*) Lemma tuple_prod_n0_eq : forall {T n} (A:T->Type) (x:T^n) (p q:tuple_prod A x), n = 0 -> p = q. simpl; intros; subst; apply tuple_prod0_eq. Qed. (*Wyckoff*) Lemma tuple_prod_n0_eq' : forall {T n} (A:T->Type) (x y:T^n) (p:tuple_prod A x) (q:tuple_prod A y) (pfe:x = y), n = 0 -> p = transfer_dep_r pfe q. simpl; intros; subst; rewrite transfer_dep_r_eq_refl; apply tuple_prod0_eq. Qed. (*Wyckoff*) Lemma tuple_prod_tt : forall {T n} {U:T->Type} {x:T^n} (w:tuple_prod U x) (pfe:n = 0), w = transfer_tt_tuple_prod U pfe x. intros; subst. rewrite transfer_tt_tuple_prod_eq_refl. destruct x. red in w. apply unit_eq. Qed. (*Wyckoff*) Definition transfer_seg_fun_to_S_app {n} (f:seg_fun (S n) Type) i (pfi:i < n) (pfi':S i < S n) (x:seg_fun_to_S f i pfi) : f (S i) pfi'. unfold seg_fun_to_S in x. erewrite f_equal_dep. apply x. auto. Defined. (*Wyckoff*) Definition transfer_to_seg_fun_to_S_app {n} (f:seg_fun (S n) Type) i (pfi:i < n) (pfi':S i < S n) (x:f (S i) pfi') : seg_fun_to_S f i pfi. unfold seg_fun_to_S. erewrite f_equal_dep in x. apply x. auto. Defined. (*Wyckoff*) Lemma transfer_seg_fun_to_S_app_undoes : forall {n} (f:seg_fun (S n) Type) i (pfi:i < n) (pfi':S i < S n) (x:f (S i) pfi'), transfer_seg_fun_to_S_app f i pfi pfi' (transfer_to_seg_fun_to_S_app f i pfi pfi' x) = x. unfold transfer_to_seg_fun_to_S_app, transfer_seg_fun_to_S_app; intros; unfold eq_rect_r, eq_rect. pose proof (proof_irrelevance _ pfi' (lt_n_S i n pfi)); subst. rewrite eq_eq_refl at 1. rewrite eq_eq_refl at 1. reflexivity. Qed. (*Wyckoff*) Lemma transfer_to_seg_fun_to_S_app_undoes : forall {n} (f:seg_fun (S n) Type) i (pfi:i < n) (pfi':S i < S n) (x:seg_fun_to_S f i pfi), transfer_to_seg_fun_to_S_app f i pfi pfi' (transfer_seg_fun_to_S_app f i pfi pfi' x) = x. unfold transfer_to_seg_fun_to_S_app, transfer_seg_fun_to_S_app; intros; unfold eq_rect_r, eq_rect. pose proof (proof_irrelevance _ pfi' (lt_n_S i n pfi)); subst. rewrite eq_eq_refl at 1. rewrite eq_eq_refl at 1. reflexivity. Qed. (*Wyckoff*) Definition transfer_seg_fun_to_S_app_iff : forall {n} (f:seg_fun (S n) Type) i (pfi:i < n) (pfi':S i < S n) (x:seg_fun_to_S f i pfi) (y:f (S i) pfi'), transfer_seg_fun_to_S_app f i pfi pfi' x = y <-> x = transfer_to_seg_fun_to_S_app f i pfi pfi' y. intros n f i hi hi' x y; split; intro h1; subst; [rewrite transfer_to_seg_fun_to_S_app_undoes; auto | rewrite transfer_seg_fun_to_S_app_undoes; auto]. Qed. (*Wyckoff*) Lemma transfer_seg_fun_to_S_app_subst : forall {n} {f:seg_fun (S n) Type} (k:seg_fun_dep (seg_fun_to_S f)) i (pfi:i < n), transfer_seg_fun_to_S_app f i pfi (lt_n_S _ _ pfi) (k i pfi) = k i pfi. intros n f k i hi. unfold transfer_seg_fun_to_S_app, eq_rect_r, eq_rect. match goal with |- context [match ?pf with | _ => _ end] => pose pf as h3 end. pose proof (proof_irrelevance _ h3 (eq_refl _)) as h4. subst. rewrite h4; auto. Qed. (*Wyckoff*) Definition transfer_dep_funT_app {T:Type} {P:T->Prop} {f:prop_dep_funT P} {x y:T} (pfe:x = y) (pfpx:P x) (pfpy:P y) : f x pfpx -> f y pfpy. intros; subst; pose proof (proof_irrelevance _ pfpx pfpy); subst; auto. Defined. (*Wyckoff*) Definition transfer_dep_funT_app_r {T:Type} {P:T->Prop} {f:prop_dep_funT P} {x y:T} (pfe:x = y) (pfpx:P x) (pfpy:P y) : f y pfpy -> f x pfpx. intros; subst; pose proof (proof_irrelevance _ pfpx pfpy); subst; auto. Defined. (*Wyckoff*) Lemma transfer_dep_funT_app_eq_refl : forall {T:Type} {P:T->Prop} {f:prop_dep_funT P} {x:T} (pfe:x = x) (pfx pfy:P x) (fx:f x pfx), transfer_dep_funT_app pfe pfx pfy fx = transfer_dep (proof_irrelevance _ pfx pfy) fx (U := fun pf:P x => f x pf). unfold transfer_dep_funT_app, eq_rect_r, eq_rect; intros ? ? ? ? h1. pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. simpl. intros pfx pfy h1. pose proof (proof_irrelevance _ pfx pfy); subst. pose proof (proof_irrelevance _ (proof_irrelevance (P x) pfy pfy) (eq_refl _)) as h2. rewrite h2; simpl; rewrite transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_dep_funT_app_r_eq_refl : forall {T:Type} {P:T->Prop} {f:prop_dep_funT P} {x:T} (pfe:x = x) (pfx pfy:P x) (fy:f x pfy), transfer_dep_funT_app_r pfe pfx pfy fy = transfer_dep (proof_irrelevance _ pfy pfx) fy (U := fun pf:P x => f x pf). unfold transfer_dep_funT_app_r, eq_rect_r, eq_rect; intros ? ? ? ? h1. pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. simpl. intros pfx pfy h1. pose proof (proof_irrelevance _ pfx pfy); subst. pose proof (proof_irrelevance _ (proof_irrelevance (P x) pfy pfy) (eq_refl _)) as h2. rewrite h2; simpl; rewrite transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma f_equal_dep2 : forall {T:Type} {P:T->Prop} {U} {f:prop_dep_funT P} {g:prop_dep_fun_nest2 f U} {x y:T} {pfpx:P x} {pfpy:P y} {fx:f x pfpx} {fy:f y pfpy} (pfe: x = y) (pffe:fx = transfer_dep_funT_app_r pfe pfpx pfpy fy), g x pfpx fx = g y pfpy fy. intros; subst; pose proof (proof_irrelevance _ pfpx pfpy); subst; f_equal. rewrite transfer_dep_funT_app_r_eq_refl, transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Definition transfer_prop_dep {T} {U:T->Type} {P:forall x:T, U x -> Prop} {x y:T} {ux:U x} {uy:U y} (pfe:x = y) (_:ux = transfer_dep_r pfe uy) : P x ux -> P y uy. subst; rewrite transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma transfer_eq_iff : forall {T U} {x:T} {y:U} (pfe:T = U), transfer pfe x = y <-> existT id T x = existT id U y. intros; subst; rewrite transfer_eq_refl; split; intro h1; subst; try apply existT_injective2 in h1; subst; auto. Qed. (*Wyckoff*) Lemma transfer_r_eq_iff : forall {T U} {x:T} {y:U} (pfe:T = U), transfer_r pfe y = x <-> existT id U y = existT id T x. intros; subst; rewrite transfer_r_eq_refl; split; intro h1; subst; try apply existT_injective2 in h1; subst; auto. Qed. (*Wyckoff*) Lemma f_equal_depT : forall {T} {U:T->Type} {x y} (pfe:x = y) (ux:U x), existT _ x ux = existT _ y (transfer_dep pfe ux). intros; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma f_equal_depT' : forall {T} {U:T->Type} x y (ux:U x) (uy:U y) (pfe:x = y), ux = transfer_dep_r pfe uy -> existT (fun x => U x) x ux = existT (fun x => U x) y uy. intros; subst; f_equal. Qed. (*Wyckoff*) Lemma f_equal_nprod : forall {T m n} (pfe:m = n) (x:T^m), existT _ m x = existT _ n (transfer_nprod pfe x). intros; subst; f_equal; rewrite transfer_nprod_eq_refl; auto. Qed. (*Wyckoff*) Lemma f_equal_nprod' : forall {T m n} {x:T^m} {y:T^n} (pfe:m = n), x = transfer_nprod_r pfe y -> existT _ m x = existT _ n y. intros; subst; f_equal; rewrite transfer_nprod_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma f_equal_prop_depT : forall {T} {U:T->Type} (P:forall x, U x -> Prop) x y (ux:U x) (uy:U y) (pfe:x = y), ux = transfer_dep_r pfe uy -> P x ux -> P y uy. intros ? ? ? ? ? ? ? ? ? h1; subst; rewrite transfer_dep_r_eq_refl in h1; auto. Qed. (*Wyckoff*) Lemma f_equal_fun_dep' : forall {T V} {U:T->Type} {P:forall x:T, U x -> Prop} {x y:T} {ux:U x} {uy:U y} (f:prop_dep_fun_dep P V) (pfe:x = y) (pfe':ux = transfer_dep_r pfe uy) (px:P x ux), let py := f_equal_prop_depT P x y ux uy pfe pfe' px in f x ux px = f y uy py. intros T V U P x y ux uy f px h1; subst. simpl. intro h1. symmetry. gen0. revert h1. rewrite transfer_dep_r_eq_refl. intros; subst; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma f_equal_fun_dep : forall {T V} {U:T->Type} {P:forall x, U x -> Prop} (f:prop_dep_fun_dep P V) (x y:T) (ux:U x) (uy:U y) (pfpx:P x ux) (pfpy:P y uy) (pfe:x = y) (pfeu:ux = transfer_dep_r pfe uy), f x ux pfpx = f y uy pfpy. intros; rewrite (f_equal_fun_dep' f pfe pfeu); repeat f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma f_equal_dep2T' : forall {T U:Type} {P:T->U->Prop} {V:T->U->Type} (f:prop_dep_fun2T' P V) (x x':T) (y y':U) (pfp:P x y) (pfp':P x' y') (pfex:x = x') (pfey:y = y'), f x y pfp = transfer_dep2_r pfex pfey (f x' y' pfp'). intros; subst; f_equal; rewrite transfer_dep2_r_eq_refl; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma proj1_sig_transfer_sig_set : forall {T} {A B:Ensemble T} (x:sig_set A) (pfe:A = B) (pfe':sig_set A = sig_set B), proj1_sig (P := B) (transfer pfe' x) = proj1_sig x. intros; subst; rewrite transfer_eq_refl; auto. Qed. (*Wyckoff*) Lemma proj1_sig_transfer_r_sig_set : forall {T} {A B:Ensemble T} (x:sig_set B) (pfe:A = B) (pfe':sig_set A = sig_set B), proj1_sig (P := A) (transfer_r pfe' x) = proj1_sig x. intros; subst; rewrite transfer_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma proj1_sig_set_transfer_dep : forall {T} {U:T->Type} {A:forall t:T, Ensemble (U t)} {t t'} (pfe:t = t') (x:sig_set (A t)), proj1_sig (transfer_dep (U := fun c => sig_set (A c)) pfe x) = transfer_dep pfe (proj1_sig x). intros; subst; rewrite transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma proj1_sig_set_transfer_dep_r : forall {T} {U:T->Type} {A:forall t:T, Ensemble (U t)} {t t'} (pfe:t = t') (x:sig_set (A t')), proj1_sig (transfer_dep_r (U := fun c => sig_set (A c)) pfe x) = transfer_dep_r pfe (proj1_sig x). intros; subst; rewrite transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma Sfun_dep_functional : forall {S} {U V:S->Type} {A:forall s:S, Ensemble (U s)} {B:forall s:S, Ensemble (V s)} (f:Sfun_dep A B) s s' (u:U s) (u':U s') (pfa: Inn (A s) u) (pfa':Inn (A s') u') (pfe:s = s') (pfe':u = transfer_dep_r pfe u'), f s u pfa = transfer_dep_r pfe (f s' u' pfa') (U := fun s => sig_set (B s)). intros; subst; rewrite transfer_dep_r_eq_refl. revert pfa. rewrite transfer_dep_r_eq_refl. intro pfa. pose proof (proof_irrelevance _ pfa pfa'); subst; auto. Qed. (*Wyckoff*) Lemma nprod_to_list_functional : forall T n n' (x:T^n) (x':T^n') (pfe:n = n'), x = transfer_dep_r pfe x' -> nprod_to_list T n x = nprod_to_list T n' x'. intros; subst; rewrite transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma nprod_to_list_S : forall T n (x:T^n) (a:T), nprod_to_list _ (S n) (a, x) = cons a (nprod_to_list _ _ x). auto. Qed. (*Wyckoff*) Lemma nprod_of_list_cons : forall {T} (l:list T) a, nprod_of_list _ (a::l) = (a, nprod_of_list _ l). auto. Qed. (*Wyckoff*) Lemma nprod_of_list_functional : forall T l l' (pfe:l = l'), nprod_of_list T l = transfer_dep_r (U := fun l => T^(length l)) pfe (nprod_of_list T l'). intros; subst; rewrite transfer_dep_r_eq_refl; auto. Qed. (*Wyckoff*) Lemma match_eq_refl_eq_transfer_dep_r_iff : forall {T} {U:T->Type} {x y z} (pfxy:x = y) (pfxz:x = z) (ux:U x) (uz:U z) (pfeu:U y = U x), match pfeu with | eq_refl => ux end = transfer_dep_r (U := U) pfxz uz <-> ux = transfer_dep_r pfxz uz. intros; subst; pose proof (proof_irrelevance _ pfeu (eq_refl _)); subst; rewrite transfer_dep_r_eq_refl; intuition. Qed. (*Wyckoff*) Lemma match_eq_refl_eq_transfer_dep_iff : forall {T} {U:T->Type} {x y z} (pfxz:x = z) (pfyz:y = z) (ux:U x) (uz:U z) (pfeu:U x = U z), match pfeu with | eq_refl => uz end = transfer_dep (U := U) pfxz ux <-> uz = transfer_dep pfxz ux. intros; subst; pose proof (proof_irrelevance _ pfeu (eq_refl _)); subst; rewrite transfer_dep_eq_refl; intuition. Qed. End Transfer. Section TupleProdAgain. Fixpoint in_tuple_prod {T n} {A:T->Type} {s:T} : forall {tpl:T^n} (pfdt:type_eq_dec T) (x:A s), tuple_prod A tpl -> Prop := match n with | 0 => fun _ _ _ _ => False | S n' => fun tpl pfdt => match (pfdt s (fst tpl)) with | left pfe => fun x p => let P := in_tuple_prod pfdt x (snd p) in let x' := transfer_dep pfe x in x' = fst p \/ P | right _ => fun x p => in_tuple_prod pfdt x (snd p) end end. Lemma in_nth_tuple_prod : forall {T n} {A:T->Type} {tpl:T^n} (pfdt:type_eq_dec T) i (pfi:i < n) (p:tuple_prod A tpl), in_tuple_prod pfdt (nth_tuple_prod p i pfi) p. intros T n. induction n as [|n ih]. intros; omega. intros A tpl. destruct tpl as [ft st]. intros hdt i. destruct i as [|i]; simpl. intros. pose proof (eq_dec_same hdt ft (hdt ft ft)) as h1. rewrite h1, transfer_dep_eq_refl. left; auto. intros hi p. destruct p as [fp sp]. lr_match_goal'. subst; rewrite transfer_dep_eq_refl; simpl. right; auto. specialize (ih A st hdt i (lt_S_n _ _ hi) sp). simpl; auto. Qed. Lemma in_tuple_prod_iff : forall {T n} {A:T->Type} {s:T} {tpl:T^n} (pfdt:type_eq_dec T) (x:A s) (p:tuple_prod A tpl), in_tuple_prod pfdt x p <-> exists i (pfi:i < n) (pfe:s = nth_prod i pfi tpl), transfer_dep pfe x = nth_tuple_prod p i pfi. intros T n. induction n as [|n ih]; simpl. intros; split; [intro; contradiction | intro h1; destruct h1 as [? [? h1]]; omega]. intros A s tpl hdt x p. destruct tpl as [ft st], p as [fp sp] ; simpl; simpl in sp, fp, ih. specialize (ih A s st hdt x sp). lr_match_goal'; subst; simpl. rewrite transfer_dep_eq_refl; split; intro h1. destruct h1 as [|h1]; subst. exists 0, (lt_0_Sn _); simpl. exists eq_refl; rewrite transfer_dep_eq_refl; auto. rewrite ih in h1. destruct h1 as [i [hi [h2 h3]]]; subst. exists (S i), (lt_n_S _ _ hi); simpl. ex_goal. f_equal; apply proof_irrelevance. exists hex. rewrite transfer_dep_eq_refl in h3; subst. termr hex. redterm1 y. pose proof (proof_irrelevance _ c hi) as he. fold c in hex; fold c. rewrite <- transfer_dep_eq_iff. gterml. redterm1 x. redterm1 c. fold c0. rewrite he; auto. destruct h1 as [i [hi [h2 h3]]]. destruct i as [|i]; simpl in h2, h3; subst. pose proof (proof_irrelevance _ h2 (eq_refl _)); subst. rewrite transfer_dep_eq_refl. left; auto. rewrite transfer_dep_eq_refl in h3; subst. right. apply in_nth_tuple_prod. rewrite ih; split; intro h1. destruct h1 as [i [h1 [h2 h3]]]; subst. rewrite transfer_dep_eq_refl in h3; subst. exists (S i), (lt_n_S _ _ h1). ex_goal. f_equal; apply proof_irrelevance. exists hex. simpl. rewrite <- transfer_dep_eq_iff. pose proof (proof_irrelevance _ h1 (lt_S_n _ _ (lt_n_S _ _ h1))) as h2. rewrite <- h2; auto. destruct h1 as [i [h1 [h2 h3]]]; subst. rewrite transfer_dep_eq_refl in h3; subst. destruct i as [|i]. simpl in hr. contradict hr; auto. exists i, (lt_S_n _ _ h1); simpl. exists (eq_refl _). rewrite transfer_dep_eq_refl; auto. Qed. Lemma in_tuple_prod_r_iff : forall {T n} {A:T->Type} {s:T} {tpl:T^n} (pfdt:type_eq_dec T) (x:A s) (p:tuple_prod A tpl), in_tuple_prod pfdt x p <-> exists i (pfi:i < n) (pfe:s = nth_prod i pfi tpl), x = transfer_dep_r pfe (nth_tuple_prod p i pfi). intros; rewrite in_tuple_prod_iff; split; intro h1; destruct h1 as [i [hi [h2 h3]]]; subst. rewrite transfer_dep_eq_refl in h3; subst. exists i, hi, eq_refl; rewrite transfer_dep_r_eq_refl; auto. exists i, hi, eq_refl; rewrite transfer_dep_eq_refl; auto. Qed. Fixpoint in_tuple_prod' {T n} {A:T->Type} {s:T} : forall {tpl:T^n} (pfdt:type_eq_dec T) (pfda:dep_type_eq_dec A) (x:A s), tuple_prod A tpl -> Prop := match n with | 0 => fun _ _ _ _ _ => False | S n' => fun tpl pfdt pfda => match (pfdt s (fst tpl)) with | left pfe => fun x p => let x' := transfer_dep pfe x in let P := in_tuple_prod' pfdt pfda x (snd p) in match (pfda _ x' (fst p)) with | left _ => True | right _ => P end | right _ => fun x p => in_tuple_prod' pfdt pfda x (snd p) end end. Lemma in_tuple_prod_compat' : forall {T n} {A:T->Type} {s:T} {tpl:T^n} (pfdt:type_eq_dec T) (pfda:dep_type_eq_dec A) (x:A s) (p:tuple_prod A tpl), in_tuple_prod' pfdt pfda x p <-> in_tuple_prod pfdt x p. intros T n; induction n as [|n ih]; simpl. tauto. intros. lr_match_goal'; subst. rewrite transfer_dep_eq_refl. lr_if_goal'; subst. tauto. split; intro h1. right; auto. rewrite <- ih; auto. apply h1. destruct h1 as [|h1]; subst. contradict hr; auto. rewrite ih; auto. auto. Qed. Lemma in_tuple_prod_iff' : forall {T n} {A:T->Type} {s:T} {tpl:T^n} (pfdt:type_eq_dec T) (pfda:dep_type_eq_dec A) (x:A s) (p:tuple_prod A tpl), in_tuple_prod' pfdt pfda x p <-> exists i (pfi:i < n) (pfe:s = nth_prod i pfi tpl), transfer_dep pfe x = nth_tuple_prod p i pfi. intros; rewrite in_tuple_prod_compat', in_tuple_prod_iff; tauto. Qed. Lemma in_tuple_prod_r_iff' : forall {T n} {A:T->Type} {s:T} {tpl:T^n} (pfdt:type_eq_dec T) (pfda:dep_type_eq_dec A) (x:A s) (p:tuple_prod A tpl), in_tuple_prod' pfdt pfda x p <-> exists i (pfi:i < n) (pfe:s = nth_prod i pfi tpl), x = transfer_dep_r pfe (nth_tuple_prod p i pfi). intros; rewrite in_tuple_prod_compat', in_tuple_prod_r_iff; tauto. Qed. Lemma in_tuple_prod_dec : forall {T n} {A:T->Type} {s:T} {tpl:T^n} (pfdt:type_eq_dec T) (pfda:dep_type_eq_dec A) (x:A s) (p:tuple_prod A tpl), { in_tuple_prod' pfdt pfda x p} + {~ in_tuple_prod' pfdt pfda x p}. intros T n. induction n as [|n ih]; simpl. intros; right; intro; contradiction. intros A s tpl hdt hda x p. lr_match_goal'; subst. lr_if_goal'; subst. left; auto. auto. auto. Qed. Definition in_tuple_prod_all {T} {A:T->Type} {s:T} {tpl:all_tuples T} (pfdt:type_eq_dec T) (x:A s) (p:tuple_prod_all A tpl) : Prop := in_tuple_prod pfdt x p. Definition in_tuple_prod_all' {T} {A:T->Type} {s:T} {tpl:all_tuples T} (pfdt:type_eq_dec T) (pfda:dep_type_eq_dec A) (x:A s) (p:tuple_prod_all A tpl) : Prop := in_tuple_prod' pfdt pfda x p. Lemma in_tuple_prod_all_dec : forall {T} {A:T->Type} {s:T} {tpl:all_tuples T} (pfdt:type_eq_dec T) (pfda:dep_type_eq_dec A) (x:A s) (p:tuple_prod_all A tpl), { in_tuple_prod_all' pfdt pfda x p} + {~ in_tuple_prod_all' pfdt pfda x p}. intros; apply in_tuple_prod_dec. Qed. Fixpoint in_tuple_prod0 {T n} {A:T->Type} : forall {x:T^n} {t:T} (z:A t) (y:tuple_prod A x), Prop := match n with | 0 => fun _ _ _ _ => False | S n' => fun x t z y => (exists (pf:t = fst x), transfer_dep pf z = fst y) \/ (@in_tuple_prod0 T n' A (snd x) t z (snd y)) end. Lemma in_nth_tuple_prod0 : forall {T n} {A:T->Type} {x:T^n} i (pfi:i < n) (y:tuple_prod A x), in_tuple_prod0 (nth_tuple_prod y i pfi) y. intros T n. induction n as [|n ih]; simpl. intros; omega. intros A x. destruct x as [xf xs]. intro i. destruct i as [|i]; simpl. intros h1 h2. left. exists (eq_refl _). rewrite transfer_dep_eq_refl; auto. intros hi y. destruct y as [yf ys]; simpl. right; auto. Qed. Lemma nth_tuple_prod_transfer : forall {S T n} {U:S->Type} {V:T->Type} {x:S^n} {y:T^n} (pfpw:pointwise_ind_eq U V x y) i (pfi:i < n) (p:tuple_prod U x), nth_tuple_prod (transfer_tuple_prod pfpw p) i pfi = transfer (pfpw i pfi) (nth_tuple_prod p i pfi). intros S T n. induction n as [|n ih]; simpl. intros; omega. intros U V x y. destruct x as [x hx]. destruct y as [y hy]. intros h1 i. destruct i as [|i]; simpl. intros h2 p. destruct p as [fp sp]; simpl; auto. f_equal; apply proof_irrelevance. intros hi p. destruct p as [fp sp]. rewrite ih; f_equal; apply proof_irrelevance. Qed. Lemma nth_tuple_prod_functional1 : forall {T n} {A:T->Type} {t:T^n} (p p':tuple_prod A t) k k' (pfk:k < n) (pfep:p = p') (pfe:k = k'), nth_tuple_prod p k pfk = transfer_dep (nth_prod_functional1 k' k (eq_lt_trans _ _ _ (eq_sym pfe) pfk) pfk t (eq_sym pfe)) (nth_tuple_prod p' k' _) (U := fun c => A c). intros; subst. pose proof (proof_irrelevance _ pfk (eq_lt_trans _ _ _ (eq_sym eq_refl) pfk)) as h2. rewrite <- h2, transfer_dep_eq_refl; auto. Qed. Lemma nth_tuple_prod_functional2 : forall {T n} {A:T->Type} {t:T^n} (p:tuple_prod A t) k k' (pfk:k < n) (pfe:k = k'), nth_tuple_prod p k pfk = transfer_dep (nth_prod_functional1 k' k (eq_lt_trans _ _ _ (eq_sym pfe) pfk) pfk t (eq_sym pfe)) (nth_tuple_prod p k' _) (U := fun c => A c). intros; subst; pose proof (proof_irrelevance _ pfk (eq_lt_trans _ _ _ (eq_sym eq_refl) pfk)) as h2; rewrite <- h2, transfer_dep_eq_refl; auto. Qed. Lemma nth_tuple_prod_functional3 : forall {T n} {U:T->Type} {t:T^n} (p:tuple_prod U t) k (pfk pfk':k < n), nth_tuple_prod p k pfk = transfer_dep (U := fun c => U c) (nth_prod_functional2 k pfk' pfk t) (nth_tuple_prod p k pfk'). intros T n U t p k hk hk'. pose proof (proof_irrelevance _ hk hk'); subst. rewrite transfer_dep_eq_refl; auto. Qed. Lemma transfer_tuple_prod_eq_iff : forall {T V n} {U:T->Type} {W:V->Type} {w z:T^n} {v:V^n} (pfew:pointwise_ind_eq U W w v) (pfez:pointwise_ind_eq U W z v) (t:tuple_prod U w) (u:tuple_prod U z) (pfewz:w = z), transfer_tuple_prod pfew t = transfer_tuple_prod pfez u <-> (forall i (pfi:i < n), transfer (pfew i pfi) (nth_tuple_prod t i pfi) = transfer (pfez i pfi) (nth_tuple_prod u i pfi)). intros T V n. induction n as [|n ih]; simpl. intros; subst. split; intros; try omega; auto. intros U W w z v hew hez t u h1; subst. pose proof (proof_irrelevance _ hez hew); subst. destruct z as [z z'], v as [v v'], u as [u u'], t as [t t']; simpl; simpl in hew, u, t, u', t'. pose proof hew as h1. apply pointwise_ind_eq_to_S in h1; simpl in h1. specialize (ih U W z' z' v' (pointwise_ind_eq_to_S hew) (pointwise_ind_eq_to_S hew) t' u' (eq_refl _)). split; intro h2. inversion h2; clear h2. rewrite ih in H1. intros i hi. destruct i as [|i]. pose proof (proof_irrelevance _ hi (lt_0_Sn _)); subst; auto. rename H1 into h2. specialize (h2 i (lt_S_n _ _ hi)). terml h2. redterm1 x. fold c in h2. gterml. redterm1 x0. fold c0. pose proof (proof_irrelevance _ c c0) as h3; rewrite <- h3. rewrite h2 at 1; auto. pose proof (h2 0 (lt_0_Sn _)) as h3; simpl in h3. rewrite h3; f_equal. apply ih. intros i hi. specialize (h2 _ (lt_n_S _ _ hi)); simpl in h2. do 2 erewrite (nth_tuple_prod_functional3 _ _ hi (lt_S_n _ _ (lt_n_S _ _ hi))). apply transfer_inj in h2. rewrite h2; auto. Qed. Fixpoint seg_fun_prod_type_to_tuple {n} : forall {f:seg_fun n Type} (g:seg_fun_prod_type f), tuple_prod (seg_fun_type_to_ind f) (full_segT_ind n) := match n with | 0 => fun _ _ => tt | S n' => fun f g => let f' := seg_fun_to_S f in let fg := transfer_r (seg_fun_type_to_ind0 f) (fst g) in let sg := transfer_tuple_prod (U := seg_fun_type_to_ind f') (pointwise_ind_eq_seg_fun_type_to_ind f) (seg_fun_prod_type_to_tuple (snd g) (n := n') (f := f')) in (fg, sg) end. Definition seg_fun_depT_nth {T n} (U:T->Type) (x:T^n) : Type := seg_fun_depT U (fun i (pfi:i < n) => nth_prod i pfi x). Definition seg_fun_depT_nth_to_S {T n} {U:T->Type} {x:T^S n} (f:seg_fun_depT_nth U x) : seg_fun_depT_nth U (snd x) := fun i (pfi:i < n) => transfer_dep (U := (fun c => U c)) (nth_prod_functional2 i (lt_S_n _ _ (lt_n_S _ _ pfi)) pfi (snd x)) (seg_fun_depT_to_S f i pfi). Lemma seg_fun_depT_nth_to_S_compat : forall {T n} {U:T->Type} (x:T^(S n)) (f:seg_fun_depT_nth U x) i (hi:S i < S n), seg_fun_depT_nth_to_S f i (lt_S_n i n hi) = f (S i) hi. unfold seg_fun_depT_nth_to_S; intros T n. intros U x f i hi. rewrite <- transfer_dep_eq_iff. unfold seg_fun_depT_to_S. pose proof (proof_irrelevance _ hi (lt_n_S _ _ (lt_S_n _ _ hi))) as h2. rewrite <- h2; auto. Qed. Fixpoint seg_fun_depT_nth_to_tuple_prod {T n} {U:T->Type} : forall (x:T^n) (g:seg_fun_depT_nth U x), tuple_prod U x := match n with | 0 => fun x g => tt | S n' => fun x g => let g' := seg_fun_depT_nth_to_S g in let sx := @seg_fun_depT_nth_to_tuple_prod T n' U (snd x) g' in (g 0 (lt_0_Sn _), sx) end. Lemma seg_fun_depT_nth_to_tuple_prod_S : forall {T n} {U:T->Type} (x:T^S n) (g:seg_fun_depT_nth U x), seg_fun_depT_nth_to_tuple_prod x g = (g 0 (lt_0_Sn _), seg_fun_depT_nth_to_tuple_prod (snd x) (seg_fun_depT_nth_to_S g)). auto. Qed. (*Wyckoff*) Definition transfer_tuple_prod_arity {T m n} {U:T->Type} {w:T^m} (pfe:m = n) (p:tuple_prod U w) : tuple_prod U (transfer_nprod pfe w) := seg_fun_depT_nth_to_tuple_prod (transfer_nprod pfe w) (fun i (pfi:i < n) => let pfi' := lt_eq_trans _ _ _ pfi (eq_sym pfe) in (transfer_dep_r (nth_prod_transfer_prod pfe w i pfi (U := U)) (nth_tuple_prod p i pfi'))). (*Wyckoff*) Definition transfer_tuple_prod_arity_r {T m n} {U:T->Type} {w:T^n} (pfe:m = n) (p:tuple_prod U w) : tuple_prod U (transfer_nprod_r pfe w) := seg_fun_depT_nth_to_tuple_prod (transfer_nprod_r pfe w) (fun i (pfi:i < m) => let pfi' := lt_eq_trans _ _ _ pfi pfe in (transfer_dep_r (nth_prod_transfer_prod_r pfe w i pfi (U := U)) (nth_tuple_prod p i pfi'))). (*Wyckoff*) Definition transfer_tuple_prod_eq_refl {T n} {U:T->Type} {w:T^n} (pfe:n = n) (p:tuple_prod U (transfer_nprod pfe w)) : tuple_prod U w := transfer_tuple_prod (eq_pointwise_ind_eq (transfer_nprod_eq_refl pfe w) U) p. (*Wyckoff*) Definition transfer_tuple_prod_arity_to_S_cast {T m n} {U:T->Type} {w:T^S m} (pfe:S m = S n) (p:tuple_prod U (transfer_nprod (S_inj _ _ pfe) (snd w))) : tuple_prod U (map_nprod (seg_fun_to_S (fun (i : nat) (pfi : i < S n) => match i as i0 return (i0 < S m -> T * T ^ m -> T) with | 0 => fun (_ : 0 < S m) (x : T * T ^ m) => fst x | S i' => fun (pflt : S i' < S m) (x : T * T ^ m) => nth_prod i' (lt_S_n i' m pflt) (snd x) end (lt_eq_trans i (S n) (S m) pfi (eq_sym pfe)) w))). unfold seg_fun_to_S; destruct w as [w w']; unfold transfer_nprod in p; simpl in p. erewrite f_equal. refine p. repeat f_equal; apply functional_extensionality_dep; intro; apply functional_extensionality_dep; intro; f_equal; apply proof_irrelevance. Defined. (*Wyckoff*) Definition transfer_tuple_prod_map_nprod_S_cast {T m n} {U:T->Type} {w:T^S m} (pfe:S m = S n) (p:tuple_prod U (map_nprod (seg_fun_to_S (fun (i : nat) (pfi : i < S n) => match i as i0 return (i0 < S m -> T * T ^ m -> T) with | 0 => fun (_ : 0 < S m) (x : T * T ^ m) => fst x | S i' => fun (pflt : S i' < S m) (x : T * T ^ m) => nth_prod i' (lt_S_n i' m pflt) (snd x) end (lt_eq_trans i (S n) (S m) pfi (eq_sym pfe)) w)))) : tuple_prod U (transfer_nprod (S_inj _ _ pfe) (snd w)). unfold seg_fun_to_S; destruct w as [w w']; unfold transfer_nprod in p; simpl in p. erewrite f_equal. refine p. unfold transfer_nprod. f_equal; apply functional_extensionality_dep; intro; apply functional_extensionality_dep; intro; unfold seg_fun_to_S; f_equal; apply proof_irrelevance. Defined. (*Wyckoff*) Lemma map_nprod_nth_prod_snd_eq_refl : forall {T n} (w:T^S n) (pfe:S n = S n), map_nprod (seg_fun_to_S (fun (i : nat) (pfi : i < S n) => match i as i0 return (i0 < S n -> T * T ^ n -> T) with | 0 => fun (_ : 0 < S n) (x : T * T ^ n) => fst x | S i' => fun (pflt : S i' < S n) (x : T * T ^ n) => nth_prod i' (lt_S_n i' n pflt) (snd x) end (lt_eq_trans i (S n) (S n) pfi (eq_sym pfe)) w)) = snd w. intros; rewrite map_nprod_eq_iff; unfold seg_fun_to_S; intros; f_equal; apply proof_irrelevance. Qed. (*Wyckoff*) Lemma transfer_dep_transfer_nprod_eq_refl_S : forall {T n} {U:T->Type} {w:T} {w':T^n} (pfe:S n = S n) (pfe':transfer_nprod pfe (w, w') = (w, w')) (p:U w), transfer_dep (U := fun c => U (fst c)) pfe' p = p. intros; erewrite f_equal_nprod_fst_fun; rewrite transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Definition transfer_tuple_prod_arity_to_S {T m n} {U:T->Type} {w:T^S m} (pfe:S m = S n) (p:tuple_prod U w) : tuple_prod U (transfer_nprod (S_inj _ _ pfe) (snd w)) := transfer_tuple_prod_arity (S_inj _ _ pfe) (snd p). (*Wyckoff*) Definition transfer_tuple_prod_arity_r_to_S {T m n} {U:T->Type} {w:T^S n} (pfe:S m = S n) (p:tuple_prod U w) : tuple_prod U (transfer_nprod_r (S_inj _ _ pfe) (snd w)) := transfer_tuple_prod_arity_r (S_inj _ _ pfe) (snd p). (*Wyckoff*) Lemma transfer_seg_fun_depT_nth_to_tuple_prod_nth_prod : forall {S T n} (w:S^n) (t:T^n), transfer_tuple_prod_const (T:=S) (U:=T) (seg_fun_depT_nth_to_tuple_prod (U := fun _ => T) w (fun (j : nat) (pfj : j < n) => nth_prod j pfj t)) = t. intros S T n. induction n as [|n ih]; simpl. intros; apply unit_eq. intros w t. destruct w as [fw sw], t as [ft st]; simpl. f_equal. specialize (ih sw st). rewrite <- ih at 1. repeat f_equal. apply functional_extensionality_dep. intro; apply functional_extensionality; intro. unfold seg_fun_depT_nth_to_S, seg_fun_depT_to_S. simpl. pose proof (proof_irrelevance _ x0 (lt_S_n _ _ (lt_n_S _ _ x0))) as h1. rewrite <- h1. gterml. redterm1 x1. rewrite transfer_dep_eq_refl; auto. Qed. (*Wyckoff*) Lemma seg_fun_depT_nth_to_tuple_prod_nth_prod : forall {S T n} (w:S^n) (t:T^n), (seg_fun_depT_nth_to_tuple_prod (U := fun _ => T) w (fun (j : nat) (pfj : j < n) => nth_prod j pfj t)) = transfer_to_tuple_prod_const w t. intros S T n w t. apply inj_transfer_tuple_prod_const. rewrite transfer_tuple_prod_const_to_compat. apply transfer_seg_fun_depT_nth_to_tuple_prod_nth_prod. Qed. (*Wyckoff*) Lemma seg_fun_depT_nth_to_tuple_prod_eq_iff : forall {T n} {U:T->Type} {x:T^n} (f:seg_fun_depT_nth U x) (w:tuple_prod U x), seg_fun_depT_nth_to_tuple_prod x f = w <-> (forall i (pfi:i < n), f i pfi = nth_tuple_prod w i pfi). intros T n. induction n as [|n ih]; simpl. intros; split; intros; subst. omega. apply unit_eq. intros U x f w. destruct x as [fx sx], w as [fw sw]; simpl. simpl in f, fw, sw. specialize (ih U sx (seg_fun_depT_nth_to_S f) sw). split; intro h1. inversion h1; subst. intro i. destruct i as [|i]. intros; f_equal; apply proof_irrelevance. intro hi. hl ih. auto. rewrite ih in hl. specialize (hl i (lt_S_n _ _ hi)). rewrite <- hl. unfold seg_fun_depT_nth_to_S; simpl. symmetry. rewrite <- transfer_dep_eq_iff. unfold seg_fun_depT_to_S. pose proof (proof_irrelevance _ hi (lt_n_S _ _ (lt_S_n _ _ hi))) as h2. rewrite <- h2; auto. pose proof (h1 0 (lt_0_Sn _)) as h2; simpl in h2. subst. f_equal. rewrite ih. intros i hi. specialize (h1 (S i) (lt_n_S _ _ hi)). simpl in h1. pose proof (proof_irrelevance _ hi (lt_S_n _ _ (lt_n_S _ _ hi))) as h3. rewrite h3, <- h1, seg_fun_depT_nth_to_S_compat; auto. Qed. End TupleProdAgain. bool2-v0-3/SetUtilities2_5.v0000644000175000017500000013467013575574055016501 0ustar dwyckoff76dwyckoff76(*Copyright (C) 2017, Daniel Wyckoff*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Version 8.4 pl4 Coq.*) Require Import SetUtilities. Require Import TypeUtilities. Require Import LogicUtilities. Require Import ProofIrrelevance. Require Import EqDec. Require Import ListUtilities. Require Import FunctionalExtensionality. Require Import SetUtilities1_5. Require Import SetUtilities2. Require Import ListUtilities2. Require Import ArithUtilities. Require Import FunctionProperties. Require Import Description. Require Import WellOrders. Section Injectors. Lemma finite_rel_classes_im_rel_sig_fun : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), Finite A -> Finite (rel_classes_im_rel_sig_fun f). intros T U A f h1. rewrite rel_classes_im_rel_eq_inv_im_singletons'. rewrite im_im. rewrite finite_full_sig_iff in h1. unfold rel_classes_im_rel_sig_fun. apply finite_image; auto. Qed. Lemma ex_same_im_subset_sig_inj : forall {T U:Type} {A:Ensemble T} (pfdt:type_eq_dec T), Finite A -> forall (f:sig_set A->U), exists (B:Ensemble T) (pf:Included B A), Inn (same_im_subsets_sig f) B /\ Injective (restriction_sig f B pf). intros T U A hdt h1 f. pose proof (finite_rel_classes_im_rel_sig_fun f h1) as h2. pose proof (finite_set_list_no_dup _ h2) as h3. destruct h3 as [l h3]. destruct h3 as [h3a h3b]. revert A h1 f h2 h3a h3b. induction l as [|B l h1]. intros A h1 f h2 h3 h4. simpl in h3. apply empty_rel_classes_im_rel_eq_inv_im_singletons in h3. subst. exists (Empty_set _). exists (inclusion_reflexive _). split. constructor. exists (inclusion_reflexive _). f_equal. apply functional_extensionality. intro x. contradict (proj2_sig x). red. intro x. contradict (proj2_sig x). intros A h2 f h3 h4 h5. pose proof (finite_mem_rel_classes_im_rel_sig_fun f h2) as hfm. assert (h16:In B (B::l)). auto with datatypes. rewrite list_to_set_in_iff in h16. rewrite <- h4 in h16. pose proof (partition_fam_rel_classes_im_rel_sig_fun f) as h17. red in h17. destruct h17 as [h17 h18]. assert (h19:Included B A). rewrite <- h18. red. intros x h19. destruct h16 as [S h16]. clear h18. subst. apply family_union_intro with (Im S (@proj1_sig _ _)). unfold rel_classes_im_rel_sig_fun. apply Im_intro with S; auto. assumption. pose proof (setminus_inc A B) as h6. assert (h7:Finite (Setminus A B)). apply Finite_downward_closed with A; auto. specialize (h1 (Setminus A B) h7 (restriction_sig f _ h6)). pose proof (finite_rel_classes_im_rel_sig_fun (restriction_sig f (Setminus A B) h6) h7) as h8. specialize (h1 h8). pose proof (no_dup_cons_nin _ _ h5) as h9. pose proof (no_dup_cons _ _ h5) as h10. assert (h11: rel_classes_im_rel_sig_fun (restriction_sig f (Setminus A B) h6) = list_to_set l). pose proof (rel_classes_im_rel_sig_fun_setminus A B f h16) as h12. pose proof (proof_irrelevance _ h6 (setminus_inc A B)) as h13. clear h18. subst. rewrite h12. rewrite h4. red in hfm. simpl in h4. rewrite h4 in hfm. rewrite (eq_refl (list_to_set (B::l))) in hfm at 1. rewrite <- p_mem_l_iff in hfm at 1. pose proof (subtract_remove_ens_compat hdt (B::l) B (Finite_downward_closed _ _ h2 _ h19) hfm) as hs. simpl in hs; rewrite sub_add_eq in hs. destruct fin_ens_eq_dec as [hf | hf]. pose proof @remove_dep_not_in. pose proof (p_mem_l_cons _ _ _ hfm) as hfm'. pose proof (remove_dep_not_in (@Finite T) (@fin_ens_eq_dec T hdt) _ B (Finite_downward_closed _ _ h2 _ h19) hfm' h9) as hr. rewrite <- hr . simpl. rewrite sub_add_eq. rewrite subtract_nin; auto. intro h6. rewrite <- list_to_set_in_iff in h6. rewrite hr in h6. contradiction. pose proof h12 as h12'. clear h12. rewrite h4 in h12'. rewrite sub_add_eq in h12'. rewrite subtract_nin in h12'. simpl. rewrite sub_add_same_nin; auto. rewrite <- list_to_set_in_iff; auto. specialize (h1 h11 h10). destruct h1 as [C h1]. destruct h1 as [h12 [h13 h14]]. pose proof (inh_rel_classes_im_rel_sig_fun f _ h16) as h20. destruct h20 as [b h20]. exists (Add C b). assert (h21:Included (Add C b) A). red. intros x h21. destruct h21 as [x h21l | x h21r]. auto with sets. destruct h21r. auto with sets. exists h21. split. constructor. exists h21. destruct h13 as [h13]. destruct h13 as [h13 h22]. rewrite im_full_sig_add. unfold restriction_sig. simpl. unfold restriction_sig in h22. simpl in h22. assert (h24: (fun x : sig_set C => f (exist (Inn A) (proj1_sig x) (h6 (proj1_sig x) (h13 (proj1_sig x) (proj2_sig x))))) = (fun x : sig_set C => f (exist (Inn A) (proj1_sig x) (h21 (proj1_sig x) (incl_add C b (proj1_sig x) (proj2_sig x)))))). apply functional_extensionality. intro x. f_equal. apply proj1_sig_injective. simpl. reflexivity. rewrite h24 in h22. rewrite <- h22. unfold restriction_sig in h14. simpl in h14. pose proof (full_sig_decompose_setminus_incl _ _ h19) as h25. rewrite h25 at 1. pose proof (im_elt_of_rel_classes_im_rel_sig_fun f B h19 h16 _ h20) as h26. rewrite im_union. do 2 rewrite im_im. rewrite h26 at 1. unfold Add. f_equal. f_equal. apply functional_extensionality. intro x. f_equal. apply proj1_sig_injective. simpl. reflexivity. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. unfold restriction_sig in h14. simpl in h14. unfold restriction_sig. simpl. red. intros x y h22. red in h14. destruct x as [x h23], y as [y h24]. simpl in h22. apply proj1_sig_injective. simpl. destruct h23 as [x h23a | x h23b], h24 as [y h24a | y h24b]. specialize (h14 (exist _ _ h23a) (exist _ _ h24a)). simpl in h14. assert (h25:h6 x (h12 x h23a) = h21 x (Union_introl T C (Singleton b) x h23a)). apply proof_irrelevance. assert (h26:h6 y (h12 y h24a) = h21 y (Union_introl T C (Singleton b) y h24a)). apply proof_irrelevance. rewrite h25, h26 in h14. specialize (h14 h22). apply exist_injective in h14. assumption. destruct h24b. assert (h24:Inn B x). symmetry in h22. pose proof (rel_classes_im_rel_sig_fun_consistent_membership f B h16 _ _ h22 h20) as h24. simpl in h24. assumption. pose proof (h12 _ h23a) as h25. destruct h25. contradiction. destruct h23b. assert (h24:Inn B y). pose proof (rel_classes_im_rel_sig_fun_consistent_membership f B h16 _ _ h22 h20) as h24. simpl in h24. assumption. pose proof (h12 _ h24a) as h26. destruct h26. contradiction. destruct h23b, h24b. reflexivity. Qed. Lemma finite_inhabited_sig_fun_injectors : forall {T U:Type} {A:Ensemble T} (pfd:type_eq_dec T) (f:sig_set A->U), Finite A -> Inhabited (sig_fun_injectors f). intros T U A h0 f h1. unfold sig_fun_injectors. pose proof (ex_same_im_subset_sig_inj h0 h1 f) as h2. destruct h2 as [B [h2 [h3 h4]]]. apply Inhabited_intro with B. constructor. exists h2. split; auto. Qed. Lemma finite_sig_fun_injectors : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), Finite A -> Finite (sig_fun_injectors f). intros T U A f h1. eapply Finite_downward_closed. eapply power_set_finite. apply h1. red; intros S h2. constructor. destruct h2 as [h2]. destruct h2; auto. Qed. Lemma finite_sig_fun_injectors_fam : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), Finite A -> forall S:Ensemble T, Inn (sig_fun_injectors f) S -> Finite S. intros T U A f h1 S h2. destruct h2 as [h2]. destruct h2 as [h2 [h3 h4]]. destruct h4 as [h4]. destruct h4 as [h4 h5]. eapply Finite_downward_closed. apply h1. assumption. Qed. Definition nat_sig_fun_injector {U:Type} {A:Ensemble nat} (f:sig_set A->U) (pf:Finite A) : Ensemble nat := can_set_nat_choice (sig_fun_injectors f) (finite_sig_fun_injectors f pf) (finite_inhabited_sig_fun_injectors nat_eq_dec f pf) (finite_sig_fun_injectors_fam f pf). Definition in_nat_sig_fun_injector : forall {U:Type} {A:Ensemble nat} (f:sig_set A->U) (pf:Finite A), Inn (sig_fun_injectors f) (nat_sig_fun_injector f pf). intros U A f h1. unfold nat_sig_fun_injector. apply (in_can_set_nat_choice (sig_fun_injectors f) (finite_sig_fun_injectors f h1) (finite_inhabited_sig_fun_injectors nat_eq_dec f h1) (finite_sig_fun_injectors_fam f h1)). Qed. Lemma incl_nat_sig_fun_injector : forall {U:Type} {A:Ensemble nat} (f:sig_set A->U) (pf:Finite A), Included (nat_sig_fun_injector f pf) A. intros U A f h1. pose proof (in_nat_sig_fun_injector f h1) as h2. destruct h2 as [h2]. destruct h2; auto. Qed. Lemma finite_nat_sig_fun_injector : forall {U:Type} {A:Ensemble nat} (f:sig_set A->U) (pf:Finite A), Finite (nat_sig_fun_injector f pf). intros U A f h1. unfold nat_sig_fun_injector. unfold can_set_nat_choice. apply list_to_set_finite. Qed. Lemma nat_sig_fun_injector_zero : forall {U:Type} (p:sig_set (seg 1)->U), nat_sig_fun_injector p (finite_seg 1) = seg 1. intros U p. pose proof seg_1_eq as h1. rewrite h1. unfold nat_sig_fun_injector, can_set_nat_choice. match goal with |- list_to_set (min_set_sortable ?x1 ?x2 ?x3 ?x4 ?x5) = _ => pose proof (min_set_sortable_compat x1 x2 x3 x4 x5) as h2 end. simpl in h2. destruct h2 as [h2 h3]. inversion h2. subst. rewrite H0 at 1. unfold sort_set_nat, sort_set. match goal with |- list_to_set (min_set_sortable ?x1 ?x2 ?x3 ?x4 ?x5) = _ => pose proof (min_set_sortable_compat x1 x2 x3 x4 x5) as h5 end. simpl in h5. destruct h5 as [h5 h6]. inversion h5. destruct H1 as [h7 h8]. rewrite <- h7 at 1. destruct x as [x h9]. simpl. destruct h9 as [h9 h10]. simpl in h7. destruct h9 as [h11 h12]. pose proof h11 as h11'. rewrite h1 in h11'. apply singleton_inc in h11'. destruct h11'; subst. symmetry in h7. apply empty_set_nil in h7. simpl in h5, h7, h8, h6. destruct h12 as [h12 h13]. destruct h13 as [h13]. destruct h13 as [h13 h14]. clear h8 h6 h5 h7 H0 H. assert (h15:Inn (seg 1) 0). constructor; auto. assert (h16:Inn (Im (full_sig (seg 1)) p) (p (exist _ _ h15))). apply Im_intro with (exist _ _ h15). constructor. reflexivity. rewrite h14 in h16. inversion h16. subst. destruct x. contradiction. reflexivity. Qed. Lemma nat_sig_fun_injector_one : forall {U:Type} (p:sig_set (seg_one 1)->U), nat_sig_fun_injector p (finite_seg_one 1) = seg_one 1. intros U p. assert (h1:seg_one 1 = Singleton 1). apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1]. assert (h2:x = 1). omega. subst. constructor. destruct h1. constructor. auto with arith. rewrite h1. unfold nat_sig_fun_injector. unfold can_set_nat_choice. match goal with |- list_to_set (min_set_sortable ?x1 ?x2 ?x3 ?x4 ?x5) = _ => pose proof (min_set_sortable_compat x1 x2 x3 x4 x5) as h2 end. simpl in h2. destruct h2 as [h2 h3]. inversion h2. subst. rewrite H0 at 1. unfold sort_set_nat, sort_set. match goal with |- list_to_set (min_set_sortable ?x1 ?x2 ?x3 ?x4 ?x5) = _ => pose proof (min_set_sortable_compat x1 x2 x3 x4 x5) as h5 end. simpl in h5. destruct h5 as [h5 h6]. inversion h5. destruct H1 as [h7 h8]. rewrite <- h7 at 1. destruct x as [x h9]. simpl. destruct h9 as [h9 h10]. simpl in h7. destruct h9 as [h11 h12]. pose proof h11 as h11'. rewrite h1 in h11'. apply singleton_inc in h11'. destruct h11'; subst. symmetry in h7. apply empty_set_nil in h7. simpl in h5, h7, h8, h6. destruct h12 as [h12 h13]. destruct h13 as [h13]. destruct h13 as [h13 h14]. clear h8 h6 h5 h7 H0 H. assert (h15:Inn (seg_one 1) 1). constructor; auto. assert (h16:Inn (Im (full_sig (seg_one 1)) p)(p (exist _ _ h15))). apply Im_intro with (exist _ _ h15). constructor. reflexivity. rewrite h14 in h16. inversion h16. subst. destruct x. contradiction. reflexivity. Qed. Definition segT_injector {U:Type} {n:nat} (f:n ---> U) : Ensemble (n @). pose (fun_st_conv f) as f'. pose (nat_sig_fun_injector f' (finite_seg _)) as S. pose proof (incl_nat_sig_fun_injector f' (finite_seg _)) as h1. pose (im_proj2_sig _ h1) as S'. refine (ens_ss_conv S'). Defined. Lemma incl_segT_injector : forall {U:Type} {n:nat} (f:n ---> U), Included (segT_injector f) (Full_set _). intros; constructor. Qed. Lemma segT_injector_compat : forall {U:Type} {n:nat} (f:n ---> U), let S := segT_injector f in Im (Full_set (n @)) f = Im S f /\ injective (restriction_fun f S). intros U n f S. pose (segT_injector f) as f'. unfold segT_injector in f'. pose proof (in_nat_sig_fun_injector (fun_st_conv f) (finite_seg n)) as h1. destruct h1 as [h1]. destruct h1 as [h1 [h2 h3]]. destruct h3 as [h3]. destruct h3 as [h3 h4]. assert (h1 = h3). apply proof_irrelevance. subst. split. apply Extensionality_Ensembles; red; split; red; intros x h5. destruct h5 as [x h5 ? h6]. subst. clear h5. destruct x as [m h5]. assert (h6:Inn (Im (full_sig (seg n)) (fun_st_conv f)) (f (exist (fun a => a < n) _ h5))). assert (h7:Inn (seg n) m). constructor; auto. apply Im_intro with (exist _ _ h7). constructor. unfold fun_st_conv. f_equal. apply proj1_sig_injective; auto. unfold S, segT_injector, ens_ss_conv. rewrite im_im. unfold im_proj2_sig. rewrite im_im. simpl. let P := type of h4 in match P with _ = ?x => pose x as A end. match goal with |- Inn ?S _ => pose S as B end. assert (h7:A = B). unfold A, B. apply im_ext_in. intros x h8. unfold restriction_sig, fun_st_conv. f_equal. apply proj1_sig_injective; auto. fold B. rewrite <- h7. rewrite h4 in h6. assumption. destruct h5 as [x h5]. subst. apply Im_intro with x. constructor. reflexivity. red. intros x y h1. red in h2. unfold S, segT_injector, ens_ss_conv in x, y. unfold im_proj2_sig in x, y. destruct x as [x h5], y as [y h6]. apply proj1_sig_injective; simpl. unfold restriction_fun in h1. simpl in h1. rewrite im_im in h5, h6. destruct h5 as [x h5], h6 as [y h6]; subst. apply proj1_sig_injective; simpl. simpl in h1. specialize (h2 x y). unfold restriction_sig, fun_sot_conv in h2. simpl in h2. assert (h7:h3 (proj1_sig x) (proj2_sig x)= incl_nat_sig_fun_injector (fun_st_conv f) (finite_seg n) (proj1_sig x) (proj2_sig x)). apply proof_irrelevance. assert (h8:h3 (proj1_sig y) (proj2_sig y)= incl_nat_sig_fun_injector (fun_st_conv f) (finite_seg n) (proj1_sig y) (proj2_sig y)). apply proof_irrelevance. rewrite <- h7, <- h8 in h1 . specialize (h2 h1). subst. reflexivity. Qed. Lemma finite_segT_injector : forall {U:Type} {n:nat} (f:n ---> U), Finite (segT_injector f). intros U n f. unfold segT_injector, ens_ss_conv, im_proj2_sig. rewrite im_im. simpl. eapply finite_image. pose proof (finite_nat_sig_fun_injector (fun_st_conv f) (finite_seg n)) as h1. apply (iff1 (finite_full_sig_iff (nat_sig_fun_injector (fun_st_conv f) (finite_seg n))) h1). Qed. Lemma segT_injector_zero : forall {U:Type} (p:1 @->U), segT_injector p = Singleton st01. intros U p. pose proof (segT_injector_compat p) as h1. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct x as [x h3]. assert (h4:x = 0). omega. subst. unfold st01. assert (h4: h3 = st01_pf). apply proof_irrelevance. subst. constructor. destruct x as [x h3]. assert (h4: x = 0). omega. subst. simpl in h1. unfold segT_injector, ens_ss_conv, im_proj2_sig. rewrite im_im. simpl. pose proof (nat_sig_fun_injector_zero (fun_st_conv p)) as h4. assert (h5:Inn (seg 1) 0). constructor. auto with arith. rewrite <- h4 in h5. apply Im_intro with (exist _ _ h5). constructor. apply proj1_sig_injective; simpl. reflexivity. Qed. Definition seg_one_t_injector {U:Type} {n:nat} (f:seg_one_t n->U) : Ensemble (seg_one_t n). pose (fun_sot_conv f) as f'. pose (nat_sig_fun_injector f' (finite_seg_one n)) as S. pose proof (incl_nat_sig_fun_injector f' (finite_seg_one n)) as h1. pose (im_proj2_sig _ h1) as S'. refine (ens_sso_conv S'). Defined. Lemma incl_seg_one_t_injector : forall {U:Type} {n:nat} (f:seg_one_t n -> U), Included (seg_one_t_injector f) (Full_set _). intros; constructor. Qed. Lemma seg_one_t_injector_compat : forall {U:Type} {n:nat} (f:seg_one_t n->U), let S := seg_one_t_injector f in Im (Full_set (seg_one_t n)) f = Im S f /\ injective (restriction_fun f S). intros U n f S. pose (seg_one_t_injector f) as f'. unfold seg_one_t_injector in f'. pose proof (in_nat_sig_fun_injector (fun_sot_conv f) (finite_seg_one n)) as h1. destruct h1 as [h1]. destruct h1 as [h1 [h2 h3]]. destruct h3 as [h3]. destruct h3 as [h3 h4]. assert (h1 = h3). apply proof_irrelevance. subst. split. apply Extensionality_Ensembles; red; split; red; intros x h5. destruct h5 as [x h5 ? h6]. subst. clear h5. destruct x as [m h5]. assert (h6:Inn (Im (full_sig (seg_one n)) (fun_sot_conv f)) (f (exist _ _ h5))). assert (h7:Inn (seg_one n) m). constructor; auto. apply Im_intro with (exist _ _ h7). constructor. unfold fun_sot_conv. f_equal. apply proj1_sig_injective; auto. unfold S, seg_one_t_injector, ens_sso_conv. rewrite im_im. unfold im_proj2_sig. rewrite im_im. simpl. let P := type of h4 in match P with _ = ?C => match goal with |- Inn ?D _ => assert (h7:C = D) end end. apply im_ext_in. intros x h8. unfold restriction_sig, fun_sot_conv. f_equal. apply proj1_sig_injective; auto. rewrite <- h7. rewrite h4 in h6. assumption. destruct h5 as [x h5]. subst. apply Im_intro with x. constructor. reflexivity. red. intros x y h1. red in h2. unfold S, seg_one_t_injector, ens_sso_conv in x, y. unfold im_proj2_sig in x, y. destruct x as [x h5], y as [y h6]. apply proj1_sig_injective; simpl. unfold restriction_fun in h1. simpl in h1. rewrite im_im in h5, h6. destruct h5 as [x h5], h6 as [y h6]; subst. apply proj1_sig_injective; simpl. simpl in h1. specialize (h2 x y). unfold restriction_sig, fun_sot_conv in h2. simpl in h2. hfirst h2. match goal with |- ?x = _ => let Q := type of h1 in match Q with ?y = _ => assert (h7:x = y) end end. f_equal. f_equal. apply proof_irrelevance. match goal with |- _ = ?x => let Q := type of h1 in match Q with _ = ?y => assert (h8:x = y) end end. f_equal. f_equal. apply proof_irrelevance. rewrite h7, h8; auto. specialize (h2 hf). subst. reflexivity. Qed. Lemma finite_seg_one_t_injector : forall {U:Type} {n:nat} (f:seg_one_t n->U), Finite (seg_one_t_injector f). intros U n f. unfold seg_one_t_injector, ens_sso_conv, im_proj2_sig. rewrite im_im. simpl. eapply finite_image. pose proof (finite_nat_sig_fun_injector (fun_sot_conv f) (finite_seg_one n)) as h1. apply (iff1 (finite_full_sig_iff (nat_sig_fun_injector (fun_sot_conv f) (finite_seg_one n))) h1). Qed. Lemma seg_one_t_injector_one : forall {U:Type} (p:seg_one_t 1->U), seg_one_t_injector p = Singleton sot_1. intros U p. pose proof (seg_one_t_injector_compat p) as h1. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct x as [x h3]. assert (h4:x = 1). omega. subst. unfold sot_1. assert (h4: h3 = conj (le_n 1) (le_n 1)). apply proof_irrelevance. subst. constructor. destruct x as [x h3]. assert (h4: x = 1). omega. subst. simpl in h1. unfold seg_one_t_injector, ens_sso_conv, im_proj2_sig. rewrite im_im. simpl. pose proof (nat_sig_fun_injector_one (fun_sot_conv p)) as h4. assert (h5:Inn (seg_one 1) 1). constructor. auto with arith. rewrite <- h4 in h5. apply Im_intro with (exist _ _ h5). constructor. apply proj1_sig_injective; simpl. reflexivity. Qed. Lemma seg_one_t_2_injector_eq_pos_evens_fun_unite_finite_odds : forall {U:Type} (f:seg_one_t 2->U), f sot12 = f sot22 -> seg_one_t_injector f = Singleton sot12 \/ seg_one_t_injector f = Singleton sot22. intros U f h1. pose proof (seg_one_t_injector_compat f) as h2. simpl in h2. destruct h2 as [h2 h3]. assert (h4 : Im (Full_set (seg_one_t 2))f = Couple (f sot12) (f sot22)). rewrite full_set_seg_one_t_2. rewrite im_couple. reflexivity. rewrite h2 in h4. rewrite h1 in h4. rewrite couple_singleton in h4. pose proof (f_equal card_fun1 h4) as h5. rewrite card_fun1_sing in h5. rewrite <- h2 in h5. pose proof (card_fun1_compat (Im (Full_set (seg_one_t 2)) f)) as ha. simpl in ha. destruct ha as [ha h7]. clear h7. assert (h6: Finite (Im (Full_set (seg_one_t 2)) f)). apply finite_image. apply finite_full_set_seg_one_t. specialize (ha h6). pose proof (card_fun1_compat (seg_one_t_injector f)) as h6'. simpl in h6'. destruct h6' as [h7 h8]. clear h8. specialize (h7 (finite_seg_one_t_injector f)). rewrite h5 in ha. pose proof (injective_preserves_card_fun1 _ h3 (Full_set (sig_set (seg_one_t_injector f)))) as h9. assert (h8 : Finite (Full_set (sig_set (seg_one_t_injector f)))). rewrite Finite_FiniteT_iff. apply Finite_ens_type. apply finite_seg_one_t_injector. specialize (h9 h8). assert (h11 : Im (Full_set (sig_set (seg_one_t_injector f))) (restriction_fun f (seg_one_t_injector f)) = Im (seg_one_t_injector f) f). apply Extensionality_Ensembles; red; split; red; intros x h11. destruct h11 as [x h11]. subst. unfold restriction_fun. apply Im_intro with (proj1_sig x). apply proj2_sig. reflexivity. destruct h11 as [x h11]. subst. apply Im_intro with (exist _ _ h11). constructor. unfold restriction_fun. simpl. reflexivity. pose proof h9 as h9'. rewrite h11 in h9'. rewrite h2 in h5. rewrite h5 in h9' at 1. pose proof (card_fun1_full_sig (seg_one_t_injector f) (finite_seg_one_t_injector _)) as h12. unfold full_sig in h12. rewrite h9' in h12. symmetry in h12. apply card_fun1_1 in h12. destruct h12 as [x h12]. destruct x as [x h13]. destruct h13 as [h13 h14]. pose proof h13 as h13'. apply le_lt_eq_dec in h13'. destruct h13' as [h15 | h15]. assert (h16: x = 2). auto with arith. subst. right. rewrite h12. f_equal. apply proj1_sig_injective; auto. subst. left. rewrite h12. f_equal. apply proj1_sig_injective; auto. Qed. End Injectors. (*In this section I develop two ways to get constructive description on finite and countable types without depending on an indefinite description. I'm sure the finite case fits the bill. It remains to be seen if any of the prerequisites for this use indefinite description.*) Section CountableDefiniteDescription. Lemma bij_seg_inh_iff : forall {T:Type} (A:Ensemble T), Finite A -> forall (f:sig_set (seg (card_fun1 A)) -> sig_set A), bijective f -> forall (P:sig_set A->Prop), (Inhabited [x:sig_set A | P x] <-> Inhabited [n : nat | exists pf : Inn (seg (card_fun1 A)) n, P (f (exist _ _ pf))]). intros T A h1 f h2 P. destruct h2 as [h2a h2b]. pose proof (card_fun1_full_sig A h1) as h3. symmetry in h3. pose proof (f_equal seg h3) as h4. pose (transfer_fun (sig_set_eq _ _ h4) f) as f'. pose (fun x => exist _ _ (Full_intro _ (f' x))) as f''. assert (h5:bijective f''). unfold f''. red. split. red. intros x y h6. apply exist_injective in h6. unfold f' in h6. rewrite <- (transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) x) in h6. rewrite <- (transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) y) in h6. do 2 rewrite transfer_fun_compat in h6. do 2 rewrite transfer_r_sig_set_eq in h6 at 1. apply h2a in h6. apply exist_injective in h6. apply proj1_sig_injective in h6. assumption. red. intro y. destruct y as [y h8]. pose proof (h2b y) as h9. destruct h9 as [x h9]. subst. exists (transfer (sig_set_eq _ _ h4) x). apply proj1_sig_injective. simpl. pose proof (transfer_sig_set_eq _ _ h4 (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) x) as h9. rewrite h9 at 1. unfold f'. rewrite transfer_fun_eq; f_equal. rewrite (transfer_r_sig_set_eq _ _ h4 (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4)) at 1. simpl. apply proj1_sig_injective; auto. pose proof (bij_seg_inh_full_iff f'' h5 (fun x:sig_set A=>P x)) as h6. rewrite h6 at 1. split. intro h7. destruct h7 as [a h7]. destruct h7 as [h7]. apply Inhabited_intro with a. constructor. destruct h7 as [h7 h8]. unfold full_sig in h4. pose proof h7 as h7'. rewrite <- h4 in h7'. exists h7'. unfold f'' in h8. simpl in h8. unfold f' in h8. rewrite transfer_fun_eq in h8. erewrite f_equal. apply h8. rewrite (transfer_r_sig_set_eq _ _ h4). repeat f_equal; apply proof_irrelevance. intro h7. destruct h7 as [n h7]. destruct h7 as [h7]. apply Inhabited_intro with n. constructor. destruct h7 as [h7 h8]. pose proof h7 as h7'. rewrite h4 in h7'. exists h7'. unfold f'', f'. simpl. pose proof (transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) (exist _ _ h7')) as h9. rewrite <- h9 at 1. rewrite transfer_fun_compat. rewrite (transfer_r_sig_set_eq _ _ h4 (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4)). simpl. match goal with |- P (f ?x) => let D := type of h8 in match D with P (f ?y) => assert (h10: x = y) end end. apply proj1_sig_injective; auto. rewrite h10. assumption. Grab Existential Variables. f_equal. assumption. f_equal. assumption. Qed. Lemma bij_seg_finite_iff : forall {T:Type} (A:Ensemble T), Finite A -> forall (f:sig_set (seg (card_fun1 A)) -> sig_set A), bijective f -> forall (P:sig_set A->Prop), (Finite [x:sig_set A | P x] <-> Finite [n : nat | exists pf : Inn (seg (card_fun1 A)) n, P (f (exist _ _ pf))]). intros T A h1 f h2 P. destruct h2 as [h2a h2b]. pose proof (card_fun1_full_sig A h1) as h3. symmetry in h3. pose proof (f_equal seg h3) as h4. pose (transfer_fun (sig_set_eq _ _ h4) f) as f'. pose (fun x => exist _ _ (Full_intro _ (f' x))) as f''. assert (h5:bijective f''). unfold f''. red. split. red. intros x y h6. apply exist_injective in h6. unfold f' in h6. rewrite <- (transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4)x) in h6. rewrite <- (transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) y) in h6. do 2 rewrite transfer_fun_compat in h6. do 2 rewrite transfer_r_sig_set_eq in h6 at 1. apply h2a in h6. apply exist_injective in h6. apply proj1_sig_injective in h6. assumption. red. intro y. destruct y as [y h8]. pose proof (h2b y) as h9. destruct h9 as [x h9]. subst. exists (transfer (sig_set_eq _ _ h4) x). apply proj1_sig_injective. simpl. pose proof (transfer_sig_set_eq _ _ h4 (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) x) as h9. rewrite h9 at 1. unfold f'. rewrite transfer_fun_eq, (transfer_r_sig_set_eq _ _ h4 (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4)). f_equal. apply proj1_sig_injective; auto. pose proof (bij_seg_finite_full_iff f'' h5 (fun x:sig_set A=>P x)) as h6. split. intro h0. rewrite h6 in h0. eapply Finite_downward_closed. apply h0. red. intros x h7. destruct h7 as [h7]. constructor. destruct h7 as [h7 h8]. pose proof h7 as h7'. rewrite h4 in h7'. exists h7'. unfold f'', f'. simpl. rewrite transfer_fun_eq, (transfer_r_sig_set_eq _ _ h4). simpl. erewrite f_equal. apply h8. repeat f_equal. apply proof_irrelevance. intro h7. rewrite h6. eapply Finite_downward_closed. apply h7. red; intros x h8. destruct h8 as [h8]. constructor. destruct h8 as [h8 h9]. pose proof h8 as h7'. unfold full_sig in h4. rewrite <- h4 in h7' at 1. exists h7'. unfold f'' in h9. simpl in h9. unfold f' in h9. rewrite transfer_fun_eq, (transfer_r_sig_set_eq _ _ h4) in h9. simpl in h9. erewrite f_equal. apply h9. repeat f_equal; apply proof_irrelevance. Grab Existential Variables. f_equal; auto. f_equal; auto. Qed. Lemma finite_constructive_definite_description : forall {T:Type} (A:Ensemble T) (pf:Finite A) (f:sig_set (seg (card_fun1 A)) -> sig_set A), bijective f -> forall (P:sig_set A->Prop), (exists x:sig_set A, P x) -> {x:sig_set A| P x}. intros T A h1 f h2 P h3. pose proof (ex_inh _ h3) as h4. pose proof (bij_seg_inh_iff _ h1 f h2 ) as h5. rewrite h5 in h4. pose [n : nat | exists pf : Inn (seg (card_fun1 A)) n, P (f (exist (Inn (seg (card_fun1 A))) n pf))] as S. pose proof (bij_seg_finite_iff _ h1 f h2) as h6. assert (ha:Finite [x : sig_set A | P x]). eapply Finite_downward_closed. rewrite <- finite_full_sig_iff. assumption. red; intros; constructor. specialize (h6 P). rewrite h6 in ha. pose (max_set_nat _ ha h4). pose proof (max_set_nat_compat _ ha h4) as h10. simpl in h10. destruct h10 as [h10 h11]. destruct h10 as [h10]. pose proof (destruct_exists_pf_set _ _ _ h10) as h12. simpl in h12. refine (exist _ _ h12). Defined. Lemma finite_constructive_definite_description_alt : forall {T:Type} (A:Ensemble T) (pf:Finite A) (f:(card_fun1 A) @-> sig_set A), bijective f -> forall (P:sig_set A->Prop), (exists x:sig_set A, P x) -> {x:sig_set A| P x}. intros T A h1 f h2 P h3. eapply finite_constructive_definite_description; auto. rewrite bij_st_conv_iff in h2 at 1. apply h2. Defined. (*This is to be used in situations that don't allow propositional case analysis, like in a function.*) Lemma surj_fin_inv : forall {T U:Type} {A:Ensemble T} (f:sig_set (seg (card_fun1 A)) -> sig_set A), Finite A -> bijective f -> forall (g:sig_set A->U), surjective g -> forall (y:U), Inn (Im (Full_set _) g) y -> {x:sig_set A | g x = y}. intros T U A f h1 h0 g h2 y h3. apply Im_inv in h3. pose proof (finite_constructive_definite_description A h1 f h0 _ h3) as h4. simpl in h4. destruct h4 as [x h4]. destruct h4 as [? h4]. refine (exist _ _ h4). Defined. Lemma bij_seg_inh_iff' : forall {T:Type} (A:Ensemble T), Finite A -> forall (f:sig_set (seg (card_fun1 A)) -> sig_set A), bijective f -> forall (P:T->Prop), (Inhabited [x:sig_set A | P (proj1_sig x)] <-> Inhabited [n : nat | exists pf : Inn (seg (card_fun1 A)) n, P (proj1_sig (f (exist _ _ pf)))]). intros T A h1 f h2 P. destruct h2 as [h2a h2b]. pose proof (card_fun1_full_sig A h1) as h3. symmetry in h3. pose proof (f_equal seg h3) as h4. pose (transfer_fun (sig_set_eq _ _ h4) f) as f'. pose (fun x => exist _ _ (Full_intro _ (f' x))) as f''. assert (h5:bijective f''). unfold f''. red. split. red. intros x y h6. apply exist_injective in h6. unfold f' in h6. rewrite <- (transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) x) in h6. rewrite <- (transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) y) in h6. do 2 rewrite transfer_fun_compat in h6. do 2 rewrite transfer_r_sig_set_eq in h6 at 1. apply h2a in h6. apply exist_injective in h6. apply proj1_sig_injective in h6. assumption. red. intro y. destruct y as [y h8]. pose proof (h2b y) as h9. destruct h9 as [x h9]. subst. exists (transfer (sig_set_eq _ _ h4) x). apply proj1_sig_injective. simpl. pose proof (transfer_sig_set_eq _ _ h4 (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) x) as h9. rewrite h9 at 1. unfold f'. rewrite transfer_fun_eq, (transfer_r_sig_set_eq _ _ h4). simpl. f_equal. apply proj1_sig_injective; auto. pose proof (bij_seg_inh_full_iff f'' h5 (fun x:sig_set A=>P (proj1_sig x))) as h6. rewrite h6 at 1. split. intro h7. destruct h7 as [a h7]. destruct h7 as [h7]. apply Inhabited_intro with a. constructor. destruct h7 as [h7 h8]. unfold full_sig in h4. pose proof h7 as h7'. rewrite <- h4 in h7'. exists h7'. unfold f'' in h8. simpl in h8. unfold f' in h8. rewrite transfer_fun_eq, (transfer_r_sig_set_eq _ _ h4) in h8; simpl in h8. erewrite f_equal. apply h8. repeat f_equal; apply proof_irrelevance. intro h7. destruct h7 as [n h7]. destruct h7 as [h7]. apply Inhabited_intro with n. constructor. destruct h7 as [h7 h8]. pose proof h7 as h7'. rewrite h4 in h7'. exists h7'. unfold f'', f'. simpl. pose proof (transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) (exist _ _ h7')) as h9. rewrite <- h9 at 1. rewrite transfer_fun_compat. rewrite (transfer_r_sig_set_eq _ _ h4 (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4)). simpl. match goal with |- P (proj1_sig (f ?x)) => let D := type of h8 in match D with P (proj1_sig (f ?y)) => assert (h10: x = y) end end. f_equal. apply proof_irrelevance. rewrite h10. assumption. Grab Existential Variables. f_equal. assumption. f_equal. assumption. Qed. Lemma bij_seg_finite_iff' : forall {T:Type} (A:Ensemble T), Finite A -> forall (f:sig_set (seg (card_fun1 A)) -> sig_set A), bijective f -> forall (P:T->Prop), (Finite [x:sig_set A | P (proj1_sig x)] <-> Finite [n : nat | exists pf : Inn (seg (card_fun1 A)) n, P (proj1_sig (f (exist _ _ pf)))]). intros T A h1 f h2 P. destruct h2 as [h2a h2b]. pose proof (card_fun1_full_sig A h1) as h3. symmetry in h3. pose proof (f_equal seg h3) as h4. pose (transfer_fun (sig_set_eq _ _ h4) f) as f'. pose (fun x => exist _ _ (Full_intro _ (f' x))) as f''. assert (h5:bijective f''). unfold f''. red. split. red. intros x y h6. apply exist_injective in h6. unfold f' in h6. rewrite <- (transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) x) in h6. rewrite <- (transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) y) in h6. do 2 rewrite transfer_fun_compat in h6. do 2 rewrite transfer_r_sig_set_eq in h6 at 1. apply h2a in h6. apply exist_injective in h6. apply proj1_sig_injective in h6. assumption. red. intro y. destruct y as [y h8]. pose proof (h2b y) as h9. destruct h9 as [x h9]. subst. exists (transfer (sig_set_eq _ _ h4) x). apply proj1_sig_injective. simpl. pose proof (transfer_sig_set_eq _ _ h4 (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) x) as h9. rewrite h9 at 1. unfold f'. match goal with |- transfer_fun _ _ ?x = _ => pose proof ((transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4)) x) as h10 end. rewrite <- h10 at 1. rewrite transfer_fun_compat. rewrite (transfer_r_sig_set_eq _ _ h4 (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4)) at 1. simpl. f_equal. apply proj1_sig_injective; auto. pose proof (bij_seg_finite_full_iff f'' h5 (fun x:sig_set A=>P (proj1_sig x))) as h6. split. intro h0. rewrite h6 in h0. eapply Finite_downward_closed. apply h0. red. intros x h7. destruct h7 as [h7]. constructor. destruct h7 as [h7 h8]. pose proof h7 as h7'. rewrite h4 in h7'. exists h7'. unfold f'', f'. simpl. match goal with |- P (proj1_sig (transfer_fun _ _ ?x)) => pose proof (transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) x) as h9 end. rewrite <- h9. rewrite transfer_fun_compat. rewrite (transfer_r_sig_set_eq _ _ h4). simpl. match goal with |- P (proj1_sig (f ?x)) => let D := type of h8 in match D with P (proj1_sig (f ?y)) => assert (h10: x = y) end end. apply proj1_sig_injective; auto. rewrite <- h10 in h8. assumption. intro h7. rewrite h6. eapply Finite_downward_closed. apply h7. red; intros x h8. destruct h8 as [h8]. constructor. destruct h8 as [h8 h9]. pose proof h8 as h7'. unfold full_sig in h4. rewrite <- h4 in h7' at 1. exists h7'. unfold f'' in h9. simpl in h9. unfold f' in h9. let C := type of h9 in match C with P (proj1_sig (transfer_fun _ _ ?c)) => pose proof ( transfer_undoes_transfer_r (sig_set_eq (seg (card_fun1 A)) (seg (card_fun1 (full_sig A))) h4) c) as h9' end. rewrite <- h9' in h9. rewrite transfer_fun_compat in h9. rewrite (transfer_r_sig_set_eq _ _ h4) in h9. simpl in h9. match goal with |- P (proj1_sig (f ?x)) => let D := type of h9 in match D with P (proj1_sig (f ?y)) => assert (h10: x = y) end end. apply proj1_sig_injective; auto. rewrite h10. assumption. Grab Existential Variables. f_equal; auto. f_equal; auto. Qed. Lemma finite_constructive_definite_description' : forall {T:Type} (A:Ensemble T) (pf:Finite A) (f:sig_set (seg (card_fun1 A)) -> sig_set A), bijective f -> forall (P:T->Prop), (exists x:sig_set A, P (proj1_sig x)) -> {x:T| P x}. intros T A h1 f h2 P h3. pose proof (ex_inh _ h3) as h4. pose proof (bij_seg_inh_iff' _ h1 f h2 ) as h5. rewrite h5 in h4. pose [n : nat | exists pf : Inn (seg (card_fun1 A)) n, P (proj1_sig (f (exist (Inn (seg (card_fun1 A))) n pf)))] as S. pose proof (bij_seg_finite_iff' _ h1 f h2) as h6. assert (ha:Finite [x : sig_set A | P (proj1_sig x)]). eapply Finite_downward_closed. rewrite <- finite_full_sig_iff. assumption. red; intros; constructor. specialize (h6 P). rewrite h6 in ha. pose (max_set_nat _ ha h4). pose proof (max_set_nat_compat _ ha h4) as h10. simpl in h10. destruct h10 as [h10 h11]. destruct h10 as [h10]. pose proof (destruct_exists_pf_set _ _ _ h10) as h12. simpl in h12. refine (exist _ _ h12). Defined. Lemma inh_im_seg_one_iff : forall {T:Type} {n:nat} (p:seg_one_t n->T), Inhabited (Im (Full_set _) p) <-> Inhabited [m:seg_one_t n | Inn (Im (Full_set _) p) (p m)]. intros T n p. split. intro h1. destruct h1 as [x h1]. destruct h1 as [m h1]. subst. apply Inhabited_intro with m. constructor. apply Im_intro with m. constructor. reflexivity. intro h1. destruct h1 as [m h1]. destruct h1 as [h1]. apply Inhabited_intro with (p m). apply Im_intro with m; auto; constructor. Qed. Lemma inh_im_seg_one_iff' : forall {T:Type} {n:nat} (p:seg_one_t n->T), Inhabited (Im (Full_set _) p) <-> Inhabited [m:nat | exists pf:1 <= m <= n, Inn (Im (Full_set _) p) (p (exist _ _ pf))]. intros T n p. rewrite inh_im_seg_one_iff. split. intro h1. destruct h1 as [m h1]. destruct m as [m h2]. destruct h1 as [h1]. apply Inhabited_intro with m. constructor. exists h2. assumption. intro h1. destruct h1 as [m h1]. destruct h1 as [h1]. destruct h1 as [h1 h2]. apply Inhabited_intro with (exist (fun s => 1 <= s <= n) _ h1). constructor. assumption. Qed. Lemma finite_index_assoc_fun_finite_aux : forall {T:Type} {n:nat} (p:seg_one_t n->T), Finite [m : nat | exists pf : 1 <= m <= n, Inn (Im (Full_set (seg_one_t n)) p) (p (exist (fun m0 : nat => 1 <= m0 <= n) m pf))]. intros T n p. eapply Finite_downward_closed. apply (finite_seg_one n). red; intros m h1. destruct h1 as [h1]. destruct h1 as [h1]. constructor; auto. Qed. Lemma finite_index_assoc_fun_in_aux : forall {T:Type} {n:nat} (p:seg_one_t n->T) (a:nat), Inn [m : nat | exists pf : 1 <= m <= n, Inn (Im (Full_set (seg_one_t n)) p) (p (exist (fun m0 : nat => 1 <= m0 <= n) m pf))] a -> 1 <= a <= n. intros T n p a h1. destruct h1 as [h1]. destruct h1; auto. Qed. Definition finite_index_assoc_fun {T U:Type} {n:nat} (p:seg_one_t n->T) (q:seg_one_t n->U) : let S := seg_one_t_injector p in sig_set (Im S p)->U. intros S y. pose proof (seg_one_t_injector_compat p) as h1. simpl in h1. destruct h1 as [h1 h2]. pose (left_inverse _ h2 (full_sig (seg_one_t_injector p))) as pi. refine (q _). pose proof (im_full_sig_restriction (seg_one_t_injector p) p) as h3. pose (transfer_fun (sig_set_eq _ _ h3) pi) as pi'. refine (proj1_sig (pi' y)). Defined. Lemma finite_index_assoc_fun_compat : forall {T U:Type} {n:nat} (p:seg_one_t n->T) (q:seg_one_t n->U) (m:seg_one_t n) (pf:Inn (seg_one_t_injector p) m), finite_index_assoc_fun p q (exist _ _ (Im_intro _ _ _ p m pf _ (eq_refl _))) = q m. intros T U n p q m h1. unfold finite_index_assoc_fun. destruct (seg_one_t_injector_compat p) as [h2 h3]. pose proof (left_inverse_compat (restriction_fun p (seg_one_t_injector p)) h3 (full_sig (seg_one_t_injector p)) (exist _ _ h1) (Full_intro _ _)) as h4. pose proof (f_equal (@proj1_sig _ _) h4) as h5. simpl in h5. rewrite <- h5 at 5. f_equal. f_equal. f_equal. match goal with |- transfer_fun ?x _ ?z = _ => rewrite <- (transfer_undoes_transfer_r x z) end. rewrite transfer_fun_compat. f_equal. match goal with |- transfer_r ?z _ = _ => rewrite (transfer_r_sig_set_eq _ _ (im_full_sig_restriction (seg_one_t_injector p) p) z) end. apply proj1_sig_injective; auto. Qed. (*This is if both the given seg_one_T functions have different integer bound variables, yet are equal.*) Definition finite_index_assoc_fun' {T U:Type} {n n':nat} (p:seg_one_t n->T) (q:seg_one_t n'->U) (pf:n = n') : let S := seg_one_t_injector p in sig_set (Im S p)->U. subst. refine (finite_index_assoc_fun p q). Defined. Lemma finite_index_assoc_fun_compat' : forall {T U:Type} {n n':nat} (p:seg_one_t n->T) (q:seg_one_t n'->U) (m:seg_one_t n) (pfeq:n = n') (pf:Inn (seg_one_t_injector p) m), finite_index_assoc_fun' p q pfeq (exist _ _ (Im_intro _ _ _ p m pf _ (eq_refl _))) = q (s_transfer (seg_one_t_functional _ _ pfeq) m). intros T U n n' p q m h1 h2. subst. pose proof (proof_irrelevance _ (seg_one_t_functional n' n' eq_refl) (eq_refl _)) as h3. rewrite h3. rewrite s_transfer_eq_refl. unfold finite_index_assoc_fun', eq_rect_r. simpl. apply finite_index_assoc_fun_compat. Qed. (*Countable*) Lemma bij_nat_inh_iff : forall {T:Type} {A:Ensemble T} (f:nat->sig_set A) (P:sig_set A->Prop), bijective f -> (Inhabited [x:sig_set A | P x] <-> Inhabited [n : nat | P (f n)]). intros T A f P h1. split. intro h2. destruct h2 as [x h2]. destruct h2 as [h2]. pose proof h1 as h1'. apply bijective_impl_invertible in h1'. pose (proj1_sig (function_inverse _ h1')) as f'. apply Inhabited_intro with (f' x). constructor. unfold f', function_inverse. destruct constructive_definite_description as [g h3]. simpl. destruct h3 as [h3 h4]. rewrite h4. assumption. intro h2. destruct h2 as [n h2]. destruct h2 as [h2]. apply Inhabited_intro with (f n). constructor. assumption. Qed. Lemma bij_nat_set_inh_iff : forall {T:Type} {A:Ensemble T} {S:Ensemble nat} (f:sig_set S->sig_set A) (P:sig_set A->Prop), bijective f -> ~Finite S -> (Inhabited [x:sig_set A | P x] <-> Inhabited [n : sig_set S | P (f n)]). intros T A S f P h1 h2. split. intro h3. destruct h3 as [x ha]. destruct ha as [ha]. pose proof h1 as h1'. apply bijective_impl_invertible in h1'. pose (proj1_sig (function_inverse _ h1')) as f'. apply Inhabited_intro with (f' x). constructor. unfold f', function_inverse. destruct constructive_definite_description as [g h3]. simpl. destruct h3 as [h3 h4]. rewrite h4. assumption. intro hb. destruct hb as [n hb]. destruct hb as [hb]. apply Inhabited_intro with (f n). constructor. assumption. Qed. Lemma countable_constructive_definite_description : forall {T:Type} {A:Ensemble T} (f:nat->sig_set A) (P:sig_set A->Prop), bijective f -> (exists x:sig_set A, P x) -> {x:sig_set A| P x}. intros T A f P h1 h2. pose proof wo_nat as h3. pose proof (ex_inh _ h2) as h4. rewrite (bij_nat_inh_iff f P h1) in h4. pose (min_set _ _ h3 _ h4) as m. pose proof (min_set_compat _ _ h3 _ h4) as h5. simpl in h5. destruct h5 as [h5 h6]. destruct h5 as [h5]. refine (exist _ _ h5). Defined. Lemma countable_set_constructive_definite_description : forall {T:Type} {S:Ensemble nat} {A:Ensemble T} (f:sig_set S->sig_set A) (P:sig_set A->Prop), ~Finite S -> bijective f -> (exists x:sig_set A, P x) -> {x:sig_set A| P x}. intros T S A f P h1 h2 h3. pose proof wo_nat as h8. pose proof (well_order_exist_rel _ h8 (fun a => Inn S a)) as h8'. pose proof (ex_inh _ h3) as h9. rewrite (bij_nat_set_inh_iff f P h2 h1) in h9. pose (min_set _ _ h8' _ h9) as n. pose proof (min_set_compat _ _ h8' _ h9) as hn. simpl in hn. destruct hn as [h10 h11]. inversion h10 as [h12]. refine (exist _ _ h12). Defined. End CountableDefiniteDescription. bool2-v0-3/ListUtilities.v0000644000175000017500000501207013575574055016345 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-2019, Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq Version 8.4 pl4.*) Require Import LogicUtilities. Require Import TypeUtilities. Require Import EnsemblesImplicit. Require Import Finite_sets_facts. Require Export List. Require Import FunctionProperties. Require Import Constructive_sets. Require Import DecidableDec. Require Import FunctionalExtensionality. Require Import SetUtilities. Require Import ArithUtilities. Require Import NaryFunctions. Require Import ProofIrrelevance. Require Import Description. Require Import EqDec. Require Import Program.Wf. Require Import Equality. Require Import Init. Require Import Bool. Require Import NPeano. Require Import Basics. Lemma length_subst : forall {T:Type} (l l':list T), l = l' -> length l = length l'. intros; subst; auto. Qed. Lemma list_eq_dec : forall {T:Type}, type_eq_dec T -> type_eq_dec (list T). intros T h0. red in h0. red. intro l. revert h0. induction l as [|a l h1]; intros h0 l'; destruct l' as [|a' l']. left; auto. right. intro h1. discriminate h1. right. intro h2. discriminate h2. destruct (h0 a a') as [h2 | h2]. subst. specialize (h1 h0 l'). destruct h1 as [h1 | h1]. subst. left; auto. right. intro h2. inversion h2; contradiction. right. intro h3. inversion h3; auto. Qed. Lemma list_nat_eq_dec : type_eq_dec (list nat). apply list_eq_dec; apply nat_eq_dec. Qed. Lemma list_nat_eq_dec_eq : forall (pf:type_eq_dec (list nat)), pf = list_nat_eq_dec. intro h1. apply functional_extensionality_dep. intro l. apply functional_extensionality_dep. intro l'. destruct (h1 l l'), (list_nat_eq_dec l l'); subst; try f_equal; try apply proof_irerlevance; auto. apply proof_irrelevance. contradict n; auto. contradict n; auto. apply proof_irrelevance. Qed. Lemma in_dec_eq : forall {T:Type} {l:list T} {a:T} (pfdt:type_eq_dec T) (pf:{In a l} + {~In a l}), in_dec pfdt a l = pf. intros T l a h1 h2. destruct h2 as [h2 | h2], in_dec as [h3 | h3]; f_equal; try apply proof_irrelevance; auto. contradict (h3 h2). contradict (h2 h3). Qed. Definition faml (T:Type) : Type := list (list T). Lemma in_subst1 : forall {T x x'} {l:list T}, x = x' -> In x l -> In x' l. intros; subst; auto. Qed. Lemma in_subst_l : forall {T:Type} (l l':list T) (x:T), l = l' -> In x l -> In x l'. intros; subst; auto. Qed. Lemma in_subst_r : forall {T:Type} (l l':list T) (x:T), l = l' -> In x l' -> In x l. intros; subst; auto. Qed. Lemma nin_subst1 : forall {T x x'} {l:list T}, x = x' -> ~In x l -> ~In x' l. intros; subst; auto. Qed. Lemma nin_subst_l : forall {T:Type} (l l':list T) (x:T), l = l' -> ~In x l -> ~In x l'. intros; subst; auto. Qed. Lemma nin_subst_r : forall {T:Type} (l l':list T) (x:T), l = l' -> ~In x l' -> ~In x l. intros; subst; auto. Qed. Lemma in_subst1_pred : forall {T:Type} (l:list T) (x x':T), x = x' -> (fun l => In x l) = (fun l => In x' l). intros; subst; auto. Qed. Lemma in_subst_pred_l : forall {T:Type} (l l':list T) (x:T), l = l' -> (fun x => In x l) = (fun x => In x l'). intros; subst; auto. Qed. Lemma in_subst_pred_r : forall {T:Type} (l l':list T) (x:T), l = l' -> (fun x => In x l') = (fun x => In x l). intros; subst; auto. Qed. Lemma nin_subst1_pred : forall {T} {x x':T}, x = x' -> (fun l => ~In x l) = (fun l => ~In x' l). intros; subst; auto. Qed. Lemma nin_subst_pred_l : forall {T:Type} {l l':list T}, l = l' -> (fun x => ~In x l) = (fun x => ~In x l'). intros; subst; auto. Qed. Lemma nin_subst_pred_r : forall {T:Type} (l l':list T) (x:T), l = l' -> (fun x => ~In x l') = (fun x => ~In x l). intros; subst; auto. Qed. Lemma in_from_cons : forall {T} (l:list T) (a x:T), In a l -> In x (a::l) -> In x l. intros ? ? ? ? ? h1; destruct h1; subst; auto. Qed. Lemma in_not_nil : forall {T:Type} (l:list T) (x:T), In x l -> l <> nil. intros; intro; subst; contradiction. Qed. Lemma not_nil_cons : forall {T:Type} (l:list T), l <> nil -> exists (a:T) (l':list T), l = a::l'. intros T l h1. destruct l as [|a l']. contradict h1. auto. exists a. exists l'. auto. Qed. Lemma cons_neq_nil : forall {T:Type} (l:list T) (a:T), a :: l <> nil. intros ? ? ? h1. discriminate h1. Qed. Lemma cons_inj : forall {T:Type} (a1 a2:T) (l1 l2:list T), a1::l1 = a2::l2 -> a1 = a2 /\ l1 = l2. intros ? ? ? ? ? h1. inversion h1. split; trivial. Qed. Lemma cons_eq : forall {T:Type} (a1 a2:T) (l1 l2:list T), a1::l1 = a2::l2 <-> a1 = a2 /\ l1 = l2. intros T a1 a2 l1 l2. split. intro h1. apply cons_inj; assumption. intro h1. destruct h1; subst. reflexivity. Qed. Lemma ex_in_cons : forall {T:Type} (l:list T) (a:T), (exists x:T, In x l) -> (exists x:T, In x (a::l)). intros T l a h1. destruct h1 as [x h1]. exists x. right; auto. Qed. Lemma nil_dec : forall {T:Type} (l:list T), {l = nil} + {exists a:T, In a l}. intros T l. induction l as [|a l ih]. left; auto. destruct ih as [h1 | h1]. subst. right. exists a. left. reflexivity. right. apply ex_in_cons; auto. Qed. Lemma ex_in_not_nil : forall {T:Type} (l:list T), (exists x:T, In x l) -> l <> nil. intros T l h1. destruct h1 as [x h1]. intro. subst. contradiction. Qed. Lemma nil_dec' : forall {T:Type} (l:list T), {l = nil} + {l <> nil}. intros T l. destruct (nil_dec l) as [h1 | h1]. left; auto. apply ex_in_not_nil in h1. right; auto. Qed. Lemma in_cons2 : forall {T} {l:list T} a b x, In x (a::l) -> In x (a::b::l). intros ? ? ? ? ? h1; destruct h1 as [|h1]; subst; auto; [left|right;right]; auto. Qed. Lemma in_cons_dec : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a x:T), In x (a::l) -> {x = a} + {In x l}. intros T h1 l a x h2. destruct (h1 x a) as [h3 | h3]. left; auto. right. destruct h2 as [|h2]; subst. contradict h3; auto. assumption. Qed. (*When you can't use [simpl]!*) Lemma length_nil : forall (T:Type), length (@nil T) = 0. auto. Qed. Lemma l_ex_dec : forall {T:Type} (pfdt:type_eq_dec T) (P:T->Prop) (l:list T), pred_dec P -> {exists x:T, In x l /\ P x} + {~exists x:T, In x l /\ P x}. intros T hdt P l. induction l as [|a l ih]; simpl. intro hp; right. intro h1. destruct h1 as [? h1]; destruct h1; auto. intro h1. specialize (ih h1). destruct ih as [ih | ih]. left. destruct ih as [x ih]. destruct ih. exists x. split; try right; auto. destruct (h1 a) as [h2 | h2]. left. exists a. split; auto. right. intro h3. destruct h3 as [x h4]. destruct h4 as [h4 h5]. destruct h4 as [|h4]; subst. contradiction. contradict ih. exists x; auto. Qed. Lemma length_cons : forall {T:Type} (l:list T) (a:T), length (a::l) = S (length l). auto. Qed. Lemma l_lt_length_cons : forall {T:Type} (l:list T) (a:T) (n:nat), S n < length (a::l) -> n < length l. intros T l a n h1. simpl in h1. omega. Qed. Lemma lt_cons_app_sing : forall {T:Type} (l:list T) (a b:T), S (length l) < length (a :: l ++ b :: nil). intros; simpl; rewrite app_length; simpl; rewrite S_compat; auto with arith. Qed. Lemma cons_app : forall {T} (l l':list T) (a:T), a::(l++l') = (a::l)++l'. auto. Qed. Lemma length_cons_eq : forall {T:Type} {l l':list T} {a:T}, length (a::l) = length l' -> length l = length (tl l'). intros T l l' a h2. simpl in h2. destruct l' as [|a' l']; simpl. simpl in h2. discriminate. simpl in h2. apply S_inj in h2; auto. Qed. Lemma length_cons_eq_iff : forall {T:Type} (l l':list T) (a:T), l' <> nil -> (length (a::l) = length l' <-> length l = length (tl l')). intros T l l' a h1. split. apply length_cons_eq. intro h2. destruct l' as [|a' l']; simpl. contradict h1; auto. simpl in h2; f_equal; auto. Qed. Lemma lt_length : forall {T:Type} (l:list T) (n:nat), n < length l -> l <> nil. intros T l n h1 h2. subst. simpl in h1. apply lt_n_0 in h1. contradiction. Qed. Lemma length_eq_S : forall {T:Type} (l:list T) (n:nat), length l = S n -> l <> nil. intros; intro; subst; discriminate. Qed. Lemma length_eq0 : forall {T:Type} (l:list T), length l = 0 -> l = nil. intros T l. destruct l as [|a l]; auto. intro h1. simpl in h1. discriminate h1. Qed. Lemma not_nil_lt : forall {T:Type} (l:list T), l <> nil -> 0 < length l. intros T l. destruct (zerop (length l)) as [h2 |]; auto. apply length_eq0 in h2. contradiction. Qed. Lemma not_nil_pred_lt : forall {T:Type} (l:list T), l <> nil -> pred (length l) < length l. intros T l h1. destruct (zerop (length l)) as [h2 |]; auto. apply length_eq0 in h2; subst. contradict h1; auto. apply (lt_pred_n _ _ l0). Qed. Lemma length_functional : forall {T:Type} (l l':list T), l = l' -> length l = length l'. intros; subst; auto. Qed. Lemma le_length_subst : forall {T:Type} (l l':list T) (m:nat), l = l' -> m <= length l -> m <= length l'. intros; subst; auto. Qed. Lemma lt_length_subst : forall {T:Type} (l l':list T) (m:nat), l = l' -> m < length l -> m < length l'. intros; subst; auto. Qed. Lemma remove_not_in' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), ~ In a l -> l = remove pfdt a l. intros T h0 l a h1. induction l. simpl. auto. simpl. destruct (h0 a a0). subst. pose proof (in_eq a0 l). contradiction. f_equal. simpl in h1. assert (h2: ~ In a l). tauto. apply IHl; auto. Qed. Lemma remove_hd_no_dup : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (def:T), NoDup l -> remove pfdt (hd def l) l = tl l. intros T h0 l. induction l as [|a l h1]. simpl. auto. simpl. intros def h2. destruct (h0 a a) as [h5 | h6]. inversion h2 as [|x l' h3 h4]. subst. symmetry. apply remove_not_in'; auto. contradiction h6. reflexivity. Qed. Lemma hd_indep : forall {T:Type} (l:list T) (d d':T), l <> nil -> hd d l = hd d' l. intros ? l; destruct l; auto; intros ? ? h1; try contradict h1; auto. Qed. Lemma hd_remove_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (r:T), r <> hd r l -> hd r (remove pfdt r l) = hd r l. intros T h0 l. destruct l as [|a l]; simpl; auto. intros r h1. destruct (h0 r a) as [h5 | h6]. subst. contradict h1; auto. simpl. reflexivity. Qed. Lemma in_hd_not_nil : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (def:T), l <> nil -> In (hd def l) l. intros T pfd l. induction l as [|a l' h1]. intros def h2. contradict h2. reflexivity. simpl. intros def h2. left. reflexivity. Qed. Lemma hd_map : forall {T U:Type} (f:T->U) (def:T) (l:list T), hd (f def) (map f l) = f (hd def l). intros T U f def l. induction l as [|a l h1]; simpl; auto. Qed. Definition hd' {T:Type} (l:list T) : l <> nil -> T := match l with | nil => (fun pf => False_rect T (pf eq_refl)) | a::l' => (fun _ => a) end. Lemma hd_compat' : forall {T:Type} (l:list T) (pf: l <> nil), let a := hd' l pf in l = a::tl l. intros T l. destruct l as [|a l]; auto. intro h1. contradict h1; auto. Qed. Lemma hd_compat : forall {T:Type} (l:list T) (pf: l <> nil), hd' l pf = hd (hd' l pf) l. intros T l. destruct l as [|a l]. intro h1. contradict h1; auto. intros h1. simpl; auto. Qed. Lemma hd_functional : forall {T:Type} (l l':list T) (pf:l = l') (pf:l <> nil) (pf':l' <> nil), hd' l pf = hd' l' pf'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma hd_functional' : forall {T:Type} (l:list T) (pf1 pf2:l <> nil), hd' l pf1 = hd' l pf2. intros; f_equal; apply proof_irrelevance. Qed. Lemma hd_hd_compat' : forall {T:Type} (l:list T) (pf:l <> nil) (def:T), hd def l = hd' l pf. intros T l. induction l as [|a l h0]; simpl. intro h1. contradict h1; auto. intros; auto. Qed. Lemma in_hd' : forall {T:Type} (l:list T) (pf:l <> nil), In (hd' l pf) l. intros T l h1. pose proof (hd_compat' l h1) as h2. simpl in h2. rewrite h2 at 2. left. reflexivity. Qed. Lemma hd_cons' : forall {T:Type} (l:list T) (a:T) (pf:a::l <> nil), hd' (a::l) pf = a. intros T l a h1. pose proof (hd_compat' (a::l) h1) as h2. simpl in h2. apply cons_inj in h2. destruct h2; auto. Qed. Lemma hd_app' : forall {T:Type} (l l':list T) (pf:l <> nil) (pf':l ++ l' <> nil), hd' (l ++ l') pf' = hd' l pf. intros T l. destruct l as [|a l]; simpl. intros ? h1; contradict h1; auto. auto. Qed. Lemma removelast_not_nil : forall {T:Type} (l:list T) (pf:removelast l <> nil), l <> nil. intros ? ? h1; intro; subst; simpl in h1; contradict h1; auto. Qed. Lemma hd_removelast' : forall {T:Type} (l:list T) (pf:removelast l <> nil), let pf' := removelast_not_nil _ pf in hd' (removelast l) pf = hd' l pf'. intros T l h1; simpl. destruct l as [|c l]; simpl in h1. contradict h1; auto. simpl. destruct l as [|d l]. contradict h1; auto. simpl; auto. Qed. Lemma impl_eq_cons : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (pf:l <> nil), tl l = l' -> l = hd' l pf ::l'. intros T h1 l. destruct l as [|a l]; simpl. intros ? h2. contradict h2; auto. intros l' h2 h3. subst; auto. Qed. Lemma impl_eq_cons' : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T) (pf:l <> nil), hd' l pf = a -> tl l = l' -> l = a::l'. intros; subst; apply impl_eq_cons; auto. Qed. Lemma cons_eq_iff : forall {T:Type} (l l':list T) (a:T), l = a::l' <-> exists pf:l <> nil, a = hd' l pf /\ l' = tl l. intros T l l' a; split; intro h1; subst. exists (cons_neq_nil _ _). simpl; auto. destruct h1 as [h1 h2]. destruct h2; subst. apply hd_compat'. Qed. Lemma in_or' : forall {T:Type} (l:list T) (x:T) (pf:In x l), x = hd' _ (in_not_nil _ _ pf) \/ In x (tl l). intros T l x h1. destruct l as [|a l]. contradict h1. simpl in h1; simpl. destruct h1; subst; [left; auto | right; auto]. Qed. Lemma remove_hd_no_dup' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pf:l <> nil), NoDup l -> remove pfdt (hd' l pf) l = tl l. intros T h1 l h2 h3. pose proof (not_nil_cons _ h2) as h4. destruct h4 as [a h4]. destruct h4 as [l' h4]. rewrite <- (hd_hd_compat' _ _ a). apply remove_hd_no_dup; auto. Qed. Lemma no_dup_cons : forall {T:Type} (l:list T) (a:T), NoDup (a::l) -> NoDup l. intros T l a h1. rewrite <- app_nil_l in h1. pose proof (NoDup_remove_1 _ _ _ h1) as h2. rewrite app_nil_l in h2. assumption. Qed. Lemma no_dup_cons2 : forall {T:Type} (l:list T) (a b:T), NoDup (a::b::l) -> NoDup l. intros T l a b h1. rewrite <- app_nil_l in h1. pose proof (NoDup_remove_1 _ _ _ h1) as h2. pose proof (NoDup_remove_1 _ _ _ h2) as h3. rewrite app_nil_l in h3. assumption. Qed. Lemma no_dup_cons_nin : forall {T:Type} (l:list T) (a:T), NoDup (a::l) -> ~In a l. intros T l a h1. rewrite <- app_nil_l in h1. pose proof (NoDup_remove_2 _ _ _ h1) as h2. rewrite app_nil_l in h2. assumption. Qed. Lemma no_dup_cons2_nin : forall {T:Type} (l:list T) (a b:T), NoDup (a::b::l) -> ~In a l. intros T l a b h1. rewrite <- app_nil_l in h1. pose proof (NoDup_remove_2 _ _ _ h1) as h2. pose proof (NoDup_remove_1 _ _ _ h1) as h3. pose proof (NoDup_remove_2 _ _ _ h3) as h4. intro h5. assert (h6: a <> b). intro; subst. rewrite app_nil_l in h2. contradict h2. left; auto. contradict h2. right; auto. Qed. Lemma no_dup_cons2' : forall {T:Type} (l:list T) (a b:T), NoDup (a::b::l) -> NoDup (a::l). intros T l a b h1. constructor. apply no_dup_cons_nin in h1; auto. contradict h1; right; auto. apply no_dup_cons2 in h1; auto. Qed. Lemma no_dup_cons2_neq : forall {T:Type} (l:list T) (a b:T), NoDup (a::b::l) -> a <> b. intros T l a b h1 h2; subst. apply no_dup_cons_nin in h1. contradict h1; left; auto. Qed. Lemma nin_cons1 : forall {T:Type} {l:list T} {a x:T}, ~In x (a::l) -> x <> a. intros; intro; subst; destruct H; auto; left; auto. Qed. Lemma nin_cons2 : forall {T:Type} {l:list T} {a x:T}, ~In x (a::l) -> ~In x l. intros; contradict H; right; auto. Qed. Lemma hd_firstn' : forall {T:Type} (l:list T) (m:nat) (pf:l <> nil) (pf':firstn m l <> nil), hd' l pf = hd' (firstn m l) pf'. intros T l. destruct l as [|a l]; simpl. intros ? h2; contradict h2; auto. intros m h2 h3. destruct m as [|m]. simpl in h3. contradict h3; auto. simpl; auto. Qed. Lemma length_eq_1 : forall {T:Type} (l:list T) (pf:length l = 1), l = hd' l (lt_length _ _ (lt_eq_trans _ _ _ (lt_0_Sn 0) (eq_sym pf)))::nil. intros T l. destruct l as [|a l]; simpl; subst. intro; discriminate. intros h1. apply S_inj in h1; subst. apply length_eq0 in h1; subst; auto. Qed. Lemma sing_eq_iff : forall {T:Type} (l:list T) (a:T), l = a::nil <-> exists pf:l <> nil, hd' l pf = a /\ tl l = nil. intros T l a; split; intro h1; subst. exists (cons_neq_nil _ _); simpl; split; auto. destruct h1 as [h1 [h2 h3]]. rewrite (hd_compat' _ h1); f_equal; auto. Qed. Lemma tl_eq : forall {T:Type} (l:list T), tl l = skipn 1 l. intros T l. destruct l; auto. Qed. Lemma tl_eq_nil_iff : forall {T:Type} (l:list T), tl l = nil <-> l = nil \/ exists x:T, l = x :: nil. intros T l. destruct l as [|a l]; simpl. tauto. split. intro; subst. right. exists a; auto. intro h1. destruct h1 as [h1 | h1]. discriminate. destruct h1 as [x h1]. apply cons_inj in h1. destruct h1; auto. Qed. Lemma app_sing : forall {T:Type} (l:list T) (a:T), a :: nil ++ l = a::l. intros; auto. Qed. Lemma app_sing_app : forall {T:Type} (l l':list T) (a:T), (l ++ a::nil) ++ l' = l ++ a::l'. intros; rewrite <- app_assoc; simpl; auto. Qed. Lemma in_app_l : forall {T:Type} {l:list T} {x:T}, In x l -> forall (l':list T), In x (l++l'). intros; rewrite in_app_iff; left; auto. Qed. Lemma in_app_r : forall {T:Type} {l':list T} {x:T}, In x l' -> forall (l:list T), In x (l++l'). intros; rewrite in_app_iff; right; auto. Qed. Lemma app_not_nil_l : forall {T:Type} (l l':list T), l <> nil -> l ++ l' <> nil. intros T l l' h1. destruct l as [|a l]. contradict h1; auto. rewrite <- app_comm_cons. intro h2. discriminate. Qed. Lemma app_not_nil_r : forall {T:Type} (l l':list T), l' <> nil -> l ++ l' <> nil. intros T l l' h1. destruct l as [|a l]; simpl; auto. intro; discriminate. Qed. Lemma lt_app : forall {T:Type} (l l':list T) (i:nat), i < length l -> i < length (l++l'). intros T l l' i h1. rewrite app_length; auto with arith. Qed. Lemma sing_not_nil : forall {T:Type} (a:T), a :: nil <> nil. intros; discriminate. Qed. Lemma app_inj_eq_length : forall {T:Type} (l1 l2 l1' l2':list T), l1 ++ l2 = l1' ++ l2' -> length l1 = length l1' -> l1 = l1' /\ l2 = l2'. intros T l1. induction l1 as [|a1 l1 ih]; simpl. intros ? ? ? ? h1; subst. pose proof length_eq0 _ (eq_sym h1); subst. split; auto. intros l2 l1'. revert l2. destruct l1' as [|a1' l1']; simpl. intros l2 l2' h3. subst. intros; discriminate. intros l2 l2' h1. inversion h1 as [h3]; subst. clear h1. intro h1. apply S_inj in h1. specialize (ih _ _ _ H h1). destruct ih; subst. auto. Qed. Lemma app_nin_length : forall {T} (l0 l0' l1 l1':list T) b, ~In b l0 -> ~In b l0' -> l0 ++ b::l1 = l0' ++ b::l1' -> length l0 = length l0'. intros T l0. induction l0 as [|a0 l0 ih]; simpl. intro l0'. destruct l0' as [|a0' l0']; simpl; auto. intros l1 l1' b h1 h2 h3. inversion h3; subst. contradict h3. contradict h2; left; auto. intro l0'. destruct l0' as [|a0' l0']. intros l1 l1' b h1 h2 h3. simpl in h3. inversion h3; subst. contradict h1; left; auto. intros l1 l1' b h1 h3 h4; simpl in h4; f_equal. specialize (ih l0' l1 l1' b). hfirst ih. contradict h1; right; auto. specialize (ih hf). hfirst ih. contradict h3; right; auto. specialize (ih hf0). inversion h4; subst. specialize (ih H1); simpl; f_equal; auto. Qed. Lemma app_nin_eq : forall {T} (l0 l0' l1 l1':list T) b, ~In b l0 -> ~In b l0' -> l0 ++ b::l1 = l0' ++ b::l1' -> l0 = l0' /\ l1 = l1'. intros T l0 l0' l1 l1' b h1 h2 h3. apply app_inj_eq_length in h3. destruct h3 as [? h3]; subst. inversion h3; tauto. apply app_nin_length in h3; auto. Qed. Lemma not_app_same_eq_r : forall {T:Type} (l l':list T), l' <> nil -> ~ l = l' ++ l. intros T l l' h1 h2. pose proof (f_equal (@length T) h2) as h3. rewrite app_length in h3. apply not_nil_lt in h1. omega. Qed. Lemma not_app_same_eq_l : forall {T:Type} (l l':list T), l' <> nil -> ~ l = l ++ l'. intros T l l' h1 h2. pose proof (f_equal (@length T) h2) as h3. rewrite app_length in h3. apply not_nil_lt in h1. omega. Qed. Lemma app_sing_eq_app_cons_cons : forall {T} (l l' l0:list T) (a b:T), ~In a (b::l0) -> l ++ a::nil = l' ++ a :: b :: l0 -> In a l. intros T l; induction l as [|a l ih]; simpl; intros l' l0 b c h0 h1. destruct l' as [|d l']. simpl in h1. inversion h1. simpl in h1. inversion h1; subst. pose proof (app_not_nil_r l' (d::c::l0) (cons_neq_nil _ _)). rewrite <- H1 in H. contradict H; auto. destruct l' as [|a' l']. simpl in h1. inversion h1; subst. left; auto. simpl in h1. inversion h1; subst. apply ih in H1. right; auto. assumption. Qed. Lemma removelast_eq : forall {T:Type} (l:list T), removelast l = firstn (pred (length l)) l. intros T l. induction l as [|a l ih]; simpl; auto. destruct l as [|b l]; simpl; auto. f_equal. simpl in ih; auto. Qed. Lemma removelast_in : forall {T:Type} (l:list T) (x:T), In x (removelast l) -> In x l. intros T l x h1. erewrite (app_removelast_last x). apply in_app_l; auto. intro; subst. simpl in h1; auto. Qed. Lemma incl_removelast : forall {T:Type} (l:list T), incl (removelast l) l. intros T l. red. intros a h1. apply removelast_in; auto. Qed. Lemma length_removelast : forall {T:Type} (l:list T), length (removelast l) = pred (length l). intros ? l; induction l as [|? l ih]; simpl; try simpl in ih; auto; destruct l;simpl; auto. Qed. Lemma removelast_cons_cons : forall {T:Type} (l:list T) (a a':T), removelast (a::a'::l) = a::removelast (a'::l). intros; simpl; auto. Qed. Lemma removelast_cons_cons_neq_nil : forall {T:Type} (l:list T) (a a':T), removelast (a::a'::l) <> nil. intros; eapply cons_neq_nil; apply removelast_cons_cons. Qed. Lemma firstn_removelast : forall {T:Type} (l:list T) (n:nat), n <= length (removelast l) -> firstn n (removelast l) = firstn n l. intros T l. induction l as [|a l ih]; simpl; auto. intros n h2. destruct l as [|b l]. simpl in h2. apply le0 in h2; subst; simpl; auto. simpl in h2. destruct n as [|n]. simpl; auto. apply le_S_n in h2. specialize (ih _ h2). simpl. f_equal. simpl in ih; auto. Qed. Lemma last_cons : forall {T:Type} (l:list T) (a:T), last (a::l) a = last l a. intros; simpl; destruct l; auto. Qed. Lemma last_cons' : forall {T:Type} (l:list T) (a d:T), l <> nil -> last (a::l) d = last l d. intros; simpl; destruct l; auto. contradict H; auto. Qed. Lemma in_last : forall {T:Type} (l:list T) (d:T), l <> nil -> In (last l d) l. intros ? l; induction l as [|c l ih]; simpl; auto. intros d h1. destruct l as [|l]. left; auto. right; apply ih; auto. apply cons_neq_nil. Qed. Lemma last_indep : forall {T:Type} (l:list T) (pf:l <> nil) (a b:T), last l a = last l b. intros T l; induction l as [|c l ih]; intro h1. contradict h1; auto. intros; simpl; auto. destruct l as [|d l]; auto. specialize (ih (cons_neq_nil _ _)). apply ih. Qed. Lemma last_app_sing : forall {T:Type} (l:list T) (a d:T), last (l++a::nil) d = a. intros T l. induction l as [|c l ih]; simpl; auto. intros a d. pose proof (in_app_r (in_eq a nil) l) as h2. rewrite (hd_compat' _ (in_not_nil _ _ h2)). simpl. specialize (ih a d). destruct l; auto. Qed. Lemma removelast_app_sing : forall {T:Type} (l:list T) (a:T), removelast (l++a::nil) = l. intros T l. induction l as [|a l ih]; simpl; auto. intro b. pose proof (in_app_r (in_eq b nil) l) as h2. rewrite (hd_compat' _ (in_not_nil _ _ h2)). f_equal. rewrite <- hd_compat'. apply ih. Qed. Lemma count_occ_nil : forall (T:Type) (pfdt:type_eq_dec T) (x:T), count_occ pfdt nil (A := T) x = 0. auto. Qed. Lemma count_occ_sing_eq : forall (T:Type) (pfdt:type_eq_dec T) (x:T), count_occ pfdt (x::nil) x = 1. simpl; intros; rewrite eq_dec_same at 1; auto. Qed. Lemma count_occ_sing_neq : forall (T:Type) (pfdt:type_eq_dec T) (x a:T), x <> a -> count_occ pfdt (a::nil) x = 0. simpl; intros; lr_if_goal'; subst; auto; contradict H; auto. Qed. Lemma count_occ_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (m:nat), m < count_occ pfdt l x -> In x l. intros ? h1 ? ? ? ? ; rewrite (count_occ_In h1); omega. Qed. Lemma count_occ_eq_S_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), count_occ pfdt l x = S n -> In x l. intros. rewrite count_occ_In. rewrite H at 1. auto with arith. Qed. Lemma count_occ_nin : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), ~ In x l -> count_occ pfdt l x = 0. intros T h1 l x h2. rewrite (count_occ_In h1) in h2; omega. Qed. Lemma count_occ_le : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), count_occ pfdt l x <= length l. intros T h1 l. induction l as [|a l ih]; simpl. intros; auto with arith. intro x. destruct (h1 a x) as [| h3]; subst. apply le_n_S. apply ih. specialize (ih x). auto with arith. Qed. Lemma count_occ_app : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T), count_occ pfdt (l++l') x = count_occ pfdt l x + count_occ pfdt l' x. intros T h1 l. induction l as [|a l ih]; simpl. intros; auto. intros l' x. destruct (h1 a x) as [|h2]; subst. rewrite plus_comm. rewrite plus_S_shift_r. f_equal. rewrite plus_comm. apply ih. apply ih. Qed. Lemma weak_bound_count_occ_app : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T) (occ:nat), count_occ pfdt l x <= occ < count_occ pfdt (l++l') x -> occ - (count_occ pfdt l x) < count_occ pfdt l' x. intros T h1 l l' x occ h2. rewrite count_occ_app in h2. omega. Qed. Lemma count_occ_firstn : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), count_occ pfdt (firstn n l) x = count_occ pfdt l x - count_occ pfdt (skipn n l) x. intros T h1 l x n. rewrite <- (firstn_skipn n l) at 2. rewrite count_occ_app; auto with arith. rewrite <- plus_minus_assoc; auto. rewrite minus_same; auto. Qed. Lemma count_occ_skipn : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), count_occ pfdt (skipn n l) x = count_occ pfdt l x - count_occ pfdt (firstn n l) x. intros T h1 l x n. rewrite <- (firstn_skipn n l) at 2. rewrite count_occ_app; auto with arith. Qed. Lemma count_occ_last : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (d:T), l <> nil -> count_occ pfdt l (last l d) = S (count_occ pfdt (removelast l) (last l d)). intros T h1 l d h2. rewrite (app_removelast_last d h2) at 1. rewrite count_occ_app. simpl. destruct (h1 (last l d) (last l d)). rewrite S_compat; auto. contradict n; auto. Qed. Lemma count_occ_removelast_neq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), l <> nil -> x <> last l x -> count_occ pfdt (removelast l) x = count_occ pfdt l x. intros T h1 l x h2 h3. rewrite (app_removelast_last x h2) at 2. rewrite count_occ_app. rewrite plus_n_O at 1. f_equal. simpl. destruct (h1 (last l x) x) as [h4 | h4]. rewrite h4 in h3. contradict h3; auto. reflexivity. Qed. Lemma count_occ_remove_neq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a x:T), In x l -> a <> x -> count_occ pfdt (remove pfdt x l) a = count_occ pfdt l a. intros T h1 l. induction l as [|a l ih]; simpl; auto. intros b x h2 h3. destruct (h1 x a) as [h4 | h4]; subst. destruct (h1 a b) as [h4 | h4]; subst. contradict h3; auto. destruct (in_dec h1 a l) as [h5 | h5]. auto. rewrite <- remove_not_in'; auto. destruct (h1 a b) as [|h5]; subst. simpl. destruct (h1 b b) as [h5 | h5]. f_equal; apply ih. destruct h2; subst; auto. contradict h3; auto. assumption. contradict h5; auto. simpl. destruct (h1 a b); try contradiction. apply ih. destruct h2; subst; auto. contradict h4; auto. assumption. Qed. Lemma count_occ_remove_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), In a l -> count_occ pfdt (remove pfdt a l) a = 0. intros T h1 l a h2. apply not_lt0_iff. intro h3. rewrite <- (count_occ_In h1 (remove h1 a l) a) in h3. apply remove_In in h3; auto. Qed. Lemma count_occs0 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), (forall x:T, count_occ pfdt l x = 0) -> l = nil. intros T h1 l h2. destruct l as [|a l]; auto. specialize (h2 a). simpl in h2. destruct (h1 a a) as [|h3]. discriminate. contradict h3; auto. Qed. Lemma count_occ_gt0 : forall (l:list nat) (pf:forall n:nat, In n l -> n > 0), count_occ nat_eq_dec l 0 = 0. intro l. induction l as [|a l ih]; simpl. intros; auto. intros h1. destruct (nat_eq_dec a 0) as [|h2]; subst. specialize (h1 0 (or_introl eq_refl)). apply lt_irrefl in h1. contradiction. apply ih. intros n h3. apply h1. right; auto. Qed. (*Counts all the elements in [l] for which [pf] is true.*) Fixpoint count_occ_prop {T:Type} (P : T -> Prop) (pf : forall x, {P x} + {~ P x}) (l:list T) : nat := match l with | nil => 0 | a :: l' => let n := count_occ_prop P pf l' in if (pf a) then S n else n end. Lemma skipn_nil : forall (T:Type) (n:nat), skipn n (@nil T) = nil. intros ? n. destruct n; auto. Qed. Lemma skipn0 : forall {T:Type} (l:list T), skipn 0 l = l. auto. Qed. Lemma skipnS_cons : forall {T:Type} (l:list T) (a:T) (n:nat), skipn (S n) (a::l)= skipn n l. intros; simpl; auto. Qed. Lemma skipn_eq_nil_iff : forall {T:Type} (l:list T) (n:nat), skipn n l = nil <-> (l = nil \/ n >= length l). intros T l. induction l as [|a l ih]; simpl. intuition. apply skipn_nil. apply skipn_nil. intros n; split; intro h1. destruct n as [|n]. simpl in h1. discriminate. rewrite skipnS_cons in h1. right. rewrite ih in h1. destruct h1; subst; auto with arith. destruct h1 as [|h1]; try discriminate. destruct n as [|n]. omega. apply le_S_n in h1. rewrite skipnS_cons. rewrite ih. right; auto. Qed. Lemma skipn_length : forall {T:Type} (l:list T), skipn (length l) l = nil. intros ? l; induction l; simpl; auto. Qed. Lemma skipn_length_app : forall {T:Type} (l l':list T), skipn (length l) (l++l') = l'. intros ? l; induction l as [|l ih]; simpl; auto. Qed. Lemma skipn_length_app_plus : forall {T:Type} (l l':list T) (n:nat), skipn (length l + n) (l++l') = skipn n l'. intros ? l; induction l as [|l ih]; simpl; auto. Qed. Lemma skipn_tl : forall {T:Type} (l:list T) (n:nat), skipn n (tl l) = skipn (S n) l. intros T l. destruct l as [|a l]; simpl; auto. intro; rewrite skipn_nil; auto. Qed. Lemma firstn_nil : forall {T:Type} (n:nat), firstn n nil = nil (A:=T). intros T n. induction n as [|n h1]; auto. Qed. Lemma firstn_sing : forall {T:Type} (n:nat) (a:T), firstn (S n) (a::nil) = a::nil. simpl; intros; f_equal; rewrite firstn_nil; auto. Qed. Lemma firstn_eq_nil : forall {T:Type} (n:nat) (l:list T), firstn n l = nil -> {n = 0} + {l = nil}. intros T n. destruct n as [|n]; simpl. intros; left; auto. intros l h1. right. destruct l; auto. discriminate. Qed. Lemma firstn_S_neq_nil : forall {T:Type} (l:list T) (n:nat), S n <= length l -> firstn (S n) l <> nil. intros T l n h1; simpl. destruct l. simpl in h1. omega. intro h2. discriminate h2. Qed. Lemma firstn_cons : forall {T:Type} (l:list T) (n:nat) (a:T), 0 < n -> firstn n (a::l) = a :: firstn (pred n) l. intros T l n. revert l. induction n as [|n h1]; simpl; intros; try omega; auto. Qed. Lemma skipn_cons : forall {T:Type} (l:list T) (n:nat) (a:T), 0 < n -> skipn n (a::l) = skipn (pred n) l. intros T l n. revert l. induction n as [|n h1]; simpl; intros; try omega; auto. Qed. Lemma in_firstn : forall {T:Type} (l:list T) (m:nat) (x:T), In x (firstn m l) -> In x l. intros ? l m. revert l. induction m; simpl. intros; contradiction. intros l x h1. destruct l; auto. destruct h1; subst; [left; auto | right; auto]. Qed. Lemma in_skipn : forall {T:Type} (l:list T) (m:nat) (x:T), In x (skipn m l) -> In x l. intros ? l m. revert l. induction m; simpl; auto. intros; destruct l; auto. right; auto. Qed. Lemma nin_firstn : forall {T:Type} (l:list T) (m:nat) (x:T), ~In x l -> ~In x (firstn m l). intros; contradict H; eapply in_firstn; apply H. Qed. Lemma nin_skipn : forall {T:Type} (l:list T) (m:nat) (x:T), ~In x l -> ~In x (skipn m l). intros; contradict H; eapply in_skipn; apply H. Qed. Lemma firstn_app_le : forall {T:Type} (l l':list T) (m:nat), m <= length l -> firstn m (l ++ l') = firstn m l. intros T l. induction l as [|a l ih]; simpl. intros l' m h1. apply le0 in h1; subst. simpl; auto. intros l' m h1. destruct m as [|m]; simpl; auto. apply le_S_n in h1. f_equal; auto. Qed. Lemma firstn_app_eq : forall {T:Type} (l l':list T), firstn (length l) (l++l') = l. intros T l. induction l as [|a l ih]; simpl; auto. intro l'. f_equal. apply ih; auto. Qed. Lemma firstn_app_ge : forall {T:Type} (l l':list T) m, m >= length l -> firstn m (l++l') = l ++ (firstn (m - length l) l'). intros ? l; induction l as [|a l ih]; simpl; auto; intro l'; destruct l'; intros m ?; destruct m as [|m]; try rewrite <- minus_n_O; auto; try omega. simpl; f_equal; eapply ih; auto with arith. rewrite app_comm_cons. simpl. f_equal. eapply ih; auto with arith. Qed. Lemma firstn_length : forall {T:Type} (l:list T), firstn (length l) l = l. intros T l. rewrite <- (app_nil_r l) at 2. apply firstn_app_eq. Qed. Lemma firstn_length_eq : forall {T:Type} (l l':list T), firstn (length l) l' = l -> l' = l ++ (skipn (length l) l'). intros T l l' h1. rewrite <- (firstn_skipn (length l) l'). rewrite h1. f_equal. rewrite skipn_length_app; auto. Qed. Lemma nin_eq_firstn_length : forall {T} l l' (a:T), ~In a l -> l = firstn (length l) (a::l') -> l = nil. intros T l. destruct l as [|b l]; simpl; auto. intros l' a h1 h2; inversion h2; revert h2 h1 H1; subst. intros ? h1. contradict h1; left; auto. Qed. Lemma firstn_length_ge : forall {T:Type} (l:list T) (m:nat), length l <= m -> firstn m l = l. intros T l. induction l as [|a l ih]; simpl. intros; rewrite firstn_nil; auto. intros m h1. rewrite firstn_cons. f_equal. apply ih. omega. omega. Qed. Lemma firstn_lt_nest : forall {T:Type} (l:list T) (m n:nat), m < n -> firstn m l = firstn m (firstn n l). intros T l m n. revert m l. induction n as [|n ih]. intros; omega. intros m l h1. destruct l as [|a l]. simpl; auto. destruct m as [|m]; simpl; auto. apply lt_S_n in h1. f_equal; auto. Qed. Lemma app_cons_assoc : forall {T} (l l':list T) (a:T), l ++ a :: l' = (l++a::nil) ++ l'. intros; rewrite <- app_assoc; auto. Qed. Lemma length_from_app_cons : forall {T} (l l0 l' l0':list T) (a:T), length (l ++ a :: l') = length (l0 ++ a :: l') -> length (l ++ l') = length (l0 ++ l'). intros ? ? ? ? ? ? h1; do 2 rewrite app_length in h1; simpl in h1; do 2 rewrite plus_S_shift_r in h1; do 2 rewrite <- app_length in h1; auto with arith. Qed. Lemma app_eq_iff : forall {T:Type} (l1 l2 l:list T), l1 ++ l2 = l <-> l1 = (firstn (length l1) l) /\ l2 = (skipn (length l1) l). intros T l1 l2 l. split; intro h1. pose proof h1 as h1'. rewrite <- (firstn_skipn (length l1) l) in h1'. apply app_inj_eq_length in h1'; auto. rewrite <- h1. rewrite firstn_app_eq; auto. destruct h1 as [h1 h2]. rewrite h1, h2. apply firstn_skipn. Qed. Lemma app_eq_same_iff_l : forall {T:Type} (l l':list T), l = l ++ l' <-> l' = nil. intros T l l'; split; intro h1. destruct (nil_dec' l') as [|h2]; auto. apply not_app_same_eq_l in h1; auto; contradiction. subst. rewrite app_nil_r; auto. Qed. Lemma app_eq_same_iff_r : forall {T:Type} (l l':list T), l = l' ++ l <-> l' = nil. intros T l l'; split; intro h1. destruct (nil_dec' l') as [|h2]; auto. apply not_app_same_eq_r in h1; auto; contradiction. subst. rewrite app_nil_l; auto. Qed. Lemma app_eq_same_iff2 : forall {T:Type} (l l' l'':list T), l = l' ++ l ++ l'' <-> l' = nil /\ l'' = nil. intros T l l' l''; split; intro h1. destruct (nil_dec' l') as [|h2]. split; auto. rewrite e in h1. rewrite app_nil_l in h1. destruct (nil_dec' l'') as [|h2]; auto. apply not_app_same_eq_l in h1. contradiction. assumption. pose proof (f_equal (@length T) h1) as h3. do 2 rewrite app_length in h3. apply not_nil_lt in h2. omega. destruct h1; subst. rewrite app_nil_l, app_nil_r; auto. Qed. Lemma eq_in_cons : forall {T:Type} (l:list T) (x a:T), x = a -> In x (a::l). intros; subst; left; auto. Qed. Lemma neq_in_cons : forall {T:Type} (l:list T) (x a:T), x <> a -> In x (a::l) -> In x l. intros ? ? ? ? h0 h1; destruct h1; subst; auto. contradict h0; auto. Qed. Lemma in_list2_comm : forall {T:Type} (l:list T) (a x y:T), In a (x::y::l) -> In a (y::x::l). intros T l a x y h1. destruct h1 as [h1a | h1b]. subst. right. left. auto. destruct h1b as [h2a | h2b]. subst. left. auto. right. right. auto. Qed. Lemma nin_remove_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), ~In a l -> remove pfdt a l = l. intros T h0 l. induction l as [|a l h1]; simpl; auto. intros x h2. apply not_or_and in h2. destruct h2 as [h2 h3]. destruct (h0 x a) as [h4 | h4]. subst. contradict h2; auto. f_equal. apply h1; auto. Qed. Lemma remove_cons_neq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a b:T), a <> b -> remove pfdt a (b::l) = b :: remove pfdt a l. intros T h0 l a b h1. simpl. destruct (h0 a b). contradiction. auto. Qed. Lemma remove_cons_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), remove pfdt a (a::l) = remove pfdt a l. intros T h0 l a. simpl. destruct (h0 a a) as [h1 | h1]; auto. contradict h1; auto. Qed. Lemma remove_cons_eq_l' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x a:T), x = a -> remove pfdt x (a::l) = remove pfdt x l. intros; subst; apply remove_cons_eq; auto. Qed. Lemma remove_cons_eq_r' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x a:T), x = a -> remove pfdt x (a::l) = remove pfdt a l. intros; subst; apply remove_cons_eq; auto. Qed. Lemma remove_a_from_2cons_nil : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a x y:T), remove pfdt a (x::y::l) = nil -> remove pfdt a (y::x::l) = nil. intros T h0 l a x y h1. simpl in h1. simpl. destruct (h0 a x); destruct (h0 a y); auto; try discriminate. Qed. Lemma remove2_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a b:T), remove pfdt a (remove pfdt b l) = remove pfdt b (remove pfdt a l). intros T h0 l a b. induction l as [|c l h1]. simpl. auto. simpl. destruct (h0 b c) as [hc | hb]; destruct (h0 a c) as [hc' | ha]. subst; auto. subst. simpl. destruct (h0 c c) as [h3 | h4]. auto. contradict h4. auto. subst. simpl. destruct (h0 c c) as [h3 | h4]. auto. contradict h4. auto. simpl. destruct (h0 a c) as [hc | ha']; destruct (h0 b c) as [hc' | hb']; subst; try contradiction; auto. contradict ha. auto. contradict hb; auto. f_equal. auto. Qed. Lemma in_remove_neq' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), In x l -> l <> remove pfdt x l. intros T hdt l x h1 h2. rewrite h2 in h1. pose proof (remove_In hdt l x). contradiction. Qed. Lemma remove_a_in_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a x:T), In x l -> x <> a -> In x (remove pfdt a l). intros T hd l a x h1 h3. destruct (in_dec hd a l) as [h2 | h0]. induction l as [|b l h7]. contradiction. simpl. destruct (hd a b) as [h4 | h5]. subst. destruct h1 as [h1a | h1b]. symmetry in h1a. contradiction. destruct (in_dec hd b l) as [h8 | h9]. apply h7; auto. pose proof (remove_not_in' hd _ _ h9) as h10. rewrite <- h10. auto. destruct h1 as [h1a | h1b]; destruct h2 as [h2a | h2b]. subst. contradiction. subst. left. auto. subst. contradict h5. auto. pose proof (h7 h1b h2b) as h8. right. auto. rewrite <- remove_not_in'; assumption. Qed. Lemma in_remove_neq_in_l : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a b:T), In a (remove pfdt b l) -> a <> b -> In a l. intros T h0 l a b h1 h2. induction l as [|c l h3]. simpl in h1. auto. simpl in h1. destruct (h0 b c). subst. specialize (h3 h1). right; auto. destruct h1 as [h1a | h1b]. subst. left. auto. specialize (h3 h1b). right. auto. Qed. Lemma in_remove_inv : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a b:T), In a (remove pfdt b l) -> In a l /\ a <> b. intros T h0 l a b h1. pose proof remove_In. destruct (h0 a b) as [h2 | h3]. subst. pose proof (remove_In h0 l b). contradiction. split; auto. eapply in_remove_neq_in_l; auto. apply h1. assumption. Qed. Lemma in_remove_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a b:T), In a (remove pfdt b l) <-> In a l /\ a <> b. intros T h0 l a b. split. intro h1. apply (in_remove_inv h0); auto. intro h1. destruct h1. apply remove_a_in_eq; auto. Qed. Lemma in_remove_impl : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a b:T), In a (remove pfdt b l) -> In a l. intros T h1 l a b h2. pose proof (in_remove_iff h1 l a b) as h3. rewrite h3 in h2. destruct h2; auto. Qed. Lemma nin_nin_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x a:T), ~In x l -> ~In x (remove pfdt a l). intros T h0 l x a h1. induction l as [|b l h2]. simpl. auto. simpl in h1. assert (h3:b <> x /\ ~ In x l). tauto. destruct h3 as [h3a h3b]. pose proof (h2 h3b) as h4. simpl. destruct (h0 a b) as [h5 | h6]. subst. auto. simpl. tauto. Qed. Lemma nin_remove_eq_nin : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a b:T), ~ In a (remove pfdt b l) -> a = b \/ ~ In a l. intros T h0 l a b h1. induction l as [|c l h2]. right. intro; auto. simpl in h1. destruct (h0 b c) as [h3 | h4]. subst. specialize (h2 h1). simpl. rewrite eq_sym_iff. rewrite eq_sym_iff. tauto. simpl in h1. assert (h3:c <> a /\ ~ In a (remove h0 b l)). tauto. destruct h3 as [h3l h3r]. specialize (h2 h3r). destruct h2; [left; auto | idtac]. right. simpl. tauto. Qed. Lemma in_tl : forall {T:Type} (l:list T) (x:T), In x (tl l) -> In x l. intros T l. induction l; simpl; auto. Qed. Lemma in_tl_lt : forall {T:Type} (l:list T) (x:T), In x (tl l) -> 1 < length l. intros T l. destruct l as [|a l]; simpl. intro; contradiction. intros ? h1 ; apply lt_n_S. apply not_nil_lt. eapply in_not_nil. apply h1. Qed. Lemma impl_in_tl : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (m:nat), S m < count_occ pfdt l x -> In x (tl l). intros T h1 l. destruct l as [|a l]; simpl. intros; omega. intros x m h2. destruct (h1 a x) as [|h3]; subst. apply lt_S_n in h2. apply count_occ_in in h2; auto. apply count_occ_in in h2; auto. Qed. Lemma length_tl : forall {T:Type} (l:list T), length (tl l) = pred (length l). intros ? l. induction l; auto. Qed. Lemma S_length_tl : forall {T:Type} (l:list T), l <> nil -> S (length (tl l)) = length l. intros T l; induction l; auto. intro h1. contradict h1; auto. Qed. Lemma lt_length_tl : forall {T:Type} (l:list T) (m:nat), S m < length l -> m < length (tl l). intros T l m h1. pose proof (lt_length _ _ h1) as h2. pose proof (hd_compat' _ h2) as h3. simpl in h3. rewrite h3 in h1. simpl in h1. apply lt_S_n in h1; auto. Qed. Lemma le_length_tl : forall {T:Type} (l:list T) (m:nat), S m <= length l -> m <= length (tl l). intros T l m h1. rewrite length_tl. omega. Qed. Lemma firstn_tl : forall {T:Type} (l:list T) (m:nat), S m <= length l -> firstn m (tl l) = tl (firstn (S m) l). intros T l m. revert l. induction m as [|a ih]; simpl. intros l h1. destruct l as [|a l]; auto. intros l h1. rewrite (hd_compat' _ (lt_length _ _ h1)). simpl; auto. Qed. Lemma tl_not_nil : forall {T:Type} (l:list T), tl l <> nil -> l <> nil. intros ? ? h1; intro; subst; simpl in h1; contradict h1; auto. Qed. Lemma tl_not_nil_lt : forall {T:Type} (l:list T), tl l <> nil -> 1 < length l. intros T l h1; destruct l; simpl. contradict h1; auto. simpl in h1. destruct l as [|a l]; simpl; auto with arith. contradict h1; auto. Qed. Lemma tl1lt : forall {T:Type} (l:list T), 1 < length l -> tl l <> nil. intros T l h1. destruct l as [|a l]. simpl in h1. apply lt_n_0 in h1. contradiction. simpl in h1. apply lt_S_n in h1. simpl. apply lt_length in h1; auto. Qed. Lemma tl_not_nil_le : forall {T:Type} (l:list T), tl l <> nil -> 2 <= length l. intros; apply tl_not_nil_lt; auto. Qed. Lemma tl_app : forall {T:Type} (l l':list T), l <> nil -> tl (l ++ l') = (tl l) ++ l'. intros T l. destruct l as [|a l]; simpl; auto. intros ? h1. contradict h1; auto. Qed. Lemma removelast_tl : forall {T:Type} (l:list T), removelast (tl l) = tl (removelast l). intros T l; destruct l as [|? l]; try destruct l; auto. Qed. Lemma tl_sing : forall {T} (l:list T) (pf:l <> nil), tl l = nil -> l = hd' l pf :: nil. intros T l h1; destruct l as [|? l]. contradict h1; auto. simpl. intro; subst; auto. Qed. Lemma tl_seq : forall (m n:nat), tl (seq m n) = seq (S m) (pred n). intro m; induction m; simpl; auto; intro n; destruct n; auto. Qed. Lemma count_occ_tl : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pf: l <> nil) (x:T), x <> hd' l pf -> count_occ pfdt (tl l) x = count_occ pfdt l x. intros T h1 l h2 x h3. rewrite (hd_compat' _ h2) at 2. simpl. destruct (h1 (hd' l h2) x); subst; auto. contradict h3; auto. Qed. Lemma count_occ_tl_dec : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), let c := count_occ pfdt (tl l) x in let c' := count_occ pfdt l x in {c = c'} + {c = pred c'}. intros T h1 l. destruct l as [|a l]; simpl. intros; left; auto. intro x. destruct (h1 a x); [right; auto | left; auto]. Qed. Definition tl2 {T:Type} (l:list T) := tl (tl l). Definition hd2 {T:Type} (l:list T) (pf:1 < length l) : T := hd' (tl l) (tl1lt l pf). Lemma length2_not_nil : forall {T:Type} (l:list T), length l = 2 -> l <> nil. intros; intro; subst. simpl in H. discriminate. Qed. Lemma length2_lt : forall {T:Type} (l:list T), length l = 2 -> 1 < length l. intros; omega. Qed. Lemma length2 : forall {T:Type} (l:list T) (pf:length l = 2), l = hd' l (length2_not_nil l pf) :: hd2 l (length2_lt _ pf) :: nil. intros T l. destruct l as [|a l]; simpl. intros; discriminate. intros h1. f_equal. unfold hd2. simpl. pose proof h1 as h2. apply S_inj in h2. rewrite (length_eq_1 _ h2) at 1. repeat f_equal. apply proof_irrelevance. Qed. Lemma skipn_plus : forall {T:Type} (l:list T) (m n:nat), skipn (m + n) l = skipn m (skipn n l). intros T l m n. revert m l. induction n as [|n ih]; simpl; auto. intro m. rewrite plus_0_r; auto. intros m l. rewrite plus_S_shift_r. destruct (nil_dec' l) as [|h1]; subst. do 2 rewrite skipn_nil; auto. rewrite (hd_compat' _ h1). simpl. apply ih. Qed. Lemma skipn_app : forall {T:Type} (l l':list T) (n:nat), n <= length l -> skipn n (l ++ l') = skipn n l ++ l'. intros T l l' n. revert l l'. induction n as [|n ih]; simpl; auto. intros l l' h1. pose proof (lt_length _ _ h1) as h2. rewrite (hd_compat' _ h2). simpl. apply ih. rewrite length_tl. omega. Qed. Lemma app_sing_inj : forall {T} (l l':list T) a b, l ++ a::nil = l' ++ b::nil -> l = l' /\ a = b. intros T l l' a b h1. pose proof (f_equal (fun m => last m a) h1) as h2; simpl in h2. do 2 rewrite last_app_sing in h2; subst. split; auto. pose proof (f_equal (firstn (length l)) h1) as h2. rewrite firstn_app_le in h2; auto with arith. rewrite firstn_length in h2. rewrite h2. pose proof (f_equal (@length _) h1) as h3. do 2 rewrite app_length in h3; simpl in h3; do 2 rewrite S_compat in h3; apply S_inj in h3. rewrite h3. rewrite firstn_app_le; auto. apply firstn_length. Qed. Lemma app_sings_eq_iff : forall {T:Type} (l l':list T) (a a':T), l ++ a::nil = l' ++ a'::nil <-> l = l' /\ a = a'. intros T l l' a a'; split; intro h1. pose proof (f_equal (@removelast _) h1) as h2. rewrite removelast_app_sing in h2. subst. rewrite removelast_app_sing in h1. rewrite removelast_app_sing. split; auto. pose proof (f_equal (fun l => (last l a)) h1) as h2. simpl in h2. do 2 rewrite last_app_sing in h2; auto. destruct h1; subst; auto. Qed. (*This can't be parameterized over type [Type] instead of [nat], because of universe inconsistencies.*) Fixpoint list_sum (P:nat->Type) (lP:forall i, list (P i)) (ln:list nat) : Type := match ln with | nil => unit | a :: ln' => (P a + list_sum P lP ln')%type end. Fixpoint list_eq {T:Type} (l:list T) : Prop := match l with | nil => True | a::l' => match l' with | nil => True | b::l0 => a = b /\ list_eq l' end end. Lemma list_eq_nil : forall (T:Type), list_eq (@nil T). simpl; auto. Qed. Lemma list_eq_rev_cons : forall {T:Type} (l:list T) (a:T) (pfn:l <> nil), list_eq (a::l) -> a = hd' l pfn /\ list_eq l. simpl; intros T l a h1 h2. rewrite (hd_compat' _ h1) in h2. destruct h2 as [? h2]; subst. rewrite <- hd_compat' in h2. split; auto. Qed. Lemma list_eq_cons : forall {T:Type} (l:list T) (a:T), list_eq l -> (exists x:T, In x l /\ x = a) -> list_eq (a::l). intros T l. induction l as [|b l ih]; simpl; auto. intros a h1 h2. destruct h2 as [x [h2 ?]]; subst. destruct h2 as [|h2]; subst. split; auto. split; auto. rewrite (hd_compat' _ (in_not_nil _ _ h2)) in h1. destruct h1 as [? h1]; subst. rewrite <- hd_compat' in h1. specialize (ih a h1). hfirst ih. exists a. split; auto. specialize (ih hf). simpl in ih. rewrite (hd_compat' _ (in_not_nil _ _ h2)) in ih. destruct ih; auto. Qed. Lemma list_eq_compat : forall {T:Type} (l:list T), list_eq l <-> (forall x y:T, In x l -> In y l -> x = y). intros T l. split. induction l as [|a l ih]; simpl. intros; contradiction. intros ha x y hb hc. destruct hb as [|hb]; subst. destruct hc as [|hc]; auto. rewrite (hd_compat' _ (in_not_nil _ _ hc)) in ha. destruct ha as [? ha]; subst. rewrite <- hd_compat' in ha. apply (ih ha); auto. apply in_hd'. destruct hc as [|hc]; subst. rewrite (hd_compat' _ (in_not_nil _ _ hb)) in ha. destruct ha as [? ha]; subst. apply ih. rewrite <- hd_compat' in ha; auto. assumption. apply in_hd'. rewrite (hd_compat' _ (in_not_nil _ _ hb)) in ha. destruct ha as [? ha]; subst. rewrite <- hd_compat' in ha. apply ih; auto. induction l as [|a l ih]; simpl; auto. intro h1. hfirst ih. intros; apply h1; right; auto. destruct (nil_dec' l) as [|h2]; subst; auto. rewrite (hd_compat' _ h2). destruct (in_dec eq_dec a l). split. apply hf; auto. apply in_hd'. rewrite <- hd_compat'. apply ih; auto. split. apply h1. left; auto. right. apply in_hd'. rewrite <- hd_compat'. apply ih; auto. Qed. Fixpoint conj_list (l:list Prop) := match l with | nil => True | P::l' => P /\ conj_list l' end. Fixpoint disj_list (l:list Prop) := match l with | nil => False | P::l' => P \/ disj_list l' end. Lemma conj_list_compat : forall (l:list Prop), conj_list l <-> (forall (P:Prop) , In P l -> P). intro l. induction l as [|Q l ih]; simpl. split; auto. intros; contradiction. split; intro h1. intros P h2. destruct h2 as [|h2]; subst. destruct h1; auto. destruct h1 as [h1 h3]. rewrite ih in h3. apply h3; auto. hr ih. intros P h2. apply h1. right; auto. rewrite <- ih in hr. split; auto. apply h1. left; auto. Qed. Lemma disj_list_compat : forall (l:list Prop), disj_list l <-> (exists (P:Prop) , In P l /\ P). intro l. induction l as [|Q l ih]; simpl. split; auto. intros; contradiction. intro h1. destruct h1 as [? [? ?]]; contradiction. split; intro h1. destruct h1 as [h1 | h1]. exists Q. split; auto. rewrite ih in h1. destruct h1 as [P [h1 h2]]. exists P. split; auto. destruct h1 as [P [[h2 | h2] h3]]; subst. left; auto. right. rewrite ih. exists P. split; auto. Qed. Definition list_to_option {T} (l:list T) : list (option T) := None :: (map (fun x => Some x) l). Lemma length_list_to_option : forall {T} (l:list T), length (list_to_option l) = S (length l). unfold list_to_option; intros; simpl; rewrite map_length; auto. Qed. (*A list of all sequence lists obtained by choosing one element from each list in [l], in the same order as the lists in [l]*) Fixpoint list_of_lists_seqs {T:Type} (l:faml T) : faml T := match l with | nil => cons nil nil | al::l' => map (fun x:(T * (list T)) => cons (fst x) (snd x)) (list_prod al (list_of_lists_seqs l')) end. Fixpoint list_to_set {T:Type} (l:list T) : Ensemble T := match l with | nil => Empty_set _ | a::l' => Add (list_to_set l') a end. Fixpoint faml_to_fam {T:Type} (ll:faml T) : Family T := match ll with | nil => Empty_set _ | la::ll' => Add (faml_to_fam ll') (list_to_set la) end. Lemma finite_faml_to_fam : forall {T:Type} (ll:faml T), Finite (faml_to_fam ll). intros T ll. induction ll as [|a ll h1]. constructor. simpl. apply Add_preserves_Finite; auto. Qed. Lemma empty_set_nil : forall {T:Type} (l:list T), list_to_set l = Empty_set _ -> l = nil. intros T l h1. induction l as [|a l]. reflexivity. simpl in h1. pose proof (Add_intro2 _ (list_to_set l) a) as h2. rewrite h1 in h2. contradiction. Qed. Lemma list_to_set_sing : forall {T:Type} (a:T), list_to_set (a::nil) = Singleton a. intros T a. simpl. symmetry. rewrite add_empty_sing. reflexivity. Qed. Lemma length_app_nprod_to_list : forall {T m n} (x:T^m) (y:T^n), length ((nprod_to_list _ _ x) ++ (nprod_to_list _ _ y)) = m + n. intros T m. induction m as [|m ih]; simpl. intros; rewrite length_nprod_to_list; auto. intros n x y. destruct x as [t x]; simpl; f_equal; auto. Qed. Lemma nprod_of_list_inj : forall {T} (l l':list T) (pfe:length l = length l'), nprod_of_list _ l = transfer_nprod_r pfe (nprod_of_list _ l') -> l = l'. intros T l. induction l as [|a l ih]; simpl . intros; subst. pose proof (eq_sym pfe) as h2. apply length_eq0 in h2; auto. intros l'. destruct l' as [|b l']; simpl. intros; discriminate. intros h1 h2. rewrite transfer_nprod_r_S in h2; simpl in h2. inversion h2; subst. rename H1 into h3. f_equal. specialize (ih l' (S_inj _ _ h1) h3); auto. Qed. Fixpoint lrep {T:Type} (x:T) (n:nat) := match n with | O => nil | S n => x :: lrep x n end. Lemma lrep_nil_iff : forall {T:Type} (a:T) (n:nat), lrep a n = nil <-> n = 0. intros T a n. destruct n as [|n]; split; intros h1; auto. simpl in h1. discriminate. discriminate. Qed. Lemma lrep_eq_iff : forall {T:Type} (a:T) (n:nat) (l:list T), lrep a n = l <-> ((forall x:T, In x l -> x = a) /\ length l = n). intros T a n. revert a. induction n as [|n ih]; simpl. intros a l; split; intro h1; subst; simpl. split; auto. intros; contradiction. destruct h1 as [? h1]. apply length_eq0 in h1; auto. intros a l; split; intro h2. destruct l as [|a' l]. discriminate. inversion h2; subst. simpl. specialize (ih a' (lrep a' n)). pose proof (iff1 ih eq_refl) as h3. destruct h3 as [h3 h4]. split; auto with arith. intros x h5. destruct h5; auto. destruct h2 as [h2 h3]. destruct l as [|a' l]. discriminate. simpl in h3. apply S_inj in h3. pose proof (h2 _ (in_eq _ _)); subst. f_equal. rewrite ih. split. intros; apply h2. right; auto. reflexivity. Qed. Lemma length_lrep : forall {T:Type} (a:T) (n:nat), length (lrep a n) = n. intros ? ? n; induction n as [|n ih]; simpl; auto. Qed. Lemma in_lrep : forall {T:Type} (a:T) (n:nat), lrep a n <> nil -> In a (lrep a n). intros T a n. destruct n; simpl. intro h1. contradict h1; auto. intros; left; auto. Qed. Lemma in_lrep' : forall {T:Type} (a b:T) (n:nat), In b (lrep a n) -> b = a. intros T a b n. induction n as [|n ih]; simpl. intros; contradiction. intro h1. destruct h1 as [|h1]; auto. Qed. Definition all_sing {T:Type} (l:list T) (a:T) := forall x : T, In x l -> x = a. Lemma all_sing_cons_iff : forall {T:Type} (l:list T) (a x:T), all_sing (a::l) x <-> x = a /\ all_sing l a. intros T l a x; split; intro h1. red in h1. pose proof (h1 a (in_eq _ _)); subst. split; auto. red; intros c h2. apply h1. right; auto. destruct h1 as [? h1]; subst. red in h1; red. intros x h2. destruct h2; auto. Qed. Lemma all_sing_cons_app : forall {T:Type} (x:T) (l:list T), all_sing l x -> x :: l = l++x::nil. intros T x l. revert l. induction l as [|a l ih]; simpl; auto. intros h1. rewrite all_sing_cons_iff in h1. destruct h1 as [? h1]; subst. specialize (ih h1). rewrite ih; auto. Qed. Lemma all_sing_count_occ : forall {T:Type} (pfdt:type_eq_dec T) (x:T) (l:list T), all_sing l x -> count_occ pfdt l x = length l. intros T h1 x l. revert x. induction l as [|a l ih]; simpl; auto. intros x h2. destruct (h1 a x) as [h3 | h3]; subst. f_equal; apply ih. red in h2; red. intros; apply h2; right; auto. red in h2. specialize (h2 _ (in_eq _ _)). contradiction. Qed. Lemma all_sing_lrep'' : forall {T:Type} (l:list T) (a:T), all_sing l a -> exists! n:nat, l = lrep a n. intros T l. induction l as [|a l ih]; simpl. intros a h1. exists 0. red; simpl; split; auto. intros n h2. symmetry in h2. rewrite lrep_nil_iff in h2; auto. intros x h1. rewrite all_sing_cons_iff in h1. destruct h1 as [h1 h2]; subst. specialize (ih _ h2). destruct ih as [n h1]. red in h1. destruct h1 as [h1 h3]; subst. exists (S n). red. split. simpl; auto. intros c h4. destruct c as [|c]. simpl in h4. discriminate. simpl in h4. inversion h4 as [h5]. specialize (h3 _ h5); subst; auto. Qed. Lemma all_sing_lrep : forall {T:Type} (l:list T) (a:T), all_sing l a -> l = lrep a (length l). intros T l. induction l as [|a l ih]; simpl. intros; auto. intros x h1. rewrite all_sing_cons_iff in h1. destruct h1 as [? h1]; subst. f_equal. apply ih; auto. Qed. Lemma all_sing_lrep' : forall {T:Type} (a:T) (n:nat), all_sing (lrep a n) a. intros T a n. revert a. induction n as [|n ih]; simpl. intros; red. intros; contradiction. intros a b h1. destruct h1 as [|h1]; subst; auto. apply in_lrep' in h1; auto. Qed. Lemma tl_lrep : forall {T:Type} (a:T) (n:nat), tl (lrep a n) = lrep a (pred n). intros T a n. destruct n as [|n]; simpl; auto. Qed. Lemma remove_lrep : forall {T:Type} (pfdt:type_eq_dec T) (a:T) (n:nat), remove pfdt a (lrep a n) = nil. intros T h1 a n. induction n; simpl; auto. destruct (h1 a a) as [|h2]; simpl; auto. contradict h2; auto. Qed. Lemma all_sing_no_dup : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), NoDup l -> l <> nil -> all_sing l a -> l = a::nil. intros T h1 l. induction l as [|b l ih]. intros ? ? h2. contradict h2; auto. intros a h2 h3 h4. red in h4. pose proof (h4 _ (in_eq _ _)) as he; subst. f_equal. destruct (nil_dec' l) as [|h5]; auto. eapply ih in h5. rewrite h5 in h2. apply no_dup_cons_nin in h2. contradict h2; left; reflexivity. apply no_dup_cons in h2; auto. red. intros; apply h4. right; auto. Qed. Lemma all_sing_eq_iff : forall {T} (l l':list T) a, all_sing l a -> all_sing l' a -> (l = l' <-> length l = length l'). intros T l. induction l as [|b l]; simpl; intro l'; destruct l' as [|b' l']; simpl; intros; intuition; try discriminate. f_equal. inversion H1; subst; auto. apply S_inj in H1. pose proof H as h2. specialize (h2 _ (in_eq _ _)); subst. pose proof H0 as h1. specialize (h1 _ (in_eq _ _)); subst. f_equal. rewrite <- IHl in H1; subst; auto. red. intros x h1. specialize (H _ (in_cons _ _ _ h1)). apply H. red. intros x h1. specialize (H0 _ (in_cons _ _ _ h1)); auto. Qed. Lemma all_sing_nprod_eq : forall {T n} (x y:T^n) a, all_sing_nprod x a -> all_sing_nprod y a -> x = y. intros ? ? ? ? ? h1 h2; apply nprod_ext; intros i hi. red in h1, h2. specialize (h1 _ (in_nth_prod i hi x)); specialize (h2 _ (in_nth_prod i hi y)); congruence. Qed. Lemma all_sing_length : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), all_sing l a -> length l = count_occ pfdt l a. intros T h1 l. induction l as [|a' l ih]; simpl; auto. intros a h2. rewrite all_sing_cons_iff in h2. destruct h2 as [? h2]; subst. destruct (h1 a' a') as [h3 | h3]; auto with arith. contradict h3; auto. Qed. Lemma all_sing_firstn : forall {T:Type} (l:list T) (a:T) (m:nat), all_sing l a -> all_sing (firstn m l) a. intros ? ? ? ? h1. red in h1; red. intros ? h2; apply h1. apply in_firstn in h2; auto. Qed. Lemma lrep1 : forall {T:Type} (a:T), a::nil = lrep a 1. auto. Qed. Lemma remove_sing : forall {T:Type} (pfdt:type_eq_dec T) (a:T), remove pfdt a (a::nil) = nil. intros; rewrite lrep1; apply remove_lrep. Qed. Lemma pred_lt_length : forall {T:Type} (l:list T) (n:nat) (pf:n < S (length l)) (pfn:l <> nil), pred n < length l. intros T l n. destruct n as [|n]; simpl. intros; apply not_nil_lt; auto. intros h1 h2. apply lt_S_n in h1; auto. Qed. Lemma in_all_sing : forall {T:Type} (l:list T) (a x:T), all_sing l a -> In x l -> x = a. intros ? ? ? ? h1; apply h1; auto. Qed. Lemma all_sing_in : forall {T:Type} (l:list T) (a:T), (l <> nil) -> all_sing l a -> In a l. intros T l a h1. destruct l as [|b l']. contradict h1; auto. intro h2. pose proof (in_eq b l') as h3. specialize (h2 _ h3). subst. auto. Qed. Lemma all_sing_cons_in : forall {T:Type} (l:list T) (a c:T), l <> nil -> all_sing (a::l) c -> In a l. intros T l a c h1 h2. pose proof (all_sing_lrep _ _ h2) as h3. simpl in h3. inversion h3 as [h4]. apply in_lrep; auto. intro h5. pose proof (not_nil_lt _ h1) as h6. rewrite (S_pred _ _ h6) in h5. simpl in h5. discriminate. Qed. Lemma list_to_set_in_iff : forall {T:Type} (l:list T) (x:T), In x l <-> Inn (list_to_set l) x. intros T l; induction l as [|a l h2]; intro x; simpl; split; intro h1; try contradiction; try destruct h1 as [|h1]; subst. right; constructor; auto. left; rewrite <- h2; auto. right; rewrite h2; auto. destruct H; left; auto. Qed. Lemma impl_all_sing : forall {T:Type} (l:list T) (a:T), list_to_set l = Singleton a -> all_sing l a. intros T l a h1 x h2. induction l as [|a' l h3]. simpl in h1. pose proof (In_singleton _ a) as h4. rewrite <- h1 in h4. contradiction. simpl in h1, h2. pose proof (Add_intro2 _ (list_to_set l) a') as h5. rewrite h1 in h5. destruct h5; subst. destruct h2 as [|h2]; subst; auto. rewrite list_to_set_in_iff in h2. pose proof (Add_intro1 _ (list_to_set l) a _ h2) as h7. rewrite h1 in h7. destruct h7; auto. Qed. Lemma all_sing_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x y:T), all_sing l x -> x <> y -> l = remove pfdt y l. intros T h1 l x y h2 h3. symmetry. apply nin_remove_eq. intro h4. red in h2. contradict h3. symmetry; auto. Qed. Lemma all_sing_remove' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), all_sing l x -> remove pfdt x l = nil. intros T h1 l x h2. apply all_sing_lrep in h2. rewrite h2. apply remove_lrep. Qed. Lemma all_sing_tl : forall {T:Type} (l:list T) (a:T), all_sing l a -> all_sing (tl l) a. intros T l a h1. red in h1; red. intros ? h2; apply h1; auto. apply in_tl; auto. Qed. Lemma all_sing_nth : forall {T:Type} (l:list T) (a:T), all_sing l a -> forall m:nat, nth m l a = a. intros T l. induction l as [|a' l ih]; simpl. intros a h1 m. destruct m; auto. intros a h1 m. rewrite all_sing_cons_iff in h1. destruct h1 as [h1 h2]. destruct m; auto. subst. apply ih; auto. Qed. Lemma all_sing_nil : forall {T:Type} (a:T), all_sing nil a. intros; red; intros; contradiction. Qed. Lemma app_eq_sing_iff : forall {T:Type} (l l':list T) (a:T), l ++ l' = a::nil <-> (l = nil /\ l' = a::nil) \/ (l' = nil /\ l = a::nil). intros T l l'. destruct l as [|a l], l' as [|b l']; simpl. intro a; split; intro h1. discriminate. destruct h1 as [h1|h1]; destruct h1; discriminate. intro a; split; intro h2. left; auto. destruct h2 as [h2|h2]; destruct h2; try discriminate; auto. intro b; split; intro h1; simpl in h1. rewrite app_nil_r in h1. right; split; auto. destruct h1 as [h1|h1]. destruct h1; discriminate. rewrite app_nil_r; destruct h1; auto. intro c; split; intro h1. inversion h1; subst. pose proof (app_cons_not_nil l l' b) as h2. rewrite H1 in h2. contradict h2; auto. destruct h1 as [h1|h1]; destruct h1; discriminate. Qed. Fixpoint app_list {T:Type} (ll:faml T) := match ll with | nil => nil | la::ll' => la ++ app_list ll' end. Lemma app_list_eq_nil_iff : forall {T:Type} (ll:faml T), app_list ll = nil <-> ll = nil \/ all_sing ll nil. intros T ll; induction ll as [|a ll ih]; simpl; try tauto. split; intro h1. right. apply app_eq_nil in h1. destruct h1; subst. rewrite ih in H0. destruct H0 as [|h1]; subst. red; intros l h1; destruct h1; auto. contradiction. red; intros l h2; destruct h2; auto. destruct h1 as [|h1]. discriminate. rewrite all_sing_cons_iff in h1. destruct h1; subst. simpl. rewrite ih; right; auto. Qed. Lemma app_list_nprod_to_list_inj : forall {T n} (x y:(list T)^n), (forall i (pfi:i < n), length (nth_prod i pfi x) = length (nth_prod i pfi y)) -> app_list (nprod_to_list _ _ x) = app_list (nprod_to_list _ _ y) -> x = y. intros T n. induction n as [|n ih]; simpl. intros; apply unit_eq. intros l l' h1 h2. destruct l as [fl sl], l' as [fl' sl']; simpl. simpl in h1. specialize (ih sl sl'). simpl in h2. apply app_inj_eq_length in h2. destruct h2 as [? h3]; subst; f_equal. hfirst ih. intros i hi. specialize (h1 (S i) (lt_n_S _ _ hi)); simpl in h1. erewrite f_equal_dep. rewrite h1 at 1; auto. reflexivity. specialize (ih hf h3). destruct ih; subst; auto; simpl. specialize (h1 0 (lt_0_Sn _)); simpl in h1; auto. Qed. Lemma app_list_all_sing : forall {T} (ll:faml T), all_sing ll nil -> app_list ll = nil. intros ? ll; induction ll as [|l ll ih]; simpl; auto; intro h1. hfirst ih. red in h1; red; intros; apply h1; left; auto. pose proof (h1 _ (in_eq _ _)); subst. symmetry; apply h1; right; auto. specialize (ih hf). rewrite ih. red in h1. specialize (h1 _ (in_eq _ _)); subst; auto. Qed. Lemma list_to_set_finite : forall {T:Type} (l:list T), Finite (list_to_set l). intros T l. induction l as [|a l h1]. simpl. auto with sets. simpl. apply Add_preserves_Finite; assumption. Qed. Lemma list_to_set_in_cons : forall {T:Type} (l:list T) (a:T), In a l -> list_to_set (a::l) = list_to_set l. intros T l a h1. simpl. apply in_add_eq. rewrite <- list_to_set_in_iff. assumption. Qed. Lemma list_to_set_list_eq : forall {T:Type} (l l':list T), list_to_set l = list_to_set l' -> list_eq l -> list_eq l'. intros T l l' h1 h2. rewrite list_eq_compat in h2. rewrite list_eq_compat. intros; apply h2. rewrite list_to_set_in_iff. rewrite h1. rewrite <- list_to_set_in_iff; auto. rewrite list_to_set_in_iff. rewrite h1. rewrite <- list_to_set_in_iff; auto. Qed. Definition length_eq {T:Type} (ll:faml T) : Prop := list_eq (map (@length T) ll). Definition slength_eq {T:Set} (ll:faml T) : Prop := list_eq (map (@length T) ll). Section Matrix. (*A [list] of [lists] all of the same length. Each [list] is thought of as a column, which convention was used to emphasize the independent nature of the columns/[list]s. And the index in the [faml] is considered the row number.*) Inductive matrix {T:Type} : faml T -> Prop := | matrix_intro : forall fl, length_eq fl -> matrix fl. Definition nat_matrix m n : Type := match m with | O => unit | S _ => match n with | O => unit | S _ => list (list nat) end end. End Matrix. Lemma map_eq_nil_iff : forall {T U:Type} (f:T->U) (l:list T), map f l = nil <-> l = nil. intros ? ? ? l. destruct l; intuition; try discriminate. Qed. Lemma no_dup_map_inj : forall {T U:Type} (l:list T) (f:T->U), NoDup l -> Injective f -> NoDup (map f l). intros T U l. induction l as [|a l h1]. simpl. intros f h1 h2. constructor. intros f h2 h3. simpl. inversion h2. subst. specialize (h1 _ H2 h3). constructor. rewrite in_map_iff. intro h4. destruct h4 as [x h5]. destruct h5 as [h5l h5r]. red in h3. specialize (h3 _ _ h5l). subst. contradiction. assumption. Qed. Lemma hd_map' : forall {T U:Type} (f:T->U) (l:list T) (pf:map f l <> nil), let pf' := (iffn1 (map_eq_nil_iff f l) pf) in hd' (map f l) pf = f (hd' l pf'). intros T U f l h1 h2. gen0. generalize h2. rewrite (hd_compat' _ h2). intros h4 h5. simpl; auto. Qed. Lemma length_eq_nil : forall (T:Type), length_eq (@nil (list T)). intro; unfold length_eq; apply list_eq_nil. Qed. Lemma length_eq_rev_cons : forall {T:Type} (ll:faml T) (la:list T) (pfn:ll <> nil), length_eq (la::ll) -> length la = length (hd' ll pfn) /\ length_eq ll. intros T ll la h1 h2; red in h2. pose proof (list_eq_rev_cons (map (@length T) ll) (length la)) as h3. hfirst h3. rewrite map_eq_nil_iff; auto. specialize (h3 hf). simpl in h2, h3. specialize (h3 h2). rewrite hd_map' in h3. destruct h3 as [h3 h4]. split; auto. rewrite h3. repeat f_equal. apply proof_irrelevance. Qed. Lemma length_eq_cons : forall {T:Type} (ll:faml T) (la:list T), length_eq ll -> (exists l:list T, In l ll /\ length l = length la) -> length_eq (la::ll). intros T ll la h1 h2. destruct h2 as [l [h2 h3]]. apply list_eq_cons. apply h1. exists (length l). generalize dependent ll. intro ll. revert h3. revert l la. induction ll as [|lb ll ih]. intros; contradiction. intros l la h1 h2 h3. simpl. split; auto. destruct h3 as [|h3]; subst. left; auto. right. apply length_eq_rev_cons in h2. specialize (ih l la h1 h2 h3). destruct ih as [ih]; auto. intro; subst; contradiction. Qed. Lemma length_eq_compat : forall {T:Type} (ll:faml T), length_eq ll <-> (forall lx ly:list T, In lx ll -> In ly ll -> length lx = length ly). intros T ll. unfold length_eq. rewrite list_eq_compat. split; intros h1 x y h2 h3. apply h1; (apply in_map; auto). rewrite in_map_iff in h2, h3. destruct h2 as [lx [h4 h5]], h3 as [ly [h6 h7]]; subst. apply h1; auto. Qed. Definition length_eq_ind {T:Type} (ll:faml T) (pfle:length_eq ll) (la lb:list T) (pfinla:In la ll) (pfinlb:In lb ll) (i:nat) (pfi:i < length la) : i < length lb := lt_eq_trans _ _ _ pfi (iff1 (length_eq_compat ll) pfle _ _ pfinla pfinlb). Lemma map_im_compat : forall {T U:Type} (f:T->U) (l:list T), list_to_set (map f l) = Im (list_to_set l) f. intros T U f l. apply Extensionality_Ensembles. red. split. red. intros y h1. rewrite <- list_to_set_in_iff in h1. rewrite in_map_iff in h1. destruct h1 as [x h2]. destruct h2 as [h2l h2r]. apply Im_intro with x. rewrite <- list_to_set_in_iff. assumption. subst. reflexivity. red. intros x h1. destruct h1 as [x h1]. subst. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h1. rewrite in_map_iff. exists x. split; auto. Qed. Lemma not_nil_inhabited : forall {T:Type} (l:list T), l <> nil -> Inhabited (list_to_set l). intros T l h1. destruct l as [|a l]. contradict h1; auto. econstructor. rewrite <- list_to_set_in_iff. left. apply eq_refl. Qed. Lemma inh_not_nil : forall {T:Type} (l:list T), Inhabited (list_to_set l) -> l <> nil. intros ? ? h1; intro; subst. apply not_inh_empty in h1; auto. Qed. Lemma in_inh : forall {T:Type} (l:list T) (a:T), In a l -> Inhabited (list_to_set l). intros T l a h1. apply Inhabited_intro with a. rewrite <- list_to_set_in_iff; auto. Qed. Lemma all_sing_list_to_set : forall {T:Type} (l:list T) (a:T), l <> nil -> all_sing l a -> list_to_set l = Singleton a. intros T l a h0 h1. apply Extensionality_Ensembles; red; split; red; intros x h2. red in h1. rewrite <- list_to_set_in_iff in h2. rewrite h1; auto. constructor; auto. destruct h2. rewrite <- list_to_set_in_iff. apply all_sing_in; auto. Qed. Lemma all_sing_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), all_sing l a <-> card_fun1 (Add (list_to_set l) a) <= 1. intros T h0 l a. split. intro h1 . destruct (nil_dec' l) as [h2 | h2]; subst. simpl. rewrite add_empty_sing. rewrite card_fun1_sing; auto with arith. pose proof (all_sing_in _ _ h2 h1) as h3. rewrite in_add_eq. erewrite all_sing_list_to_set. rewrite card_fun1_sing; auto with arith. assumption. apply h1. rewrite <- list_to_set_in_iff; auto. intro h1. apply le_lt_eq_dec in h1. destruct h1 as [h1 | h1]. destruct (in_dec h0 a l) as [h2 | h2]. rewrite in_add_eq in h1. assert (h3:card_fun1 (list_to_set l) = 0). omega. apply card_fun1_O in h3. destruct h3 as [h3 | h3]. rewrite list_to_set_in_iff in h2. rewrite h3 in h2. contradiction. contradict h3; apply list_to_set_finite. rewrite <- list_to_set_in_iff; auto. rewrite card_add_nin' in h1. apply lt_S_n in h1. omega. apply list_to_set_finite. rewrite <- list_to_set_in_iff; auto. destruct (in_dec h0 a l) as [h2 | h2]. rewrite in_add_eq in h1. apply card_fun1_1 in h1. destruct h1 as [a' h1]. rewrite list_to_set_in_iff in h2. rewrite h1 in h2. destruct h2. red. intros x h2. rewrite list_to_set_in_iff in h2. rewrite h1 in h2. destruct h2; auto. rewrite <- list_to_set_in_iff; auto. rewrite card_add_nin' in h1. apply S_inj in h1. apply card_fun1_O in h1. destruct h1 as [h1 | h1]. apply empty_set_nil in h1; subst. apply all_sing_nil. contradict h1; apply list_to_set_finite. apply list_to_set_finite. rewrite <- list_to_set_in_iff; auto. Qed. Lemma all_sing_dec : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), In a l -> {all_sing l a} + {~all_sing l a}. intros T h1 l a hi. destruct (list_eq_dec h1 l (lrep a (length l))) as [h2 | h2]. left. rewrite h2. apply all_sing_lrep'. right. contradict h2. apply all_sing_lrep in h2; auto. Qed. Lemma remove_eq_nil_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), remove pfdt x l = nil <-> (l = nil \/ all_sing l x). intros T h1 l. induction l as [|a l ih]; simpl. tauto. intro x; split; intro h2. destruct (h1 x a) as [|h3]; subst. rewrite ih in h2. right. destruct h2 as [|h2]; subst. apply impl_all_sing; simpl. rewrite add_empty_sing; auto. apply impl_all_sing. destruct (nil_dec' l) as [|hn]; subst. simpl. apply add_empty_sing. apply all_sing_list_to_set in h2; auto. simpl. rewrite h2. apply add_sing_same. discriminate. destruct (h1 x a) as [|h3]; subst. destruct h2 as [h2 | h2]. discriminate. rewrite all_sing_cons_iff in h2. destruct h2 as [? h2]. clear H. rewrite ih. right; auto. destruct h2 as [|h2]; try discriminate. rewrite all_sing_cons_iff in h2. destruct h2; contradiction. Qed. Lemma incl_nil : forall {T:Type} (l:list T), incl nil l. intros; red; intros; contradiction. Qed. Lemma incl_nil' : forall {T:Type} (l:list T), incl l nil -> l = nil. intros T l h1. destruct l as [|a l]; simpl; auto. red in h1. specialize (h1 a (in_eq _ _)). contradiction. Qed. Lemma incl_subst : forall {T:Type} {l l1 l2:list T}, l1 = l2 -> incl l l1 -> incl l l2. intros; subst; auto. Qed. Lemma incl_sing_iff : forall {T:Type} (l:list T) (a:T), incl l (a::nil) <-> l = nil \/ all_sing l a. intros T l a; split; intro h1. destruct (nil_dec' l) as [|h2]; subst. left; auto. right. apply impl_all_sing. red in h1. apply Extensionality_Ensembles; red; split; red; intros x h3. rewrite <- list_to_set_in_iff in h3. apply h1 in h3. destruct h3; subst. constructor; auto. contradiction. destruct h3. destruct l as [|b l]. contradict h2; auto. specialize (h1 _ (in_eq _ _)). destruct h1; subst. rewrite <- list_to_set_in_iff; left; auto. contradiction. destruct h1 as [|h1]; subst. apply incl_nil. red in h1; red. intros x h2. apply h1 in h2; subst. left; auto. Qed. Lemma incl_sing_no_dup_iff : forall {T:Type} (l:list T) (a:T), NoDup l -> (incl l (a::nil) <-> l = nil \/ l = a::nil). intros T l a h1. rewrite incl_sing_iff. split; intro h2; destruct h2 as [|h2]; subst. left; auto. destruct (nil_dec' l) as [h3 | h3]. left; auto. right. destruct l as [|b l]. contradict h3; auto. red in h2; pose proof (h2 _ (in_eq _ _)); subst; auto. f_equal. destruct l as [|b l]; auto. specialize (h2 b). hfirst h2. right; left; auto. apply h2 in hf. subst. apply no_dup_cons_nin in h1. contradict h1; left; auto. left; auto. right. red. intros x h2. destruct h2; auto. contradiction. Qed. Inductive sub_list {T:Type} (l:list T) : list T->Prop := | sub_list_intro : forall l1 l2:list T, sub_list l (l1++l++l2). Lemma sub_list_intro_eq : forall {T:Type} {l l1 l2:list T} (pf:sub_list l (l1 ++ l ++ l2)), pf = sub_list_intro l l1 l2. intros; apply proof_irrelevance. Qed. Definition sub_list_hd {T:Type} (l l':list T) : Prop := firstn (length l) l' = l. Definition sub_list_tl {T:Type} (l l':list T) : Prop := skipn (length l) l' = l. Lemma sub_list_nil1 : forall {T:Type} (l:list T), sub_list nil l. intros T l. do 2 rewrite <- app_nil_r. rewrite <- app_assoc. constructor. Qed. Lemma sub_list_nil2_iff : forall {T:Type} (l:list T), sub_list l nil <-> l = nil. intros T l; split; intro h1. dependent induction h1. apply app_eq_nil in x. destruct x as [? h1]; subst. apply app_eq_nil in h1. destruct h1; auto. subst. rewrite <- app_nil_r. rewrite <- app_nil_l. constructor. Qed. Lemma length_sub_list : forall {T:Type} (l l':list T), sub_list l l' -> length l <= length l'. intros ? ? ? h1; inversion h1; do 2 rewrite app_length; omega. Qed. Lemma sub_list_refl : forall (T:Type), Reflexive _ (@sub_list T). intros T; red; intro l. rewrite <- app_nil_r. rewrite <- app_nil_l. constructor. Qed. Lemma sub_list_trans : forall (T:Type), Transitive _ (@sub_list T). intro T; red; intros l1 l2 l3 h4 h5. destruct h4 as [l1' l1'']. dependent induction h5. rewrite app_assoc. rewrite app_assoc. rewrite <- app_assoc. rewrite <- (app_assoc l1 l1''). constructor. Qed. Lemma sub_list_antisym : forall {T:Type} (l l':list T), sub_list l l' -> sub_list l' l -> l = l'. intros T l l' h1. destruct h1 as [l0 l1]. intro h2. dependent induction h2. rewrite <- app_assoc in x. rewrite <- (app_assoc l2 l3 l1) in x. rewrite (app_assoc l0 l4) in x. rewrite app_eq_same_iff2 in x. destruct x as [h1 h2]. apply app_eq_nil in h1. destruct h1; subst. apply app_eq_nil in h2. destruct h2; subst. repeat rewrite app_nil_l, app_nil_r; auto. Qed. Lemma sub_list_incl : forall {T:Type} (l l':list T), sub_list l l' -> incl l l'. intros T l l' h1. destruct h1. red. intros a h1. apply in_app_r. apply in_app_l; auto. Qed. Lemma sub_list_in : forall {T:Type} (l l':list T) (a:T), In a l -> sub_list l l' -> In a l'. intros T l l'. revert l. induction l' as [|a' l' ih]; simpl. intros l a h1 h2. apply sub_list_nil2_iff in h2. subst. contradiction. intros l a h1 h2. dependent induction h2. destruct l1 as [|a1 l1]. rewrite app_nil_l in x. destruct l as [|b l]. rewrite app_nil_l in x. destruct l2 as [|a2 l2]. discriminate. inversion x. subst. contradiction. rewrite <- app_comm_cons in x. inversion x. subst. destruct h1 as [|h1]; subst. left; auto. right. apply in_app_l; auto. rewrite <- app_comm_cons in x. inversion x. subst. specialize (ih l a h1 (sub_list_intro _ _ _)). right; auto. Qed. Lemma sub_list_sing_dec : forall {T:Type} (l:list T) (a:T), sub_list l (a::nil) -> {l = nil} + {l = a::nil}. intros T l. destruct l as [|b l]. intros; left; auto. intros a h1. right. inversion h1. rewrite app_eq_iff in H0. destruct H0 as [h2 h3]. destruct l1 as [|c l1]. simpl in h2, h3. inversion h3. apply app_eq_nil in H1. destruct H1; subst. repeat rewrite app_nil_l, app_nil_r; auto. simpl in h2, h3. rewrite firstn_nil in h2. rewrite skipn_nil in h3. inversion h2; subst. discriminate. Qed. Lemma not_sub_list_sing : forall {T:Type} (l:list T) (a:T), l <> nil -> l <> a::nil -> ~sub_list l (a::nil). intros T l a h1 h2 h3. apply sub_list_sing_dec in h3. tauto. Qed. Lemma sub_list_sing_sing : forall {T:Type} (a a':T), sub_list (a::nil) (a'::nil) -> a = a'. intros T a a' h1. inversion h1 as [l1 l2 h3]. rewrite app_eq_iff in h3; simpl in h3. destruct h3 as [h3 h4]. destruct l1 as [|a1 l1]. simpl in h3, h4. inversion h4; auto. simpl in h3, h4. rewrite skipn_nil in h4. discriminate. Qed. Lemma sub_list_sing_cons : forall {T:Type} (l:list T) (a:T), sub_list (a::nil) (a::l). intros T l a; apply (sub_list_intro (a::nil) nil l). Qed. Lemma sub_list_sing_cons_dec : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a a':T), sub_list (a::nil) (a'::l) -> {a = a'} + {In a l}. intros T h1 l a a' h2. destruct (h1 a a') as [|h3]; subst. left; auto. right. inversion h2. destruct l1 as [|a1 l1]. rewrite app_nil_l in H0. inversion H0; subst. contradict h3; auto. rewrite <- app_comm_cons in H0. inversion H0; subst. apply in_app_r. left; auto. Qed. Lemma sub_list_cons : forall {T:Type} (l l':list T) (a:T), sub_list l l' -> sub_list l (a::l'). intros T l l' a h1. destruct h1; subst. rewrite app_comm_cons. constructor. Qed. Lemma sub_list_firstn : forall {T:Type} (l:list T) (n:nat), sub_list (firstn n l) l. intros T l n. rewrite <- (firstn_skipn n). rewrite <- app_nil_l. constructor. Qed. Lemma sub_list_skipn : forall {T:Type} (l:list T) (n:nat), sub_list (skipn n l) l. intros T l n. rewrite <- (firstn_skipn n). rewrite <- app_nil_r. rewrite <- app_assoc. constructor. Qed. Lemma sub_list_firstn_le : forall {T:Type} (l:list T) (m n:nat), m <= n -> sub_list (firstn m l) (firstn n l). intros T l m n h1. apply le_lt_eq_dec in h1. destruct h1 as [h1 | h1]. erewrite firstn_lt_nest. apply sub_list_firstn. assumption. subst. apply sub_list_refl. Qed. Lemma sub_list_hd_compat : forall {T:Type} (l l':list T), sub_list_hd l l' -> sub_list l l'. intros T l l' h1. red in h1. rewrite <- h1. apply sub_list_firstn. Qed. Lemma sub_list_tl_compat : forall {T:Type} (l l':list T), sub_list_tl l l' -> sub_list l l'. intros T l l' h1. red in h1. rewrite <- h1. apply sub_list_skipn. Qed. Lemma sub_list_rev_cons1 : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T), sub_list (a::l) l' -> exists l1 l2:list T, l' = l1 ++ l ++ l2 /\ last l1 a = a. intros T h1 l l' a h2. dependent induction h2. exists (l1 ++ a::nil), l2; simpl. rewrite <- app_assoc, last_app_sing; split; auto. Qed. Lemma sub_list_cons1 : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T), sub_list (a::l) l' -> sub_list l l'. intros T h1 l l' a h2. inversion h2; subst. rewrite app_assoc. rewrite <- app_sing. pose proof (app_assoc l1 (a::nil) l) as h3. simpl in h3. rewrite h3 at 1. rewrite <- app_assoc. constructor. Qed. Lemma sub_list_rev_2cons_dec : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T), sub_list (a::l) (a::l') -> {firstn (length l) l' = l} + {sub_list (a::l) l'}. intros T hdt l l' a h1. destruct (list_eq_dec hdt (a::l) (firstn (S (length l)) (a::l'))) as [h2 | h2]. left. simpl in h2. inversion h2 as [he]. rewrite <- he; auto. right. simpl in h2. assert (h3:l <> firstn (length l) l'). contradict h2; f_equal; auto. revert h3 h2. dependent induction h1. intros h3 h4. rewrite app_eq_iff in x. destruct x as [h5 h6]. simpl in h5, h6. destruct l1 as [|b l1]. simpl in h6. inversion h6; subst. rewrite firstn_app_eq in h3. contradict h3; auto. simpl in h5, h6. inversion h5 as [h7]. rewrite <- (firstn_skipn (length l1)). rewrite <- h6. apply (sub_list_intro (a::l) (firstn (length l1) l') l2). Qed. Lemma sub_list_rev_2cons_neq : forall {T:Type} (l l':list T) (a a':T), a <> a' -> sub_list (a::l) (a'::l') -> sub_list l l'. intros T l l' a a' h2 h3. dependent induction h3. destruct l1 as [|a1 l1]. rewrite app_nil_l in x. rewrite <- app_comm_cons in x. inversion x; contradiction. rewrite <- app_comm_cons in x. inversion x; subst. rewrite app_comm_cons. rewrite <- (app_sing l a). rewrite app_assoc. pose proof (app_assoc l1 (a::nil) l) as h3. simpl. simpl in h3. rewrite h3. rewrite <- app_assoc. constructor. Qed. Lemma neq_sub_list_cons : forall {T:Type} (l l':list T) (a':T), l <> firstn (length l) (a'::l') -> sub_list l (a'::l') -> sub_list l l'. intros T l l' a' h1 h2. revert h1. dependent induction h2. intro h3. destruct l1 as [|a l1]. rewrite app_nil_l in x. contradict h3. rewrite <- x. rewrite firstn_app_eq; auto. rewrite <- app_comm_cons in x. inversion x; subst. constructor. Qed. Lemma sub_list_rev_cons_dec : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a':T), sub_list l (a'::l') -> {sub_list l l'} + {exists pf: l <> nil, hd' l pf = a' /\ tl l = firstn (pred (length l)) l'}. intros T h1 l l' a h2. destruct l as [|b l]; simpl. left; apply sub_list_nil1. destruct (h1 a b) as [|h3]; subst. apply sub_list_rev_2cons_dec in h2. destruct h2 as [h2 | h2]. right. exists (cons_neq_nil _ _ ). split; auto. left; auto. assumption. apply neq_sym in h3. left. inversion h2 as [l1 l2 h4]. destruct l1 as [|c l1]. rewrite app_nil_l in h4. inversion h4. contradiction. rewrite <- app_comm_cons in h4. inversion h4; subst. rewrite app_comm_cons. constructor. Qed. Lemma sub_list_dec : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), {sub_list l l'} + {~sub_list l l'}. intros T hdt l l'. revert l. induction l' as [|a l' ih]. intro l. destruct (nil_dec' l) as [|h1]; subst. left; apply sub_list_nil1. right. intro h2. rewrite sub_list_nil2_iff in h2. contradiction. intro l. specialize (ih l). destruct ih as [h1 | h1]. left. apply sub_list_cons; auto. destruct (list_eq_dec hdt l (firstn (length l) (a::l'))) as [h2 | h2]. left. rewrite h2. apply sub_list_firstn. right. intro h3. inversion h3 as [l1 l2 h4]. destruct l1 as [|a1 l1]. rewrite app_nil_l in h4. destruct l2 as [|a2 l2]. rewrite app_nil_r in h4. subst. rewrite firstn_length in h2. contradict h2; auto. rewrite <- h4 in h2. rewrite firstn_app_eq in h2. contradict h2; auto. rewrite <- app_comm_cons in h4. inversion h4; subst. contradict h1. constructor. Qed. Lemma sub_list_hd_dec : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), {sub_list_hd l l'} + {~sub_list_hd l l'}. intros T h1 l l'. unfold sub_list_hd. apply list_eq_dec. apply h1. Qed. Lemma sub_list_tl_dec : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), {sub_list_tl l l'} + {~sub_list_tl l l'}. intros T h1 l l'. unfold sub_list_tl. apply list_eq_dec. apply h1. Qed. Lemma le_count_occ_sub_list : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T) (pfsl:sub_list l l'), count_occ pfdt l x <= count_occ pfdt l' x. intros T h1 l l'. revert l. induction l' as [|a' l' ih]. intros l x h2. rewrite sub_list_nil2_iff in h2; subst; simpl; auto with arith. intros l x h2. destruct (in_dec h1 x l) as [hin | hin]. pose proof (sub_list_in _ _ _ hin h2) as hi. pose proof h2 as h2'. apply sub_list_rev_cons_dec in h2; auto. destruct h2 as [h2 | h2]. pose proof (ih _ x h2) as ih'. simpl. destruct (h1 a' x) as [|h3]; subst; auto. destruct h2 as [h2 h3]. simpl. destruct h3 as [? h4]; subst. destruct (h1 (hd' l h2) x) as [|h5]; subst. rewrite (hd_compat' _ h2) at 1. simpl. destruct (h1 (hd' l h2) (hd' l h2)) as [|h5]. apply le_n_S. apply ih. rewrite h4. apply sub_list_firstn. contradict h5; auto. destruct hi as [|hi]; subst. contradict h5; auto. apply sub_list_rev_cons_dec in h2'. destruct h2' as [h6 | h6]. apply ih; auto. destruct h6 as [h6 [h7 h8]]. specialize (ih (tl l) x). hfirst ih. rewrite h8. apply sub_list_firstn. specialize (ih hf). erewrite count_occ_tl in ih; auto. apply neq_sym. apply h5. assumption. rewrite (count_occ_In h1) in hin; omega. Qed. Lemma le_count_occ_firstn : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), le (count_occ pfdt (firstn n l) x) (count_occ pfdt l x). intros; apply le_count_occ_sub_list; apply sub_list_firstn. Qed. Lemma le_count_occ_skipn : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), le (count_occ pfdt (skipn n l) x) (count_occ pfdt l x). intros; apply le_count_occ_sub_list; apply sub_list_skipn. Qed. Lemma count_occ_firstn_eq_whole_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), count_occ pfdt (firstn n l) x = count_occ pfdt l x <-> ~In x (skipn n l). intros T h1 l x n. revert l x. induction n as [|n ih]; simpl. intros; split; intro h2; [rewrite (count_occ_In h1); omega | rewrite (count_occ_In h1) in h2; omega]. intros l x. destruct l as [|a l]. intros; split; auto. simpl. destruct (h1 a x) as [|h2]; subst. split; intro h2. apply S_inj in h2. rewrite ih in h2; auto. f_equal. rewrite ih; auto. apply ih. Qed. Lemma count_occ_skipn_eq_whole_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), count_occ pfdt (skipn n l) x = count_occ pfdt l x <-> ~In x (firstn n l). intros T h1 l x n. revert l x. induction n as [|n ih]; simpl. tauto. intros l x; split; intro h0. destruct l as [|a l]. tauto. simpl; simpl in h0. destruct (h1 a x) as [|h2]; subst. pose proof (le_count_occ_skipn h1 l x n); omega. intro h3. destruct h3 as [|h3]. contradiction. rewrite ih in h0. contradiction. destruct l as [|a l]; auto. simpl. destruct (h1 a x) as [|h2]; subst. contradict h0; left; auto. rewrite ih. contradict h0. right; auto. Qed. Section NatAnalogues. (*In this section will be non-polymorphic [nat]-boolean-analogues to many of the [list] functions and propositions involving [type_eq_dec], so we can capitalize on [nat]'s "easy" computational abilities.*) Fixpoint In_nat (l:list nat) (m:nat) : bool := match l with | nil => false | a :: l' => if (beq_nat a m) then true else In_nat l' m end. Fixpoint remove_nat (l:list nat) (m:nat) : list nat := match l with | nil => nil | a :: l' => let ln := remove_nat l' m in let be := beq_nat a m in if (In_nat l' m) then if be then ln else a :: ln else if be then l' else a :: l' end. End NatAnalogues. Definition lS (l:list nat) := map S l. Definition lpred (l:list nat) := map pred l. Lemma lpred_lS : forall l:list nat, lpred (lS l) = l. unfold lpred, lS; intro. rewrite map_map. rewrite <- map_id. f_equal. Qed. Lemma length_lS : forall (l:list nat), length (lS l) = length l. unfold lS; intros; apply map_length. Qed. Lemma lS_eq_nil_iff : forall (l:list nat), lS l = nil <-> l = nil. intro l. destruct l as [|a l]; simpl. tauto. split; intro h1; discriminate. Qed. Lemma count_occ_lS : forall (l:list nat) (m:nat), m > 0 -> count_occ nat_eq_dec (lS l) m = count_occ nat_eq_dec l (pred m). intros l m. induction l as [|a l ih]; simpl; auto. destruct (nat_eq_dec (S a) m) as [h1 | h1]. subst. destruct (nat_eq_dec a (pred (S a))) as [h2 | h2]. f_equal; auto. assert (h3:a = 0). omega. subst. simpl in h2. contradict h2; auto. destruct (nat_eq_dec a (pred m)) as [|h2]; subst; auto. assert (h2:m = 0). omega. subst. simpl in ih; simpl. intro; omega. Qed. Lemma firstn_lS : forall (l:list nat) (n:nat), firstn n (lS l) = lS (firstn n l). intros l n. revert l. induction n as [|n ih]; simpl; auto. intro l. destruct l as [|a l]; simpl; auto. f_equal; auto. Qed. Lemma skipn_lS : forall (l:list nat) (n:nat), skipn n (lS l) = lS (skipn n l). intros l n. revert l. induction n as [|n ih]; simpl; auto. intro l. destruct l as [|a l]; simpl; auto. Qed. Lemma firstn_lpred : forall (l:list nat) (n:nat), firstn n (lpred l) = lpred (firstn n l). intros l n. revert l. induction n as [|n ih]; simpl; auto. intro l. destruct l as [|a l]; simpl; auto. f_equal; auto. Qed. Lemma skipn_lpred : forall (l:list nat) (n:nat), skipn n (lpred l) = lpred (skipn n l). intros l n. revert l. induction n as [|n ih]; simpl; auto. intro l. destruct l as [|a l]; simpl; auto. Qed. Lemma length_lpred : forall (l:list nat), length (lpred l) = length l. unfold lpred; intros; apply map_length. Qed. Lemma neq_nil_lpred : forall (l:list nat), l <> nil -> lpred l <> nil. intros l h1. apply not_nil_cons in h1. destruct h1 as [a h2]; subst. destruct h2 as [l' h2]; subst. simpl. apply cons_neq_nil. Qed. Lemma lpred_neq_nil : forall (l:list nat), lpred l <> nil -> l <> nil. intros l h1 h2; subst. simpl in h1. contradict h1; auto. Qed. Lemma hd_lpred' : forall (l:list nat) (pf:lpred l <> nil), hd' (lpred l) pf = pred (hd' l (lpred_neq_nil _ pf)). intros l. destruct l as [|a l]; simpl. intro h1; contradict h1; auto. intros; auto. Qed. Lemma tl_lpred : forall (l:list nat), tl (lpred l) = lpred (tl l). intro l. destruct l; auto. Qed. Lemma in_lS_iff : forall (l:list nat) (m:nat), In m (lS l) <-> exists m':nat, In m' l /\ m = S m'. intros l m. unfold lS. rewrite in_map_iff. split; intro h1; destruct h1 as [n hm]; destruct hm as [? h2]; subst; exists n; tauto. Qed. Lemma nin0_lS : forall (l:list nat), ~In 0 (lS l). intros l h1. rewrite in_lS_iff in h1. destruct h1 as [? [? ?]]. discriminate. Qed. Lemma in_lpred_iff : forall (l:list nat) (m:nat), In m (lpred l) <-> exists m':nat, In m' l /\ m = pred m'. intros l m. unfold lpred. rewrite in_map_iff. split; intro h1; destruct h1 as [n hm]; destruct hm as [? h2]; subst; exists n; tauto. Qed. Lemma lS_lpred_iff : forall (l:list nat), ~ In 0 l <-> l = lS (lpred l). intro l; split; intro h1. induction l as [|a l ih]; simpl; auto. destruct (zerop a) as [|h2]; subst. contradict h1; left; auto. rewrite <- (S_pred _ _ h2). f_equal; apply ih. contradict h1; right; auto. intro h2. rewrite h1 in h2. rewrite in_lS_iff in h2. destruct h2 as [? [h4 h5]]. discriminate. Qed. Lemma no_dup_lS : forall (l:list nat), NoDup l -> NoDup (lS l). intros l h1. unfold lS. apply no_dup_map_inj; auto. red; apply S_inj. Qed. Lemma no_dup_lS_impl : forall (l:list nat), NoDup (lS l) -> NoDup l. intro l. induction l as [|a l ih]; simpl; auto. intro h1. pose proof (no_dup_cons _ _ h1) as h2. specialize (ih h2). pose proof (no_dup_cons_nin _ _ h1) as h3. constructor; auto. contradict h3. apply in_map; auto. Qed. Lemma no_dup_lpred : forall (l:list nat), (In 0 l -> ~In 1 l) -> (In 1 l -> ~In 0 l) -> NoDup l -> NoDup (lpred l). intro l. induction l as [|c l ih]; simpl. intros; constructor. intros h0 h1 h2. hfirst ih. intuition. pose proof (no_dup_cons _ _ h2) as h4. pose proof (no_dup_cons_nin _ _ h2) as h5. constructor. contradict h5. unfold lpred in h5. rewrite in_map_iff in h5. destruct h5 as [c' h5]. destruct h5 as [h5 h7]. destruct (lt_eq_lt_dec c 1) as [h8 | h8]. destruct h8 as [h8 | h8]. assert (h9:c = 0). omega. subst. simpl in h5. apply pred0_dec in h5. destruct h5; subst; auto. intuition. subst. simpl in h5. apply pred0_dec in h5. destruct h5; subst. intuition. assumption. assert (h9:c' > 0). omega. apply pred_inj0 in h5; auto with arith. subst; auto. apply ih; auto. intuition. Qed. Lemma no_dup_lpred_impl : forall (l:list nat), NoDup (lpred l) -> NoDup l. intro l. induction l as [|a l ih]; simpl; auto. intro h1. pose proof (no_dup_cons _ _ h1) as h2. specialize (ih h2). pose proof (no_dup_cons_nin _ _ h1) as h3. constructor; auto. contradict h3. apply in_map; auto. Qed. (*Bit Vector*) Definition bv : Type := list bool. Definition trues (n:nat) : bv := lrep true n. Definition falses (n:nat) : bv := lrep false n. (*The indexes of [bv] which are [false]*) Fixpoint bv0s_aux (v:bv) (acc:nat) : list nat:= match v with | nil => nil | a::l => let n := bv0s_aux l (S acc) in match a with | false => acc::n | true => n end end. Definition bv0s (v:bv) := bv0s_aux v 0. Lemma length_bv0s_aux_le_S : forall (v:bv) (n:nat), length (bv0s_aux v n) <= length (bv0s_aux v (S n)). intro v. induction v as [|b v ih]; simpl; auto with arith. intro n. destruct b; simpl; simpl in ih; auto. apply le_n_S; auto. Qed. Lemma length_bv0s_le : forall (v:bv) (n:nat), length (bv0s v) <= length (bv0s_aux v n). intros v n. revert v. induction n as [|n ih]; simpl; auto. intro v. eapply le_trans. apply ih. apply length_bv0s_aux_le_S. Qed. Lemma length_bv0s_aux_le : forall (v:bv) (m n:nat), m <= n -> length (bv0s_aux v m) <= length (bv0s_aux v n). intro v. induction v as [|b v ih]; simpl; auto. intros m n h1. destruct b; simpl. apply ih; auto with arith. apply le_n_S; auto. apply ih; auto with arith. Qed. Lemma bv0s_aux_S : forall (v:bv) (m:nat), bv0s_aux v (S m) = lS (bv0s_aux v m). intro v. induction v as [|b v ih]; simpl; auto with arith. intro m. destruct b. specialize (ih (S m)). simpl in ih; auto. simpl. f_equal. apply ih. Qed. Lemma length_bv0s_le' : forall (v:bv), length (bv0s v) <= length v. intro v. unfold bv0s. induction v as [|b v ih]; auto with arith. simpl. destruct b; simpl; rewrite bv0s_aux_S; rewrite length_lS; auto with arith. Qed. Lemma in_bv0s_lt : forall (v:bv) (i:nat), In i (bv0s v) -> i < length v. intro v. induction v as [|b v ih]; simpl. intros; contradiction. intros i h1. unfold bv0s in h1; simpl in h1. destruct b; auto with arith. rewrite bv0s_aux_S in h1. rewrite in_lS_iff in h1. destruct h1 as [m [h2 h3]]; subst; auto with arith. destruct h1 as [|h1]; subst; auto with arith. rewrite bv0s_aux_S in h1. rewrite in_lS_iff in h1. destruct h1 as [m [h2 h3]]; subst; auto with arith. Qed. Lemma no_dup_bv0s : forall (v:bv), NoDup (bv0s v). intro v. unfold bv0s. induction v as [|b v ih]; simpl. constructor. destruct b. rewrite bv0s_aux_S. apply no_dup_lS; auto. rewrite bv0s_aux_S. constructor. intro h2. rewrite in_lS_iff in h2. destruct h2 as [? [? ?]]; discriminate. apply no_dup_lS; auto. Qed. (*The indexes of [bv] which are [true]*) Fixpoint bv1s_aux (v:bv) (acc:nat) : list nat:= match v with | nil => nil | a::l => let n := bv1s_aux l (S acc) in match a with | false => n | true => acc::n end end. Definition bv1s (v:bv) := bv1s_aux v 0. Lemma length_bv1s_aux_le_S : forall (v:bv) (n:nat), length (bv1s_aux v n) <= length (bv1s_aux v (S n)). intro v. induction v as [|b v ih]; simpl; auto with arith. intro n. destruct b; simpl; simpl in ih; auto. apply le_n_S; auto. Qed. Lemma length_bv1s_le : forall (v:bv) (n:nat), length (bv1s v) <= length (bv1s_aux v n). intros v n. revert v. induction n as [|n ih]; simpl; auto. intro v. eapply le_trans. apply ih. apply length_bv1s_aux_le_S. Qed. Lemma length_bv1s_aux_le : forall (v:bv) (m n:nat), m <= n -> length (bv1s_aux v m) <= length (bv1s_aux v n). intro v. induction v as [|b v ih]; simpl; auto. intros m n h1. destruct b; simpl. apply le_n_S; auto. apply ih; auto with arith. apply ih; auto with arith. Qed. Lemma bv1s_aux_S : forall (v:bv) (m:nat), bv1s_aux v (S m) = lS (bv1s_aux v m). intro v. induction v as [|b v ih]; simpl; auto with arith. intro m. destruct b. specialize (ih (S m)). simpl in ih; auto. simpl. f_equal; auto. apply ih. Qed. Lemma length_bv1s_le' : forall (v:bv), length (bv1s v) <= length v. intro v. unfold bv1s. induction v as [|b v ih]; auto with arith. simpl. destruct b; simpl; rewrite bv1s_aux_S; rewrite length_lS; auto with arith. Qed. Lemma in_bv1s_lt : forall (v:bv) (i:nat), In i (bv1s v) -> i < length v. intro v. induction v as [|b v ih]; simpl. intros; contradiction. intros i h1. unfold bv1s in h1; simpl in h1. destruct b; auto with arith. rewrite bv1s_aux_S in h1. destruct h1 as [|h1]; subst; auto with arith. rewrite in_lS_iff in h1. destruct h1 as [m [h2 h3]]; subst; auto with arith. rewrite bv1s_aux_S in h1. rewrite in_lS_iff in h1. destruct h1 as [m [h2 h3]]; subst; auto with arith. Qed. Lemma no_dup_bv1s : forall (v:bv), NoDup (bv1s v). intro v. unfold bv1s. induction v as [|b v ih]; simpl. constructor. rewrite bv1s_aux_S. destruct b. constructor. intro h2. rewrite in_lS_iff in h2. destruct h2 as [? [? ?]]; discriminate. apply no_dup_lS; auto. apply no_dup_lS; auto. Qed. (*Typically, [l <> nil]*) (*The bit vector of the [tl] in its given list*) Definition bv_tl {T:Type} (l:list T) : bv := match l with | nil => nil | _::l' => false :: lrep true (length l') end. Lemma length_bv_tl : forall {T:Type} (l:list T), length (bv_tl l) = length l. intros T l. destruct l as [|a l]; simpl; auto. f_equal; rewrite length_lrep; auto. Qed. Lemma bv_count_occ : forall(v:bv), length v = count_occ bool_eq_dec v true + count_occ bool_eq_dec v false. intro v. induction v as [|a v]; simpl; auto. lr_if_goal; fold hlr; destruct hlr as [hl | hr]; subst; lr_if_goal; fold hlr; destruct hlr as [hl' | hr']; subst; try discriminate. rewrite plus_S_shift_l. f_equal; auto. rewrite plus_S_shift_r. f_equal; auto. destruct a. contradict hr; auto. contradict hr'; auto. Qed. Lemma bv0s_cons_true : forall v, bv0s (true::v) = lS (bv0s v). intro v. unfold bv0s. induction v as [|b v ih]; simpl; auto. destruct b. simpl in ih. rewrite bv0s_aux_S in ih. rewrite bv0s_aux_S; auto. simpl. f_equal. rewrite bv0s_aux_S; auto. Qed. Lemma bv0s_cons_false : forall v, bv0s (false::v) = 0 :: lS (bv0s v). intro v. unfold bv0s. induction v as [|b v ih]; simpl; auto. destruct b. simpl in ih. rewrite bv0s_aux_S in ih. rewrite bv0s_aux_S; auto. simpl. f_equal. rewrite bv0s_aux_S; auto. Qed. Lemma bv1s_cons_true : forall v, bv1s (true::v) = 0:: lS (bv1s v). intro v. unfold bv1s. induction v as [|b v ih]; simpl; auto. destruct b. simpl in ih. rewrite bv1s_aux_S in ih. rewrite bv1s_aux_S; auto. simpl. f_equal. rewrite bv1s_aux_S; auto. Qed. Lemma bv1s_cons_false : forall v, bv1s (false::v) = lS (bv1s v). intro v. unfold bv1s. induction v as [|b v ih]; simpl; auto. destruct b. simpl in ih. rewrite bv1s_aux_S in ih. rewrite bv1s_aux_S; auto. simpl. f_equal. rewrite bv1s_aux_S; auto. Qed. Lemma all_sing_true_iff : forall (v:bv), all_sing v true <-> ~In false v. intro v; induction v as [|v]; simpl; intuition; try (red; intros; contradiction); subst; try (red in H1; specialize (H1 _ (in_eq _ _)); discriminate); try (red; intros; destruct H; subst); destruct x; auto. Qed. Lemma all_sing_false_iff : forall (v:bv), all_sing v false <-> ~In true v. intro v; induction v as [|v]; simpl; intuition; try (red; intros; contradiction); subst; try (red in H1; specialize (H1 _ (in_eq _ _)); discriminate); try (red; intros; destruct H; subst); destruct x; try specialize (H2 (eq_refl _)); try contradiction; auto. Qed. Lemma nin_true_bv1s : forall (v:bv), ~In true v -> bv1s v = nil. intro v. induction v as [|b v ih]; simpl; auto. intro h1. destruct b. contradict h1; left; auto. rewrite bv1s_cons_false. rewrite ih; auto. Qed. Lemma nin_false_bv0s : forall (v:bv), ~In false v -> bv0s v = nil. intro v. induction v as [|b v ih]; simpl; auto. intro h1. destruct b. rewrite bv0s_cons_true, ih; auto. contradict h1; left; auto. Qed. Lemma bv0s_eq_nil_iff : forall (v:bv), bv0s v = nil <-> ~In false v. intro v; induction v as [|b v ih]; simpl. intuition. destruct b. split; intro h1. rewrite bv0s_cons_true in h1. rewrite lS_eq_nil_iff in h1. rewrite ih in h1. contradict h1. destruct h1; try discriminate; auto. hr ih. contradict h1. right; auto. rewrite bv0s_cons_true. rewrite lS_eq_nil_iff. tauto. split; intro h1. rewrite bv0s_cons_false in h1. discriminate. contradict h1; left; auto. Qed. Lemma bv1s_eq_nil_iff : forall (v:bv), bv1s v = nil <-> ~In true v. unfold bv1s; intro v; induction v as [|b v ih]; simpl. intuition. destruct b. split; intro h1; [discriminate|contradict h1; left; auto]. rewrite bv1s_aux_S; split; intro h1. rewrite lS_eq_nil_iff in h1. rewrite ih in h1. contradict h1. destruct h1; try discriminate; auto. rewrite lS_eq_nil_iff. rewrite ih. contradict h1. right; auto. Qed. Fixpoint bv_lt_eq_len (v:bv) : forall w, length v = length w -> Prop := match v with | nil => fun _ _ => False | b::v' => fun w => match w with | nil => fun pfe:length (b::v') = length nil => False_rect _ (disc_nat (length v') pfe) | c::w' => fun pfe => match (tso_dec_bool_lt b c) with | inleft (left _) => True | inleft (right _) => bv_lt_eq_len v' w' (S_inj _ _ pfe) | inright _ => False end end end. Definition bv_lt (v w:bv) : Prop := match (lt_eq_lt_dec (length v) (length w)) with | inleft (left _) => True | inleft (right pfe) => bv_lt_eq_len v w pfe | inright _ => False end. Lemma tso_dec_bv_lt_eq_len : forall (v w:bv) (pfe:length v = length w), {bv_lt_eq_len v w pfe} + {v = w} + {bv_lt_eq_len w v (eq_sym pfe)}. intro v. induction v as [|b v ih]; simpl. intros w h1. pose proof (length_eq0 _ (eq_sym h1)); subst. left; right; auto. intro w. destruct w as [|c w]. intro; discriminate. simpl; intro h1. specialize (ih _ (S_inj _ _ h1)). destruct ih as [h2|h2]. destruct h2 as [h2|h2]. lr_match_sum_goal'. destruct hl; subst. left; auto. left; left. gterm0. fold c0. term0 h2. pose proof (proof_irrelevance _ c0 c1) as h3. rewrite h3; auto. lr_match_sum_goal'. destruct hl as [h3|]; subst. right; auto. apply irrefl_bool_lt in hr; contradiction. pose proof (transitive_bool_lt _ _ _ hr hr0) as h3. apply irrefl_bool_lt in h3; contradiction. subst. lr_match_sum_goal'. destruct hl as [h2|]; subst. left; auto. lr_match_sum_goal'. destruct hl. apply irrefl_bool_lt in b; contradiction. left; right; auto. apply irrefl_bool_lt in hr; contradiction. right. lr_match_sum_goal'. destruct hl as [h2|]; subst; auto. apply irrefl_bool_lt in hr; contradiction. pose proof (transitive_bool_lt _ _ _ hr hr0) as h3. apply irrefl_bool_lt in h3; contradiction. lr_match_sum_goal'. destruct hl as [h3|]; subst. left; auto. lr_match_sum_goal'. destruct hl as [h3|]; subst. apply irrefl_bool_lt in h3; contradiction. right. gterm0. fold c0. term0 h2. pose proof (proof_irrelevance _ c0 c1) as h3. rewrite h3; auto. apply irrefl_bool_lt in hr; contradiction. right. lr_match_sum_goal'. destruct hl as [h3|]; subst; auto. apply irrefl_bool_lt in hr; contradiction. pose proof (transitive_bool_lt _ _ _ hr hr0) as h3. apply irrefl_bool_lt in h3; contradiction. Qed. Lemma total_strict_order_bv_lt_eq_len : forall (v w:bv) (pfe:length v = length w), bv_lt_eq_len v w pfe \/ v = w \/ bv_lt_eq_len w v (eq_sym pfe). intro v. induction v as [|b v ih]; simpl. intros w h1. pose proof (length_eq0 _ (eq_sym h1)); subst. right; left; auto. intro w. destruct w as [|c w]. intro; discriminate. simpl; intro h1. lr_match_sum_goal'. destruct hl. left; auto. subst. specialize (ih _ (S_inj _ _ h1)). destruct ih as [h2|h2]. left. gterm0. fold c0. term0 h2. pose proof (proof_irrelevance _ c0 c1) as h3. rewrite h3; auto. destruct h2 as [|h2]; subst. right; left; auto. lr_match_sum_goal'. destruct hl. apply irrefl_bool_lt in b; contradiction. right; right. gterm0. fold c0. term0 h2. pose proof (proof_irrelevance _ c0 c1) as h3. rewrite h3; auto. apply irrefl_bool_lt in hr; contradiction. right. lr_match_sum_goal'. destruct hl; subst. right; auto. apply irrefl_bool_lt in hr; contradiction. pose proof (transitive_bool_lt _ _ _ hr hr0) as h2. apply irrefl_bool_lt in h2; contradiction. Qed. Lemma irrefl_bv_lt_eq_len : forall (v:bv) (pfe:length v = length v), ~bv_lt_eq_len v v pfe. intro v. induction v as [|b v ih]; simpl; auto. intro. lr_match_sum_goal'. destruct hl as [hl|hl]; auto. apply irrefl_bool_lt in hl; contradiction. apply irrefl_bool_lt in hr; contradiction. Qed. Lemma transitive_bv_lt_eq_len : forall (v w x:bv) (pfe: length v = length w) (pfe':length w = length x), bv_lt_eq_len v w pfe -> bv_lt_eq_len w x pfe' -> bv_lt_eq_len v x (eq_trans pfe pfe'). intro v. induction v as [|b v ih]; simpl. intros; contradiction. intro w; destruct w as [|c w]. intro; discriminate. intro x; destruct x as [|d x]; simpl. intro; discriminate. lr_match_sum_goal'. destruct hl as [hl|hl]. lr_match_sum_goal'. destruct hl0 as [hl0|hl0]. lr_match_sum_goal'. destruct hl1 as [hl1|hl1]; auto. subst. pose proof (transitive_bool_lt _ _ _ hl hl0) as h1. apply irrefl_bool_lt in h1; contradiction. pose proof (transitive_bool_lt _ _ _ hl hl0) as h1. pose proof (transitive_bool_lt _ _ _ hr h1) as h2. apply irrefl_bool_lt in h2; contradiction. subst. lr_match_sum_goal'. destruct hl0 as [hl0|]; subst; auto. apply irrefl_bool_lt in hl; contradiction. pose proof (transitive_bool_lt _ _ _ hl hr) as h1. apply irrefl_bool_lt in h1; contradiction. intros; contradiction. lr_match_sum_goal'. destruct hl0 as [hl0|]; subst; auto. lr_match_sum_goal'. destruct hl as [hl|]; subst; auto. apply irrefl_bool_lt in hl0; contradiction. pose proof (transitive_bool_lt _ _ _ hl0 hr) as h1. apply irrefl_bool_lt in h1; contradiction. lr_match_sum_goal'. destruct hl as [hl|hl]. apply irrefl_bool_lt in hl; contradiction. intros h1 h2 h3 h4. specialize (ih _ _ (S_inj _ _ h1) (S_inj _ _ h2) h3 h4). gterm0. fold c. term0 ih. fold c0 in ih. pose proof (proof_irrelevance _ c c0) as h5. rewrite h5. subst; auto. apply irrefl_bool_lt in hr; contradiction. intros; contradiction. intros; contradiction. Qed. Lemma tso_dec_bv_lt : tso_dec bv_lt. unfold bv_lt; red; intros v w. lr_match_sum_goal'. destruct hl as [hl|hl]. left; auto. lr_match_sum_goal'. destruct hl0 as [h0|h0]. omega. pose proof (proof_irrelevance _ h0 (eq_sym hl)); subst. apply tso_dec_bv_lt_eq_len. omega. lr_match_sum_goal'; try omega. destruct hl; try omega; auto. Qed. Lemma total_strict_order_bv_lt : total_strict_order bv_lt. unfold bv_lt; red; intros v w. lr_match_sum_goal'. destruct hl as [h2|h2]. left; auto. lr_match_sum_goal'. destruct hl as [h3|h3]. omega. pose proof (proof_irrelevance _ h3 (eq_sym h2)); subst. apply total_strict_order_bv_lt_eq_len. omega. lr_match_sum_goal'; try omega. destruct hl as [h2|h2]; try omega; auto. Qed. Lemma irrefl_bv_lt : irrefl bv_lt. unfold bv_lt; red; intro; lr_match_sum_goal'. destruct hl as [hl|hl]. omega. apply irrefl_bv_lt_eq_len. omega. Qed. Lemma transitive_bv_lt : Transitive _ bv_lt. unfold bv_lt; red. intros v w x. lr_match_sum_goal'. destruct hl as [hl|hl]. lr_match_sum_goal'. destruct hl0 as [hl0|hl0]. lr_match_sum_goal'. destruct hl1 as [hl1|hl1]; auto. omega. omega. lr_match_sum_goal'. destruct hl1 as [hl1|hl1]; auto. omega. omega. intros; contradiction. lr_match_sum_goal'. destruct hl0 as [hl0|hl0]; auto. lr_match_sum_goal'. destruct hl1 as [hl1|hl1]; auto. omega. omega. lr_match_sum_goal'. destruct hl1 as [hl1|hl1]; auto. pose proof (proof_irrelevance _ hl1 (eq_trans hl hl0)); subst. apply transitive_bv_lt_eq_len. omega. intros; contradiction. intro; contradiction. Qed. Lemma sortdec_bv_lt : sort_dec bv_lt. constructor. apply tso_dec_bv_lt. apply transitive_bv_lt. apply irrefl_bv_lt. Qed. Lemma sortable_bv_lt : sortable bv_lt. constructor. apply total_strict_order_bv_lt. apply transitive_bv_lt. apply irrefl_bv_lt. Qed. (*Removes exactly one element from a list corresponding to a given index.*) Fixpoint remove_ind {T} (l:list T) (i:nat) : list T := match l with | nil => nil | a :: l' => match i with | 0 => l' | S i' => a :: remove_ind l' i' end end. Lemma in_remove_ind : forall {T} (l:list T) (x:T) (i:nat), In x (remove_ind l i) -> In x l. intros T l; induction l as [|? ? ih]; auto; intros; simpl; simpl in H; destruct i as [|i]; [right|destruct H as [|?]; subst; [left|apply ih in H; right]]; auto. Qed. Lemma remove_ind_ge : forall {T} (l:list T) (i:nat), i >= length l -> remove_ind l i = l. intros T l; induction l as [|a l ih]; simpl; auto; intros i; destruct i as [|i]; try omega; intro h1; apply ge_S_n in h1; f_equal; auto. Qed. Lemma incl_iff : forall {T:Type} (l l':list T), incl l l' <-> Included (list_to_set l) (list_to_set l'). intros T l l'. split. intro h1. red; intros x h2. rewrite <- list_to_set_in_iff in h2. apply h1 in h2. rewrite <- list_to_set_in_iff. assumption. intro h1. red; intros x h2. rewrite list_to_set_in_iff in h2. apply h1 in h2. rewrite list_to_set_in_iff; auto. Qed. Lemma list_to_set_incl : forall {T:Type} (l l':list T), list_to_set l = list_to_set l' -> incl l l'. intros; rewrite incl_iff, H. apply inclusion_reflexive. Qed. Lemma refl_incl : forall (T:Type), Reflexive _ (@incl T). intro; red; red; intros; auto. Qed. Lemma trans_incl : forall (T:Type), Transitive _ (@incl T). intro T. red. intros lx ly lz h1 h2. red in h1, h2; red. intros x h3. apply h2; apply h1; auto. Qed. Lemma incl_cons : forall {T:Type} (l:list T) (a:T), incl l (a::l). intros; red; intros; right; auto. Qed. Lemma incl_in_cons : forall {T} (l l':list T) (a:T), incl l l' -> In a l' -> incl (a::l) l'. intros ? ? ? ? h1 h2; red; intros x h3; destruct h3 as [|h3]; subst; auto. Qed. Lemma incl_from_cons : forall {T} (l l':list T) (a:T), incl l (a::l') -> In a l' -> incl l l'. intros ? ? ? ? h1 ?; red; intros ? h3. apply h1 in h3. destruct h3; subst; auto. Qed. Lemma incl_cons_nin : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T), ~In a l -> incl l (a::l') -> incl (remove pfdt a l) l'. intros T h1 l l' a h2 h3. red in h3; red. intros x h4. rewrite in_remove_iff in h4. destruct h4 as [h4 h5]. apply h3 in h4. destruct h4; subst; auto; try contradict h5; auto. Qed. Lemma incl_cons2_same : forall {T:Type} (l l':list T) (a:T), ~In a l -> ~In a l' -> incl (a::l) (a::l') -> incl l l'. intros T l l' a h1 h2 h3. red in h3; red. intros x h4. specialize (h3 x (in_cons _ _ _ h4)). destruct h3 as [|h3]; subst. contradiction. assumption. Qed. Lemma incl_app_l : forall {T:Type} (l l':list T) , incl l (l ++ l'). intros; red; intros; apply in_app_l; auto. Qed. Lemma incl_app_r : forall {T:Type} (l l':list T) , incl l' (l ++ l'). intros; red; intros; apply in_app_r; auto. Qed. Lemma anti_sym_incl_list_to_set : forall {T:Type} (l l':list T), incl l l' -> incl l' l -> list_to_set l = list_to_set l'. intros T l l' h1 h2. apply Extensionality_Ensembles; red; split; red; intros x h3. rewrite <- list_to_set_in_iff in h3. rewrite <- list_to_set_in_iff. apply h1; auto. rewrite <- list_to_set_in_iff in h3. rewrite <- list_to_set_in_iff. apply h2; auto. Qed. Lemma list_to_set_eq_iff : forall {T:Type} (l l':list T), list_to_set l = list_to_set l' <-> incl l l' /\ incl l' l. intros T l l'; split; intro h1. split; red; intros x h2; rewrite list_to_set_in_iff in h2; rewrite list_to_set_in_iff. rewrite <- h1; auto. rewrite h1; auto. destruct h1; apply anti_sym_incl_list_to_set; auto. Qed. Definition disjointl {T:Type} (l l':list T) : Prop := disjoint (list_to_set l) (list_to_set l'). Lemma sym_disjointl : forall {T:Type} (l l':list T), disjointl l l' -> disjointl l' l. intros T l l' h1. do 2 red in h1. do 2 red. rewrite Intersection_commutative; auto. Qed. Fixpoint list_minus {T:Type} (pfdt:type_eq_dec T) (l l':list T) : list T := match l with | nil => nil | a::al => if (in_dec pfdt a l') then (list_minus pfdt al l') else a::(list_minus pfdt al l') end. Lemma list_minus_nil2 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), list_minus pfdt l nil = l. intros T h0 l. induction l as [|a l h1]; simpl; auto. rewrite h1; auto. Qed. Lemma list_minus_nin_in_remove : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T), ~In a l -> In a l' -> list_minus pfdt l (remove pfdt a l') = list_minus pfdt l l'. intros T h0 l. induction l as [|b l h1]; simpl; auto. intros l' a h2 h4. apply not_or_and in h2. destruct h2 as [h2 h3]. destruct in_dec as [h5 | h5]. destruct in_dec as [h6 | h6]. apply h1; auto. contradict h6. eapply in_remove_neq_in_l. apply h5. assumption. destruct in_dec as [h6 | h6]. contradict h5. rewrite in_remove_iff. split; auto. f_equal. apply h1; auto. Qed. Lemma list_minus_cons_remove : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T), list_minus pfdt l (a::l') = remove pfdt a (list_minus pfdt l l'). intros T h1 l. induction l as [|c l ih]; simpl; auto. intros l' a. destruct (h1 a c) as [h2 | h2]; simpl; subst. destruct in_dec as [h2 | h2]. apply ih. simpl. destruct (h1 c c) as [h3 | h3]. apply ih. contradict h3; auto. destruct in_dec as [h3 | h3]. apply ih. simpl. destruct (h1 a c) as [h4 | h4]; try contradiction. f_equal. apply ih. Qed. Lemma list_minus_cons_remove2 : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T), In a l' -> list_minus pfdt (a::l) l' = list_minus pfdt l l'. intros T h1 l l' a h2; simpl. destruct in_dec as [h3 | h3]; auto. contradiction. Qed. Lemma list_minus_sing : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), list_minus pfdt l (a::nil) = remove pfdt a l. intros T hdt l. induction l as [|al l ih]; simpl; auto. intro a. destruct (hdt a al) as [h1 | h1]; subst; simpl; auto. f_equal. apply ih. Qed. Lemma list_minus_same : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), list_minus pfdt l l = nil. intros T h1 l. induction l as [|a l ih]; simpl; auto. destruct (h1 a a) as [h2 | h2]; auto. rewrite list_minus_cons_remove. rewrite ih. simpl. reflexivity. contradict h2; auto. Qed. Lemma impl_in_list_minus : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T), In x l -> ~In x l' -> In x (list_minus pfdt l l'). intros T h0 l. induction l as [|a l h1]; simpl. intros; contradiction. intros l' x h2 h3. destruct in_dec as [h4 | h4]. destruct h2 as [h2 | h2]; subst. contradiction. apply h1; auto. destruct h2 as [h2 | h2]; subst. left; auto. right. apply h1; auto. Qed. Lemma in_list_minus_impl : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T), In x (list_minus pfdt l l') -> (In x l /\ ~In x l'). intros T h0 l. induction l as [|a l ih]; simpl. intros; contradiction. intros l' x h1. destruct (in_dec h0 a l') as [h2 | h2]. specialize (ih _ _ h1). destruct ih as [h3 h4]. split. right; auto. assumption. simpl in h1. destruct h1 as [h1 | h1]; subst. split. left; auto. assumption. apply ih in h1. destruct h1 as [h1 h3]. split. right; auto. assumption. Qed. Lemma in_list_minus_impl' : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T), In x (list_minus pfdt l l') -> In x l. intros T h1 l l' x h2. pose proof (in_list_minus_impl h1 _ _ _ h2) as h3. destruct h3; auto. Qed. Lemma impl_list_minus_nil : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), incl l l' -> list_minus pfdt l l' = nil. intros T h1 l. induction l as [|a l ih]; simpl; auto. intros l' h2. destruct in_dec as [h3 | h3]. apply ih. red in h2; red. intros; apply h2. right; auto. contradict h3. red in h2. apply h2. left; auto. Qed. Lemma list_minus_eq_nil_iff : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), list_minus pfdt l l' = nil <-> incl l l'. intros T h1 l l'; split; intro h2. red. intros x h3. destruct (in_dec h1 x l') as [h4|h4]; auto. pose proof (impl_in_list_minus h1 l l' x h3 h4) as h5. rewrite h2 in h5. contradiction. apply impl_list_minus_nil; auto. Qed. Lemma list_to_set_list_minus : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), list_to_set (list_minus pfdt l l') = Setminus (list_to_set l) (list_to_set l'). intros T h0 l. induction l as [|a l h1]. intros. simpl. rewrite setminus_empty. reflexivity. intro l'. simpl. destruct in_dec as [h2 | h2]. specialize (h1 l'). rewrite setminus_add. rewrite h1. symmetry. rewrite Union_commutative. rewrite <- inclusion_iff_union. red; intros x h3. destruct h3 as [h3 h4]. destruct h3. contradict h4. rewrite <- list_to_set_in_iff; auto. simpl. rewrite setminus_add. rewrite h1. unfold Add. f_equal. rewrite disjoint_setminus_compat; auto. red. apply eq_empty. red; intros x h3. destruct h3 as [x h4 h5]. destruct h4. rewrite <- list_to_set_in_iff in h5. contradiction. Qed. Lemma no_dup_list_minus : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), NoDup l -> NoDup (list_minus pfdt l l'). intros T h1 l. induction l as [|a l ih]; simpl; auto. intros l' h2. destruct (in_dec h1 a l') as [h3 | h3]. apply ih. apply no_dup_cons in h2; auto. constructor. contradict h3. apply in_list_minus_impl in h3. destruct h3 as [h3]. apply no_dup_cons_nin in h2; contradiction. apply ih. apply no_dup_cons in h2; auto. Qed. Lemma incl_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), incl (remove pfdt x l) l. intros T h1 l x. red; intros a h2. rewrite in_remove_iff in h2. destruct h2; auto. Qed. Lemma incl_list_minus : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), incl (list_minus pfdt l l') l. intros T h0 l l'. rewrite incl_iff. rewrite list_to_set_list_minus. apply setminus_inc. Qed. Lemma in_incl : forall {T:Type} (l:list T) (x:T), In x l -> incl (x::nil) l. intros; red; intros x' h1; destruct h1; subst; try contradiction; auto. Qed. Lemma incl_dec : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), {incl l l'} + {~incl l l'}. intros T h1 l. induction l as [|a l ih]; simpl. intro; left; apply incl_nil. intros l'. specialize (ih l'). destruct ih as [ih | ih]. destruct (in_dec h1 a l') as [h2 | h2]. left. red; intros x h3. destruct h3 as [|h3]; subst; auto. right. intro h3. red in h3. specialize (h3 _ (in_eq _ _)). contradiction. right. contradict ih. eapply trans_incl. apply incl_cons. apply ih. Qed. Lemma lneq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), l <> nil -> l' <> nil -> (list_to_set l <> list_to_set l' <-> exists x:T, In x l /\ ~In x l' \/ In x l'/\ ~In x l). intros T h0 l l' h1 h2; split; intro h3. rewrite list_to_set_eq_iff in h3. destruct (incl_dec h0 l l') as [h4 | h4]. assert (h5:~ incl l' l). intro; contradict h3; split; auto. rewrite <- (list_minus_eq_nil_iff h0) in h5. exists (hd' _ h5). right. split. eapply incl_list_minus. apply in_hd'. intro h6. pose proof (in_hd' (list_minus h0 l' l) h5) as h7. apply in_list_minus_impl in h7. destruct h7; contradiction. rewrite <- (list_minus_eq_nil_iff h0) in h4. exists (hd' _ h4). left. split. eapply incl_list_minus. apply in_hd'. intro h6. pose proof (in_hd' (list_minus h0 l l') h4) as h7. apply in_list_minus_impl in h7. destruct h7; contradiction. rewrite list_to_set_eq_iff. intro h4. destruct h4 as [h4 h5]. destruct h3 as [x h6]. destruct h6 as [h6 | h6]. destruct h6 as [h6 h7]. apply h4 in h6; contradiction. destruct h6 as [h6 h7]. apply h5 in h6; contradiction. Qed. Lemma finite_mem_faml : forall {T:Type} (ll:faml T), finite_mem (faml_to_fam ll). intros T ll. induction ll as [|al lll h1]; simpl. do 2 red; intros; contradiction. intros A h2. destruct h2 as [A h2|A h2]; subst. apply h1; auto. inversion h2. subst. apply list_to_set_finite. Qed. Lemma in_faml_to_fam_impl : forall {T:Type} (ll:faml T) (A:Ensemble T) (pf:Inn (faml_to_fam ll) A), exists l:list T, In l ll /\ list_to_set l = A. intros T ll. induction ll as [|al ll h1]; simpl. intros; contradiction. intros A h2. destruct h2 as [A h2 | A h2]. specialize (h1 _ h2). destruct h1 as [l h1]. exists l. tauto. inversion h2. subst. clear h2. exists al. tauto. Qed. Lemma in_faml_impl : forall {T:Type} (ll:faml T) (l:list T), In l ll -> Inn (faml_to_fam ll) (list_to_set l). intros T ll. induction ll as [|al ll h1]; simpl. intros; contradiction. intros l h2. destruct h2 as [|h2]; subst. right. constructor. specialize (h1 _ h2). left; auto. Qed. Lemma not_nil_inh_faml_to_fam : forall {T:Type} (ll:faml T), ll <> nil -> Inhabited (faml_to_fam ll). intros T ll h1. pose proof (not_nil_inhabited _ h1) as h2. destruct h2 as [l h2]. apply Inhabited_intro with (list_to_set l). apply in_faml_impl. rewrite list_to_set_in_iff; auto. Qed. Lemma faml_to_fam_eq : forall {T:Type} (l:faml T), faml_to_fam l = list_to_set (map list_to_set l). intros T ll. apply Extensionality_Ensembles; red; split; red; intros A h1. rewrite <- list_to_set_in_iff. rewrite in_map_iff. apply in_faml_to_fam_impl in h1. destruct h1 as [l h1]. exists l. tauto. rewrite <- list_to_set_in_iff in h1. rewrite in_map_iff in h1. destruct h1 as [l h1]. destruct h1 as [h1 h2]. subst. apply in_faml_impl; auto. Qed. Lemma faml_to_fam_decompose : forall {T:Type} (pfdt:type_eq_dec (list T)) (ll:faml T) (l:list T), In l ll -> faml_to_fam ll = Add (faml_to_fam (remove pfdt l ll)) (list_to_set l). intros T hdt ll l h1. apply Extensionality_Ensembles; red; split; red; intros lx h2. apply in_faml_to_fam_impl in h2. destruct h2 as [la [h2 h3]]; subst. destruct (hdt l la) as [h3 | h3]; subst. right; auto. constructor. left. apply in_faml_impl. rewrite in_remove_iff; auto. destruct h2 as [lx h2 | lx h2]. apply in_faml_to_fam_impl in h2. destruct h2 as [la [h2 h3]]; subst. rewrite in_remove_iff in h2. destruct h2 as [h2 h3]. apply in_faml_impl; auto. destruct h2. apply in_faml_impl; auto. Qed. Definition pairwise_disjointl {T:Type} (ll:faml T) : Prop := forall l l':list T, (list_to_set l) <> (list_to_set l') -> In l ll -> In l' ll -> disjointl l l'. Lemma pairwise_disjotntl_nil : forall (T:Type), pairwise_disjointl (@nil (list T)). intro; red; simpl; intros; contradiction. Qed. Lemma pairwise_disjointl_compat : forall {T:Type} (ll:faml T), pairwise_disjointl ll <-> pairwise_disjoint (faml_to_fam ll). intros T ll; split; intro h1; red in h1; red. intros A B h2 h3 h4. apply in_faml_to_fam_impl in h3. apply in_faml_to_fam_impl in h4. destruct h3 as [l [h3 h5]]; subst. destruct h4 as [l' [h4 h6]]; subst. apply (h1 _ _ h2 h3 h4). intros l l' h2 h3 h4. apply (h1 _ _ h2 (in_faml_impl _ _ h3) (in_faml_impl _ _ h4)). Qed. Lemma pairwise_disjointl_incl : forall {T:Type} (ll lm:faml T), pairwise_disjointl ll -> incl lm ll -> pairwise_disjointl lm. intros T ll lm h1 h2. red in h1. red. intros l l' h3 h4 h5. apply h2 in h4. apply h2 in h5. apply h1; auto. Qed. Lemma pairwise_disjointl_cons : forall {T:Type} (ll:faml T), pairwise_disjointl ll -> forall l:list T, (forall l', In l' ll -> disjointl l l') -> pairwise_disjointl (l::ll). intros T ll h1 l' h2. red in h1. red. intros l1 l2 h3 h4 h5. destruct h4 as [|h4], h5 as [|h5]; subst; subst. contradict h3; auto. apply h2; auto. specialize (h2 _ h4). apply sym_disjoint; auto. apply h1; auto. Qed. Lemma pairwise_disjointl_cons_rev : forall {T:Type} (ll:faml T) (l:list T), pairwise_disjointl (l::ll) -> pairwise_disjointl ll. intros T ll l h1. red in h1. red. intros l1 l2 h3 h4 h5. apply h1; auto; right; auto. Qed. Definition partition_faml {T:Type} (ll:faml T) (l:list T) : Prop := partition_fam (faml_to_fam ll) (list_to_set l). Lemma partition_faml_nil : forall {T:Type} (l:list T), partition_faml nil l -> l = nil. intros T l h1. do 2 red in h1; simpl in h1. destruct h1 as [h1 h2]. rewrite empty_family_union in h2. symmetry in h2. apply empty_set_nil in h2; auto. Qed. Lemma partition_faml_incl : forall {T:Type} (ll:faml T) (l lx:list T), partition_faml ll l -> In lx ll -> incl lx l. intros T ll l lx h1 h2. do 2 red in h1. destruct h1 as [h1 h3]. red; intros x h4. rewrite list_to_set_in_iff. rewrite <- h3. econstructor. apply in_faml_impl. apply h2. rewrite <- list_to_set_in_iff; auto. Qed. Lemma partition_faml_cons : forall {T:Type} (pfdt:type_eq_dec T) (ll:faml T) (l l':list T), ~ Inn (faml_to_fam ll) (list_to_set l) -> partition_faml (l::ll) l' -> partition_faml ll (list_minus pfdt l' l). intros T h1 ll l l' h2 h3. do 2 red in h3. do 2 red. destruct h3 as [h3 h4]. simpl in h3. simpl. split. apply pairwise_disjoint_add_rev in h3; auto. simpl in h4. rewrite family_union_add in h4. rewrite list_to_set_list_minus. rewrite <- h4. rewrite Union_commutative. rewrite setminus_union. rewrite disjoint_setminus_compat; auto. red. apply pairwise_disjoint_add_family_union; auto. Qed. Lemma partition_faml_eq : forall {T:Type} (ll:faml T) (la lb lc:list T) (a:T), partition_faml ll la -> In lb ll -> In lc ll -> In a lb -> In a lc -> list_to_set lb = list_to_set lc. intros T ll la lb lc a h1 h2 h3 h4 h5. apply in_faml_impl in h2. apply in_faml_impl in h3. rewrite list_to_set_in_iff in h4, h5. eapply partition_fam_eq. apply h1. assumption. assumption. apply h4. assumption. Qed. Lemma partition_faml_dec : forall {T:Type} (ll:faml T) (la lb lu:list T), partition_faml ll lu -> In la ll -> In lb ll -> {list_to_set la = list_to_set lb} + {disjointl la lb}. intros T ll la lb lu h1 h2 h3. unfold disjointl. apply in_faml_impl in h2. apply in_faml_impl in h3. eapply partition_fam_dec. apply h1. assumption. assumption. Qed. Fixpoint plusl (l:list nat) := match l with | nil => 0 | a::l' => a + (plusl l') end. Fixpoint timesl (l:list nat) := match l with | nil => 1 | a::l' => a * (timesl l') end. Lemma plusl_app : forall l l', plusl (l ++ l') = plusl l + plusl l'. intro l. induction l as [|a l ih]; simpl; auto. intro l'. rewrite ih, plus_assoc; auto. Qed. Lemma timesl_app : forall l l', timesl (l ++ l') = timesl l * timesl l'. intro l. induction l as [|a l ih]; simpl; auto. intro l'. rewrite ih, mult_assoc; auto. Qed. Lemma plusl_eq0_iff : forall (l:list nat), plusl l = 0 <-> all_sing l 0. intro l; induction l as [|a l ih]; simpl; split;intro h1; auto; try red; try (intros; contradiction). apply plus_is_O in h1; destruct h1; subst. rename H0 into h2. rewrite ih in h2; red in h2. intros ? h3; destruct h3; subst; auto. hr ih. red in h1; red; intros c hc; specialize (h1 _ (in_cons _ _ _ hc)); auto. rewrite <- ih in hr. rewrite hr, <- plus_n_O. red in h1; specialize (h1 _ (in_eq _ _)); auto. Qed. Lemma plusl_eq1_iff : forall (l:list nat), plusl l = 1 <-> exists! l0 l1, l = l0 ++ 1::nil ++ l1 /\ all_sing l0 0 /\ all_sing l1 0. intro l; induction l as [|a l ih]; simpl. split; intro h1. discriminate. destruct h1 as [l0 [h1 h2]]. destruct h1 as [l1 h1]. red in h1. destruct h1 as [[h1 [h3 h5]] h4]. symmetry in h1. apply app_eq_nil in h1. destruct h1; discriminate. split; intro h1. apply plus_is_one in h1. destruct h1 as [h1 | h1]; destruct h1 as [? h1]; subst. rewrite ih in h1. destruct h1 as [l0 [h1 h2]]; destruct h1 as [l1 h1]. red in h1. destruct h1 as [[h1 [h3 h5]] h4]; subst; simpl in h4. exists (0::l0); red; split. exists l1; red; split. split; simpl; auto. split; red; intros x h6. destruct h6 as [|h6]; subst; auto. specialize (h5 _ h6); auto. simpl. intros l h6. destruct h6 as [h6 h7]. apply cons_inj in h6. destruct h6 as [h6 h8]; clear h6. apply app_inj_eq_length in h8; auto. destruct h8 as [h8 h9]. clear h8. inversion h9; auto. intros l' h6. specialize (h2 (tl l')). hfirst h2. destruct h6 as [l2 [[h6 [h7 h9]] h8]]; simpl in h6. exists l2; simpl. red. split. split. pose proof (f_equal (@tl _) h6) as h10; simpl in h10. rewrite h10. rewrite tl_app; auto. intro; subst. inversion h6. split; auto. apply all_sing_tl; auto. intros l h11. destruct h11 as [h11 [h12 h13]]. apply h8. split. simpl. rewrite h11. assert (h14:l' <> nil). intro; subst; inversion h6. rewrite (hd_compat' _ h14); simpl. f_equal. red in h7. symmetry. apply h7. apply in_hd'. split; auto. specialize (h2 hf); subst. destruct h6 as [l h6]. red in h6. destruct h6 as [h6 h7]. destruct h6 as [h6 [h8 h9]]. assert (h10:l' <> nil). intro; subst; inversion h6. rewrite (hd_compat' _ h10); simpl. f_equal. symmetry. red in h8. apply h8. apply in_hd'. rewrite plusl_eq0_iff in h1. exists nil. red. split; simpl. exists l. red; split. split; auto. split; auto. red. intros; contradiction. intros l' h2. destruct h2 as [h2 [h3 h4]]. inversion h2; auto. intros l' h2. destruct h2 as [l0 h3]. red in h3. destruct h3 as [h3 h4]. destruct h3 as [h3 [h6 h7]]. destruct l' as [|a' l']; auto. simpl in h3. inversion h3; subst. red in h6. specialize (h6 _ (in_eq _ _)). discriminate. destruct h1 as [l0 h1]. red in h1. destruct h1 as [h1 h2]. destruct h1 as [l1 h1]. red in h1. destruct h1 as [[h1 [h3 h5]] h4]. pose proof (f_equal plusl h1) as h6. simpl in h6. rewrite plusl_app in h6; simpl in h6. rewrite h6. rewrite plus_S_shift_r. f_equal. rewrite <- plusl_eq0_iff in h3, h5. rewrite h3, h5; auto with arith. Qed. Lemma timesl_eq0_iff : forall (l:list nat), timesl l = 0 <-> In 0 l. intro l; induction l as [|a l ih]; simpl; split; intro h1; try discriminate; try contradiction. apply mult_is_O in h1; destruct h1 as [|h1]; subst. left; auto. rewrite ih in h1; right; auto. destruct h1 as [|h1]; subst. apply mult_0_l. rewrite <- ih in h1; rewrite h1, mult_0_r; auto. Qed. Lemma timesl_eq1_iff : forall (l:list nat), timesl l = 1 <-> all_sing l 1. intro l; induction l as [|a l ih]; simpl; split;intro h1; auto; try red; try (intros; contradiction). apply mult_is_one in h1; destruct h1; subst. rename H0 into h2. rewrite ih in h2; red in h2. intros ? h3; destruct h3; subst; auto. hr ih. red in h1; red; intros c hc; specialize (h1 _ (in_cons _ _ _ hc)); auto. rewrite <- ih in hr. rewrite hr, mult_1_r. red in h1; specialize (h1 _ (in_eq _ _)); auto. Qed. Lemma plusl_nprod_eq : forall (l:list nat), plusl l = plus_nprod (nprod_of_list _ l). intro l; induction l; simpl; auto. Qed. Lemma timesl_nprod_eq : forall (l:list nat), timesl l = times_nprod (nprod_of_list _ l). intro l; induction l; simpl; auto. Qed. Lemma plus_nprod_eq : forall {n} (x:nat^n), plus_nprod x = plusl (nprod_to_list _ _ x). intro n; induction n; simpl; auto; intro x; destruct x; simpl; auto. Qed. Lemma times_nprod_l : forall {n} (x:nat^n), times_nprod x = timesl (nprod_to_list _ _ x). intro n; induction n; simpl; auto; intro x; destruct x; simpl; auto. Qed. Lemma length_app_list_nprod_to_list : forall {T n} (x:(list T)^n), length (app_list (nprod_to_list _ _ x)) = plus_nprod (map_nprod (fun i pf => length (nth_prod i pf x))). intros T n; induction n as [|n ih]; simpl; auto. intro x; destruct x as [l x]; simpl; specialize (ih x); rewrite app_length; f_equal; rewrite ih; repeat f_equal; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; unfold seg_fun_to_S; repeat f_equal; apply proof_irrelevance. Qed. Definition list_to_set_injable {T:Type} (P:list T->Prop) : Prop := forall (l l':list T), P l -> P l' -> list_to_set l = list_to_set l' -> l = l'. Lemma list_to_set_injable_in_cons : forall {T:Type} (ll:faml T) (la:list T), list_to_set_injable (fun lx => In lx (la::ll)) -> list_to_set_injable (fun lx => In lx ll). intros T ll la h1. red in h1; red. intros lb lc h2 h3. specialize (h1 _ _ (in_cons _ _ _ h2) (in_cons _ _ _ h3)). intro h4. apply h1 in h4; auto. Qed. Lemma list_to_set_injable_in_remove : forall {T:Type} (pfdt:type_eq_dec (list T)) (ll:faml T) (la:list T), list_to_set_injable (fun lx => In lx ll) -> list_to_set_injable (fun lx => In lx (remove pfdt la ll)). intros T h1 ll la h2. red in h2; red; intros l l' h3 h4 h5. rewrite in_remove_iff in h3, h4. destruct h3 as [h3 ?], h4 as [h4 ?]. apply h2; auto. Qed. Lemma partition_faml_remove : forall {T:Type} (pfdt:type_eq_dec T) (ll:faml T) (l l':list T), In l' ll -> list_to_set_injable (fun xl => In xl ll) -> partition_faml ll l -> partition_faml (remove (list_eq_dec pfdt) l' ll) (list_minus pfdt l l'). intros T h1 ll l l' h2 h3 h4. pose proof (faml_to_fam_decompose (list_eq_dec h1) ll l' h2) as h6. destruct h4 as [h4 h5]. rewrite h6 in h4, h5; clear h6. do 2 red in h4; do 2 red. split. apply pairwise_disjoint_add_rev in h4; auto. rewrite family_union_add in h5. rewrite list_to_set_list_minus. rewrite <- h5. rewrite Union_commutative. rewrite setminus_union. rewrite disjoint_setminus_compat; auto. red. apply pairwise_disjoint_add_family_union; auto. intro h6. apply in_faml_to_fam_impl in h6. destruct h6 as [la [h6 h7]]. rewrite in_remove_iff in h6. destruct h6 as [h6 h8]. apply h3 in h7; subst. contradict h8; auto. assumption. assumption. Qed. Lemma list_to_set_bij : forall {T:Type} (l:list T), exists f:{x:T | In x l} -> {x:T | Inn (list_to_set l) x}, bijective f. intros T l. pose proof (list_to_set_in_iff l) as h1. assert (h2:forall x : T, In x l -> Inn (list_to_set l) x). intros. rewrite <- h1; assumption. exists (fun t:{x : T | In x l} => exist _ (proj1_sig t) (h2 _ (proj2_sig t))). red. split. red. intros ? ? h3. apply exist_injective in h3. apply proj1_sig_injective. assumption. red. intros y. destruct (h1 (proj1_sig y)) as [? h1r]. specialize (h1r (proj2_sig y)). exists (exist _ (proj1_sig y) h1r). simpl. rewrite unfold_sig. apply existTexist. apply subsetT_eq_compat. reflexivity. Qed. Lemma list_to_set_finitet : forall {T:Type} (l:list T), FiniteT {x:T | In x l}. intros T l. pose proof (list_to_set_bij l) as h1. pose proof (list_to_set_finite l) as h2. pose proof (Finite_ens_type _ h2) as h3. destruct h1 as [f h1]. pose proof (bijective_impl_invertible _ h1) as h4. pose proof (invertible_impl_inv_invertible _ h4) as h5. apply (bij_finite _ _ _ h3 h5). Qed. (*See [finite_set_list_no_dup] for a stronger version.*) Lemma finite_set_list : forall {T:Type} (E:Ensemble T), Finite E -> exists (l:list T), E = list_to_set l. intros T E h1. induction h1 as [| E h2 h3 t h4]. exists nil. unfold list_to_set. reflexivity. destruct h3 as [l h5]. exists (t::l). unfold list_to_set. unfold list_to_set in h5. rewrite <- h5. reflexivity. Qed. Lemma list_to_set_app : forall {T:Type} (l l':list T), list_to_set (l ++ l') = Union (list_to_set l) (list_to_set l'). intros T l. induction l as [|a l h1]; simpl. intro l'. destruct l' as [|b l']; simpl. rewrite empty_union. reflexivity. rewrite empty_union. reflexivity. intro l'. rewrite h1. unfold Add. rewrite <- Union_associative. rewrite (Union_commutative (list_to_set l') (Singleton a)). rewrite Union_associative. reflexivity. Qed. Lemma add_cons_compat : forall {T:Type} (l:list T) (a:T), Add (list_to_set l) a = list_to_set (a::l). intros. simpl. auto. Qed. Lemma subtract_remove_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), Subtract (list_to_set l) x = list_to_set (remove pfdt x l). intros T h0 l. induction l; simpl. apply sub_empty. intro x. simpl. destruct (h0 x a) as [h2 | h3]. subst. destruct (in_dec h0 a l) as [h4 | h5]. rewrite list_to_set_in_iff in h4. pose proof (sub_add_same_in _ _ h4) as h6. rewrite h6. apply IHl. rewrite <- (remove_not_in' h0 l a). rewrite list_to_set_in_iff in h5. apply sub_add_same_nin; auto. auto. rewrite sub_add_comm. rewrite IHl. apply add_cons_compat. apply neq_sym_iff; auto. Qed. Lemma list_to_set_map_remove : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:T->U) (l:list T) (x:T) (pf:In x l), Injective f -> list_to_set (map f (remove pfdt x l)) = Subtract (list_to_set (map f l)) (f x). intros T U hdt hdu f l x h2 hinj. rewrite (subtract_remove_compat hdu). apply Extensionality_Ensembles; red; split; red; intros c h3; rewrite <- list_to_set_in_iff in h3; rewrite <- list_to_set_in_iff. rewrite in_map_iff in h3. destruct h3 as [a [h3 h4]]; subst. rewrite in_remove_iff in h4. rewrite in_remove_iff. destruct h4 as [h4 h5]. split. apply in_map; auto. contradict h5. apply hinj in h5; auto. rewrite in_remove_iff in h3. destruct h3 as [h3 h4]. rewrite in_map_iff in h3. destruct h3 as [a [h3 h5]]; subst. apply in_map. rewrite in_remove_iff. split; auto. intro; subst. contradict h4; auto. Qed. Definition p_mem_l {T:Type} (P:T->Prop) (l:list T) : Prop := forall x:T, In x l -> P x. Lemma p_mem_l_iff : forall {T:Type} (P:T->Prop) (l:list T), p_mem_l P l <-> p_mem P (list_to_set l). intros T P l. split. intro h1. red in h1. red. intros x h2. apply h1. rewrite list_to_set_in_iff. assumption. intro h1. red in h1. red. intros x h2. apply h1. rewrite <- list_to_set_in_iff. assumption. Qed. Lemma p_mem_l_cons : forall {T:Type} (P:T->Prop) (l:list T) (a:T), p_mem_l P (a::l) -> p_mem_l P l. intros T P l a h1. red in h1. red. intros x h2. apply h1; right; auto. Qed. Lemma p_mem_l_type_dep_eq_dec_impl : forall {T:Type} (P:T->Prop) (l:list T), p_mem_l P l -> type_dep_eq_dec P -> type_dep_eq_dec (fun x => P x /\ In x l). intros T P l. induction l as [|c l h1]; simpl. intros h1 h2. red. intros ? ? h3. destruct h3; contradiction. intros h2 h3. apply p_mem_l_cons in h2. specialize (h1 h2 h3). red in h1. red. intros a b h4 h5. destruct h4 as [h4 h6], h5 as [h5 h7]. red in h3. specialize (h3 a b h4 h5). assumption. Qed. Definition no_dup_mem {T:Type} (ll:faml T) := p_mem_l (@NoDup T) ll. Lemma no_dup_mem_cons : forall {T:Type} (ll:faml T) (la:list T), no_dup_mem (la::ll) -> no_dup_mem ll. intros ? ? ? h1; do 2 red in h1; do 2 red; intros; apply h1; right; auto. Qed. Lemma no_dup_mem_remove : forall {T:Type} (pfdt:type_eq_dec (list T)) (ll:faml T) (la:list T), no_dup_mem ll -> no_dup_mem (remove pfdt la ll). intros ? ? ? ? h1; do 2 red in h1; do 2 red; intros ? h0; apply h1. rewrite in_remove_iff in h0. destruct h0; auto. Qed. (*Are indices of some list.*) Inductive are_inds (T:Type) (ln:list nat) : Prop := | are_inds_intro : forall l:list T, p_mem_l (fun i => i < length l) ln -> are_inds T ln. Definition lexcl_mid {T:Type} (l:list T) (P:T->Prop) : Type := forall x:T, In x l -> {P x} + {~P x}. Lemma in_dep_dec : forall {T:Type} {P:T->Prop}, type_dep_eq_dec P -> forall (l:list T), p_mem_l P l -> forall (x:T) (pfx:P x), {In x l} + {~In x l}. intros T P h1 l. induction l as [|a l ih]; simpl. intros; right; intro; contradiction. intros h2 x h3. hfirst ih. red. red in h2. intros; apply h2. right; auto. specialize (ih hf). specialize (ih x h3). red in h2. specialize (h2 a (in_eq _ _)). red in h1. specialize (h1 _ _ h2 h3). destruct h1 as [h1 | h1]; subst. left; left; auto. destruct ih as [h4 | h4]. left; right; auto. right. intro h5. destruct h5 as [h5 | h5]; subst. contradict h1; auto. contradiction. Qed. Lemma in_neq : forall {T} (l:list T) x a, In x l -> ~In a l -> x <> a. intros; intro; subst; auto. Qed. Lemma in_cons_neq : forall {T} (l:list T) x a, In x (a::l) -> x <> a -> In x l. intros ? ? ? ? h1; destruct h1; subst; auto; intro h1; contradict h1; auto. Qed. Fixpoint in_dep {T:Type} {P:T->Prop} (x:T) (pfx:P x) (l:list T) {struct l} : Prop := match l with | nil => False | a::l' => P a /\ x = a \/ (in_dep x pfx l') end. Lemma impl_in_dep : forall {T:Type} {P:T->Prop} (x:T) (pfx:P x) (l:list T), In x l -> in_dep x pfx l. intros T P x h1 l. revert h1; revert x. induction l as [|a l ih]; simpl; auto. intros ? ? h2. destruct h2; [left | right]; auto; subst; auto. Qed. Lemma in_dep_iff : forall {T:Type} {P:T->Prop} (x:T) (pfx:P x) (l:list T), in_dep x pfx l <-> (In x l /\ P x). intros T P x h1 l. split. induction l as [|l a ih]; simpl; auto. intro h3. destruct h3 as [h3 | h3]; subst. split; try left; auto. destruct h3; auto. specialize (ih h3). destruct ih as [h2 h5]. split; auto. intro h2; destruct h2 as [h2 h3]; apply impl_in_dep; auto. Qed. Fixpoint remove_dep {T:Type} (P:T->Prop) (pfe:type_dep_eq_dec P) (l:list T) (x:T) (pfp:P x) {struct l} : p_mem_l P l -> list T := match l with | nil => (fun _ => nil) | a::l' => (fun pfpm => let r := remove_dep P pfe l' x pfp (p_mem_l_cons _ _ _ pfpm) in if (pfe x a pfp (pfpm a (in_eq _ _))) then r else a::r) end. Lemma remove_dep_not_in : forall {T:Type} (P:T->Prop) (pfe:type_dep_eq_dec P) (l:list T) (x:T) (pfp:P x) (pfpm:p_mem_l P l), ~In x l -> remove_dep P pfe l x pfp pfpm = l. intros T P h1 l. induction l as [|a l ih]; simpl; auto. intros x h2 h3 h4. destruct_head_match as [h5 | h5]; subst. contradict h4. left; auto. f_equal. rewrite ih; auto. Qed. Lemma add_list_to_set_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), In a l -> Add (list_to_set (remove pfdt a l)) a = list_to_set l. intros T h0 l a h1. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2 as [x h2 | x h2]. rewrite <- subtract_remove_compat in h2. destruct h2; auto. destruct h2. rewrite <- list_to_set_in_iff; auto. rewrite <- subtract_remove_compat. rewrite add_sub_compat_in; auto. rewrite <- list_to_set_in_iff; auto. Qed. Definition seg_list (n:nat) := seq 0 n. Lemma seg_list0 : seg_list 0 = nil. auto. Qed. Lemma length_seg_list : forall (n:nat), length (seg_list n) = n. intro; unfold seg_list; rewrite seq_length; auto. Qed. Lemma seg_list_eq_nil : forall (n:nat), seg_list n = nil -> n = 0. intro n. destruct n; auto. intro h1. unfold seg_list, seq in h1. discriminate. Qed. Lemma seq_eq_seq_0 : forall (m n:nat), seq m n = map (plus m) (seg_list n). intros m n. revert m. induction n as [|n h1]; simpl; auto. intro m. rewrite plus_comm. simpl. f_equal. rewrite <- (seq_shift n 0) at 1. rewrite map_map. rewrite h1. f_equal. apply functional_extensionality. intros; omega. Qed. Lemma seg_list_seq_cons : forall m:nat, 0 < m -> seg_list m = 0::seq 1 (pred m). intro m. destruct m as [|m]. intro h1; apply lt_irrefl in h1. contradiction. intro h1. destruct (zerop m) as [|h2]; subst. unfold seg_list; simpl; auto. unfold seg_list. simpl. auto. Qed. Lemma list_to_set_seg_list_eq : forall (n:nat), list_to_set (seg_list n) = seg n. intro n. induction n as [|n h1]; simpl. rewrite seg_0_eq. reflexivity. rewrite seq_eq_seq_0. rewrite map_im_compat at 1. rewrite h1. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2 as [x h2 | x h2]. destruct h2 as [x h2]. subst. destruct h2 as [h2]. constructor. omega. destruct h2. constructor. omega. destruct h2 as [h2]. destruct (zerop x) as [h3 | h3]. subst. right. constructor. left. apply Im_intro with (pred x). constructor. omega. omega. Qed. Lemma list_to_set_eq_seg_in : forall (l:list nat) (m n:nat), list_to_set l = seg n -> m < n -> In m l. intros l m n h1 h2. pose proof (in_seg _ _ h2) as h3. rewrite <- h1 in h3. rewrite <- list_to_set_in_iff in h3; auto. Qed. Lemma list_to_set_eq_seg_lt : forall (l:list nat) (m n:nat), list_to_set l = seg n -> In m l -> m < n. intros l m n h1 h2. rewrite list_to_set_in_iff in h2. rewrite h1 in h2. destruct h2; auto. Qed. Lemma in_seg_list_iff : forall (m n:nat), In m (seg_list n) <-> m < n. intros m n. rewrite list_to_set_in_iff. rewrite list_to_set_seg_list_eq. split. intro h1. destruct h1; auto. intros; constructor; auto. Qed. Lemma in_seq_iff : forall i m n, In i (seq m n) <-> m <= i < m + n. intros i m n. revert i m. induction n as [|n ih]; simpl; intros i m; split; intro h1. contradiction. omega. destruct h1 as [|h1]; subst. omega. specialize (ih i (S m)). rewrite ih in h1; omega. destruct h1 as [h1 h2]. apply le_lt_eq_dec in h1. destruct h1 as [h1|]; subst. right. rewrite ih. omega. left; auto. Qed. Lemma hd_seg_list' : forall (n:nat) (pf:seg_list n <> nil), hd' (seg_list n) pf = 0. intro n. destruct n as [|n]; simpl; auto. intro h1; contradict h1; auto. Qed. Lemma tl_seg_list : forall (n:nat), tl (seg_list n) = seq 1 (pred n). intro n; destruct n; simpl; auto. Qed. Lemma seq_S : forall (m n:nat), 0 < m -> seq n m = n::seq (S n) (pred m). intro m. induction m as [|m ih]; simpl. intros; omega. intros; simpl; auto. Qed. Lemma seq_nil : forall (m n:nat), seq m n = nil -> n = 0. intros m n. revert m. destruct n as [|n]; auto. intros m h1. rewrite seq_S in h1. discriminate. auto with arith. Qed. Lemma seq_app : forall (m n:nat), 0 < m -> seq n m = seq n (pred m) ++ n+m-1::nil. intro m. induction m as [|m ih]; simpl. intros; omega. intros n h1; simpl. destruct m as [|m]; simpl. f_equal. omega. f_equal. red in h1. apply le_S_n in h1. apply le_lt_eq_dec in h1. destruct h1 as [h1 | h1]. simpl in ih. specialize (ih (S n) h1). rewrite ih. f_equal. f_equal. omega. discriminate. Qed. Lemma seg_list_seq_app : forall m n:nat, m <= n -> seg_list n = seg_list m ++ seq m (n - m). intro m. unfold seg_list. induction m as [|m ih]. simpl. intros; f_equal; auto with arith. intros n h1. destruct (zerop m) as [|h2]; subst. simpl. rewrite seg_list_seq_cons. f_equal. simpl. f_equal; omega. omega. rewrite ih; try omega. simpl. rewrite (S_pred _ _ h2) at 1. simpl. f_equal. pose proof @seq_S. rewrite (seq_S (n - m) m); try omega. rewrite (seq_app _ _ h2). simpl. rewrite <- app_assoc. rewrite <- minus_n_O. simpl. repeat f_equal; try omega. Qed. Lemma seg_list_S : forall n, seg_list (S n) = seg_list n ++ (n :: nil). unfold seg_list; intro; simpl. destruct (zerop n) as [|h1]; subst; simpl; auto. symmetry. rewrite seq_S; simpl. f_equal. symmetry. rewrite seq_app; simpl. f_equal. f_equal; auto with arith. assumption. assumption. Qed. Lemma firstn_seq : forall (m n r:nat), firstn m (seq n r) = seq n (min m r). intros m n r. destruct (le_lt_dec m r) as [h1 | h1]. rewrite min_l; auto. revert h1 n. revert r. induction m as [|m ih]; simpl; auto. intros r h1 n. specialize (ih (pred r)). hfirst ih. apply le_S_n. erewrite <- S_pred; auto. assert (h2:0 < r). omega. apply h2. specialize (ih hf). rewrite seq_S. f_equal. apply ih. omega. rewrite min_r; auto with arith. pose proof (seq_length r n) as h2. rewrite <- h2 in h1. apply firstn_length_ge. omega. Qed. Lemma seg_list_app : forall (n:nat), 0 < n -> seg_list n = seg_list (pred n) ++ n-1::nil. intros; unfold seg_list; rewrite seq_app; auto. Qed. Lemma seg_list_not_nil : forall (n:nat), 0 < n -> seg_list n <> nil. intros; rewrite seg_list_app; auto. apply neq_sym. apply app_cons_not_nil. Qed. Lemma intersection_seg_sing : forall (n:nat) (m:nat), m >= n -> Intersection (list_to_set (seg_list n)) (Singleton m) = Empty_set _. intros m n h1; apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2 as [x h2 h3]. destruct h3. rewrite <- list_to_set_in_iff in h2. rewrite in_seg_list_iff in h2. omega. destruct h2; contradiction. Qed. Lemma seg_list_eq_sing : forall (n a:nat), seg_list n = a::nil -> n = 1 /\ a = 0. intros n a. destruct n as [|n]; simpl. rewrite seg_list0. intro; discriminate. intro h1. rewrite seg_list_S in h1. rewrite app_eq_iff in h1. destruct h1 as [h1 h2]. rewrite length_seg_list in h1, h2. destruct n as [|n]. simpl in h1, h2. inversion h2; auto. simpl in h1, h2. rewrite skipn_nil in h2. discriminate. Qed. Lemma seg_list_eq_couple : forall (n a b:nat), seg_list n = a::b::nil -> n = 2 /\ a = 0 /\ b = 1. intros n a b. destruct n as [|n]; simpl. rewrite seg_list0. intro; discriminate. destruct n as [|n]; simpl. rewrite seg_list_S. rewrite seg_list0. simpl. intro h1. inversion h1; discriminate. do 2 rewrite seg_list_S. intro h1. destruct a as [|a], b as [|b]; auto with arith. simpl in h1. rewrite app_eq_iff in h1. destruct h1 as [h1 h2]. rewrite app_length in h2. rewrite length_seg_list in h2. simpl in h2. rewrite S_compat in h2. simpl in h2. destruct n as [|n]. simpl in h2. inversion h2. simpl in h2. rewrite skipn_nil in h2. discriminate. pose proof (f_equal (fun l => last l 0) h1) as h2. simpl in h2. rewrite last_app_sing in h2. apply S_inj in h2; subst. pose proof (f_equal (@removelast _) h1) as h2. simpl in h2. rewrite removelast_app_sing in h2. destruct b as [|b]; auto. rewrite app_eq_iff in h2. destruct h2 as [h2 h3]. rewrite length_seg_list in h3. simpl in h3. rewrite skipn_nil in h3. discriminate. rewrite app_eq_iff in h1. destruct h1 as [h1 h2]. rewrite app_length in h2. rewrite length_seg_list in h2. simpl in h2. rewrite S_compat in h2. simpl in h2. destruct n as [|n]. simpl in h2. inversion h2. simpl in h2. rewrite skipn_nil in h2. discriminate. rewrite app_eq_iff in h1. destruct h1 as [h1 h2]. rewrite app_length, length_seg_list in h2; simpl in h2. rewrite S_compat in h2; simpl in h2. destruct n as [|n]. simpl in h2. inversion h2. subst. simpl in h1. inversion h1. simpl in h2. rewrite skipn_nil in h2. discriminate. Qed. Lemma seg_list_eq_cons : forall (l:list nat) (n a:nat), seg_list n = a::l -> a = 0 /\ l = seq 1 (pred n). intros l n. revert l. induction n as [|n ih]. intros l a h1. rewrite seg_list0 in h1. discriminate. intros l a h1. rewrite seg_list_S in h1. destruct (nil_dec' l) as [|h2]; subst. rewrite app_eq_iff in h1. destruct h1 as [h1 h2]. rewrite length_seg_list in h1. rewrite length_seg_list in h2. symmetry in h2. destruct (zerop n) as [|h3]; subst. simpl in h2; simpl. inversion h2; auto. rewrite (S_pred _ _ h3) in h2. simpl in h2. rewrite skipn_nil in h2. discriminate. simpl. rewrite (app_removelast_last a h2) in h1 at 1. rewrite app_comm_cons in h1. rewrite app_sings_eq_iff in h1. destruct h1 as [h1 ?]; subst. rewrite seq_app. simpl. rewrite <- minus_n_O. apply ih in h1. destruct h1 as [? h1]; subst. split; auto. rewrite <- h1. apply app_removelast_last; auto. destruct (zerop (last l a)) as [h3 | h3]. rewrite h3 in h1. rewrite seg_list0 in h1. discriminate. assumption. Qed. Fixpoint inv_iml {T U:Type} (pfdu:type_eq_dec U) (f:T->U) (l:list T) (u:U) : list T := match l with | nil => nil | a::l' => let il := inv_iml pfdu f l' u in if (pfdu u (f a)) then a :: il else il end. Lemma inv_iml_functional : forall {T U:Type} (pfdu:type_eq_dec U) (f f':T->U) (l l':list T) (u u':U), f = f' -> l = l' -> u = u' -> inv_iml pfdu f l u = inv_iml pfdu f' l' u'. intros; subst; auto. Qed. Lemma in_inv_iml_iff : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (u:U) (f:T->U) (x:T), In x (inv_iml pfdu f l u) <-> In x l /\ f x = u. intros T U h1 l. induction l as [|a l ih]; simpl. intros; split. intro; contradiction. intro h2. destruct h2; auto. intros u f x. lr_if_goal. fold hlr. destruct hlr as [h2 | h2]; subst. split. intro h2. destruct h2 as [|h2]; subst. intuition. specialize (ih (f a) f x). rewrite ih in h2. destruct h2 as [h2 h3]. intuition. intro h2. destruct h2 as [h2 h3]. destruct h2 as [|h2]; subst. left; auto. right. rewrite ih. split; auto. split. intros h3. rewrite ih in h3. destruct h3 as [h3 h4]; subst. intuition. intro h3. destruct h3 as [h3]; subst. destruct h3 as [|h3]; subst. contradict h2; auto. rewrite ih. intuition. Qed. Lemma incl_inv_iml : forall {T U:Type} (pfdu:type_eq_dec U) (f:T->U) (l:list T) (u:U), incl (inv_iml pfdu f l u) l. intros; red; intros. rewrite in_inv_iml_iff in H. destruct H; auto. Qed. Lemma inv_iml_eq_nil_iff : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:T->U) (l:list T) (b:U), inv_iml pfdu f l b = nil <-> (l = nil \/ (forall a, In a l -> f a <> b)). intros T U h1 h1' f l. induction l as [|a l ih]; simpl. intuition. intros b. destruct (h1' b (f a)) as [|h2]; subst. split; intro h2. discriminate. destruct h2 as [|h2]. discriminate. specialize (h2 a (or_introl eq_refl)). contradict h2; f_equal; apply proof_irrelevance. specialize (ih b). rewrite ih. split; intro h3. destruct h3 as [|h3]; subst. right. intros a' h3. destruct h3; subst; try contradiction. contradict h2. rewrite <- h2. f_equal; apply proof_irrelevance. right. intros a' h4. destruct (h1 a a') as [|h5]; subst. contradict h2; rewrite <- h2; f_equal; apply proof_irrelevance. destruct h4 as [|h4]; subst. contradict h5; auto. apply h3; auto. destruct h3 as [|h3]. discriminate. destruct (nil_dec' l) as [h4 | h4]. left; auto. right. intros a' h5. apply h3; auto. Qed. Lemma inv_iml_sing_eq : forall {T U:Type} {l:list T} {a:T} (pfdu:type_eq_dec U) (f:T->U) (u:U), u = f a -> inv_iml pfdu f (a::nil) u = a::nil. intros T U l a h1 f u h2. subst. simpl. lr_if_goal. fold hlr. destruct hlr as [h2 | h2]; auto. contradict h2; auto. Qed. Lemma count_occ_map : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:T->U) (l:list T) (y:U), count_occ pfdu (map f l) y = length (inv_iml pfdu f l y). intros T U h1 h2 f l. induction l as [|a l ih]; simpl; auto. intro y. rewrite switch_eq_dec_eq at 1. destruct (h2 y (f a)) as [h4 | h4]; subst. unfold switch_eq_dec; simpl. f_equal; apply ih. unfold switch_eq_dec; auto. Qed. (*See way below for [lS_app]*) Definition famlS (ll:faml nat) := map lS ll. Definition faml_lpred (ll:faml nat) := map lpred ll. Lemma famlS_nil : famlS nil = nil. auto. Qed. Lemma famlS_eq_nil_iff : forall (ll:faml nat), famlS ll = nil <-> ll = nil. intro ll; split; destruct ll; simpl; intros; try discriminate; auto. Qed. Lemma famlS_cons : forall (ll:faml nat) (al:list nat), famlS (al::ll) = (lS al) :: famlS ll. auto. Qed. Lemma length_famlS : forall (ll:faml nat), length (famlS ll) = length ll. intro ll; induction ll; simpl; auto. Qed. Lemma in_famlS_iff : forall (ll:faml nat) (ln:list nat), In ln (famlS ll) <-> exists ln':list nat, In ln' ll /\ ln = lS ln'. intros ll ln. unfold famlS. rewrite in_map_iff. split; intro h1; destruct h1 as [n hm]; destruct hm as [? h2]; subst; exists n; tauto. Qed. Lemma in_famlS_iff' : forall (ll:faml nat) (ln:list nat), In ln (famlS ll) <-> (In (lpred ln) ll /\ ~In 0 ln). unfold famlS; intros ll ln. rewrite in_map_iff. split; intro h1. destruct h1 as [ln' [h1 h2]]; subst. split. rewrite lpred_lS; auto. apply nin0_lS. destruct h1 as [h1 h2]. exists (lpred ln). split. rewrite lS_lpred_iff in h2. rewrite h2. f_equal. rewrite lpred_lS; auto. assumption. Qed. Lemma length_faml_lpred : forall (ll:faml nat), length (faml_lpred ll) = length ll. intro ll; induction ll; simpl; auto. Qed. Lemma in_faml_lpred_iff : forall (ll:faml nat) (lm:list nat), In lm (faml_lpred ll) <-> exists ln:list nat, In ln ll /\ lm = lpred ln. unfold faml_lpred. intros ll lm; split; intro h2. rewrite in_map_iff in h2. destruct h2 as [ln [h2 h3]]; subst. exists ln; auto. destruct h2 as [ln [h3 h4]]; subst. rewrite in_map_iff. exists ln; auto. Qed. Lemma map_inj : forall {T U:Type} (f:T->U), Injective f -> forall (l l':list T), map f l = map f l' -> l = l'. intros T U f h1 l. induction l as [|a l ih]; simpl. intros ? h2. symmetry in h2; apply map_eq_nil in h2; auto. intros l' h2. destruct l' as [|a' l']. simpl in h2; discriminate. simpl in h2. inversion h2. apply h1 in H0; subst. f_equal; auto. Qed. Lemma inj_lS : Injective lS. red; intros l l' h1. unfold lS in h1. apply map_inj in h1; auto. red; apply S_inj. Qed. Lemma inj_famlS : forall (ll ll':faml nat), famlS ll = famlS ll' -> ll = ll'. intros ll ll' h1. unfold famlS in h1. apply map_inj in h1; auto. apply inj_lS. Qed. Lemma map_dep_sig_functional : forall {T U:Type} {l l':list T} (pf:l = l'), let V := (fun al => list {a | In a al}) in forall (df:dep_fun V) (f:{a | In a l}->U), map f (df l) = map (transfer_dep (U:=fun al:list T=> {a|In a al}->U) pf f) (df l'). intros; subst. rewrite transfer_dep_eq_refl; auto. Qed. Definition fun_in_dep {T:Type} (l:list T) (U:Type) : Type := prop_dep_fun (fun x => In x l) U. Definition fun_in_dep_prop {T:Type} (l:list T) (P:T->Prop) (U:Type) : Type := prop_dep_fun2 (fun x => In x l) P U. Definition fun_in_dep_nil (T U:Type) : fun_in_dep (@nil T) U := fun x pf => False_rect _ pf. Lemma inj_dep_nil : forall (T U:Type), inj_dep (fun_in_dep_nil T U). red; intros; contradiction. Qed. Lemma surj_dep_nil : forall (T:Type), surj_dep (fun_in_dep_nil T False). red; intros; contradiction. Qed. Lemma bij_dep_nil : forall (T:Type), bij_dep (fun_in_dep_nil T False). red; intros; split; [apply inj_dep_nil | apply surj_dep_nil]. Qed. Lemma fun_in_dep_ext : forall {T:Type} (l:list T) (U:Type) (f g:fun_in_dep l U), (forall (x:T) (pf:In x l), f x pf = g x pf) -> f = g. intros; apply prop_dep_fun_ext; auto. Qed. Lemma fun_in_dep_functional : forall {T U:Type} {l:list T} (f:fun_in_dep l U) (x x':T) (pfx:In x l) (pfx':In x' l), x = x' -> f x pfx = f x' pfx'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma fun_in_dep_prop_ext : forall {T:Type} (l:list T) (P:T->Prop) (U:Type) (f g:fun_in_dep_prop l P U), (forall (x:T) (pf:In x l) (pfp:P x), f x pf pfp = g x pf pfp) -> f = g. intros; apply prop_dep_fun2_ext; auto. Qed. Definition fun_in_dep_cons {T U} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l U) (a:T) (b:U) : fun_in_dep (a::l) U := fun x (pf:In x (a::l)) => match (in_cons_dec pfdt _ _ _ pf) with | left pfe => b | right pfin => f x pfin end. Definition seg_fun_in_dep {U:Type} {n:nat} (f:seg_fun n U) : fun_in_dep (seg_list n) U. red in f; do 2 red. intros m h1; eapply f. apply (iff1 (in_seg_list_iff _ _) h1). Defined. Definition remove_ind_in_dep {T U} {l:list T} m (f:fun_in_dep l U) : fun_in_dep (remove_ind l m) U := fun x pf => f x (in_remove_ind _ _ _ pf). Lemma inj_remove_ind_in_dep : forall {T U} {l:list T} m (f:fun_in_dep l U), inj_dep f -> inj_dep (remove_ind_in_dep m f). intros ? ? ? ? ? h1; red in h1; red; intros ? ? ? ? h2; unfold remove_ind_in_dep in h2; apply h1 in h2; auto. Qed. Definition prop_in_dep {T:Type} (l:list T) : Type := prop_dep_prop (fun x => In x l). Definition prop_in_dep_u {T:Type} (l:list T) (U:Type) : Type := prop_dep_prop_u (fun x => In x l) U. Definition set_in_dep {T:Type} (U:T->Set) (l:list T) : Type := prop_dep_set_dep (fun x => In x l) U. Definition transfer_fun_in_dep {T U:Type} {l l':list T} (f:fun_in_dep l U) (pf:l = l') : fun_in_dep l' U := transfer_dep pf f (U:= (fun xl => fun_in_dep xl U)). Definition transfer_fun_in_dep_prop {T U:Type} {l l':list T} {P:T->Prop} (f:fun_in_dep_prop l P U) (pf:l = l') : fun_in_dep_prop l' P U := transfer_dep pf f (U:= (fun xl => fun_in_dep_prop xl P U)). Definition transfer_fun_in_dep_ran {T U V:Type} {l:list T} (pfe:U = V) (f:fun_in_dep l U) : fun_in_dep l V := fun x pf => transfer pfe (f x pf). Definition transfer_prop_in_dep {T:Type} {l l':list T} (f:prop_in_dep l) (pf:l = l') : prop_in_dep l' := transfer_dep pf f. Definition transfer_prop_in_dep_u {T U:Type} {l l':list T} (f:prop_in_dep_u l U) (pf:l = l') : prop_in_dep_u l' U := transfer_dep pf f (U:= (fun xl => prop_in_dep_u xl U)). Definition transfer_set_in_dep {T:Type} {U:T->Set} {l l':list T} (f:set_in_dep U l) (pf:l = l') : set_in_dep U l' := transfer_dep (U:=fun xl =>set_in_dep U xl) pf f. Lemma transfer_fun_in_dep_eq_refl : forall {T U:Type} {l:list T} (f:fun_in_dep l U), transfer_fun_in_dep f eq_refl = f. intros; unfold transfer_fun_in_dep; rewrite transfer_dep_eq_refl; auto. Qed. Lemma transfer_fun_in_dep_prop_eq_refl : forall {T U:Type} {l:list T} {P:T->Prop} (f:fun_in_dep_prop l P U), transfer_fun_in_dep_prop f eq_refl = f. intros; unfold transfer_fun_in_dep_prop; rewrite transfer_dep_eq_refl; auto. Qed. Lemma transfer_prop_in_dep_eq_refl : forall {T:Type} {l:list T} (f:prop_in_dep l), transfer_prop_in_dep f eq_refl = f. intros; unfold transfer_prop_in_dep; rewrite transfer_dep_eq_refl; auto. Qed. Lemma transfer_prop_in_dep_u_eq_refl : forall {T U:Type} {l:list T} (f:prop_in_dep_u l U), transfer_prop_in_dep_u f eq_refl = f. intros; unfold transfer_prop_in_dep_u; rewrite transfer_dep_eq_refl; auto. Qed. Lemma transfer_set_in_dep_eq_refl : forall {T:Type} {l:list T} {U:T->Set} (f:set_in_dep U l), transfer_set_in_dep f eq_refl = f. intros; unfold transfer_set_in_dep; rewrite transfer_dep_eq_refl; auto. Qed. Lemma transfer_fun_in_dep_compat : forall {T U:Type} {l l':list T} (f:fun_in_dep l U) (x:T) (pfin:In x l') (pfeq:l = l'), transfer_fun_in_dep f pfeq x pfin = f x (transfer_dep_r (U:=fun xl => In x xl) pfeq pfin). intros; subst; rewrite transfer_fun_in_dep_eq_refl, transfer_dep_r_eq_refl; auto. Qed. Lemma transfer_fun_in_dep_prop_compat : forall {T U:Type} {l l':list T} {P:T->Prop} (f:fun_in_dep_prop l P U) (x:T) (pfin:In x l') (pfp:P x) (pfeq:l = l'), transfer_fun_in_dep_prop f pfeq x pfin pfp = f x (transfer_dep_r (U:=fun xl => In x xl) pfeq pfin) pfp. intros; subst; rewrite transfer_fun_in_dep_prop_eq_refl, transfer_dep_r_eq_refl; auto. Qed. Lemma transfer_prop_in_dep_compat : forall {T:Type} {l l':list T} (f:prop_in_dep l) (x:T) (pfin:In x l') (pfeq:l = l'), transfer_prop_in_dep f pfeq x pfin = f x (transfer_dep_r (U:=fun xl => In x xl) pfeq pfin). intros; subst; rewrite transfer_prop_in_dep_eq_refl, transfer_dep_r_eq_refl; auto. Qed. Lemma transfer_prop_in_dep_u_compat : forall {T U:Type} {l l':list T} (f:prop_in_dep_u l U) (x:T) (pfin:In x l') (pfeq:l = l'), transfer_prop_in_dep_u f pfeq x pfin = f x (transfer_dep_r (U:=fun xl => In x xl) pfeq pfin). intros; subst; rewrite transfer_prop_in_dep_u_eq_refl, transfer_dep_r_eq_refl; auto. Qed. Lemma transfer_set_in_dep_compat : forall {T:Type} {U:T->Set} {l l':list T} (f:set_in_dep U l) (x:T) (pfin:In x l') (pfeq:l = l'), transfer_set_in_dep f pfeq x pfin = f x (transfer_dep_r (U:=fun xl => In x xl) pfeq pfin). intros; subst; rewrite transfer_set_in_dep_eq_refl, transfer_dep_r_eq_refl; auto. Qed. Definition compose_in_dep_prop {T U} {l:list U} (f:fun_in_dep l T) (R:T->Prop) : prop_in_dep l := fun_in_dep_fun_compose f R. Definition compose_sig_prop {T U} {l:list T} (f:U->{x:T|In x l}) (R:{x:T| In x l}->Prop) : U -> Prop := compose R f. Definition compose_sig_prop_in_dep {T U} {l:list U} (f:T->{x:U | In x l}) (R:prop_in_dep l) : T->Prop := fun x => R (proj1_sig (f x)) (proj2_sig _). Definition unq_ex_ex_to_fun_in_dep {T U:Type} {l:list T} {f:list T->list U} {P:prop_in_dep_u l U} (pf:forall y:U, In y (f l) -> exists! x:T, exists pfin:In x l, P x pfin y) : fun_in_dep (f l) T. do 2 red. intros y h1. specialize (pf y h1). refine (proj1_sig (constructive_definite_description _ pf)). Defined. Lemma unq_ex_ex_to_fun_in_dep_compat : forall {T U:Type} {l:list T} {f:list T->list U} {P:prop_in_dep_u l U} (pf:forall (y:U) (pf:In y (f l)), exists! x:T, exists pfin:In x l, P x pfin y) (y:U) (pfiny:In y (f l)), let f := unq_ex_ex_to_fun_in_dep pf in let x := f y pfiny in exists pfin:In x l, P x pfin y. unfold unq_ex_ex_to_fun_in_dep; intros; destruct constructive_definite_description; auto. Qed. Lemma unq_ex_ex_to_fun_in_dep_functional : forall {T U:Type} {l:list T} {f:list T->list U} {P:prop_in_dep_u l U} (pf:forall (y:U) (pf:In y (f l)), exists! x:T, exists pfin:In x l, P x pfin y) (y y':U) (pfiny:In y (f l)) (pfiny':In y' (f l)), y = y' -> unq_ex_ex_to_fun_in_dep pf y pfiny = unq_ex_ex_to_fun_in_dep pf y' pfiny'. intros; subst. pose proof (proof_irrelevance _ pfiny' pfiny); subst; auto. Qed. Definition id_in_dep {T:Type} (l:list T) : fun_in_dep l T := id_prop_dep (fun x => In x l). Lemma id_in_dep_compat : forall {T:Type} (l:list T) (x:T) (pf:In x l), id_in_dep l x pf = x. intros; unfold id_in_dep, id_prop_dep; auto. Qed. Lemma inj_id_in_dep : forall {T} (l:list T), inj_dep (id_in_dep l). intros; red; intros; auto. Qed. Definition id_in_dep_sig {T:Type} (l:list T) : {x:T|In x l} -> T := id_prop_dep_sig (fun x => In x l). Definition id_in_dep_sig' {T:Type} (l:list T) : fun_in_dep l {x:T|In x l} := id_prop_dep_sig' (fun x => In x l). Definition sigl_from_cons {T U l a} (f:{x:T | In x (a::l)} -> U) : {x:T|In x l} -> U := fun x => f (exist (fun c => In c (a::l)) _ (in_cons _ _ _ (proj2_sig x))). Definition sigl_with_cons {T U l a} (f:{x:T | In x (a::l)} -> U) : {x:T|In x (removelast (a::l)) } -> U := fun x => f (exist (fun c => In c (a::l)) _ (removelast_in _ _ (proj2_sig x))). Definition const_in_dep {T U:Type} (l:list T) (val:U) : fun_in_dep l U := fun _ _ => val. Definition fun_from_cons {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U) : fun_in_dep l U := fun x (pf:In x l) => f x (in_cons _ _ _ pf). Lemma fun_from_cons_functional : forall {T U:Type} {l:list T} {a:T} (f g:fun_in_dep (a::l) U) (x:T), f x = g x-> fun_from_cons f x = fun_from_cons g x. intros T U l a f g x h1. apply functional_extensionality. intro h2. unfold fun_from_cons. rewrite h1; auto. Qed. Lemma fun_from_cons_compat : forall {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U) (x:T) (pf:In x l), let f' := fun_from_cons f in f' x pf = f x (in_cons _ _ _ pf). intros; unfold f', fun_from_cons; auto; f_equal; apply proof_irrelevance. Qed. Definition fun_from_cons2 {T U} {l:list T} {a b:T} (f:fun_in_dep (a::b::l) U) : fun_in_dep (a::l) U := fun x pfx => f x (in_cons2 a b x pfx). Lemma inj_from_cons : forall {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U), inj_dep f -> inj_dep (fun_from_cons f). intros T U l a f h1. red in h1; red. intros x y hx hy h2. apply (h1 x y (in_cons _ _ _ hx) (in_cons _ _ _ hy)). unfold fun_from_cons in h2. erewrite fun_in_dep_functional. symmetry. erewrite fun_in_dep_functional. symmetry. apply h2. reflexivity. reflexivity. Qed. Lemma surj_from_cons_in : forall {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U), In a l -> surj_dep f -> surj_dep (fun_from_cons f). intros T U l a f h1 h2; red in h2; red; intro y. specialize (h2 y); destruct h2 as [x [h2 h3]]; subst. destruct h2 as [|h2]; subst. exists x, h1; unfold fun_from_cons; f_equal; apply proof_irrelevance. exists x, h2; unfold fun_from_cons; f_equal; apply proof_irrelevance. Qed. Lemma bij_from_cons_in : forall {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U), In a l -> bij_dep f -> bij_dep (fun_from_cons f). intros ? ? ? ? ? ? h1; red in h1; red; destruct h1; split; [apply inj_from_cons | apply surj_from_cons_in]; auto. Qed. Lemma inj_from_cons2 : forall {T U} {l:list T} {a b:T} (f:fun_in_dep (a::b::l) U), inj_dep f -> inj_dep (fun_from_cons2 f). intros T U l a b f h1. red in h1; red. intros x y hx hy h2. apply (h1 x y (in_cons2 _ _ _ hx) (in_cons2 _ _ _ hy)). unfold fun_from_cons in h2. erewrite fun_in_dep_functional. symmetry. erewrite fun_in_dep_functional. symmetry. apply h2. reflexivity. reflexivity. Qed. Lemma surj_from_cons2_in : forall {T U} {l:list T} {a b:T} (f:fun_in_dep (a::b::l) U), In b l -> surj_dep f -> surj_dep (fun_from_cons2 f). unfold fun_from_cons2; intros ? ? ? ? ? ? hi h1; red in h1; red; intro y; specialize (h1 y); destruct h1 as [x [h1 h3]]; subst; destruct h1 as [|h1]; subst; [exists x, (in_eq _ _); f_equal; apply proof_irrelevance |]; destruct h1 as [|h1]; subst; [exists x, (in_cons _ _ _ hi); f_equal; apply proof_irrelevance | exists x, (in_cons _ _ _ h1); f_equal; apply proof_irrelevance]. Qed. Lemma bij_from_cons2_in : forall {T U:Type} {l:list T} {a b:T} (f:fun_in_dep (a::b::l) U), In b l -> bij_dep f -> bij_dep (fun_from_cons2 f). intros ? ? ? ? ? ? ? h1; red in h1; red; destruct h1; split; [apply inj_from_cons2 | apply surj_from_cons2_in]; auto. Qed. Definition fun_from_cons_ran {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U) (pfinj:inj_dep f) (pfn : ~In a l) : fun_in_dep l {y:U | y <> f a (in_eq _ _)} := fun x (pfx:In x l) => let pfx' := in_cons _ _ _ pfx in let pfneq := inj_dep_impl_contra f pfinj _ _ pfx' (in_eq _ _) (in_neq _ _ _ pfx pfn) in exist (fun c => c <> f a (in_eq _ _)) (f x pfx') pfneq. Lemma fun_from_cons_ran_functional : forall {T U:Type} {l:list T} {a:T} (f g:fun_in_dep (a::l) U) (pfnin:~In a l) (pfinjf:inj_dep f) (pfinjg:inj_dep g) (x:T) (pfx:In x l) (pfx':In x (a::l)), f x pfx' = g x pfx' -> proj1_sig (fun_from_cons_ran f pfinjf pfnin x pfx) = proj1_sig (fun_from_cons_ran g pfinjg pfnin x pfx). unfold fun_from_cons_ran; intros; simpl; erewrite f_equal_dep; auto; rewrite H at 1; f_equal; apply proof_irrelevance. Qed. Lemma fun_from_cons_ran_compat : forall {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U) (pfnin:~In a l) (pfinjf:inj_dep f) (x:T) (pfx:In x l), let pfx' := in_cons _ _ _ pfx in proj1_sig (fun_from_cons_ran f pfinjf pfnin x pfx) = f x pfx'. unfold fun_from_cons_ran; intros; simpl; auto. Qed. Lemma inj_from_cons_ran : forall {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U) (pfinj:inj_dep f) (pfnin : ~In a l), inj_dep (fun_from_cons_ran f pfinj pfnin). intros T U l a f h1 h2; red in h1; red; intros x y ha ha' h3. unfold fun_from_cons_ran in h3; apply exist_injective in h3; apply h1 in h3; auto. Qed. Lemma inj_surj_from_cons_ran : forall {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U) (pfinj:inj_dep f) (pfnin : ~In a l), surj_dep f -> surj_dep (fun_from_cons_ran f pfinj pfnin). intros T U l a f hinj hnin h1; red in h1; red; intro y; destruct y as [y h2]. specialize (h1 y). destruct h1 as [x [h1 ?]]; subst. destruct h1 as [|h1]; subst. contradict h2; f_equal; apply proof_irrelevance. exists x, h1. unfold fun_from_cons_ran. apply f_equal_dep; f_equal; apply proof_irrelevance. Qed. Lemma inj_surj_bij_from_cons_ran : forall {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U) (pfinj:inj_dep f) (pfnin : ~In a l), surj_dep f -> bij_dep (fun_from_cons_ran f pfinj pfnin). unfold fun_from_cons_ran; intros; red; split; red; [intros x y hx hy h1; eapply exist_injective in h1; eapply pfinj in h1| intro y; destruct y as [y hy]; red in H; specialize (H y); destruct H as [c [hc ?]]; subst; destruct hc as [|hc]; [subst; contradict hy; f_equal; apply proof_irrelevance| exists c, hc; do 2 apply f_equal_dep] ]; auto. Qed. Definition tl_in_dep {T U:Type} {l:list T} (f:fun_in_dep l U) : fun_in_dep (tl l) U := fun x pf => f x (in_tl _ _ pf). Lemma tl_in_dep_cons : forall {T U:Type} {l:list T} {a} (pfdt:type_eq_dec T) (f:fun_in_dep (a::l) U), tl_in_dep f = fun_from_cons f. unfold tl_in_dep, fun_from_cons; intros; simpl; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; f_equal; apply proof_irrelevance. Qed. Lemma inj_tl : forall {T U} {l:list T} (f:fun_in_dep l U), inj_dep f -> inj_dep (tl_in_dep f). intros ? ? ? ? h1; red in h1; red; intros x y hx hy h2. unfold tl_in_dep in h2. apply h1 in h2; auto. Qed. Lemma surj_tl : forall {T U} {l:list T} (f:fun_in_dep l U) (pfn:l <> nil) (pfin:In (hd' l pfn) (tl l)), surj_dep f -> surj_dep (tl_in_dep f). intros ? ? ? ? h1 h2 h3; red in h3; red; intro y. specialize (h3 y). destruct h3 as [x [h3 ?]]; subst. pose proof h3 as h4. rewrite (hd_compat' _ h1) in h4. destruct h4 as [|h4]; subst. exists _, h2; unfold tl_in_dep; f_equal; apply proof_irrelevance. exists _, h4. unfold tl_in_dep; f_equal; apply proof_irrelevance. Qed. Lemma bij_tl : forall {T U} {l:list T} (f:fun_in_dep l U) (pfn:l <> nil) (pfin:In (hd' l pfn) (tl l)), bij_dep f -> bij_dep (tl_in_dep f). intros ? ? ? ? ? ? h1; red in h1. red. split; destruct h1 as [h1 h2]; [apply inj_tl; auto | eapply surj_tl; auto; apply pfin]. Qed. Definition tl_in_dep_ran {T U} {l:list T} (f:fun_in_dep l U) (pfn:l <> nil) (pfinj:inj_dep f) (pfnin : ~In (hd' l pfn) (tl l)) : fun_in_dep (tl l) {y:U | y <> f (hd' l pfn) (in_hd' _ _)} := fun x (pfx:In x (tl l)) => let pfx' := in_tl _ _ pfx in let pfneq := inj_dep_impl_contra f pfinj _ _ pfx' (in_hd' _ _) (in_neq _ _ _ pfx pfnin) in exist (fun c => c <> f (hd' l pfn) (in_hd' _ _)) (f x pfx') pfneq. Lemma proj1_sig_tl_in_dep_ran : forall {T U} {l:list T} (f:fun_in_dep l U) (pfn:l <> nil) (pfinj:inj_dep f) (pfnin : ~In (hd' l pfn) (tl l)), proj1_sig_fun_dep (tl_in_dep_ran f pfn pfinj pfnin) = tl_in_dep f. unfold proj1_sig_fun_dep, tl_in_dep_ran, tl_in_dep, fun_in_dep_fun_compose; intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intros; simpl; auto. Qed. Lemma inj_tl_ran: forall {T U} {l:list T} (f:fun_in_dep l U) (pfn:l <> nil) (pfinj:inj_dep f) (pfnin : ~In (hd' l pfn) (tl l)), inj_dep (tl_in_dep_ran f pfn pfinj pfnin). unfold tl_in_dep_ran; intros ? ? ? ? ? hinj ?; red. intros ? ? ? ? h1. apply exist_injective in h1. apply hinj in h1; auto. Qed. Lemma surj_tl_ran : forall {T U} {l:list T} (f:fun_in_dep l U) (pfn:l <> nil) (pfinj:inj_dep f) (pfnin : ~In (hd' l pfn) (tl l)), surj_dep f -> surj_dep (tl_in_dep_ran f pfn pfinj pfnin). unfold tl_in_dep_ran; intros ? ? ? ? hn hinj ? hs; red; intro y; destruct y as [y hy]. specialize (hs y). destruct hs as [x [hx ?]]; subst. exists x. ex_goal. pose proof hx as hx'. rewrite (hd_compat' _ hn) in hx'. destruct hx' as [|h1]; subst; auto. contradict hy. f_equal; apply proof_irrelevance. exists hex. apply proj1_sig_injective; simpl; f_equal; apply proof_irrelevance. Qed. Lemma bij_tl_ran : forall {T U} {l:list T} (f:fun_in_dep l U) (pfn:l <> nil) (pfinj:inj_dep f) (pfnin : ~In (hd' l pfn) (tl l)), bij_dep f -> bij_dep (tl_in_dep_ran f pfn pfinj pfnin). intros ? ? ? ? ? ? ? h1; red in h1; red. destruct h1 as [h1 h2]; split. apply inj_tl_ran; auto. apply surj_tl_ran; auto. Qed. (*This function removes a point in the domain list of f, yet DOESN'T remove other equal values at different points.*) (*Contrast with [fun_in_dep_remove_in], which does remove points at equal values.*) Definition fun_in_dep_remove {T U:Type} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l U) (x:T) : fun_in_dep (remove pfdt x l) U. do 2 red; intros ? h1; eapply f. apply in_remove_impl in h1. apply h1. Defined. Lemma fun_in_dep_remove_functional : forall {T U:Type}{l:list T} (pfdt:type_eq_dec T) (f g:fun_in_dep l U) (a x:T), f x = g x-> fun_in_dep_remove pfdt f a x = fun_in_dep_remove pfdt g a x. intros T U l h1 f g a x h2. apply functional_extensionality. intro h3. unfold fun_in_dep_remove. rewrite h2; auto. Qed. Lemma fun_in_dep_remove_compat : forall {T U:Type}{l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l U) (a x:T) (pf:In x (remove pfdt a l)), let pf' := in_remove_impl pfdt _ _ _ pf in fun_in_dep_remove pfdt f a x pf = f x pf'. simpl; intros. unfold fun_in_dep_remove; auto. Qed. Lemma inj_dep_remove : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l U) (x:T), inj_dep f -> inj_dep (fun_in_dep_remove pfdt f x). intros T U l h1 f x h2. red in h2; red. intros a b h3 h4. unfold fun_in_dep_remove. intro h5. apply h2 in h5; auto. Qed. Definition fun_in_dep_comm_cons {T U:Type} {l:list T} {a b:T} (f:fun_in_dep (a::b::l) U) : fun_in_dep (b::a::l) U. do 2 red in f; do 2 red. intros x h1. apply in_list2_comm in h1. refine (f _ h1). Defined. Lemma fun_in_dep_comm_cons_compat : forall {T U:Type} {l:list T} {a b:T} (f:fun_in_dep (a::b::l) U) (x:T) (pf:In x (a::b::l)), let f' := fun_in_dep_comm_cons f in let pf' := in_list2_comm _ _ _ _ pf in f x pf = f' x pf'. simpl; unfold fun_in_dep_comm_cons; intros; f_equal; apply proof_irrelevance. Qed. Lemma fun_from_cons2_fun_in_dep_comm_cons : forall {T U:Type} {l:list T} {a b:T} (f:fun_in_dep (a::b::l) U), fun_from_cons (fun_from_cons f) = fun_from_cons (fun_from_cons (fun_in_dep_comm_cons f)). intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro. do 4 rewrite fun_from_cons_compat. rewrite fun_in_dep_comm_cons_compat at 1. f_equal. apply proof_irrelevance. Qed. Fixpoint inv_im_in_dep {T U:Type} (pfdu:type_eq_dec U) (l:list T) (u:U) : fun_in_dep l U -> list T := match l with | nil => fun _ => nil | a::l' => fun f => let il := inv_im_in_dep pfdu l' u (fun_from_cons f) in if (pfdu u (f a (in_eq _ _))) then a :: il else il end. Lemma inv_im_in_dep_functional : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (u u':U) (f f':fun_in_dep l U), u = u' -> f = f' -> inv_im_in_dep pfdu l u f = inv_im_in_dep pfdu l u' f'. intros; subst; auto. Qed. Lemma inv_im_in_dep_transfer_functional : forall {T U:Type} (pfdu:type_eq_dec U) (l l':list T) (u u':U) (f:fun_in_dep l U) (pf:l = l'), u = u' -> let f' := transfer_fun_in_dep f pf in inv_im_in_dep pfdu l u f = inv_im_in_dep pfdu l' u' f'. simpl; intros; subst; rewrite transfer_fun_in_dep_eq_refl; auto. Qed. Lemma in_inv_im_in_dep_iff : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (u:U) (f:fun_in_dep l U) (x:T), In x (inv_im_in_dep pfdu l u f) <-> exists pf:In x l, f x pf = u. intros T U h1 l. induction l as [|a l ih]; simpl. intros; split. intro; contradiction. intro h2. destruct h2; auto. intros u f x. lr_if_goal. fold hlr. destruct hlr as [h2 | h2]; subst. split. intro h2. destruct h2 as [|h2]; subst. exists (or_introl eq_refl). f_equal. apply proof_irrelevance. specialize (ih (f a (in_eq a l)) (fun_from_cons f) x). rewrite ih in h2. destruct h2 as [h2 h3]. exists (or_intror h2). rewrite fun_from_cons_compat in h3. rewrite <- h3. f_equal. apply proof_irrelevance. intro h2. destruct h2 as [h2 h3]. destruct h2 as [|h2]; subst. left; auto. right. rewrite ih. exists h2. rewrite <- h3. rewrite fun_from_cons_compat. f_equal. apply proof_irrelevance. split. intros h3. rewrite ih in h3. destruct h3 as [h3 h4]; subst. exists (or_intror h3). rewrite fun_from_cons_compat. f_equal. apply proof_irrelevance. intro h3. destruct h3 as [h3]; subst. destruct h3 as [|h3]; subst. contradict h2. f_equal. apply proof_irrelevance. rewrite ih. exists h3. rewrite fun_from_cons_compat. f_equal. apply proof_irrelevance. Qed. Lemma incl_inv_im_in_dep : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (u:U) (f:fun_in_dep l U), incl (inv_im_in_dep pfdu l u f) l. intros T U h1 l u f. red. intros x h2. rewrite in_inv_im_in_dep_iff in h2. destruct h2; auto. Qed. Lemma inv_im_in_dep_eq_nil_iff : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (b:U) (f:fun_in_dep l U), inv_im_in_dep pfdu l b f = nil <-> (l = nil \/ (forall (a:T) (pf:In a l), f a pf <> b)). intros T U h1 h1' l. induction l as [|a l ih]; simpl. intuition. intros b f. destruct (h1' b (f a (in_eq a l))) as [|h2]; subst. split; intro h2. discriminate. destruct h2 as [|h2]. discriminate. specialize (h2 a (or_introl eq_refl)). contradict h2; f_equal; apply proof_irrelevance. specialize (ih b (fun_from_cons f)). rewrite ih. split; intro h3. destruct h3 as [|h3]; subst. right. intros a' h3. destruct h3; subst; try contradiction. contradict h2. rewrite <- h2. f_equal; apply proof_irrelevance. right. intros a' h4. destruct (h1 a a') as [|h5]; subst. contradict h2; rewrite <- h2; f_equal; apply proof_irrelevance. destruct h4 as [|h4]; subst. contradict h5; auto. specialize (h3 _ h4). unfold fun_from_cons in h3. erewrite f_equal_dep. apply h3. reflexivity. destruct h3 as [|h3]. discriminate. destruct (nil_dec' l) as [h4 | h4]. left; auto. right. intros a' h5. specialize (h3 _ (in_cons _ _ _ h5)). contradict h3. rewrite <- h3. unfold fun_from_cons. f_equal. Qed. Lemma inv_im_in_dep_eq_nil : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (b:U) (f:fun_in_dep l U), inv_im_in_dep pfdu l b f = nil -> l = nil \/ (exists (a:T) (pf:In a l), f a pf <> b). intros T U h1 l. induction l as [|a l ih]; simpl. intros; left; auto. intros b f h2. right. destruct (h1 b (f a (in_eq a l))) as [h3 | h3]. discriminate. specialize (ih _ _ h2). destruct ih as [ih | ih]; subst. exists a, (or_introl eq_refl). contradict h3; subst. f_equal. apply proof_irrelevance. destruct ih as [a' [h4 h5]]. exists a', (in_cons _ _ _ h4). contradict h5; subst. unfold fun_from_cons. f_equal. Qed. Lemma inv_im_in_dep_sing_eq : forall {T U:Type} {l:list T} {a:T} (pfdu:type_eq_dec U) (f:fun_in_dep (a::nil) U) (u:U), u = f a (in_eq _ _) -> inv_im_in_dep pfdu _ u f = a::nil. intros T U l a h1 f u h2. subst. simpl. lr_if_goal. fold hlr. destruct hlr as [h2 | h2]; auto. contradict h2; auto. Qed. Definition list_pred {T} (P:T->Prop) (l:list T) : Type := forall x, In x l -> P x. Lemma list_pred_irrelevance : forall {T} {P:T->Prop} (l:list T) (pfl pfl':list_pred P l), pfl = pfl'. intros. apply functional_extensionality_dep. intro x. apply functional_extensionality_dep. intro h1. apply proof_irrelevance. Qed. Definition list_pred_type {T} (l:list T) : Type := forall x, In x l -> Prop. Definition list_pred_type_transfer {T} {l l':list T} (pfe:l = l') (P:list_pred_type l) : list_pred_type l' := transfer_dep pfe P. Lemma list_pred_type_transfer_eq_refl : forall {T} {l:list T} (P:list_pred_type l), list_pred_type_transfer eq_refl P = P. unfold list_pred_type_transfer; intros; rewrite transfer_dep_eq_refl; auto. Qed. Definition list_pred_dec {T} {l:list T} (P:list_pred_type l) : Type := forall x (pfin:In x l), { P x pfin} + {~P x pfin}. Lemma list_pred_dec_irrelevance : forall {T} {l:list T} {P:list_pred_type l} (pfl pfl':list_pred_dec P), pfl = pfl'. intros T l P hl hl'. apply functional_extensionality_dep. intro i. apply functional_extensionality_dep. intro j. destruct (hl i j), (hl' i j); f_equal; try apply proof_irrelevance; try contradiction. Qed. Lemma list_pred_dec_type_transfer : forall {T} {l l':list T} (pfe:l = l') (P:list_pred_type l), list_pred_dec P -> list_pred_dec (list_pred_type_transfer pfe P). intros; subst; rewrite list_pred_type_transfer_eq_refl; auto. Qed. Definition list_pred_incl {T} {P:T->Prop} {l l':list T} {a:T} (P':list_pred P l') (pf:incl l l') : list_pred P l := fun x (pfin:In x l) => P' x (pf _ pfin). Definition list_pred_cons {T} {P:T->Prop} {l:list T} {a:T} (P':list_pred P (a::l)) : list_pred P l := fun x (pfin:In x l) => P' x (in_cons _ _ _ pfin). Definition list_pred_app_l {T} {P:T->Prop} {l l':list T} (P':list_pred P (l++l')) : list_pred P l := fun x (pfin:In x l) => P' x (in_app_l pfin l'). Definition list_pred_app_r {T} {P:T->Prop} {l l':list T} (P':list_pred P (l++l')) : list_pred P l' := fun x (pfin:In x l') => P' x (in_app_r pfin l). Definition list_pred_type_incl {T} {l l':list T} {a:T} (P:list_pred_type l') (pf:incl l l') : list_pred_type l := fun x (pfin:In x l) => P x (pf _ pfin). Definition list_pred_type_cons {T} {l:list T} {a:T} (P:list_pred_type (a::l)) : list_pred_type l := fun x (pfin:In x l) => P x (in_cons _ _ _ pfin). Definition list_pred_type_app_sing {T} {l:list T} {a:T} (P:list_pred_type (l++a::nil)) : list_pred_type l := fun x pf => P x (in_app_l pf _). Definition list_pred_type_app_l {T} {l l':list T} (P:list_pred_type (l++l')) : list_pred_type l := fun x pf => P x (in_app_l pf _). Definition list_pred_type_app_r {T} {l l':list T} (P:list_pred_type (l++l')) : list_pred_type l' := fun x pf => P x (in_app_r pf _). Lemma list_pred_dec_cons : forall {T} {l:list T} {a:T} {P:list_pred_type (a::l)}, list_pred_dec P -> list_pred_dec (list_pred_type_cons P). intros T l a h1 h2. red in h2; red. intros x h3. specialize (h2 x (in_cons _ _ _ h3)). unfold list_pred_type_cons; auto. Qed. Lemma list_pred_dec_app_sing : forall {T} {l:list T} {a:T} {P:list_pred_type (l++a::nil)}, list_pred_dec P -> list_pred_dec (list_pred_type_app_sing P). intros T l a h1 h2. red in h2; red. intros x h3. specialize (h2 _ (in_app_l h3 _)). unfold list_pred_type_app_sing; auto. Qed. Lemma list_pred_dec_app_l : forall {T} {l l':list T} {P:list_pred_type (l++l')}, list_pred_dec P -> list_pred_dec (list_pred_type_app_l P). intros T l a h1 h2. red in h2; red. intros x h3. specialize (h2 _ (in_app_l h3 _)). unfold list_pred_type_app_l; auto. Qed. Lemma list_pred_dec_app_r : forall {T} {l l':list T} {P:list_pred_type (l++l')}, list_pred_dec P -> list_pred_dec (list_pred_type_app_r P). intros T l a h1 h2. red in h2; red. intros x h3. specialize (h2 _ (in_app_r h3 _)). unfold list_pred_type_app_r; auto. Qed. Definition seg_pred_list_pred {n} (P:nat->Prop) (pfsp:seg_pred n P) : list_pred P (seg_list n) := fun x (pf:In x (seg_list n)) => let pflt := iff1 (in_seg_list_iff _ _) pf in pfsp x pflt. Definition seg_pred_type_list_pred {n} (pfsp:seg_pred_type n) : list_pred_type (seg_list n) := fun x (pf:In x (seg_list n)) => let pflt := iff1 (in_seg_list_iff _ _) pf in pfsp x pflt. Definition seg_pred_dec_list_pred {n} (pfsp:seg_pred_type n) (pfd:seg_pred_dec pfsp) : list_pred_dec (seg_pred_type_list_pred pfsp) := fun x (pfi:In x (seg_list n)) => let pflt := iff1 (in_seg_list_iff _ _) pfi in pfd x pflt. Definition list_pred_seg_pred {n} {P:nat->Prop} (pflp:list_pred P (seg_list n)) : seg_pred n P := fun i (pf:i < n) => let pflt := iff2 (in_seg_list_iff _ _) pf in pflp i pflt. Definition list_pred_type_seg_pred {n} (pflp:list_pred_type (seg_list n)) : seg_pred_type n := fun i (pf:i < n) => let pflt := iff2 (in_seg_list_iff _ _) pf in pflp i pflt. Definition list_pred_dec_seg_pred {n} (pflp:list_pred_type (seg_list n)) (pfd:list_pred_dec pflp) : seg_pred_dec (list_pred_type_seg_pred pflp) := fun i (pfi:i < n) => let pflt := iff2 (in_seg_list_iff _ _) pfi in pfd i pflt. (*Returns first index for which [pflpt] is true, and [length l] if it's never true.*) Fixpoint first_l_pred {T} {l:list T} : forall {P:list_pred_type l} (pfd:list_pred_dec P), nat := match l with | nil => fun _ _ => 0 | a::l' => fun P pfd => match (pfd a (in_eq _ _)) with | left pfe => 0 | right pfne => let P' := list_pred_type_cons P in let pfld' := list_pred_dec_cons pfd in S (first_l_pred pfld') end end. Lemma first_l_pred_functional : forall {T} {l l':list T} {P:list_pred_type l} {P':list_pred_type l'} (pfd:list_pred_dec P) (pfd':list_pred_dec P') (pfe:l = l'), P' = transfer_dep pfe P -> first_l_pred pfd = first_l_pred pfd'. intros; subst. revert pfd pfd'. rewrite transfer_dep_eq_refl. intros; f_equal. apply list_pred_dec_irrelevance. Qed. Lemma first_l_pred_dec_type_transfer : forall {T} {l l':list T} {P:list_pred_type l} (pfe:l = l') (pfd:list_pred_dec P), first_l_pred (list_pred_dec_type_transfer pfe P pfd) = first_l_pred pfd. intros; subst; gterml. redterm0 x. fold c. apply (first_l_pred_functional c pfd (eq_refl _) (transfer_dep_eq_refl _ _ (eq_refl _) P)). Qed. Lemma le_first_l_pred : forall {T} {l:list T} {P:list_pred_type l} (pfd:list_pred_dec P), first_l_pred pfd <= length l. intros T l. induction l as [|a l ih]; simpl; auto with arith. intros h1 h2. lr_if_goal; fold hlr; destruct hlr; auto with arith. Qed. Lemma first_l_pred_eq_length_iff : forall {T} {l:list T} {P:list_pred_type l} (pfd:list_pred_dec P), first_l_pred pfd = length l <-> forall (x:T) (pf:In x l), ~P x pf. intros T l P hd; split. revert hd. revert P. induction l as [|a l ih]; simpl. intros; contradiction. intros P hd h1 x h2. lr_if h1; fold hlr in h1; destruct hlr as [hl | hr]. discriminate. destruct h2 as [|h2]; subst. contradict hr. erewrite f_equal_dep; auto. apply hr. apply S_inj in h1. specialize (ih (list_pred_type_cons P) (list_pred_dec_cons hd) h1 x h2). unfold list_pred_type_cons in ih. contradict ih. erewrite f_equal_dep; auto. apply ih. revert hd. revert P. induction l as [|a l ih]. intros; simpl; auto. intros P hd h1. simpl. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. specialize (h1 a (in_eq _ _)). contradiction. f_equal. apply ih. unfold list_pred_type_cons. intros; apply h1. Qed. Lemma lt_first_l_pred : forall {T} {l:list T} {P:list_pred_type l} (pfd:list_pred_dec P), first_l_pred pfd < length l -> exists y (pf:In y l), P y pf. intros T l. induction l as [|a l ih]; simpl. intros ? ? h1. apply lt_irrefl in h1. contradiction. intros P hd h1. pose proof (le_first_l_pred (list_pred_dec_cons hd)) as h2. apply le_lt_eq_dec in h2. lr_if h1; fold hlr in h1; destruct hlr as [hl | hr]. specialize (ih (list_pred_type_cons P) (list_pred_dec_cons hd)). destruct h2 as [h2 | h2]. specialize (ih h2). destruct ih as [y [h3 h4]]. exists y. ex_goal. right; auto. exists hex. unfold list_pred_type_cons in h4. erewrite f_equal_dep; auto. apply h4. pose proof (first_l_pred_eq_length_iff (list_pred_dec_cons hd)) as h3. pose proof h2 as h4. rewrite h3 in h4. exists a. ex_goal. left; auto. exists hex. erewrite f_equal_dep; auto. apply hl. apply lt_S_n in h1. specialize (ih (list_pred_type_cons P) (list_pred_dec_cons hd) h1). destruct ih as [x [hx h3]]. exists x. ex_goal. right; auto. exists hex. unfold list_pred_type_cons in h3. erewrite f_equal_dep; auto. apply h3. Qed. Lemma first_l_pred_cons_P : forall {T} {l:list T} {a:T} {P:list_pred_type (a::l)} (pfd:list_pred_dec P) (pfd':list_pred_dec (list_pred_type_cons P)), P a (in_eq _ _) -> first_l_pred pfd = 0. intros T l a P hd hd' h1. unfold first_l_pred. lr_if_goal; fold hlr; destruct hlr as [hl | hr]; auto. contradiction. Qed. Lemma first_l_pred_cons_nP : forall {T} {l:list T} {a:T} {P:list_pred_type (a::l)} (pfd:list_pred_dec P) (pfd':list_pred_dec (list_pred_type_cons P)), ~ P a (in_eq _ _) -> first_l_pred pfd = S (first_l_pred pfd'). intros T l a P hd hd' h1. simpl. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. contradiction. f_equal. f_equal. apply list_pred_dec_irrelevance. Qed. (*Returns the first number < n such that [pf] is true, and n if it's not true*) Definition first_seg_pred {n} {pf:seg_pred_type n} (pfs:seg_pred_dec pf) := let pf' := seg_pred_type_list_pred pf in let pfs' := seg_pred_dec_list_pred pf pfs in first_l_pred pfs'. Lemma le_first_seg_pred : forall {n} {P:seg_pred_type n} (pfd:seg_pred_dec P), first_seg_pred pfd <= n. unfold first_seg_pred. intros. eapply le_trans. apply le_first_l_pred. rewrite length_seg_list; auto with arith. Qed. Lemma eq_first_seg_pred : forall {n} {P:seg_pred_type n} (pfd:seg_pred_dec P), first_seg_pred pfd = n -> forall i (pf:i < n), ~P i pf. intros n P hd h1. pose proof @first_l_pred_eq_length_iff. unfold first_seg_pred in h1. rewrite <- (length_seg_list n) in h1. rewrite first_l_pred_eq_length_iff in h1. intros i hi. specialize (h1 i). hfirst h1. rewrite in_seg_list_iff; auto. specialize (h1 hf). contradict h1. unfold seg_pred_type_list_pred. erewrite f_equal_dep. apply h1. reflexivity. Qed. Lemma lt_first_seg_pred : forall {n} {P:seg_pred_type n} (pfd:seg_pred_dec P), first_seg_pred pfd < n -> exists i (pf:i < n), P i pf. intros n P hd h1. pose proof (lt_first_l_pred (seg_pred_dec_list_pred _ hd)) as h2. hfirst h2. rewrite length_seg_list; auto. specialize (h2 hf). destruct h2 as [i h2]. destruct h2 as [h2 h3]. pose proof h2 as h4. rewrite in_seg_list_iff in h4. exists _, h4. unfold seg_pred_type_list_pred in h3. erewrite f_equal_dep. apply h3. reflexivity. Qed. (*This is the analogue of [count_occ_prop] for the [seg] structure, whereby this function counts all elements less than [n] for which P is true.*) Fixpoint count_occ_prop_seg {n:nat} : forall {P:seg_pred_type n} (pfd:seg_pred_dec P), nat := match n with | 0 => fun _ _ => 0 | S n' => fun (P:seg_pred_type (S n')) (pfd:seg_pred_dec P) => let P' := seg_pred_type_from_S P in let pfd' := seg_pred_dec_from_S pfd in let k := count_occ_prop_seg pfd' in if (pfd n' (lt_n_Sn _)) then S k else k end. (*This function removes a point in the domain list of f, and also removes other points of the same value.*) Definition fun_in_dep_remove_in {T U:Type} {l:list T} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:fun_in_dep l U) (x:T) (pf:In x l) : fun_in_dep (list_minus pfdt l (inv_im_in_dep pfdu l (f x pf) f)) U. do 2 red in f; do 2 red. intros x' h2. pose proof (in_list_minus_impl' pfdt _ _ _ h2) as h3. refine (f _ h3). Defined. Lemma fun_in_dep_remove_in_compat : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:fun_in_dep l U) (a x:T) (pfa:In a l) (pfx':In x (list_minus pfdt l (inv_im_in_dep pfdu l (f a pfa) f))), let pfx := in_list_minus_impl' pfdt _ _ _ pfx' in let f' := fun_in_dep_remove_in pfdt pfdu f a pfa in f x pfx = f' x pfx'. unfold fun_in_dep_remove_in; intros; auto. Qed. Lemma fun_in_dep_remove_in_compat' : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:fun_in_dep l U) (a x:T) (pfa:In a l) (pfx':In x (list_minus pfdt l (inv_im_in_dep pfdu l (f a pfa) f))) (pfx:In x l), let f' := fun_in_dep_remove_in pfdt pfdu f a pfa in f x pfx = f' x pfx'. unfold fun_in_dep_remove_in; intros; auto. f_equal. apply proof_irrelevance. Qed. (*Assumes constant l to preserve functional equality.*) Lemma fun_in_dep_remove_in_functional : forall {T U:Type} {l:list T} (pfdt pfdt':type_eq_dec T) (pfdu pfdu':type_eq_dec U) (f f':fun_in_dep l U) (a x a' x':T) (pfa:In a l) (pfa':In a' l) (pfx: In x (list_minus pfdt l (inv_im_in_dep pfdu l (f a pfa) f))) (pfx':In x' (list_minus pfdt' l (inv_im_in_dep pfdu' l (f' a' pfa') f'))), f = f' -> a = a' -> x = x' -> fun_in_dep_remove_in pfdt pfdu f a pfa x pfx = fun_in_dep_remove_in pfdt' pfdu' f' a' pfa' x' pfx'. intros T U l h1 h1' h2 h2' f f' a x a' x' h3 h3' h4 h4' h5 h6 h7. subst. pose proof (type_eq_dec_eq h1 h1'). pose proof (type_eq_dec_eq h2 h2'). pose proof (proof_irrelevance _ h3 h3'); subst. pose proof (proof_irrelevance _ h4 h4'); subst. reflexivity. Qed. (*Doesn't assume [NoDup lu], so there will be duplicates in the function value, if there are duplicates in [lu].*) Fixpoint inv_im_in_dep_l {T U:Type} (pfdu:type_eq_dec U) (l:list T) (f:fun_in_dep l U) (lu:list U) : list T := match lu with | nil => nil | b::l' => (inv_im_in_dep pfdu l b f) ++ (inv_im_in_dep_l pfdu l f l') end. Lemma inv_im_in_dep_l_functional : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (f f':fun_in_dep l U) (lu lu':list U), f = f' -> lu = lu' -> inv_im_in_dep_l pfdu l f lu = inv_im_in_dep_l pfdu l f' lu'. intros; subst; auto. Qed. Lemma inv_im_in_dep_l_transfer_functional : forall {T U:Type} (pfdu:type_eq_dec U) (l l':list T) (f:fun_in_dep l U) (lu lu':list U) (pf:l = l'), lu = lu' -> let f' := transfer_fun_in_dep f pf in inv_im_in_dep_l pfdu l f lu = inv_im_in_dep_l pfdu l' f' lu'. simpl; intros; subst; rewrite transfer_fun_in_dep_eq_refl; auto. Qed. Lemma in_inv_im_in_dep_l_iff : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (f:fun_in_dep l U) (lu: list U) (x:T), In x (inv_im_in_dep_l pfdu l f lu) <-> exists pf:In x l, In (f x pf) lu. intros T U h1 l f lu x; split; intro h2. revert h2. revert x f. revert l. induction lu as [|u lu ih]; simpl. intros; contradiction. intros l x f h2. apply in_app_or in h2. destruct h2 as [h2 | h2]. rewrite in_inv_im_in_dep_iff in h2. destruct h2 as [h2]; subst. exists h2. left; auto. specialize (ih _ _ _ h2). destruct ih as [h3 h4]. exists h3. right; auto. revert h2. revert x. revert f. revert l. induction lu as [|u lu ih]; simpl. intros l h2 h3 h4. destruct h4; auto. intros l f x h2. destruct h2 as [h2 h3]. destruct h3 as [h3 | h3]. symmetry in h3. pose proof (ex_intro (fun c => f x c = u) h2 h3) as h4. rewrite <- (in_inv_im_in_dep_iff h1) in h4. apply in_app_l; auto. specialize (ih l f x). hfirst ih. exists h2; auto. specialize (ih hf). apply in_app_r; auto. Qed. Lemma inv_im_in_dep_l_incl : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (lu:list U) (f:fun_in_dep l U), incl (inv_im_in_dep_l pfdu l f lu) l. intros T U h1 l lu f. red; intros x h2. rewrite in_inv_im_in_dep_l_iff in h2. destruct h2; auto. Qed. (*See past [map_in_dep] below, for a function [fun_in_dep_list_minus_incl], analogous to [fun_in_dep_remove_in]*) Definition fun_prop_from_cons {T U:Type} {l:list T} {a:T} {P:T->Prop} (f:fun_in_dep_prop (a::l) P U) : fun_in_dep_prop l P U. do 2 red; intros ? h1; eapply f; auto; right; apply h1. Defined. Lemma fun_prop_from_cons_compat : forall {T U:Type} {l:list T} {a:T} {P:T->Prop} (f:fun_in_dep_prop (a::l) P U) (x:T) (pf:In x l) (pfp:P x), let f' := fun_prop_from_cons f in f' x pf pfp = f x (in_cons _ _ _ pf) pfp. intros; unfold f', fun_prop_from_cons; auto; f_equal; apply proof_irrelevance. Qed. Definition set_from_cons {T:Type} {U:T->Set} {l:list T} {a:T} (f:set_in_dep U (a::l)) : set_in_dep U l. do 2 red; intros ? h1; eapply f; auto; right; apply h1. Defined. Lemma set_from_cons_compat : forall {T:Type} {U:T->Set} {l:list T} {a:T} (f:set_in_dep U (a::l)) (x:T) (pf:In x l), let f' := set_from_cons f in f' x pf = f x (in_cons _ _ _ pf). intros; unfold f', set_from_cons; auto; f_equal; apply proof_irrelevance. Defined. Definition fun_from_app1 {T U:Type} {l l':list T} (f:fun_in_dep (l++l') U) : fun_in_dep l U. do 2 red; intros ? h1. eapply f. rewrite in_app_iff. left. apply h1. Defined. Definition fun_from_app2 {T U:Type} {l l':list T} (f:fun_in_dep (l++l') U) : fun_in_dep l' U. do 2 red; intros ? h1. eapply f. rewrite in_app_iff. right. apply h1. Defined. Lemma fun_from_app1_compat : forall {T U:Type} {l l':list T} (f:fun_in_dep (l++l') U) (x:T) (pf:In x l), let f' := fun_from_app1 f in f' x pf = f x (in_app_l pf l'). intros; unfold f', fun_from_app1; f_equal; apply proof_irrelevance. Qed. Lemma fun_from_app2_compat : forall {T U:Type} {l l':list T} (f:fun_in_dep (l++l') U) (x:T) (pf:In x l'), let f' := fun_from_app2 f in f' x pf = f x (in_app_r pf l). intros; unfold f', fun_from_app2; f_equal; apply proof_irrelevance. Qed. Definition fun_in_dep_incl {T U:Type} (l l':list T) (f:fun_in_dep l U) (pfi:incl l' l) : fun_in_dep l' U. do 2 red. intros x h1. do 2 red in f. apply pfi in h1. specialize (f _ h1). refine f. Defined. Lemma fun_in_dep_incl_compat : forall {T U:Type} (l l':list T) (f:fun_in_dep l U) (pfi:incl l' l), let f' := fun_in_dep_incl l l' f pfi in forall (x:T) (pf:In x l'), f' x pf = f x (pfi x pf). intros; unfold fun_in_dep_incl; auto. Qed. Lemma inj_dep_incl : forall {T U:Type} {l:list T} (f:fun_in_dep l U) (l':list T) (pf:incl l' l), let f' := fun_in_dep_incl _ _ f pf in inj_dep f -> inj_dep f'. intros T U l f l' h1 f' h2. red in h2; red. intros ? ? h3 h4. apply (h2 _ _ (h1 _ h3) (h1 _ h4)). Qed. Definition fun_in_dep_map_to_sig_set_im {T U V:Type} (g:T->U) (l:list T) (f:fun_in_dep (map g l) V) : sig_set (Im (list_to_set l) g)->V := fun x:sig_set (Im (list_to_set l) g) => let pf := proj2_sig x in let pf' := in_eq_set _ _ (eq_sym (map_im_compat g l)) (proj1_sig x) pf in let pf0 := iff2 (list_to_set_in_iff _ _) pf' in f _ pf0. Definition sig_set_im_to_fun_in_dep {T U V:Type} (g:T->U) (l:list T) (f:sig_set (Im (list_to_set l) g)->V) : fun_in_dep (map g l) V := fun (x:U) (pf:In x (map g l)) => let pf' := iff1 (list_to_set_in_iff _ _) pf in let pf0 := in_eq_set _ _ (map_im_compat g l) x pf' in f (exist _ _ pf0). Lemma fun_in_dep_map_to_seg_set_im_undoes : forall {T U V:Type} (g:T->U) (l:list T) (f:sig_set (Im (list_to_set l) g)->V), fun_in_dep_map_to_sig_set_im g l (sig_set_im_to_fun_in_dep g l f) = f. intros T U V g l f. unfold fun_in_dep_map_to_sig_set_im, sig_set_im_to_fun_in_dep. apply functional_extensionality. intro x. f_equal. apply proj1_sig_injective; auto. Qed. Lemma sig_set_im_to_fun_in_dep_undoes : forall {T U V:Type} (g:T->U) (l:list T) (f:fun_in_dep (map g l) V), (sig_set_im_to_fun_in_dep g l (fun_in_dep_map_to_sig_set_im g l f)) = f. intros T U V g l f. unfold fun_in_dep_map_to_sig_set_im, sig_set_im_to_fun_in_dep. simpl. apply functional_extensionality_dep. intro y. apply functional_extensionality. intro h1. f_equal. apply proof_irrelevance. Qed. Fixpoint map_seg {U:Type} {n:nat} : seg_fun n U -> list U := match n with | 0 => fun _ => nil | S n => fun f => (map_seg (seg_fun_from_S f)) ++ (f n (lt_n_Sn _)) ::nil end. Lemma map_seg_S : forall {U:Type} (n:nat) (pf:n < S n) (f:seg_fun (S n) U), map_seg (seg_fun_from_S f) ++ (f n pf) :: nil = map_seg f. intros; simpl; repeat f_equal; apply proof_irrelevance. Qed. Lemma map_seg_functional : forall {U:Type} {n n':nat} (pf:n = n') (f:seg_fun n U), map_seg f = map_seg (transfer_seg_fun pf f). intros U n n' h1; subst. intro f. rewrite transfer_seg_fun_eq_refl; auto. Qed. Lemma length_map_seg : forall {U:Type} {n:nat} (f:seg_fun n U), length (map_seg f) = n. intros ? n. induction n as [|n ih]; simpl; auto. intro; rewrite app_length; rewrite ih; f_equal; simpl; omega. Qed. Lemma map_seg_length_eq : forall {U V:Type} {n:nat} (f:seg_fun n U) (g:seg_fun n V), length (map_seg f) = length (map_seg g). intros; do 2 rewrite length_map_seg; auto. Qed. Lemma map_seg_eq_nil : forall {U:Type} {n:nat} (f:seg_fun n U), map_seg f = nil -> n = 0. intros U n. destruct n; simpl; auto. intros f h1. apply app_eq_nil in h1. destruct h1; discriminate. Qed. Lemma map_seg_eq_nil_iff : forall {U:Type} {n:nat} (f:seg_fun n U), map_seg f = nil <-> n = 0. intros; split; try apply map_seg_eq_nil; intro; subst; simpl; auto. Qed. Lemma in_map_seg : forall {U:Type} {n:nat} (f:seg_fun n U) (i:nat) (pf:i < n), In (f i pf) (map_seg f). intros U n. induction n as [|n ih]; simpl. intros; omega. intros f i h1. pose proof h1 as h1'. apply le_lt_eq_dec in h1'. destruct h1' as [h2 | h2]. apply lt_S_n in h2. specialize (ih (seg_fun_from_S f) i h2). apply in_app_l. unfold seg_fun_from_S in ih. erewrite f_equal_dep. apply ih. reflexivity. apply S_inj in h2; subst. apply in_app_r. constructor. f_equal; apply proof_irrelevance. Qed. Lemma in_map_seg_iff : forall {U:Type} {n:nat} (f:seg_fun n U) (b:U), In b (map_seg f) <-> exists (m:nat) (pf:m < n), f m pf = b. intros U n. induction n as [|n ih]; simpl. intros f b. split; intro h1; try contradiction. destruct h1 as [m h1]. destruct h1; omega. intros f b. split; intro h1. apply in_app_or in h1. destruct h1 as [h1 | h1]. rewrite ih in h1. destruct h1 as [m h1]. destruct h1 as [h1 h2]. subst. exists m. ex_goal. auto with arith. exists hex. unfold seg_fun_from_S. f_equal. apply proof_irrelevance. simpl in h1. destruct h1 as [h1|]; try contradiction; subst. exists n. exists (lt_n_Sn _); auto. destruct h1 as [m [h1 ?]]; subst. pose proof h1 as h3. apply le_S_n in h3. apply le_lt_eq_dec in h3. destruct h3 as [h5 | h5]; subst. apply in_app_l. rewrite ih. exists m, h5. unfold seg_fun_from_S; f_equal; apply proof_irrelevance. apply in_app_r; simpl. left. f_equal; apply proof_irrelevance. Qed. Lemma map_seg_ext : forall {U:Type} {n:nat} (f f':seg_fun n U), (forall (m:nat) (pf:m < n), f m pf = f' m pf) -> map_seg f = map_seg f'. intros U n. induction n as [|n ih]; simpl; auto. intros f f' h1. f_equal. apply ih. intros; apply h1. rewrite h1; auto. Qed. Lemma map_seg_lt : forall {U:Type} {n:nat} (f:seg_fun n U), map_seg f <> nil -> pred n < n. intros U n f h1. destruct n as [|n]. simpl in h1; contradict h1; auto. simpl; auto. Qed. Lemma map_seg_lt0 : forall {U:Type} {n:nat} (f:seg_fun n U), map_seg f <> nil -> 0 < n. intros U n f h1. destruct n as [|n]. simpl in h1; contradict h1; auto. auto with arith. Qed. Fixpoint map_seq {U:Type} {n w:nat} : seq_fun n w U -> list U := match w with | 0 => fun _ => nil | S w' => fun f => map_seq (seq_fun_from_S f) ++ f (n+w') (seq_S_w n w') :: nil end. Lemma length_map_seq : forall {U:Type} {n w:nat} (f:seq_fun n w U), length (map_seq f) = w. intros U n w. revert n. induction w as [|w ih]; simpl; auto. intros n h1. rewrite app_length. simpl. rewrite S_compat. f_equal. apply ih; auto. Qed. Lemma map_seg0_decompose : forall {U:Type} {n:nat} (f:seg_fun (S n) U), map_seg f = (f 0 (lt_0_Sn _)) :: map_seq (seg_S_tl f). intros U n. induction n as [|n ih]; simpl. intro. f_equal. f_equal. apply proof_irrelevance. intro f. specialize (ih (seg_fun_from_S f)). simpl in ih. simpl. rewrite ih. rewrite <- app_comm_cons. unfold seg_fun_from_S. f_equal. f_equal. apply proof_irrelevance. unfold seg_S_tl. f_equal. unfold seq_fun_from_S. f_equal. apply functional_extensionality_dep. intro m. apply functional_extensionality. intro. f_equal. apply proof_irrelevance. f_equal. f_equal. apply proof_irrelevance. Qed. Lemma list_to_set_map_seg : forall {U:Type} {n:nat} (f:seg_fun n U), list_to_set (map_seg f) = im_in_ens (seg_fun_to_ens f). intros U n f; apply Extensionality_Ensembles; red; split; red; intros x h1. rewrite <- list_to_set_in_iff in h1. rewrite in_map_seg_iff in h1. destruct h1 as [m [h1 h2]]; subst. apply im_in_ens_intro with m (in_seg _ _ h1). rewrite seg_fun_to_ens_compat. f_equal. apply proof_irrelevance. rewrite <- list_to_set_in_iff. destruct h1 as [i h1]; subst. destruct h1. rewrite seg_fun_to_ens_compat. apply in_map_seg. Qed. Lemma list_to_set_map_seg_ : forall {U:Type} (f:nat->U) (n:nat), list_to_set (map_seg (fun_to_seg f n)) = Im (seg n) f. intros U f n; apply Extensionality_Ensembles; red; split; red; intros x h1. rewrite <- list_to_set_in_iff in h1. rewrite in_map_seg_iff in h1. destruct h1 as [m [h1 h2]]; subst. unfold fun_to_seg. apply Im_intro with m; auto. constructor; auto. destruct h1 as [x h1]; subst. destruct h1 as [h1]. rewrite <- list_to_set_in_iff. rewrite in_map_seg_iff. exists x, h1. unfold fun_to_seg; auto. Qed. Lemma in_im_in_ens_nat : forall {U:Type} {n:nat} (f:seg_fun n U) (y:U), Inn (im_in_ens f) y -> In y (map_seg f). intros U n f y h1. destruct h1; subst. red in pf. apply in_map_seg. Qed. Lemma in_im_seg : forall {U:Type} {n:nat} (f:nat->U) (y:U), Inn (Im (seg n) f) y -> In y (map_seg (fun i (pf:i < n) => f i)). intros U n f y h1. destruct h1; subst. destruct H. rewrite in_map_seg_iff. exists x, H; auto. Qed. (*The image of a [seg_fun] in natural [list] form.*) Fixpoint im_seg_fun {n} : seg_fun n nat -> list nat := match n with | 0 => fun _ => nil | S n' => fun f => let f' := seg_fun_from_S f in (im_seg_fun f') ++ (f n' (lt_n_Sn _)) :: nil end. Lemma in_map_seq_iff : forall {U:Type} {n w:nat} (f:seq_fun n w U) (y:U), In y (map_seq f) <-> exists i (pf:n <= i < n + w), f i pf = y. intros U n w. revert n. induction w as [|w ih]; simpl. intros n f y; split; intro h1. contradiction. destruct h1 as [i [h1]]. omega. intros n f y; split; intro h1. apply in_app_or in h1. destruct h1 as [h1 | h1]. rewrite ih in h1. destruct h1 as [i [h1 h2]]. subst. destruct h1 as [h1 h2]. pose proof (lt_n_Sn (n + w)) as h3. pose proof (lt_trans _ _ _ h2 h3) as h4. rewrite <- plus_S_shift_r in h4. exists i. ex_goal. split; auto. exists hex. rewrite seq_fun_from_S_compat; f_equal; apply proof_irrelevance. simpl in h1. destruct h1 as [h1 | h1]. exists (n + w). ex_goal. omega. exists hex. rewrite <- h1. f_equal. apply proof_irrelevance. contradiction. destruct h1 as [i [h1 h2]]. subst. pose proof h1 as h2. destruct h2 as [h2 h3]. apply lt_dec in h3. destruct h3 as [h3 | h3]. rewrite plus_S_shift_r in h3. simpl in h3. apply in_app_l. rewrite ih. exists _, (conj h2 h3). rewrite seq_fun_from_S_compat. f_equal. apply proof_irrelevance. subst. apply in_app_r. constructor. apply f_equal_dep. omega. Qed. Fixpoint map_in_dep {T U:Type} {l:list T} : fun_in_dep l U ->list U := match l with | nil => fun _ => nil | a::l' => fun f => f a (in_eq a l') :: map_in_dep (fun_from_cons f) end. Lemma map_in_dep_ext : forall {T U:Type} {l:list T} (f g:fun_in_dep l U), (forall x:T, In x l -> f x = g x) -> map_in_dep f = map_in_dep g. intros T U l. induction l as [|a l ih]; simpl. intros; auto. intros f g h1. rewrite h1. f_equal. apply ih. intros x h2. specialize (h1 x (or_intror _ h2)). apply fun_from_cons_functional; auto. left; auto. Qed. (*Search for [map_in_dep_cons] for the more standard properties of the above.*) Fixpoint map_in_dep_prop_l {T U:Type} {P:T->Prop} {l:list T} : (forall x:T, In x l -> {P x} + {~ P x}) -> fun_in_dep_prop l P U -> list U := match l with | nil => fun _ _=> nil | a::l' => fun pf f => match (pf a (in_eq a l')) with | left pfp => f a (in_eq a l') pfp :: (map_in_dep_prop_l (set_from_cons pf) (fun_prop_from_cons f)) | right _ => map_in_dep_prop_l (set_from_cons pf) (fun_prop_from_cons f) end end. Fixpoint map_in_dep_prop_r {T U:Type} {P:T->Prop} {l:list T} : (forall x:T, In x l -> {P x} + {~P x}) -> fun_in_dep_prop l (not_pred P) U -> list U := match l with | nil => fun _ _=> nil | a::l' => fun pf f => match (pf a (in_eq a l')) with | left _ => map_in_dep_prop_r (set_from_cons pf) (fun_prop_from_cons f) | right pfq => f a (in_eq a l') pfq :: (map_in_dep_prop_r (set_from_cons pf) (fun_prop_from_cons f)) end end. (*Use this when all elements in [l] satisfy [P]*) Definition fun_in_dep_prop_all {T U:Type} {P:T->Prop} {l:list T} (f:fun_in_dep_prop l P U) (pfl:forall x:T, In x l -> P x) : fun_in_dep l U := fun (x:T) (pf:In x l) => f x pf (pfl x pf). Lemma map_in_dep_prop_l_ext : forall {T U:Type} {P:T->Prop} {l:list T} (pf pf':lexcl_mid l P) (f g:fun_in_dep_prop l P U), (forall x:T, In x l -> P x -> f x = g x) -> map_in_dep_prop_l pf f = map_in_dep_prop_l pf' g. intros T U P l h0. induction l as [|a l ih]; simpl; auto. intros h1 f g h2. destruct (h0 a (in_eq a l)) as [h3 | h3], (h1 a (in_eq a l)) as [h3' | h3']; try contradiction. rewrite h2. f_equal. f_equal. apply proof_irrelevance. f_equal. f_equal. simpl in h0. apply prop_dep_set_dep_ext. apply unit_set_pred_fexcl_mid. apply functional_extensionality_dep. intro x. apply functional_extensionality. intro h4. apply functional_extensionality. intro h5. do 2 rewrite fun_prop_from_cons_compat. rewrite h2. f_equal. right; auto. assumption. left; auto. assumption. erewrite ih. reflexivity. intros x h4 h5. apply functional_extensionality. intro h6. apply functional_extensionality. intro h7. do 2 rewrite fun_prop_from_cons_compat. rewrite h2; auto. Qed. Lemma map_in_dep_prop_l_all : forall {T U:Type} {P:T->Prop} {l:list T} (pf:forall x:T, In x l -> {P x} + {~ P x}) (f:fun_in_dep_prop l P U) (pfl:forall x:T, In x l -> P x), let f' := fun_in_dep_prop_all f pfl in map_in_dep_prop_l pf f = map_in_dep f'. intros T U P l. simpl. induction l as [|a l ih]; simpl; auto. intros h1 f h2. lr_match_goal. fold hlr. destruct hlr as [h3 | h3]. unfold fun_in_dep_prop_all. f_equal. f_equal. apply proof_irrelevance. hfirst ih. intros x h4. apply h1. right; auto. specialize (ih hf (fun_prop_from_cons f)). hfirst ih. intros; apply h2. right; auto. specialize (ih hf0). unfold set_from_cons. gterml. terml ih. assert (h4:x = x0). unfold x, x0. apply map_in_dep_prop_l_ext. intros; auto. fold x. rewrite h4. unfold x0. rewrite ih. apply map_in_dep_ext. intros c h5. unfold fun_in_dep_prop_all, fun_from_cons. apply functional_extensionality. intros h6. unfold fun_prop_from_cons. f_equal; apply proof_irrelevance. pose proof (h2 a (in_eq _ _)); contradiction. Qed. Lemma map_in_dep_prop_l_none : forall {T U:Type} {P:T->Prop} {l:list T} (pf:forall x:T, In x l -> {P x} + {~ P x}) (f:fun_in_dep_prop l P U), (forall x:T, In x l -> ~ P x) -> map_in_dep_prop_l pf f = nil. intros T U P l. induction l as [|a l ih]; simpl; auto. intros h1 f h2. destruct (h1 a (in_eq a l)) as [h3 | h3]. contradict h3. apply h2; auto. apply ih; auto. Qed. Lemma map_in_dep_prop_r_ext : forall {T U:Type} {P:T->Prop} {l:list T} (pf pf':lexcl_mid l P) (f g:fun_in_dep_prop l (not_pred P) U), (forall x:T, In x l -> ~ P x -> f x = g x) -> map_in_dep_prop_r pf f = map_in_dep_prop_r pf' g. intros T U P l h0. induction l as [|a l ih]; simpl; auto. intros h1 f g h2. destruct (h0 a (in_eq a l)) as [h3 | h3], (h1 a (in_eq a l)) as [h3' | h3']; try contradiction. f_equal. f_equal. apply prop_dep_set_dep_ext. apply unit_set_pred_fexcl_mid. apply functional_extensionality_dep. intro x. apply functional_extensionality. intro h4. apply functional_extensionality. intro h5. simpl in h5. do 2 rewrite fun_prop_from_cons_compat. rewrite h2. f_equal. right; auto. assumption. rewrite h2. f_equal. f_equal. apply functional_extensionality_dep. intro. contradiction. erewrite ih. reflexivity. intros x h4 h5. apply functional_extensionality. intro h6. apply functional_extensionality. intro h7. do 2 rewrite fun_prop_from_cons_compat. rewrite h2; auto. left; auto. assumption. Qed. Lemma map_in_dep_prop_r_all : forall {T U:Type} {P:T->Prop} {l:list T} (pf:forall x:T, In x l -> {P x} + {~ P x}) (f:fun_in_dep_prop l (not_pred P) U) (pfl:forall x:T, In x l -> ~ P x), let f' := fun_in_dep_prop_all f pfl in map_in_dep_prop_r pf f = map_in_dep f'. intros T U P l. simpl. induction l as [|a l ih]; simpl; auto. intros h1 f h2. lr_match_goal. fold hlr. destruct hlr as [h3 | h3]. contradict h3. apply h2. left; auto. unfold fun_in_dep_prop_all. f_equal. f_equal. apply proof_irrelevance. hfirst ih. intros x h4. apply h1. right; auto. specialize (ih hf (fun_prop_from_cons f)). hfirst ih. intros; apply h2. right; auto. specialize (ih hf0). unfold set_from_cons. gterml. terml ih. assert (h4:x = x0). unfold x, x0. apply map_in_dep_prop_r_ext. intros; auto. fold x. rewrite h4. unfold x0. rewrite ih. apply map_in_dep_ext. intros c h5. unfold fun_in_dep_prop_all, fun_from_cons. apply functional_extensionality. intros h6. unfold fun_prop_from_cons. f_equal; apply proof_irrelevance. Qed. Lemma map_in_dep_prop_r_none : forall {T U:Type} {P:T->Prop} {l:list T} (pf:forall x:T, In x l -> {P x} + {~ P x}) (f:fun_in_dep_prop l (not_pred P) U), (forall x:T, In x l -> P x) -> map_in_dep_prop_r pf f = nil. intros T U P l. induction l as [|a l ih]; simpl; auto. intros h1 f h2. destruct (h1 a (in_eq a l)) as [h3 | h3]. apply ih; auto. contradict h3. apply h2; auto. Qed. Lemma in_map_in_dep_prop_l : forall {T U:Type} {P:T->Prop} {l:list T} (pfd:(forall x:T, In x l -> {P x} + {~P x})) (f:fun_in_dep_prop l P U) (x:T) (pfin:In x l) (pfp:P x), In (f x pfin pfp) (map_in_dep_prop_l pfd f). intros T U P l. induction l as [|a l ih]; simpl. intros; contradiction. intros h1 f x h3 h4. destruct h3 as [h3 | h3]; subst. destruct (h1 x (in_eq x l)) as [h5 | h5]. simpl. left. f_equal. apply proof_irrelevance. apply proof_irrelevance. contradiction. specialize (ih (set_from_cons h1) (fun_prop_from_cons f)). destruct (h1 a (in_eq a l)) as [h5 | h5]. right. apply ih. apply ih. Qed. Lemma in_map_in_dep_prop_r : forall {T U:Type} {P:T->Prop} {l:list T} (pfd:(forall x:T, In x l -> {P x} + {~P x})) (f:fun_in_dep_prop l (fun x => ~P x) U) (x:T) (pfin:In x l) (pfp:~P x), In (f x pfin pfp) (map_in_dep_prop_r pfd f). intros T U P l. induction l as [|a l ih]; simpl. intros; contradiction. intros h1 f x h3 h4. destruct h3 as [h3 | h3]; subst. destruct (h1 x (in_eq x l)) as [h5 | h5]. simpl. contradiction. left. f_equal. apply proof_irrelevance. apply proof_irrelevance. specialize (ih (set_from_cons h1) (fun_prop_from_cons f)). destruct (h1 a (in_eq a l)) as [h5 | h5]. apply ih. right. apply ih. Qed. Lemma in_map_in_dep_prop_l_iff : forall {T U:Type} {P:T->Prop} {l:list T} (pfd:(forall x:T, In x l -> {P x} + {~P x})) (f:fun_in_dep_prop l P U) (y:U), In y (map_in_dep_prop_l pfd f) <-> (exists (x:T) (pf:In x l) (pfp:P x), f x pf pfp = y). intros T U P l. induction l as [|a l h1]; simpl. intros. split. intro; contradiction. intro h1. destruct h1 as [? [ ]]. contradiction. split. intro h2. destruct (pfd a (in_eq a l)) as [h3 | h3]. destruct h2 as [h2 | h2]; subst. exists a. exists (or_introl eq_refl). exists h3. f_equal. apply proof_irrelevance. specialize (h1 (set_from_cons pfd) (fun_prop_from_cons f) y). rewrite h1 in h2. destruct h2 as [x h2]. destruct h2 as [h2 h4]. destruct h4 as [h4 h5]; subst. unfold fun_prop_from_cons. exists x. exists (or_intror h2). exists h4. f_equal. specialize (h1 (set_from_cons pfd) (fun_prop_from_cons f) y). rewrite h1 in h2. destruct h2 as [x h2]. destruct h2 as [h2 h4]. destruct h4 as [h4 h5]; subst. unfold fun_prop_from_cons. exists x. exists (or_intror h2). exists h4. f_equal. simpl. intro h2. destruct h2 as [x h2]. destruct h2 as [h2 h3]. destruct h3 as [h3 h4]; subst. destruct h2 as [h2 | h2]; subst. destruct (pfd x (in_eq x l)) as [h5 | h5]. simpl. left. f_equal. apply proof_irrelevance. apply proof_irrelevance. contradiction. destruct (pfd a (in_eq a l)) as [h4 | h4]. right. specialize (h1 (set_from_cons pfd) (fun_prop_from_cons f) (f x (or_intror h2) h3)). rewrite h1. exists x, h2, h3. unfold fun_prop_from_cons. f_equal. specialize (h1 (set_from_cons pfd) (fun_prop_from_cons f) (f x (or_intror h2) h3)). rewrite h1. exists x, h2, h3. unfold fun_prop_from_cons. f_equal. Qed. Lemma in_map_in_dep_prop_r_iff : forall {T U:Type} {P:T->Prop} {l:list T} (pfd:(forall x:T, In x l -> {P x} + {~P x})) (f:fun_in_dep_prop l (not_pred P) U) (y:U), In y (map_in_dep_prop_r pfd f) <-> (exists (x:T) (pf:In x l) (pfp:~ P x), f x pf pfp = y). intros T U P l. induction l as [|a l h1]; simpl. intros. split. intro; contradiction. intro h1. destruct h1 as [? [ ]]. contradiction. split. intro h2. destruct (pfd a (in_eq a l)) as [h3 | h3]. specialize (h1 (set_from_cons pfd) (fun_prop_from_cons f) y). rewrite h1 in h2. destruct h2 as [x h2]. destruct h2 as [h2 h4]. destruct h4 as [h4 h5]; subst. unfold fun_prop_from_cons. exists x. exists (or_intror h2). exists h4. f_equal. destruct h2 as [h2 | h2]; subst. exists a, (or_introl eq_refl), h3. f_equal. apply proof_irrelevance. specialize (h1 (set_from_cons pfd) (fun_prop_from_cons f) y). rewrite h1 in h2. destruct h2 as [x [h2 [h4 h5]]]. exists x, (or_intror h2), h4. rewrite <- h5 at 1. unfold fun_prop_from_cons. f_equal. intro h2. destruct h2 as [x h2]. destruct h2 as [h2 h3]. destruct h3 as [h3 h4]; subst. destruct h2 as [h2 | h2]; subst. destruct (pfd x (in_eq x l)) as [h4 | h4]. contradiction. left. f_equal. apply proof_irrelevance. apply proof_irrelevance. destruct (pfd a (in_eq a l)) as [h4 | h4]. specialize (h1 (set_from_cons pfd) (fun_prop_from_cons f) (f x (or_intror h2) h3)). rewrite h1. exists x, h2, h3. unfold fun_prop_from_cons. f_equal. right. specialize (h1 (set_from_cons pfd) (fun_prop_from_cons f) (f x (or_intror h2) h3)). rewrite h1. exists x, h2, h3. unfold fun_prop_from_cons. f_equal. Qed. Lemma map_in_dep_cons : forall {T U:Type} {l:list T} (a:T) (f:fun_in_dep (a::l) U), map_in_dep f = (f a (in_eq a l)) :: map_in_dep (fun_from_cons f). intros; auto. Qed. Lemma map_in_dep_functional : forall {T U:Type} {l l':list T} (f:fun_in_dep l U) (pf: l = l'), map_in_dep f = map_in_dep (transfer_fun_in_dep f pf). intros; subst; rewrite transfer_fun_in_dep_eq_refl; auto. Qed. Lemma map_in_dep_sing : forall {T U:Type} {a:T} (f:fun_in_dep (a::nil) U), map_in_dep f = f a (in_eq _ _) :: nil. intros; auto. Qed. Lemma in_map_in_dep : forall {T U:Type} {l:list T} (f:forall x:T, In x l->U) (x:T) (pf:In x l), In (f x pf) (map_in_dep f). intros T U l. induction l as [|a l h1]. intros; contradiction. intros f x h2. destruct h2 as [h2 | h2]; subst. unfold map_in_dep. simpl. left. f_equal. apply proof_irrelevance. specialize (h1 (fun_from_cons f) x h2). rewrite fun_from_cons_compat in h1. simpl. right. assert (h3:in_cons a x l h2 = or_intror h2). apply proof_irrelevance. rewrite <- h3; auto. Qed. Lemma in_map_in_dep_iff : forall {T U:Type} {l:list T} (f:forall x:T, In x l->U) (y:U), In y (map_in_dep f) <-> exists (x:T) (pf:In x l), f x pf = y. intros T U l. induction l as [|a l ih]; simpl. intros. split. intro; contradiction. intro h1. destruct h1 as [x h1]. destruct h1; contradiction. intros f y. specialize (ih (fun_from_cons f) y). split. intros h1. destruct h1 as [h1 | h1]. subst. exists a. exists (or_introl eq_refl). f_equal. apply proof_irrelevance. rewrite ih in h1. clear ih. destruct h1 as [x [h1 h2]]. subst. exists x. exists (or_intror h1). rewrite fun_from_cons_compat. f_equal. apply proof_irrelevance. intro h1. destruct h1 as [x h1]. destruct h1 as [h1 h2]. subst. destruct h1 as [h1|h1]. subst. left. f_equal. apply proof_irrelevance. destruct ih as [h2 h3]. hfirst h3. exists x, h1. rewrite fun_from_cons_compat. f_equal. apply proof_irrelevance. specialize (h3 hf). clear hf. right. assumption. Qed. Lemma length_map_in_dep : forall {T U:Type} {l:list T} (f:fun_in_dep l U), length (map_in_dep f) = length l. intros ? ? l; induction l; simpl; auto. Qed. Lemma map_in_dep_eq_nil : forall {T U:Type} {l:list T} (f:fun_in_dep l U), map_in_dep f = nil -> l = nil. intros T U l. destruct l as [|a l]; auto. intros f h1. pose proof (in_map_in_dep f a (in_eq a l)) as h2. rewrite h1 in h2. contradiction. Qed. Lemma map_in_dep_eq_nil_iff : forall {T U:Type} {l:list T} (f:fun_in_dep l U), map_in_dep f = nil <-> l = nil. intros; split; intro h1; try apply map_in_dep_eq_nil in h1; auto; subst; auto. Qed. Lemma map_in_dep_id : forall {T:Type} (l:list T), map_in_dep (id_in_dep l) = l. intros T l; induction l; simpl; auto. rewrite id_in_dep_compat. f_equal; auto. Qed. Lemma map_const_in_dep : forall {T U:Type} (l:list T) (val:U), map_in_dep (const_in_dep l val) = lrep val (length l). intros T U l. induction l as [|a l ih]; simpl; auto. intro y. unfold const_in_dep. f_equal. apply ih. Qed. Lemma map_map_in_dep_ext : forall {T U:Type} (l:list T) (f:T->U) (g:fun_in_dep l U), (forall (x:T) (pf:In x l), f x = g x pf) -> map f l = map_in_dep g. intros T U l. induction l as [|a l ih]; simpl; auto. intros f g h1. pose proof (h1 a (or_introl _ eq_refl)) as h2. rewrite h2. f_equal. f_equal. apply proof_irrelevance. apply ih. intros; apply h1. Qed. (*See past defintiion of [nth_lt] for a similar lemma [map_in_dep_eq_iff]*) Lemma map_in_dep_app : forall {T U:Type} {l l':list T} (f:fun_in_dep (l++l') U), map_in_dep f = (map_in_dep (fun_from_app1 f)) ++ (map_in_dep (fun_from_app2 f)). intros T U. induction l as [|a l h1]; simpl. intros l' f. unfold fun_from_app2. f_equal. apply fun_in_dep_ext. intros. f_equal. apply proof_irrelevance. intros l' f. rewrite fun_from_app1_compat. f_equal. f_equal. apply proof_irrelevance. rewrite h1. f_equal. f_equal. apply fun_in_dep_ext. intros. unfold fun_from_cons. do 2 rewrite fun_from_app1_compat. f_equal. apply proof_irrelevance. specialize (h1 _ (fun_from_cons f)). unfold fun_from_app2. f_equal. apply fun_in_dep_ext. intros. unfold fun_from_cons. f_equal. apply proof_irrelevance. Qed. Lemma map_in_dep_map : forall {T U V:Type} (la:list T) (f:T->U) (g:fun_in_dep (map f la) V), map_in_dep (fun (b:U) (pfb:In b (map f la)) => g b pfb) = map_in_dep (fun (a:T) (pfa:In a la) => g (f a) (in_map _ _ _ pfa)). intros T U V la. induction la as [|a la ih]; simpl; auto. intros f g. f_equal. f_equal. apply proof_irrelevance. specialize (ih f (fun_from_cons g)). let P := type of ih in match P with ?x = _ => match goal with |- ?y = _ => assert (h1:x = y) end end. apply map_in_dep_ext. intros b h1. apply functional_extensionality. intros. unfold fun_from_cons. f_equal. rewrite h1 in ih. rewrite ih. apply map_in_dep_ext. intros x h3. apply functional_extensionality. intro h4. unfold fun_from_cons. f_equal. apply proof_irrelevance. Qed. Lemma map_map_in_dep : forall {T U V:Type} (la:list T) (f:fun_in_dep la U) (g:U->V), map g (map_in_dep f) = map_in_dep (fun (a:T) (pfa:In a la) => g (f a pfa)). intros T U V la. induction la as [|a la ih]; simpl; auto. intros f g. f_equal. specialize (ih (fun_from_cons f) g). rewrite ih. unfold fun_from_cons; auto. Qed. Lemma map_in_dep_map_in_dep : forall {T U V:Type} (la:list T) (f:fun_in_dep la U) (g:fun_in_dep (map_in_dep f) V), map_in_dep (fun (b:U) (pfb:In b (map_in_dep f)) => g b pfb) = map_in_dep (fun (a:T) (pfa:In a la) => g (f a pfa) (in_map_in_dep _ _ pfa)). intros T U V la. induction la as [|a la ih]; simpl; auto. intros f g. f_equal. f_equal. apply proof_irrelevance. specialize (ih (fun_from_cons f) (fun_from_cons g)). unfold fun_from_cons. unfold fun_from_cons in ih. rewrite ih. apply map_in_dep_ext. intros b h1. apply functional_extensionality. intros; f_equal. apply proof_irrelevance. Qed. Lemma map_eq_map_in_dep : forall {T U:Type} (l:list T) (f:T->U), map f l = map_in_dep (fun (x:T) (_:In x l) => f x). intros ? ? l. induction l as [|a l h1]; simpl; auto. intro; f_equal. auto. unfold fun_from_cons; auto. Qed. Lemma hd_map_in_dep' : forall {T U:Type} {l:list T} (f:fun_in_dep l U) (pf:map_in_dep f <> nil), hd' (map_in_dep f) pf = f (hd' l (iffn1 (map_in_dep_eq_nil_iff f) pf)) (in_hd' _ _). intros T U l. destruct l as [|a l]; simpl. intros ? h1; contradict h1; auto. intros; f_equal; apply proof_irrelevance. Qed. Lemma list_to_set_map_in_dep : forall {T U:Type} {l:list T} (f:fun_in_dep l U), list_to_set (map_in_dep f) = im_in_ens (fun i (pf:Inn (list_to_set l) i) => let pf' := iff2 (list_to_set_in_iff _ _) pf in f i pf'). intros T U l f. apply Extensionality_Ensembles; red; split; red; intros y h1. rewrite <- list_to_set_in_iff in h1. rewrite in_map_in_dep_iff in h1. destruct h1 as [x [h1 h2]]; subst. pose proof h1 as h1'. rewrite list_to_set_in_iff in h1'. apply im_in_ens_intro with x h1'. f_equal. apply proof_irrelevance. destruct h1 as [y h1]; subst. rewrite <- list_to_set_in_iff. apply in_map_in_dep. Qed. Definition fun_in_dep_decompose1 {T U:Type} (pfdt:type_eq_dec T) (xl l:list T) (pfincl:incl xl l) (f:fun_in_dep l U) : fun_in_dep xl U := fun x (pfi:In x xl) => f x (pfincl _ pfi). Definition fun_in_dep_decompose2 {T U:Type} (pfdt:type_eq_dec T) (xl l:list T) (pfincl:incl xl l) (f:fun_in_dep l U) : fun_in_dep (list_minus pfdt l xl) U := fun x (pfi:In x (list_minus pfdt l xl)) => let (pfi', _) := in_list_minus_impl _ _ _ _ pfi in f x pfi'. (*These use [sub_list] instead of [incl]*) Definition fun_in_dep_decompose1' {T U:Type} (pfdt:type_eq_dec T) (xl l:list T) (pf:sub_list xl l) (f:fun_in_dep l U) : fun_in_dep xl U := fun_in_dep_decompose1 pfdt xl l (sub_list_incl _ _ pf) f. Definition fun_in_dep_decompose2' {T U:Type} (pfdt:type_eq_dec T) (xl l:list T) (pf:sub_list xl l) (f:fun_in_dep l U) : fun_in_dep (list_minus pfdt l xl) U := fun_in_dep_decompose2 pfdt xl l (sub_list_incl _ _ pf) f. Lemma list_to_set_map_in_dep_decompose : forall {T U:Type} (pfdt:type_eq_dec T) (l l':list T) (f:fun_in_dep l' U) (pf:incl l l'), let f1 := fun_in_dep_decompose1 pfdt l l' pf f in let f2 := fun_in_dep_decompose2 pfdt l l' pf f in list_to_set (map_in_dep f) = Union (list_to_set (map_in_dep f1)) (list_to_set (map_in_dep f2)). intros T U h0 l l' f h1. simpl. apply Extensionality_Ensembles; red; split; red; intros y h2. rewrite <- list_to_set_in_iff in h2. rewrite in_map_in_dep_iff in h2. destruct h2 as [x [h2 h3]]; subst. destruct (in_dec h0 x l) as [h3 | h3]. left. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. exists x, h3. unfold fun_in_dep_decompose1; f_equal; apply proof_irrelevance. right. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. exists x, (impl_in_list_minus _ _ _ _ h2 h3). unfold fun_in_dep_decompose2; simpl. destruct (in_list_minus_impl h0 l' l x (impl_in_list_minus _ _ _ _ h2 h3)). f_equal. apply proof_irrelevance. destruct h2 as [y h2 | y h2]. rewrite <- list_to_set_in_iff in h2. rewrite in_map_in_dep_iff in h2. destruct h2 as [x [h2 h3]]. subst. rewrite <- list_to_set_in_iff. unfold fun_in_dep_decompose1. rewrite in_map_in_dep_iff. exists x, (h1 x h2); auto. rewrite <- list_to_set_in_iff in h2. rewrite in_map_in_dep_iff in h2. destruct h2 as [x [h2 h3]]. subst. rewrite <- list_to_set_in_iff. unfold fun_in_dep_decompose2. destruct (in_list_minus_impl _ _ _ _ h2). apply in_map_in_dep_iff. exists x, i; auto. Qed. (*This is kind of weak, but the next lemma strengthens it.*) Lemma list_to_set_map_in_dep_remove_inj : forall {T U:Type} (pfdt:type_eq_dec T) (l:list T) (f:fun_in_dep l U) (x:T) (pf:In x l), inj_dep f -> list_to_set (map_in_dep (fun_in_dep_remove pfdt f x)) = Subtract (list_to_set (map_in_dep f)) (f x pf). intros T U h1 l f x h2 hi. apply Extensionality_Ensembles; red; split; red; intros y h3. rewrite <- list_to_set_in_iff in h3. rewrite in_map_in_dep_iff in h3. destruct h3 as [a [h3 h4]]; subst. pose proof h3 as h3'. rewrite in_remove_iff in h3'. destruct h3' as [h4 h5]. constructor. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. exists a, h4. rewrite fun_in_dep_remove_compat. f_equal. apply proof_irrelevance. rewrite in_sing_iff. intro h6. rewrite fun_in_dep_remove_compat in h6. red in hi. specialize (hi x a h2 (in_remove_impl h1 l a x h3) h6). subst. contradict h5; auto. destruct h3 as [h3 h4]. rewrite in_sing_iff in h4. rewrite <- list_to_set_in_iff in h3. rewrite in_map_in_dep_iff in h3. destruct h3 as [a [h3 h5]]; subst. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. exists a. ex_goal. rewrite in_remove_iff. split; auto. intro; subst. contradict h4. f_equal. apply proof_irrelevance. exists hex. rewrite fun_in_dep_remove_compat. f_equal. apply proof_irrelevance. Qed. Lemma list_to_set_map_in_dep_remove_in : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (f:fun_in_dep l U) (x:T) (pf:In x l), list_to_set (map_in_dep (fun_in_dep_remove_in pfdt pfdu f x pf)) = Subtract (list_to_set (map_in_dep f)) (f x pf). intros T U h1 h2 l f a h3. apply Extensionality_Ensembles; red; split; red; intros y h4. rewrite <- list_to_set_in_iff in h4. rewrite in_map_in_dep_iff in h4. destruct h4 as [x [h4 h5]]. pose proof (in_list_minus_impl h1 _ _ _ h4) as h6. destruct h6 as [h6 h7]. constructor. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. exists x, h6. subst. rewrite <- fun_in_dep_remove_in_compat. f_equal. apply proof_irrelevance. contradict h7. rewrite in_sing_iff in h7; subst. erewrite <- fun_in_dep_remove_in_compat in h7. rewrite in_inv_im_in_dep_iff. exists h6. rewrite h7. f_equal. apply proof_irrelevance. destruct h4 as [h4 h5]. rewrite <- list_to_set_in_iff in h4. rewrite in_map_in_dep_iff in h4. destruct h4 as [x [h4 h6]]. subst. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. exists x. ex_goal. eapply impl_in_list_minus; auto. contradict h5. rewrite in_sing_iff. rewrite in_inv_im_in_dep_iff in h5. destruct h5 as [h5 h6]. pose proof (proof_irrelevance _ h4 h5); subst. rewrite h6; auto. exists hex. erewrite <- fun_in_dep_remove_in_compat'. reflexivity. Qed. Lemma list_to_set_map_in_dep_decompose_sumbool : forall {T U:Type} (l:list T) (Q:T->Prop) (pfs:forall x:T, In x l -> {Q x} + {~Q x}) (fq:fun_in_dep_prop l Q U) (fr:fun_in_dep_prop l (fun x => ~ Q x) U), let f := fun (x:T) (pf:In x l) => match (pfs x pf) with | left pfq => fq x pf pfq | right pfr => fr x pf pfr end in list_to_set (map_in_dep f) = Union (list_to_set (map_in_dep_prop_l pfs fq)) (list_to_set (map_in_dep_prop_r pfs fr)). intros T U l Q h1 fq fr. simpl. apply Extensionality_Ensembles; red; split; red; intros y h2. rewrite <- list_to_set_in_iff in h2. rewrite in_map_in_dep_iff in h2. destruct h2 as [x h2]. destruct h2 as [h2 h3]; subst. destruct (h1 x h2) as [h4 | h4]; subst. left. rewrite <- list_to_set_in_iff. apply in_map_in_dep_prop_l. right. rewrite <- list_to_set_in_iff. apply in_map_in_dep_prop_r. destruct h2 as [y h2 | y h2]. rewrite <- list_to_set_in_iff in h2. rewrite in_map_in_dep_prop_l_iff in h2. destruct h2 as [x [h2 [h3 h4]]]; subst. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. exists x, h2. destruct (h1 x h2) as [h4 | h4]. f_equal. apply proof_irrelevance. contradiction. rewrite <- list_to_set_in_iff in h2. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_prop_r_iff in h2. destruct h2 as [x [h2 [h3 h4]]]; subst. rewrite in_map_in_dep_iff. exists x, h2. destruct (h1 x h2) as [h4 | h4]. contradiction. f_equal. apply proof_irrelevance. Qed. Lemma lS_map_in_dep : forall {T:Type} (l:list T) (f:fun_in_dep l nat), lS (map_in_dep f) = map_in_dep (fun c pf => S (f c pf)). intros T l. induction l as [|a l ih]; simpl; auto. intro; f_equal; apply ih. Qed. (*Returns a list of all elemnts in [l] which satisfy [P]*) Fixpoint map_filter {T U:Type} (P:U -> Prop) (f:T->U) (pfp:forall u:U, {P u} + {~P u}) (l:list T) : list T := match l with | nil => nil | a::l' => let ml := map_filter P f pfp l' in if (pfp (f a) ) then a :: ml else ml end. Fixpoint map_filter_not {T U:Type} (P:U -> Prop) (f:T->U) (pfr:forall u:U, {P u} + {~P u}) (l:list T) : list T := match l with | nil => nil | a::l' => let ml := map_filter P f pfr l' in if (pfr (f a) ) then ml else a :: ml end. (*Filters out all elements whose [f] values don't match [u]*) Definition map_filter_eq {T U:Type} (pfdu:type_eq_dec U) (f:T->U) (u:U) (l:list T) := map_filter (fun c => u = c) f (pfdu u) l. Definition map_filter_lt {T U:Type} (pfdu:type_eq_dec U) (R:Relation U) (pfrlem:Rlem_dec R) (f:T->U) (u:U) (l:list T) := map_filter (fun c => R c u) f (fun c => (pfrlem c u)) l. Definition map_filter_gt {T U:Type} (pfdu:type_eq_dec U) (R:Relation U) (pfrlem:Rlem_dec R) (f:T->U) (u:U) (l:list T) := map_filter (fun c => R u c) f (pfrlem u) l. Definition map_filter_le {T U:Type} (pfdu:type_eq_dec U) (R:Relation U) (pfrlem:Rlem_dec R) (f:T->U) (u:U) (l:list T) := map_filter (fun c => R c u) f (fun c => (pfrlem c u)) l. Definition map_filter_ge {T U:Type} (pfdu:type_eq_dec U) (R:Relation U) (pfrwlem:Rwlem_dec R) (f:T->U) (u:U) (l:list T) := map_filter (fun c => Rw R u c) f (pfrwlem u) l. Lemma map_filter_eq_nil_iff : forall {T U:Type} (P:U -> Prop) (f:T->U) (pfp:forall u:U, {P u} + {~P u}) (l:list T), map_filter P f pfp l = nil <-> l = nil \/ forall x:T, In x l -> ~P (f x). intros T U P f h1 l. revert h1. revert f P. induction l as [|c l ih]; simpl. intuition. intros f P h1; split; intro h2. destruct (h1 (f c)) as [h6 | h6]. discriminate. right. intros x h4. destruct h4 as [|h4]; subst; auto. specialize (ih f P h1). rewrite ih in h2. destruct h2 as [| h2]; subst. contradiction. apply h2; auto. destruct h2 as [h2 | h2]. discriminate. destruct (h1 (f c)) as [h6 | h6]. specialize (h2 c (or_introl eq_refl)). contradiction. rewrite ih. right; auto. Qed. Lemma inv_im_in_dep_l_cons_eq_nil_impl : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (a:T) (f:fun_in_dep (a::l) U) (lu:list U), inv_im_in_dep_l pfdu (a::l) f lu = nil -> inv_im_in_dep_l pfdu l (fun_from_cons f) lu = nil. intros T U h1 l a f lu. revert f. revert a l h1. induction lu as [|b lu ih]; simpl. intros; auto. intros a l h1 v h2. destruct (h1 b (v a (in_eq _ _))) as [h3 | h3]; subst. apply app_eq_nil in h2. destruct h2 as [h2 h3]. discriminate. apply app_eq_nil in h2. destruct h2 as [h2 h4]. rewrite h2, ih; auto. Qed. Lemma inv_im_in_dep_l_eq_nil : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (f:fun_in_dep l U) (lu:list U), inv_im_in_dep_l pfdu l f lu = nil -> l = nil \/ lu = nil \/ ~incl (map_in_dep f) lu. intros T U h1 l f lu. revert f. revert l h1. induction lu as [|b lu ih]; simpl. intros; right; left; auto. intros l h1 f h2. apply app_eq_nil in h2. destruct h2 as [h2 h3]. apply inv_im_in_dep_eq_nil in h2. destruct h2 as [h2 | h2]. left; auto. destruct h2 as [a [h4 h5]]. right; right. specialize (ih _ _ _ h3). destruct ih as [ih | [ih | ih]]. pose proof h4 as h4'. rewrite ih in h4'. contradiction. subst. simpl in h3. intro h6. red in h6. specialize (h6 _ (in_map_in_dep _ _ h4)). simpl in h6. destruct h6; subst. pose proof (h5 eq_refl); auto. contradiction. intro h7. red in h7. specialize (h7 _ (in_map_in_dep _ _ h4)). inversion h7 as [h10 | h10]; subst. contradict h5; auto. pose proof (in_inv_im_in_dep_l_iff h1 l f lu a) as h8. let P := type of h8 in match P with _ <-> ?pf => assert (h11:pf) end. exists h4; auto. rewrite <- h8 in h11. rewrite h3 in h11. contradiction. Qed. Definition fun_in_dep_list_minus_incl {T U:Type} {l:list T} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:fun_in_dep l U) (lx: list T) (pf:incl lx l) : fun_in_dep (list_minus pfdt l (inv_im_in_dep_l pfdu l f (map_in_dep (fun_in_dep_incl _ _ f pf)))) U. do 2 red; do 2 red in f; intros x h1. pose proof (in_list_minus_impl' pfdt _ _ _ h1) as h2. refine (f _ h2). Defined. Lemma fun_in_dep_list_minus_incl_compat : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:fun_in_dep l U) (la: list T) (pfa:incl la l) (x:T) (pfx':In x (list_minus pfdt l (inv_im_in_dep_l pfdu l f (map_in_dep (fun_in_dep_incl _ _ f pfa))))) (pfx:In x l), let f' := fun_in_dep_list_minus_incl pfdt pfdu f la pfa in f x pfx = f' x pfx'. intros T U l h1 h2 f la h3 x h4 h5; simpl. unfold fun_in_dep_list_minus_incl. f_equal. apply proof_irrelevance. Qed. Lemma incl_inv_im_in_dep_l : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (la:list T) (f:fun_in_dep l U) (pfa:incl la l), incl la (inv_im_in_dep_l pfdu l f (map_in_dep f)). intros T U h1 l la f h2; red; intros x h3. rewrite in_inv_im_in_dep_l_iff. exists (h2 _ h3). rewrite in_map_in_dep_iff. exists x, (h2 _ h3); auto. Qed. Lemma list_to_set_map_in_dep_list_minus_incl : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (f:fun_in_dep l U) (lx:list T) (pf:incl lx l), let pf' := inv_im_in_dep_l_incl pfdu _ (map_in_dep (fun_in_dep_incl _ _ f pf)) f in let f' := fun_in_dep_incl _ _ f pf' in list_to_set (map_in_dep (fun_in_dep_list_minus_incl pfdt pfdu f lx pf)) = Setminus (list_to_set (map_in_dep f)) (list_to_set (map_in_dep f')). intros T U h1 h2 l f lx h3. simpl. apply Extensionality_Ensembles; red; split; red; intros y h4. rewrite <- list_to_set_in_iff in h4. rewrite in_map_in_dep_iff in h4. destruct h4 as [x [h4 h5]]; subst. pose proof (in_list_minus_impl h1 _ _ _ h4) as h5. destruct h5 as [h5 h6]. constructor. rewrite <- list_to_set_in_iff. unfold fun_in_dep_list_minus_incl. apply in_map_in_dep. intro h7. rewrite <- list_to_set_in_iff in h7. rewrite in_map_in_dep_iff in h7. destruct h7 as [x' [h7 h8]]. unfold fun_in_dep_incl, fun_in_dep_list_minus_incl in h8. contradict h6. rewrite in_inv_im_in_dep_l_iff. exists h5. rewrite in_map_in_dep_iff. pose proof h7 as h7'. rewrite in_inv_im_in_dep_l_iff in h7'. destruct h7' as [h9 h10]. rewrite in_map_in_dep_iff in h10. destruct h10 as [x'' [h11 h12]]. exists x'', h11. rewrite h12. let P := type of h8 in match P with f x' ?pf = _ => assert (h13:pf = h9) end. apply proof_irrelevance. subst. rewrite h8. f_equal. apply proof_irrelevance. destruct h4 as [h4 h5]. rewrite <- list_to_set_in_iff in h4. rewrite <- list_to_set_in_iff in h5. rewrite in_map_in_dep_iff in h5. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff in h4. destruct h4 as [x [h4 h6]]; subst. rewrite in_map_in_dep_iff. exists x. ex_goal. eapply impl_in_list_minus; auto. intro h6. contradict h5. exists x, h6. unfold fun_in_dep_incl. f_equal. apply proof_irrelevance. exists hex. unfold fun_in_dep_list_minus_incl. f_equal. apply proof_irrelevance. Qed. Lemma incl_map_in_dep_incl : forall {T U:Type} (l' l:list T) (f:fun_in_dep l U) (pf:incl l' l), incl (map_in_dep (fun_in_dep_incl l l' f pf)) (map_in_dep f). intros T U l' l f h1; red; intros x h2. rewrite in_map_in_dep_iff in h2. rewrite in_map_in_dep_iff. destruct h2 as [x' [h2 h3]]; subst. exists x', (h1 _ h2); auto. Qed. Lemma in_match_in_dec_iff : forall {T U:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (A:In a l ->Ensemble U) (B: ~In a l->Ensemble U) (x:U), Inn match in_dec pfdt a l with | left pfa => A pfa | right pfb => B pfb end x <-> (exists pfin : In a l, Inn (A pfin) x) \/ (exists pfnin : ~In a l, Inn (B pfnin) x). intros T' U h1 l a A B x. destruct (in_dec h1 a l) as [h2 | h2]; split; intros h3. left. exists h2; auto. destruct h3 as [h3 | h3]. destruct h3 as [h3 h4]. pose proof (proof_irrelevance _ h2 h3); subst; auto. destruct h3; contradiction. right. exists h2; auto. destruct h3 as [h3 | h3]. destruct h3; contradiction. destruct h3 as [h3 h4]. pose proof (proof_irrelevance _ h2 h3); subst; auto. Qed. Lemma list_decompose : forall {T:Type} (l:list T) (def:T), l <> nil -> l = (hd def l) :: (tl l). intros T l. induction l; auto. intros def h1. contradict h1; auto. Qed. (*The last element (like [last]), and unlike [last], it requires no default value, just a guarantee the list is non-empty.*) Definition lst {T:Type} (l:list T) : l <> nil -> T := match l with | nil => fun pf => False_rect T (pf eq_refl) | a::l' => fun pf => last l (hd a l) end. Lemma lst_functional : forall {T:Type} (l l':list T) (pf:l <> nil) (pf':l' <> nil), l = l' -> lst l pf = lst l' pf'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma lst_indep : forall {T:Type} (l:list T) (pf pf':l <> nil), lst l pf = lst l pf'. intros; f_equal; apply proof_irrelevance. Qed. Lemma lst_compat : forall {T:Type} (l:list T) (pf:l <> nil), let a := lst l pf in l = (removelast l)++(a::nil). intros T l. induction l as [|a l ih]. intro h1; contradict h1; auto. intros h1. simpl. destruct l as [|c l]. rewrite app_nil_l; auto. simpl. f_equal. simpl in ih. specialize (ih (cons_neq_nil _ _)). rewrite ih. f_equal. destruct l; auto. erewrite last_indep. reflexivity. auto with datatypes. Qed. Lemma lst_compat' : forall {T:Type} (l:list T) (pf:l <> nil), lst l pf = last l (hd' l pf). intros ? l; destruct l; auto. Qed. Lemma last_compat : forall {T:Type} (l:list T) (def:T) (pf:l <> nil), last l def = lst l pf. intros; erewrite lst_functional. erewrite lst_compat' at 1. apply last_indep; auto. reflexivity. Grab Existential Variables. assumption. Qed. Lemma in_lst : forall {T:Type} (l:list T) (pf:l <> nil), In (lst l pf) l. intros T l. destruct l as [|a l]. intro h1. contradict h1; auto. intro; apply in_last; auto. Qed. Lemma lst_tl : forall {T:Type} (l:list T) (pf: tl l <> nil), let pf' := tl_not_nil _ pf in lst (tl l) pf = lst l pf'. intros ? l; destruct l; simpl. intro h1; contradict h1; auto. intro. rewrite (hd_compat' _ pf) at 2. rewrite lst_compat'. apply last_indep; auto. Qed. (*The analogue of the standard library*) Lemma app_removelast_lst : forall {T:Type} (l:list T) (pf:l <> nil), l = removelast l ++ (lst l pf) :: nil. intros T l. destruct l as [|a l]; simpl. intro h1. contradict h1; auto. intro h1. apply (app_removelast_last a (cons_neq_nil l a)). Qed. Lemma removelast_not_nil' : forall {T:Type} (l:list T), removelast l <> nil -> pred (pred (length l)) < length l. intros T l h1. rewrite <- length_removelast. pose proof (not_nil_lt _ h1) as h2. destruct l as [|a l]. simpl in h1; auto. erewrite length_removelast. simpl. omega. Qed. Lemma removelast_eq_iff : forall {T:Type} (l l':list T) (pf:l <> nil), removelast l = l' <-> l = l' ++ lst l pf::nil. intros T l l' h1; split; intro h2. rewrite <- h2. apply app_removelast_lst. rewrite h2. rewrite removelast_app_sing; auto. Qed. Definition map_proj1_sig {T} {P:T->Prop} (l:list {x:T | P x}) : list T := map (@proj1_sig _ P) l. Lemma in_map_proj1_sig_iff : forall {T} {P:T->Prop} (l:list {x:T | P x}) (x:T), In x (map_proj1_sig l) <-> exists (pfx:P x), In (exist _ x pfx) l. unfold map_proj1_sig; intros T P l x; rewrite in_map_iff; split; intro h1. destruct h1 as [y [? hy]]; subst. destruct y as [y hp]. simpl; exists hp; auto. destruct h1 as [h2 h3]. exists (exist _ _ h2); simpl. split; auto. Qed. Lemma map_proj1_sig_incl : forall {T} (l':list T) (l:list {x:T | In x l'}), incl (map_proj1_sig l) l'. intros; red; intros x h1. rewrite (in_map_proj1_sig_iff (P := fun c => In c l')) in h1. destruct h1; auto. Qed. Definition proj1_sig_leq {T} {P Q:T->Prop} (l: list {x:T | P x}) (l':list {x:T | Q x}) : Prop := map_proj1_sig l = map_proj1_sig l'. Lemma proj1_sig_leq_nil1 : forall {T} (P Q:T->Prop) (l': list {x:T | P x}), proj1_sig_leq nil l' (P:=P) -> l' = nil. unfold proj1_sig_leq; intros ? ? ? l. destruct l; auto; simpl. intro; discriminate. Qed. Lemma proj1_sig_leq_nil2 : forall {T} (P Q:T->Prop) (l: list {x:T | P x}), proj1_sig_leq l nil (Q:=Q)-> l = nil. unfold proj1_sig_leq; intros ? ? ? l. destruct l; auto; simpl. intro; discriminate. Qed. Lemma proj1_sig_leq_cons : forall {T} {P Q:T->Prop} (l: list {x:T | P x}) (l':list {x:T | Q x}) (a:T) (pfp:P a) (pfq:Q a), proj1_sig_leq l l' -> proj1_sig_leq (exist _ _ pfp :: l) (exist _ _ pfq :: l'). unfold proj1_sig_leq; intros; simpl; f_equal; auto. Qed. Lemma proj1_sig_leq_cons' : forall {T} {P Q:T->Prop} (l: list {x:T | P x}) (l':list {x:T | Q x}) (a:{x:T | P x}) (b:{x:T | Q x}), proj1_sig_leq (a :: l) (b :: l') -> proj1_sig a = proj1_sig b /\ proj1_sig_leq l l'. unfold proj1_sig_leq; intros ? ? ? ? ? ? ? h1; simpl in h1. inversion h1; auto. Qed. Definition map_sig {T:Type} (l:list T) : list {x:T | In x l} := map_in_dep (fun x (pfx:In x l) => (exist (fun c => In c l) _ pfx)). Lemma map_sig_nil : forall T, @map_sig T nil = nil. unfold map_sig; intros; apply map_in_dep_eq_nil_iff; auto. Qed. Lemma map_sig_in : forall {T:Type} (l:list T) (i:{x:T | In x l}), In i (map_sig l). unfold map_sig; intros T l i; rewrite in_map_in_dep_iff; exists (proj1_sig i), (proj2_sig i); destruct i; auto. Qed. Lemma length_map_sig : forall {T} (l:list T), length (map_sig l) = length l. unfold map_sig; intros; apply length_map_in_dep. Qed. Lemma in_map_sig : forall {T:Type} (l:list T) (x:{t:T | In t l}), In x (map_sig l) -> In (proj1_sig x) l. intros; apply proj2_sig. Qed. Lemma in_map_sig_iff : forall {T:Type} (l:list T) (x:{t:T | In t l}), In x (map_sig l) <-> In (proj1_sig x) l. intros; split; intro h1. apply in_map_sig; auto. unfold map_sig. rewrite in_map_in_dep_iff. exists (proj1_sig x), (proj2_sig x). apply proj1_sig_injective; auto. Qed. Definition map_sig_incl {T} {l l':list T} (pfinc:incl l l') : list {x:T | In x l'} := map_in_dep (fun x (pfin:In x l) => exist (fun c => In c l') x (pfinc _ pfin)). Lemma length_map_sig_incl : forall {T} {l l':list T} (pfinc:incl l l'), length (map_sig_incl pfinc) = length l. unfold map_sig_incl; intros; rewrite length_map_in_dep; auto. Qed. Definition map_sig' {T} {l:list T} {P:T->Prop} (Q:forall x, In x l -> P x) : list {x:T | P x} := map_in_dep (fun x (pf:In x l) => exist P x (Q x pf)). Lemma length_map_sig' : forall {T} {l:list T} {P:T->Prop} (Q:forall x, In x l -> P x), length (map_sig' Q) = length l. unfold map_sig'; intros; rewrite length_map_in_dep; auto. Qed. Lemma in_map_sig_iff' : forall {T} {l:list T} {P:T->Prop} (Q:forall x, In x l -> P x) (y:{x:T | P x}), In y (map_sig' Q) <-> In (proj1_sig y) l. unfold map_sig'; intros T l P Q y; split; intro h1. rewrite in_map_in_dep_iff in h1. destruct h1 as [? [? ?]]; subst; simpl. auto. rewrite in_map_in_dep_iff. exists (proj1_sig y), h1. apply proj1_sig_injective; auto. Qed. Definition map_sig_cons {T} (l:list T) a : list {x:T | In x (a::l)} := map_sig' (fun x (pf:In x l) => in_cons _ _ _ pf). Lemma in_map_sig_cons_iff : forall {T} (l:list T) a x, In x (map_sig_cons l a) <-> In (proj1_sig x) l. unfold map_sig_cons; intros; rewrite in_map_sig_iff'; tauto. Qed. Lemma length_map_sig_cons : forall {T} (l:list T) a, length (map_sig_cons l a) = length l. unfold map_sig'; intros; apply length_map_sig'; auto. Qed. Definition map_sig_from_in_cons {T} {l:list T} {a} (pfa:In a l) (l':list {x:T | In x (a::l)}) : list {x:T | In x l} := map (fun x => exist _ (proj1_sig x) (in_from_cons _ _ _ pfa (proj2_sig x))) l'. Lemma in_map_sig_from_in_cons_iff : forall {T} {l:list T} {a} (pfa:In a l) (l':list {x:T | In x (a::l)}) (y:{x:T | In x l}), In y (map_sig_from_in_cons pfa l') <-> In (exist (fun c => In c (a::l)) (proj1_sig y) (in_cons _ _ _ (proj2_sig y))) l'. unfold map_sig_from_in_cons; intros; rewrite in_map_iff; split; intro h1. destruct h1 as [x [? h1]]; subst; simpl. eapply in_subst1. apply proj1_sig_injective; simpl. reflexivity. assumption. exists (exist (fun c => In c (a::l)) _ (in_cons _ _ _ (proj2_sig y))); simpl. split. apply proj1_sig_injective; auto. assumption. Qed. Lemma in_map_sig_from_in_cons_iff' : forall {T} {l:list T} {a} (pfa:In a l) (l':list {x:T | In x (a::l)}) (y:{x:T | In x l}), In y (map_sig_from_in_cons pfa l') <-> In (proj1_sig y) (map_proj1_sig l'). intros; rewrite in_map_sig_from_in_cons_iff, (in_map_proj1_sig_iff (P := fun c => In c (a::l))); destruct y; simpl; split; intro h1. exists (in_cons _ _ _ i); auto. destruct h1 as [? h2]. erewrite f_equal_dep. apply h2. reflexivity. Qed. Definition map_sig_from_neq_cons {T} {l:list T} {a} (l':list {x:T | In x (a::l)}) (pfn:forall y, In y l' -> proj1_sig y <> a) : list {x:T | In x l} := map_in_dep (fun y (pfy:In y l') => exist (fun c => In c l) (proj1_sig y) (in_cons_neq _ _ _ (proj2_sig y) (pfn y pfy))). Lemma in_map_sig_from_neq_cons : forall {T} {l:list T} {a} (l':list {x:T | In x (a::l)}) (pfn:forall y, In y l' -> proj1_sig y <> a) (y:{x:T | In x l}), In y (map_sig_from_neq_cons l' pfn) <-> In (exist (fun c => In c (a::l)) (proj1_sig y) (in_cons _ _ _ (proj2_sig y))) l'. unfold map_sig_from_neq_cons; intros; rewrite in_map_in_dep_iff; split; intro h1. destruct h1 as [x [h1 h2]]; subst; simpl. eapply in_subst1. apply proj1_sig_injective; simpl; reflexivity. assumption. destruct y as [y hy]; simpl in h1. exists (exist (fun c => In c (a::l)) _ (in_cons _ _ _ hy)), h1; apply proj1_sig_injective; simpl; auto. Qed. Lemma in_map_sig_from_neq_cons': forall {T} {l:list T} {a} (l':list {x:T | In x (a::l)}) (pfn:forall y, In y l' -> proj1_sig y <> a) (y:{x:T | In x l}), In y (map_sig_from_neq_cons l' pfn) <-> In (proj1_sig y) (map_proj1_sig l'). intros; rewrite in_map_sig_from_neq_cons, (in_map_proj1_sig_iff (P := fun c => In c (a::l))); destruct y; split; intro h1; simpl in h1; simpl. exists (in_cons _ _ _ i); auto. destruct h1 as [h1 h2]. erewrite f_equal_dep. apply h2. reflexivity. Qed. Lemma in_from_sig_cons_remove : forall {T} {l:list T} {a} (pfdt:type_eq_dec {x:T | In x (a::l)}) (l':list {x:T | In x (a::l)}) (x:{t:T | In t (a::l)}), In x (remove pfdt (exist (fun c => In c (a::l)) a (in_eq _ _)) l') -> In (proj1_sig x) l. intros T l a hdt l' x h2. rewrite in_remove_iff in h2. destruct h2 as [h2 h3]. destruct x as [x hx]. simpl. destruct hx as [|hx]; subst; auto. contradict h3. f_equal; apply proof_irrelevance. Qed. Definition list_from_sig_cons_remove {T} {l:list T} {a} (pfdt:type_eq_dec T) (l':list {x:T | In x (a::l)}) : list {x:T | In x l} := map_in_dep (fun c (pfc:In c (remove (sig_eq_dec pfdt (fun c => In c (a::l))) (exist (fun c => In c (a::l)) a (in_eq _ _)) l')) => exist _ (proj1_sig c) (in_from_sig_cons_remove _ l' c pfc)). Lemma in_list_from_sig_cons_remove_iff : forall {T} {l:list T} {a} (pfdt:type_eq_dec T) (l':list {x:T | In x (a::l)}) (y:{x:T | In x l}), In y (list_from_sig_cons_remove pfdt l') <-> In (exist (fun c => (In c (a::l))) (proj1_sig y) (in_cons _ _ _ (proj2_sig y))) l' /\ (proj1_sig y <> a). unfold list_from_sig_cons_remove; intros; rewrite in_map_in_dep_iff; split; intro h1. destruct h1 as [x [h1 ?]]; subst; simpl. pose proof h1 as h1'. rewrite in_remove_iff in h1'. destruct h1' as [h2 h3]. split. eapply in_subst1. apply proj1_sig_injective; simpl; reflexivity. assumption. destruct x; simpl; contradict h3; subst; f_equal; apply proof_irrelevance. destruct h1 as [h1 h2]. exists (exist (fun c => (In c (a::l))) (proj1_sig y) (in_cons _ _ _ (proj2_sig y))). ex_goal. rewrite in_remove_iff. split; auto. intro h3. apply exist_injective in h3; subst. contradict h2; auto. exists hex. simpl. destruct y; simpl; f_equal; apply proof_irrelevance. Qed. (*This is an induction principle to tie in the left-recursive/ right-recursive incompatibility of induction on [nat] versus induction on [list]*) Lemma list_nat_ind : forall {T} (P:list T->nat->Prop) (l:list T) n, length l = n -> P nil 0 -> (forall a l' n', length l' = n' -> P l' n' -> P (l' ++ a::nil) (S n')) -> P l n. intros T P l n. revert l P. induction n as [|n ih]; simpl. intros l P h1. apply length_eq0 in h1; subst; auto. intros l P h1 h2 h3. pose proof (lt_eq_trans _ _ _ (lt_0_Sn n) (eq_sym h1)) as h4. apply lt_length in h4. rewrite (app_removelast_lst _ h4) in h1 at 1. rewrite app_length, S_compat in h1. apply S_inj in h1; subst. rewrite length_removelast. rewrite <- (S_pred _ _ (not_nil_lt _ h4)). rewrite (app_removelast_lst _ h4), app_length; simpl; rewrite S_compat. apply h3; auto. apply ih; auto. Qed. Lemma im_full_set_segT_eq : forall {T:Type} {n:nat} (f:n ---> T) (def:T), Im (Full_set (n @)) f = list_to_set (map (segT_fun_to_nat_fun f def) (seg_list n)). intros T n f def. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]. subst. rewrite <- list_to_set_in_iff. rewrite in_map_iff. exists (proj1_sig x). split. erewrite segT_fun_to_nat_fun_compat. f_equal. pose proof (unfold_sig _ x) as h2. symmetry in h2. apply h2. rewrite list_to_set_in_iff. rewrite list_to_set_seg_list_eq. constructor. apply proj2_sig. rewrite <- list_to_set_in_iff in h1. rewrite in_map_iff in h1. destruct h1 as [m h1]. destruct h1 as [h1 h2]. rewrite in_seg_list_iff in h2. apply Im_intro with (exist (fun a => a < n) _ h2). constructor. subst. rewrite (segT_fun_to_nat_fun_compat _ _ _ h2). reflexivity. Qed. Lemma in_list_power1 : forall {T1 T2:Type} {l1:list T1} {l2:list T2} (pl:list (T1 * T2)), In pl (list_power l1 l2) -> forall (pr:T1*T2), In pr pl -> In (fst pr) l1. intros T1 T2 l1 l2. induction l1 as [|b l1 h4]. intros pl h1. simpl in h1. destruct h1 as [h1l | h1r]. subst. contradiction. contradiction. intros pl h1. simpl in h1. rewrite in_flat_map in h1. destruct h1 as [pl' h5]. destruct h5 as [h5l h5r]. rewrite in_map_iff in h5r. destruct h5r as [t2 h6]. destruct h6 as [h6l h6r]. specialize (h4 pl' h5l). intros pr h7. subst. destruct h7 as [h7l | h7r]. subst. simpl. left. reflexivity. specialize (h4 pr h7r). right. assumption. Qed. Lemma in_list_power2 : forall {T1 T2:Type} {l1:list T1} {l2:list T2} (pl:list (T1 * T2)), In pl (list_power l1 l2) -> forall (pr:T1*T2), In pr pl -> In (snd pr) l2. intros T1 T2 l1 l2. induction l1 as [|b l1 h4]. intros pl h1. simpl in h1. intros pr h2. destruct h1 as [h1l | h1r]. subst. contradiction. contradiction. intros pl h2. simpl in h2. rewrite in_flat_map in h2. destruct h2 as [pl' h3]. destruct h3 as [h3l h3r]. rewrite in_map_iff in h3r. destruct h3r as [t2 h5]. destruct h5 as [h5l h5r]. subst. intros pr h6. destruct h6 as [h6l | h6r]. subst. simpl. assumption. specialize (h4 _ h3l _ h6r). assumption. Qed. Lemma no_dup_fun_well_defined_fst : forall {T U:Type} (S:Ensemble (T*U)) (l:list (T*U)), S = list_to_set l -> NoDup l ->fun_well_defined S -> NoDup (map (@fst _ _) l). intros T U S l h0 h1 h2. generalize dependent S. revert h1. induction l as [|pr l h1]. simpl. intros; constructor. intros h2 S h3 h4. simpl. constructor. assert (h5:Inn S pr). rewrite h3. rewrite <- list_to_set_in_iff. left. reflexivity. intro h6. rewrite in_map_iff in h6. destruct h6 as [pr' h6]. destruct h6 as [h6l h6r]. assert (h6:Inn S pr'). rewrite h3. rewrite <- list_to_set_in_iff. right. assumption. rewrite surjective_pairing in h5. rewrite surjective_pairing in h6. rewrite h6l in h6. pose proof (fp_functional h4 _ _ _ h5 h6) as h7. assert (h8:pr = pr'). apply injective_projections; auto. subst. pose proof (no_dup_cons_nin _ _ h2). contradiction. simpl in h3. pose proof (f_equal (fun S=>Subtract S pr) h3) as h5. simpl in h5. rewrite sub_add_same_nin in h5. pose proof (no_dup_cons _ _ h2) as h6. pose proof (incl_subtract S pr) as h7. pose proof (fun_well_defined_incl _ _ h7 h4) as h8. specialize (h1 h6 _ h5 h8). assumption. intro h6. rewrite <- list_to_set_in_iff in h6. pose proof (no_dup_cons_nin _ _ h2). contradiction. Qed. Inductive functionally_paired_l {T U:Type} (la:list T) (lb:list U) (lp:list (T*U)) : Prop := functional_pairs_intro : (forall (x:T), In x la -> (exists! y:U, In y lb /\ In (x, y) lp)) -> (forall (pr:T*U), In pr lp -> In (fst pr) la /\ In (snd pr) lb) -> functionally_paired_l la lb lp. Lemma fpl_in_dom : forall {T U:Type} {la:list T} {lb:list U} {lp:list (T*U)}, functionally_paired_l la lb lp -> forall pr:T*U, In pr lp -> In (fst pr) la. intros T U la lb lp h1 pr h2. destruct h1 as [h1a h1b]. specialize (h1b _ h2). destruct h1b; assumption. Qed. Lemma fpl_in_ran : forall {T U:Type} {la:list T} {lb:list U} {lp:list (T*U)}, functionally_paired_l la lb lp -> forall pr:T*U, In pr lp -> In (snd pr) lb. intros T U la lb lp h1 pr h2. destruct h1 as [h1a h1b]. specialize (h1b _ h2). destruct h1b; assumption. Qed. Lemma fpl_functional : forall {T U:Type} {la:list T} {lb:list U} {lp:list (T*U)}, functionally_paired_l la lb lp -> forall (x:T) (y1 y2:U), In (x, y1) lp -> In (x, y2) lp -> y1 = y2. intros T U la lb lp h1 x y1 y2 h2 h3. inversion h1 as [h4 h5]. pose proof (h5 _ h2) as h6. pose proof (h5 _ h3) as h7. simpl in h6. simpl in h7. destruct h6 as [h6a h6b]. destruct h7 as [h7a h7b]. pose proof (h4 _ h6a) as h8. destruct h8 as [y h9]. red in h9. destruct h9 as [h9a h9b]. pose proof (h9b _ (conj h6b h2)). pose proof (h9b _ (conj h7b h3)). subst. assumption. Qed. Lemma fpl_empty1 : forall (T U:Type) (lb:list U), functionally_paired_l (@nil T) lb nil. intros; constructor; intros; contradiction. Qed. Lemma fpl_empty1_s : forall {T U:Type} (lb:list U) (lp:list (T*U)), functionally_paired_l (@nil T) lb lp -> lp = nil. intros T U lb lp h1. destruct h1 as [h2 h3]. destruct lp as [|pr lp]. reflexivity. pose proof (in_eq pr lp) as h4. specialize (h3 _ h4). destruct h3; contradiction. Qed. Lemma no_fpl_empty2 : forall {T U:Type} (la:list T) (lp:list (T*U)), functionally_paired_l la (@nil U) lp -> la = nil. intros T U la. destruct la as [|a la]. intros; reflexivity. intros lp h2. destruct h2 as [h2 h3]. pose proof (in_eq a la) as h4. specialize (h2 _ h4). destruct h2 as [y h2]. red in h2. destruct h2 as [[]]; contradiction. Qed. Lemma no_fpl_empty2' : forall {T U:Type} (la:list T) (lp:list (T*U)), functionally_paired_l la (@nil U) lp -> lp = nil. intros T U la lp. destruct lp as [|pr lp]. intros; reflexivity. intro h1. pose proof (no_fpl_empty2 _ _ h1). subst. pose proof (fpl_empty1_s _ _ h1). discriminate. Qed. Lemma fp_fpl_compat : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (S:Ensemble (T*U)) (la:list T) (lb:list U) (lp:list (T*U)), list_to_set la = A -> list_to_set lb = B -> list_to_set lp = S -> ((functionally_paired A B S) <-> (functionally_paired_l la lb lp)). intros T U A B S la lb lp h1 h2 h3. split. intro h4. destruct h4 as [h5 h6]. constructor. intros x h7. rewrite list_to_set_in_iff in h7. rewrite h1 in h7. specialize (h5 _ h7). destruct h5 as [y h5]. exists y. red in h5. red. destruct h5 as [h5l h5r]. rewrite <- h2 in h5l. rewrite <- h3 in h5l. do 2 rewrite <- list_to_set_in_iff in h5l. split; try assumption. intros b h8. do 2 rewrite list_to_set_in_iff in h8. rewrite h2 in h8. rewrite h3 in h8. specialize (h5r _ h8). assumption. intros pr h7. rewrite list_to_set_in_iff in h7. rewrite h3 in h7. specialize (h6 _ h7). rewrite <- h1 in h6. rewrite <- h2 in h6. do 2 rewrite <- list_to_set_in_iff in h6. assumption. intro h4. destruct h4 as [h4 h5]. constructor. intros x h6. rewrite <- h1 in h6. rewrite <- list_to_set_in_iff in h6. specialize (h4 _ h6). destruct h4 as [y h4]. exists y. red in h4. red. destruct h4 as [h4l h4r]. split. do 2 rewrite list_to_set_in_iff in h4l. rewrite h2 in h4l. rewrite h3 in h4l. assumption. intros b h7. rewrite <- h2 in h7. rewrite <- h3 in h7. do 2 rewrite <- list_to_set_in_iff in h7. apply h4r; assumption. intros pr h6. rewrite <- h3 in h6. rewrite <- list_to_set_in_iff in h6. specialize (h5 _ h6). do 2 rewrite list_to_set_in_iff in h5. rewrite h1 in h5. rewrite h2 in h5. assumption. Qed. Lemma fpl_to_fun_in_dep_ex : forall {T U:Type} {la:list T} {lb:list U} {lp:list (T*U)}, functionally_paired_l la lb lp -> exists! f:fun_in_dep la U, forall (a:T) (pf:In a la), In (a, f a pf) lp. intros T U la lb lp h1. inversion h1 as [h0 h2]. pose (fun (x:T) (pf:In x la) => proj1_sig (constructive_definite_description _ (h0 x pf))) as f. exists f. unfold f. red. split. intros a h3. destruct constructive_definite_description as [u h4]. destruct h4 as [h4 h5]. simpl; auto. intros f' h3. apply fun_in_dep_ext. intros a h4. destruct constructive_definite_description as [u h5]. simpl. specialize (h3 _ h4). destruct h5 as [h5 h6]. pose proof h0 as h0'. specialize (h0' _ h4). destruct h0' as [u' h7]. red in h7. destruct h7 as [h7 h8]. destruct h7 as [h7 h9]. pose proof (h8 u (conj h5 h6)) as h10. subst. apply h8. split; auto. apply (fpl_in_ran h1 _ h3). Qed. Definition fpl_to_fun_in_dep {T U:Type} {la:list T} {lb:list U} {lp:list (T*U)} (pf:functionally_paired_l la lb lp) : fun_in_dep la U := proj1_sig (constructive_definite_description _ (fpl_to_fun_in_dep_ex pf)). Lemma fpl_to_fun_in_dep_compat : forall {T U:Type} {la:list T} {lb:list U} {lp:list (T*U)} (pf:functionally_paired_l la lb lp), let f := fpl_to_fun_in_dep pf in forall (a:T) (pf:In a la), In (a, f a pf) lp. unfold fpl_to_fun_in_dep; intros; destruct constructive_definite_description; auto. Qed. Lemma fpl_to_fun_in_dep_in_eq : forall {T U:Type} (la:list T) (lb:list U) (lp:list (T*U)) (pffp:functionally_paired_l la lb lp) (pr:T*U) (pfin:In pr lp), let f := fpl_to_fun_in_dep pffp in let pfina := fpl_in_dom pffp pr pfin in snd pr = f (fst pr) pfina. intros T U la lb lp h1 pr h2. simpl. symmetry. gen0. intro h3. pose proof (fpl_to_fun_in_dep_compat h1 _ h3) as h4. rewrite (surjective_pairing pr) in h2. eapply fpl_functional. apply h1. apply h4. assumption. Qed. Definition im1l {T U V:Type} {la:list T} {lb:list U} {lc:list V} {lp:list (T*U*V)} (pf:functionally_paired_l (list_prod la lb) lc lp) (x:T) (pfx:In x la) : list V := let f := fpl_to_fun_in_dep pf in map_in_dep (fun (y:U) (pfy:In y lb) => f (x, y) (in_prod _ _ _ _ pfx pfy)). Definition im2l {T U V:Type} {la:list T} {lb:list U} {lc:list V} {lp:list (T*U*V)} (pf:functionally_paired_l (list_prod la lb) lc lp) (y:U) (pfy:In y lb) : list V := let f := fpl_to_fun_in_dep pf in map_in_dep (fun (x:T) (pfx:In x la) => f (x, y) (in_prod _ _ _ _ pfx pfy)). Lemma no_dup_sing : forall {T:Type} (a:T), NoDup (a::nil). intros; constructor. intro; contradiction. constructor. Qed. Lemma no_dup_comm_cons : forall {T:Type} (l:list T) (a b:T), NoDup (a::b::l) ->NoDup (b::a::l). intros T l a b h1. pose proof (no_dup_cons _ _ h1) as h2. pose proof (no_dup_cons _ _ h2) as h3. pose proof (no_dup_cons_nin _ _ h1) as h4. constructor. contradict h4. destruct h4 as [|h4]; subst. left; auto. pose proof (no_dup_cons_nin _ _ h2); contradiction. constructor. contradict h4. right; auto. assumption. Qed. Lemma no_dup_remove_singleton : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), NoDup l -> remove pfdt a l = nil -> l = nil \/ l = a::nil. intros T h0 l. induction l as [|a l h1]; simpl; auto. intros x h2 h3. right. destruct (h0 x a) as [h4 | h5]. subst. pose proof (no_dup_cons _ _ h2) as h4. specialize (h1 _ h4 h3). destruct h1 as [h1 | h1]. subst; auto. rewrite h1 in h2. inversion h2. contradict H1. left. reflexivity. discriminate h3. Qed. Lemma no_dup_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), NoDup l -> NoDup (remove pfdt x l). intros T h0 l. induction l as [|a l h1]; simpl; auto. intros x h2. destruct (h0 x a) as [h3 | h4]. subst. apply no_dup_cons in h2. specialize (h1 a h2). assumption. pose proof (no_dup_cons _ _ h2) as h5. pose proof (no_dup_cons_nin _ _ h2) as h6. clear h2. specialize (h1 x h5). constructor; auto. apply nin_nin_remove. assumption. Qed. Lemma no_dup_tl : forall {T:Type} (l:list T), NoDup l -> NoDup (tl l). intros T l. induction l as [|a l h1]. simpl. auto. intro h2. simpl. apply no_dup_cons with a. assumption. Qed. Lemma count_occ_no_dup : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), NoDup l -> forall x:T, In x l -> count_occ pfdt l x = 1. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros h2 x h3. destruct h3 as [|h3]; subst. destruct (h1 x x) as [h3 | h3]. f_equal. pose proof (no_dup_cons_nin _ _ h2) as h4. rewrite (count_occ_In h1) in h4; omega. contradict h3; auto. destruct (h1 a x) as [h4 | h4]; subst. f_equal. pose proof (no_dup_cons_nin _ _ h2) as h4. rewrite (count_occ_In h1) in h4; omega. apply ih; auto. eapply no_dup_cons. apply h2. Qed. Lemma count_occ_map_inj : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:T->U) (l:list T) (x:T) (pfx:In x l), Injective f -> count_occ pfdu (map f l) (f x) = count_occ pfdt l x. intros T U h1 h2 f l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h3 h4. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. apply h4 in hl; subst. erewrite eq_dec_same at 1. f_equal. destruct (in_dec h1 x l) as [h5 | h5]. apply ih; auto. pose proof (count_occ_In h1 l x) as h6. rewrite h6 in h5. pose proof (count_occ_In h2 (map f l) (f x)) as h7. apply ngt_nlt in h5. rewrite not_lt0_iff in h5. rewrite h5. rewrite <- not_lt0_iff. apply ngt_nlt. rewrite <- (count_occ_In h2). intro h8. rewrite in_map_iff in h8. destruct h8 as [x' [h8 h9]]. apply h4 in h8; subst. rewrite h6 in h9. omega. assumption. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']; subst. contradict (hr eq_refl). destruct h3 as [|h3]; try contradiction. apply ih; auto. Qed. Lemma nin_hd_tl_no_dup : forall {T:Type} (l:list T) (def:T), NoDup l -> ~In (hd def l) (tl l). intros T l. induction l as [|a l h1]; simpl; auto. intros ? h2 h3. pose proof (no_dup_cons_nin _ _ h2) as h5. contradiction. Qed. Lemma nin_hd_tl_no_dup' : forall {T:Type} (l:list T) (pf:l <> nil), NoDup l -> ~In (hd' l pf) (tl l). intros T l h1. pose proof (not_nil_cons l h1) as h2. destruct h2 as [a [l' h2]]; subst. intro h2. simpl. apply no_dup_cons_nin; auto. Qed. Definition no_dup_faml {T:Type} (ll:faml T) : Prop := forall l:list T, In l ll -> NoDup l. Inductive dup {T:Type} : list T -> Prop := | dup_intro_in : forall (x:T) (l:list T), In x l -> dup (x::l) | dup_intro_cons: forall (a:T) (l:list T), dup l -> dup (a::l). Lemma not_dup_nil (T:Type): ~ (dup (@nil T)). intro h1; inversion h1. Qed. Lemma not_dup_sing : forall {T:Type} (a:T), ~dup (a::nil). intros T a h1. inversion h1 as [x l h2 h3 | x l h2 h3]; subst. red in h2; auto. inversion h2. Qed. Lemma not_dup_cons : forall {T:Type} (l:list T) (a:T), ~dup (a::l) -> ~ dup l. intros T l a h1 h2. contradict h1. eapply dup_intro_cons; auto. Qed. Lemma not_dup_cons_nin : forall {T:Type} (l:list T) (a:T), ~dup (a::l) -> ~In a l. intros T l a h1 h2. contradict h1. constructor; auto. Qed. Lemma no_dup_iff : forall {T:Type} (l:list T), NoDup l <-> ~ dup l. intros T l. induction l as [|a l ih]. split. intros h1 h2. inversion h2. intro; constructor. split. intros h1 h2. inversion h2 as [x l' h3 h4|]; subst. pose proof (no_dup_cons_nin _ _ h1) as h4. contradiction. pose proof (no_dup_cons _ _ h1) as h3. rewrite ih in h3. contradiction. intro h1. pose proof (not_dup_cons _ _ h1) as h2. pose proof (not_dup_cons_nin _ _ h1) as h3. constructor; auto. rewrite ih; auto. Qed. Lemma dup_dec : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), {dup l} + {NoDup l}. intros T h0 l. induction l as [|a l ih]. right. constructor. destruct ih as [ih | ih]. left. apply dup_intro_cons; auto. destruct (in_dec h0 a l) as [h1 | h1]. left. constructor; auto. right. constructor; auto. Qed. Lemma dup_iff : forall {T:Type} (pf:type_eq_dec T) (l:list T), dup l <-> ~ NoDup l. intros T h0 l. induction l as [|a l ih]; simpl. split. intro h1. inversion h1. intro h1. contradict h1. constructor. split. intro h1. inversion h1 as [x l' h2 h3| xl' h2 h3]; subst. intro h3. pose proof (no_dup_cons_nin _ _ h3) as h4. contradiction. rewrite ih in h3. intro h4. contradict h3. eapply no_dup_cons. apply h4. intro h1. destruct (dup_dec h0 (a::l)); auto. contradiction. Qed. Lemma dup_cons_iff : forall {T:Type} (l:list T) (a:T), dup (a::l) <-> In a l \/ dup l. intros T l h1. split. intro h2. inversion h2; subst. left; auto. right; auto. intro h2. destruct h2 as [h2 | h2]; simpl. left; auto. right; auto. Qed. Lemma dup_cons : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), dup (a::l) -> {In a l} + {dup l}. intros T h0 l a h1. destruct (dup_dec h0 l) as [h2 | h2]. right; auto. left. rewrite dup_cons_iff in h1. destruct h1 as [|h1]; auto. rewrite dup_iff in h1. contradiction. assumption. Qed. Lemma dup_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), dup (remove pfdt a l) -> dup l. intros T h1 l. induction l as [|a l ih]; simpl; auto. intros a' h2. destruct (h1 a' a) as [h3 | h3]; subst. rewrite dup_cons_iff. apply ih in h2. right; auto. rewrite dup_cons_iff in h2. destruct h2 as [h2 | h2]. left. rewrite in_remove_iff in h2. destruct h2; auto. apply ih in h2. right; auto. Qed. Lemma dup_list_minus : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (l':list T), dup (list_minus pfdt l l') -> dup l. intros T h1 l. induction l as [|a l ih]; simpl; auto. intros a' h2. destruct (in_dec h1 a a') as [h3 | h3]. right. eapply ih. apply h2. rewrite dup_cons_iff in h2. destruct h2 as [h2 | h2]. apply in_list_minus_impl in h2. destruct h2 as [h2 h4]. left. assumption. right. eapply ih. apply h2. Qed. Lemma list_to_set_eq_sing : forall {T:Type} (l:list T) (a:T), NoDup l -> list_to_set l = Singleton a -> l = a::nil. intros T l. induction l as [|c l h1]; simpl. intros a h1 h2. symmetry in h2. apply sing_not_empty in h2. contradiction. intros a h2 h3. symmetry in h3. apply sing_eq_add in h3. destruct h3 as [h3 h4]. subst. destruct h4 as [h4 | h4]. apply empty_set_nil in h4. subst; auto. pose proof (no_dup_cons _ _ h2) as h3. specialize (h1 _ h3 h4). subst. inversion h2. contradict H1. left; auto. Qed. Lemma list_to_set_eq_couple : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a b:T), NoDup l -> list_to_set l = Couple a b -> a = b \/ l = a::b::nil \/ l = b::a::nil. intros T hdt l a b h1 h2. destruct l as [|c l]. simpl in h2. symmetry in h2. pose proof (inh_couple a b) as h3. rewrite h2 in h3. destruct h3; contradiction. pose proof (Couple_l _ a b) as h3. rewrite <- h2 in h3. rewrite <- list_to_set_in_iff in h3. destruct h3 as [|h3]; subst. pose proof (Couple_r _ a b) as h3. rewrite <- h2 in h3. rewrite <- list_to_set_in_iff in h3. destruct h3 as [|h3]; subst. left; auto. right. left. f_equal. pose proof (no_dup_cons_nin _ _ h1) as h4. simpl in h2. rewrite couple_comm, couple_add_sing in h2. apply add_nin_inj_eq in h2. apply list_to_set_eq_sing in h2; auto. apply no_dup_cons in h1; auto. rewrite <- list_to_set_in_iff; auto. rewrite in_sing_iff; intro; subst; contradiction. pose proof (Couple_r _ a b) as h4. simpl in h2. rewrite <- h2 in h4. destruct h4 as [b h4 | b h4]. rewrite <- list_to_set_in_iff in h4. assert (h5:c <> a). intro; subst. apply no_dup_cons_nin in h1; contradiction. assert (h6:c <> b). intro; subst. apply no_dup_cons_nin in h1; contradiction. pose proof (Add_intro2 _ (list_to_set l) c) as h7. rewrite h2 in h7. destruct h7. contradict h5; auto. contradict h6; auto. destruct h4. pose proof (no_dup_cons_nin _ _ h1) as h4. rewrite couple_add_sing in h2. apply add_nin_inj_eq in h2. apply list_to_set_eq_sing in h2; subst. right; right; auto. apply no_dup_cons in h1. assumption. rewrite <- list_to_set_in_iff; auto. rewrite in_sing_iff; intro; subst; contradiction. Qed. Lemma lt_length_removelast : forall {T:Type} (l:list T) (m:nat), S m < length l -> m < length (removelast l). intros T l m h1. pose proof (lt_length _ _ h1) as h2. pose proof (lst_compat _ h2) as h3. simpl in h3. rewrite h3 in h1. simpl in h1. rewrite app_length in h1. simpl in h1. omega. Qed. Lemma le_length_removelast : forall {T:Type} (l:list T) (m:nat), S m <= length l -> m <= length (removelast l). intros T l m h1. pose proof (lt_length _ _ h1) as h2. pose proof (lst_compat _ h2) as h3. simpl in h3. rewrite h3 in h1. simpl in h1. rewrite app_length in h1. simpl in h1. omega. Qed. Lemma S_length_remove_no_dup : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), NoDup l -> In x l -> S (length (remove pfdt x l)) = length l. intros T h0 l. induction l as [|a l h1]. intros; contradiction. intros x h2 h3. pose proof (no_dup_cons_nin _ _ h2) as h4. simpl in h3. destruct h3 as [h3 | h3]. subst. simpl. destruct (h0 x x) as [h3 | h3]. f_equal. rewrite <- remove_not_in'; auto. contradict h3; auto. pose proof (no_dup_cons _ _ h2) as h5. specialize (h1 _ h5 h3). simpl. f_equal. destruct (h0 x a) as [h6 | h6]. subst. contradiction. simpl. assumption. Qed. Lemma length_remove_no_dup : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), NoDup l -> In a l -> length (remove pfdt a l) = pred (length l). intros T h0 l a h1 h2. assert (h3:l <> nil). intro; subst. simpl in h2. contradiction. pose proof (list_decompose _ a h3) as h4. rewrite h4. simpl. destruct (h0 a (hd a l)) as [h5 | h5]. rewrite h5 at 1. pose proof (nin_hd_tl_no_dup _ a h1) as h6. rewrite <- (remove_not_in' h0 (tl l) (hd a l)); auto. simpl. assert (h6:In a (tl l)). rewrite h4 in h2. simpl in h2. destruct h2 as [h2 | h2]. symmetry in h2. contradiction. assumption. apply S_length_remove_no_dup; auto. rewrite h4 in h1. apply (no_dup_cons _ _ h1). Qed. Lemma incl_no_dup_le : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), NoDup l -> NoDup l' -> incl l l' -> length l <= length l'. intros T hdt l. induction l as [|a l ih]; simpl; auto with arith. intros l' h1 h2 h3. destruct l' as [|b l']. apply incl_nil' in h3. discriminate. simpl. apply le_n_S. destruct (hdt a b) as [|h4]; subst. apply incl_cons2_same in h3. apply ih. apply no_dup_cons in h1. apply no_dup_cons in h2. assumption. apply no_dup_cons in h2; auto. assumption. apply no_dup_cons_nin in h1; auto. apply no_dup_cons_nin in h2; auto. specialize (ih (remove hdt a (b::l')) (no_dup_cons _ _ h1)). hfirst ih. apply no_dup_remove; auto. specialize (ih hf). hfirst ih. red in h3; red. intros x h6. specialize (h3 _ (in_cons _ _ _ h6)). destruct h3 as [|h3]; subst. rewrite in_remove_iff. split. left; auto. auto. rewrite in_remove_iff. split. right; auto. intro; subst. apply no_dup_cons_nin in h1. contradiction. specialize (ih hf0). rewrite length_remove_no_dup in ih. simpl in ih; auto. assumption. right. red in h3. specialize (h3 _ (in_eq _ _)). destruct h3; subst; auto. contradict h4; auto. Qed. Lemma incl_cons_length : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T), incl (a::l) l' -> incl l' l -> NoDup (a::l) -> NoDup l' -> length l' - length l = 1. intros T hdt l l' a h2 h3 h4 h5. pose proof (incl_no_dup_le hdt _ _ h4 h5 h2) as h6. pose proof (incl_no_dup_le hdt _ _ h5 (no_dup_cons _ _ h4) h3) as h7. simpl in h6. omega. Qed. Lemma partition_faml_le : forall {T:Type} (pfdt:type_eq_dec T) (ll:faml T) (l lx:list T), no_dup_mem ll -> NoDup l -> partition_faml ll l -> In lx ll -> length lx <= length l. intros T hdt ll l lx hnd hnd' h1 h2. pose proof (partition_faml_incl _ _ _ h1 h2) as h3. apply incl_no_dup_le in h3; auto. Qed. Lemma no_dup_list_to_set_eq_length : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), NoDup l -> NoDup l' -> list_to_set l = list_to_set l' -> length l = length l'. intros T hdt l. induction l as [|a l ih]; simpl. intros ? ? ? h1. symmetry in h1. apply empty_set_nil in h1; subst; auto. intro l'. destruct l' as [|a' l']. intros ? ? h1. simpl in h1. apply add_not_empty in h1; contradiction. intros h1 h2 h3. simpl in h3; simpl. destruct (hdt a a') as [|h4]; subst. pose proof (f_equal (fun S => Subtract S a') h3) as h4. simpl in h4. do 2 rewrite sub_add_same_nin in h4. f_equal. apply ih. apply no_dup_cons in h1; auto. apply no_dup_cons in h2; auto. assumption. pose proof (no_dup_cons_nin _ _ h2) as h5. rewrite <- list_to_set_in_iff; auto. pose proof (no_dup_cons_nin _ _ h1) as h5. rewrite <- list_to_set_in_iff; auto. pose proof (no_dup_cons_nin _ _ h1) as h5. rewrite <- list_to_set_in_iff; auto. f_equal. pose proof (f_equal (fun S => Subtract S a) h3) as h5. simpl in h5. rewrite sub_add_same_nin in h5. specialize (ih (remove hdt a (a'::l')) (no_dup_cons _ _ h1)). hfirst ih. apply no_dup_remove; auto. specialize (ih hf). simpl in ih. destruct (hdt a a') as [h6 | h6]. contradiction. simpl in ih. rewrite sub_add_comm in h5. rewrite (subtract_remove_compat hdt) in h5. specialize (ih h5). rewrite ih. rewrite length_remove_no_dup. pose proof (Add_intro2 _ (list_to_set l) a) as h7. rewrite h3 in h7. destruct h7 as [a h7 | a h7]. rewrite <- list_to_set_in_iff in h7. apply in_not_nil in h7. apply not_nil_lt in h7. omega. destruct h7. contradict h4; auto. apply no_dup_cons in h2; auto. pose proof (Add_intro2 _ (list_to_set l) a) as h7. rewrite h3 in h7. destruct h7 as [a h7 | a h7]. rewrite <- list_to_set_in_iff in h7; auto. destruct h7; contradict h4; auto. auto. apply no_dup_cons_nin in h1. rewrite <- list_to_set_in_iff; auto. Qed. Lemma list_minus_remove : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a':T), NoDup l -> In a' l' -> ~In a' l -> list_minus pfdt l (remove pfdt a' l') = list_minus pfdt l l'. intros T h1 l. induction l as [|a l ih]; simpl; auto. intros l' a' hnd h2 hnin. destruct in_dec as [h3 | h3]. destruct in_dec as [h4 | h4]. apply ih; auto. apply no_dup_cons in hnd; auto. rewrite in_remove_iff in h3. destruct h3; contradiction. destruct in_dec as [h4 | h4]. destruct (h1 a a') as [|h5]; subst. contradict hnin; left; auto. contradict h3. rewrite in_remove_iff. split; auto. f_equal. apply ih. apply no_dup_cons in hnd; auto. assumption. contradict hnin. right; auto. Qed. Lemma no_dup_length_list_minus : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), NoDup l -> NoDup l' -> incl l' l -> length (list_minus pfdt l l') = length l - length l'. intros T hdt l. induction l as [|a l ih]; simpl; auto. intros l' h1 h2 h3. destruct l' as [|a' l']. destruct in_dec. contradiction. simpl. f_equal. rewrite ih; simpl; auto with arith. apply no_dup_cons in h1; auto. apply incl_nil. destruct (hdt a a') as [|h4]; subst. destruct in_dec as [h4 | h4]. rewrite list_minus_cons_remove. rewrite <- (remove_not_in' hdt (list_minus hdt l l') a'). simpl. apply ih. apply no_dup_cons in h1; auto. apply no_dup_cons in h2; auto. pose proof (no_dup_cons_nin _ _ h1) as h5. apply incl_cons2_same in h3; auto. apply no_dup_cons_nin in h2; auto. apply no_dup_cons_nin in h1. contradict h1. apply in_list_minus_impl in h1. destruct h1; auto. contradict h4. left; auto. destruct in_dec as [h5 | h5]. destruct h5 as [|h5]; subst. contradict h4; auto. pose proof (h3 _ (in_eq _ _)) as hi. destruct hi as [|hi]; subst. contradict h4; auto. simpl. rewrite list_minus_cons_remove. rewrite length_remove_no_dup. specialize (ih (remove hdt a (a'::l')) (no_dup_cons _ _ h1) (no_dup_remove hdt _ _ h2)). rewrite remove_cons_neq in ih. simpl in ih. rewrite list_minus_cons_remove in ih. rewrite length_remove_no_dup in ih. simpl. rewrite length_remove_no_dup in ih. pose proof (in_not_nil _ _ h5) as h7. apply not_nil_lt in h7. rewrite <- (S_pred _ _ h7) in ih; auto. rewrite <- ih. f_equal. f_equal. symmetry. apply list_minus_remove; auto. apply no_dup_cons in h1; auto. apply no_dup_cons_nin in h1; auto. apply no_dup_cons in h2; auto. red in h3; red. intros x ha. destruct ha as [|ha]; subst; auto. rewrite in_remove_iff in ha. destruct ha as [ha hb]. specialize (h3 _ (in_cons _ _ _ ha)). destruct h3 as [|h3]; subst. contradict hb; auto. assumption. apply no_dup_cons in h2; auto. assumption. apply no_dup_list_minus. apply no_dup_cons in h1; auto. apply impl_in_list_minus; auto. intro ha. rewrite in_remove_iff in ha. destruct ha as [ha hb]. apply no_dup_cons_nin in h2. contradiction. assumption. apply no_dup_list_minus. apply no_dup_cons in h1; auto. apply impl_in_list_minus; auto. apply no_dup_cons_nin in h2; auto. simpl. rewrite list_minus_cons_remove. destruct (in_dec hdt a' l) as [hi | hi]. specialize (ih l' (no_dup_cons _ _ h1) (no_dup_cons _ _ h2)). hfirst ih. red. intros x ha. destruct (hdt x a) as [|hb]; subst. contradict h5; right; auto. red in h3. specialize (h3 _ (in_cons _ _ _ ha)). destruct h3 as [|h3]; subst. contradict hb; auto. assumption. specialize (ih hf). rewrite length_remove_no_dup. pose proof (impl_in_list_minus hdt l l' _ hi (no_dup_cons_nin _ _ h2)) as h6. apply in_not_nil in h6. apply not_nil_lt in h6. rewrite <- (S_pred _ _ h6); auto. apply no_dup_list_minus. apply no_dup_cons in h1; auto. apply impl_in_list_minus; auto. apply no_dup_cons_nin in h2; auto. contradict hi. red in h3. specialize (h3 _ (in_eq _ _)). destruct h3 as [|h3]; subst. contradict h4; auto. assumption. Qed. Lemma no_dup_partition_faml_nat_list_to_set_eq : forall (ll ll':faml nat) (lm ln:list nat), NoDup ll -> NoDup ll' -> no_dup_mem ll -> no_dup_mem ll' -> list_to_set_injable (fun lx => In lx ll) -> list_to_set_injable (fun lx => In lx ll') -> list_to_set ll = list_to_set ll' -> NoDup lm -> NoDup ln -> partition_faml ll lm -> partition_faml ll' ln -> length lm = length ln. intro ll. induction ll as [|la ll ih]; simpl. intros ? ? ? ? ? ? ? ? ? h1. symmetry in h1. apply empty_set_nil in h1; subst. intros ? ? h1 h2. apply partition_faml_nil in h1. apply partition_faml_nil in h2. subst; auto. intro ll'. destruct ll' as [|la' ll']. intros ? ? ? ? ? ? ? ? h1 ? ?. simpl in h1. apply add_not_empty in h1. contradiction. intros lm ln h1 h2 hndpmm hndpmn hli hli' h3 hnm hnn. simpl in h3. destruct (list_nat_eq_dec la la') as [|h4]; subst. pose proof (f_equal (fun S => Subtract S la') h3) as h4. simpl in h4. do 2 rewrite sub_add_same_nin in h4. intros h5 h6. pose proof h5 as h5'. pose proof h6 as h6'. apply (partition_faml_cons nat_eq_dec _ la') in h5. apply (partition_faml_cons nat_eq_dec _ la') in h6. specialize (ih ll' (list_minus nat_eq_dec lm la') (list_minus nat_eq_dec ln la') (no_dup_cons _ _ h1) (no_dup_cons _ _ h2) (no_dup_mem_cons _ _ hndpmm) (no_dup_mem_cons _ _ hndpmn) (list_to_set_injable_in_cons _ _ hli) (list_to_set_injable_in_cons _ _ hli') h4 (no_dup_list_minus nat_eq_dec _ _ hnm) (no_dup_list_minus nat_eq_dec _ _ hnn) h5 h6). pose proof (hndpmm _ (in_eq _ _)) as h7. pose proof (hndpmn _ (in_eq _ _)) as h8. do 2 rewrite no_dup_length_list_minus in ih; auto. pose proof (partition_faml_incl _ _ _ h5' (in_eq _ _)) as h9. pose proof (partition_faml_incl _ _ _ h6' (in_eq _ _)) as h10. pose proof h9 as h9'. pose proof h10 as h10'. apply incl_no_dup_le in h9'; auto. apply incl_no_dup_le in h10'; auto. omega. apply nat_eq_dec. apply nat_eq_dec. eapply partition_faml_incl in h6'. apply h6'. left; auto. eapply partition_faml_incl in h5'. apply h5'. left; auto. eapply partition_faml_incl in h5'. apply h5'. left; auto. eapply partition_faml_incl in h5'. apply h5'. left; auto. eapply partition_faml_incl in h5'. apply h5'. left; auto. intro h7. apply in_faml_to_fam_impl in h7. destruct h7 as [l [h7 h8]]. pose proof (no_dup_cons_nin _ _ h2) as h9. red in hli'. pose proof (hli' _ _ (in_cons _ _ _ h7) (in_eq _ _)) as h10. apply h10 in h8; subst. contradiction. intro h7. apply in_faml_to_fam_impl in h7. destruct h7 as [l [h7 h8]]. pose proof (no_dup_cons_nin _ _ h1) as h9. pose proof (hli _ _ (in_cons _ _ _ h7) (in_eq _ _)) as h10. apply h10 in h8; subst. contradiction. apply no_dup_cons_nin in h2. rewrite <- list_to_set_in_iff; auto. apply no_dup_cons_nin in h1. rewrite <- list_to_set_in_iff; auto. apply no_dup_cons_nin in h1. rewrite <- list_to_set_in_iff; auto. intros h5 h6. specialize (ih (remove list_nat_eq_dec la (la'::ll')) (list_minus nat_eq_dec lm la) (list_minus nat_eq_dec ln la) (no_dup_cons _ _ h1) (no_dup_remove _ _ _ h2) (no_dup_mem_cons _ _ hndpmm) (no_dup_mem_remove _ _ _ hndpmn) (list_to_set_injable_in_cons _ _ hli) (list_to_set_injable_in_remove _ _ _ hli')). hfirst ih. rewrite <- subtract_remove_compat. simpl. rewrite <- h3. rewrite sub_add_same_nin; auto. apply no_dup_cons_nin in h1. rewrite <- list_to_set_in_iff; auto. specialize (ih hf (no_dup_list_minus _ _ _ hnm) (no_dup_list_minus _ _ _ hnn)). pose proof h5 as h5'. pose proof h6 as h6'. apply (partition_faml_cons nat_eq_dec _ la) in h5. pose proof (partition_faml_remove nat_eq_dec (la' :: ll') ln la) as h7. hfirst h7. right. pose proof (Add_intro2 _ (list_to_set ll) (la)) as h8. rewrite h3 in h8. destruct h8 as [la h8 | la h8]. rewrite list_to_set_in_iff; auto. destruct h8. contradict h4; auto. pose proof (partition_faml_incl _ _ _ h5' (in_eq _ _)) as h9. pose proof (partition_faml_incl _ _ _ h6' hf0) as h10. pose proof h9 as h9'. pose proof h10 as h10'. apply incl_no_dup_le in h9; auto. apply incl_no_dup_le in h10; auto. specialize (h7 hf0 hli' h6'). pose proof (list_nat_eq_dec_eq (list_eq_dec nat_eq_dec)) as h8. rewrite h8 in h7. specialize (ih h5 h7). do 2 rewrite no_dup_length_list_minus in ih; auto. omega. apply nat_eq_dec. apply nat_eq_dec. pose proof (no_dup_cons_nin _ _ h1) as h7. contradict h7. apply in_faml_to_fam_impl in h7. destruct h7 as [l [h7 h8]]. apply hli in h8; subst; auto. Qed. Lemma partition_faml_length : forall {T:Type} (pfdt:type_eq_dec T) (ll:faml T) (l:list T), NoDup l -> NoDup ll -> no_dup_mem ll -> list_to_set_injable (fun lx => In lx ll) -> partition_faml ll l -> plusl (map (@length _) ll) = length l. intros T hdt ll. induction ll as [|al ll ih]; simpl. intros l hnd0 hnd hndm h0 h1. apply partition_faml_nil in h1; subst; auto. intros l hnd0 hnd hndm h1. specialize (ih (list_minus hdt l al) (no_dup_list_minus _ _ _ hnd0) (no_dup_cons _ _ hnd) (no_dup_mem_cons _ _ hndm) (list_to_set_injable_in_cons _ _ h1)). pose proof (partition_faml_remove hdt (al::ll) l _ (in_eq _ _) h1) as h2. intro h3. specialize (h2 h3). rewrite remove_cons_eq in h2. erewrite <- remove_not_in' in h2. apply ih in h2. rewrite h2. rewrite no_dup_length_list_minus; auto with arith. pose proof (partition_faml_le hdt (al::ll) _ _ hndm hnd0 h3 (in_eq _ _)) as h4. omega. eapply hndm. left; auto. red. intros a h4. do 2 red in h3. simpl in h3. destruct h3 as [h3 h5]. rewrite family_union_add in h5. rewrite list_to_set_in_iff. rewrite <- h5. left. rewrite <- list_to_set_in_iff; auto. apply no_dup_cons_nin in hnd; auto. Qed. Lemma partition_faml_same_length_eq : forall (ll ll':faml nat) (ln:list nat), NoDup ll -> NoDup ll' -> no_dup_mem ll -> no_dup_mem ll' -> list_to_set_injable (fun lx => In lx ll) -> list_to_set_injable (fun lx => In lx ll') -> list_to_set ll = list_to_set ll' -> NoDup ln -> partition_faml ll ln -> partition_faml ll' ln -> length ll = length ll'. intro ll. induction ll as [|al ll ih]; simpl. intros ? ? ? ? ? ? ? ? h1. symmetry in h1. apply empty_set_nil in h1; subst; auto. intro ll'. destruct ll' as [|al' ll']. intros ? ? ? ? ? ? ? h1. simpl in h1. apply add_not_empty in h1; contradiction. intros ln h1 h2 h3 h4 h5 h6 h7 h8 h9 h10. pose proof (no_dup_cons_nin _ _ h1) as hnin. pose proof (no_dup_cons_nin _ _ h2) as hnin'. pose proof (no_dup_cons _ _ h1) as h1'. pose proof (no_dup_cons _ _ h2) as h2'. pose proof (no_dup_mem_cons _ _ h3) as h11. pose proof (no_dup_mem_cons _ _ h4) as h12. pose proof (list_to_set_injable_in_cons ll al h5) as h13. pose proof (list_to_set_injable_in_cons ll' al' h6) as h14. pose proof (partition_faml_cons nat_eq_dec ll al ln) as h15. hfirst h15. intro h16. apply in_faml_to_fam_impl in h16. destruct h16 as [l [h16 h17]]. red in h5. pose proof (h5 l al (in_cons _ _ _ h16) (in_eq _ _) h17); subst. contradiction. specialize (h15 hf h9). simpl; f_equal. destruct (list_nat_eq_dec al al') as [h16 | h16]; subst. simpl in h7. pose proof (f_equal (fun S => Subtract S al') h7) as h17. simpl in h17. do 2 rewrite sub_add_same_nin in h17; auto. pose proof (partition_faml_cons nat_eq_dec ll' al' ln) as h15'. hfirst h15'. intro h16. apply in_faml_to_fam_impl in h16. destruct h16 as [l [h16 h17']]. red in h6. pose proof (h6 l al' (in_cons _ _ _ h16) (in_eq _ _) h17'); subst. contradiction. specialize (h15' hf0 h10). specialize (ih ll' (list_minus nat_eq_dec ln al') h1' h2' h11 h12 h13 h14 h17 (no_dup_list_minus nat_eq_dec _ _ h8) h15 h15'); auto. rewrite <- list_to_set_in_iff; auto. rewrite <- list_to_set_in_iff; auto. rewrite <- list_to_set_in_iff; auto. specialize (ih (remove list_nat_eq_dec al (al'::ll')) (list_minus nat_eq_dec ln al) h1' (no_dup_remove _ _ _ h2) h11 (no_dup_mem_remove _ _ _ h4) h13 (list_to_set_injable_in_remove _ _ _ h6)). hfirst ih. rewrite <- subtract_remove_compat; simpl. simpl in h7; rewrite <- h7. rewrite sub_add_same_nin; auto. rewrite <- list_to_set_in_iff; auto. specialize (ih hf0 (no_dup_list_minus _ _ _ h8) h15). pose proof (partition_faml_remove nat_eq_dec (al' :: ll') ln al) as h17. hfirst h17. right. pose proof (Add_intro2 _ (list_to_set ll) al) as h18. rewrite h7 in h18. destruct h18 as [al h18 | al h18]. rewrite list_to_set_in_iff; auto. destruct h18. contradict h4; auto. specialize (h17 hf1 h6 h10). rewrite (list_nat_eq_dec_eq (list_eq_dec nat_eq_dec)) in h17. specialize (ih h17). rewrite length_remove_no_dup in ih; auto. Qed. Lemma le_length_app_l : forall {T:Type} (l l':list T), length l <= length (l ++ l'). intros; rewrite app_length; auto with arith. Qed. Lemma le_length_app_r : forall {T:Type} (l l':list T), length l' <= length (l ++ l'). intros; rewrite app_length; auto with arith. Qed. Lemma lt_length_app : forall {T:Type} (l l':list T) (a:T), length l < length (l++(a::l')). intros; rewrite app_length; simpl; omega. Qed. Lemma lt_length_app_sing_dec : forall {T:Type} (l:list T) (a:T) (n:nat), n < length (l++(a::nil)) -> {n < length l} + {n = length l}. intros T l a n h1. rewrite app_length in h1. simpl in h1. rewrite S_compat in h1. apply lt_S_dec in h1; auto. Qed. Lemma lt_length_app_sing_S : forall {T:Type} (l:list T) (a:T) (n:nat), S n < length (l++(a::nil)) -> n < length l. intros T l a n h1. rewrite app_length in h1. simpl in h1. rewrite S_compat in h1. apply lt_S_n in h1; auto. Qed. Lemma le_length_app_sing_S : forall {T:Type} (l:list T) (a:T) (n:nat), S n <= length (l++(a::nil)) -> n <= length l. intros T l a n h1. rewrite app_length in h1. simpl in h1. rewrite S_compat in h1. apply le_S_n in h1; auto. Qed. Lemma lt_length_app_ge_impl : forall {T:Type} (l l':list T) (n:nat), n >= length l -> n < length (l++l') -> n - length l < length l'. intros ? ? ? ? ? h1; rewrite app_length in h1; omega. Qed. Lemma lt_length_app_dec : forall {T:Type} (l l' :list T) (n:nat), n < length (l ++ l') -> {n < length l} + {n - length l < length l'}. intros T l l' n h1. destruct (lt_eq_lt_dec n (length l)) as [h2 | h2]. destruct h2 as [h2 | h2]; subst. left; auto. right. rewrite minus_same. destruct (nil_dec' l') as [h2 | h2]; subst. rewrite app_nil_r in h1. apply lt_irrefl in h1; contradiction. apply not_nil_lt; auto. right. apply lt_length_app_ge_impl; auto with arith. Qed. Fixpoint list_singularize_aux {T:Type} (pfdt:type_eq_dec T) (l:list T) (acc:list T) := match l with | nil => nil | a::l' => if (in_dec pfdt a l') then (list_singularize_aux pfdt l' acc) else (a::(list_singularize_aux pfdt l' (a::acc))) end. (*This returns the list without the duplicate entries, in the same order as each new element appearing starting backwards, i.e. [list_singularize (3::4::2::4::2::3::3::3::nil) = 4::2::3].*) Definition list_singularize {T:Type} (pfdt:type_eq_dec T) (l:list T) := list_singularize_aux pfdt l nil. Lemma list_singularize_not_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (acc:list T), (forall (x:T), In x acc -> ~In x l) -> list_singularize_aux pfdt l acc = list_singularize_aux pfdt l nil. intros T h0 l. induction l as [|a l h1]. intros. simpl. auto. intros acc h2. simpl. destruct (in_dec h0 a l) as [h3 | h4]. assert (h4:forall x:T, In x acc -> ~ In x l). intros x h4. specialize (h2 x h4). simpl in h2. tauto. apply h1; auto. assert (h5: forall x:T, In x (a::acc) -> ~In x l). intros x h6. destruct h6 as [|h6]. subst. auto. specialize (h2 _ h6). simpl in h2. tauto. pose proof (h1 _ h5) as h7. rewrite h7. assert (h8:forall x:T, In x (a::nil) -> ~In x l). intros x h9. destruct h9. subst. auto. contradiction. rewrite (h1 _ h8). reflexivity. Qed. Lemma list_singularize_nil : forall {T:Type} (pfdt:type_eq_dec T), list_singularize pfdt nil = nil. intros; simpl; auto. Qed. Lemma list_singularize_nil_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), list_singularize pfdt l = nil <-> l = nil. unfold list_singularize; intros T h1 l; split; intro h2. revert h2. induction l as [|a l ih]; simpl; auto. lr_if_goal; fold hlr; destruct hlr as [hl | hr]; simpl. intro h2. apply ih in h2; subst. contradict hl. intro; discriminate. subst; auto. Qed. (* Returns the list [l] with [x] at the head of it, and all other occurrences of [x] removed.*) Definition new_head {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) : list T := x::remove pfdt x l. Lemma list_singularize_new_head_not_in_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), ~In x l -> list_singularize pfdt l = remove pfdt x (list_singularize pfdt l). unfold list_singularize; intros T h0 l. induction l as [|a l h1]. intros x h1. simpl. auto. intros x h2. simpl. simpl in h2. destruct (in_dec h0 a l) as [h4 | h5]. apply h1; auto; simpl. pose proof (remove_cons_neq h0 (list_singularize_aux h0 l (a::nil)) x a) as h6. hfirst h6. contradict h2; left; auto. specialize (h6 hf). rewrite h6. f_equal. rewrite list_singularize_not_in. specialize (h1 x). hfirst h1. contradict h2; right; auto. specialize (h1 hf0). rewrite h1; auto. f_equal; auto. intros y hy; destruct hy; subst; auto. Qed. Lemma subtract_new_head_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), Subtract (list_to_set l) x = list_to_set (remove pfdt x l). intros T h0 l. induction l as [|a l h1]. intros x. simpl. apply Extensionality_Ensembles. red. split. red. intros y h1. destruct h1; contradiction. auto with sets. intro x. simpl. destruct (h0 x a) as [h2 | h3]. subst. pose proof (in_dec h0 a l) as h4. destruct h4 as [h4l | h4r]. rewrite list_to_set_in_iff in h4l. rewrite sub_add_same_in. apply h1. auto. rewrite list_to_set_in_iff in h4r. rewrite sub_add_same_nin. rewrite nin_remove_eq; auto. rewrite list_to_set_in_iff. auto. auto. simpl. rewrite sub_add_comm. rewrite h1. auto. auto. Qed. Lemma add_list_to_set_to_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (A:Ensemble T) (x:T), ~ Inn A x -> Add A x = list_to_set (list_singularize pfdt l) -> A = list_to_set (remove pfdt x (list_singularize pfdt l)). unfold list_singularize; intros T h0 l. induction l as [|a l h1]. intros A x h1 h2. simpl in h2. pose proof (add_not_empty A x). contradiction. intros A x h2 h3. simpl in h3. destruct (in_dec h0 a l) as [h4 | h5]. rewrite (h1 _ x). simpl. destruct (in_dec h0 a l) as [h5|?]. reflexivity. contradiction. assumption. assumption. simpl. destruct (in_dec h0 a l) as [h6|h7]. contradiction. simpl. destruct (h0 x a) as [h8 | h9]. subst. assert (h6:forall c:T, In c (a::nil) -> ~In c l). intros y h8. destruct h8; [subst | contradiction]; auto. simpl in h3. pose proof (add_nin_sub_compat _ _ _ h2 h3) as h8. rewrite (subtract_new_head_compat h0) in h8. auto. assert (h6:forall c:T, In c (a::nil) -> ~In c l). intros y h8. destruct h8; [subst | contradiction]; auto. rewrite list_singularize_not_in in h3. rewrite list_singularize_not_in. assert (h10:Subtract (Add A x) x = Subtract (list_to_set (a::list_singularize_aux h0 l nil)) x). f_equal. auto. rewrite sub_add_same_nin in h10. rewrite h10. simpl. rewrite sub_add_comm. rewrite (subtract_new_head_compat h0). auto. auto. auto. intros y h8. destruct h8; [subst | contradiction]; auto. intros y h8. destruct h8; [subst | contradiction]; auto. Qed. Lemma remove_list_singularize_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), remove pfdt x (list_singularize pfdt l) = list_singularize pfdt (remove pfdt x l). unfold list_singularize; intros T h0 l. induction l as [|a l h1]. intros x. simpl. auto. intros x. simpl. destruct (in_dec h0 a l) as [h2 | h3]; destruct (h0 x a) as [h4 | h5]. subst. auto. apply neq_sym in h5. simpl. destruct (in_dec h0 a (remove h0 x l)) as [h6 | h7]; auto. pose proof (nin_remove_eq_nin h0 l a x h7) as h8. destruct h8; contradiction. subst. simpl. rewrite eq_dec_eq at 1; auto. rewrite list_singularize_not_in; auto. intros x hx; destruct hx; subst; auto. simpl. rewrite eq_dec_neq_r at 1; auto. destruct (in_dec h0 a (remove h0 x l)). rewrite in_remove_iff in i. destruct i; contradiction. rewrite list_singularize_not_in. rewrite (list_singularize_not_in h0 (remove h0 x l) (a::nil)); auto. f_equal; auto. intros c hc; destruct hc; subst; auto. intros c hc; destruct hc; subst; auto. Grab Existential Variables. auto. Qed. Lemma list_singularize_in_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), In x l <-> In x (list_singularize pfdt l). unfold list_singularize; intros T h0 l. induction l as [|a l h1]. (* nil *) split. (* -> *) intro; contradiction. simpl. tauto. (* <- *) (* split *) intro x. split. (* -> *) intros h2. destruct h2 as [h2 | h3]. subst. simpl. destruct (in_dec h0 x l) as [h2 | h3]. apply h1. assumption. constructor. reflexivity. simpl. destruct (in_dec h0 a l) as [h4 | h5]. apply h1. assumption. rewrite list_singularize_not_in. right. apply h1. assumption. intros y h6. destruct h6; try subst. assumption. contradiction. intro h2. simpl in h2. destruct (in_dec h0 a l) as [h3 | h4]. right. apply h1. assumption. rewrite list_singularize_not_in in h2. destruct h2 as [h2a | h2b]. subst. constructor. reflexivity. right. apply h1. assumption. intros y h3. destruct h3; subst. assumption. contradiction. Qed. Lemma list_singularize_lrep : forall {T:Type} (pfdt:type_eq_dec T) (x:T) (n:nat), list_singularize pfdt (lrep x (S n)) = x::nil. unfold list_singularize; intros T h1 x n. revert x. induction n as [|n ih]; simpl; auto. intro x. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. specialize (ih x). simpl in ih. lr_if_goal; fold hlr; fold hlr in ih; destruct hlr as [hl' | hr']; auto. f_equal. contradict hr. left; auto. Qed. Lemma list_singularize_sing : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), list_to_set l = Singleton x -> list_singularize pfdt l = x::nil. unfold list_singularize; intros T h1 l x h2. pose proof h2 as h2'. apply impl_all_sing in h2. apply all_sing_lrep in h2. rewrite h2. pose proof (In_singleton _ x) as h3. rewrite <- h2' in h3. rewrite <- list_to_set_in_iff in h3. pose proof (in_not_nil _ _ h3) as h4. apply S_length_tl in h4. rewrite <- h4. apply list_singularize_lrep. Qed. Lemma remove_app : forall {T:Type} (pfdt:type_eq_dec T) (a:T) (l l':list T), remove pfdt a l ++ (remove pfdt a l') = remove pfdt a (l ++ l'). intros T h0 a l. revert a. induction l as [|b l h1]. simpl. intros; auto. intros a l'. simpl. destruct (h0 a b) as [h3 | h4]. subst. apply h1. simpl. rewrite h1. reflexivity. Qed. Lemma list_to_set_singularize_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), list_to_set l = list_to_set (list_singularize pfdt l). unfold list_singularize; intros T h0 l. induction l as [|a l h1]. simpl. reflexivity. simpl. destruct (in_dec h0 a l) as [h2 | h3]. rewrite list_to_set_in_iff in h2. rewrite h1 in h2. rewrite h1. rewrite in_add_eq. reflexivity. assumption. simpl. rewrite h1. assert (h9:forall c:T, In c (a::nil) -> ~In c l). intros y h10. destruct h10; [subst | contradiction]; auto. rewrite (list_singularize_not_in _ _ _ h9). reflexivity. Qed. Lemma list_singularize_rev_cons : forall {T:Type} (pfdt:type_eq_dec T) (a:T) (l:list T), list_singularize pfdt(l++(a::nil)) = (remove pfdt a (list_singularize pfdt l))++(a::nil). unfold list_singularize; intros T hdt a l. revert a. induction l as [|b l h0]. intro a. simpl. reflexivity. simpl. intro a. destruct (in_dec hdt b (l++a::nil)) as [h1|h2]; destruct (in_dec hdt b l) as [h3|h4]. rewrite h0. reflexivity. pose proof (in_app_or _ _ _ h1) as h5. destruct h5 as [h5a | h5b]. contradiction. destruct h5b; subst. assert (h6:forall x:T, In x (b::nil) -> ~In x l). intros x h7. destruct h7; subst. assumption. contradiction. pose proof (list_singularize_not_in hdt _ _ h6) as h7. rewrite h7. simpl. destruct (hdt b b) as [h8 | h9]. apply h0. contradict h9. reflexivity. contradiction. contradict h2. apply in_or_app. left; assumption. assert (h5:a <> b). intro h5. subst. contradict h2. apply in_or_app. right; constructor. reflexivity. rewrite list_singularize_not_in. pose proof (list_singularize_not_in hdt l (b::nil)) as h6. assert (h7:(forall x:T, In x (b::nil) -> ~In x l)). intros x h7. destruct h7; subst. assumption. contradiction. specialize (h6 h7); clear h7. rewrite h6. pose proof (app_comm_cons (list_singularize_aux hdt l nil) (a::nil) b) as h7. rewrite remove_cons_neq. rewrite <- app_comm_cons. rewrite <- h0. reflexivity. assumption. intros x h6. destruct h6; subst. assumption. contradiction. Qed. Lemma list_singularize_in_cons : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), In a l -> list_singularize pfdt (a::l) = list_singularize pfdt l. unfold list_singularize; intros T h0 l. induction l as [|b l h1]; simpl. intros; contradiction. intros a h2. destruct h2 as [h2 |h2]; subst. destruct (h0 a a) as [h2 | h2]. reflexivity. contradict h2. reflexivity. destruct (h0 b a) as [h3 |h3]. reflexivity. destruct (in_dec h0 a l) as [h4 | h4]. reflexivity. contradiction. Qed. Lemma list_singularize_nin : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf: ~In a l), list_singularize pfdt (a::l) = a::list_singularize pfdt l. unfold list_singularize; intros T h0 l a h1; simpl. destruct in_dec as [h2 | h2]. contradiction. rewrite list_singularize_not_in; auto. intros x h3. destruct h3; subst; auto. Qed. Lemma no_dup_list_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), NoDup (list_singularize pfdt l). intros T h0 l. induction l as [|a l h1]. simpl. constructor. simpl. destruct (in_dec h0 a l) as [h2 | h3]. rewrite list_singularize_in_cons; auto. rewrite list_singularize_nin. constructor. rewrite <- list_singularize_in_iff. assumption. assumption. assumption. Qed. Lemma list_singularize_no_dup : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), NoDup l <-> list_singularize pfdt l = l. intros T h0 l. split. (* -> *) intro h1. induction h1 as [|x l h2 h3 h4]. simpl. reflexivity. simpl. destruct (in_dec h0 x l) as [h6 | h7]. contradiction. rewrite list_singularize_nin. rewrite h4. reflexivity. assumption. intro h1. rewrite <- h1. apply no_dup_list_singularize. Qed. Lemma incl_list_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), incl (list_singularize pfdt l) l. intros T h1 l; red; intros x h2. rewrite <- list_singularize_in_iff in h2; auto. Qed. Lemma incl_list_singularize_compat : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), incl l l' -> incl (list_singularize pfdt l) (list_singularize pfdt l'). intros T h1 l l' h2; red; intros x h3. rewrite <- list_singularize_in_iff in h3. rewrite <- list_singularize_in_iff; auto. Qed. Lemma finite_set_list_no_dup : forall {T:Type} (E:Ensemble T), Finite E -> exists (l:list T), E = list_to_set l /\ NoDup l. intros T E h1. induction h1 as [|E h2 h3 a h5]. exists nil; simpl. split; auto; try constructor. destruct h3 as [l h3]. destruct h3 as [h3 h4]. exists (a::l); simpl. subst. split; auto. constructor. rewrite list_to_set_in_iff; auto. assumption. Qed. Lemma length_list_singularize_cons_gt_0 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), 0 < length (list_singularize pfdt (a::l)). intros T h0 l. induction l as [|b l h1]; simpl; auto with arith. intro a. destruct (h0 b a) as [h2 | h2]; subst. destruct (in_dec h0 a l) as [h2 | h2]. rewrite list_singularize_in_cons. rewrite list_singularize_in_cons. rewrite list_singularize_in_iff in h2. apply in_not_nil in h2. apply not_nil_lt in h2; auto. apply h2. assumption. left; auto. rewrite list_singularize_in_cons; simpl. rewrite list_singularize_nin; simpl; auto with arith. left; auto. destruct (in_dec h0 a l) as [h3 | h3]. destruct (in_dec h0 b l) as [h4 | h4]. rewrite list_singularize_in_cons. rewrite list_singularize_in_cons. rewrite (list_singularize_in_iff h0) in h3. apply in_not_nil, not_nil_lt in h3; auto. assumption. right; auto. rewrite list_singularize_in_cons. rewrite (list_singularize_in_iff h0) in h3. apply in_not_nil, not_nil_lt in h3; auto. right; auto. rewrite list_singularize_nin. simpl; auto with arith. intro h4. destruct h4; subst. contradict h2; auto. contradiction. Qed. Lemma list_singularize_nin_app_nil : forall {T:Type} (pfdt:type_eq_dec T) (x:T) (l:list T), ~In x l -> list_singularize pfdt (l ++ x::nil) = list_singularize pfdt l ++ x :: nil. intros T h1 x l. revert x. induction l as [|a l ih]; simpl. intros; auto. intros x h2. destruct (in_dec h1 a (l++x::nil)) as [hl | hr]. destruct (in_dec h1 a l) as [hl' | hr']. destruct (in_dec h1 x l) as [h3 | h3]. contradict h2. right; auto. rewrite list_singularize_in_cons. rewrite list_singularize_in_cons. apply ih; auto. assumption. apply in_app_l; auto. apply in_app_or in hl. destruct hl as [|hl]. contradiction. destruct hl; subst. contradict h2. left; auto. contradiction. rewrite list_singularize_nin; auto. destruct (in_dec h1 a l) as [hl' | hr']. contradict hr. apply in_app_l; auto. rewrite list_singularize_nin; auto. rewrite <- app_comm_cons. f_equal; auto. Qed. Lemma list_singularize_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), list_singularize pfdt (remove pfdt x l) = remove pfdt x (list_singularize pfdt l). intros T h1 l. induction l as [|a l ih]; simpl; auto. intros x h2. destruct (h1 x a) as [|h3]; subst. destruct (in_dec h1 a l) as [h3 | h3]. rewrite list_singularize_in_cons; auto. rewrite list_singularize_nin; auto. rewrite <- remove_not_in'; auto. rewrite remove_cons_eq. rewrite <- remove_not_in'; auto. rewrite <- list_singularize_in_iff; auto. destruct h2 as [|h2]. subst; contradict h3; auto. destruct (in_dec h1 a (remove h1 x l)) as [h4 | h4]. rewrite list_singularize_in_cons. pose proof h4 as h4'. rewrite in_remove_iff in h4'. destruct h4' as [h4' h5]. rewrite list_singularize_in_cons; auto. assumption. rewrite list_singularize_nin; auto. assert (h5:~In a l). contradict h4. rewrite in_remove_iff; auto. rewrite list_singularize_nin; auto. rewrite remove_cons_neq; auto. f_equal; auto. Qed. Lemma list_singularize_in_app_nil : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), In a l -> list_singularize pfdt (l++a::nil) = list_singularize pfdt (remove pfdt a l) ++ a::nil. intros T h1 l. induction l as [|b l ih]; simpl. intros; contradiction. intros a h2. lr_if_goal; fold hlr; destruct hlr as [hl | hr]; subst. rewrite list_singularize_in_cons. destruct (in_dec h1 b l) as [h3 | h3]. apply ih; auto. rewrite <- remove_not_in'; auto. rewrite list_singularize_nin_app_nil; auto. apply in_app_r; constructor; auto. destruct h2 as [h2 | h2]. subst. contradict hr; auto. simpl. destruct (in_dec h1 b l) as [hl' | hr']. rewrite list_singularize_in_cons; auto. rewrite list_singularize_in_cons; auto. rewrite in_remove_iff. split; auto. apply in_app_l; auto. rewrite list_singularize_nin. rewrite list_singularize_nin. simpl. f_equal; auto. intro h3. rewrite in_remove_iff in h3. destruct h3 as [h3 h4]. contradiction. intro h3. apply in_app_or in h3. destruct h3 as [|h3]; subst. contradiction. destruct h3; subst. contradict hr; auto. contradiction. Qed. Lemma list_to_sets_eq_length : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), list_to_set l = list_to_set l' -> length (list_singularize pfdt l) = length (list_singularize pfdt l'). intros T h1 l. induction l as [|a l ih]; simpl. intros ? h2. symmetry in h2. apply empty_set_nil in h2; subst; simpl; auto. intros l'. destruct l' as [|a' l']. intro h2. apply add_not_empty in h2. contradiction. simpl. intro h2. destruct (h1 a a') as [|h3]; subst. destruct (in_dec h1 a' l) as [h4 | h4], (in_dec h1 a' l') as [h5 | h5]. rewrite list_singularize_in_cons; auto. rewrite list_singularize_in_cons; auto. apply ih. do 2 rewrite in_add_eq in h2; auto. rewrite <- list_to_set_in_iff; auto. rewrite <- list_to_set_in_iff; auto. rewrite <- list_to_set_in_iff; auto. rewrite in_add_eq in h2. simpl. pose proof (ih (a'::l') h2) as h6. simpl in h6. rewrite list_singularize_in_cons; auto. rewrite <- list_to_set_in_iff; auto. rewrite list_singularize_nin; auto. rewrite list_singularize_in_cons; auto. simpl. rewrite (in_add_eq (list_to_set l')) in h2. pose proof (f_equal (fun S => Subtract S a') h2) as h6. simpl in h6. rewrite sub_add_same_nin in h6. specialize (ih (remove h1 a' l')). hfirst ih. rewrite <- subtract_remove_compat; auto. specialize (ih hf). rewrite list_singularize_remove, length_remove_no_dup in ih. pose proof h5 as h5'. rewrite (list_singularize_in_iff h1) in h5'. apply in_not_nil in h5'. apply not_nil_lt in h5'. omega. apply no_dup_list_singularize. rewrite <- list_singularize_in_iff; auto. assumption. rewrite <- list_to_set_in_iff; auto. rewrite <- list_to_set_in_iff; auto. rewrite list_singularize_nin; auto. rewrite list_singularize_nin; auto; simpl. f_equal; apply ih. pose proof (f_equal (fun S => Subtract S a') h2) as h6. simpl in h6. rewrite sub_add_same_nin in h6. rewrite sub_add_same_nin in h6; auto. rewrite <- list_to_set_in_iff; auto. rewrite <- list_to_set_in_iff; auto. destruct (in_dec h1 a l) as [h4 | h4], (in_dec h1 a' l') as [h5 | h5]. do 2 rewrite in_add_eq in h2; auto. rewrite list_singularize_in_cons; auto. rewrite list_singularize_in_cons; auto. rewrite <- list_to_set_in_iff; auto. rewrite <- list_to_set_in_iff; auto. rewrite <- list_to_set_in_iff; auto. rewrite list_singularize_in_cons; auto. rewrite list_singularize_nin; auto. simpl. rewrite in_add_eq in h2. specialize (ih (a'::l')). apply ih in h2. rewrite list_singularize_nin in h2; auto. rewrite <- list_to_set_in_iff; auto. rewrite list_singularize_nin; auto. rewrite list_singularize_in_cons; auto. simpl. rewrite (in_add_eq (list_to_set l')) in h2. pose proof (f_equal (fun S => Subtract S a) h2) as h6. simpl in h6. rewrite sub_add_same_nin in h6. specialize (ih (remove h1 a l')). hfirst ih. rewrite <- subtract_remove_compat; auto. specialize (ih hf). rewrite ih. rewrite list_singularize_remove. rewrite length_remove_no_dup. rewrite (list_singularize_in_iff h1) in h5. apply in_not_nil in h5. apply not_nil_lt in h5. omega. apply no_dup_list_singularize. rewrite <- list_singularize_in_iff; auto. pose proof (Add_intro2 _ (list_to_set l) a) as h7. rewrite h2 in h7. rewrite list_to_set_in_iff; auto. pose proof (Add_intro2 _ (list_to_set l) a) as h7. rewrite h2 in h7. rewrite list_to_set_in_iff; auto. rewrite <- list_to_set_in_iff; auto. rewrite <- list_to_set_in_iff; auto. rewrite list_singularize_nin; auto. rewrite list_singularize_nin; auto. simpl; f_equal. pose proof (Add_intro2 _ (list_to_set l) a) as h6. rewrite h2 in h6. inversion h6; subst. rewrite <- list_to_set_in_iff in H. pose proof (Add_intro2 _ (list_to_set l') a') as h6'. rewrite <- h2 in h6'. inversion h6; subst. rewrite <- list_to_set_in_iff in H0. simpl. pose proof (f_equal (fun S => Subtract S a) h2) as h7. simpl in h7. rewrite sub_add_same_nin in h7. specialize (ih (remove h1 a (a'::l'))). hfirst ih. rewrite <- subtract_remove_compat. simpl; auto. specialize (ih hf). rewrite list_singularize_remove in ih. rewrite list_singularize_nin in ih; auto. rewrite length_remove_no_dup in ih. simpl in ih; auto. constructor. rewrite <- list_singularize_in_iff; auto. apply no_dup_list_singularize. right. rewrite <- list_singularize_in_iff; auto. right; auto. rewrite <- list_to_set_in_iff; auto. destruct H0. contradiction. destruct H. contradict h3; auto. Qed. Lemma all_sing_list_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), l <> nil -> all_sing l a -> list_singularize pfdt l = a::nil. intros T h0 l. induction l as [|c l ih]; simpl. intros ? h1. contradict h1; auto. intros a h1 h2. destruct (nil_dec' l) as [hn | hn]; subst. rewrite (list_singularize_sing h0 _ c). f_equal. red in h2. apply h2. left; auto. simpl; auto. rewrite add_empty_sing; auto. destruct (in_dec h0 c l) as [h3 | h3]. rewrite list_singularize_in_cons; auto. apply ih; auto. red in h2; red; intros; apply h2; auto. right; auto. rewrite list_singularize_nin; auto. pose proof h2 as h2'. red in h2'. specialize (h2' c (in_eq _ _)); subst. apply all_sing_cons_in in h2; auto. contradiction. Qed. Definition singular_fun_in_dep {T U:Type} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l U) : fun_in_dep (list_singularize pfdt l) U := fun x pf => f x (iff2 (list_singularize_in_iff pfdt l x) pf). Lemma singular_fun_in_dep_compat : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l U) (x:T) (pf:In x (list_singularize pfdt l)), singular_fun_in_dep pfdt f x pf = f x (iff2 (list_singularize_in_iff pfdt l x) pf). unfold singular_fun_in_dep; intros; auto. Qed. Lemma singular_fun_in_dep_compat' : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l U) (x:T) (pf:In x l), f x pf = singular_fun_in_dep pfdt f x (iff1 (list_singularize_in_iff pfdt l x) pf). unfold singular_fun_in_dep; intros; f_equal; apply proof_irrelevance. Qed. Definition fun_in_dep_in_ens {T U:Type} {l:list T} (f:fun_in_dep l U) : fun_in_ens (list_to_set l) U := fun x pf => f x (iff2 (list_to_set_in_iff l x) pf). Definition fun_in_dep_to_sig {T U:Type} {l:list T} (f:fun_in_dep l U) : sig_set (list_to_set l) -> U := fun x => f (proj1_sig x) (iff2 (list_to_set_in_iff l (proj1_sig x)) (proj2_sig x)). Lemma fun_in_dep_in_ens_compat : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l U) (x:T) (pf:In x l), fun_in_dep_in_ens f x (iff1 (list_to_set_in_iff l x) pf) = f x pf. unfold fun_in_dep_in_ens; intros; f_equal; apply proof_irrelevance. Qed. Lemma fun_in_dep_to_sig_compat : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l U) (x:T) (pf:In x l), fun_in_dep_to_sig f (exist _ _ (iff1 (list_to_set_in_iff l x) pf)) = f x pf. unfold fun_in_dep_to_sig; intros; f_equal; apply proof_irrelevance. Qed. Definition sing_unq {T:Type} (ll:faml T) : Prop := forall l l', In l ll -> In l' ll -> list_to_set l = list_to_set l' -> l = l'. Lemma list_power_no_dup : forall {T U:Type} (la:list T) (lb:list U) (lp:list (T*U)), NoDup la -> NoDup lb -> (In lp (list_power la lb)) -> NoDup lp. intros T U la. induction la as [|a la h1]. simpl. intros lb lp h1 h2 h3. destruct h3 as [h3l | h3r]. rewrite <- h3l. constructor. contradiction. intros lb lp h2 h3 h4. rewrite <- app_nil_l in h2. pose proof (NoDup_remove_1 _ _ _ h2) as h5. rewrite app_nil_l in h5. simpl in h4. rewrite in_flat_map in h4. destruct h4 as [lp' h6]. destruct h6 as [h6l h6r]. specialize (h1 _ _ h5 h3 h6l). rewrite in_map_iff in h6r. destruct h6r as [b h7]. destruct h7 as [h7l h7r]. rewrite <- h7l. constructor. intro h8. rewrite app_nil_l in h2. inversion h2. pose proof (in_list_power1 _ h6l _ h8) as h9. simpl in h9. contradiction. assumption. Qed. Lemma no_dup_map_in_dep_inj : forall {T U:Type} {l:list T} (f:fun_in_dep l U), NoDup l -> inj_dep f -> NoDup (map_in_dep f). intros T U l. induction l as [|a l ih]; simpl. intros; constructor. intros f h1 h2. pose proof (no_dup_cons _ _ h1) as h3. pose proof (no_dup_cons_nin _ _ h1) as h4. specialize (ih (fun_from_cons f) h3 (inj_from_cons _ h2)). constructor. contradict h4. rewrite in_map_in_dep_iff in h4. destruct h4 as [a' [h4 h5]]. unfold fun_from_cons in h5. apply h2 in h5; subst; auto. assumption. Qed. Lemma length_list_singularize_map_in_dep_inj : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:fun_in_dep l U), inj_dep f -> length (list_singularize pfdu (map_in_dep f)) = length (list_singularize pfdt l). intros T U l h1 h2. induction l as [|a l ih]; simpl; auto. intros f h3. destruct (in_dec h2 (f a (in_eq _ _)) (map_in_dep (fun_from_cons f))) as [hl | hr]. rewrite list_singularize_in_cons; auto. destruct (in_dec h1 a l) as [hl' | hr']. rewrite list_singularize_in_cons; auto. rewrite in_map_in_dep_iff in hl. destruct hl as [a' h4]. destruct h4 as [h4 h5]. apply ih. apply inj_from_cons; auto. red in h3. rewrite in_map_in_dep_iff in hl. destruct hl as [a' hl]. destruct hl as [hl h4]. unfold fun_from_cons in h4. apply h3 in h4; subst; contradiction. rewrite list_singularize_nin; auto. destruct (in_dec h1 a l) as [hl' | hr']. rewrite list_singularize_in_cons; auto. contradict hr. rewrite in_map_in_dep_iff. exists a, hl'. unfold fun_from_cons; f_equal; apply proof_irrelevance. rewrite list_singularize_nin; auto. simpl; f_equal; apply ih. apply inj_from_cons; auto. Qed. Lemma no_dup_app : forall {T:Type} (l l':list T), NoDup l -> NoDup l' -> Intersection (list_to_set l) (list_to_set l') = Empty_set _ -> NoDup (l ++ l'). intros T l1. induction l1 as [|a l1 h1]. simpl. intros. assumption. intros l2 h2 h3 h4. rewrite <- app_comm_cons. constructor. intro h5. pose proof (in_app_or _ _ _ h5) as h6. destruct h6 as [h6l | h6r]. inversion h2. contradiction. rewrite list_to_set_in_iff in h6r. pose proof (in_eq a l1) as h7. rewrite list_to_set_in_iff in h7. pose proof (Intersection_intro _ _ _ _ h7 h6r) as h8. rewrite h4 in h8. contradiction. assert (h5:Included (list_to_set l1) (list_to_set (a::l1))). red. simpl. intros x h5. left. assumption. pose proof (intersection_preserves_inclusion _ _ (list_to_set l2) h5) as h6. rewrite Intersection_commutative in h6. rewrite (Intersection_commutative _ (list_to_set l2) (list_to_set (a::l1))) in h6. rewrite h4 in h6. assert (h7:(Intersection (list_to_set l1) (list_to_set l2)) = (Empty_set T)). apply Extensionality_Ensembles. red; split; auto with sets. inversion h2. apply h1; assumption. Qed. Lemma no_dup_app_sing_nin : forall {T:Type} (l:list T) (a:T), NoDup (l++a::nil) -> ~In a l. intros T l. induction l as [|b l ih]; simpl. intros; auto. intros a h1 h2. pose proof (no_dup_cons _ _ h1) as h3. pose proof (no_dup_cons_nin _ _ h1) as h4. specialize (ih _ h3). destruct h2 as [|h2]; subst. contradict h4. apply in_app_r; left; auto. contradiction. Qed. Lemma no_dup_seg_list : forall (n:nat), NoDup (seg_list n). intro n. induction n as [|n ih]; simpl. rewrite seg_list0. constructor. rewrite seg_list_S. apply no_dup_app; auto. apply no_dup_sing. simpl. rewrite add_empty_sing. apply intersection_seg_sing; auto with arith. Qed. Lemma no_dup_seq : forall (m n:nat), NoDup (seq m n). intros m n. revert m. induction n as [|n ih]; simpl. intros; constructor. intro m. constructor. intro h1. rewrite in_seq_iff in h1. omega. apply ih. Qed. Lemma no_dup_map_seg_inj : forall {U:Type} {n:nat} (f:seg_fun n U), inj_seg f -> NoDup (map_seg f). intros U n. induction n as [|n ih]; simpl. intros; constructor. intros f h1. apply no_dup_app. apply ih. apply inj_from_S; auto. constructor. intro; contradiction. constructor. simpl. rewrite add_empty_sing. rewrite <- nin_iff. intro h2. rewrite <- list_to_set_in_iff in h2. rewrite in_map_seg_iff in h2. destruct h2 as [m [h2 h3]]. rewrite seg_fun_from_S_compat in h3. apply h1 in h3; subst. apply lt_irrefl in h2; auto. Qed. Lemma list_to_set_eq_seg : forall (l:list nat) (n:nat), NoDup l -> list_to_set l = seg n -> n = length l. intros l n. revert l. induction n as [|n ih]; simpl. intros ? ? h1. rewrite seg_0_eq in h1. apply empty_set_nil in h1; subst; auto. intros l h1 h2. destruct (nil_dec' l) as [|h3]; subst. simpl in h2. pose proof (in_seg _ _ (lt_0_Sn n)) as h3. rewrite <- h2 in h3. contradiction. pose proof (lt_n_Sn n) as h4. apply in_seg in h4. rewrite <- h2 in h4. rewrite <- list_to_set_in_iff in h4. specialize (ih (remove nat_eq_dec n l)). hfirst ih. apply no_dup_remove; auto. specialize (ih hf). hfirst ih. rewrite <- subtract_remove_compat. rewrite h2. rewrite seg_Sn. rewrite sub_add_same_nin; auto. apply nin_seg. specialize (ih hf0). rewrite ih. rewrite length_remove_no_dup; auto. apply not_nil_lt in h3. rewrite <- (S_pred _ _ h3); auto. Qed. Lemma list_to_set_eq_seg_lt' : forall (l:list nat) (m n:nat), NoDup l -> list_to_set l = seg n -> In m l -> m < length l. intros l m n h1 h2 h3. pose proof (list_to_set_eq_seg_lt l m n h2 h3) as hl. rewrite <- list_to_set_seg_list_eq in h2. pose proof (incl_no_dup_le nat_eq_dec (seg_list n) l (no_dup_seg_list n) h1) as h4. hfirst h4. red; intros a h5. rewrite list_to_set_in_iff in h5. rewrite <- h2 in h5. rewrite <- list_to_set_in_iff in h5; auto. apply h4 in hf. rewrite length_seg_list in hf. eapply lt_le_trans. apply hl. assumption. Qed. Lemma no_dup_partition_faml_nat_seg_list_to_set_eq : forall (ll ll':faml nat) (m n:nat), NoDup ll -> NoDup ll' -> no_dup_mem ll -> no_dup_mem ll' -> list_to_set_injable (fun lx => In lx ll) -> list_to_set_injable (fun lx => In lx ll') -> list_to_set ll = list_to_set ll' -> partition_faml ll (seg_list m) -> partition_faml ll' (seg_list n) -> m = n. intros ll ll' m n h1 h2 h3 h4 h5 h6 h7 h8 h9. pose proof (no_dup_partition_faml_nat_list_to_set_eq ll ll' (seg_list m) (seg_list n) h1 h2 h3 h4 h5 h6 h7 (no_dup_seg_list _) (no_dup_seg_list _) h8 h9) as h10. do 2 rewrite length_seg_list in h10; auto. Qed. Lemma no_dup_list_prod : forall {T U:Type} (la:list T) (lb:list U), NoDup la -> NoDup lb -> NoDup (list_prod la lb). intros T U la. induction la as [|a la h1]. simpl. intros lb h1 h2. constructor. simpl. intros lb h2 h3. inversion h2. subst. specialize (h1 _ H2 h3). assert (h4: Injective (fun y:U => (a, y))). red. intros x1 x2 h4. inversion h4. subst. reflexivity. pose proof (no_dup_map_inj _ _ h3 h4) as h5. apply no_dup_app; try assumption. apply Extensionality_Ensembles. red. split. red. intros x h6. destruct h6 as [pr h6l h6r]. rewrite <- list_to_set_in_iff in h6l. rewrite <- list_to_set_in_iff in h6r. rewrite in_map_iff in h6l. destruct h6l as [x h7]. destruct h7; subst. rewrite in_prod_iff in h6r. destruct h6r; contradiction. auto with sets. Qed. Lemma list_of_lists_seqs_no_dup : forall {T:Type} (ll:faml T), no_dup_faml ll -> NoDup (list_of_lists_seqs ll). intros T ll. induction ll as [|al ll ih]; simpl. intro; apply no_dup_sing. intro h2. hfirst ih. red in h2; red. intros; apply h2; right; auto. specialize (ih hf). apply no_dup_map_inj. apply no_dup_list_prod. red in h2. apply h2; left; auto. assumption. red. intros p1 p2 p3. inversion p3. rewrite surjective_pairing. rewrite <- H0, <- H1. rewrite <- surjective_pairing; auto. Qed. Lemma list_of_lists_seqs_length : forall {T:Type} (ll:faml T) (l:list T), In l (list_of_lists_seqs ll) -> length l = length ll. intros T ll. induction ll as [|al ll ih]; simpl. intros; subst; auto. destruct H; subst; auto; try contradiction. intros l h1. rewrite in_map_iff in h1. destruct h1 as [p h1]. destruct h1 as [h1 h2]. subst. simpl. f_equal. apply ih. destruct p as [x y]. rewrite in_prod_iff in h2. simpl. destruct h2; auto. Qed. Definition ens_list_to_fam {T:Type} (E:Ensemble (list T)) : Family T := Im E (@list_to_set _). Definition list_reps {T:Type} (E:Ensemble T) : Ensemble (list T) := [l:list T | E = list_to_set l /\ NoDup l]. Lemma in_list_reps_eq : forall {T:Type} (E:Ensemble T) (l:list T), Inn (list_reps E) l -> list_to_set l = E. intros T E l h1. destruct h1 as [h1]. destruct h1; auto. Qed. Lemma list_reps_empty : forall (T:Type), list_reps (Empty_set T) = Singleton nil. intro T. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1]. destruct h1 as [h1 h2]. symmetry in h1. apply empty_set_nil in h1. subst. constructor. destruct h1. constructor; simpl. split; try constructor; auto. Qed. Lemma list_reps_add : forall {T:Type} (E:Ensemble T) (a:T), ~Inn E a -> list_reps (Add E a) = FamilyUnion ( Im (list_reps E) (fun l => [m:list T | NoDup m /\ In a m /\ remove eq_dec a m = l])). intros T E a h1. apply Extensionality_Ensembles; red; split; red; intros l h3. destruct h3 as [h3]. destruct h3 as [h3 h4]. assert (h5:Inn (list_reps E) (remove eq_dec a l)). constructor. split. rewrite <- subtract_remove_compat. rewrite <- h3. rewrite sub_add_same_nin; auto. apply no_dup_remove; auto. apply family_union_intro with ((fun l0 : list T => [m : list T | NoDup m /\ In a m /\ remove eq_dec a m = l0]) (remove eq_dec a l)). apply Im_intro with (remove eq_dec a l); auto. constructor. repeat split; auto. pose proof (Add_intro2 _ E a) as h6. rewrite h3 in h6. rewrite list_to_set_in_iff; auto. destruct h3 as [S l h3 h4]. destruct h3 as [S h3]. subst. destruct h4 as [h4]. destruct h4 as [h4 [h5 h6]]. subst. destruct h3 as [h3]. destruct h3 as [h3 h6]. subst. constructor. split; auto. rewrite <- subtract_remove_compat. rewrite add_sub_compat_in; auto. rewrite <- list_to_set_in_iff; auto. Qed. Definition interp {T:Type} (l:list T) (x:T) (n:nat) := (firstn n l) ++ (x::nil) ++ (skipn n l). Lemma inh_list_reps : forall {T:Type}, forall (E:Ensemble T), Finite E -> Inhabited (list_reps E). intros T E h1. pose proof (finite_set_list_no_dup E h1) as h2. destruct h2 as [l h2]. apply Inhabited_intro with l; constructor; auto. Qed. Lemma inh_not_nil_in_list_reps : forall {T:Type} (E:Ensemble T) (l:list T), Inhabited E -> Inn (list_reps E) l -> l <> nil. intros T E l h1 h2 h3. subst. destruct h2 as [h2]. destruct h2 as [h2 h3]. subst. simpl in h1. destruct h1 as [x h1]. contradiction. Qed. Lemma in_list_reps : forall {T:Type} (E:Ensemble T) (l:list T), Inn (list_reps E) l -> forall x:T, In x l -> Inn E x. intros T E l h1 x h2. destruct h1 as [h1]. destruct h1 as [h1 h3]. subst. rewrite <- list_to_set_in_iff. assumption. Qed. (*See way below for finite_list_reps*) Lemma list_reps_sing : forall {T:Type} (x:T), list_reps (Singleton x) = Singleton (x::nil). intros T a. apply Extensionality_Ensembles; red; split; red; intros l h1. destruct h1 as [h1 h2]. destruct h1 as [h1 h2]. symmetry in h1. apply list_to_set_eq_sing in h1. subst. constructor. assumption. destruct h1. constructor. simpl. split. rewrite add_empty_sing; auto. constructor; simpl. intro h1. contradiction. constructor. Qed. Lemma fun_list_to_set_eq : forall {T:Type} {U:Ensemble T->T->Set} (A:Ensemble T), Finite A -> (forall (x:T) (u u' : U A x), u = u') -> (forall (x:T) (l:list T), list_to_set l = A -> U A x) -> (forall x:T, U A x). intros T U A h1 h0. intros h2 x. pose proof (finite_inh_or_empty_dec A h1) as h5. specialize (h2 x). pose proof (inh_list_reps _ h1) as h3. eapply fun_in_set_inh_const. refine (fun (l:list T) (pf:Inn (list_reps A) l) => (h2 _ (in_list_reps_eq _ _ pf))). apply (h0 x). assumption. Qed. Lemma fun_dep_list_to_set_eq : forall {T:Type} {P:T->Prop} {U:forall (C:Ensemble T) (x:T), P x -> Set} (A:Ensemble T), Finite A -> (forall (x:T) (pfx:P x) (u u' : U A x pfx), u = u') -> (forall (x:T) (pfx:P x) (l:list T), list_to_set l = A -> U A x pfx) -> (forall (x:T) (pfx:P x), U A x pfx). intros T P U A h1 h0. intros h2 x hx. pose proof (finite_inh_or_empty_dec A h1) as h5. specialize (h2 x hx). pose proof (inh_list_reps _ h1) as h3. eapply fun_in_set_inh_const. refine (fun (l:list T) (pf:Inn (list_reps A) l) => (h2 _ (in_list_reps_eq _ _ pf))). apply (h0 x). assumption. Qed. Lemma fun_finite_list_to_set_eq : forall {T:Type}, let P := (@Finite T) in forall {U:forall (A B:Ensemble T), P A -> P B -> Set} (A:Ensemble T) (pfa:P A), (forall (B:Ensemble T) (pfb:P B) (u u' : U A B pfa pfb), u = u') -> (forall (B:Ensemble T) (pfb:P B) (l:list T), list_to_set l = A -> U A B pfa pfb) -> (forall (B:Ensemble T) (pfb:P B), U A B pfa pfb). intros T. simpl. intros U A h1 h0 h2 B h4. specialize (h2 B h4). pose proof (inh_list_reps _ h1) as h6. cut (forall (l:list T) (pf:Inn (list_reps A) l), U A B h1 h4); auto. intro h7. eapply fun_in_set_inh_const. refine (fun (l:list T) (pf:Inn (list_reps A) l) => (h2 _ (in_list_reps_eq _ _ pf))). apply h0. assumption. intros l h7. eapply h2. destruct h7 as [h7]. destruct h7 as [h7 h8]. symmetry. apply h7. Qed. Lemma list_prod_nil : forall {T U:Type} (l:list T), list_prod l (@nil U) = nil. intros ? ? l. induction l; auto. Qed. Lemma list_prod_eq_nil : forall {T U:Type} (l:list T) (l':list U), list_prod l l' = nil -> l = nil \/ l' = nil. intros T U l. destruct l as [|a l]; simpl. intros; left; auto. intros l' h1. right. apply app_eq_nil in h1. destruct h1 as [h1]. apply (map_eq_nil _ _ h1). Qed. Lemma list_prod_cart_prod_compat : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (la:list T) (lb:list U), list_to_set la = A -> list_to_set lb = B -> list_to_set (list_prod la lb) = cart_prod A B. intros T U A B la. induction la as [|a la h1]. simpl. intros lb h1 h2. rewrite <- h1. rewrite cart_prod_empty. reflexivity. simpl. intros lb h2 h3. apply Extensionality_Ensembles. red. split. (* <= *) red. intros pr h4. rewrite <- list_to_set_in_iff in h4. pose proof (in_app_or _ _ _ h4) as h5. destruct h5 as [h5l | h5r]. rewrite in_map_iff in h5l. destruct h5l as [b h6]. destruct h6 as [h6l h6r]. subst. constructor. split. simpl. apply Add_intro2. simpl. rewrite <- list_to_set_in_iff. assumption. rewrite (surjective_pairing pr) in h5r. rewrite in_prod_iff in h5r. constructor. do 2 rewrite list_to_set_in_iff in h5r. destruct h5r as [h5a h5b]. pose proof (Add_intro1 _ _ a _ h5a) as h6. rewrite h2 in h6. rewrite h3 in h5b. split; assumption. (* >= *) red. intros pr h4. rewrite <- list_to_set_in_iff. apply in_or_app. rewrite <- h2 in h4. rewrite <- h3 in h4. destruct h4 as [h4]. destruct h4 as [h4l h4r]. inversion h4l as [? h5|? h6]. subst. right. rewrite (surjective_pairing pr). rewrite <- list_to_set_in_iff in h4r. rewrite <- list_to_set_in_iff in h5. apply in_prod; assumption. subst. inversion h6. rewrite (surjective_pairing pr). left. simpl. rewrite in_map_iff. exists (snd pr). subst. split; auto. rewrite list_to_set_in_iff. assumption. Qed. Lemma nth_0 : forall {T:Type} (l:list T) (def:T), nth 0 l def = hd def l. intros ? l. induction l; simpl; auto. Qed. Lemma nth_indep' : forall {T:Type} (l:list T) (n:nat) (pf:n < length l) (def:T), nth n l def = nth n l (hd' l (lt_length _ _ pf)). intros; apply nth_indep; auto. Qed. Definition nth_lt {T:Type} (l:list T) (n:nat) : n < length l -> T := match l with | nil => (fun pf => (False_rect _ (lt_n_0 _ pf))) | a::l' => (fun _ => nth n (a::l') a) end. Definition indl {T} (l:list T) i : option T := match lt_le_dec i (length l) with | left pflt => Some (nth_lt l i pflt) | right _ => None end. Notation "l '\_' i" := (indl l i) (at level 30, right associativity). Lemma in_nth_lt : forall {T:Type} (l:list T) (n:nat) (pf:n < length l), In (nth_lt l n pf) l. intros T l. destruct l as [|a l]; simpl. intros; omega. intros n h1. destruct n as [|n]. left; auto. apply lt_S_n in h1. right. apply nth_In; auto. Qed. Lemma in_infl : forall {T:Type} (l:list T) (n:nat), n < length l -> match (l \_ n) with | Some x => In x l | None => True end. intros T l n h1. unfold indl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. apply in_nth_lt. auto. Qed. Lemma nth_lt_compat : forall {T:Type} (l:list T) (n:nat) (pflt:n < length l), let x' := nth_lt l n pflt in nth_lt l n pflt = nth n l x'. intros T l n h1; unfold nth_lt. simpl. destruct l as [|a l]. simpl in h1. pose proof (lt_n_0 _ h1). contradiction. destruct n as [|n]; simpl; auto. apply nth_indep. apply lt_S_n in h1; auto. Qed. Lemma nth_lt_cons : forall {T:Type} (l:list T) (a:T) (n:nat) (pflt:S n < S (length l)), let pflt' := lt_S_n _ _ pflt in nth_lt (a::l) (S n) pflt = nth_lt l n pflt'. intros T l a n h1; simpl; unfold nth_lt; simpl. destruct l as [|l c]; simpl. destruct n; destruct lt_n_0. destruct n; auto. apply nth_indep. simpl in h1. do 2 apply lt_S_n in h1; auto. Qed. Lemma nth_inj : forall {T:Type} (l:list T), NoDup l -> forall (m n:nat) (d1 d2:T), m < length l -> n < length l -> nth m l d1 = nth n l d2 -> m = n. intros T l. induction l as [|a l h1]; simpl. intros; omega. intros h2 m n d1 d2 h3 h4 h5. pose proof (no_dup_cons_nin _ _ h2) as hnin. pose proof (no_dup_cons _ _ h2) as hdup. destruct (le_lt_eq_dec _ _ h3) as [h6 | h7]; destruct (le_lt_eq_dec _ _ h4) as [h8 | h9]. apply lt_S_n in h6. apply lt_S_n in h8. apply (h1 hdup _ _ d1 d2 h6 h8). destruct m; destruct n; auto. apply nth_indep. assumption. apply lt_S_n in h4. pose proof (nth_In _ d1 h4) as h11. subst. pose proof (nth_indep). pose proof (nth_indep _ d1 d2 h4) as hi. rewrite hi in h11. contradiction. subst. apply lt_S_n in h3. pose proof (nth_In _ d2 h3) as h12. pose proof (nth_indep _ d1 d2 h3) as hi. rewrite <- hi in h12. contradiction. apply h1 in h5. subst. apply nth_indep; auto. assumption. apply lt_S_n. assumption. apply lt_S_n. assumption. apply S_inj in h9. apply lt_S_n in h6. subst. destruct (zerop (length l)) as [h8 | h9]. omega. pose proof (S_pred _ _ h9) as h10. rewrite h10 in h5. destruct m as [|m]. assert (h11:pred (length l) < length l). omega. pose proof (nth_In _ d2 h11) as h12. rewrite <- h5 in h12. contradiction. assert (h11:pred (length l) < length l). omega. apply lt_S_n in h3. pose proof (h1 hdup _ _ d1 d2 h3 h11 h5). omega. apply lt_S_n in h8. apply S_inj in h7. subst. destruct (zerop (length l)) as [h9 | h10]. omega. rewrite (S_pred _ _ h10) in h5. destruct n as [|n]. assert (h11:pred (length l) < length l). omega. pose proof (nth_In _ d1 h11) as h12. rewrite h5 in h12. contradiction. apply lt_S_n in h4. assert (h11:pred (length l) < length l). omega. pose proof (h1 hdup _ _ d1 d2 h11 h4 h5) as h12. omega. apply S_inj in h7. apply S_inj in h9. congruence. Qed. Lemma nth_lt_inj : forall {T:Type} (l:list T) (pfnd:NoDup l) (m n:nat) (pfm: m < length l) (pfn: n < length l), nth_lt l m pfm = nth_lt l n pfn -> m = n. intros T l h1 m n h2 h3 h4. pose proof (nth_lt_compat _ _ h2) as h5. pose proof (nth_lt_compat _ _ h3) as h6. simpl in h5, h6. rewrite h6, h5 in h4. eapply nth_inj; auto. apply h1. apply h2. apply h3. apply h4. Qed. Lemma inj_seg_nth_lt : forall {T:Type} (l:list T), NoDup l -> inj_seg (nth_lt l). intros T l h1. red; intros x y h2 h3 h4. apply nth_lt_inj in h4; auto. Qed. Lemma nth_lt_0 : forall {T:Type} (l:list T) (a:T), NoDup (a::l) -> forall (n:nat) (pfl:n < length (a::l)), nth_lt (a::l) n pfl = a -> n = 0. intros T l a h1 n h2 h3. destruct n; auto. rewrite nth_lt_cons in h3. pose proof (nth_lt_compat l n (lt_S_n n (length l) h2)) as h4. simpl in h4. simpl in h2. subst. pose proof (in_nth_lt l n (lt_S_n _ _ h2)) as h3. pose proof (lt_S_n _ _ h2) as h5. pose proof (no_dup_cons_nin _ _ h1) as h8. contradiction. Qed. Lemma nth_lt0 : forall {T:Type} (l:list T) (pf:l <> nil), let pf' := not_nil_lt l pf in nth_lt l 0 pf' = hd' l pf. intros T l. destruct l as [|a l]; simpl; auto. intro h1; contradict h1; auto. Qed. Lemma nth_lt0' : forall {T:Type} (l:list T) (pf:0 < length l), let pf' := lt_length _ _ pf in nth_lt l 0 pf = hd' l pf'. intros T l. destruct l as [|a l]. simpl; intros h1; contradict (lt_irrefl _ h1). intros; simpl; auto. Qed. Lemma nth_lt_functional1 : forall {T:Type} (l l':list T) (n:nat) (pfeq: l = l') (pf: n < length l) (pf' : n < length l'), nth_lt l n pf = nth_lt l' n pf'. intros T l l' n h0 h1 h2. subst. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. Lemma nth_lt_functional2 : forall {T:Type} (l:list T) (n n':nat) (pfeq: n = n') (pf: n < length l) (pf' : n' < length l), nth_lt l n pf = nth_lt l n' pf'. intros T l n n' h0 h1 h2. subst. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. Lemma nth_lt_functional3 : forall {T:Type} (l:list T) (n:nat) (pf1 pf2: n < length l), nth_lt l n pf1 = nth_lt l n pf2. intros T l n h1 h2. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. Lemma nth_lt_0' : forall {T:Type} (l:list T) (a:T) (pf:0 < length (a::l)), nth_lt (a::l) 0 pf = a. intros; simpl; auto. Qed. Lemma nth_lt_0_hd' : forall {T:Type} (l:list T) (pf:0 < length l), nth_lt l 0 pf = hd' l (lt_length _ _ pf). intros T l h1. destruct l as [|a l]; simpl. f_equal. apply proof_irrelevance. reflexivity. Qed. Lemma nth_lt_cons' : forall {T:Type} (a:T) (l:list T) (n:nat) (pf:n < S (length l)) (pflt : 0 < n), nth_lt (a::l) n pf = nth_lt l (pred n) (mono_pred_lt _ _ pf pflt). intros T a l n h1 h2. simpl. destruct n as [|n]; simpl. pose proof (lt_irrefl _ h2). contradiction. symmetry. gen0. intro h3. simpl in h3. rewrite nth_lt_compat. apply nth_indep; auto. Qed. Lemma nth_lt_S : forall {T:Type} (l:list T) (n:nat) (pf:S n < length l), nth_lt l (S n) pf = nth_lt (tl l) n (lt_length_tl _ _ pf). intros T l n h1. pose proof (hd_compat' l (lt_length _ _ h1)) as h3. simpl in h3. generalize h1. rewrite h3. intro. rewrite (nth_lt_cons' _ _ _ _ (lt_0_Sn _)) at 1. simpl. f_equal. apply proof_irrelevance. Qed. Lemma nth_lt_cons_in : forall {T:Type} (l:list T) (a:T) (n:nat) (pf0: 0 < n) (pfl:n < S (length l)), nth_lt (a::l) n pfl = a -> In a l. intros T l a n h1 h2 h3. pose proof (nth_lt_cons' a l n h2 h1) as h4. rewrite h3 in h4. pose proof (nth_lt_compat l (pred n) (mono_pred_lt n (S (length l)) h2 h1)) as h5. simpl in h5. destruct h5 as [h5 h6]. subst. apply in_nth_lt. Qed. Lemma nth_lt_1 : forall {T:Type} (l:list T) (pf: 1 < length l) (pftl:tl l <> nil), nth_lt l 1 pf = hd' (tl l) pftl. intros T l. induction l as [|a l h1]; simpl. intros; omega. intros h2 h3. rewrite nth_0. erewrite hd_hd_compat'. reflexivity. Qed. Lemma nth_lt_sing : forall {T:Type} (a:T) (m:nat) (pf:m < length (a::nil)), nth_lt (a::nil) m pf = a. intros T a m h1. simpl in h1. assert (h2:m = 0). omega. subst. rewrite nth_lt_0_hd'. rewrite hd_cons'; auto. Qed. Lemma nth_lt_dup : forall {T:Type} (l:list T) (m n:(length l) @), m <> n -> nth_lt _ _ (proj2_sig m) = nth_lt _ _ (proj2_sig n) -> ~NoDup l. intros T l m n h1 h2 h3. pose proof (nth_lt_compat _ _ (proj2_sig m)) as h4. simpl in h4. pose proof (nth_lt_compat _ _ (proj2_sig n)) as h5. simpl in h5. rewrite h2 in h4. rewrite h4 in h5 at 1. pose proof (nth_inj _ h3 _ _ _ _ (proj2_sig m) (proj2_sig n) h5) as h6. apply proj1_sig_injective in h6. contradict h1; auto. Qed. Lemma nth_lt_dup' : forall {T:Type} (l:list T) (m n:nat) (pfm : m < length l) (pfn : n < length l), m <> n -> nth_lt _ _ pfm = nth_lt _ _ pfn -> ~NoDup l. intros T l m n pfm pfn h1 h2 h3. pose proof (nth_lt_compat _ _ pfm) as h4. simpl in h4. pose proof (nth_lt_compat _ _ pfn) as h5. simpl in h5. rewrite h2 in h4. rewrite h4 in h5 at 1. pose proof (nth_inj _ h3 _ _ _ _ pfm pfn h5); subst. contradict h1; auto. Qed. Lemma nth_lt_S_neq : forall {T:Type} (l:list T) (n:nat) (pfnil: l <> nil) (pf: n < length l), nth_lt l n pf <> hd' l pfnil -> In (nth_lt l n pf) (tl l). intros T l. induction l as [|a l h1]; simpl; intros n h2 h3 h6; try omega. revert h6. revert h2 h1 h3. revert a l. induction n as [|n h5]. intros a l h2 h0 h3 hn. contradict hn; auto. intros; apply nth_In. apply lt_S_n in h3; auto. Qed. Lemma nth_lt_map : forall {T U:Type} (l:list T) (f:T->U) (n:nat) (pf:n < length l) (pf': n < length (map f l)), nth_lt (map f l) n pf' = f (nth_lt l n pf). intros T U l. induction l as [|a l ih]; simpl. intros ? ? h1. pose proof (lt_n_0 _ h1). contradiction. intros f n h1 h2. destruct n as [|n]; simpl; auto. apply lt_S_n in h1. apply lt_S_n in h2. specialize (ih f n h1 h2). rewrite nth_lt_compat in ih. erewrite nth_indep. rewrite ih at 1. f_equal. rewrite nth_lt_compat. apply nth_indep; auto. assumption. Qed. Lemma nth_lt_map' : forall {T U:Type} (l:list T) (f:T->U) (n:nat) (pf': n < length (map f l)), let pf := lt_eq_trans _ _ _ pf' (map_length f l) in nth_lt (map f l) n pf' = f (nth_lt l n pf). intros T U l f n h1 h2. erewrite nth_lt_map. reflexivity. Qed. Lemma map_eq_iff : forall {T U:Type} (l:list T) (f:T->U) (l':list U), map f l = l' <-> (exists pfeq:length l = length l', (forall (n:nat) (pflt:n < length l), f (nth_lt l n pflt) = nth_lt l' n (lt_eq_trans _ _ _ pflt pfeq))). intros T U l. induction l as [|a l ih]; simpl. intros. split; intro h1. subst; simpl. exists eq_refl. intros; omega. destruct h1 as [h1 h2]. pose proof (length_eq0 _ (eq_sym h1)); auto. intros f l'. split; intro h1. pose proof (f_equal (@length _) h1) as h2. simpl in h2. rewrite map_length in h2. exists h2. intros n h3. destruct n as [|n]; simpl. rewrite nth_lt_0_hd'. erewrite hd_compat' in h1. inversion h1 as [h4]. rewrite h4 at 1. reflexivity. pose proof (lt_S_n _ _ h3) as h4. symmetry. gen0. rewrite <- h1. intro h5; simpl. rewrite map_nth; auto. destruct h1 as [h1 h2]. destruct l' as [|a' l']. simpl in h1; discriminate. simpl in h1. pose proof (S_inj _ _ h1) as h3. pose proof (h2 0 (lt_0_Sn _)) as h4. simpl in h4; subst. f_equal. specialize (ih f l'). rewrite ih. exists h3. intros n h4. specialize (h2 (S n)). hfirst h2. apply lt_n_S; auto. specialize (h2 hf). simpl in h2. rewrite nth_lt_compat. erewrite nth_indep. rewrite h2 at 1. rewrite nth_lt_compat. apply nth_indep. omega. assumption. Qed. Lemma maps_eq_iff : forall {T U V:Type} (l:list T) (l':list U) (f:T->V) (g:U->V), map f l = map g l' <-> exists (pfe:length l = length l'), forall i (pflt:i < length l), let pflt' := lt_eq_trans _ _ _ pflt pfe in f (nth_lt l i pflt) = g (nth_lt l' i pflt'). intros T U V l l' f g. rewrite map_eq_iff. split; intro h1; destruct h1 as [h1 h2]. pose proof h1 as h1'. rewrite map_length in h1'. exists h1'. intros i h3 h4. rewrite h2. rewrite (nth_lt_map _ _ _ h4); auto. ex_goal. rewrite map_length; auto. exists hex. intros n h3. rewrite (nth_lt_map _ _ _ (lt_eq_trans _ _ _ h3 h1)). rewrite h2; auto. Qed. Lemma nth_lt_app1 : forall {T:Type} (l l':list T) (n:nat) (pfn:n < length l) (pfn' : n < (length (l++l'))), nth_lt (l++l') n pfn' = nth_lt l n pfn. intros T l. induction l as [|a l ih]; simpl. intros l' n h1. pose proof (lt_n_0 _ h1). contradiction. intros l' n h1 h2. destruct n as [|n]; simpl; auto. apply lt_S_n in h1. apply lt_S_n in h2. specialize (ih _ _ h1 h2). rewrite nth_lt_compat in ih. erewrite nth_indep. rewrite ih at 1. rewrite nth_lt_compat. apply nth_indep; auto. assumption. Qed. Lemma nth_lt_app2 : forall {T:Type} (l l':list T) (n:nat) (pfn:n >= length l) (pfn' : n < (length (l++l'))), nth_lt (l++l') n pfn' = nth_lt l' (n - length l) (lt_length_app_ge_impl _ _ _ pfn pfn'). intros T l. induction l as [|a l ih]; simpl. intros l' n h1 h2. symmetry. gen0. simpl. pose proof (minus_n_O n) as h3. rewrite <- h3. intro h4. f_equal. apply proof_irrelevance. intros l' n h1 h2. destruct n as [|n]; simpl. omega. pose proof (lt_S_n _ _ h2) as h2'. pose proof (le_S_n _ _ h1) as h1'. specialize (ih _ _ h1' h2'). rewrite nth_lt_compat in ih. erewrite nth_indep. rewrite ih at 1. f_equal. apply proof_irrelevance. apply lt_S_n in h2. assumption. Qed. Lemma nth_lt_pred_length : forall {T:Type} (l:list T) (pf:pred (length l) < length l), nth_lt l (pred (length l)) pf = lst l (lt_length _ _ pf). intros T l h1. pose proof (lst_compat l (lt_length _ _ h1)) as h2. simpl in h2. pose proof (nth_lt_functional1 _ _ (pred (length l)) h2 h1) as h3. hfirst h3. rewrite <- h2; auto. specialize (h3 hf). rewrite h3. erewrite nth_lt_app2. simpl. rewrite length_removelast. rewrite minus_same; auto. Grab Existential Variables. rewrite length_removelast; auto. Qed. Lemma lst_eq : forall {T:Type} (l:list T) (pfn:l <> nil), lst l pfn = nth_lt l (pred (length l)) (not_nil_pred_lt _ pfn). intros; rewrite nth_lt_pred_length; f_equal; apply proof_irrelevance. Qed. Lemma all_sing_nth_lt : forall {T:Type} (l:list T) (a:T), all_sing l a -> forall (m:nat) (pfm:m < length l), nth_lt l m pfm = a. intros T l. destruct l as [|a l]; simpl. intros; omega. intros a' h1 m h2. rewrite all_sing_cons_iff in h1. destruct h1 as [? h1]; subst. destruct m as [|m]; auto. apply all_sing_nth; auto. Qed. Lemma nth_lt_seg_list : forall (m n:nat) (pf:m < length (seg_list n)), nth_lt (seg_list n) m pf = m. intro m. induction m as [|m ih]; simpl; auto; intros n h1. revert h1. induction n as [|n]; simpl; auto. intros; omega. rewrite nth_lt_S. unfold seg_list. gen0. rewrite tl_seq. intro h2. pose proof h2 as h2'. rewrite seq_length in h2'. specialize (ih (pred n)). hfirst ih. rewrite length_seg_list; auto. specialize (ih hf). gen0. rewrite seq_eq_seq_0. intros h3; simpl. erewrite nth_lt_map. rewrite ih at 1. reflexivity. Qed. Lemma seg_list_eq_n_seg_list : forall m n, seg_list m = n :: seg_list n -> m = 1. intros m. destruct m as [|m]; simpl. intros n h1. pose proof (in_eq n (seg_list n)) as h2. rewrite <- h1, seg_list0 in h2. simpl in h2; contradiction. intros n h1. f_equal. pose proof (lt_0_Sn m) as h2. pose proof (nth_lt_functional1 _ _ 0 h1) as h3. hfirst h3. rewrite length_seg_list; auto with arith. specialize (h3 hf (lt_0_Sn _)). simpl in h3; subst. rewrite seg_list0 in h1. pose proof (lt_n_Sn m) as h3. apply in_seg in h3. rewrite <- list_to_set_seg_list_eq in h3. rewrite h1 in h3. simpl in h3. rewrite add_empty_sing in h3. rewrite in_sing_iff in h3; auto. Qed. Lemma nth_lt_seq : forall (m n r:nat) (pf:r < length (seq m n)), nth_lt (seq m n) r pf = m + r. intros m n. revert m. induction n as [|n ih]. simpl; intros; omega. intros m r h1. pose proof h1 as h1'. rewrite seq_length in h1'. gen0. rewrite seq_S; auto with arith. intros h2. destruct r as [|r]. erewrite nth_lt_functional3, nth_lt0. simpl; auto. simpl. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite ih at 1; omega. rewrite seq_length. omega. Grab Existential Variables. rewrite seq_length; omega. intro; discriminate. Qed. Lemma nth_lt_map_seg : forall {U:Type} {n:nat} (f:seg_fun n U) (m:nat) (pf:m < n), let pf' := lt_eq_trans _ _ _ pf (eq_sym (length_map_seg f)) in nth_lt (map_seg f) m pf' = f m pf. intros U n. induction n as [|n ih]; simpl. intros; omega. intros f m h1. simpl in ih. red in h1. pose proof h1 as h1'. apply le_S_n in h1'. destruct (le_lt_eq_dec _ _ h1') as [h2 | h2]. specialize (ih (seg_fun_from_S f) m h2). erewrite nth_lt_app1. unfold seg_fun_from_S in ih. rewrite ih at 1. f_equal. apply proof_irrelevance. subst. erewrite nth_lt_app2. simpl. rewrite length_map_seg. rewrite minus_same. f_equal. apply proof_irrelevance. Grab Existential Variables. rewrite length_map_seg; auto with arith. Qed. Lemma nth_lt_map_seg' : forall {U:Type} {n:nat} (f:seg_fun n U) (m:nat) (pf:m < length (map_seg f)), let pf' := lt_eq_trans _ _ _ pf (length_map_seg f) in nth_lt (map_seg f) m pf = f m pf'. intros U n f m h1. simpl. erewrite nth_lt_functional3. erewrite nth_lt_map_seg. reflexivity. Qed. Lemma lst_map_seg : forall {U:Type} {n:nat} {f:seg_fun n U} (pfn:map_seg f <> nil), lst (map_seg f) pfn = f (pred n) (map_seg_lt _ pfn). intros U n f h1. rewrite lst_eq. rewrite nth_lt_map_seg'. apply f_equal_dep. rewrite length_map_seg; auto. Qed. Lemma last_map_seg : forall {U:Type} {n:nat} (def:U) (f:seg_fun n U) (pf:0 < n), last (map_seg f) def = f (pred n) (lt_pred_n _ _ pf). intros; erewrite last_compat, lst_eq, nth_lt_map_seg'; apply f_equal_dep. rewrite length_map_seg; auto. Grab Existential Variables. intro h1. apply map_seg_eq_nil in h1. subst. apply lt_irrefl in pf; contradiction. Qed. Lemma map_to_map_seg : forall {T U:Type} (f:T->U) (l:list T), map f l = map_seg (fun i (pf:i < length l) => f (nth_lt _ _ pf)). intros T U f l. rewrite map_eq_iff. ex_goal. rewrite length_map_seg; auto. exists hex. intros n h1. symmetry. erewrite nth_lt_functional3. rewrite nth_lt_map_seg. reflexivity. Qed. Lemma nth_lt_map_in_dep : forall {T U:Type} {l:list T} (f:fun_in_dep l U) (n:nat) (pf:n < length l) (pf':n < length (map_in_dep f)), nth_lt (map_in_dep f) n pf' = f (nth_lt l n pf) (in_nth_lt _ _ _). intros T U l. induction l as [|a l h1]; simpl. intros; omega. intros f n. destruct n as [|n]; simpl. intros h2 h3. f_equal. apply proof_irrelevance. intros h2 h3. pose proof h2 as h2'. pose proof (lt_S_n _ _ h2) as h4. pose proof (lt_S_n _ _ h3) as h5. specialize (h1 (fun_from_cons f) _ h4 h5). rewrite nth_lt_compat in h1 at 1. erewrite nth_indep. rewrite h1 at 1. rewrite fun_from_cons_compat. apply f_equal_prop_dep. rewrite nth_lt_compat. apply nth_indep; auto. assumption. Qed. Lemma nth_lt_map_in_dep' : forall {T U:Type} {l:list T} (f:fun_in_dep l U) (n:nat)(pf':n < length (map_in_dep f)), let pf := lt_eq_trans _ _ _ pf' (length_map_in_dep f) in nth_lt (map_in_dep f) n pf' = f (nth_lt l n pf) (in_nth_lt _ _ _). intros; erewrite nth_lt_map_in_dep. reflexivity. Qed. Lemma map_seg_list : forall {U:Type} (f:nat->U) (n:nat), map f (seg_list n) = map_seg (fun_to_seg f n). intros U f n. rewrite map_eq_iff. ex_goal. rewrite length_seg_list. rewrite length_map_seg; auto. exists hex. intros m h1. symmetry. erewrite nth_lt_functional3. rewrite nth_lt_map_seg at 1. unfold fun_to_seg. f_equal. rewrite nth_lt_seg_list; auto. Grab Existential Variables. rewrite length_seg_list in h1; auto. Qed. Lemma map_in_dep_eq_iff : forall {T U:Type} {l:list T} (f:fun_in_dep l U) (l':list U), map_in_dep f = l' <-> (exists pfeq:length l = length l', forall (n:nat) (pflt:n < length l), f (nth_lt l n pflt) (in_nth_lt _ _ _) = nth_lt l' n (lt_eq_trans _ _ _ pflt pfeq)). intros T U l. induction l as [|a l ih]; simpl. intros f l'. split; intro h1. subst. simpl. exists eq_refl. intros; omega. destruct h1 as [h1]. pose proof (length_eq0 _ (eq_sym h1)); auto. intros f l'. split; intro h1. ex_goal. pose proof (f_equal (@length _) h1) as h2. simpl in h2. rewrite length_map_in_dep in h2; auto. exists hex. intros n h2. destruct n as [|n]; simpl. rewrite nth_lt_0_hd'. symmetry. gen0. rewrite <- h1. intro h3. simpl. f_equal. apply proof_irrelevance. pose proof h2 as h2'. apply lt_S_n in h2'. destruct l' as [|a' l']. simpl in hex; discriminate. simpl in hex. pose proof hex as hex'. apply S_inj in hex'. simpl. inversion h1 as [h3]; subst. erewrite nth_indep. erewrite <- nth_lt_compat at 1. rewrite nth_lt_map_in_dep at 1. unfold fun_from_cons. apply f_equal_prop_dep. rewrite nth_lt_compat. apply nth_indep; auto. rewrite length_map_in_dep; auto. destruct h1 as [h1 h2]. pose proof (h2 0 (lt_0_Sn _)) as h2'. simpl in h2'. rewrite nth_lt_0_hd' in h2'. erewrite hd_compat'. rewrite <- h2' at 1. f_equal. f_equal. apply proof_irrelevance. specialize (ih (fun_from_cons f) (tl l')). rewrite ih. destruct l' as [|a' l']; simpl. simpl in h1; discriminate. simpl in h1. pose proof h1 as h1'. apply S_inj in h1'. exists h1'. intros n h3. specialize (h2 (S n) (lt_n_S _ _ h3)). simpl in h2. unfold fun_from_cons. rewrite nth_lt_compat. erewrite nth_indep. rewrite <- h2 at 1. apply f_equal_prop_dep. rewrite nth_lt_compat. apply nth_indep; auto. eapply lt_eq_trans. apply h3. assumption. Grab Existential Variables. assumption. rewrite length_map_in_dep; auto. Qed. Lemma map_in_dep_to_map_seg : forall {T U:Type} {l:list T} (f:fun_in_dep l U), map_in_dep f = map_seg (fun i (pf:i < length l) => f _ (in_nth_lt _ _ pf)). intros T U l f. rewrite map_in_dep_eq_iff. ex_goal. rewrite length_map_seg; auto. exists hex. intros n h1. symmetry. erewrite nth_lt_functional3. rewrite nth_lt_map_seg. reflexivity. Qed. Lemma map_in_deps_eq_iff : forall {T U V:Type} (l:list T) (l':list U) (f:fun_in_dep l V) (g:fun_in_dep l' V), map_in_dep f = map_in_dep g <-> exists (pfe:length l = length l'), forall i (pflt:i < length l), let pflt' := lt_eq_trans _ _ _ pflt pfe in f (nth_lt l i pflt) (in_nth_lt _ _ _) = g (nth_lt l' i pflt') (in_nth_lt _ _ _). intros T U V l l' f g. rewrite map_in_dep_eq_iff. split; intro h1; destruct h1 as [h1 h2]. ex_goal. pose proof h1 as h1'. rewrite length_map_in_dep in h1'; auto. exists hex. intros i h3 h4. rewrite h2. rewrite nth_lt_map_in_dep'. apply f_equal_dep. apply nth_lt_functional3. ex_goal. rewrite length_map_in_dep; auto. exists hex. intros n h3. rewrite nth_lt_map_in_dep'. rewrite h2. apply f_equal_dep. apply nth_lt_functional3. Qed. Lemma nth_lt_map_sig : forall {T} (l:list T) i (pf:i < length (map_sig l)), nth_lt (map_sig l) i pf = exist (fun c => In c l) (nth_lt l i (lt_eq_trans _ _ _ pf (length_map_sig _))) (in_nth_lt _ _ _). unfold map_sig; intros; rewrite nth_lt_map_in_dep'; apply proj1_sig_injective; simpl; f_equal; apply proof_irrelevance. Qed. Lemma map_sig_compat : forall {T:Type} (l:list T), map (@proj1_sig _ _) (map_sig l) = l. intros T l; rewrite map_eq_iff; exists (length_map_sig _). intros; rewrite nth_lt_map_sig; simpl; auto. Qed. Lemma map_map_sig_functional : forall {T U:Type} (l l':list T) (f:{t:T | In t l}->U) (pf:l = l'), map f (map_sig l) = map (transfer_dep (U:= (fun xl:list T => {t:T | In t xl}->U)) pf f) (map_sig l'). intros; subst; rewrite transfer_dep_eq_refl; auto. Qed. Lemma list_of_lists_seqs_not_nil : forall {T:Type} (ll:faml T), p_mem_l (fun l => l <> nil) ll -> list_of_lists_seqs ll <> nil. intros T ll. induction ll as [|la ll ih]; simpl. intros; intro; discriminate. intros h1. hfirst ih. eapply p_mem_l_cons. apply h1. specialize (ih hf). intro h2. apply map_eq_nil_iff in h2. apply list_prod_eq_nil in h2. destruct h2 as [|h2]; subst. red in h1. specialize (h1 nil (in_eq _ _)). contradict h1; auto. contradiction. Qed. (*Elementary binary vector of width [n] and support [m].*) Definition ebv (m n:nat) : bv := map_seg (fun i (pf:i < n) => if (nat_eq_dec i m) then true else false). (*Complement of a binary vector.*) Definition nbv (v:bv) : bv := map negb v. Definition tbv {v v':bv} (pfeq:length v = length v') : bv := map_seg (fun i (pfi:i < length v) => let pfi' := lt_eq_trans _ _ _ pfi pfeq in (andb (nth_lt v i pfi) (nth_lt v' i pfi'))). Definition pbv {v v':bv} (pfeq:length v = length v') : bv := map_seg (fun i (pfi:i < length v) => let pfi' := lt_eq_trans _ _ _ pfi pfeq in (orb (nth_lt v i pfi) (nth_lt v' i pfi'))). Definition xbv {v v':bv} (pfeq:length v = length v') : bv := map_seg (fun i (pfi:i < length v) => let pfi' := lt_eq_trans _ _ _ pfi pfeq in (xorb (nth_lt v i pfi) (nth_lt v' i pfi'))). (*Complement of an elementary binary vector.*) Definition cebv (m n:nat) : bv := nbv (ebv m n). Definition cons_mask (n:nat) : bv := cebv 0 n. Definition inds_to_bv {T:Type} (l:list T) (ln:list nat): bv := map_seg (fun i (pfi:i < length l) => if (in_dec nat_eq_dec i ln) then true else false). Lemma in_bv0s : forall (v:bv) (m:nat) (pfm:In m (bv0s v)), let pfv := in_bv0s_lt _ _ pfm in nth_lt v _ pfv = false. intro v. induction v as [|b v ih]; simpl. intros; contradiction. intros m h1. unfold bv0s in h1; simpl in h1. destruct b. rewrite bv0s_aux_S in h1. rewrite in_lS_iff in h1. destruct h1 as [? [h1 ?]]; subst. specialize (ih _ h1). simpl in ih. rewrite <- ih. rewrite nth_lt_compat. apply nth_indep. pose proof h1 as h1'. apply in_bv0s_lt in h1'; auto. destruct m as [|m]; auto. destruct h1 as [|h1]. discriminate. rewrite bv0s_aux_S in h1. rewrite in_lS_iff in h1. destruct h1 as [? [h1 h2]]. apply S_inj in h2; subst. specialize (ih _ h1). simpl in ih. rewrite <- ih. rewrite nth_lt_compat. apply nth_indep. pose proof h1 as h1'. apply in_bv0s_lt in h1'; auto. Qed. Lemma in_bv1s : forall (v:bv) (m:nat) (pfm:In m (bv1s v)), let pfv := in_bv1s_lt _ _ pfm in nth_lt v _ pfv = true. intro v. induction v as [|b v ih]; simpl. intros; contradiction. intros m h1. unfold bv1s in h1; simpl in h1. destruct b. destruct m as [|m]; auto. destruct h1 as [|h1]. discriminate. rewrite bv1s_aux_S in h1. rewrite in_lS_iff in h1. destruct h1 as [? [h1 h2]]. apply S_inj in h2; subst. specialize (ih _ h1). simpl in ih. rewrite <- ih. rewrite nth_lt_compat. apply nth_indep. pose proof h1 as h1'. apply in_bv1s_lt in h1'; auto. rewrite bv1s_aux_S in h1. rewrite in_lS_iff in h1. destruct h1 as [? [h1 ?]]; subst. specialize (ih _ h1). simpl in ih. rewrite <- ih. rewrite nth_lt_compat. apply nth_indep. pose proof h1 as h1'. apply in_bv1s_lt in h1'; auto. Qed. (*Think of [coarse_list] as kind of a "coarsening" so that for instance [coarse_list [1, 2, 5, 0, 8] [3, 1, 4, 2, 5, 6, 0, 8, 9 10]], i.e. the first argument is contained in the second argument, albeit interspersed by other elements.*) Inductive coarse_list {T:Type} : list T -> list T -> Prop := | coarse_list_nil : forall l:list T, coarse_list nil l | coarse_list_cons1 : forall (l:list T) (l':list T) (a:T), coarse_list l l' -> coarse_list l (a::l') | coarse_list_cons2 : forall (l:list T) (l':list T) (a:T), coarse_list l l' -> coarse_list (a::l) (a::l'). Lemma coarse_list_refl : forall {T:Type} (l:list T), coarse_list l l. intros T l. induction l as [|a l ih]; simpl. constructor. apply coarse_list_cons2; auto. Qed. Lemma length_coarse_list : forall {T:Type} (l l':list T), coarse_list l l' -> length l <= length l'. intros ? ? ? h1; induction h1; simpl; auto with arith. Qed. Lemma coarse_list_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), coarse_list (remove pfdt a l) l. intros T h1 l. induction l as [|a l ih]; simpl. intros; constructor. intro a'. destruct (h1 a' a) as [|h2]; subst. apply coarse_list_cons1. apply ih. apply coarse_list_cons2. apply ih. Qed. Lemma le_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), length (remove pfdt a l) <= length l. intros; apply length_coarse_list; apply coarse_list_remove. Qed. Lemma lt_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), In x l -> length (remove pfdt x l) < length l. intros T h1 l. induction l as [|a l ih]. intros; contradiction. intros x h2. simpl. simpl in h2. destruct (h1 x a) as [|h3]; subst. pose proof (le_remove h1 l a) as h4. omega. destruct h2 as [|h2]; subst. contradict h3; auto. simpl. apply lt_n_S. apply ih; auto. Qed. Lemma coarse_list_nil2_iff : forall {T:Type} (l:list T), coarse_list l nil <-> l = nil. intros ? l; destruct l; split; intro h2; try inversion h2; subst; auto; try constructor. Qed. Lemma not_coarse_list_cons_nil : forall {T:Type} (l:list T) (a:T), ~coarse_list (a::l) nil. intros ? ? ? h1; inversion h1. Qed. Lemma coarse_list_sing_cons : forall {T:Type} (l:list T) (a:T), coarse_list (a::nil) (a::l). intros T l a. apply coarse_list_cons2. constructor. Qed. Lemma coarse_list_sing1_iff : forall {T:Type} (l:list T) (a:T), coarse_list (a::nil) l <-> In a l. intros T l; induction l as [|b l ih]; simpl; split; intro h2. rewrite coarse_list_nil2_iff in h2. discriminate. contradiction. inversion h2; subst. right. rewrite <- ih; auto. left; auto. destruct h2 as [|h2]; subst. apply coarse_list_sing_cons. rewrite <- ih in h2. apply coarse_list_cons1; auto. Qed. Lemma coarse_list_rev_cons1 : forall {T:Type} (l l':list T) (a:T), coarse_list (a::l) l' -> coarse_list l l'. intros ? ? ? ? h1. dependent induction h1. apply coarse_list_cons1; auto. apply coarse_list_cons1; auto. Qed. Lemma coarse_list_rev_cons_impl : forall {T:Type} (l l':list T) (a:T), coarse_list (a::l) (a::l') -> coarse_list (a::l) l' \/ coarse_list l l'. intros T l l' a h1. inversion h1; subst. left; auto. right; auto. Qed. Lemma coarse_list_tl : forall {T:Type} (l:list T), coarse_list (tl l) l. destruct l as [|a l]; simpl. constructor. apply coarse_list_cons1. apply coarse_list_refl. Qed. Lemma coarse_list_cons1_iff : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T), coarse_list (a::l) l' <-> exists pf':l' <> nil, (a = hd' l' pf' /\ coarse_list l (tl l')) \/ (a <> hd' l' pf' /\ coarse_list (a::l) (tl l')). intros T h1 l l' a; split; intro h2. dependent induction h2; subst; simpl. exists (cons_neq_nil _ _). destruct (h1 a a0) as [|h3]; subst. left. split; auto. eapply coarse_list_rev_cons1. apply h2. right. split; auto. exists (cons_neq_nil _ _). left. split; auto. destruct h2 as [h2 h3]. rewrite (hd_compat' _ h2). destruct h3 as [h3 | h3]. destruct h3 as [h3 h4]. rewrite <- h3. apply coarse_list_cons2; auto. destruct h3 as [h3 h4]. apply coarse_list_cons1; auto. Qed. Lemma not_cons_coarse_list_same : forall {T:Type} (l:list T) (a:T), ~ coarse_list (a::l) l. intros T l a h1. pose proof (length_coarse_list _ _ h1) as h2. simpl in h2. omega. Qed. Lemma coarse_list_antisym : forall {T:Type} (l l':list T), coarse_list l l' -> coarse_list l' l -> l = l'. intros T l l' h1. induction h1. intro h2. rewrite coarse_list_nil2_iff in h2; auto. intro h2. dependent induction h2. pose proof (coarse_list_rev_cons1 _ _ _ h1) as h3. pose proof (coarse_list_rev_cons1 _ _ _ h2) as h4. specialize (IHh2 h3). hfirst IHh2. intro h5. hfirst IHh1. constructor; auto. specialize (IHh1 hf). subst. apply not_cons_coarse_list_same in h5. contradiction. specialize (hf h4); subst. hfirst IHh1. constructor; auto. specialize (IHh1 hf). pose proof (f_equal (@length _) IHh1) as h5. simpl in h5. omega. f_equal. hfirst IHh1. constructor; auto. specialize (IHh1 hf). subst. apply not_cons_coarse_list_same in h2; contradiction. intro h2. f_equal. apply IHh1. inversion h2; subst. apply length_coarse_list in h1. apply length_coarse_list in H1. simpl in H1. omega. assumption. Qed. Lemma in_coarse_list : forall {T:Type} (l l':list T) (a:T), In a l -> coarse_list l l' -> In a l'. intros T l l' a h1 h2. revert h1. revert a. induction h2 as [|l0 l1 a0 h5 ih|l0 l1 a0 h5 ih]; subst. intros; contradiction. intros; right; apply ih; auto. intros a h6. destruct h6 as [|h6]; subst. left; auto. right. apply ih; auto. Qed. Lemma coarse_list_cons_in : forall {T:Type} (l l':list T) (a:T), coarse_list (a::l) l' -> In a l'. intros; eapply in_coarse_list. apply in_eq. apply H. Qed. Lemma coarse_list_incl : forall {T:Type} (l l':list T), coarse_list l l' -> incl l l'. intros T l l' h1. induction h1 as [|l l' a ih h2|l l' a ih h2]. apply incl_nil. red in h2; red. intros x h3. specialize (h2 _ h3). right; auto. red in h2; red. intros x h3. destruct h3 as [|h3]; subst. left; auto. specialize (h2 _ h3). right; auto. Qed. Lemma coarse_list_cons1_not_nil : forall {T:Type} (l l':list T) (a:T), coarse_list (a::l) l' -> l' <> nil. intros ? ? ? ? h1; inversion h1; subst; apply cons_neq_nil. Qed. (*Scroll past definition of [lind_occ] for [coarse_list_rev_cons1_iff]*) Lemma coarse_list_rev_cons2_iff : forall {T:Type} (l l':list T) (a':T), coarse_list l (a'::l') <-> (l = nil \/ (exists (pf:l <> nil), hd' l pf = a' /\ coarse_list (tl l) l') \/ coarse_list l l'). intros T l l' a'; split; intro h1. dependent induction h1. left; auto. right; right; auto. right. left. exists (cons_neq_nil _ _). simpl. split; auto. destruct h1 as [|h1]; subst. constructor. destruct h1 as [h1 | h1]. destruct h1 as [h1 [h2 h3]]. subst. rewrite (hd_compat' _ h1) at 1. apply coarse_list_cons2; auto. apply coarse_list_cons1; auto. Qed. Lemma coarse_list_trans : forall (T:Type) (pfdt:type_eq_dec T), Transitive _ (@coarse_list T). intros T hdt; red; intros l1 l2 l3 h1. revert l3. induction h1. intros; constructor. intros ? h7; apply IHh1. apply coarse_list_rev_cons1 in h7; auto. intros l0 h2. generalize dependent l0. intro l0. revert IHh1 h1. revert a l' l. induction l0 as [|a0 l0 ih]. intros. apply not_coarse_list_cons_nil in h2. contradiction. intros a l' l h1 h2 h3. destruct (hdt a a0) as [|h4]; subst. apply coarse_list_cons2. pose proof h3 as h3'. apply coarse_list_rev_cons_impl in h3'. destruct h3' as [h3' | h3']. specialize (ih a0 l' l h1 h2 h3'). apply coarse_list_rev_cons1 in ih; auto. apply h1; auto. specialize (ih a l' l h1 h2). rewrite coarse_list_rev_cons2_iff in h3. destruct h3 as [h3 | h3]. discriminate. destruct h3 as [h3 | h3]. destruct h3 as [h3 h5]. simpl in h5. destruct h5; contradiction. specialize (ih h3). apply coarse_list_cons1; auto. Qed. Lemma coarse_list_rev_cons2_dec : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a':T), coarse_list l (a'::l') -> {l = nil} + {exists (pf:l <> nil), hd' l pf = a' /\ coarse_list (tl l) l'} + {coarse_list l l'}. intros T hdt l l' a h1. destruct (nil_dec' l) as [|h2]; subst. left; auto. destruct (hdt (hd' l h2) a) as [|h3]; subst. left; right. exists h2. split; auto. rewrite coarse_list_rev_cons2_iff in h1. destruct h1 as [|[|h1]]; subst. contradict h2; auto. destruct H as [? [? ?]]; auto. eapply coarse_list_trans; auto. apply coarse_list_tl. assumption. right. rewrite coarse_list_rev_cons2_iff in h1. destruct h1 as [|[|h1]]; subst. contradict (h2 (eq_refl _)). destruct H as [? [? ?]]; auto. pose proof (proof_irrelevance _ x h2); subst. contradict h3; auto. assumption. Qed. Lemma coarse_list_app_l : forall {T:Type} (l l1 l2:list T), coarse_list l l1 -> coarse_list l (l1 ++ l2). intros T l l1 l2 h1. revert l2. induction h1 as [l|l1 l2 a h2 ih|l1 l2 a h2 ih]. intros; constructor. intros l'; simpl. apply coarse_list_cons1; auto. intro l'. rewrite <- app_comm_cons. apply coarse_list_cons2; auto. Qed. Lemma coarse_list_app_r : forall {T:Type} (l l1 l2:list T), coarse_list l l2 -> coarse_list l (l1 ++ l2). intros T l l1. revert l. induction l1 as [|a1 l1 ih]. intros; rewrite app_nil_l; auto. intros l l2 h2. specialize (ih _ _ h2). rewrite <- app_comm_cons. apply coarse_list_cons1; auto. Qed. Lemma coarse_list_app_interp_compat : forall {T:Type} (pfdt:type_eq_dec T) (l l' l0 l1:list T), coarse_list l (l0++l') -> coarse_list l (l0++l1++l'). intros T h0 l l' l0 l1 h1. revert l1. dependent induction h1. intros; constructor. destruct l0 as [|a0 l0]. rewrite app_nil_l in x. subst. intro l1. specialize (IHh1 l'0 nil). rewrite app_nil_l in IHh1. specialize (IHh1 eq_refl (l1++a::nil)). rewrite app_nil_l. rewrite app_nil_l in IHh1; simpl in IHh1. rewrite <- app_assoc in IHh1. simpl in IHh1; auto. intros l1. rewrite <- app_comm_cons in x. inversion x; subst. specialize (IHh1 l' l0 eq_refl l1). rewrite <- app_comm_cons. apply coarse_list_cons1; auto. intro l1. destruct l0 as [|a0 l0]. rewrite app_nil_l in x; subst. rewrite app_nil_l. rewrite coarse_list_cons1_iff. ex_goal. rewrite app_eq_iff; simpl. intro h2. destruct h2 as [h2 h3]. rewrite skipn_nil in h3. discriminate. exists hex. destruct l1 as [|b l1]. left. generalize hex. rewrite app_nil_l. intro hex'. simpl. split; auto. simpl. destruct (h0 a b) as [|h2]; subst. left. split; auto. apply coarse_list_app_r. apply coarse_list_cons1; auto. right; split; auto. apply coarse_list_app_r. apply coarse_list_cons2; auto. assumption. rewrite <- app_comm_cons in x. inversion x; subst. rewrite <- app_comm_cons. apply coarse_list_cons2; auto. Qed. Lemma coarse_list_app2 : forall {T:Type} (la la' lb lb':list T), coarse_list la lb -> coarse_list la' lb' -> coarse_list (la ++ la') (lb ++ lb'). intros T la la' lb lb' h1. induction h1. intros; apply coarse_list_app_r; auto. intro h2; specialize (IHh1 h2). rewrite <- app_comm_cons. apply coarse_list_cons1; auto. intro h2; specialize (IHh1 h2). do 2 rewrite <- app_comm_cons. apply coarse_list_cons2; auto. Qed. (*This removes the coarse sub-list from the super-list.*) Fixpoint remove_coarse_list {T} {l l':list T} (pfdt:type_eq_dec T) : coarse_list l l' -> list T := match l' with | nil => fun _ => nil | a'::l' => fun (pf:coarse_list l (a'::l')) => match (coarse_list_rev_cons2_dec pfdt _ _ _ pf) with | inleft (left _) => nil | inleft (right pfe) => let (_, pfc') := constructive_definite_proof pfe in let (_, pfc) := pfc' in @remove_coarse_list _ (tl l) l' pfdt pfc | inright pfc => @remove_coarse_list _ l l' pfdt pfc end end. Lemma in_remove_coarse_list : forall {T} {l l':list T} (pfdt:type_eq_dec T) (pfc:coarse_list l l') x, In x (remove_coarse_list pfdt pfc) -> In x l'. intros T l l'. revert l. induction l' as [|a' l' ih]; simpl. intros; contradiction. intros l hdt h1 x h2. lr_match_sum' h2. destruct hl as [|hl]; subst. contradiction. destruct hl as [h3 [h4 h5]]; subst. destruct constructive_definite_description; simpl in h2. destruct a as [h4 h6]; subst. specialize (ih _ hdt h6 x h2). right; auto. specialize (ih _ hdt hr x h2). right; auto. Qed. Definition coarse_lists {T:Type} (l:list T) := [l':list T | coarse_list l' l]. Lemma inh_coarse_lists : forall {T:Type} (l:list T), Inhabited (coarse_lists l). intros; econstructor; constructor; apply coarse_list_nil. Qed. Lemma coarse_lists_nil : forall T:Type, (coarse_lists (@nil T)) = Singleton nil. intro T. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1]. apply coarse_list_nil2_iff in h1. subst. constructor; auto. destruct h1. constructor. apply coarse_list_nil. Qed. Lemma coarse_lists_cons : forall {T:Type} (l:list T) (a:T), coarse_lists (a::l) = Union (coarse_lists l) (Im (coarse_lists l) (cons a)). intros T l a; apply Extensionality_Ensembles; red; split; red; intros lx h1. destruct h1 as [h1]. rewrite coarse_list_rev_cons2_iff in h1. destruct h1 as [|h1]; subst. left. constructor. apply coarse_list_nil. destruct h1 as [h1 | h1]. destruct h1 as [h1 [h2 h3]]; subst. right. apply Im_intro with (tl lx). constructor; auto. apply hd_compat'. left. constructor; auto. destruct h1 as [lx h1 | lx h1]. destruct h1 as [h1]. constructor. apply coarse_list_cons1; auto. destruct h1 as [lx h1]; subst. destruct h1 as [h1]. constructor. apply coarse_list_cons2; auto. Qed. Lemma finite_coarse_lists : forall {T:Type} (l:list T), Finite (coarse_lists l). intros T l. induction l as [|a l ih]; simpl. rewrite coarse_lists_nil. apply Singleton_is_finite. rewrite coarse_lists_cons. apply Union_preserves_Finite; auto. apply finite_image; auto. Qed. Definition coarse_seg_list (l:list nat) n : Prop := coarse_list l (seg_list n). Lemma coarse_list_app_seg_list_n : forall {n l}, coarse_seg_list l n -> coarse_seg_list (l++n::nil) (S n). intros n l h1; red in h1; red. rewrite seg_list_app; auto with arith. simpl. rewrite <- minus_n_O. apply coarse_list_app2; auto. apply coarse_list_refl. Qed. Lemma coarse_list_inv_im_in_dep : forall {T U:Type} {l:list T} (pfdu:type_eq_dec U) (f:fun_in_dep l U) (u:U), coarse_list (inv_im_in_dep pfdu l u f) l. intros T U l. induction l as [|a l ih]; simpl. intros; constructor. intros h1 f u. destruct (h1 u (f a (in_eq a l))) as [|h2]; subst. apply coarse_list_cons2; auto. apply coarse_list_cons1; auto. Qed. Lemma coarse_list_inv_iml : forall {T U:Type} {l:list T} (pfdu:type_eq_dec U) (f:T->U) (u:U), coarse_list (inv_iml pfdu f l u) l. intros T U l. induction l as [|a l ih]; simpl. intros; constructor. intros h1 f u. destruct (h1 u (f a)) as [|h2]; subst. apply coarse_list_cons2; auto. apply coarse_list_cons1; auto. Qed. Lemma coarse_list_removelast : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), coarse_list (removelast l) l. intros T hdt l. destruct (nil_dec' l) as [|h1]; subst; simpl. constructor. rewrite (app_removelast_lst _ h1) at 2. apply coarse_list_app_l. apply coarse_list_refl. Qed. Lemma no_dup_coarse_list : forall {T:Type} (l l':list T), coarse_list l l' -> NoDup l' -> NoDup l. intros T l l' h1. induction h1. intro; constructor. intro; apply IHh1. apply no_dup_cons in H; auto. intro. constructor. intro h2. pose proof (in_coarse_list _ _ _ h2 h1) as h3. pose proof (no_dup_cons_nin _ _ H); contradiction. apply IHh1. apply no_dup_cons in H; auto. Qed. Lemma no_dup_removelast : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), NoDup l -> NoDup (removelast l). intros; eapply no_dup_coarse_list; auto. apply coarse_list_removelast; auto. assumption. Qed. Lemma no_dup_inv_im_in_dep : forall {T U:Type} (pfdu:type_eq_dec U) (l:list T) (u:U) (f:fun_in_dep l U), NoDup l -> NoDup (inv_im_in_dep pfdu l u f). intros; eapply no_dup_coarse_list. apply coarse_list_inv_im_in_dep. assumption. Qed. Lemma no_dup_app' : forall {T:Type} (l l':list T), NoDup (l ++ l') -> NoDup l /\ NoDup l'. intros T l l'. revert l. induction l' as [|a' l' ih]; simpl. intros ? h1. rewrite app_nil_r in h1. split; auto. constructor. intros l h1. specialize (ih (l++a'::nil)). hfirst ih. rewrite <- app_assoc. rewrite <- app_comm_cons. simpl; auto. specialize (ih hf). destruct ih as [h2 h3]. pose proof (no_dup_coarse_list l (l++a'::nil)) as h4. hfirst h4. apply coarse_list_app_l. apply coarse_list_refl. specialize (h4 hf0). specialize (h4 h2). split; auto. pose proof (no_dup_coarse_list (a'::l') (l++a'::l')) as h5. hfirst h5. apply coarse_list_app_r. apply coarse_list_refl. specialize (h5 hf1 h1); auto. Qed. Lemma no_dup_app_intersection : forall {T:Type} (l l':list T), NoDup (l ++ l') -> Intersection (list_to_set l) (list_to_set l') = Empty_set _. intros T l l'. revert l. induction l' as [|a' l' ih]; simpl. intros. rewrite Intersection_commutative. apply empty_intersection. intros l h1. pose proof (ih (l++a'::nil)) as ih'. hfirst ih'. rewrite <- app_assoc, <- app_comm_cons; simpl; auto. specialize (ih' hf). rewrite list_to_set_app in ih'. simpl in ih'. rewrite add_empty_sing in ih'. rewrite Intersection_commutative in ih'. rewrite <- ih'. unfold Add. rewrite Distributivity. rewrite Distributivity. rewrite Intersection_commutative. f_equal. pose proof (no_dup_app' _ _ h1) as h1'. destruct h1' as [h2 h3]. pose proof (ih (a'::nil) h3) as ih0. simpl in ih0. rewrite add_empty_sing in ih0. rewrite Intersection_commutative in ih0. rewrite ih0. apply no_dup_app' in hf. destruct hf as [h4 h5]. rewrite <- nin_iff. rewrite <- list_to_set_in_iff. apply no_dup_app_sing_nin; auto. Qed. Lemma no_dup_app_nin : forall {T:Type} (l l':list T), NoDup (l ++ l') -> forall x:T, ~(In x l /\ In x l'). intros T l l' h1 x h2. pose proof (no_dup_app_intersection l l' h1) as h3. destruct h2 as [h2 h4]. rewrite list_to_set_in_iff in h2, h4. pose proof (Intersection_intro _ _ _ _ h2 h4) as h5. rewrite h3 in h5. contradiction. Qed. Lemma no_dup_map_seg_inj' : forall {U:Type} {n:nat} (f:seg_fun n U), NoDup (map_seg f) -> inj_seg f. intros U n. induction n as [|n ih]; simpl. intros; red; intros; omega. intros f h1. specialize (ih (seg_fun_from_S f)). pose proof (no_dup_app' _ _ h1) as h2. destruct h2 as [h2 h3]. specialize (ih h2). red in ih; red. intros i j h5 h6. pose proof h5 as h5'. pose proof h6 as h6'. apply le_lt_eq_dec in h5'. apply le_lt_eq_dec in h6'. destruct h5' as [h7 | h7], h6' as [h8 | h8]. intros; eapply ih. unfold seg_fun_from_S. erewrite f_equal_dep. rewrite H at 1. f_equal. apply proof_irrelevance. reflexivity. apply S_inj in h8; subst. intro h8. pose proof (no_dup_app_sing_nin _ _ h1) as h9. contradict h9. erewrite f_equal_dep. rewrite <- h8 at 1. rewrite in_map_seg_iff. exists i, (lt_S_n _ _ h7). unfold seg_fun_from_S; f_equal; apply proof_irrelevance. reflexivity. apply S_inj in h7; subst. intro h9. pose proof (no_dup_app_sing_nin _ _ h1) as h10. contradict h10. erewrite f_equal_dep. rewrite h9 at 1. rewrite in_map_seg_iff. exists j, (lt_S_n _ _ h8). unfold seg_fun_from_S. f_equal. apply proof_irrelevance. reflexivity. omega. Grab Existential Variables. apply lt_S_n in h8; auto. apply lt_S_n in h7; auto. Qed. Definition mask_to_coarse_list {T:Type} {l:list T} {v:bv} (pfeq:length l = length v) : list T := map_in_dep (fun (i:nat) (pfi:In i (bv1s v)) => let pfi' := lt_eq_trans _ _ _ (in_bv1s_lt _ _ pfi) (eq_sym pfeq) in nth_lt l i pfi'). Lemma length_mask_to_coarse_list : forall {T:Type} {l:list T} {v:bv} (pfeq:length l = length v), length (mask_to_coarse_list pfeq) = length (bv1s v). intros; unfold mask_to_coarse_list; rewrite length_map_in_dep; auto. Qed. Lemma mask_to_coarse_list_eq_nil_iff : forall {T:Type} {l:list T} {v:bv} (pfeq:length l = length v), mask_to_coarse_list pfeq = nil <-> bv1s v = nil. unfold mask_to_coarse_list; intros; rewrite map_in_dep_eq_nil_iff, bv1s_eq_nil_iff; tauto. Qed. (*Is True if [l] can be gotten from [l'] via [v], when [coarse_list l l'] is constructed via the "mask" [v].*) (*There are two variants.*) (*The first version is for easy construction, given essential components. The more verbose version (primed) is easier to prove in a goal, when you are applying the [constructor] tactic. Think of the difference between the two as being analogous to [Inductive Im], as it's done in the library, and without the redundant parameteters. The latter is easier when constructing [In (Im A f) (f x)], the former (library version) is easier when using [econstructor] in a goal.*) Inductive coarse_list_mask {T} : forall (l l':list T), coarse_list l l' -> forall (v:bv), length l' = length v -> Prop := | coarse_list_mask_nil : forall l, coarse_list_mask nil l (coarse_list_nil l) (lrep false (length l)) (eq_sym (length_lrep false (length l))) | coarse_list_mask_cons1 : forall l l' a (pfc:coarse_list l l') v (pfe:length l' = length v), coarse_list_mask l l' pfc v pfe -> coarse_list_mask l (a::l') (coarse_list_cons1 l l' a pfc) (false::v) (f_equal S pfe) | coarse_list_mask_cons2 : forall l l' a (pfc:coarse_list l l') v (pfe:length l' = length v), coarse_list_mask l l' pfc v pfe -> coarse_list_mask (a::l) (a::l') (coarse_list_cons2 l l' a pfc) (true::v) (f_equal S pfe). Inductive coarse_list_mask' {T} : forall (l l':list T), coarse_list l l' -> forall (v:bv), length l' = length v -> Prop := | coarse_list_mask_nil' : forall l v (pfc:coarse_list nil l) (pfe:length l = length v), all_sing v false -> coarse_list_mask' nil l pfc v pfe | coarse_list_mask_cons1' : forall l l' a (pfc:coarse_list l l') (pfc':coarse_list l (a::l')) v (pfe:length l' = length v) (pfe':S (length l') = S (length v)) (b:bool), b = false -> coarse_list_mask' l l' pfc v pfe -> coarse_list_mask' l (a::l') pfc' (b::v) pfe' | coarse_list_mask_cons2': forall l l' a (pfc:coarse_list l l') (pfc':coarse_list (a::l) (a::l')) v (pfe:length l' = length v) (pfe':S (length l') = S (length v)) (b:bool), b = true -> coarse_list_mask' l l' pfc v pfe -> coarse_list_mask' (a::l) (a::l') pfc' (b::v) pfe'. Lemma coarse_list_mask_functional : forall {T} (l l' l0 l0':list T) v v' (pfle: l = l0) (pfle': l' = l0') (pfc:coarse_list l l') (pfc':coarse_list l0 l0') (pfv:v = v') (pfe:length l' = length v) (pfe':length l0' = length v'), coarse_list_mask l l' pfc v pfe -> coarse_list_mask l0 l0' pfc' v' pfe'. intros; subst; pose proof (proof_irrelevance _ pfc pfc'); pose proof (proof_irrelevance _ pfe pfe'); subst; auto. Qed. Lemma coarse_list_mask_functional' : forall {T} (l l' l0 l0':list T) v v' (pfle: l = l0) (pfle': l' = l0') (pfc:coarse_list l l') (pfc':coarse_list l0 l0') (pfv:v = v') (pfe:length l' = length v) (pfe':length l0' = length v'), coarse_list_mask' l l' pfc v pfe -> coarse_list_mask' l0 l0' pfc' v' pfe'. intros; subst; pose proof (proof_irrelevance _ pfc pfc'); pose proof (proof_irrelevance _ pfe pfe'); subst; auto. Qed. Lemma coarse_list_mask_iff' : forall {T} {l l':list T} {v:bv} (pfc:coarse_list l l') (pfe:length l' = length v), coarse_list_mask l l' pfc v pfe <-> coarse_list_mask' l l' pfc v pfe. intros; split; intro h1; induction h1; subst. constructor. red; intros x h1; apply in_lrep' in h1; auto. eapply coarse_list_mask_cons1'; auto; apply IHh1; auto. eapply coarse_list_mask_cons2'; auto; apply IHh1; auto. pose proof (coarse_list_mask_nil l) as h2. pose proof (all_sing_lrep _ _ H) as h1. pose proof (proof_irrelevance _ pfc (coarse_list_nil l)) as h4. gen0. rewrite h1. rewrite <- pfe. rewrite h4. intro pfe'. pose proof (proof_irrelevance _ pfe' (eq_sym (length_lrep false (length l)))) as h5. rewrite h5; auto. pose proof (coarse_list_mask_cons1 _ _ a pfc v pfe IHh1) as h2. pose proof (proof_irrelevance _ pfc' (coarse_list_cons1 l l' a pfc)). pose proof (proof_irrelevance _ pfe' (f_equal S pfe)); subst; constructor; auto. pose proof (coarse_list_mask_cons2 _ _ a pfc v pfe IHh1) as h2. pose proof (proof_irrelevance _ pfc' (coarse_list_cons2 l l' a pfc)). pose proof (proof_irrelevance _ pfe' (f_equal S pfe)); subst; constructor; auto. Qed. Lemma coarse_list_mask_nil_iff : forall {T} {l':list T} {v:bv} (pfc:coarse_list nil l') (pfe:length l' = length v), coarse_list_mask nil l' pfc v pfe <-> v = lrep false (length l'). intros; split; intro h3; dependent induction h3; subst; auto. rewrite coarse_list_mask_iff'; constructor. red. intros x h1. apply in_lrep' in h1; auto. Qed. Lemma coarse_list_mask_cons_true_not_nil : forall {T} {l l':list T} {v:bv} {a'} (pfc:coarse_list l (a'::l')) (pfe:length (a'::l') = length (true::v)), coarse_list_mask l (a'::l') pfc (true::v) pfe -> l <> nil. intros; intro; subst; inversion H. Qed. Lemma coarse_list_mask_cons_true_iff : forall {T} {l l':list T} {v:bv} {a'} (pfc:coarse_list l (a'::l')) (pfe:length (a'::l') = length (true::v)), coarse_list_mask l (a'::l') pfc (true::v) pfe <-> exists (pfn:l <> nil), a' = hd' l pfn /\ exists pfc':coarse_list (tl l) l', coarse_list_mask (tl l) l' pfc' v (S_inj _ _ pfe). intros T l l' v a' hc h1; split; intro h2. exists (coarse_list_mask_cons_true_not_nil _ _ h2). inversion h2; subst; simpl; auto. split; auto. exists pfc. eapply coarse_list_mask_functional in H1; auto. apply H1. destruct h2 as [h2 [h3 h4]]. revert h4. subst. intro h4. destruct h4 as [h4 h5]. pose proof (coarse_list_mask_cons2 (tl l) l' (hd' l h2) h4 v (S_inj _ _ h1) h5) as h6. eapply coarse_list_mask_functional in h6. apply h6. rewrite (hd_compat' _ h2); auto. reflexivity. reflexivity. Qed. Lemma coarse_list_mask_same_nil : forall {T} (a c:T) b (pfeq':length (a::nil) = length (b::nil)) (pfc':coarse_list (a::nil) (c::nil)), coarse_list_mask (a :: nil) (c :: nil) pfc' (b::nil) pfeq' -> a = c /\ b = true. intros T a c b h1 h2 h3. inversion h2; subst. rewrite coarse_list_nil2_iff in H1. discriminate. split; auto. inversion h3; subst; auto. pose proof (coarse_list_nil2_iff (c::nil)) as h5. pose proof pfc as h6. rewrite h5 in h6. discriminate. Qed. Lemma coarse_list_mask_cons_false_iff : forall {T} {l l':list T} {v:bv} {a'} (pfc:coarse_list l (a'::l')) (pfe:length (a'::l') = length (false::v)), coarse_list_mask l (a'::l') pfc (false::v) pfe <-> exists pfc':coarse_list l l', coarse_list_mask l l' pfc' v (S_inj _ _ pfe). intros T l l'. destruct l' as [|a' l']. simpl. intros v a h1 h2. pose proof (S_inj _ _ h2) as h3. symmetry in h3. apply length_eq0 in h3; subst. inversion h1; subst. split; intro h3. exists (coarse_list_nil nil). rewrite coarse_list_mask_iff'. constructor. red; intros; contradiction. rewrite coarse_list_mask_iff'. constructor. red; intros x hx; destruct hx; auto. contradiction. rewrite coarse_list_nil2_iff in H1; subst. split; intro h3. exists (coarse_list_nil nil). rewrite coarse_list_mask_iff'. constructor. red; intros; contradiction. rewrite coarse_list_mask_iff'. constructor. red; intros x hx; destruct hx; auto. contradiction. rewrite coarse_list_nil2_iff in H1; subst. split; intro h3. apply coarse_list_mask_same_nil in h3. destruct h3; discriminate. destruct h3 as [h3 h4]. pose proof h3 as h5. rewrite coarse_list_nil2_iff in h5; subst. discriminate. intros v a hc he. split; intro h1. inversion h1; subst. exists (coarse_list_nil (a'::l')). rewrite coarse_list_mask_iff'. constructor. red; intros x h2. destruct h2 as [|h2]; auto. apply in_lrep' in h2; auto. exists pfc. gterm0. pose proof (proof_irrelevance _ c pfe) as h3. fold c. rewrite h3; auto. destruct h1 as [h1 h2]. erewrite coarse_list_mask_iff'. econstructor. reflexivity. rewrite <- coarse_list_mask_iff'. apply h2. Qed. (*The set of all bit-masks that can construct a given coarse-sub-list.*) Definition coarse_list_masks {T} {l l':list T} (pfc:coarse_list l l') : Ensemble bv := [v:bv | exists pfe:length l' = length v, coarse_list_mask l l' pfc v pfe]. Lemma finite_bv_eq_length : forall {T} (l:list T), Finite [v:bv | length l = length v]. intros T l. induction l as [|a l ih]; simpl. gterm0. fold c. assert (h1:c = Singleton nil). apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1]. symmetry in h1. apply length_eq0 in h1; subst. constructor. destruct h1. constructor. auto. rewrite h1. apply Singleton_is_finite. gterm0. fold c. term0 ih. fold c0 in ih. assert (h1:c = Union (Im c0 (fun l => true::l)) (Im c0 (fun l => false::l))). apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [h1]. destruct x as [|b v]; try discriminate. simpl in h1. apply S_inj in h1. destruct b. left. apply Im_intro with v; auto. constructor; auto. right. apply Im_intro with v; auto. constructor; auto. destruct h1 as [x h1 | x h1]; destruct h1 as [x h1]; subst; destruct h1 as [h1]; constructor; simpl; auto with arith. rewrite h1. apply Union_preserves_Finite; apply finite_image; auto. Qed. Lemma finite_coarse_lists_masks : forall {T} {l l':list T} (pfc:coarse_list l l'), Finite (coarse_list_masks pfc). intros; eapply Finite_downward_closed. apply (finite_bv_eq_length l'). red; intros x h1. destruct h1 as [h1]. destruct h1 as [h1 h2]. constructor; auto. Qed. Lemma inh_coarse_lists_masks : forall {T} {l l':list T} (pfc:coarse_list l l'), Inhabited (coarse_list_masks pfc). intros T l l'; revert l. induction l' as [|a l' ih]. intros. pose proof pfc as h1. rewrite coarse_list_nil2_iff in h1; subst. apply Inhabited_intro with nil. constructor. exists (eq_refl _). rewrite coarse_list_mask_iff'. constructor. red; intros x h1. contradiction. intros l h1. inversion h1; subst. apply Inhabited_intro with (lrep false (S (length l'))). constructor. exists (eq_sym (length_lrep _ _)). simpl. rewrite coarse_list_mask_iff'. constructor. red; intros x h2. destruct h2 as [|h2]; auto. apply in_lrep' in h2; auto. specialize (ih l H1). destruct ih as [v h3]. destruct h3 as [h3]. destruct h3 as [h3 h4]. apply Inhabited_intro with (false::v). constructor. exists (f_equal S h3). pose proof (coarse_list_mask_cons1 l l' a _ _ h3 h4) as h5. pose proof (proof_irrelevance _ h1 (coarse_list_cons1 l l' a H1)); subst; auto. specialize (ih _ H1). destruct ih as [v h2]. destruct h2 as [h2]. destruct h2 as [h2 h3]. pose proof (coarse_list_mask_cons2 l0 l' a H1 v h2 h3) as h4. apply Inhabited_intro with (true::v). constructor. exists (f_equal S h2). pose proof (proof_irrelevance _ h1 (coarse_list_cons2 l0 l' a H1)); subst; auto. Qed. Definition coarse_list_to_mask {T} {l l':list T} (pfc:coarse_list l l') : bv := min_set_sortable (coarse_list_masks pfc) (finite_coarse_lists_masks pfc) (inh_coarse_lists_masks pfc) bv_lt sortable_bv_lt. Lemma coarse_list_to_mask_compat : forall {T} {l l':list T} (pfc:coarse_list l l'), exists pfe:length l' = length (coarse_list_to_mask pfc), coarse_list_mask l l' pfc (coarse_list_to_mask pfc) pfe /\ forall (v:bv) (pfe':length l' = length v), v <> coarse_list_to_mask pfc -> coarse_list_mask l l' pfc v pfe' -> bv_lt (coarse_list_to_mask pfc) v. intros T l l' h1. pose proof (min_set_sortable_compat (coarse_list_masks h1) (finite_coarse_lists_masks h1) (inh_coarse_lists_masks h1) bv_lt sortable_bv_lt) as h2. simpl in h2. destruct h2 as [h2 h3]. destruct h2 as [h2]. destruct h2 as [h2 h4]. exists h2. split; auto. intros v he h5 h6. specialize (h3 v). apply h3. constructor. constructor. exists he; auto. contradict h5. inversion h5; auto. Qed. Lemma length_cons_count_occ_false : forall {T:Type} (l:list T) (a:T) (v:bv) (pfleq:length (a::l) = count_occ bool_eq_dec (false::v) false), length l = count_occ bool_eq_dec v false. intros T l a v h1. simpl in h1. lr_if h1; fold hlr in h1; destruct hlr; auto with arith. contradict n; auto. Qed. Lemma length_count_occ_true_false : forall {T:Type} (l:list T) (v:bv) (pfleq:length l = count_occ bool_eq_dec (true::v) false), length l = count_occ bool_eq_dec v false. intros T l v h1. simpl in h1. lr_if h1; fold hlr in h1; destruct hlr; auto with arith. discriminate. Qed. (*Inserts an [x] in between the elements of [l] so that the final indices of [x] in the returned list correspond to the [trues] in [v]. Think of it as an inverse of [remove]. Assumes [~In x l], althought it doesn't use it in the computation, only in the compatibility*) Fixpoint interp_bv_aux {T:Type} (x:T) (acc : list T) (v:bv) : forall (l:list T), length l = count_occ bool_eq_dec v false -> list T := match v with | nil => fun _ _ => acc | b::v' => match b with | true => fun l (pfleq:length l = count_occ bool_eq_dec (true::v') false) => let acc' := acc ++ x::nil in let pfleq' := length_count_occ_true_false _ _ pfleq in interp_bv_aux x acc' v' l pfleq' | false => fun l => match l with | nil => (fun _ => acc) | a::l' => (fun (pfleq:length (a::l') = count_occ bool_eq_dec (false::v') false) => let acc' := acc ++ a::nil in let pfleq' := length_cons_count_occ_false _ _ _ pfleq in interp_bv_aux x acc' v' l' pfleq') end end end. Lemma interp_bv_aux_functional : forall {T:Type} (x x':T) (acc acc': list T) (v v':bv) (l l':list T) (pf:length l = count_occ bool_eq_dec v false) (pf':length l' = count_occ bool_eq_dec v' false), x = x' -> acc = acc' -> v = v' -> l = l' -> interp_bv_aux x acc v l pf = interp_bv_aux x acc' v' l' pf'. intros; subst; f_equal; apply proof_irrelevance. Qed. Definition interp_bv {T:Type} (x:T) (v:bv) (l:list T) (pfe:length l = count_occ bool_eq_dec v false) := interp_bv_aux x nil v l pfe. Lemma interp_bv_functional : forall {T:Type} (x x':T) (v v':bv) (l l':list T) (pf:length l = count_occ bool_eq_dec v false) (pf':length l' = count_occ bool_eq_dec v' false), x = x' -> v = v' -> l = l' -> interp_bv x v l pf = interp_bv x v' l' pf'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma length_interv_bv_aux : forall {T:Type} (x:T) (v:bv) (l:list T) (pfe:length l = count_occ bool_eq_dec v false) (acc:list T), length (interp_bv_aux x acc v l pfe) = length v + length acc. intros T x v. revert x. induction v as [|b v ih]; simpl; auto. intros x l h1. pose proof h1 as h1'. lr_if h1'; fold hlr in h1'; destruct hlr as [hl | hr]; subst. intro acc. destruct l as [|a l]. simpl in h1'. discriminate. specialize (ih x l (length_cons_count_occ_false l a v h1) (acc ++ a::nil)). rewrite ih, app_length, S_compat, <- plus_S_shift_r; auto. intro acc. destruct b. rewrite ih, app_length, S_compat, <- plus_S_shift_r; auto. contradict hr; auto. Qed. Lemma length_interp_bv : forall {T:Type} (x:T) (v:bv) (l:list T) (pfe:length l = count_occ bool_eq_dec v false), length (interp_bv x v l pfe) = length v. unfold interp_bv; intros; rewrite length_interv_bv_aux; simpl; auto. Qed. Lemma length_count_occ_false_tl : forall {T:Type} (x:T) (l:list T) (v:bv), length l = count_occ bool_eq_dec (false::v) false -> length (tl l) = count_occ bool_eq_dec v false. intros T x l v h1. destruct l as [|a l]; simpl. simpl in h1. lr_if h1; fold hlr in h1; destruct hlr as [h3 | h3]; subst. discriminate. contradict h3; auto. simpl in h1. lr_if h1; fold hlr in h1; destruct hlr as [h3 | h3]; subst; auto with arith. contradict h3; auto. Qed. Lemma length_count_occ_true : forall {T:Type} (x:T) (l:list T) (v:bv), length l = count_occ bool_eq_dec (true::v) false -> length l = count_occ bool_eq_dec v false. intros. rewrite count_occ_cons_neq in H; auto. intro; discriminate. Qed. Lemma length_count_occ_true_true : forall {T:Type} (x:T) (l:list T) (v:bv), length l = count_occ bool_eq_dec (true::(true::v)) false -> length l = count_occ bool_eq_dec v false. intros. rewrite count_occ_cons_neq in H; auto. rewrite count_occ_cons_neq in H; auto. intro; discriminate. intro; discriminate. Qed. Lemma interp_bv_aux_false : forall {T:Type} (x:T) (v:bv) (l acc:list T) (pfle:length l = count_occ bool_eq_dec (false::v) false) (pfn:l <> nil), let pfle' := length_count_occ_false_tl x _ _ pfle in interp_bv_aux x acc (false :: v) l pfle = acc ++ hd' l pfn :: interp_bv_aux x nil v (tl l) pfle'. intros T x v. revert x. induction v as [|b v ih]. intros ? l acc h1 h2. destruct l as [|a l]. contradict h2; auto. simpl; auto. destruct b. intros x l acc h1 h2 h3. simpl. destruct l as [|a l]. contradict h2; auto. pose proof ih as ih'. specialize (ih x (x::l) (acc++a::nil)). pose proof h1 as h1'. simpl in h1'. lr_if h1'; fold hlr in h1'; destruct hlr as [hl | hr]. lr_if h1'; fold hlr in h1'; destruct hlr as [hl' | hr']. discriminate. apply S_inj in h1'. hfirst ih. simpl. lr_if_goal; fold hlr; destruct hlr as [hl0 | hr0]. f_equal; auto. contradict hr0; auto. specialize (ih hf (cons_neq_nil _ _)). simpl in ih. erewrite f_equal_dep. rewrite ih at 1. rewrite <- app_assoc. simpl. do 2 f_equal. specialize (ih' x (x::l) nil). simpl in ih'. hfirst ih'. lr_if_goal; fold hlr; destruct hlr as [hl0 | hr0]. f_equal; auto. contradict hr0; auto. specialize (ih' hf0 (cons_neq_nil _ _)). symmetry. erewrite f_equal_dep. rewrite ih' at 1. repeat f_equal. apply proof_irrelevance. reflexivity. reflexivity. contradict hr; auto. intros x l acc h1 h2 h3. pose proof ih as ih'. pose proof h1 as h1'. simpl in h1'. lr_if h1'; fold hlr in h1'; destruct hlr as [hl | hr]. specialize (ih x (tl l) (acc++hd' l h2::nil)). hfirst ih. omega. specialize (ih hf). hfirst ih. destruct l as [|a l]. simpl in h1'. discriminate. simpl. destruct l as [|l]. simpl in h1'. apply S_inj in h1'. discriminate. apply cons_neq_nil. specialize (ih hf0). simpl. simpl in ih. destruct l as [|b l]. simpl in h1'. discriminate. destruct l as [|c l]. simpl; auto. simpl. simpl in ih. erewrite f_equal_dep. rewrite ih at 1. rewrite <- app_assoc. simpl. do 2 f_equal. specialize (ih' x (c::l) nil). hfirst ih'. simpl. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. simpl in h1'. apply S_inj in h1'; auto. contradict hr'; auto. specialize (ih' hf1 (cons_neq_nil _ _)). simpl in ih'. symmetry. erewrite f_equal_dep. rewrite ih' at 1. repeat f_equal. apply proof_irrelevance. reflexivity. reflexivity. contradict hr; auto. Qed. Lemma interp_bv_false_nil : forall {T:Type} (x:T) (v:bv) (pfle:0 = count_occ bool_eq_dec (false::v) false), interp_bv x (false::v) nil pfle = nil. unfold interp_bv; intros; simpl; auto. Qed. Lemma interp_bv_false : forall {T:Type} (x:T) (v:bv) (l:list T) (pfle:length l = count_occ bool_eq_dec (false::v) false) (pfn:l <> nil), let pfle' := length_count_occ_false_tl x _ _ pfle in interp_bv x (false::v) l pfle = (hd' _ pfn)::interp_bv x v (tl l) pfle'. intros T x v l h1 h2 h3. unfold interp_bv. erewrite interp_bv_aux_false. simpl. reflexivity. Qed. Lemma interp_bv_aux_true_true : forall {T:Type} (x:T) (v:bv) (l acc:list T) (pfle:length l = count_occ bool_eq_dec (true::true::v) false) (pfn:l <> nil), (forall (x : T) (l acc : list T) (pfle : length l = count_occ bool_eq_dec (true :: v) false), l <> nil -> let pfle' := length_count_occ_true x l v pfle in interp_bv_aux x acc (true :: v) l pfle = acc ++ x :: interp_bv_aux x nil v l pfle') -> let pfle' := length_count_occ_true x _ _ pfle in interp_bv_aux x acc (true :: true :: v) l pfle = acc ++ x :: interp_bv_aux x nil (true :: v) l pfle'. intros T x v l acc h1 h2 ih h4. pose proof h1 as h1'. simpl in h1'. lr_if h1'; fold hlr in h1'; destruct hlr as [hl | hr]. discriminate. pose proof ih as ih'. specialize (ih x l (acc++x::nil)). hfirst ih. rewrite count_occ_cons_neq; auto. specialize (ih hf h2). simpl in ih. simpl. erewrite f_equal_dep. rewrite ih at 1; simpl. rewrite <- app_assoc. simpl. do 2 f_equal. specialize (ih' x l nil). hfirst ih'. simpl. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. discriminate. simpl in h1'; auto. specialize (ih' hf0 h2). simpl in ih'. symmetry. erewrite f_equal_dep. rewrite ih' at 1. repeat f_equal. apply proof_irrelevance. reflexivity. reflexivity. Qed. Lemma interp_bv_aux_nil : forall {T:Type} (x:T) (v:bv) (acc:list T) (pf:count_occ bool_eq_dec v false = 0), interp_bv_aux x acc v nil (eq_sym pf) = acc ++ lrep x (length v). intros T x v. revert x. induction v as [|b v ih]; simpl. intros. rewrite app_nil_r; auto. intros x acc h1. pose proof h1 as h1'. lr_if h1'; fold hlr in h1'; destruct hlr as [hl | hr]. subst. discriminate. destruct b. specialize (ih x (acc++x::nil) h1'). pose proof (proof_irrelevance _ (eq_sym h1') (length_count_occ_true_false (@nil T) v (eq_sym h1))) as h8. rewrite <- h8. rewrite <- app_assoc in ih; auto. contradict hr; auto. Qed. Lemma interp_bv_aux_true : forall {T:Type} (x:T) (v:bv) (l acc:list T) (pfle:length l = count_occ bool_eq_dec (true::v) false) (pfn:l <> nil), let pfle' := length_count_occ_true x _ _ pfle in interp_bv_aux x acc (true :: v) l pfle = acc ++ x :: interp_bv_aux x nil v l pfle'. intros T x v. revert x. induction v as [|b v ih]. intros; auto. destruct b. intros x l acc h1 h2 h3. rewrite interp_bv_aux_true_true; simpl. repeat f_equal. assumption. assumption. intros x l acc h1 h2 h3. destruct l as [|a l]; auto. pose proof h1 as h1'. rewrite count_occ_cons_neq in h1'. assert (h4:interp_bv_aux x (acc++x::nil) (false::v) (a::l) h1' = interp_bv_aux x acc (true :: false :: v) (a :: l) h1). simpl. repeat f_equal. apply proof_irrelevance. rewrite <- h4. do 2 erewrite interp_bv_aux_false. simpl. rewrite <- app_assoc. repeat f_equal. simpl. repeat f_equal. apply proof_irrelevance. intro; discriminate. Grab Existential Variables. apply cons_neq_nil. apply cons_neq_nil. Qed. Lemma interp_bv_true : forall {T:Type} (x:T) (v:bv) (l:list T) (pfle:length l = count_occ bool_eq_dec (true::v) false), let pfle' := eq_trans pfle (count_occ_cons_neq bool_eq_dec v diff_true_false) in interp_bv x (true::v) l pfle = x::interp_bv x v l pfle'. unfold interp_bv; intros T x v l h1. destruct (nil_dec' l) as [|h2]. subst. pose proof (interp_bv_aux_nil x (true::v) nil (eq_sym h1)) as h2. pose proof (proof_irrelevance _ h1 (eq_sym (eq_sym h1))) as h3. rewrite h3 at 1. simpl; simpl in h2. rewrite h2. f_equal. pose proof (interp_bv_aux_nil x v nil (eq_sym (eq_trans h1 (count_occ_cons_neq bool_eq_dec v diff_true_false)))) as h4. simpl in h4. simpl. pose proof (proof_irrelevance _ (eq_trans h1 (count_occ_cons_neq bool_eq_dec v diff_true_false)) (eq_sym (eq_sym (eq_trans h1 (count_occ_cons_neq bool_eq_dec v diff_true_false))))) as h5. rewrite h5 at 1. rewrite h4 at 1. reflexivity. rewrite interp_bv_aux_true; simpl; auto. repeat f_equal. apply proof_irrelevance. Qed. Lemma in_interp_bv : forall {T:Type} (x:T) (v:bv) (l:list T) (pf:length l = count_occ bool_eq_dec v false), In true v -> In x (interp_bv x v l pf). intros T x v. revert x. induction v as [|b v ih]. intros; contradiction. intros x l h1 h2. destruct h2 as [|h2]; subst. rewrite interp_bv_true. left; auto. destruct b. rewrite interp_bv_true. left; auto. pose proof h1 as h1'. rewrite count_occ_cons_eq in h1'. pose proof h1' as h3. apply length_eq_S in h3. rewrite (interp_bv_false _ _ _ _ h3). right. apply ih; auto. reflexivity. Qed. Lemma nin_false_interp_bv : forall {T:Type} (x:T) (v:bv) (l:list T) (pf:length l = count_occ bool_eq_dec v false), ~ In false v -> interp_bv x v l pf = lrep x (length v). intros T x v. revert x. induction v as [|b v ih]. intros x l h1. pose proof h1 as h1'. apply length_eq0 in h1'. subst. intro. unfold interp_bv; auto. intros x l h1 h2. destruct b. pose proof h1 as h1'. rewrite count_occ_cons_neq in h1'; auto. rewrite interp_bv_true. simpl. f_equal; apply ih. contradict h2; right; auto. intro; discriminate. contradict h2; left; auto. Qed. Lemma nin_true_interp_bv : forall {T:Type} (x:T) (v:bv) (l:list T) (pf:length l = count_occ bool_eq_dec v false), ~ In true v -> interp_bv x v l pf = l. intros T x v l h1 h2. pose proof (bv_count_occ v) as h3. pose proof h2 as hnin. rewrite (count_occ_In bool_eq_dec) in h2. apply ngt_nlt in h2. rewrite not_lt0_iff in h2. rewrite h2 in h3. rewrite plus_0_l in h3. rewrite <- h1 in h3. revert h3 h1. revert l x. induction v as [|b v ih]. intros ? ? ? h1. symmetry in h3. apply length_eq0 in h3; subst. unfold interp_bv; simpl; auto. intros l x h1 h4. destruct b. contradict hnin; left; auto. erewrite interp_bv_false. symmetry. erewrite hd_compat' at 1. f_equal. symmetry. apply ih. rewrite count_occ_cons_neq in h2; auto. contradict hnin; right; auto. rewrite length_tl. simpl in h1. omega. Grab Existential Variables. simpl in h1. symmetry in h1. apply length_eq_S in h1; auto. Qed. Lemma remove_interp_bv : forall {T:Type} (pfdt:type_eq_dec T) (x:T) (v:bv) (l:list T) (pf:length l = count_occ bool_eq_dec v false), ~In x l -> remove pfdt x (interp_bv x v l pf) = l. intros T h1 x v. revert x. induction v as [|b v ih]; simpl. intros ? ? h2. apply length_eq0 in h2; auto. intros x l h3 h4. destruct b. rewrite interp_bv_true. rewrite remove_cons_eq. apply ih; auto. pose proof h3 as h3'. lr_if h3'; fold hlr in h3'; destruct hlr as [hl | hr]. pose proof h3' as h5. apply length_eq_S in h5. rewrite (interp_bv_false _ _ _ _ h5). rewrite remove_cons_neq. rewrite ih. symmetry. apply hd_compat'. contradict h4. rewrite (hd_compat' _ h5). right; auto. intro; subst. contradict h4. apply in_hd'. contradict hr; auto. Qed. Lemma count_occ_interp_bv : forall {T:Type} (pfdt:type_eq_dec T) (x:T) (v:bv) (l:list T) (pfe:length l = count_occ bool_eq_dec v false), ~In x l -> count_occ pfdt (interp_bv x v l pfe) x = length (bv1s v). intros T h1 x v. revert x. induction v as [|b v ih]; auto. intros x l h2. destruct b. pose proof h2 as h2'. rewrite count_occ_cons_neq in h2'. intro h3. rewrite bv1s_cons_true. rewrite interp_bv_true. rewrite count_occ_cons_eq; auto. simpl. f_equal. rewrite length_lS. rewrite ih; auto. intro; discriminate. intro h3. rewrite bv1s_cons_false. destruct l as [|a l]. pose proof h2 as h2'. rewrite count_occ_cons_eq in h2'. discriminate. reflexivity. rewrite (interp_bv_false _ _ _ _ (cons_neq_nil _ _)). simpl. destruct (h1 a x) as [|h4]; subst. contradict h3; left; auto. rewrite ih. rewrite length_lS; auto. contradict h3; right; auto. Qed. Lemma nth_lt_interp_bv : forall {T:Type} (pfdt:type_eq_dec T) (x:T) (l:list T) (v:list bool) (n:nat) (pfe:length l = count_occ bool_eq_dec v false) (pf:In x (interp_bv x v l pfe)) (pfn:n < length (interp_bv x v l pfe)) (pfinn:In n (bv1s v)), ~In x l -> nth_lt (interp_bv x v l pfe) n pfn = x. intros T h1 x l v. revert l x. induction v as [|b v ih]; simpl. intros; contradiction. intros l x n h2 h3 h4 h5 h6. pose proof h2 as h2'. lr_if h2'; fold hlr in h2'; destruct hlr as [hl | hr]; subst. rewrite bv1s_cons_false in h5. rewrite in_lS_iff in h5. destruct h5 as [m' [hi h5]]; subst. gen0. erewrite interp_bv_false. intro h7. rewrite nth_lt_cons. rewrite ih; auto. pose proof h2' as h8. apply length_eq_S in h8. pose proof (interp_bv_false x v l) as h9. specialize (h9 h2 h8). simpl in h9. pose proof h3 as h3'. rewrite h9 in h3'. destruct h3' as [h10 | h10]; subst. apply in_interp_bv. pose proof hi as h10. pose proof (in_bv1s _ _ h10) as h11. simpl in h11. rewrite <- h11. apply in_nth_lt. assumption. contradict h6. apply in_tl in h6; auto. destruct b. gen0. rewrite interp_bv_true. intro h7. rewrite bv1s_cons_true in h5. destruct h5 as [|h5]; subst. rewrite nth_lt_0_hd'; auto. rewrite in_lS_iff in h5. destruct h5 as [m [h8 h9]]; subst. rewrite nth_lt_cons. pose proof h7 as h9. simpl in h9. apply lt_S_n in h9. pose proof h8 as h10. pose proof (in_bv1s _ _ h10) as h12. simpl in h12. pose proof (in_nth_lt v m (in_bv1s_lt v m h10)) as h13. rewrite h12 in h13. apply ih; auto. apply in_interp_bv; auto. contradict hr; auto. Grab Existential Variables. apply length_eq_S in h2'; auto. Qed. Lemma in_firstn_iff : forall {T:Type} (l:list T) (x:T) (n:nat) (pfn:n <= length l), In x (firstn n l) <-> exists (m:nat) (pfmn:m < n), nth_lt l m (lt_le_trans _ _ _ pfmn pfn) = x. intros T l x n. revert l x. induction n as [|n ih]; simpl. intuition. destruct H as [? [? ?]]; omega. intros l x h1; split; intro h2. destruct l as [|a l]. contradiction. simpl in h1, h2. pose proof h1 as h1'. apply le_S_n in h1'. destruct h2 as [|h2]; subst. exists 0, (lt_0_Sn _). simpl; auto. rewrite (ih _ _ h1') in h2. destruct h2 as [m [h2 h3]]. exists (S m). exists (lt_n_S _ _ h2). simpl. subst. rewrite nth_lt_compat. apply nth_indep. eapply lt_le_trans. apply h2. assumption. destruct l as [|a l]. simpl in h1; omega. simpl in h1. pose proof h1 as h1'. apply le_S_n in h1'. destruct h2 as [m [h2 h3]]. simpl in h3. destruct m as [|m]; subst. left; auto. simpl. apply lt_S_n in h2. specialize (ih l (nth m l a) h1'). hr ih. exists m, h2. rewrite nth_lt_compat; apply nth_indep. eapply lt_le_trans. apply h2. assumption. rewrite <- ih in hr. right; auto. Qed. Lemma firstn_S_eq : forall {T:Type} (l:list T) (m:nat) (pflt:m < length l), firstn (S m) l = firstn m l ++ ((nth_lt l m pflt)::nil). intros T l. induction l as [| a l h1]; simpl. intros; omega. intros m h2. destruct m as [|m]; simpl; auto. f_equal. apply lt_S_n in h2. specialize (h1 _ h2). rewrite nth_lt_compat in h1. erewrite nth_indep. rewrite <- h1 at 1. reflexivity. assumption. Qed. Lemma skipn_S_eq : forall {T:Type} (l:list T) (m:nat), skipn (S m) l = tl (skipn m l). intros ? l; simpl. induction l; simpl. intros; rewrite skipn_nil; auto. intro m; destruct m; simpl; auto. Qed. Lemma tl_skipn : forall {T:Type} (l:list T) (m:nat), tl (skipn m l) = skipn m (tl l). intros T l m. revert l. induction m as [|m ih]; simpl; auto. intro l. destruct l as [|a l]; simpl; auto. rewrite <- skipn_S_eq. simpl; auto. Qed. Lemma length_firstn_S_eq : forall {T:Type} (l:list T) (m:nat) (pflt:m < length l), length (firstn (S m) l) = S (length (firstn m l)). intros T l m h1. rewrite (firstn_S_eq _ _ h1). rewrite app_length; simpl. rewrite S_compat; auto. Qed. Lemma incl_firstn : forall {T:Type} (l:list T) (n:nat), n <= length l -> incl (firstn n l) l. intros; red; intros x h1. rewrite in_firstn_iff in h1. destruct h1 as [h1 [h2 h3]]. rewrite <- h3. apply in_nth_lt. Grab Existential Variables. assumption. Qed. Lemma incl_firstn_le : forall {T:Type} (l:list T) (m n:nat), m <= n -> incl (firstn m l) (firstn n l). intros T l m n h0; red; intros ? h1. destruct (le_lt_dec n (length l)) as [h2 | h2]; subst. rewrite (in_firstn_iff _ _ _ (le_trans _ _ _ h0 h2)) in h1. rewrite (in_firstn_iff _ _ _ h2). destruct h1 as [r h3]. destruct h3 as [h3 h4]; subst. exists r. ex_goal. eapply lt_le_trans. apply h3. assumption. exists hex. apply nth_lt_functional3. rewrite firstn_length_ge. apply in_firstn in h1; auto with arith. auto with arith. Qed. Lemma skipn_sing : forall {T:Type} (l:list T) (x:T) (n:nat), x::nil = skipn n l <-> exists pf:length l = S n, let pf' := lt_eq_trans _ _ _ (lt_n_Sn _) (eq_sym pf) in nth_lt l n pf' = x. intros T l. induction l as [|a l ih]; simpl. intros ? ?. split; intro h1. rewrite skipn_nil in h1. discriminate. destruct h1; discriminate. intros x n. split; intro h1. destruct n as [|n]. simpl in h1. inversion h1; subst. ex_goal; auto. exists hex; auto. simpl in h1. apply ih in h1; auto. destruct h1 as [h1 h2]. simpl in h2. ex_goal. f_equal; auto. exists hex. rewrite nth_lt_compat in h2. erewrite nth_indep. rewrite h2 at 1; auto. rewrite h1. apply lt_n_Sn. destruct h1 as [h1 h2]. apply S_inj in h1. destruct n as [|n]; simpl. subst; auto. apply length_eq0 in h1. subst; auto. subst. rewrite ih. exists h1. simpl. rewrite nth_lt_compat. apply nth_indep. rewrite h1; auto with arith. Qed. Lemma firstn_map : forall {T U:Type} (f:T->U) (l:list T) (n:nat), n <= length l -> firstn n (map f l) = map f (firstn n l). intros T U f l n. revert f l. induction n as [|n ih]; auto. intros f l h0. erewrite firstn_S_eq. erewrite firstn_S_eq. rewrite map_app. rewrite ih. simpl. rewrite nth_lt_map'. reflexivity. auto with arith. Grab Existential Variables. rewrite map_length; auto with arith. Qed. Lemma firstn_map_seg : forall {U:Type} {n:nat} (f:seg_fun n U) (m:nat) (pfm:m <= n), firstn m (map_seg f) = map_seg (fun i pf => f i (lt_le_trans _ _ _ pf pfm)). intros U n. induction n as [|n ih]; simpl. intros ? ? h1. pose proof h1 as h1'. apply le0 in h1'. subst. simpl; auto. intros f m h1. pose proof h1 as h1'. apply le_lt_eq_dec in h1'. destruct h1' as [h2|]; subst. apply lt_n_Sm_le in h2. specialize (ih (seg_fun_from_S f) m h2). rewrite firstn_app_le. rewrite ih at 1. f_equal. apply functional_extensionality_dep. intro x. apply functional_extensionality. intro y. unfold seg_fun_from_S. f_equal. apply proof_irrelevance. rewrite length_map_seg; auto with arith. rewrite firstn_length_ge. simpl. f_equal. f_equal. apply functional_extensionality_dep. intro x. unfold seg_fun_from_S. apply functional_extensionality. intro; f_equal; apply proof_irrelevance. repeat f_equal. apply proof_irrelevance. rewrite app_length. rewrite length_map_seg. simpl; rewrite S_compat; auto with arith. Qed. Lemma firstn_eq : forall {T:Type} (l:list T) (n:nat) (pflt:pred n < (length l)) (pfnd:NoDup l), 0 < n -> firstn n l = firstn (pred n) l ++ ((nth_lt _ _ pflt) :: nil). intros T l n. revert l. induction n as [|n h1]; simpl; auto. intros; omega. intros l h2 h3 h4. destruct (zerop n) as [h5 | h5]; subst. simpl. destruct l. simpl in h2. omega. pose proof (nth_lt_compat (t :: l) 0 h2) as h5. simpl in h5. simpl. reflexivity. destruct l as [|a l]. destruct n. simpl. pose proof (nth_lt_compat nil 0 h2) as h6. simpl in h6. omega. simpl in h2. omega. simpl in h2. assert (h6:pred n < length l). omega. pose proof (no_dup_cons _ _ h3) as h7. specialize (h1 l h6 h7 h5). rewrite h1. rewrite firstn_cons. simpl. f_equal. f_equal. assert (h8:S (pred n) < S (length l)). omega. pose proof (nth_lt_cons l a (pred n) h8) as h9. simpl in h9. assert (h10: n = S (pred n)). omega. pose proof (subsetT_eq_compat _ (fun m => m < (S (length l))) _ _ h2 h8 h10) as h11. assert (h12:nth_lt (a::l) (S (pred n)) h8 = nth_lt (a::l) n h2). f_equal. dependent rewrite <-h11. auto. simpl in h12. rewrite <- h12. rewrite h9. f_equal. f_equal. apply proof_irrelevance. assumption. Qed. Lemma firstn_eq_firstn_cons_iff : forall {T:Type} (m:nat) (l:list T) (a:T), m <= length l -> (firstn m l = firstn m (a::l) <-> all_sing (firstn m l) a). intros T m. induction m as [|m ih]; simpl. intros; split; intro; auto. apply all_sing_nil. intros l a hl; split; intro h1. destruct l as [|b l]. discriminate. inversion h1 as [h2]; subst. rewrite all_sing_cons_iff. split; auto. simpl in hl. rewrite <- ih; auto with arith. destruct l as [|b l]. simpl in hl; omega. simpl in hl. apply le_S_n in hl; auto. rewrite all_sing_cons_iff in h1. destruct h1 as [? h1]; subst. f_equal. rewrite ih; auto. Qed. Lemma length_firstn : forall {T:Type} (l:list T) (m:nat), m <= length l -> length (firstn m l) = m. intros T l m. revert l. induction m as [|m ih]; simpl; auto. intros l h1. destruct l as [|a l]; simpl. simpl in h1; omega. simpl in h1. apply le_S_n in h1. f_equal. apply ih; auto. Qed. Lemma length_firstn_S : forall {T:Type} (l:list T) (m:nat), m < length l -> length (firstn (S m) l) = S m. intros T l m. revert l. induction m as [|m ih]; simpl; auto. intros l h1. destruct l as [|a l]; simpl. simpl in h1; omega. reflexivity. intros l h1. destruct l as [|a l]. simpl in h1; omega. simpl in h1. apply lt_S_n in h1. simpl. f_equal; apply ih; auto. Qed. Lemma length_firstn_le : forall {T:Type} (l:list T) (m:nat), length (firstn m l) <= length l. intros T l. induction l as [|a l h1]; simpl. intro m. destruct m; simpl; auto with arith. intro m. destruct m; simpl; auto with arith. Qed. Lemma length_eq_firstn : forall {T U:Type} (l:list T) (l':list U) (m:nat), length l = length l' -> length (firstn m l) = length (firstn m l'). intros T U l l' m h1. destruct (le_dec m (length l)) as [h2 | h2]. rewrite length_firstn; auto. rewrite length_firstn; auto. eapply le_eq_trans. apply h2. assumption. erewrite firstn_length_ge. erewrite firstn_length_ge; auto. omega. omega. Qed. Lemma eq_firstn_iff : forall {T} (l l' l0:list T), length l = length l' -> (l = firstn (length l) (l' ++ l0) <-> l = l'). intros T l; induction l as [|a l ih]; simpl. intros l l' h1. symmetry in h1. apply length_eq0 in h1; subst. tauto. intro l'. destruct l' as [|a' l']. simpl. intro; discriminate. intro l0; destruct l0 as [|a0 l0]; simpl. intro h1. apply S_inj in h1. pose proof h1 as h2. eapply ih in h2. split; intro h3. inversion h3. rewrite h2 in H1; auto; subst; auto. inversion h3; subst. f_equal. rewrite h2; auto. intro h1. apply S_inj in h1. split; intro h2. inversion h2. revert h1 h2; revert ih H1; revert l; subst. intros l h2 ih h3 h4. f_equal. pose proof ih as h5. rewrite h2 in ih; auto; subst. symmetry; rewrite h2; auto. inversion h2; auto; subst. f_equal. rewrite ih; auto. Qed. Lemma lt_length_firstn : forall {T:Type} (l:list T) (m n:nat), m < n -> m < length l -> m < length (firstn n l). intros T l m n. revert m l. induction n as [|n ih]; simpl; auto. intros m ? ? h2. destruct l; auto. simpl in h2; simpl. destruct m; auto with arith. Qed. Lemma firstn_idem : forall {T:Type} (l:list T) (m:nat), firstn m (firstn m l) = firstn m l. intros T l m. destruct (le_dec m (length l)) as [h1 | h1]. apply firstn_length_ge. rewrite length_firstn; auto with arith. apply not_le in h1. rewrite firstn_length_ge; auto with arith. rewrite firstn_length_ge; auto with arith. Qed. Lemma firstn_le_nest : forall {T:Type} (l:list T) (m n:nat), m <= n -> firstn m l = firstn m (firstn n l). intros T l m n h1. apply le_lt_eq_dec in h1. destruct h1; subst. apply firstn_lt_nest; auto. symmetry. apply firstn_idem. Qed. Lemma length_skipn : forall {T:Type} (l:list T) (m:nat), length (skipn m l) = length l - m. intros ? l m; revert l. induction m as [|m ih]; simpl; auto with arith. intro l; destruct l; simpl; auto. Qed. Lemma skipn_firstn_nil : forall {T:Type} (l:list T) (n:nat), skipn n (firstn n l) = nil. intros T l n. revert l. induction n; simpl; auto. intro l; destruct l as [|a l]; simpl; auto. Qed. Lemma eq_firstn_iff' : forall {T} (l l' l0:list T), l = firstn (length l) (l' ++ l0) <-> (length l < length l' -> (l = firstn (length l) l' /\ skipn (length l) l' <> nil)) /\ (length l = length l' -> l = l') /\ (length l > length l' -> l = l' ++ firstn (length l - length l') l0). intros T l l' l0; split; intros h1; try split; try intro h2. pose proof (firstn_length_eq _ _ (eq_sym h1)) as h3. split. rewrite firstn_app_le in h1; auto with arith. intro h4. rewrite skipn_eq_nil_iff in h4. destruct h4; revert h1; subst. simpl in h2; contradict (lt_n_O _ h2). intro h4. omega. split. intro h2. apply eq_firstn_iff in h1; auto. intros h2; revert h1; subst; intro h3. rewrite firstn_app_ge in h3; auto with arith. destruct h1 as [h1 [h2 h3]]. destruct (lt_eq_lt_dec (length l) (length l')) as [[h4|h4]|h4]. specialize (h1 h4). destruct h1 as [h1 h5]. rewrite firstn_app_le; auto with arith. apply h2 in h4; subst. rewrite firstn_app_eq; auto. specialize (h3 h4). rewrite firstn_app_ge; auto with arith. Qed. Lemma in_nth_lt_firstn : forall {T:Type} (l:list T) (m n:nat) (pfn:n <= length l) (pfm:m < n), In (nth_lt (firstn n l) m (lt_eq_trans _ _ _ pfm (eq_sym (length_firstn _ _ pfn)))) (firstn n l). intros; apply in_nth_lt. Qed. Lemma nth_lt_firstn_eq : forall {T:Type} (l:list T) (m n:nat) (pf:m < length (firstn n l)), nth_lt (firstn n l) m pf = nth_lt l m ( lt_le_trans _ _ _ pf (length_firstn_le l n)). intros T l m n h1. simpl. pose proof (firstn_skipn n l) as h2. symmetry. gen0. generalize h1. rewrite <- h2; auto. intros h3 h4. rewrite (nth_lt_app1 _ _ _ h1). generalize h3. rewrite h2; intros h5. f_equal. apply proof_irrelevance. Qed. Lemma nth_lt_skipn_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m n:nat) (pfm:m < length (skipn n l)) (pfmn:m + n < length l), nth_lt (skipn n l) m pfm = nth_lt l (m + n) pfmn. intros T hdt l m n. revert l m. induction n as [|n ih]. simpl; intros; apply nth_lt_functional2; auto with arith. intros l m h1 h2. gen0. rewrite skipn_S_eq. intro h3. symmetry. gen0. rewrite plus_S_shift_r. intro. rewrite nth_lt_S. erewrite <- ih. apply nth_lt_functional1. rewrite tl_skipn; auto. Grab Existential Variables. rewrite length_skipn. rewrite length_tl in h3. rewrite length_skipn in h3. rewrite length_tl; auto. rewrite length_skipn in h1. rewrite pred_minus in h3; auto with arith. omega. Qed. Lemma all_sing_app_eq : forall {T} (l0 l0' l1 l1':list T) a b, a <> b -> all_sing l0 a -> all_sing l0' a -> l0 ++ b :: l1 = l0' ++ b :: l1' -> l0 = l0' /\ l1 = l1'. intros T l0 l0' l1 l1' a b h1 h2 h3 h4. red in h2, h3. apply neq_sym in h1. apply app_nin_eq in h4; auto. Qed. Lemma in_nth_lt_firstn' : forall {T:Type} (l:list T) (m n:nat) (pfn:n <= length l) (pfm:m < n), In (nth_lt l m (lt_le_trans _ _ _ pfm pfn)) (firstn n l). intros. erewrite nth_lt_functional3. rewrite <- nth_lt_firstn_eq at 1. apply in_nth_lt_firstn. Grab Existential Variables. assumption. assumption. Qed. Lemma nth_lt_removelast : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pf':n < length (removelast l)) (pf:n < length l), nth_lt (removelast l) n pf' = nth_lt l n pf. intros; gen0; rewrite removelast_eq; intro; rewrite nth_lt_firstn_eq. apply nth_lt_functional3. Qed. Lemma lst_removelast : forall {T:Type} {hdt:type_eq_dec T} (l:list T) (pf:removelast l <> nil), let pflt := removelast_not_nil' _ pf in lst (removelast l) pf = nth_lt l (pred (pred (length l))) pflt. intros; erewrite lst_eq, nth_lt_removelast; auto. gen0. rewrite length_removelast; intros; apply nth_lt_functional3; apply proof_irrelevance. Grab Existential Variables. rewrite length_removelast; auto. Qed. Lemma firstn_map_in_dep_firstn : forall {T U:Type} {l:list T} (f:fun_in_dep l U) (m n:nat) (pfm:m <= n) (pfn:n <= length l), firstn m (map_in_dep (fun x (pf:In x (firstn n l)) => f x (incl_firstn _ _ pfn _ pf))) = firstn m (map_in_dep f). intros T U l f m. revert f. revert l. induction m as [|m ih]; auto. intros l f n h1 h2. erewrite firstn_S_eq. erewrite firstn_S_eq. rewrite nth_lt_map_in_dep'. rewrite nth_lt_map_in_dep'. f_equal. apply ih. auto with arith. f_equal. gen0. intro h3. apply f_equal_dep. rewrite nth_lt_firstn_eq. apply nth_lt_functional3. Grab Existential Variables. rewrite length_map_in_dep. omega. rewrite length_map_in_dep, length_firstn; auto. Qed. Lemma in_firstn_S_dec : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pf:S m <= length l) (x:T), In x (firstn (S m) l) -> {In x (firstn m l)} + {x = nth_lt l m pf}. intros T hdt l m h1 x h2. destruct (hdt x (nth_lt l m h1)); subst. right; auto. left. rewrite (in_firstn_iff _ _ _ h1) in h2. destruct h2 as [r [h2 h3]]. rewrite <- h3. rewrite in_firstn_iff. exists r. ex_goal. red in h2. pose proof h2 as h2'. apply le_S_n in h2'. apply le_lt_eq_dec in h2'. destruct h2'; subst; auto. contradict n. apply nth_lt_functional3. exists hex. apply nth_lt_functional3. Grab Existential Variables. auto with arith. Qed. Lemma impl_lt_length_firstn : forall {T:Type} (l:list T) (m n:nat), n <= length l -> m < n -> m < (length (firstn n l)). intros; rewrite length_firstn; auto. Qed. Lemma lt_length_firstn_impl : forall {T:Type} (l:list T) (m n:nat), n <= length l -> m < (length (firstn n l)) -> m < n. intros ? ? ? ? ? h1; rewrite length_firstn in h1; auto. Qed. Lemma nth_firstn : forall {T:Type} (l:list T) (m n:nat) (def:T), m < n -> nth m (firstn n l) def = nth m l def. intros T l. induction l as [|a l h1]; simpl. intros m n def h3. destruct n as [|n]. omega. rewrite firstn_nil; auto. intros m n def h2. destruct n as [|n]; simpl. omega. red in h2. apply le_S_n in h2. apply le_lt_eq_dec in h2. destruct m as [|m]; auto. destruct h2 as [h2 | h2]. assert (h3:m < n). omega. specialize (h1 _ _ def h3); auto. subst. assert (h2: m < S m). auto with arith. specialize (h1 _ _ def h2). assumption. Qed. Lemma nth_firstn' : forall {T:Type} (l:list T) (m n:nat) (d1 d2:T), m < length l -> m < n -> nth m (firstn n l) d1 = nth m l d2. intros T l m n d1 d2 h0 h1. rewrite nth_firstn; auto. apply nth_indep; auto. Qed. Lemma nth_lt_map_seq : forall {U:Type} {n w:nat} (f:seq_fun n w U) (m:nat) (pf:m < w), let pf' := lt_eq_trans _ _ _ pf (eq_sym (length_map_seq f)) in let pf'' := lt_seq _ _ _ pf in nth_lt (map_seq f) m pf' = f _ pf''. intros U n w. revert n. induction w as [|w ih]; simpl. intros. omega. intros n f m h1. simpl in ih. red in h1. pose proof h1 as h1'. apply le_S_n in h1'. destruct (le_lt_eq_dec _ _ h1') as [h2 | h2]. specialize (ih n (seq_fun_from_S f) m h2). erewrite nth_lt_app1. unfold seq_fun_from_S in ih. rewrite ih at 1. f_equal. apply proof_irrelevance. subst. erewrite nth_lt_app2. simpl. rewrite length_map_seq. rewrite minus_same. f_equal. apply proof_irrelevance. Grab Existential Variables. rewrite length_map_seq; auto with arith. Qed. Lemma nth_lt_lrep : forall {T:Type} (a:T) (m n:nat) (pf:m < length (lrep a n)), nth_lt (lrep a n) m pf = a. intros T a m n. revert m a. induction n as [|n ih]; simpl. intros; omega. intros m a h1. destruct m as [|m]; auto. apply lt_S_n in h1. specialize (ih _ _ h1). rewrite <- ih at 3. rewrite nth_lt_compat at 1. apply nth_indep; auto. Qed. Lemma count_occ_lrep : forall {T:Type} (pfdt:type_eq_dec T) (a:T) (n:nat), count_occ pfdt (lrep a n) a = n. intros T h1 a n. revert a. induction n as [|n ih]; simpl; auto. intro a. erewrite eq_dec_same at 1. f_equal. apply ih. assumption. Qed. Lemma count_occ_lrep' : forall {T:Type} (pfdt:type_eq_dec T) (a x:T) (n:nat), In x (lrep a n) -> count_occ pfdt (lrep a n) x = n. intros T h1 a x n h2. apply in_lrep' in h2; subst. apply count_occ_lrep. Qed. Lemma sub_list_firstn_removelast_le : forall {T:Type} (l:list T) (n:nat), sub_list (firstn n (removelast l)) (firstn n l). intros T l n. destruct (nil_dec' l) as [|h1]; subst. simpl. apply sub_list_refl. pose proof (app_removelast_last (hd' _ h1) h1) as h2. rewrite h2 at 2. destruct (le_dec n (length (removelast l))) as [h3 | h3]. rewrite firstn_app_le; auto with arith. apply sub_list_refl. apply not_le in h3. rewrite firstn_length_ge; auto with arith. rewrite firstn_length_ge; auto with arith. rewrite <- (app_nil_l (removelast l ++ last l (hd' l h1) :: nil)). constructor. rewrite app_length. simpl. omega. Qed. Lemma count_occ_nth_lt0 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pf:0 < length l), count_occ pfdt l (nth_lt l 0 pf) = S (count_occ pfdt (tl l) (nth_lt l 0 pf)). intros T h1 l. destruct l as [|a l]; simpl. intro h2; apply lt_irrefl in h2; contradiction. intro h2. destruct (h1 a a) as [|h4]; auto. contradict h4; auto. Qed. Lemma length_list_singularize_map_inj : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (f:T->U), Injective f -> length (list_singularize pfdu (map f l)) = length (list_singularize pfdt l). intros T U h1 h2 l. induction l as [|a l ih]; simpl; auto. intros f h3. destruct (in_dec h2 (f a) (map f l)) as [hl | hr]. rewrite list_singularize_in_cons; auto. destruct (in_dec h1 a l) as [hl' | hr']. rewrite list_singularize_in_cons; auto. rewrite in_map_iff in hl. destruct hl as [a' h4]. destruct h4 as [h4 h5]. apply h3 in h4; subst; contradiction. rewrite list_singularize_nin; auto. destruct (in_dec h1 a l) as [hl' | hr']. rewrite list_singularize_in_cons; auto. contradict hr. rewrite in_map_iff. exists a; tauto. rewrite list_singularize_nin; auto. simpl; auto. Qed. Lemma length_list_singularize_no_dup : forall {T} (pfdt:type_eq_dec T) l, NoDup l -> length (list_singularize pfdt l) = length l. intros T hdt l h1. rewrite list_singularize_no_dup in h1. rewrite h1; auto. Qed. Lemma nth_lt_list_singularize_no_dup : forall {T} (pfdt:type_eq_dec T) (l:list T) (pfndp:NoDup l) i (pfi:i < length (list_singularize pfdt l)), let pfi' := lt_eq_trans _ _ _ pfi (length_list_singularize_no_dup pfdt l pfndp) in nth_lt (list_singularize pfdt l) i pfi = nth_lt l i pfi'. intros; apply nth_lt_functional1; apply list_singularize_no_dup; auto. Qed. Lemma nth_lt_list_singularize_map : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (f:T->U) (i:nat) (pfinj:Injective f) (pflt:i < length (list_singularize pfdt l)), let pflt' := lt_eq_trans _ _ _ pflt (eq_sym (length_list_singularize_map_inj pfdt pfdu l f pfinj)) in nth_lt (list_singularize pfdu (map f l)) i pflt' = f (nth_lt (list_singularize pfdt l) i pflt). intros T U h1 h2 l. induction l as [|a l ih]; simpl. intros; omega. intros f i h3 h4. destruct (in_dec h2 (f a) (map f l)) as [hl | hr]. pose proof hl as hl'. rewrite in_map_iff in hl'. destruct hl' as [a' h5]. destruct h5 as [h5 h5']. apply h3 in h5; subst. gen0. generalize h4. rewrite list_singularize_in_cons. rewrite list_singularize_in_cons; auto. intros h5 h6. specialize (ih f i h3 h5). simpl in ih. rewrite <- ih. apply nth_lt_functional3. assumption. gen0. rewrite list_singularize_nin; auto. intro h5. assert (h6:~In a l). contradict hr. rewrite in_map_iff. exists a; tauto. generalize h4. rewrite list_singularize_nin; auto. intro h7. simpl. destruct i as [|i]; auto. erewrite nth_indep. erewrite <- nth_lt_compat. erewrite nth_indep. erewrite <- nth_lt_compat. apply ih. simpl in h7; auto with arith. simpl in h7; auto with arith. Grab Existential Variables. assumption. simpl in h7; auto with arith. Qed. Lemma nth_lt_list_singularize_map_in_dep : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:fun_in_dep l U) (i:nat) (pfinj:inj_dep f) (pflt:i < length (list_singularize pfdt l)), let pflt' := lt_eq_trans _ _ _ pflt (eq_sym (length_list_singularize_map_in_dep_inj pfdt pfdu f pfinj)) in nth_lt (list_singularize pfdu (map_in_dep f)) i pflt' = f _ ((iff2 (list_singularize_in_iff pfdt l (nth_lt (list_singularize pfdt l) i pflt))) (in_nth_lt (list_singularize pfdt l) i pflt)). intros T U l h1 h2. induction l as [|a l ih]; simpl. intros; omega. intros f i h3 h4. destruct (in_dec h2 (f a (in_eq _ _)) (map_in_dep (fun_from_cons f))) as [hl | hr]. pose proof hl as hl'. rewrite in_map_in_dep_iff in hl'. destruct hl' as [a' h5]. destruct h5 as [h5 h5']. pose proof h3 as h3'. red in h3'. unfold fun_from_cons in h5'. apply h3' in h5'; subst. gen0. generalize h4. rewrite (list_singularize_in_cons h2 _ _ hl); auto. intros h5' h6. specialize (ih (fun_from_cons f) i). hfirst ih. apply inj_from_cons; auto. specialize (ih hf). hfirst ih. rewrite list_singularize_in_cons in h4; auto. specialize (ih hf0). simpl in ih. erewrite nth_lt_functional3. rewrite ih at 1. unfold fun_from_cons; apply f_equal_dep. symmetry. gen0. rewrite list_singularize_in_cons. intro h7. apply nth_lt_functional3; auto. assumption. gen0. rewrite list_singularize_nin; auto. intro h5. assert (h6:~In a l). contradict hr. rewrite in_map_in_dep_iff. exists a, hr. unfold fun_from_cons; f_equal; apply proof_irrelevance. symmetry. gen0. generalize h4. rewrite (list_singularize_nin h1 l a); auto. intros h7 h10. simpl. destruct i as [|i]; auto. f_equal; apply proof_irrelevance. symmetry. specialize (ih (fun_from_cons f) i). hfirst ih. apply inj_from_cons; auto. specialize (ih hf). hfirst ih. rewrite (list_singularize_nin h1 _ _ h6) in h4. simpl in h4; auto with arith. specialize (ih hf0). simpl in ih. rewrite nth_lt_compat in ih. erewrite nth_indep. rewrite ih at 1. unfold fun_from_cons. apply f_equal_dep. erewrite nth_lt_compat. apply nth_indep. assumption. simpl in h5; auto with arith. Qed. Lemma list_singularize_map_inj : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (f:T->U), Injective f -> list_singularize pfdu (map f l) = map f (list_singularize pfdt l). intros T U h1 h2 l f h3. symmetry. rewrite map_eq_iff. exists (eq_sym (length_list_singularize_map_inj _ _ _ _ h3)). intros n h4. symmetry; apply nth_lt_list_singularize_map; auto. Qed. Lemma list_singularize_map_in_dep_inj : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:fun_in_dep l U), inj_dep f -> list_singularize pfdu (map_in_dep f) = map_in_dep (fun x (pf:In x (list_singularize pfdt l)) => f x (iff2 (list_singularize_in_iff pfdt l x) pf)). intros T U l h1 h2 f h3. symmetry. rewrite map_in_dep_eq_iff. exists (eq_sym (length_list_singularize_map_in_dep_inj _ _ _ h3)). intros n h4. symmetry; apply nth_lt_list_singularize_map_in_dep; auto. Qed. Lemma list_singularize_firstn_S : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pfsm: S m <= length l), list_singularize pfdt (firstn (S m) l) = remove pfdt (nth_lt l m pfsm) (list_singularize pfdt (firstn m l)) ++ nth_lt l m pfsm :: nil. intros T h1 l m h2. erewrite firstn_S_eq. apply list_singularize_rev_cons. Qed. Lemma incl_list_singularize_firstn_S_app_sing : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (n:nat), S n <= length (l++a::nil) -> incl (list_singularize pfdt (firstn n l)) (list_singularize pfdt (firstn (S n) (l++a::nil))). intros T hdt l a n hn; apply incl_list_singularize_compat; red; intros x h1. rewrite app_length in hn; rewrite S_compat in hn. apply le_S_n in hn. rewrite (in_firstn_iff _ _ _ hn) in h1. destruct h1 as [m [h1 h2]]; subst. erewrite in_firstn_iff. exists m. exists (lt_trans _ _ _ h1 (lt_n_Sn _)). erewrite nth_lt_app1. apply nth_lt_functional3. Grab Existential Variables. eapply lt_le_trans. apply h1. assumption. rewrite app_length; rewrite S_compat; auto with arith. Qed. Lemma count_occ_map_in_dep : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (f:fun_in_dep l U) (y:U), count_occ pfdu (map_in_dep f) y = length (inv_im_in_dep pfdu l y f). intros T U h1 h2 l. induction l as [|a l ih]; simpl; auto. intros f y. rewrite switch_eq_dec_eq at 1. destruct (h2 y (f a (in_eq a l))) as [h4 | h4]; subst. unfold switch_eq_dec; simpl. f_equal; apply ih; auto. unfold switch_eq_dec. apply ih; apply inj_from_cons; auto. Qed. Lemma count_occ_map_in_dep_inj : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:fun_in_dep l U) (x:T) (pfx:In x l), inj_dep f -> count_occ pfdu (map_in_dep f) (f x pfx) = count_occ pfdt l x. intros T U l h1 h2. induction l as [|a l ih]; simpl. intros; contradiction. intros f x h3 h4. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. apply h4 in hl; subst. erewrite eq_dec_same at 1. f_equal. destruct (in_dec h1 x l) as [h5 | h5]. specialize (ih (fun_from_cons f) x h5). hfirst ih. apply inj_from_cons; auto. specialize (ih hf). unfold fun_from_cons in ih at 2. pose proof (proof_irrelevance _ h3 (or_intror h5)); subst. rewrite <- ih; repeat f_equal; apply proof_irrelevance. pose proof (count_occ_In h1 l x) as h6. rewrite h6 in h5. pose proof (count_occ_In h2 (map_in_dep (fun_from_cons f)) (f x h3)) as h7. apply ngt_nlt in h5. rewrite not_lt0_iff in h5. rewrite h5. rewrite <- not_lt0_iff. apply ngt_nlt. rewrite <- (count_occ_In h2). intro h8. rewrite in_map_in_dep_iff in h8. destruct h8 as [x' [h8 h9]]. unfold fun_from_cons in h9. apply h4 in h9; subst. rewrite h6 in h8. omega. assumption. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']; subst. contradict hr. f_equal; apply proof_irrelevance. destruct h3 as [|h3]; try contradiction. specialize (ih (fun_from_cons f) x h3). hfirst ih. apply inj_from_cons; auto. specialize (ih hf). unfold fun_from_cons in ih at 2; auto. rewrite <- ih. repeat f_equal; apply proof_irrelevance. Qed. Lemma map_seg_eq_in_dep : forall {U:Type} {n:nat} (f:seg_fun n U), map_seg f = map_in_dep (seg_fun_in_dep f). intros U n f. symmetry. rewrite map_in_dep_eq_iff. ex_goal. rewrite length_seg_list, length_map_seg; auto. exists hex. intros m h1. unfold seg_fun_in_dep. gen0. rewrite nth_lt_seg_list. intro h2. erewrite nth_lt_functional3. rewrite nth_lt_map_seg. reflexivity. Qed. Lemma count_occs_eq_list_to_set : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), (forall x:T, count_occ pfdt l x = count_occ pfdt l' x) -> list_to_set l = list_to_set l'. intros T h1 l l' h2. apply Extensionality_Ensembles; red; split; red; intros x h3; rewrite <- list_to_set_in_iff in h3; rewrite <- list_to_set_in_iff. specialize (h2 x). apply (count_occ_In h1) in h3. rewrite h2 in h3. apply (count_occ_In h1); auto. specialize (h2 x). apply (count_occ_In h1) in h3. rewrite <- h2 in h3. apply (count_occ_In h1); auto. Qed. Lemma count_occs_eq_impl : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), (forall x:T, count_occ pfdt l x = count_occ pfdt l' x) -> length (list_singularize pfdt l) = length (list_singularize pfdt l'). intros T h1 l l' h2. apply count_occs_eq_list_to_set in h2. apply list_to_sets_eq_length; auto. Qed. Lemma map_seg_eq_iff : forall {U:Type} (n:nat) (f:seg_fun n U) (l:list U), map_seg f = l <-> (exists pfeq:n = length l, forall (m:nat) (pflt:m < n), let pflt' := lt_eq_trans _ _ _ pflt pfeq in f m pflt = nth_lt l m pflt'). intros U n. induction n as [|n ih]; simpl. intros. split; intro h1. subst. simpl. exists eq_refl. intros; omega. destruct h1 as [h1 h2]. symmetry. apply length_eq0; auto. intros f l. rewrite app_eq_iff. rewrite length_map_seg. split; intro h1. destruct h1 as [h1 h2]. apply skipn_sing in h2. destruct h2 as [h2 h2']. simpl in h2'. exists (eq_sym h2). intros m h4. pose proof h4 as h4'. apply le_S_n in h4'. apply le_lt_eq_dec in h4'. destruct h4' as [h5 | h5]. rewrite ih in h1. destruct h1 as [h1 h6]. specialize (h6 _ h5). simpl in h6. unfold seg_fun_from_S in h6. rewrite nth_lt_firstn_eq in h6. erewrite nth_lt_functional3. rewrite <- h6 at 1. f_equal. apply proof_irrelevance. subst. erewrite nth_lt_functional3. rewrite h2' at 1. f_equal. apply proof_irrelevance. destruct h1 as [h1 h2]. split. rewrite ih. ex_goal. rewrite length_firstn; auto. rewrite <- h1; auto with arith. exists hex. intros m h3; simpl. unfold seg_fun_from_S. rewrite h2. rewrite nth_lt_firstn_eq. apply nth_lt_functional3. rewrite skipn_sing. exists (eq_sym h1). simpl. rewrite h2. apply nth_lt_functional3. Qed. Lemma map_segs_eq_iff : forall {U:Type} (n:nat) (f g:seg_fun n U), map_seg f = map_seg g <-> forall i (pflt:i < n), f i pflt = g i pflt. intros U n f g. rewrite map_seg_eq_iff. split; intro h1. destruct h1 as [h1 h2]. simpl in h2. intros i h3. specialize (h2 i h3). rewrite h2. erewrite nth_lt_functional3. rewrite nth_lt_map_seg. reflexivity. ex_goal. rewrite length_map_seg; auto. exists hex. intros i h2; simpl. specialize (h1 _ h2). rewrite h1. erewrite nth_lt_functional3. rewrite nth_lt_map_seg. reflexivity. Qed. Lemma map_seg_id_seg : forall n, map_seg (id_seg n) = seg_list n. intro n. unfold id_seg. rewrite map_seg_eq_iff. ex_goal. rewrite length_seg_list; auto. exists hex. intros m h1 h2. rewrite nth_lt_seg_list; auto. Qed. Lemma map_seg_nth_lt : forall {T:Type} (l:list T), map_seg (nth_lt l) = l. intros; rewrite map_seg_eq_iff. exists eq_refl. intros; f_equal; apply proof_irrelevance. Qed. Lemma removelast_map_seg : forall {U:Type} {n:nat} (pfdu:type_eq_dec U) (f:seg_fun n U) (pf:0 < n), removelast (map_seg f) = map_seg (seg_fun_rest f (lt_pred_n _ _ pf)). intros U hdu n f h1. symmetry. rewrite map_seg_eq_iff. ex_goal. rewrite length_removelast, length_map_seg; auto. exists hex. intros m h2 h3. erewrite nth_lt_removelast. rewrite seg_fun_rest_compat, nth_lt_map_seg. reflexivity. assumption. Qed. Lemma map_seg_last : forall {U:Type} {n:nat} (f:seg_fun (S n) U), map_seg f = map_seg (seg_fun_from_S f) ++ (f n (lt_n_Sn _)) ::nil. intros U n f. rewrite map_seg_eq_iff. ex_goal. rewrite app_length; simpl. rewrite S_compat. f_equal. rewrite length_map_seg; auto. exists hex. intros m h1 h2. unfold seg_fun_from_S. erewrite nth_lt_functional3. pose proof h1 as h1'. apply lt_S_dec in h1'. destruct h1' as [h3 | h3]. erewrite nth_lt_app1 at 1. erewrite nth_lt_map_seg. f_equal; apply proof_irrelevance. subst. erewrite nth_lt_app2 at 1. symmetry. gen0. rewrite length_map_seg. intro h3. simpl. rewrite minus_same. f_equal; apply proof_irrelevance. Grab Existential Variables. rewrite length_map_seg; auto with arith. rewrite app_length; simpl. rewrite S_compat. rewrite length_map_seg; auto with arith. assumption. Qed. Lemma map_seq_fun_to_seg : forall {U:Type} {n w:nat} (f:seq_fun n w U), map_seq (fun m pf => let pf' := proj2 (lt_seq' m n w pf) in seq_fun_to_seg f (m - n) pf') = map_seg (seq_fun_to_seg f). intros U n w f. symmetry. rewrite map_seg_eq_iff. ex_goal. rewrite length_map_seq; auto. exists hex. intros m h1. simpl. unfold seq_fun_to_seg. erewrite nth_lt_functional3. rewrite (nth_lt_map_seq _ _ h1) at 1. apply f_equal_prop_dep. omega. Qed. Lemma map_seq_eq_iff : forall {U:Type} {n w:nat} (f:seq_fun n w U) (l:list U), map_seq f = l <-> (exists pfeq:w = length l, forall (m:nat) (pflt: m < w), let pflt' := lt_seq _ _ _ pflt in let pflt'' := lt_eq_trans _ _ _ pflt pfeq in f _ pflt' = nth_lt l m pflt''). intros U n w f l. split; intro h1. rewrite (seq_fun_eq_seg f) in h1. rewrite map_seq_fun_to_seg in h1. rewrite map_seg_eq_iff in h1. simpl in h1; simpl. destruct h1 as [h1 h2]. exists h1. intros m h3. specialize (h2 m h3). unfold seq_fun_to_seg in h2. assumption. rewrite (seq_fun_eq_seg f). rewrite map_seq_fun_to_seg. rewrite map_seg_eq_iff. simpl in h1; simpl. destruct h1 as [h1 h2]. exists h1. intros m h3. specialize (h2 _ h3). unfold seq_fun_to_seg. assumption. Qed. Lemma map_seg_decompose : forall {U:Type} {m n:nat} (f:seg_fun n U) (pfm:m < n), map_seg f = (map_seg (seg_fun_rest f pfm)) ++ map_seq (seg_fun_rest_seq f pfm). intros U m n f h1. rewrite map_seg_eq_iff. ex_goal. rewrite app_length. rewrite length_map_seg, length_map_seq; auto with arith. exists hex. intros m' h2 h3. destruct (lt_le_dec m' m) as [h4 | h4]. erewrite nth_lt_app1. unfold h3; clear h3. symmetry. gen0. intro h3. erewrite nth_lt_functional3. rewrite length_map_seg in h3. pose proof (nth_lt_map_seg (seg_fun_rest f h1) _ h3) as h5. simpl in h5. erewrite nth_lt_functional3 in h5. rewrite h5 at 1. unfold seg_fun_rest. f_equal. apply proof_irrelevance. unfold h3; clear h3. symmetry. gen0. intro h3. pose proof h3 as h3'. rewrite app_length in h3'. rewrite length_map_seg, length_map_seq in h3'. erewrite nth_lt_app2. gen0. rewrite length_map_seg. intro h10. pose proof h10 as h10'. rewrite length_map_seq in h10'. erewrite nth_lt_functional3. pose (seg_fun_rest_seq f h1). pose proof (nth_lt_map_seq (seg_fun_rest_seq f h1) (m' - m) h10') as h11. simpl in h11. rewrite h11 at 1. unfold seg_fun_rest_seq. apply f_equal_dep. omega. Grab Existential Variables. rewrite length_map_seg; auto. rewrite length_map_seg; auto. rewrite length_map_seg; auto. Qed. Definition seg_pred_f_eq_dec {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) i (pfi:i < n) : seg_pred_dec (fun j (pfj: j < n) => f i pfi = f j pfj) := fun j pfj => pfdu (f i pfi) (f j pfj). Definition seg_pred_f_eq_dec_S_eq {U n} {f:seg_fun (S n) U} {i} {pfi':i < S n} (pfdu:type_eq_dec U) (pfi: i < n) (pffs:seg_pred_dec (fun j (pfj:j < S n) => f i pfi' = f j pfj)) : seg_pred_dec (fun j (pfj:j < n) => (seg_fun_from_S f) i pfi = (seg_fun_from_S f) j pfj). unfold seg_pred_dec in pffs; unfold seg_pred_dec. intros j h2. unfold seg_fun_from_S. specialize (pffs j (lt_S _ _ h2)). destruct pffs as [h3 | h3]. left. erewrite f_equal_dep. rewrite h3 at 1; auto. reflexivity. right. erewrite f_equal_dep. apply h3. reflexivity. Defined. Definition map_ind {T:Type} (l:list T) : list (nat * T) := map_seg (fun (m:nat) (pf:m < length l) => (m, nth_lt l m pf)). Definition list_eq_prop {T:Type} (l l':list T) (pfeq:length l = length l') : list Prop := map_seg (fun (n:nat) (pf:n < length l) => nth_lt l n pf = nth_lt l' n (lt_eq_trans _ _ _ pf pfeq)). (*This is a binary analogue of [remove].*) Fixpoint keep_list_pr {T U:Type} (pfdu:type_eq_dec U) (l:list (T * U)) (u:U) : list T := match l with | nil => nil | ap::l' => let lk := keep_list_pr pfdu l' u in if (pfdu (snd ap) u) then (fst ap) :: lk else lk end. Lemma bv_tl_compat : forall {T:Type} (l:list T), l <> nil -> mask_to_coarse_list (eq_sym (length_bv_tl l)) = tl l. intros T l h1. induction l as [|a l ih]. contradict h1; auto. simpl. unfold mask_to_coarse_list. destruct l as [|b l]; simpl; auto. simpl in ih. f_equal. specialize (ih (cons_neq_nil _ _)). rewrite <- ih at 3. unfold fun_from_cons, mask_to_coarse_list; simpl. simpl. unfold bv1s. simpl. rewrite map_in_dep_eq_iff. ex_goal. rewrite length_map_in_dep. rewrite bv1s_aux_S. rewrite length_lS; auto. exists hex. intros n h2. pose proof h2 as h2'. rewrite bv1s_aux_S in h2'. rewrite length_lS in h2'. rewrite (nth_lt_map_in_dep _ _ h2'). pose proof (bv1s_aux_S (lrep true (length l)) 1) as h3. pose proof (nth_lt_functional1 _ _ n h3 h2) as h4. hfirst h4. rewrite <- bv1s_aux_S; auto. specialize (h4 hf). rewrite h4. unfold lS. rewrite (nth_lt_map _ _ _ h2'). pose proof (in_nth_lt (bv1s_aux (lrep true (length l)) 1) n h2') as h5. destruct ( nth_lt (bv1s_aux (lrep true (length l)) 1) n h2'); auto. apply nth_indep. rewrite bv1s_aux_S in h5. rewrite in_lS_iff in h5. destruct h5 as [m [h5 h6]]. apply S_inj in h6. rewrite h6. apply in_bv1s_lt in h5. rewrite length_lrep in h5; auto. Qed. Lemma mask_to_coarse_list_cons_false : forall {T:Type} {l:list T} {a:T} {v:bv} (pfeq:length (a::l) = length (false::v)), mask_to_coarse_list pfeq = mask_to_coarse_list (S_inj _ _ pfeq). intros T l a v h1. pose proof h1 as h1'. simpl in h1'. apply S_inj in h1'. unfold mask_to_coarse_list; simpl. unfold bv1s. simpl. rewrite map_in_dep_eq_iff. ex_goal. rewrite length_map_in_dep. rewrite bv1s_aux_S. rewrite length_lS; auto. exists hex. intros n h2. unfold fun_from_cons. erewrite nth_lt_map_in_dep. pose proof (bv1s_aux_S v 0) as h3. pose proof (nth_lt_functional1 _ _ _ h3 h2) as h4. hfirst h4. rewrite <- h3; auto. specialize (h4 hf). rewrite h4. unfold lS. erewrite nth_lt_map. symmetry. rewrite nth_lt_compat. apply nth_indep. pose proof (in_nth_lt (bv1s v) n) as h5. hfirst h5. pose proof h2 as h2'. rewrite bv1s_aux_S in h2'. rewrite length_lS in h2'; auto. specialize (h5 hf0). apply in_bv1s_lt in h5. rewrite h1'. erewrite nth_lt_functional3. apply h5. Grab Existential Variables. rewrite bv1s_aux_S in h2. rewrite length_lS in h2; auto. Qed. Lemma mask_to_coarse_list_cons_true : forall {T:Type} {l:list T} {a:T} {v:bv} (pfeq:length (a::l) = length (true::v)), mask_to_coarse_list pfeq = a :: mask_to_coarse_list (S_inj _ _ pfeq). intros T l a v h1. unfold mask_to_coarse_list; simpl. f_equal. rewrite map_in_dep_eq_iff. ex_goal. rewrite length_map_in_dep. rewrite bv1s_aux_S. rewrite length_lS. reflexivity. exists hex. intros n h2. unfold fun_from_cons. erewrite nth_lt_map_in_dep. pose proof (bv1s_aux_S v 0) as h3. pose proof (nth_lt_functional1 _ _ _ h3 h2) as h4. hfirst h4. rewrite <- h3; auto. specialize (h4 hf). rewrite h4. unfold lS. erewrite nth_lt_map. symmetry. rewrite nth_lt_compat. apply nth_indep. pose proof (in_nth_lt (bv1s v) n) as h5. hfirst h5. pose proof h2 as h2'. rewrite bv1s_aux_S in h2'. rewrite length_lS in h2'; auto. specialize (h5 hf0). apply in_bv1s_lt in h5. pose proof h1 as h1'. simpl in h1'. apply S_inj in h1'. rewrite h1'. erewrite nth_lt_functional3. apply h5. Grab Existential Variables. erewrite <- length_lS. rewrite <- bv1s_aux_S; auto. Qed. Lemma mask_to_coarse_list_compat : forall {T:Type} {l:list T} {v:bv} (pfdt:type_eq_dec T) (pfeq:length l = length v), coarse_list (mask_to_coarse_list pfeq) l. intros T l. induction l as [|a l ih]. intros v hdt h1. pose proof h1 as h1'. symmetry in h1'. apply length_eq0 in h1'; subst. simpl in h1. unfold mask_to_coarse_list. constructor. intros v hdt h1. destruct v as [|b v]. simpl in h1; discriminate. pose proof h1 as h1'. apply S_inj in h1'. specialize (ih _ hdt h1'). destruct b. pose proof (mask_to_coarse_list_cons_true h1) as h2. rewrite h2. apply coarse_list_cons2. match goal with |- coarse_list ?l'' _ => pose l'' as l' end. assert (h3:l' = mask_to_coarse_list h1'). unfold l', mask_to_coarse_list. apply map_in_dep_ext. intros i h3. apply functional_extensionality. intro h3'. apply nth_lt_functional3. subst. rewrite h3; auto. pose proof (mask_to_coarse_list_cons_false h1) as h2. rewrite h2. match goal with |- coarse_list ?l'' _ => pose l'' as l' end. assert (h3:l' = mask_to_coarse_list h1'). unfold l', mask_to_coarse_list. apply map_in_dep_ext. intros i h3. apply functional_extensionality. intro h3'. apply nth_lt_functional3. subst. rewrite h3; auto. eapply coarse_list_trans. assumption. apply ih. apply coarse_list_cons1. apply coarse_list_refl. Qed. Lemma coarse_list_mask_to_coarse_list : forall {T} {l:list T} {v:bv} (pfdt:type_eq_dec T) (pfe:length l = length v), let pfc := mask_to_coarse_list_compat pfdt pfe in coarse_list_mask (mask_to_coarse_list pfe) l pfc v pfe. intros T l. induction l as [|a l ih]. intros v hdt he. pose proof (length_eq0 _ (eq_sym he)); subst. simpl. rewrite coarse_list_mask_iff'. constructor. red. intros; contradiction. intro v. destruct v as [|b v]. intros hdt he hc. simpl in he. discriminate. intros hdt he hc. pose proof he as he'. simpl in he'. specialize (ih v hdt (S_inj _ _ he')). simpl in ih. term2 ih. destruct b. pose proof (coarse_list_mask_cons2 _ _ a c _ (S_inj _ _ he') ih) as h4. eapply coarse_list_mask_functional in h4. apply h4. rewrite mask_to_coarse_list_cons_true. repeat f_equal; apply proof_irrelevance. reflexivity. reflexivity. pose proof (coarse_list_mask_cons1 _ _ a c _ (S_inj _ _ he') ih) as h4. eapply coarse_list_mask_functional in h4. apply h4. rewrite mask_to_coarse_list_cons_false. repeat f_equal; apply proof_irrelevance. reflexivity. reflexivity. Qed. (*Finds the index of the [S occ]'s occureence of [x] in [l]. if [occ >= pred (count_occ)] then the value is the same as [pred(count_occ)]'s *) Fixpoint lind_occ {T:Type} (pfdt:type_eq_dec T) (l:list T) (occ:nat) (x:T) : In x l -> nat := match l with | nil => fun pf => False_rect nat pf | a::l' => fun pf => match (pfdt x a) with | left _ => match occ with | 0 => 0 | S n => match in_dec pfdt x l' with | left pfs => S (lind_occ pfdt l' n x pfs) | right _ => O end end | right pfn => let pf' := neq_in_cons _ _ _ pfn pf in S (lind_occ pfdt l' occ x pf') end end. (*Minimum list index of a given element in a list.*) Definition lind {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l):= lind_occ pfdt l 0 x pf. Lemma lind_functional : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x y:T) (pfx:In x l) (pfy:In y l), x = y -> lind pfdt l x pfx = lind pfdt l y pfy. intros T h0 l x y h2 h3 h4. subst. pose proof (proof_irrelevance _ h2 h3); subst; auto. Qed. Lemma lind_functional' : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x y:T) (pfx:In x l) (pfy:In y l'), l = l' -> x = y -> lind pfdt l x pfx = lind pfdt l' y pfy. intros T h0 l l' x y h2 h3 h4 h5. subst. pose proof (proof_irrelevance _ h2 h3); subst; auto. Qed. Lemma lind_functional3 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfx pfx':In x l), lind pfdt l x pfx = lind pfdt l x pfx'. intros; f_equal; apply proof_irrelevance. Qed. Lemma lind_occ_functional : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (occ occ':nat) (x x':T) (pf:In x l) (pf':In x' l'), l = l' -> occ = occ' -> x = x' -> lind_occ pfdt l occ x pf = lind_occ pfdt l' occ' x' pf'. intros T h1 l l' occ occ' x x' h2 h3 h4 h5 h6. subst. pose proof (proof_irrelevance _ h2 h3); subst; auto. Qed. Lemma lind_occ_functional3 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (occ:nat) (x:T) (pf pf':In x l), lind_occ pfdt l occ x pf = lind_occ pfdt l occ x pf'. intros; f_equal; apply proof_irrelevance. Qed. Lemma lt_lind_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (occ:nat) (x:T) (pf:In x l), lind_occ pfdt l occ x pf < length l. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros occ x h3. destruct (h1 x a) as [h4 | h4]; subst. destruct occ as [|occ]; simpl; auto with arith. destruct (in_dec h1 a l) as [h4 | h4]. apply lt_n_S. apply ih. auto with arith. apply lt_n_S. apply ih. Qed. Lemma lind_occ_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (occ:nat) (x:T) (pf:In x l), nth_lt l (lind_occ pfdt l occ x pf) (lt_lind_occ _ _ _ _ _) = x. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros occ x h2. destruct (h1 x a) as [h3 | h3]; simpl; subst. destruct occ as [|occ]; simpl; auto. destruct (in_dec h1 a l) as [h3 | h3]; simpl. specialize (ih occ a h3). simpl in ih. rewrite <- ih. erewrite nth_lt_functional3. rewrite nth_lt_compat at 1. apply nth_indep. apply lt_lind_occ. reflexivity. specialize (ih occ x (neq_in_cons _ _ _ h3 h2)). simpl in ih. rewrite <- ih. erewrite nth_lt_functional3. rewrite nth_lt_compat at 1. apply nth_indep. apply lt_lind_occ. Grab Existential Variables. apply lt_lind_occ. apply lt_lind_occ. Qed. Lemma ex_lind_occ_nth_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pf:n < length l), exists (occ:nat), occ < count_occ pfdt l (nth_lt l n pf) /\ lind_occ pfdt l occ (nth_lt l n pf) (in_nth_lt _ _ _) = n. intros T h1 l. induction l as [|a l ih]. intros ? h2; simpl in h2; omega. intros n h2. simpl in h2. destruct n as [|n]. exists 0. simpl. destruct (h1 a a) as [h3 | h3]. split; auto with arith. contradict h3; auto. pose proof (lt_S_n _ _ h2) as h2'. specialize (ih n h2'). destruct ih as [occ [h3 h4]]. simpl. rewrite switch_eq_dec_eq. destruct (h1 (nth n l a) a) as [h5 | h5]. unfold switch_eq_dec. destruct in_dec as [h6 | h6]. exists (S occ). split. apply lt_n_S. rewrite nth_lt_compat in h3. erewrite nth_indep. apply h3. assumption. f_equal. rewrite <- h4. apply lind_occ_functional; auto. rewrite nth_lt_compat. apply nth_indep. assumption. contradict h6. apply nth_In; auto. unfold switch_eq_dec. exists occ. split; auto. rewrite nth_lt_compat in h3. erewrite nth_indep in h3. apply h3. assumption. f_equal. rewrite <- h4. apply lind_occ_functional; auto. rewrite nth_lt_compat. apply nth_indep. assumption. Qed. Lemma lt_lind : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l), lind pfdt l x pfin < length l. unfold lind; intros; apply lt_lind_occ. Qed. Lemma lind_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:In a l), nth_lt l (lind pfdt l a pf) (lt_lind pfdt l a pf) = a. unfold lind; intros; erewrite nth_lt_functional3; rewrite lind_occ_compat at 1; auto. Qed. (*See [lind_cons_same] for [occ = 0]*) Lemma lind_occ_cons_in_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) occ (pf:In x l) (pf':In x (x::l)), 0 < occ -> lind_occ pfdt (x::l) occ x pf' = S (lind_occ pfdt l (pred occ) x pf). intros T h1 l x occ h2 h3 h4. simpl. rewrite eq_dec_eq at 1; auto. rewrite (S_pred _ _ h4). lr_match_goal'. f_equal. apply lind_occ_functional; auto. contradiction. Qed. Lemma lind_occ_cons_in_neq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a x:T) (pf:In x l) (occ:nat), x <> a -> lind_occ pfdt (a::l) occ x (in_cons _ _ _ pf) = S (lind_occ pfdt l occ x pf). intros T h1 l a x h2 occ h3. simpl. destruct (h1 x a); subst. contradict h3; auto. do 2 f_equal. apply proof_irrelevance. Qed. Lemma nth_lt_lind_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l), let pf := lt_lind pfdt l x pfin in nth_lt l (lind pfdt l x pfin) pf = x. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h3. unfold lind. simpl. destruct (h1 x a) as [h4 | h4]; auto. simpl in ih. pose proof (neq_in_cons _ _ _ h4 h3) as h5. specialize (ih _ h5). rewrite <- ih. pose proof (nth_lt_compat l (lind h1 l x h5) (lt_lind h1 l x h5)) as h6. simpl in h6. rewrite h6. match goal with |- nth ?x _ _ = nth ?y _ _ => assert (h7:x = y) end. unfold lind. f_equal. apply proof_irrelevance. rewrite h7. apply nth_indep. apply lt_lind. Qed. Lemma lind_occ_inj : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x y:T) (pfx:In x l) (pfy:In y l) (m m':nat), lind_occ pfdt l m x pfx = lind_occ pfdt l m' y pfy -> x = y. intros T h0 l x y h2 h3 m m' h4. pose proof (nth_lt_functional2 _ _ _ h4 (lt_lind_occ _ _ _ _ _) (lt_lind_occ _ _ _ _ _)) as h5. do 2 rewrite lind_occ_compat in h5; auto. Qed. Lemma lind_occ_inj' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfx pfx':In x l) (m m':nat), m < count_occ pfdt l x -> m' < count_occ pfdt l x -> lind_occ pfdt l m x pfx = lind_occ pfdt l m' x pfx' -> m = m'. intros T hdt l. induction l as [|a l ih]; simpl. intros; omega. intros x hx hx' m m'. destruct (hdt a x) as [|h7]; subst. rewrite eq_dec_same at 1. rewrite eq_dec_same at 1. destruct m as [|m], m' as [|m']; auto. intros h1 h2 h3. apply lt_S_n in h2. pose proof (count_occ_in hdt _ _ _ h2) as h4. destruct in_dec as [h5 | h5]. discriminate. contradiction. intro h1. apply lt_S_n in h1. pose proof (count_occ_in hdt _ _ _ h1) as h4. destruct in_dec as [h5 | h5]. discriminate. contradiction. intros h1 h2. apply lt_S_n in h1. apply lt_S_n in h2. pose proof (count_occ_in hdt _ _ _ h1) as h4. pose proof (count_occ_in hdt _ _ _ h2) as h5. destruct in_dec as [h6 | h6]. intro h7. apply S_inj in h7. f_equal; auto. eapply ih. apply h1. assumption. apply h7. contradiction. assumption. assumption. intros h8 h9. destruct hx as [|h10]; subst. contradiction (h7 eq_refl). destruct (hdt x a) as [h11 | h11]; subst. contradict h7; auto. intro h12. apply S_inj in h12. eapply ih. apply h8. assumption. apply h12. Qed. Lemma lind_inj_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x y:T) (pfx:In x l) (pfy:In y l), lind pfdt l x pfx = lind pfdt l y pfy <-> x = y. intros T h0 l x y h2 h3; split; intro h4. pose proof (nth_lt_lind_eq h0 l x h2) as h5. pose proof (nth_lt_lind_eq h0 l y h3) as h6. simpl in h5, h6. rewrite <- h5, <- h6. gen0. rewrite h4. intro; f_equal. apply proof_irrelevance. subst. apply lind_functional; auto. Qed. Lemma inj_seg_nth_lt' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), inj_seg (nth_lt l) -> NoDup l. intros T hdt l. induction l as [|a l ih]; simpl. intros; constructor. intro h1. constructor. intro h2. red in h1. specialize (h1 0 (S (lind hdt _ _ h2)) (lt_0_Sn _)). hfirst h1. apply lt_n_S. apply lt_lind. specialize (h1 hf). simpl in h1. hfirst h1. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite lind_compat; auto. apply lt_lind. apply h1 in hf0. discriminate. apply ih. red in h1; red. intros i j h2 h3 h4. specialize (h1 (S i) (S j)). hfirst h1. apply lt_n_S; auto. specialize (h1 hf). hfirst h1. apply lt_n_S; auto. specialize (h1 hf0). simpl in h1. hfirst h1. erewrite nth_indep. symmetry. erewrite nth_indep. erewrite <- nth_lt_compat. rewrite <- h4 at 1. apply nth_lt_compat. assumption. assumption. apply h1 in hf1. apply S_inj in hf1; auto. Qed. Lemma lind_sing : forall {T:Type} (pfdt:type_eq_dec T) (a x:T) (pf:In x (a::nil)), lind pfdt (a::nil) x pf = 0. intros T hd a x h0. destruct h0 as [h0 | h0]. subst. unfold lind. simpl. destruct (hd x x) as [hx | hx]. auto. contradict hx; auto. simpl in h0. contradiction. Qed. Lemma lind_firstn : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (m:nat) (pf:In x (firstn m l)), let pf' := in_firstn _ _ _ pf in lind pfdt (firstn m l) x pf = lind pfdt l x pf'. intros T h1 l. induction l as [|a l ih]; simpl. intros ? ? h2. pose proof h2 as h2'. rewrite firstn_nil in h2'. contradiction. intros x m h2. unfold lind; simpl. destruct (h1 x a) as [h4 | h4]; subst. destruct (zerop m) as [h4 | h4]; subst. simpl in h2. contradiction. gen0. rewrite (firstn_cons l m a h4). intro h5. simpl. destruct (h1 a a) as [h6 | h6]; auto. contradict h6; auto. simpl in ih. destruct (zerop m) as [h5 | h5]; subst. simpl in h2. contradiction. gen0. rewrite (S_pred _ _ h5). intro h0. simpl. simpl in h0. destruct h0 as [|h0]; subst. contradict h4; auto. destruct (h1 x a) as [|ha]; subst. contradict h4; auto. f_equal. gen0. intro h3. specialize (ih _ _ h3). unfold lind in ih. rewrite ih. f_equal. apply proof_irrelevance. Qed. Lemma lind_cons_same : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:In a (a::l)), lind pfdt (a::l) a pf = 0. intros T h1 l a. unfold lind. simpl. destruct (h1 a a) as [h2 | h2]; auto. contradict h2. reflexivity. Qed. Lemma lind_hd' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pf:l <> nil), lind pfdt l (hd' _ pf) (in_hd' _ _) = 0. intros T h1 l. destruct l as [|a l]. intro h2; contradict h2; auto. intro h2. simpl. erewrite lind_functional. rewrite lind_cons_same; auto. reflexivity. Grab Existential Variables. left; auto. Qed. Lemma lind_occ_hd : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (occ:nat) (pfin:In x (x::l)) (pfo:occ < count_occ pfdt l x), let pfin' := count_occ_in _ _ _ _ pfo in lind_occ pfdt (x :: l) (S occ) x pfin = S (lind_occ pfdt l occ x pfin'). intros T h1 l x occ h2 h3 h4; simpl. destruct (h1 x x) as [h6|h6]. destruct in_dec as [h7|h7]. repeat f_equal; try apply proof_irrelevance. contradiction. contradict h6; auto. Qed. Lemma lind_occ_hd' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pf:l <> nil) (occ:nat) (pfo:S occ < count_occ pfdt l (hd' _ pf)), let pfin' := impl_in_tl pfdt _ _ _ pfo in lind_occ pfdt l (S occ) (hd' l pf) (in_hd' _ _) = S (lind_occ pfdt (tl l) occ _ pfin'). intros T h1 l h2 occ h3 h4. destruct l as [|a l]. pose proof (h2 eq_refl); contradiction. simpl. rewrite (eq_dec_same h1) at 1. destruct in_dec as [h5 | h5]. repeat f_equal. apply proof_irrelevance. simpl in h3. pose proof h3 as h3'. rewrite (eq_dec_same h1) in h3' at 1. simpl in h3'. apply lt_S_n in h3'. contradict h5. apply count_occ_in in h3'; auto. Qed. Lemma lind_eq0_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l), lind pfdt l x pfin = 0 <-> hd' l (in_not_nil _ _ pfin) = x. intros T h1 l x h2; split; intro h3. destruct l as [|a l]; simpl. contradiction. unfold lind in h3. simpl in h3. destruct (h1 x a) as [h4 | h4]; auto. discriminate. generalize h2. rewrite <- h3. intro h2'. erewrite lind_functional. apply lind_hd'. reflexivity. Qed. Lemma lind_occ_eq0_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (occ:nat) (pfin:In x l), occ < count_occ pfdt l x -> (lind_occ pfdt l occ x pfin = 0 <-> (occ = 0 /\ hd' l (in_not_nil _ _ pfin) = x)). intros T h1 l. destruct l as [|a l]; simpl. intros; contradiction. intros x occ h2 hlt. destruct h2 as [|h2]; subst. destruct occ as [|occ]. destruct (h1 x x) as [h2 | h2]. tauto. contradict h2; auto. destruct (h1 x x) as [h2 | h2]. apply lt_S_n in hlt. destruct (in_dec h1 x l) as [h3 | h3]. intuition. contradict h3. eapply count_occ_in. apply hlt. contradict h2; auto. rewrite switch_eq_dec_eq in hlt. destruct (h1 x a) as [|h3]; subst. unfold switch_eq_dec in hlt. destruct occ as [|occ]. tauto. apply lt_S_n in hlt. destruct in_dec as [h3 | h3]. intuition. contradict h3. eapply count_occ_in. apply hlt. unfold switch_eq_dec in hlt. intuition. Qed. Lemma lind_occ_eq0_iff' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (occ:nat) (pfin:In x l), (lind_occ pfdt l occ x pfin = 0 <-> ((occ = 0 \/ count_occ pfdt l x = 1) /\ hd' l (in_not_nil _ _ pfin) = x)). intros T h1 l. destruct l as [|a l]; simpl. intros; contradiction. intros x occ h2. destruct h2 as [|h2]; subst. destruct occ as [|occ]. destruct (h1 x x) as [h2 | h2]. tauto. contradict h2; auto. destruct (h1 x x) as [h2 | h2]. destruct (in_dec h1 x l) as [h3 | h3]. intuition. apply S_inj in H. pose proof (count_occ_In h1 l x) as h4. pose proof h3 as h3'. rewrite h4 in h3'. omega. rewrite (count_occ_In h1 l x) in h3. intuition. contradict h2; auto. rewrite switch_eq_dec_eq. destruct (h1 a x) as [|h3]; subst. unfold switch_eq_dec. destruct occ as [|occ]. tauto. destruct in_dec as [h3 | h3]. intuition. apply S_inj in H. rewrite (count_occ_In h1 l x) in h2. omega. contradiction. unfold switch_eq_dec. intuition. Qed. Lemma lind_cons_neq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a x:T) (pfix:In x l) (pfn:x <> a), lind pfdt (a::l) x (or_intror pfix) = S (lind pfdt l x pfix). intros T h1 l a x h3. unfold lind. simpl. destruct (h1 x a) as [|h4]; subst. intro h4; contradict h4; auto. intro h5. f_equal. f_equal. apply proof_irrelevance. Qed. Lemma nth_lt_eq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pflt:n < length l) (x:T), nth_lt l n pflt = x <-> exists (pfin:In x l) (occ:nat), occ < count_occ pfdt l x /\ lind_occ pfdt l occ x pfin = n. intros T h1 l n h2 ?; split; intro ha. subst. exists (in_nth_lt _ _ _). revert h2. revert n. induction l as [|a l ih]; simpl. intros; omega. intros n h2. destruct n as [|n]. destruct (h1 a a) as [|h3]. exists 0. split; auto with arith. contradict h3; auto. pose proof h2 as h2'. apply lt_S_n in h2'. specialize (ih n h2'). destruct ih as [occ ih]. destruct ih as [h3 h4]. rewrite (switch_eq_dec_eq h1 a (nth n l a)). destruct (h1 (nth n l a) a) as [h5 | h5]. unfold switch_eq_dec. destruct in_dec as [h6 | h6]. exists (S occ). rewrite nth_lt_compat in h3. split. apply lt_n_S. erewrite nth_indep. apply h3. apply lt_S_n in h2; auto. f_equal. generalize h6. erewrite nth_indep. rewrite <- nth_lt_compat. intro h6'. erewrite lind_occ_functional. apply h4. reflexivity. reflexivity. reflexivity. apply lt_S_n in h2; auto. contradict h6. apply nth_In. apply lt_S_n in h2; auto. unfold switch_eq_dec. exists occ. split. erewrite nth_indep. rewrite <- nth_lt_compat. apply h3. apply lt_S_n in h2; auto. f_equal. erewrite lind_occ_functional. apply h4. reflexivity. reflexivity. rewrite nth_lt_compat. apply nth_indep. apply lt_S_n in h2; auto. destruct ha as [h3 [occ [h4 h5]]]. subst. erewrite nth_lt_functional3. rewrite lind_occ_compat; auto. Qed. Lemma nth_lts_eq_length_list_singularize : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l'), (forall (m n : nat) (pfm : m < length l) (pfn : n < length l), let pfm' := lt_eq_trans m (length l) (length l') pfm pfe in let pfn' := lt_eq_trans n (length l) (length l') pfn pfe in nth_lt l m pfm = nth_lt l n pfn <-> nth_lt l' m pfm' = nth_lt l' n pfn') -> length (list_singularize pfdt l) = length (list_singularize pfdu l'). intros T U hdt hdu l. induction l as [|a l ih]. intros l' h1. pose proof h1 as h1'. symmetry in h1'. apply length_eq0 in h1'; subst; simpl; auto. intro l'. destruct l' as [|a' l']. simpl; intros; discriminate. intros h1 h2. pose proof h1 as h1'. simpl in h1'. apply S_inj in h1'. destruct (in_dec hdt a l) as [h3 | h3], (in_dec hdu a' l') as [h4 | h4]. rewrite list_singularize_in_cons; auto. rewrite list_singularize_in_cons; auto. eapply ih. intros m n h5 h6 h7 h8. specialize (h2 (S m) (S n)). hfirst h2. simpl; auto with arith. specialize (h2 hf). hfirst h2. simpl; auto with arith. specialize (h2 hf0). simpl in h2. erewrite nth_lt_compat. erewrite (nth_lt_compat l n). erewrite (nth_lt_compat l' m). erewrite (nth_lt_compat l' n). erewrite nth_indep; auto. erewrite (nth_indep l _ _ h6); auto. erewrite (nth_indep l'); auto. erewrite (nth_indep l' _ _ h8); auto. apply h2. specialize (h2 (S (lind hdt _ _ h3)) 0). hfirst h2. simpl. eapply lt_n_S. apply lt_lind. specialize (h2 hf). hfirst h2. simpl; auto with arith. specialize (h2 hf0). simpl in h2. erewrite (nth_indep l) in h2. erewrite <- nth_lt_compat in h2. rewrite lind_compat in h2. pose proof (iff1 h2 eq_refl) as h5. pose proof (@nth_In _ (lind hdt l a h3) l' a') as h6. hfirst h6. eapply lt_eq_trans. apply lt_lind. auto. specialize (h6 hf1). rewrite h5 in h6. contradiction. apply lt_lind. specialize (h2 (S (lind hdu _ _ h4)) 0). hfirst h2. simpl. eapply lt_n_S. rewrite h1'. apply lt_lind. specialize (h2 hf). hfirst h2. simpl; auto with arith. specialize (h2 hf0). simpl in h2. erewrite (nth_indep l') in h2. erewrite <- nth_lt_compat in h2. rewrite lind_compat in h2. pose proof (iff2 h2 eq_refl) as h5. pose proof (@nth_In _ (lind hdu l' a' h4) l a) as h6. hfirst h6. eapply lt_eq_trans. apply lt_lind. auto. specialize (h6 hf1). rewrite h5 in h6. contradiction. apply lt_lind. rewrite list_singularize_nin; auto. rewrite list_singularize_nin; auto. simpl; f_equal. eapply ih. intros m n h5 h6 h7 h8. specialize (h2 (S m) (S n)). hfirst h2. simpl; auto with arith. specialize (h2 hf). hfirst h2. simpl; auto with arith. specialize (h2 hf0). simpl in h2. erewrite nth_lt_compat. erewrite (nth_lt_compat l n). erewrite (nth_lt_compat l' m). erewrite (nth_lt_compat l' n). erewrite nth_indep; auto. erewrite (nth_indep l _ _ h6); auto. erewrite (nth_indep l'); auto. erewrite (nth_indep l' _ _ h8); auto. apply h2. Grab Existential Variables. assumption. assumption. Qed. Lemma leq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), l = l' <-> exists (pfe:length l = length l'), forall (i:nat) (pfi:i < length l), let pfi' := lt_eq_trans _ _ _ pfi pfe in nth_lt l i pfi = nth_lt l' i pfi'. intros T h0 l l'; split; intro h1. subst. exists eq_refl. intros i h1 h2. apply nth_lt_functional3. destruct h1 as [h1 h2]. generalize dependent l'. induction l as [|a l ih]; simpl. intros l' h1. pose proof h1 as h1'. symmetry in h1'. apply length_eq0 in h1'; subst; auto. intros l' h1. destruct l' as [|a' l']. simpl in h1. discriminate. simpl in h1. pose proof h1 as h1'. apply S_inj in h1'. intro h2. destruct (h0 a a') as [|h3]; subst. f_equal. eapply ih. intros i h3. specialize (h2 (S i)). hfirst h2. apply lt_n_S; auto. specialize (h2 hf). simpl in h2. simpl. rewrite nth_lt_compat. symmetry. rewrite nth_lt_compat. symmetry. erewrite nth_indep. rewrite h2 at 1. apply nth_indep. rewrite <- h1'; auto. assumption. specialize (ih _ h1'); simpl in ih. specialize (h2 0 (lt_0_Sn _)). simpl in h2. contradiction. Grab Existential Variables. assumption. Qed. Lemma app_cons_eq_sing_iff : forall {T} (l l':list T) a, l ++ a :: l' = a :: nil <-> l = nil /\ l' = nil. intros T l l' a; split; intro h1; destruct l, l'; auto; simpl in h1; inversion h1; subst; try discriminate. apply app_not_nil_l in H1; contradict H1. subst. rewrite app_nil_l. intro; discriminate. apply app_not_nil_l in H1; contradict H1; subst. simpl. intro; discriminate. Qed. Lemma app_cons_nil_eq_iff : forall {T} {l l':list T} a a', ~In a l -> ~In a l' -> (l ++ a :: nil = l' ++ a' :: nil <-> (a = a' /\ l = l')). intros T l. induction l as [|b l ih]; subst. intro l'; destruct l' as [|b l']; simpl; intros; split; intro h1. inversion h1; auto. destruct h1; subst; auto. symmetry in h1. inversion h1; subst. contradict H0; left; auto. destruct h1; discriminate. intro l'. destruct l' as [|b' l']. intros; simpl; split; intro h1. inversion h1; subst. apply app_not_nil_r in H3; contradict H3. discriminate. destruct h1; discriminate. intros a c h1 h2; split; intro h3. inversion h3; subst. inversion h3. rewrite ih in H0. destruct H0; subst; auto. contradict h1; right; auto. contradict h2; right; auto. destruct h3 as [? h3]; subst. inversion h3; subst; auto. Qed. Lemma app_cons_nil_eq_iff' : forall {T} {l l0 l':list T} a, ~In a l -> ~In a l' -> ~In a l0 -> (l ++ a :: nil = l' ++ a :: l0 <-> l0 = nil /\ l = l'). intros T l l0 l' a h2 h3; split; intro h4. destruct l0; simpl in h4. rewrite app_cons_nil_eq_iff in h4; tauto. rewrite app_eq_iff in h4. destruct h4 as [h4 h5]. symmetry in h5. rewrite sing_eq_iff in h5. destruct h5 as [h5 [h6 h7]]. destruct (nil_dec' l') as [h8|h8], (nil_dec' l) as [h9|h9]; revert h7 h6 h4; subst; simpl. intro; discriminate. intros h6 h7 h8. pose proof (nin_eq_firstn_length _ (t::l0) a h2 h8); contradiction. intros h9 h10. erewrite hd_app' in h10. rewrite <- h10 in h3. contradict h3. apply in_hd'. intros h10 h11. pose proof (tl_sing _ h5 h10) as h12. rewrite h11 in h12. intro h13. pose proof (firstn_skipn (length l) (l'++a::t::l0)) as h14. rewrite <- h13 in h14. rewrite h12 in h14. apply app_sing_eq_app_cons_cons in h14. contradiction. assumption. destruct h4; subst; auto. Grab Existential Variables. assumption. Qed. Lemma app_cons_eq_length : forall {T} (l l0 l' l0':list T) (a:T), ~In a l -> ~In a l0 -> ~In a l' -> ~In a l0' -> l ++ a :: l0 = l' ++ a :: l0' -> length l = length l'. intros T l. induction l as [|b l ih]; simpl. intros l0 l'. destruct l' as [|b' l']; simpl; auto. intros l0' a h1 h2 h3 h4 h5. inversion h5; subst. contradict h3; left; auto. intros l0 l'; destruct l' as [|b' l']. intros l' a h1 h2 h3 h4 h5. simpl in h5; inversion h5; subst; contradict h1; left; auto. intros l0' a h1 h2 h3 h4 h5. simpl in h5. inversion h5; subst. inversion h5. simpl; f_equal. eapply ih. contradict h1. right; apply h1. apply h2. contradict h3; right; auto. apply h4. assumption. Qed. Lemma app_cons_eq_iff : forall {T} (l l0 l' l0':list T) (a:T), ~In a l -> ~In a l0 -> ~In a l' -> ~In a l0' -> (l ++ a :: l0 = l' ++ a :: l0' <-> l = l' /\ l0 = l0'). intros T l l0 l' l0' a h1 h2 h3 h4; split; intro h5. pose proof (f_equal (firstn (length l)) h5) as h6. rewrite firstn_app_le, firstn_length in h6. pose proof (f_equal (@length _) h5) as h7. do 2 rewrite app_length in h7; simpl in h7. do 2 rewrite plus_S_shift_r in h7. apply S_inj in h7. pose proof (f_equal (firstn (length l')) (eq_sym h5)) as h6'. rewrite firstn_app_le, firstn_length in h6'; auto. rewrite h5 in h6'. erewrite app_cons_eq_length in h6. split. rewrite h6; rewrite h6' at 2. reflexivity. pose proof (f_equal (skipn (length l)) h5) as h16. rewrite skipn_length_app in h16. erewrite app_cons_eq_length in h16. rewrite skipn_length_app in h16. clear h6 h6'. inversion h16; subst; auto. apply h1. apply h2. apply h3. apply h4. assumption. apply h1. apply h2. assumption. apply h4. assumption. apply le_refl. destruct h5; subst; auto. Qed. Lemma le_lind_nth_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pflt:n < length l), lind pfdt l (nth_lt l n pflt) (in_nth_lt _ _ _) <= n. intros T h1 l n. revert l. induction n as [|n ih]; simpl. intros l h2. pose proof (nth_lt0 l (lt_length _ _ h2)) as h3. pose proof (proof_irrelevance _ h2 (not_nil_lt _ (lt_length l 0 h2))) as h4. simpl in h3. rewrite <- h4 in h3. generalize (in_nth_lt l 0 h2). rewrite h3. intro h5. erewrite lind_functional. rewrite lind_hd'; auto with arith. reflexivity. intros l h2. pose proof h2 as h2'. rewrite (S_pred _ _ h2') in h2'. apply lt_S_n in h2'. specialize (ih (tl l)). hfirst ih. rewrite length_tl; auto. specialize (ih hf). pose proof (lt_length _ _ h2) as h3. pose proof (hd_compat' _ h3) as h4. simpl in h4. pose proof (nth_lt_functional1 _ _ (S n) h4 h2) as h5. simpl in h5. hfirst h5. apply lt_n_S. rewrite length_tl; auto. specialize (h5 hf0). pose proof (@lind_functional). erewrite nth_indep in h5. rewrite <- (nth_lt_compat _ _ hf) in h5 at 1. pose proof (in_nth_lt (tl l) n hf) as h6. pose proof (in_cons (hd' l h3) _ _ h6) as h7. pose proof (lind_functional' h1 _ _ _ _ (in_nth_lt l (S n) h2) h7 h4 h5) as h8. rewrite h8. simpl. destruct (h1 (hd' l h3) (nth_lt (tl l) n hf)) as [h9 | h9]. generalize h7. rewrite <- h9. intro h7'. erewrite lind_functional. rewrite lind_cons_same; auto with arith. reflexivity. apply neq_sym in h9. pose proof (lind_cons_neq h1 _ _ _ h6 h9) as h10. pose proof (proof_irrelevance _ h7 (or_intror h6)) as h11. rewrite <- h11 in h10. rewrite h10. apply le_n_S. erewrite lind_functional. apply ih. reflexivity. assumption. Grab Existential Variables. left; auto. Qed. Lemma lind_app_in1 : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T) (pf:In a l), let pf' := in_app_l pf l' in lind pfdt (l++l') a pf' = lind pfdt l a pf. intros T h1 l. induction l as [|b l ih]; simpl. intros; contradiction. intros l' a h2. destruct h2 as [|h2]; subst. erewrite lind_functional. rewrite lind_cons_same; auto. erewrite lind_functional. rewrite lind_cons_same; auto. reflexivity. reflexivity. specialize (ih l' a h2). simpl in ih. destruct (h1 a b) as [|h3]; subst. erewrite lind_functional. rewrite lind_cons_same. erewrite lind_functional. rewrite lind_cons_same. reflexivity. reflexivity. reflexivity. pose proof (lind_cons_neq h1 (l++l') b a) as h4. hfirst h4. apply in_app_l; auto. specialize (h4 hf h3). rewrite lind_cons_neq; auto. erewrite lind_functional. rewrite h4 at 1. f_equal. erewrite lind_functional. rewrite ih at 1; auto. reflexivity. reflexivity. Grab Existential Variables. left; auto. left; auto. left; auto. left; auto. Qed. Lemma lind_app_in2 : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a':T) (pfin':In a' l') (pfnin: ~In a' l), let pfin := in_app_r pfin' l in lind pfdt (l++l') a' pfin = length l + lind pfdt l' a' pfin'. intros T h1 l. induction l as [|a l ih]; simpl. intros; f_equal; apply proof_irrelevance. intros l' a' h2 h3. unfold lind. simpl. destruct (h1 a' a) as [|h4]; subst. contradict h3; left; auto. f_equal. specialize (ih l' a' h2). hfirst ih. contradict h3; right; auto. specialize (ih hf). simpl in ih. unfold lind in ih. rewrite <- ih. f_equal. apply proof_irrelevance. Qed. Lemma lind_occ_app_lt_in1 : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T) (occ:nat) (pflt:occ < count_occ pfdt l x), let pfin := count_occ_in pfdt _ _ _ pflt in let pfin' := in_app_l pfin l' in lind_occ pfdt (l++l') occ x pfin' = lind_occ pfdt l occ x pfin. intros T h1 l. induction l as [|a l ih]; simpl. intros; omega. intros l' x occ h3. pose proof h3 as h3'. rewrite switch_eq_dec_eq in h3'. destruct (h1 x a) as [|h4]; subst. unfold switch_eq_dec in h3'. destruct occ as [|occ] ;auto. apply lt_S_n in h3'. destruct (in_dec h1 a l) as [h4 | h4]. destruct in_dec as [h5 | h5]. f_equal. erewrite lind_occ_functional. rewrite ih at 1. f_equal. apply proof_irrelevance. reflexivity. reflexivity. reflexivity. contradict h5. apply in_app_l; auto. contradict h4. rewrite (count_occ_In h1); omega. unfold switch_eq_dec in h3'. f_equal. erewrite lind_occ_functional. rewrite ih; auto. f_equal. apply proof_irrelevance. reflexivity. reflexivity. reflexivity. Grab Existential Variables. assumption. assumption. Qed. Lemma lind_occ_app_ge_in2 : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T) (pf':In x l') (occ:nat), count_occ pfdt l x <= occ -> let pfa := in_app_r pf' l in lind_occ pfdt (l++l') occ x pfa = length l + lind_occ pfdt l' (occ - count_occ pfdt l x) x pf'. intros T h1 l. induction l as [|a l ih]; simpl. intros; f_equal; auto with arith; apply proof_irrelevance. intros l' x h2 occ h3. rewrite (switch_eq_dec_eq h1 x a). destruct (h1 a x) as [|h4]; subst. unfold switch_eq_dec. assert (h4:occ = S (pred occ)). omega. rewrite h4. destruct in_dec as [h5 | h5]. f_equal. rewrite <- h4. erewrite lind_occ_functional. rewrite ih at 1. f_equal. f_equal. assert (h6:pred occ - count_occ h1 l x = occ - S (count_occ h1 l x)). omega. rewrite h6 at 1. omega. omega. reflexivity. reflexivity. reflexivity. contradict h5; apply in_app_r; auto. unfold switch_eq_dec. f_equal. erewrite lind_occ_functional. rewrite ih. f_equal. assumption. reflexivity. reflexivity. reflexivity. Qed. Lemma lind_app_cons_nin_eq : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T), ~In a l -> l <> nil -> let pfin := in_app_r (in_eq _ _) l in lind pfdt (l++a::l') a pfin = length l. intros T h1 l. induction l as [|c l ih]. intros l' a h2 h3. contradict h3; auto. intros l' a h2 h3. unfold lind. simpl. destruct (h1 a c) as [h4 | h4]. subst. contradict h2; left; auto. f_equal. assert (h5:~In a l). contradict h2. right; auto. specialize (ih l' _ h5). destruct (nil_dec' l) as [|h6]; subst. gen0. rewrite app_nil_l. intro h6. simpl. destruct (h1 a a) as [|h7]; auto. contradict h7; auto. specialize (ih h6). simpl in ih. erewrite lind_functional in ih. unfold lind in ih. rewrite ih at 1. reflexivity. reflexivity. Qed. Lemma lind_app_nin : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T) (pf: ~In a l) (pf': In a l'), let pfin := in_app_r pf' l in lind pfdt (l++l') a pfin = length l + lind pfdt l' a pf'. intros T h1 l l' a h2 h3. simpl. destruct (nil_dec' l) as [|ha]; subst. gen0. rewrite app_nil_l. intros; simpl; f_equal; apply proof_irrelevance. revert ha h2. revert l. revert h3. revert a. induction l' as [|a' l' ih]. intros; contradiction. intros a h3 l ha h2. simpl. destruct h3 as [|h3]; subst. unfold lind at 2. simpl. destruct (h1 a a) as [|h3]. rewrite plus_0_r. gen0. rewrite <- (app_sing l' a). pose proof (app_assoc l (a::nil) l') as h4. simpl in h4. simpl. rewrite h4. intro h5. erewrite lind_functional. rewrite lind_app_in1 at 1. rewrite lind_app_cons_nin_eq at 1; auto. reflexivity. contradict h3; auto. unfold lind at 2. simpl. destruct (h1 a a') as [|h4]; subst. rewrite plus_0_r. erewrite lind_functional. rewrite lind_app_cons_nin_eq at 1; auto. reflexivity. specialize (ih a h3 (l++a'::nil)). hfirst ih. intro h5. apply app_eq_nil in h5. destruct h5; discriminate. specialize (ih hf). hfirst ih. intro h6. rewrite in_app_iff in h6. destruct h6 as [| h6]; try contradiction. destruct h6; subst. contradict h4; auto. contradiction. specialize (ih hf0). revert ih. Gen0. rewrite <- app_assoc. simpl. intros h6 ih. erewrite lind_functional. rewrite ih at 1. rewrite app_length. simpl. unfold lind. rewrite <- plus_assoc. rewrite (plus_comm 1). rewrite S_compat. f_equal. f_equal. f_equal. apply proof_irrelevance. reflexivity. Qed. Lemma skipn_lind_app_cons : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T) (pf:In a (l++a::l')), ~In a l -> skipn (lind pfdt (l++a::l') a pf) (l++a::l') = a::l'. intros T h1 l l' a h2 h3. erewrite lind_functional. rewrite lind_app_nin at 1. rewrite skipn_plus. rewrite lind_cons_same. rewrite skipn0. rewrite skipn_length_app; auto. assumption. reflexivity. Grab Existential Variables. left; auto. Qed. Lemma skipn_lind_sub_list_not_nil : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a:T) (pf:sub_list (a::l) l'), let pfin := sub_list_in (a::l) l' a (in_eq _ _) pf in let n := lind pfdt _ _ pfin in skipn n l' <> nil. intros T h1 l l' a h2. simpl. inversion h2 as [l1 l2]; subst. destruct (in_dec h1 a l1) as [h3 | h3]. pose proof (lind_app_in1 h1 l1 ((a::l) ++ l2) a h3) as h4. simpl in h4. erewrite lind_functional. rewrite h4 at 1. rewrite skipn_app. intro h5. apply app_eq_nil in h5. destruct h5 as [? h5]. apply app_eq_nil in h5. destruct h5; discriminate. apply lt_le_weak. apply lt_lind. reflexivity. destruct (nil_dec' l2) as [| h4]; subst. rewrite app_nil_r at 3. match goal with |- skipn (lind _ _ _ ?pf) _ <> _ => generalize pf end. rewrite app_nil_r. intro h4. rewrite skipn_lind_app_cons. intro; discriminate. assumption. pose proof (lind_app_nin h1 l1 ((a::l) ++ l2) a h3) as h5. hfirst h5. apply in_app_l; left; auto. specialize (h5 hf). simpl in h5. erewrite lind_functional. rewrite h5 at 1. rewrite skipn_length_app_plus. erewrite lind_functional. rewrite lind_cons_same. rewrite skipn0. intro h6. apply app_eq_nil in h6. destruct h6; discriminate. reflexivity. reflexivity. Grab Existential Variables. left; auto. Qed. Lemma count_occ_firstn_eq0_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), count_occ pfdt (firstn n l) x = 0 <-> (forall pf:In x l, n <= lind pfdt l x pf). intros T h1 l x n. revert l x. induction n as [|n ih]; simpl. intros l x; split; intros; auto with arith. intros l x; split; intro h2. intro hi. destruct l as [|a l]. contradiction. simpl in h2. destruct (h1 a x) as [|h3]; subst. discriminate. rewrite ih in h2. unfold lind; simpl. apply neq_sym in h3. destruct (h1 x a) as [|h5]. contradiction. apply le_n_S. destruct hi as [|hi]; subst. contradict h5; auto. specialize (h2 hi). unfold lind in h2. erewrite lind_occ_functional. apply h2. reflexivity. reflexivity. reflexivity. destruct l as [|a l]; simpl; auto. destruct (h1 a x) as [|h3]; subst. specialize (h2 (in_eq _ _)). unfold lind in h2. simpl in h2. destruct (h1 x x) as [|h4]. omega. pose proof (h4 eq_refl); contradiction. rewrite ih. intro hi. specialize (h2 (in_cons _ _ _ hi)). unfold lind in h2. simpl in h2. destruct (h1 x a) as [|h5]; subst. contradict h3; auto. apply le_S_n in h2. unfold lind. erewrite lind_occ_functional. apply h2. reflexivity. reflexivity. reflexivity. Qed. Lemma count_occ_skipn_eq0_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), count_occ pfdt (skipn n l) x = 0 <-> (forall (pf:In x l) (occ:nat), occ < count_occ pfdt l x -> lind_occ pfdt l occ x pf < n). intros T h1 l x n. revert x l. induction n as [|n ih]; simpl. intros x l; split; intro h2. intros h3 occ h4. pose proof h3 as h3'. rewrite (count_occ_In h1) in h3'. omega. destruct (in_dec h1 x l) as [h3 | h3]. specialize (h2 h3 0). hfirst h2. rewrite <- (count_occ_In h1 l x); auto. specialize (h2 hf). omega. rewrite (count_occ_In h1) in h3. omega. intros x l; split; intro h2. intros h3 occ h4. destruct l as [|a l]. simpl in h4. omega. simpl in h4, h3. simpl. rewrite (switch_eq_dec_eq h1 a x) in h4. destruct (h1 x a) as [|h5]; subst. unfold switch_eq_dec in h4. destruct occ as [|occ]; auto with arith. destruct (in_dec h1 a l) as [h5 | h5]. apply lt_n_S. apply ih; auto with arith. auto with arith. unfold switch_eq_dec in h4. destruct h3 as [|h3]; subst. contradict h5; auto. apply lt_n_S. apply ih; auto. destruct l as [|a l]; simpl; auto. rewrite ih. intros h3 occ h4. destruct (h1 x a) as [|h5]; subst. specialize (h2 (in_eq _ _) (S occ)). simpl in h2. destruct (h1 a a) as [|h5]. hfirst h2; auto with arith. specialize (h2 hf). destruct in_dec as [|h6]; auto with arith. pose proof (proof_irrelevance _ i h3); subst; auto with arith. contradiction. pose proof (h5 eq_refl); contradiction. specialize (h2 (in_cons _ _ _ h3)). specialize (h2 occ). simpl in h2. destruct (h1 a x) as [|h6]; subst. contradict h5; auto. specialize (h2 h4). destruct (h1 x a) as [|h7]; subst. contradict h5; auto. apply lt_S_n in h2. erewrite lind_occ_functional. apply h2. reflexivity. reflexivity. reflexivity. Qed. Lemma count_occ_skipn_S_lind : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), count_occ pfdt (skipn (S (lind pfdt l x pf)) l) x = pred (count_occ pfdt l x). intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2. destruct h2 as [|h2]; subst. unfold lind; simpl. rewrite eq_dec_same at 1 2; auto. destruct (h1 a x) as [|h3]; subst. erewrite lind_functional. rewrite lind_hd' at 1; simpl; auto. reflexivity. erewrite <- ih. repeat f_equal. unfold lind; simpl. destruct (h1 x a) as [|h4]; subst. contradict h3; auto. f_equal. f_equal. apply proof_irrelevance. Grab Existential Variables. assumption. apply cons_neq_nil. Qed. Lemma lt_lind_S_occ : forall {T:Type} (h1:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (occ:nat), lind h1 l x pf < lind_occ h1 l (S occ) x pf \/ 1 = count_occ h1 l x. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2 occ. destruct (h1 x a) as [|h3]; subst. destruct (h1 a a) as [|h4]. destruct in_dec as [h3 | h3]. unfold lind. simpl. destruct (h1 a a) as [|h4]. left; auto with arith. contradict h4; auto. right. simpl. rewrite (count_occ_In h1) in h3. omega. contradict h4; auto. unfold lind. simpl. destruct h2 as [|h2]; subst. contradict h3; auto. specialize (ih _ h2). destruct (h1 a x) as [|he]; subst. contradict h3; auto. destruct (h1 x a) as [|he']; subst. left; auto with arith. match goal with |- S (lind_occ _ _ _ _ ?pf) < _ \/ _ => assert (h5: h2 = pf) end. apply proof_irrelevance. rewrite <- h5; auto. fold (lind h1 l x h2). specialize (ih occ). destruct ih as [h4 | h4]. left. apply lt_n_S. pose proof (proof_irrelevance _ he' h3) as h6. rewrite <- h6. rewrite <- h5; auto. right; auto. Qed. Lemma lt_lind_occ_S : forall {T:Type} (h1:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (occ:nat), S occ < count_occ h1 l x -> lind_occ h1 l occ x pf < lind_occ h1 l (S occ) x pf. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2 occ h3. destruct h2 as [|h2]; subst. destruct (h1 x x) as [h4 | h4]. apply lt_S_n in h3. pose proof (count_occ_in h1 _ _ _ h3) as h5. destruct occ as [|occ]. destruct in_dec as [| h6]; auto with arith. contradiction. destruct in_dec as [h6 | h6]. apply lt_n_S. apply ih; auto. contradiction. apply lt_n_S. apply ih; auto. rewrite switch_eq_dec_eq in h3. destruct (h1 x a) as [|h4]; subst. unfold switch_eq_dec in h3. apply lt_S_n in h3. destruct in_dec as [|h4]. destruct occ as [|occ]; auto with arith. contradiction. unfold switch_eq_dec in h3. apply lt_n_S. apply ih; auto. Qed. Lemma lt_lind_occ_lt : forall {T:Type} (h1:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (occ occ':nat), occ < occ' -> occ' < count_occ h1 l x -> lind_occ h1 l occ x pf < lind_occ h1 l occ' x pf. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2 occ occ' h3 h4. destruct h2 as [|h2]; subst. destruct (h1 x x) as [h5 | h5]. destruct occ' as [|occ']. apply lt_n_0 in h3; contradiction. apply lt_S_n in h4. destruct in_dec as [h6 | h76]. destruct occ as [|occ]; auto with arith. apply count_occ_in in h4; contradiction. contradict h5; auto. rewrite switch_eq_dec_eq in h4. destruct (h1 x a) as [|h5]; subst. simpl in h4. destruct occ' as [|occ']. apply lt_n_0 in h3. contradiction. destruct occ as [|occ]. destruct in_dec as [h5 | h5]; auto with arith. contradiction. apply lt_S_n in h3. apply lt_S_n in h4. destruct in_dec as [h5 | h5]. apply lt_n_S. apply ih; auto. contradiction. simpl in h4. apply lt_n_S. apply ih; auto. Qed. Lemma in_list_of_lists_seqs_iff : forall {T:Type} (ll:faml T) (l:list T), In l (list_of_lists_seqs ll) <-> (length l = length ll /\ (forall (i:nat) (pfi:i < length l), exists pfi':i < length ll, In (nth_lt l i pfi) (nth_lt ll i pfi'))). intros T ll. induction ll as [|al ll ih]; simpl. intros l. split; intro h1. destruct h1 as [|h1]; subst; try contradiction. simpl. split; auto. intros ? h1. pose proof (lt_n_0 _ h1). contradiction. destruct l as [|a l]. left; auto. destruct h1 as [h1]. simpl in h1. discriminate. intro l. split. intros h1. rewrite in_map_iff in h1. destruct h1 as [p [h1 h2]]. subst. simpl. destruct p as [x l]. simpl. rewrite in_prod_iff in h2. destruct h2 as [h2 h3]. rewrite ih in h3. destruct h3 as [h3 h4]. split. f_equal; auto. intros i h5. rewrite h3 in h5. exists h5. destruct i as [|i]. assumption. apply lt_S_n in h5. rewrite <- h3 in h5. specialize (h4 _ h5). destruct h4 as [h4 h6]. rewrite nth_lt_compat in h6. rewrite (nth_lt_compat ll) in h6. erewrite nth_indep, (nth_indep ll). apply h6. assumption. assumption. intro h1. destruct h1 as [h1 h2]. destruct l as [|a l]. simpl in h1. discriminate. simpl in h1. apply S_inj in h1. specialize (ih l). hr ih. split; auto. intros i h3. specialize (h2 (S i)). hfirst h2. simpl. apply lt_n_S; auto. specialize (h2 hf). destruct h2 as [h2 h4]. simpl in h2. apply lt_S_n in h2. exists h2. simpl in h4. erewrite nth_lt_compat, (nth_lt_compat ll). erewrite nth_indep, (nth_indep ll). apply h4. assumption. assumption. rewrite <- ih in hr. rewrite in_map_iff. exists (a, l). simpl. split; auto. rewrite in_prod_iff. split; auto. specialize (h2 0 (lt_0_Sn _)). simpl in h2. destruct h2; auto. Qed. Lemma in_list_of_lists_seqs_length : forall {T:Type} (ll:faml T) (l:list T), In l (list_of_lists_seqs ll) -> length l = length ll. intros T ll l h1. rewrite in_list_of_lists_seqs_iff in h1. destruct h1; auto. Qed. Lemma in_list_of_lists_seqs_lt : forall {T:Type} (ll:faml T) (l:list T) (i:nat), In l (list_of_lists_seqs ll) -> i < length l -> i < length ll. intros; rewrite <- (in_list_of_lists_seqs_length ll l); auto. Qed. Lemma coarse_list_in_list_of_lists_seqs : forall {T:Type} (ll:faml T) (l:list T) (v:bv) (pfeq:length l = length v) (pfi:In l (list_of_lists_seqs ll)), let pfeq' := eq_trans (eq_sym (in_list_of_lists_seqs_length _ _ pfi)) pfeq in let fl := map in In (mask_to_coarse_list pfeq) (list_of_lists_seqs (mask_to_coarse_list pfeq')). intros T ll l v h1 h2; simpl. rewrite in_list_of_lists_seqs_iff. split. do 2 rewrite length_mask_to_coarse_list; auto. intros i h3. ex_goal. rewrite length_mask_to_coarse_list. rewrite length_mask_to_coarse_list in h3; auto. exists hex. pose proof h2 as h2'. rewrite in_list_of_lists_seqs_iff in h2'. destruct h2' as [h4 h5]. pose proof h3 as h3'. rewrite length_mask_to_coarse_list in h3'. specialize (h5 (nth_lt _ _ h3')). hfirst h5. rewrite h1. apply in_bv1s_lt. apply in_nth_lt. specialize (h5 hf). destruct h5 as [h5 h6]. unfold mask_to_coarse_list. erewrite nth_lt_map_in_dep. erewrite nth_lt_map_in_dep. erewrite nth_lt_functional3, (nth_lt_functional3 ll). apply h6. Qed. Lemma eq_hd_lind_occ_tl : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x (tl l)) (occ:nat), x = hd x l -> let pf' := in_tl _ _ pf in lind_occ pfdt (tl l) occ x pf = pred (lind_occ pfdt l (S occ) x pf'). intros T h1 l. destruct l as [|a l]; simpl. intros; contradiction. intros x h2 occ ?; subst. destruct (h1 a a) as [h3 | h3]. destruct in_dec as [h4 | h4]. simpl. unfold lind; f_equal; apply proof_irrelevance. contradiction. contradict h3; auto. Qed. Lemma neq_hd_lind_occ_tl : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x (tl l)) (occ:nat), hd x l <> x -> let pf' := in_tl _ _ pf in lind_occ pfdt (tl l) occ x pf = pred (lind_occ pfdt l occ x pf'). intros T h1 l. destruct l as [|a l]; simpl. intros; contradiction. intros x h2 h3 hn. destruct (h1 x a) as [|h4]; subst. contradict h3; auto. simpl. f_equal. apply proof_irrelevance. Qed. Lemma eq_hd_lind_tl : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x (tl l)), x = hd x l -> let pf' := in_tl _ _ pf in lind pfdt (tl l) x pf = pred (lind_occ pfdt l 1 x pf'). unfold lind; intros; apply eq_hd_lind_occ_tl; auto. Qed. Lemma neq_hd_lind_tl : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x (tl l)), hd x l <> x -> let pf' := in_tl _ _ pf in lind pfdt (tl l) x pf = pred (lind pfdt l x pf'). unfold lind; intros; apply neq_hd_lind_occ_tl; auto. Qed. Lemma in_tl_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), In x (tl l) <-> (exists pf:In x l, x = hd' l (in_not_nil _ _ pf) -> count_occ pfdt l x > 1). intros T h1 l x; split; intro h2. exists (in_tl _ _ h2). intro h3. destruct l as [|a l]. contradict h2. simpl in h2, h3. subst. simpl. destruct (h1 a a) as [h3 | h3]. apply lt_n_S. rewrite <- (count_occ_In h1 l a); auto. contradict h3; auto. destruct h2 as [h2 h3]. pose proof h2 as h2'. rewrite (hd_compat' _ (in_not_nil _ _ h2)) in h2'. destruct h2' as [h4 | h4]. symmetry in h4. specialize (h3 h4). destruct l as [|a l]. simpl in h3. contradiction. simpl in h2, h3, h4. simpl. destruct (h1 a x) as [h5 | h5]. apply lt_S_n in h3. rewrite (count_occ_In h1); auto. subst. contradict h5; auto. assumption. Qed. Lemma in_tl_iff' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), In x (tl l) <-> (exists (n:nat) (pf:S n < length l), nth_lt l (S n) pf = x). intros T h1 l x; split; intro h2. pose proof (lt_lind h1 _ _ h2) as h3. exists (lind h1 (tl l) x h2). ex_goal. rewrite length_tl in h3; omega. exists hex. rewrite nth_lt_S. erewrite nth_lt_functional3. rewrite lind_compat; auto. destruct h2 as [n [h2 h3]]; subst. rewrite nth_lt_S. apply in_nth_lt. Qed. Lemma length_inv_im_in_dep_inj : forall {T U:Type} (pfdu:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T), NoDup l -> forall (u:U) (f:fun_in_dep l U), inj_dep f -> let l' := inv_im_in_dep pfdu l u f in l' <> nil -> length l' = 1. intros T U h1 h1' l hnd u f h2 l' h3. destruct (nil_dec' (tl l')) as [h4 | h4]. rewrite (hd_compat' _ h3). rewrite h4; auto. pose proof (in_hd' _ h3) as h5. pose proof (in_hd' _ h4) as h6. rewrite (in_tl_iff h1) in h6. destruct h6 as [h6 h7]. unfold l' in h5, h6. pose proof h6 as h6'. rewrite in_inv_im_in_dep_iff in h5, h6'. destruct h5 as [h5 h5'], h6' as [h8 h8']. rewrite <- h5' in h8'. apply h2 in h8'. fold l' in h8'. pose proof (proof_irrelevance _ h3 (in_not_nil l' (hd' (tl l') h4) h6)) as h9. rewrite <- h9 in h7. specialize (h7 h8'). rewrite count_occ_no_dup in h7; auto. omega. unfold l'. apply no_dup_inv_im_in_dep; auto. Qed. Lemma in_firstn_iff' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (m:nat), In x (firstn m l) <-> (exists pf:In x l, lind pfdt l x pf < m). intros T h1 l. induction l as [|a l ih]; simpl. intros; rewrite firstn_nil. simpl. split; intro; auto; try contradiction. destruct H; auto. intros x m; split; intro h2. destruct m as [|m]. simpl in h2. contradiction. simpl in h2. destruct h2 as [|h2]; subst. exists (or_introl eq_refl). unfold lind. simpl. destruct (h1 x x); auto with arith. contradict n; auto. rewrite ih in h2. destruct h2 as [h2 h3]. exists (or_intror h2). unfold lind; simpl. destruct (h1 x a) as [|h4]; auto with arith. apply lt_n_S. erewrite lind_occ_functional. unfold lind in h3. apply h3. reflexivity. reflexivity. reflexivity. destruct h2 as [h2 h3]. destruct h2 as [|h2]; subst. unfold lind in h3. simpl in h3. destruct (h1 x x) as [|h4]. rewrite (S_pred _ _ h3). simpl. left; auto. pose proof (h4 eq_refl); contradiction. unfold lind in h3. simpl in h3. destruct (h1 x a) as [|h4]; subst. rewrite (S_pred _ _ h3). simpl. left; auto. rewrite (S_pred _ _ h3). simpl. right. rewrite ih. exists h2. let P := type of h3 in match P with S (lind_occ _ _ _ _ ?pf) < _ => pose pf as hyp end. fold hyp in h3. fold (lind h1 l x hyp) in h3. pose proof (proof_irrelevance _ h2 hyp) as h5. rewrite h5. omega. Qed. Lemma in_nth_lt_skipn : forall {T:Type} (l:list T) (n:nat) (pflt:S n <= length l), In (nth_lt l n pflt) (skipn n l). intros T l n. revert l. induction n as [|n ih]; simpl. intros; apply in_nth_lt. intros l h1. destruct l as [|a l]. simpl in h1. omega. simpl in h1. pose proof h1 as h1'. apply le_S_n in h1'. simpl. erewrite nth_indep. rewrite <- nth_lt_compat. apply ih. omega. Grab Existential Variables. omega. Qed. Lemma in_skipn_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (m:nat), In x (skipn m l) <-> (exists (pf:In x l) (occ:nat), occ < count_occ pfdt l x /\ lind_occ pfdt l occ x pf >= m). intros T h1 l x m; split; intro h2. pose proof (in_skipn _ _ _ h2) as h3. exists h3. revert h3 h2. revert x l. induction m as [|m ih]; simpl. intros; auto with arith. exists 0; auto with arith. split; auto with arith. rewrite (count_occ_In h1) in h2; auto with arith. intros x l h2. destruct l as [|a l]. contradiction. intro h3. pose proof (in_skipn _ _ _ h3) as h4. specialize (ih _ _ h4 h3). destruct ih as [occ ih]. destruct h2 as [|h2]; subst. exists (S occ). simpl. destruct (h1 x x) as [h5 | h5]. destruct in_dec as [h6 | h6]. pose proof (proof_irrelevance _ h4 h6); subst; auto with arith. destruct ih as [ih ih']. split; auto with arith. contradiction. contradict h5; auto. destruct ih as [ih ih']. destruct (h1 a x) as [|h5]; subst. exists (S occ). simpl. destruct (h1 x x) as [h1x | h1x]. destruct in_dec as [h5 | h5]. pose proof (proof_irrelevance _ h4 h5); subst. split; auto with arith. contradiction. contradict h1x; auto. exists occ; simpl. rewrite (switch_eq_dec_eq h1 a x). destruct (h1 x a) as [|h6]; subst. contradict h5; auto. unfold switch_eq_dec. pose proof (proof_irrelevance _ h4 (neq_in_cons l x a h6 (or_intror h2))); subst. split; auto with arith. destruct h2 as [h2 [occ [h3 h4]]]. revert h2 occ h3 h4. revert x l. induction m as [|m ih]; simpl. intros; eapply count_occ_in. apply h3. intros x l h2 occ h3 h4. destruct l as [|a l]. contradiction. destruct h2 as [|h2]; subst. rewrite count_occ_cons_eq in h3; auto. destruct occ as [|occ]. simpl in h4. destruct (h1 x x) as [h5 | h5]. omega. pose proof (h5 eq_refl); contradiction. apply lt_S_n in h3. simpl in h4. destruct (h1 x x) as [h5 | h5]. destruct in_dec as [h6 | h6]. eapply ih. apply h3. apply le_S_n in h4. apply h4. omega. pose proof (h5 eq_refl); contradiction. simpl in h3, h4. rewrite switch_eq_dec_eq in h3. destruct (h1 x a) as [|h5]; subst. unfold switch_eq_dec in h3. destruct occ as [|occ]. omega. destruct in_dec as [h5 | h5]. apply lt_S_n in h3. apply le_S_n in h4. eapply ih. apply h3. apply h4. contradiction. unfold switch_eq_dec in h3. apply le_S_n in h4. eapply ih. apply h3. apply h4. Qed. Lemma in_skipn_le : forall {T:Type} (m n:nat) (l:list T) (x:T), m <= n -> In x (skipn n l) -> In x (skipn m l). intros T m n. revert m. induction n as [|n ih]; simpl. intros m l x h1. apply le0 in h1; subst; simpl; auto. intros m l x h1 h2. destruct l as [|a l]; simpl. contradiction. destruct m as [|m]; simpl. right. apply in_skipn in h2; auto. apply le_S_n in h1; auto. Qed. Lemma in_firstn_le : forall {T:Type} (m n:nat) (l:list T) (x:T), m <= n -> In x (firstn m l) -> In x (firstn n l). intros T m n. revert m. induction n as [|n ih]; simpl. intros m l x h1. apply le0 in h1; subst; simpl; auto. intros m l x h1 h2. destruct l as [|a l]; simpl. rewrite firstn_nil in h2. contradiction. destruct m as [|m]; simpl. simpl in h2. contradiction. simpl in h2; auto. destruct h2 as [|h2]. left; auto. right. apply le_S_n in h1. specialize (ih _ _ _ h1 h2); auto. Qed. Lemma nin_firstn_le_lind : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (m:nat) (pf:In x l), ~In x (firstn m l) -> m <= lind pfdt l x pf. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x m h2. destruct h2 as [|h2]; subst. intro h2. destruct m as [|m]; auto with arith. unfold lind. simpl. destruct (h1 x x) as [|h3]. contradict h2. simpl. left; auto. contradict h3; auto. intro h3. unfold lind. simpl. destruct (h1 x a) as [|h4]; subst. destruct m as [|m]; auto. contradict h3. simpl. left; auto. specialize (ih x (pred m) h2). hfirst ih. contradict h3. destruct m as [|m]. simpl in h3. contradiction. simpl. simpl in h3. right; auto. specialize (ih hf). pose proof (proof_irrelevance _ h2 (neq_in_cons _ _ a h4 (or_intror h2))) as h6. rewrite <- h6. unfold lind in ih; omega. Qed. Lemma in_firstn_nin_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (i j:nat), i <= length l -> In x (firstn i l) -> ~In x (firstn j l) -> j < i. intros T hdt l x i j hle h1 h2. rewrite in_firstn_iff in h1. destruct h1 as [k [h3 h4]]. rewrite <- h4 in h2. eapply (nin_firstn_le_lind hdt _ _ _ (in_nth_lt _ _ _)) in h2. pose proof (le_lind_nth_lt hdt l k (lt_le_trans _ _ _ h3 hle)) as h5. eapply le_lt_trans. eapply le_trans. apply h2. apply h5. assumption. Qed. Lemma lind_occ_firstn : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (m:nat) (pf:In x (firstn m l)) (occ:nat), occ < count_occ pfdt (firstn m l) x -> lind_occ pfdt (firstn m l) occ x pf = lind_occ pfdt l occ x (in_firstn _ _ _ pf). intros T h1 l. induction l as [|a l ih]; simpl. intros ? ? h2. pose proof h2 as h2'. rewrite firstn_nil in h2'. contradiction. intros x m h2 occ hlt. destruct m as [|m]. simpl in h2. contradiction. simpl. simpl in h2, hlt. rewrite switch_eq_dec_eq in hlt. destruct (h1 x a) as [h4 | h4]; subst. unfold switch_eq_dec in hlt. destruct occ as [|occ]; auto. apply lt_S_n in hlt; auto. destruct in_dec as [h3 | h3]. destruct in_dec as [h4 | h4]. f_equal. rewrite ih; auto; f_equal; apply proof_irrelevance. contradict h4. apply in_firstn in h3; auto. contradict h3. rewrite (count_occ_In h1); omega. unfold switch_eq_dec in hlt. destruct h2 as [|h2]; subst. contradict h4; auto. f_equal. rewrite ih; auto. f_equal. apply proof_irrelevance. Qed. Lemma ex_lind_occ_nth_lt_firstn : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m n:nat) (pfle:m <= length l) (pflt:n < length l) (pfnm:n < m) (pfin:In (nth_lt l n pflt) (firstn m l)), exists (occ:nat), occ < count_occ pfdt (firstn m l) (nth_lt l n pflt) /\ lind_occ pfdt (firstn m l) occ (nth_lt l n pflt) pfin = n. intros T h1 l. induction l as [|a l ih]; simpl. intros; omega. intros m n h2 h3 hnm h4. destruct n as [|n]; simpl. destruct m as [|m]; simpl. simpl in h4. contradiction. apply le_S_n in h2. simpl in h4. destruct (h1 a a) as [h5 | h5]. destruct in_dec as [h6 | h6]. exists 0. split; auto with arith. exists 0. split; auto with arith. contradict h5; auto. apply lt_S_n in h3. destruct m as [|m]. simpl in h4. contradiction. apply le_S_n in h2. simpl in h4. simpl. rewrite switch_eq_dec_eq. destruct (h1 (nth n l a) a) as [h5 | h5]. unfold switch_eq_dec. destruct in_dec as [h6 | h6]. apply le_S_n in hnm. specialize (ih _ _ h2 h3 hnm). hfirst ih. rewrite nth_lt_compat. erewrite nth_indep. apply h6. assumption. specialize (ih hf). destruct ih as [occ [h7 h8]]. exists (S occ). split. apply lt_n_S. erewrite nth_lt_compat in h7. erewrite nth_indep in h7. apply h7. assumption. f_equal. rewrite <- h8. apply lind_occ_functional; auto. rewrite nth_lt_compat. apply nth_indep. assumption. apply lt_S_n in hnm. contradict h6. rewrite in_firstn_iff. exists n, hnm. rewrite nth_lt_compat. apply nth_indep. assumption. apply lt_S_n in hnm. unfold switch_eq_dec. destruct h4 as [h4 | h4]. pose proof h5 as h5'. rewrite h4 in h5' at 2. contradict h5'; auto. specialize (ih m n h2 h3 hnm). hfirst ih. rewrite nth_lt_compat. erewrite nth_indep. apply h4. omega. specialize (ih hf). destruct ih as [occ [h6 h7]]. exists occ. split. rewrite nth_lt_compat in h6. erewrite nth_indep in h6. apply h6. omega. f_equal. rewrite <- h7. apply lind_occ_functional; auto. erewrite nth_lt_compat. apply nth_indep. assumption. Grab Existential Variables. assumption. Qed. Lemma count_occ_eq1_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), count_occ pfdt l x = 1 <-> exists pf:In x l, ~ In x (skipn (S (lind pfdt l x pf)) l). intros T h1 l x; split; intro h2. pose proof (count_occ_In h1 l x) as h3. rewrite h2 in h3. let P := type of h3 in match P with _ <-> ?hyp => assert (h4:hyp) end. auto with arith. rewrite <- h3 in h4. exists h4. intro h5. rewrite (in_skipn_iff h1 l x) in h5. destruct h5 as [h5 [occ [h6 h7]]]. rewrite h2 in h6. assert (h8:occ = 0). omega. subst. unfold lind in h7. pose proof (proof_irrelevance _ h4 h5); subst. omega. destruct h2 as [h2 h3]. rewrite (in_skipn_iff h1 l x) in h3. pose proof (count_occ_In h1 l x) as h4. pose proof h2 as h2'. rewrite h4 in h2'. rewrite (S_pred _ _ h2'). f_equal. rewrite <- not_not0_iff. contradict h3. exists h2. exists 1. split. omega. pose proof (lt_lind_S_occ h1 l x h2 0) as h5. destruct h5 as [h5 | h5]. omega. rewrite <- h5 in h3. simpl in h3. contradict h3; auto. Qed. Lemma nin_skipn_S_lind_occ_pred_count_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), ~ In x (skipn (S (lind_occ pfdt l (pred (count_occ pfdt l x)) x pf)) l). intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2. destruct h2 as [|h2]; subst. destruct (h1 x x) as [h3 | h3]. simpl. destruct (in_dec h1 x l) as [h4 | h4]. pose proof h4 as h4'. rewrite (count_occ_In h1) in h4'. rewrite (S_pred _ _ h4'). apply ih. rewrite (count_occ_In h1) in h4. rewrite (not_lt0_iff (count_occ h1 l x)) in h4 at 1. rewrite h4. simpl. intro h5. rewrite (count_occ_In h1) in h5. omega. contradict h3; auto. destruct (h1 x a) as [|h3]; subst. destruct (h1 a a) as [h3 | h3]. simpl. rewrite (count_occ_In h1) in h2. rewrite (S_pred _ _ h2). destruct in_dec as [h4 | h4]. apply ih. simpl; auto. contradict h3; auto. destruct (h1 a x) as [|h4]; subst. contradict h3; auto. apply ih. Qed. Lemma lind_occ_S : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (occ:nat), S occ < count_occ pfdt l x -> let l' := skipn (S (lind pfdt l x pf)) l in forall (pf':In x l'), let lo := lind_occ pfdt l' occ x pf' in lind_occ pfdt l (S occ) x pf = lind pfdt l x pf + S lo. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2 occ h3. destruct h2 as [|h2]; subst. destruct (h1 x x) as [|h4]; subst. apply lt_S_n in h3. destruct in_dec as [h4 | h4]. erewrite lind_functional. rewrite lind_hd'. simpl. f_equal. intros; repeat f_equal; apply proof_irrelevance. simpl; auto. contradict h4. apply (count_occ_in h1) in h3; auto. contradict h4; auto. intros h4. destruct (h1 x a) as [|h5]; subst. destruct (h1 a a) as [h10 | h10]. apply lt_S_n in h3. destruct in_dec as [h5 | h5]. generalize h4. erewrite lind_functional. rewrite lind_hd'. simpl. intro h6. repeat f_equal; apply proof_irrelevance. simpl. reflexivity. contradiction. contradict h10; auto. unfold lind. simpl. unfold lind in h4. simpl in h4. destruct (h1 a x) as [|h6]; subst. contradict h5; auto. assert (h7:h1 x a = right (x = a) h5). destruct (h1 x a). contradiction. f_equal; apply proof_irrelevance. rewrite h7 at 1. pose proof h4 as h4'. rewrite h7 in h4' at 1. simpl in h4'. rewrite plus_S_shift_r. f_equal. specialize (ih _ _ _ h3 h4'). rewrite ih at 1. rewrite plus_S_shift_r, plus_S_shift_l. f_equal. f_equal. generalize h4. rewrite h7. intro h8. f_equal. apply proof_irrelevance. Grab Existential Variables. apply cons_neq_nil. apply cons_neq_nil. Qed. Lemma in_firstn_S_lind : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), In x (firstn (S (lind pfdt l x pf)) l). intros T h1 l x h2. erewrite firstn_S_eq. apply in_app_r. simpl. left. apply lind_compat. Qed. Lemma lind_occ_lt_mono : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l) (occ occ':nat), occ < occ' -> occ' < count_occ pfdt l x -> lind_occ pfdt l occ x pfin < lind_occ pfdt l occ' x pfin. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2 occ occ' h3 h4. destruct h2 as [|h2]; subst. destruct (h1 x x) as [h5 | h5]. destruct occ' as [|occ']. omega. destruct occ as [|occ]. destruct in_dec as [h6 | h6]; auto with arith. contradict h6. apply lt_S_n in h4. apply count_occ_in in h4; auto. destruct in_dec as [h6 | h6]; auto. apply lt_n_S. apply ih; auto with arith. contradict h6. apply lt_S_n in h4. apply count_occ_in in h4; auto. contradict h5; auto. rewrite switch_eq_dec_eq in h4. destruct (h1 x a) as [|h5]; subst. simpl in h4. destruct occ' as [|occ']. omega. destruct occ as [|occ]. destruct in_dec as [h5 | h5]; auto with arith. contradiction. destruct in_dec as [h5 | h5]. apply lt_n_S. apply ih; auto with arith. contradiction. simpl in h4. apply lt_n_S. apply ih; auto. Qed. Lemma lind_occ_plus_occ_in_skip : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l) (occ occ':nat) (pflt'':occ + (S occ') < count_occ pfdt l x), let lo := lind_occ pfdt l occ x pfin in forall (pfins:In x (skipn (S lo) l)), let lo' := lind_occ pfdt (skipn (S lo) l) occ' x pfins in lind_occ pfdt l (occ + (S occ')) x pfin = lo + (S lo'). intros T h1 l x h2 occ. revert h2. revert l x. induction occ as [|occ ih]; simpl. intros l x h2 occ' h3 hs. destruct l as [|a l]. contradiction. simpl. simpl in h2, h3, hs. destruct h2 as [|h2]; subst. rewrite (eq_dec_same h1) at 1 2. destruct (h1 x x) as [h5 | h5]. destruct in_dec as [h6 | h6]. simpl. repeat f_equal; apply proof_irrelevance. simpl in hs. contradiction. pose proof (h5 eq_refl); contradiction. rewrite (switch_eq_dec_eq h1) in h3. destruct (h1 x a) as [h4 | h4]; subst. simpl in h3, hs. destruct in_dec as [h8 | h8]; simpl. repeat f_equal; apply proof_irrelevance. contradiction. simpl in h3. rewrite plus_S_shift_r. f_equal. simpl. rewrite <- plus_S_shift_r. fold (lind h1 l x (neq_in_cons l x a h4 (or_intror h2))). erewrite lind_occ_functional. rewrite lind_occ_S at 1. f_equal. f_equal. assumption. reflexivity. reflexivity. reflexivity. intros l x h2 occ' h3 h5. destruct l as [|a l]. contradiction. simpl in h2, h3, h5. simpl. destruct h2 as [|h2]; subst. destruct (h1 x x) as [h6 | h6]. apply lt_S_n in h3. destruct in_dec as [h7 | h7]. specialize (ih l x h7 occ' h3). specialize (ih h5). simpl in ih. rewrite ih. rewrite plus_S_shift_l. reflexivity. simpl in h5. contradiction. pose proof (h6 eq_refl); contradiction. rewrite switch_eq_dec_eq in h3. destruct (h1 x a) as [h6 |]; subst. simpl in h3. apply lt_S_n in h3. destruct in_dec as [h6 | h6]. pose proof (proof_irrelevance _ h2 h6); subst. specialize (ih l a h6 _ h3 h5). simpl in ih. rewrite ih. rewrite plus_S_shift_l. reflexivity. simpl in h5. contradiction. simpl in h3. rewrite plus_S_shift_l. f_equal. pose proof (lind_occ_S h1 l x (neq_in_cons l x a n (or_intror h2)) (occ + S occ') h3) as h6. pose proof (lind_occ_S h1 l x (neq_in_cons l x a n (or_intror h2)) occ) as h7. hfirst h7. omega. specialize (h7 hf). pose proof (in_skipn_le (S (lind h1 l x (neq_in_cons l x a n (or_intror h2)))) (S (lind_occ h1 l (S occ) x (neq_in_cons l x a n (or_intror h2)))) l x) as h9. hfirst h9. apply le_S. pose proof (lt_lind_S_occ h1 l x (neq_in_cons l x a n (or_intror h2)) occ) as h10. destruct h10 as [h10 | h10]. apply h10. rewrite <- h10 in h3. omega. specialize (h9 hf0 h5). specialize (h6 h9). specialize (h7 h9). rewrite h6 at 1. rewrite h7 at 1. rewrite <- plus_assoc. f_equal. rewrite plus_S_shift_l. f_equal. pose proof (ih _ _ h9 occ). pose proof (count_occ_skipn_S_lind h1 l x (neq_in_cons l x a n (or_intror h2))) as h8. pose proof (ih (skipn (S (lind h1 l x (neq_in_cons l x a n (or_intror h2)))) l) x h9 occ') as ih'. hfirst ih'. rewrite h8. omega. specialize (ih' hf1). assert (h10: In x (skipn (S (lind_occ h1 (skipn (S (lind h1 l x (neq_in_cons l x a n (or_intror h2)))) l) occ x h9)) (skipn (S (lind h1 l x (neq_in_cons l x a n (or_intror h2)))) l))). rewrite (in_skipn_iff h1). exists h9. exists (occ + S occ'). split; auto. apply lind_occ_lt_mono; omega. specialize (ih' h10). rewrite ih'. f_equal. f_equal. generalize h10. rewrite <- skipn_plus. rewrite plus_comm. intro hf2'. generalize h5. erewrite lind_occ_S. intro h5'. rewrite h8 in hf1. pose proof (proof_irrelevance _ h5' hf2'); subst; auto. assumption. Qed. Lemma lind_occ_plus_occ_nin_skip : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l) (occ occ':nat), let lo := lind_occ pfdt l occ x pfin in forall (pfins: ~In x (skipn (S lo) l)), lind_occ pfdt l (occ + (S occ')) x pfin = lo. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2 occ occ' h4. destruct h2 as [|h2]; subst. destruct (h1 x x) as [h5 | h5]. destruct occ as [|occ]. simpl in h4. simpl. destruct in_dec as [h6 | h6]. contradiction. reflexivity. rewrite plus_S_shift_l. destruct in_dec as [h6 | h6]. f_equal. specialize (ih x h6 occ occ' h4); auto. simpl in h4. reflexivity. pose proof (h5 eq_refl); contradiction. destruct (h1 x a) as [|h5]; subst. destruct occ as [|occ]. simpl in h4. contradiction. rewrite plus_S_shift_l. destruct in_dec as [h5 | h5]. f_equal. specialize (ih _ h5 occ occ' h4); auto. reflexivity. f_equal. specialize (ih x (neq_in_cons l x a h5 (or_intror h2)) occ occ' h4); auto. Qed. Lemma lind_occ_terminates : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (occ:nat), occ >= pred (count_occ pfdt l x) -> lind_occ pfdt l occ x pf = lind_occ pfdt l (pred (count_occ pfdt l x)) x pf. intros T h1 l x h2 occ h3. apply le_lt_eq_dec in h3. destruct h3 as [h3 | h3]. apply lt_decompose in h3. destruct h3; subst. rewrite lind_occ_plus_occ_nin_skip; auto. apply nin_skipn_S_lind_occ_pred_count_occ. rewrite h3; auto. Qed. Lemma le_lind_occ : forall {T:Type} (h1:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (occ:nat), lind h1 l x pf <= lind_occ h1 l occ x pf. intros T h1 l x h2 occ. destruct occ as [|occ]. unfold lind; auto. pose proof (lt_lind_S_occ h1 l x h2 occ) as h3. destruct h3 as [| h3]; auto with arith. symmetry in h3. rewrite count_occ_eq1_iff in h3. destruct h3 as [h3 h4]. rewrite (in_skipn_iff h1 l x) in h4. apply not_lt. contradict h4. pose proof (lt_lind_S_occ h1 l x h2 occ) as h5. destruct h5 as [|h5]. omega. assert (h6:S occ >= count_occ h1 l x). omega. rewrite lind_occ_terminates in h4. rewrite <- h5 in h4. simpl in h4. unfold lind in h4. apply lt_irrefl in h4; contradiction. omega. Qed. Lemma in_removelast_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), In x (removelast l) <-> ((x <> last l x -> In x l) /\ (x = last l x -> 1 < count_occ pfdt l x)). intros T h1 l x. destruct (nil_dec' l) as [|h2]; subst. simpl. split; intro h2. contradiction. destruct h2 as [h2 h3]. specialize (h3 eq_refl). omega. split; intro h3. split. intro h4. rewrite (app_removelast_last x). apply in_app_l; auto. assumption. intro h4. rewrite h4. rewrite count_occ_last. apply lt_n_S. rewrite (count_occ_In h1) in h3. rewrite <- h4; auto with arith. assumption. destruct h3 as [h3 h4]. destruct (h1 x (last l x)) as [h5 | h5]. specialize (h4 h5). rewrite (app_removelast_last x h2) in h4. rewrite count_occ_app in h4. simpl in h4. destruct (h1 (last l x) x) as [h6 | h6]. rewrite S_compat in h4. apply lt_S_n in h4. rewrite (count_occ_In h1); auto. rewrite h5 in h6 at 2. contradict h6; auto. specialize (h3 h5). rewrite (app_removelast_last x h2) in h3. apply in_app_or in h3. destruct h3 as [|h3]; auto. destruct h3 as [h3|]. rewrite <- h3 in h5 at 1. contradict h5; auto. contradiction. Qed. Lemma lind_eq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (n:nat) (pflt:n < length l), lind pfdt l x pf = n <-> (nth_lt l n pflt = x /\ ~In x (firstn n l)). intros T h1 l. induction l as [|a l ih]; simpl. intros; omega. intros x h2 n h3. unfold lind. simpl. destruct (h1 x a) as [|h4]; subst. split; intro h4; subst; simpl. split; auto. destruct n as [|n]; auto. destruct h4 as [h4 h5]. contradict h5. simpl. left; auto. split; intro h5. destruct n as [|n]. discriminate. apply S_inj in h5. subst. split. erewrite nth_indep. rewrite <- nth_lt_compat. apply lt_S_n in h3. destruct h2 as [|h2]; subst. pose proof (h4 eq_refl). contradiction. pose (lt_lind h1 _ _ h2) as h5. pose proof (lind_compat h1 l x h2) as h6. unfold lind in h6. pose proof (proof_irrelevance _ h2 (neq_in_cons _ _ a h4 (or_intror h2))) as h7. rewrite h7 in h6. rewrite h6 at 1; auto. apply lt_S_n in h3; auto. intro h5. simpl in h5. destruct h5 as [|h5]; subst. pose proof (h4 eq_refl); contradiction. destruct h2 as [|h2]; subst. pose proof (h4 eq_refl); contradiction. apply lt_S_n in h3. specialize (ih _ h2 _ h3). unfold lind in ih. pose proof (proof_irrelevance _ h2 (neq_in_cons _ _ a h4 (or_intror h2))) as h7. rewrite <- h7 in ih at 1. pose proof (eq_refl (lind_occ h1 l 0 x h2)) as h8. rewrite ih in h8. destruct h8 as [h8 h9]. contradiction. destruct h2 as [|h2]; subst. contradict h4; auto. destruct n as [|n]; subst. destruct h5 as [? h5]; subst. contradict h4; auto. apply lt_S_n in h3. destruct h5 as [h5 h6]. f_equal. specialize (ih _ h2 n h3). let P := type of ih in match P with _ <-> ?hyp => assert (h7:hyp) end. rewrite nth_lt_compat. split. erewrite nth_indep. rewrite h5 at 1; auto. assumption. contradict h6. simpl. right; auto. rewrite <- ih in h7. rewrite <- h7. unfold lind. f_equal. apply proof_irrelevance. Qed. Definition lind_singularize {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pf:n < length l) : nat := let pf' := iff1 (list_singularize_in_iff pfdt _ _) (in_nth_lt l n pf) in lind pfdt (list_singularize pfdt l) (nth_lt l n pf) pf'. Lemma lt_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pf:n < length l), lind_singularize pfdt l n pf < length (list_singularize pfdt l). intros; apply lt_lind. Qed. Lemma lind_singularize_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pf:n < length l), nth_lt l n pf = nth_lt (list_singularize pfdt l) (lind_singularize pfdt l n pf) (lt_lind_singularize pfdt _ _ _). intros T h1 l n h2. unfold lind_singularize. erewrite nth_lt_functional3. rewrite lind_compat; auto. Qed. Lemma lind_singularize_eq_nth_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i j:nat) (pfi:i < length l) (pfj:j < length l), lind_singularize pfdt l i pfi = lind_singularize pfdt l j pfj -> nth_lt l i pfi = nth_lt l j pfj. intros T hdt l i j h1 h2 h3. unfold lind_singularize in h3. apply lind_inj_iff in h3; auto. Qed. Lemma in_nth_lt_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (pflt:i < length l), Inn (list_to_set l) (nth_lt (list_singularize pfdt l) (lind_singularize pfdt l i pflt) (lt_lind_singularize pfdt l i pflt)). intros T h1 l i h2. rewrite (list_to_set_singularize_compat h1). rewrite <- list_to_set_in_iff. apply in_nth_lt. Qed. Lemma inj_ind_fun : forall {T:Type} (pfdt:type_eq_dec T) (m:nat) (l:list T) (pflt:pred m < length l), Injective (fun i : (S (lind_singularize pfdt l (pred m) pflt)) @ => let pflt' := le_lt_trans (proj1_sig i) (lind_singularize pfdt l (pred m) pflt) _ (lt_n_Sm_le _ _ (proj2_sig i)) (lt_lind_singularize _ _ _ _) in nth_lt (list_singularize pfdt l) _ pflt'). intros T h1 m l h2; red; intros x y h4. apply nth_lt_inj in h4. apply proj1_sig_injective in h4; auto. apply no_dup_list_singularize. Qed. Lemma lind_singularize_cons_in0 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pflt: 0 < S (length l)) (pfi:In a l), let pfi' := iff1 (list_singularize_in_iff pfdt _ _) pfi in lind_singularize pfdt (a::l) 0 pflt = lind pfdt _ _ pfi'. intros T hdt l a h1 h2. unfold lind_singularize. gen0. rewrite nth_lt_0_hd'; simpl. rewrite list_singularize_in_cons; auto. intros; apply lind_functional; auto. Qed. Lemma lind_singularize_cons_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (a:T) (pf:i < S (length l)) (pflt:i > 0), In a l -> let pf' := mono_pred_lt _ _ pf pflt in lind_singularize pfdt (a::l) i pf = lind_singularize pfdt l (pred i) pf'. intros T hdt l i a h1 h2 h3 h4. unfold lind_singularize. gen0. rewrite list_singularize_in_cons; auto. erewrite nth_lt_cons'. intro h5. apply lind_functional. apply nth_lt_functional3. Grab Existential Variables. assumption. Qed. Lemma lind_singularize_cons_nin0 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pflt: 0 < S (length l)), ~In a l -> lind_singularize pfdt (a::l) 0 pflt = 0. intros T hdt l a h1 h2. unfold lind_singularize. gen0. rewrite list_singularize_nin; auto. rewrite nth_lt_0_hd'; simpl. intros h3. erewrite lind_functional3. rewrite lind_cons_same; auto. Grab Existential Variables. left; auto. Qed. Lemma lind_singularize_cons_nin : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (a:T) (pf:i < S (length l)) (pflt:i > 0) (pfni: ~In a l), let pf' := mono_pred_lt _ _ pf pflt in lind_singularize pfdt (a::l) i pf = S (lind_singularize pfdt l _ pf'). intros T hdt l i a h1 h2 h3. unfold lind_singularize. gen0. erewrite nth_lt_cons'. rewrite list_singularize_nin; auto. intro h4. pose proof h4 as h4'. destruct h4' as [h5 | h5]. contradict h3. rewrite h5. apply in_nth_lt. erewrite lind_functional3. pose proof @lind_cons_neq. pose proof (lind_cons_neq hdt (list_singularize hdt l) a _ h5) as h6. hfirst h6. intro h7. pose proof h5 as h5'. rewrite h7 in h5'. rewrite <- list_singularize_in_iff in h5'; contradiction. specialize (h6 hf). rewrite h6. repeat f_equal. apply proof_irrelevance. Qed. Definition im_lind_singularize {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pf:n <= length l) : Ensemble nat := im_in_ens (fun i (pfi:Inn (seg n) i) => let pfilt := seg_in _ _ pfi in let pfilt' := lt_le_trans _ _ _ pfilt pf in lind_singularize pfdt l i pfilt'). Lemma lt_im_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i n:nat) (pf:n <= length l), Inn (im_lind_singularize pfdt l n pf) i -> i < length (list_singularize pfdt l). intros T hdt l i n h1 h2. destruct h2 as [y h2]; subst. simpl. destruct h2 as [h2]. apply lt_lind_singularize. Qed. Lemma finite_im_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pf:n <= length l), Finite (im_lind_singularize pfdt l n pf). intros T hdt l n h1. unfold im_lind_singularize. apply finite_im_in_ens. apply finite_seg. Qed. Definition map_seg_lind_singularize {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pf:n <= length l) : list nat := map_seg (fun i (pfi:i < n) => let pfilt' := lt_le_trans _ _ _ pfi pf in lind_singularize pfdt l i pfilt'). Lemma lt_map_seg_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i n:nat) (pf:n <= length l), In i (map_seg_lind_singularize pfdt l n pf) -> i < length (list_singularize pfdt l). intros T hdt l i n h1 h2. unfold map_seg_lind_singularize in h2. rewrite in_map_seg_iff in h2. destruct h2 as [m [h2 h3]]; subst. apply lt_lind_singularize. Qed. Lemma length_map_seg_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pf:n <= length l), length (map_seg_lind_singularize pfdt l n pf) = n. unfold map_seg_lind_singularize; intros; rewrite length_map_seg; auto. Qed. Definition inv_map_seg_lind_singularize {T:Type} (pfdt:type_eq_dec T) (l:list T) (i n:nat) (pfle:n <= length l) (pfi:In i (map_seg_lind_singularize pfdt l n pfle)) : nat := lind nat_eq_dec _ _ pfi. Lemma inv_map_seg_lind_singularize_functional : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i i' n:nat) (pfle:n <= length l) (pfi: In i (map_seg_lind_singularize pfdt l n pfle)) (pfi':In i' (map_seg_lind_singularize pfdt l n pfle)), i = i' -> inv_map_seg_lind_singularize pfdt l i n pfle pfi = inv_map_seg_lind_singularize pfdt l i' n pfle pfi'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma lt_inv_map_seg_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i n:nat) (pfle:n <= length l) (pfi:In i (map_seg_lind_singularize pfdt l n pfle)), inv_map_seg_lind_singularize pfdt l i n pfle pfi < n. intros T hdt l i n h1 h2. unfold inv_map_seg_lind_singularize, map_seg_lind_singularize. eapply lt_eq_trans. apply lt_lind. apply length_map_seg. Qed. Lemma inv_map_seg_lind_singularize_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i n:nat) (pfle:n <= length l) (pfi:In i (map_seg_lind_singularize pfdt l n pfle)), let pflt := lt_inv_map_seg_lind_singularize pfdt l i n pfle pfi in let pflt' := lt_le_trans _ _ _ pflt pfle in lind_singularize pfdt l (inv_map_seg_lind_singularize pfdt l i n pfle pfi) pflt' = i. intros T hdt l i n h1 h2 h3 h4. unfold lind_singularize, inv_map_seg_lind_singularize, map_seg_lind_singularize. pose proof h2 as h2'. unfold map_seg_lind_singularize in h2'. rewrite in_map_seg_iff in h2'. destruct h2' as [m [h5 h6]]. rewrite <- h6. unfold lind_singularize. apply lind_functional. rewrite (nth_lt_eq_iff hdt). exists (in_nth_lt _ _ _). exists 0. split. apply gt_lt. rewrite <- (count_occ_In hdt). apply in_nth_lt. symmetry. pose proof @lind_eq_iff. pose proof (lind_eq_iff nat_eq_dec (map_seg_lind_singularize hdt l n h1) i h2 (lind hdt l _ (in_nth_lt _ _ (lt_le_trans _ _ _ h5 h1)))) as hl. hfirst hl. rewrite length_map_seg_lind_singularize; auto. eapply le_lt_trans. eapply le_lind_nth_lt. assumption. specialize (hl hf). unfold map_seg_lind_singularize, lind_singularize in hl. rewrite hl. split. rewrite nth_lt_map_seg'. gen0. erewrite nth_lt_functional3. rewrite lind_compat. intros h7. rewrite <- h6. unfold lind_singularize. apply lind_functional; auto. intro h7. rewrite in_firstn_iff in h7. destruct h7 as [r [h7 h8]]. rewrite nth_lt_map_seg' in h8. rewrite <- h6 in h8. unfold lind_singularize in h8. rewrite lind_inj_iff in h8. revert h8. gen0. intros h9 h10. pose proof (le_lind_nth_lt hdt l r h9) as h11. revert h11. generalize (in_nth_lt _ _ h9). rewrite h10. intros h11 h12. pose proof (proof_irrelevance _ h11 (in_nth_lt _ _ _)); subst. omega. Grab Existential Variables. rewrite length_map_seg. pose proof (le_lind_nth_lt hdt l m (lt_le_trans _ _ _ h5 h1)) as h7. unfold lind in h7. apply lt_impl_le. eapply le_lt_trans. apply h7. assumption. Qed. Lemma list_to_set_map_seg_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pf:n <= length l), list_to_set (map_seg_lind_singularize pfdt l n pf) = im_lind_singularize pfdt l n pf. intros T pfdt l n h1. unfold im_lind_singularize, map_seg_lind_singularize. rewrite list_to_set_map_seg. reflexivity. Qed. Lemma in_im_lind_singularize_map_seg : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i n:nat) (pf:n <= length l), Inn (im_lind_singularize pfdt l n pf) i -> In i (map_seg_lind_singularize pfdt l n pf). intros; rewrite list_to_set_in_iff, list_to_set_map_seg_lind_singularize; auto. Qed. Definition inv_im_lind_singularize {T:Type} (pfdt:type_eq_dec T) (l:list T) (i n:nat) (pfle:n <= length l) (pfi:Inn (im_lind_singularize pfdt l n pfle) i) := let pfi' := in_im_lind_singularize_map_seg pfdt l i n pfle pfi in inv_map_seg_lind_singularize pfdt l i n pfle pfi'. Lemma lt_inv_im_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i n:nat) (pfle:n <= length l) (pfi:Inn (im_lind_singularize pfdt l n pfle) i), let pfi' := in_im_lind_singularize_map_seg pfdt l i n pfle pfi in inv_map_seg_lind_singularize pfdt l i n pfle pfi' < n. unfold inv_map_seg_lind_singularize; intros. eapply lt_eq_trans. apply lt_lind. apply length_map_seg. Qed. Lemma inv_im_lind_singularize_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i n:nat) (pfle:n <= length l) (pfi:Inn (im_lind_singularize pfdt l n pfle) i), let pflt := lt_le_trans _ _ _ (lt_inv_im_lind_singularize pfdt l i n pfle pfi) pfle in lind_singularize pfdt l (inv_im_lind_singularize pfdt l i n pfle pfi) pflt = i. intros T hdt l i n h1 h2 h3. unfold inv_im_lind_singularize. pose proof (inv_map_seg_lind_singularize_compat hdt l i n h1 (in_im_lind_singularize_map_seg hdt l i n h1 h2)) as h4. simpl in h4. rewrite <- h4. apply f_equal_dep; auto. Qed. Definition im_im_lind_singularize {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) : Ensemble T := im_in_ens (fun (i : nat) (pf : Inn (im_lind_singularize pfdt l n pfle) i) => let pf' := lt_im_lind_singularize pfdt l i n pfle pf in nth_lt (list_singularize pfdt l) i pf'). Lemma finite_im_im_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l), Finite (im_im_lind_singularize pfdt l n pfle). unfold im_im_lind_singularize; intros. apply finite_im_in_ens. apply finite_im_lind_singularize. Qed. Definition map_in_dep_map_seg_lind_singularize {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) : list T := map_in_dep (fun (i : nat) (pf : In i (map_seg_lind_singularize pfdt l n pfle)) => let pf' := lt_map_seg_lind_singularize pfdt l i n pfle pf in nth_lt (list_singularize pfdt l) i pf'). Lemma length_map_in_dep_map_seg_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l), length (map_in_dep_map_seg_lind_singularize pfdt l n pfle) = n. unfold map_in_dep_map_seg_lind_singularize. intros; rewrite length_map_in_dep. apply length_map_seg_lind_singularize. Qed. Lemma in_map_in_dep_map_seg_lind_singularize' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T), In x (map_in_dep_map_seg_lind_singularize pfdt l n pfle) -> In x l. intros T hdt l n h1 x h2. unfold map_in_dep_map_seg_lind_singularize in h2. rewrite in_map_in_dep_iff in h2. destruct h2 as [c [h2 h3]]; subst. rewrite list_singularize_in_iff. apply in_nth_lt. Qed. Lemma list_to_set_map_in_dep_map_seg_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l), list_to_set (map_in_dep_map_seg_lind_singularize pfdt l n pfle) = im_im_lind_singularize pfdt l n pfle. intros T hdt l n h1. unfold map_in_dep_map_seg_lind_singularize, im_im_lind_singularize. rewrite list_to_set_map_in_dep. pose proof (list_to_set_map_seg_lind_singularize hdt l n h1) as h2. rewrite (im_in_ens_transfer _ h2). f_equal. apply functional_extensionality_dep. intro i. apply functional_extensionality. intro h3. rewrite fun_in_ens_transfer_compat. simpl. f_equal. apply proof_irrelevance. Qed. Lemma in_map_in_dep_map_seg_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T), Inn (im_im_lind_singularize pfdt l n pfle) x -> In x (map_in_dep_map_seg_lind_singularize pfdt l n pfle). intros T hdt l n pfle x h1. rewrite <- list_to_set_map_in_dep_map_seg_lind_singularize in h1. rewrite <- list_to_set_in_iff in h1; auto. Qed. Definition inv_map_in_dep_map_seg_lind_singularize {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T) (pf:In x (map_in_dep_map_seg_lind_singularize pfdt l n pfle)) : nat := let pflt := lt_eq_trans _ _ _ (lt_eq_trans _ _ _ (lt_lind pfdt _ _ _) (length_map_in_dep_map_seg_lind_singularize pfdt l n pfle)) (eq_sym (length_map_seg_lind_singularize pfdt l n pfle)) in nth_lt (map_seg_lind_singularize pfdt l n pfle) (lind pfdt _ x pf) pflt. Lemma lt_inv_map_in_dep_map_seg_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T) (pf:In x (map_in_dep_map_seg_lind_singularize pfdt l n pfle)), inv_map_in_dep_map_seg_lind_singularize pfdt l n pfle x pf < length (list_singularize pfdt l). intros T hdt l n h1 x h2. unfold inv_map_in_dep_map_seg_lind_singularize. unfold map_seg_lind_singularize. rewrite nth_lt_map_seg'. unfold lind_singularize. apply lt_lind. Qed. Lemma inv_map_in_dep_map_seg_lind_singularize_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T) (pf:In x (map_in_dep_map_seg_lind_singularize pfdt l n pfle)) (pflt:inv_map_in_dep_map_seg_lind_singularize pfdt l n pfle x pf < length (list_singularize pfdt l)), nth_lt _ _ pflt = x. intros T hdt l n hle x h1 h2. rewrite nth_lt_eq_iff. ex_goal. rewrite <- list_singularize_in_iff. eapply in_map_in_dep_map_seg_lind_singularize'. apply h1. exists hex. exists 0. split. apply gt_lt. apply (count_occ_In hdt). assumption. symmetry. unfold inv_map_in_dep_map_seg_lind_singularize, map_seg_lind_singularize. rewrite nth_lt_map_seg'. unfold lind_singularize. fold (lind hdt _ _ hex). apply lind_functional. rewrite nth_lt_eq_iff. ex_goal. rewrite <- list_singularize_in_iff in hex; auto. exists hex0. exists 0. split. apply gt_lt. apply (count_occ_In hdt); auto. unfold lind. unfold map_in_dep_map_seg_lind_singularize. symmetry. change (lind hdt (map_in_dep (fun (i : nat) (pf : In i (map_seg_lind_singularize hdt l n hle)) => nth_lt (list_singularize hdt l) i (lt_map_seg_lind_singularize hdt l i n hle pf))) x h1 = lind hdt l x hex0). rewrite lind_eq_iff. split. rewrite nth_lt_map_in_dep'. rewrite nth_lt_eq_iff. exists hex, 0. split. apply gt_lt. apply (count_occ_In hdt); auto. unfold map_seg_lind_singularize. rewrite nth_lt_map_seg'. unfold lind_singularize. apply lind_occ_functional; auto. erewrite nth_lt_functional3. rewrite lind_compat; auto. intro h3. rewrite in_firstn_iff in h3. destruct h3 as [m [h3 h4]]. rewrite nth_lt_map_in_dep' in h4. rewrite nth_lt_eq_iff in h4. destruct h4 as [h4 [occ [h5 h6]]]. unfold map_seg_lind_singularize in h6. rewrite nth_lt_map_seg' in h6. unfold lind_singularize in h6. apply lind_occ_inj in h6. symmetry in h6. revert h6. gen0. intros h6 h7; subst. pose proof (le_lind_nth_lt hdt l m h6) as h7. pose proof (proof_irrelevance _ hex0 (in_nth_lt _ _ _ )); subst. omega. Grab Existential Variables. rewrite length_map_in_dep, length_map_seg_lind_singularize. unfold map_in_dep_map_seg_lind_singularize in h1. pose proof h1 as h1'. rewrite in_map_in_dep_iff in h1'. destruct h1' as [c [h3 h4]]; subst. unfold map_seg_lind_singularize in h3. pose proof h3 as h3'. rewrite in_map_seg_iff in h3'. destruct h3' as [m [h4 h6]]; subst. unfold lind_singularize. generalize hex0. unfold lind_singularize. erewrite nth_lt_functional3. rewrite lind_compat. intro h7. eapply lt_impl_le. eapply le_lt_trans. erewrite lind_functional3. apply le_lind_nth_lt. assumption. rewrite length_map_in_dep, length_map_seg_lind_singularize. unfold map_in_dep_map_seg_lind_singularize in h1. pose proof h1 as h1'. rewrite in_map_in_dep_iff in h1'. destruct h1' as [c [h3 h4]]; subst. unfold map_seg_lind_singularize in h3. pose proof h3 as h3'. rewrite in_map_seg_iff in h3'. destruct h3' as [m [h4 h6]]; subst. unfold lind_singularize. generalize hex0. unfold lind_singularize. erewrite nth_lt_functional3. rewrite lind_compat. intro h7. eapply le_lt_trans. erewrite lind_functional3. apply le_lind_nth_lt. assumption. Qed. Lemma in_inv_map_in_dep_map_seg_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T) (pf:In x (map_in_dep_map_seg_lind_singularize pfdt l n pfle)), In (inv_map_in_dep_map_seg_lind_singularize pfdt l n pfle x pf) (map_seg_lind_singularize pfdt l n pfle). intros T hdt l n h1 x h2. unfold map_seg_lind_singularize. rewrite in_map_seg_iff. pose proof h2 as h2'. unfold map_in_dep_map_seg_lind_singularize in h2'. rewrite in_map_in_dep_iff in h2'. destruct h2' as [i [h3 h4]]. unfold map_seg_lind_singularize in h3. pose proof h3 as h5. rewrite in_map_seg_iff in h5; subst. destruct h5 as [m [h5 h6]]. exists m, h5. rewrite h6. subst. unfold lind_singularize. eapply nth_lt_inj. eapply no_dup_list_singularize. rewrite inv_map_in_dep_map_seg_lind_singularize_compat. apply nth_lt_functional2; auto. Grab Existential Variables. apply lt_inv_map_in_dep_map_seg_lind_singularize. apply lt_lind. Qed. Definition inv_inv_map_in_dep_map_seg_lind_singularize {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T) (pf:In x (map_in_dep_map_seg_lind_singularize pfdt l n pfle)) : nat := let pf' := in_inv_map_in_dep_map_seg_lind_singularize pfdt l n pfle x pf in inv_map_seg_lind_singularize pfdt l _ n pfle pf'. Lemma lt_inv_inv_map_in_dep_map_seg_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T) (pf:In x (map_in_dep_map_seg_lind_singularize pfdt l n pfle)), inv_inv_map_in_dep_map_seg_lind_singularize pfdt l n pfle x pf < n. unfold inv_inv_map_in_dep_map_seg_lind_singularize; intros. apply lt_inv_map_seg_lind_singularize. Qed. Definition inv_inv_im_im_lind_singularize {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T) (pf:Inn (im_im_lind_singularize pfdt l n pfle) x) : nat := let pf' := in_map_in_dep_map_seg_lind_singularize pfdt l n pfle x pf in inv_inv_map_in_dep_map_seg_lind_singularize pfdt l n pfle x pf'. Lemma lt_inv_inv_im_im_lind_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T) (pf:Inn (im_im_lind_singularize pfdt l n pfle) x), inv_inv_im_im_lind_singularize pfdt l n pfle x pf < n. unfold inv_inv_im_im_lind_singularize; intros; apply lt_inv_inv_map_in_dep_map_seg_lind_singularize. Qed. Lemma in_im_im_lind_singularize_inv : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i n:nat) (pfle:n <= length l) (pfin:i < n) (pfi:Inn (im_im_lind_singularize pfdt l n pfle) (nth_lt l i (lt_le_trans i n (length l) pfin pfle))), Inn (im_lind_singularize pfdt l n pfle) (lind_singularize pfdt l i (lt_le_trans i n (length l) pfin pfle)). intros T hdt l i n hle hin hi. unfold im_lind_singularize. econstructor. apply f_equal_dep. reflexivity. Grab Existential Variables. constructor; auto. Qed. (*Search below for [inv_inv_map_in_dep_map_seg_lind_singularize_compat]*) Lemma nth_lt_tl : forall {T} (l:list T) i (pfi:i < length (tl l)) (pfn:l <> nil), let pfi0 := lt_eq_trans _ _ _ pfi (length_tl l) in let pfi' := lt_eq_trans _ _ _ (lt_n_S _ _ pfi0) (eq_sym (S_pred _ _ (not_nil_lt l pfn))) in nth_lt (tl l) i pfi = nth_lt l (S i) pfi'. intros T l i hi hn; simpl. destruct l as [|a l]; simpl. contradict hn; auto. rewrite nth_lt_compat. apply nth_indep. simpl in hi; auto. Qed. Lemma tl_map_seg : forall {U:Type} {n:nat} (pfdu:type_eq_dec U) (f:seg_fun n U), tl (map_seg f) = map_seq (seg_fun_tl f). intros U n hdu f. symmetry. rewrite map_seq_eq_iff. ex_goal. rewrite length_tl, length_map_seg; auto. exists hex. intro m; simpl; intro h1. unfold seg_fun_tl, seq_fun_tl. rewrite seg_fun_to_seq_compat. symmetry. rewrite (nth_lt_eq_iff hdu). ex_goal. rewrite in_tl_iff'. exists m. ex_goal. rewrite length_map_seg; omega. exists hex0. rewrite nth_lt_map_seg'. f_equal; apply proof_irrelevance. auto. exists hex0. pose proof h1 as h1'. rewrite hex in h1'. pose proof (ex_lind_occ_nth_lt hdu _ _ h1') as h2. destruct h2 as [occ [h2 h3]]. exists occ. split. pose proof (lt_length _ _ h1') as h4. pose proof (nth_lt_tl _ _ h1') as h5. hfirst h5. rewrite map_seg_eq_nil_iff. omega. rewrite (nth_lt_tl _ _ h1' hf) , nth_lt_map_seg' in h2 at 1. specialize (h5 hf); simpl in h5. erewrite f_equal_dep. apply h2. reflexivity. rewrite <- h3. apply lind_occ_functional; auto. rewrite nth_lt_tl, nth_lt_map_seg' at 1. f_equal; apply proof_irrelevance. Grab Existential Variables. rewrite map_seg_eq_nil_iff. omega. Qed. Lemma nth_lt_lind_occ_same : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (occ occ':nat), let i := lind_occ pfdt l occ x pf in let i' := lind_occ pfdt l occ' x pf in nth_lt l i (lt_lind_occ _ _ _ _ _) = nth_lt l i' (lt_lind_occ _ _ _ _ _). intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2 occ occ'. destruct (h1 a x), (h1 x a); try contradiction; simpl. clear e. subst. destruct occ as [|occ], occ' as [|occ']; simpl; auto. destruct (h1 a a) as [ha | ha]. destruct (in_dec h1 a l) as [h5 | h5]. erewrite nth_indep. rewrite <- (nth_lt_compat _ _ (lt_lind_occ _ _ _ _ _)) at 1. rewrite lind_occ_compat; auto. apply lt_lind_occ. reflexivity. destruct (in_dec h1 a l) as [h5 | h5]. contradict ha; auto. reflexivity. destruct (in_dec h1 a l) as [h5 | h5]. erewrite nth_indep. rewrite <- (nth_lt_compat _ _ (lt_lind_occ _ _ _ _ _)) at 1. rewrite lind_occ_compat; auto. apply lt_lind_occ. reflexivity. destruct (in_dec h1 a l) as [h5 | h5]; simpl. erewrite nth_indep. rewrite <- (nth_lt_compat _ _ (lt_lind_occ _ _ _ _ _)) at 1. symmetry. erewrite nth_indep. rewrite <- (nth_lt_compat _ _ (lt_lind_occ _ _ _ _ _)) at 1. apply ih; auto. apply lt_lind_occ. apply lt_lind_occ. reflexivity. subst. contradict n; auto. subst. contradict n; auto. pose proof (neq_in_cons _ _ _ n0 h2) as h5. erewrite nth_indep. rewrite <- (nth_lt_compat _ _ (lt_lind_occ _ _ _ _ _)) at 1. symmetry. erewrite nth_indep. rewrite <- (nth_lt_compat _ _ (lt_lind_occ _ _ _ _ _)) at 1. apply ih; auto. apply lt_lind_occ. apply lt_lind_occ. Qed. Lemma nth_lt_lind_lind_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (occ:nat) (x:T) (pf:In x l), nth_lt l (lind_occ pfdt l occ x pf) (lt_lind_occ _ _ _ _ _) = nth_lt l (lind pfdt l x pf) (lt_lind _ _ _ _). unfold lind; intros; erewrite nth_lt_lind_occ_same; apply nth_lt_functional3. Qed. Lemma nth_lt_lind_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (occ:nat) (x:T) (pf:In x l), nth_lt l (lind_occ pfdt l occ x pf) (lt_lind_occ _ _ _ _ _) = x. intros; rewrite nth_lt_lind_lind_occ, lind_compat; auto. Qed. Lemma lind_occ_map_inj : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:T->U) (l:list T) (x:T) (pfx:In x l) (pffx:In (f x) (map f l)) (occ:nat), Injective f -> lind_occ pfdu (map f l) occ (f x) pffx = lind_occ pfdt l occ x pfx. intros T U h1 h2 f l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h3 h4 occ h5. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. apply h5 in hl; subst. erewrite eq_dec_same at 1. destruct occ as [|occ]; auto. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. f_equal. apply ih; auto. pose proof hl as hl'. rewrite in_map_iff in hl'. destruct hl' as [a' [h6 h7]]. apply h5 in h6; subst. contradiction. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. contradict hr; apply in_map; auto. reflexivity. assumption. destruct h4 as [h4 | h4]. pose proof h4 as h4'. symmetry in h4'; contradiction. destruct h3 as [|h3]; subst. contradict hr; auto. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. subst. contradict hr; auto. f_equal. apply ih; auto. Qed. Lemma lind_occ_map_in_dep_inj : forall {T U:Type} {l:list T} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (f:fun_in_dep l U) (x:T) (pfx:In x l) (pffx:In (f x pfx) (map_in_dep f)) (occ:nat), inj_dep f -> lind_occ pfdu (map_in_dep f) occ (f x pfx) pffx = lind_occ pfdt l occ x pfx. intros T U l h1 h2. induction l as [|a l ih]; simpl. intros; contradiction. intros f x h3 h4 occ h5. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. apply h5 in hl; subst. erewrite eq_dec_same at 1. destruct occ as [|occ]; auto. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. f_equal. specialize (ih (fun_from_cons f) a hl'). hfirst ih. apply in_map_in_dep. specialize (ih hf occ). hfirst ih. apply inj_from_cons; auto. specialize (ih hf0). unfold fun_from_cons in ih at 2. rewrite <- ih; apply f_equal_dep; f_equal; apply proof_irrelevance. pose proof hl as hl'. rewrite in_map_in_dep_iff in hl'. destruct hl' as [a' [h6 h7]]. unfold fun_from_cons in h7. red in h5. apply h5 in h7; subst; contradiction. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. contradict hr. rewrite in_map_in_dep_iff. exists a, hl'. unfold fun_from_cons; f_equal; apply proof_irrelevance. reflexivity. assumption. destruct h4 as [h4 | h4]. pose proof h4 as h4'. symmetry in h4'; contradiction. destruct h3 as [|h3]; subst. contradict hr; f_equal; apply proof_irrelevance. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. subst. contradict hr; f_equal; apply proof_irrelevance. f_equal. specialize (ih (fun_from_cons f) x h3 (in_map_in_dep _ _ _) occ). hfirst ih. apply inj_from_cons; auto. specialize (ih hf). unfold fun_from_cons in ih at 2; auto. symmetry. erewrite f_equal_dep; auto. rewrite <- ih at 1. apply f_equal_dep; f_equal; apply proof_irrelevance. Qed. Lemma lind_map_seg_fun0 : forall {U:Type} {n:nat} (pfdu:type_eq_dec U) (f:seg_fun n U) (pflt: 0 < n) (pfi:In (f 0 pflt) (map_seg f)), lind pfdu (map_seg f) (f 0 pflt) pfi = 0. intros U n hdu f h1 h2. rewrite lind_eq_iff. split. rewrite nth_lt_map_seg'. f_equal. apply proof_irrelevance. simpl. intro; contradiction. Grab Existential Variables. rewrite length_map_seg; auto. Qed. Lemma nat_fun_lind_map_seg : forall {U:Type} {n:nat} (pfdu:type_eq_dec U) (f:seg_fun n U) (i:nat) (pfi:i < n) (pffi:In (f i pfi) (map_seg f)) (pflt:lind pfdu (map_seg f) (f i pfi) pffi < n), f (lind pfdu (map_seg f) (f i pfi) pffi) pflt = f i pfi. intros U n. induction n as [|n ih]; simpl. intros; contradiction. intros hdu f i h2 h3 h4. destruct i as [|i]; simpl. pose proof h3 as h3'. apply in_app_or in h3'. destruct h3' as [h5 | h5]. generalize h4. erewrite lind_functional3. erewrite lind_app_in1. intro h6. destruct (zerop n) as [|h7]; subst. pose proof h6 as h6'. apply lt1 in h6'. gen0. rewrite h6'; intros; f_equal; apply proof_irrelevance. pose proof (lind_map_seg_fun0 hdu (seg_fun_from_S f) h7) as h8. hfirst h8. apply in_map_seg. specialize (h8 hf). unfold seg_fun_from_S in h8. unfold seg_fun_from_S. pose proof (f_equal_dep f _ 0 h6 h2) as h9. hfirst h9. rewrite <- h8. apply lind_functional. f_equal. apply proof_irrelevance. specialize (h9 hf0). rewrite h9 at 1. reflexivity. simpl in h5. destruct h5 as [h5 | h5]. destruct (zerop n) as [|h6]; subst. gen0. erewrite lind_functional. simpl. rewrite lind_sing. intros; f_equal; apply proof_irrelevance. reflexivity. gen0. erewrite lind_functional. erewrite lind_app_in1. erewrite lind_map_seg_fun0. intros; f_equal; apply proof_irrelevance. unfold seg_fun_from_S. f_equal. apply proof_irrelevance. contradiction. pose proof h2 as h2'. apply lt_S_n in h2'. pose proof h3 as h3'. apply in_app_or in h3'. pose proof h2 as hn. red in hn. apply le_S_n in hn. apply le_lt_eq_dec in hn. destruct hn as [hn | hn]. destruct h3' as [h5 | h5]. gen0. erewrite lind_functional. erewrite lind_app_in1. intro h6. specialize (ih hdu (seg_fun_from_S f) (S i) hn). hfirst ih. rewrite seg_fun_from_S_compat; auto. rewrite in_map_seg_iff. exists (S i), hn. rewrite seg_fun_from_S_compat; auto. specialize (ih hf). hfirst ih. eapply lt_eq_trans. apply lt_lind. rewrite length_map_seg; auto. specialize (ih hf0). do 2 rewrite seg_fun_from_S_compat in ih. erewrite f_equal_dep in ih. rewrite ih at 1. f_equal. apply proof_irrelevance. apply lind_functional. reflexivity. rewrite seg_fun_from_S_compat. f_equal. apply proof_irrelevance. destruct h5 as [h5|]. rewrite <- h5. gen0. generalize h3. rewrite <- h5. intros h6. erewrite lind_functional. rewrite lind_app_in1. intro h7. rewrite h5. specialize (ih hdu (seg_fun_from_S f) (S i) hn). hfirst ih. apply in_map_seg. specialize (ih hf). hfirst ih. eapply lt_eq_trans. apply lt_lind. apply length_map_seg. specialize (ih hf0). do 2 rewrite seg_fun_from_S_compat in ih. symmetry. erewrite f_equal_dep. rewrite <- ih at 1. apply f_equal_dep. apply lind_functional. reflexivity. reflexivity. rewrite seg_fun_from_S_compat. rewrite h5. f_equal. apply proof_irrelevance. contradiction. subst. destruct (in_dec hdu (f (S i) h2) (map_seg (seg_fun_from_S f))) as [hi | hi]. pose proof hi as hj. rewrite in_map_seg_iff in hj. destruct hj as [j [hj h6]]. specialize (ih hdu (seg_fun_from_S f) j hj). hfirst ih. apply in_map_seg. specialize (ih hf). hfirst ih. eapply lt_eq_trans. apply lt_lind. apply length_map_seg. specialize (ih hf0). rewrite seg_fun_from_S_compat in ih. revert ih. Gen0. generalize hf. rewrite h6. intros h7 h8 h9. rewrite <- h9. apply f_equal_dep. erewrite lind_functional. rewrite lind_app_in1. reflexivity. reflexivity. gen0. erewrite lind_functional. rewrite lind_app_in2, lind_sing, plus_0_r, length_map_seg. intros; f_equal; apply proof_irrelevance. apply hi. reflexivity. Grab Existential Variables. constructor. f_equal; apply proof_irrelevance. apply in_map_seg. apply in_map_seg. assumption. apply in_map_seg. apply in_app_r. constructor. assumption. assumption. Qed. Lemma nth_lt_lS : forall (l:list nat) i (pfi:i < (length (lS l))), nth_lt (lS l) i pfi = S (nth_lt l i (lt_eq_trans _ _ _ pfi (length_lS _))). intros l i hi. gtermr. redterm0 y. rewrite (nth_lt_eq_iff nat_eq_dec). ex_goal. rewrite in_lS_iff. exists c; split. unfold c; apply in_nth_lt. reflexivity. exists hex. pose proof hi as hi'. rewrite length_lS in hi'. pose proof (ex_lind_occ_nth_lt nat_eq_dec _ _ hi') as h1. destruct h1 as [occ [h1 h2]]. exists occ; split; auto. rewrite count_occ_lS; simpl. erewrite nth_lt_functional3. apply h1. auto with arith. rewrite <- h2. unfold lS. erewrite lind_occ_map_inj. apply lind_occ_functional; auto. apply nth_lt_functional3. red; intros ? ? h3. apply S_inj in h3; auto. Grab Existential Variables. apply in_nth_lt. Qed. Lemma nat_fun_lind_map_seg' : forall {U:Type} (pfdu:type_eq_dec U) (f:nat->U) (i n:nat) (pfi:i < n) (pffi:In (f i) (map_seg (fun_to_seg f n))), f (lind pfdu (map_seg (fun_to_seg f n)) (f i) pffi) = f i. intros U hdu f i n h1 h2. pose proof (nat_fun_lind_map_seg hdu (fun_to_seg f n) i h1 (in_map_seg _ _ h1)) as h3. hfirst h3. eapply lt_eq_trans. apply lt_lind. apply length_map_seg. specialize (h3 hf). do 2 rewrite fun_to_seg_compat in h3. rewrite <- h3. f_equal. apply lind_functional. rewrite fun_to_seg_compat; auto. Qed. Lemma lt_lind_count_occ_S : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:count_occ pfdt l x = 1) (n:nat), let pfin := count_occ_eq_S_in pfdt l x _ pf in lind pfdt l x pfin = lind_occ pfdt l (S n) x pfin. intros T h1 l. induction l as [|a l ih]; simpl. intros; omega. intros x h2 n. destruct (h1 x a) as [|h3]; subst. pose proof h2 as h2'. terml h2'. assert (h4:x = S (count_occ h1 l a)). unfold x. destruct (h1 a a); auto. contradict n0; auto. destruct in_dec as [h3 | h3]. fold x in h2'. rewrite h2' in h4. apply S_inj in h4. pose proof h3 as h3'. rewrite (count_occ_In h1) in h3'. omega. unfold lind. simpl. assert (h5:h1 a a = left eq_refl). destruct (h1 a a); auto. f_equal. apply proof_irrelevance. contradict n0; auto. rewrite h5; auto. pose proof (neq_sym _ _ h3) as h3'. assert (h5:h1 a x = right h3'). destruct (h1 a x). contradiction. f_equal; apply proof_irrelevance. pose proof h2 as h2'. rewrite h5 in h2'. specialize ( ih x h2'). simpl in ih. unfold lind. simpl. assert (h5':h1 x a = right h3). destruct (h1 x a). contradiction. f_equal; apply proof_irrelevance. rewrite h5'. f_equal. specialize (ih n). match goal with |- lind_occ _ _ _ _ ?pf = _ => generalize pf end. intro h6. fold (lind h1 l x h6). erewrite lind_functional. rewrite ih at 1. f_equal. apply proof_irrelevance. reflexivity. Qed. Lemma lind_occ_pred_count_occ_lt_length_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l), lind_occ pfdt l (pred (count_occ pfdt l x)) x pfin < length l. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2. rewrite (switch_eq_dec_eq h1 a x). destruct (h1 x a) as [|h3]; subst. unfold switch_eq_dec. simpl. destruct (in_dec h1 a l) as [h3 | h3]. pose proof h3 as h3'. rewrite (count_occ_In h1) in h3'. rewrite (S_pred _ _ h3'). apply lt_n_S; auto. rewrite (count_occ_In h1) in h3. rewrite (not_lt0_iff (count_occ h1 l a)) in h3. rewrite h3; auto with arith. apply lt_n_S. unfold switch_eq_dec. apply ih. Qed. Lemma lind_occ_pred_nth_lt_pred : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pflt:pred (length l) < length l) (pfi:In (nth_lt l _ pflt) l), lind_occ pfdt l (pred (count_occ pfdt l (nth_lt l (pred (length l)) pflt))) (nth_lt _ _ pflt) pfi = pred (length l). intros T hdt l. induction l as [|a l ih]. intro; contradiction. destruct (nil_dec' l) as [|h2]; subst. intros ? ?. simpl; do 2 rewrite (eq_dec_same hdt a) at 1; simpl; auto. intros h1 h3. pose proof h1 as h4. simpl in h4. pose proof (mono_pred_lt _ _ h4 (not_nil_lt _ h2)) as h5. simpl in h5. pose proof h3 as h3'. simpl in h3'. rewrite (S_pred _ _ (not_nil_lt _ h2)) in h3'. erewrite nth_indep in h3'. erewrite <- nth_lt_compat in h3'. destruct h3' as [h6 | h6]. simpl. specialize (ih h5 (in_nth_lt _ _ _)). lr_match_goal; fold hlr; destruct hlr as [hl | hr]. rewrite (S_pred _ _ (not_nil_lt _ h2)) in hl. rewrite (S_pred _ _ (not_nil_lt _ h2)). lr_if_goal; fold hlr; destruct hlr as [hl' | hr']; simpl. lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]; simpl. pose proof hl0 as h7. rewrite (count_occ_In hdt) in h7. rewrite (S_pred _ _ h7). f_equal. rewrite <- ih. apply lind_occ_functional; auto. repeat f_equal. rewrite nth_lt_compat. apply nth_indep. eapply lt_pred_n. apply (not_nil_lt _ h2). rewrite nth_lt_compat. apply nth_indep. eapply lt_pred_n. apply (not_nil_lt _ h2). contradict hr0. erewrite nth_indep. erewrite <- nth_lt_compat at 1. apply in_nth_lt. eapply lt_pred_n. apply (not_nil_lt _ h2). contradict hr'. rewrite h6 at 1. erewrite nth_indep. erewrite <- nth_lt_compat at 1. reflexivity. eapply lt_pred_n. apply (not_nil_lt _ h2). pose proof hr as h7. rewrite (S_pred _ _ (not_nil_lt _ h2)) in h7. contradict h7. rewrite h6 at 2. rewrite nth_lt_compat. apply nth_indep. eapply lt_pred_n. apply (not_nil_lt _ h2). simpl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. rewrite (S_pred _ _ (not_nil_lt _ h2)) in hl. rewrite (S_pred _ _ (not_nil_lt _ h2)). rewrite hl. rewrite (eq_dec_same hdt a) at 1. simpl. pose proof h6 as h7. rewrite nth_lt_compat in h7. erewrite nth_indep in h7. rewrite hl in h7 at 1. pose proof h7 as h8. rewrite (count_occ_In hdt) in h8. rewrite (S_pred _ _ h8). lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]; simpl. f_equal. erewrite <- ih. apply lind_occ_functional; auto. do 2 f_equal. rewrite <- hl at 1. rewrite nth_lt_compat. apply nth_indep. eapply lt_pred_n. eapply not_nil_lt. assumption. rewrite <- hl at 1. rewrite nth_lt_compat. apply nth_indep. eapply lt_pred_n. eapply not_nil_lt. assumption. contradiction. eapply lt_pred_n. eapply not_nil_lt. assumption. pose proof hr as h7. rewrite (S_pred _ _ (not_nil_lt _ h2)) in h7. generalize (neq_in_cons _ _ _ hr h3). rewrite (S_pred _ _ (not_nil_lt _ h2)). intro h8. lr_if_goal; fold hlr; destruct hlr as [hl0 | hr0]; simpl. contradict h7; auto. f_equal. erewrite <- ih. apply lind_occ_functional; auto. repeat f_equal. rewrite nth_lt_compat. apply nth_indep. eapply lt_pred_n. apply not_nil_lt. assumption. rewrite nth_lt_compat. apply nth_indep. eapply lt_pred_n. apply not_nil_lt. assumption. eapply lt_pred_n. apply not_nil_lt. assumption. Grab Existential Variables. apply in_nth_lt. eapply lt_pred_n. apply not_nil_lt; auto. apply in_nth_lt. eapply lt_pred_n. apply not_nil_lt; auto. eapply lt_pred_n. apply not_nil_lt; auto. eapply lt_pred_n. apply not_nil_lt; auto. Qed. Fixpoint dup_witness {T:Type} {l:list T} (pfdt:type_eq_dec T) : dup l -> nat*nat := match l with | nil => fun pf => (False_rect _ (not_dup_nil _ pf)) | a :: l' => fun (pf:dup (a::l')) => match (in_dec pfdt a l') with | left pfia => (0, (S (lind pfdt _ _ pfia))) | right pfnia => let pfdc := dup_cons pfdt _ _ pf in let pfd := dec_not_l pfdc pfnia in let p := dup_witness pfdt pfd in (S (fst p), S (snd p)) end end. Lemma lt_dup_witness1 : forall {T:Type} {l:list T} (pfdt:type_eq_dec T) (pfd:dup l), fst (dup_witness pfdt pfd) < length l. intros T l. induction l as [|a l ih]; simpl. intros ? h1. pose proof h1 as h1'. apply not_dup_nil in h1'. contradiction. intros h1 h2. destruct in_dec as [h3 | h3]; simpl; auto with arith. Qed. Lemma lt_dup_witness2 : forall {T:Type} {l:list T} (pfdt:type_eq_dec T) (pfd:dup l), snd (dup_witness pfdt pfd) < length l. intros T l. induction l as [|a l ih]; simpl. intros ? h1. pose proof h1 as h1'. apply not_dup_nil in h1'. contradiction. intros h1 h2. destruct in_dec as [h3 | h3]; simpl; auto with arith. apply lt_n_S. apply lt_lind. Qed. Lemma dup_witness2_pos : forall {T:Type} {l:list T} (pfdt:type_eq_dec T) (pfd:dup l), snd (dup_witness pfdt pfd) > 0. intros T l. destruct l as [|a l]. intros ? h1. contradict (not_dup_nil _ h1). intros h1 h2. simpl. destruct in_dec; simpl; auto with arith. Qed. Lemma dup_witness_neq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pfd:dup l), fst (dup_witness pfdt pfd) <> snd (dup_witness pfdt pfd). intros T h1 l. induction l as [|a l ih]. intro h2. contradict (not_dup_nil _ h2). intros h2. pose proof (dup_cons h1 _ _ h2) as h3. destruct h3 as [h3 | h3]. simpl. destruct in_dec as [h4 | h4]; simpl. intro; discriminate. contradiction. specialize (ih h3). simpl. destruct in_dec as [h4 | h4]. simpl; intro; discriminate. simpl. intro h5. apply S_inj in h5. pose proof (proof_irrelevance _ h3 (dec_not_l (dup_cons h1 l a h2) h4) ); subst; contradiction. Qed. Lemma dup_witness_compat : forall {T:Type} {l:list T} (pfdt:type_eq_dec T) (pfd:dup l), let pf := dup_witness pfdt pfd in nth_lt l _ (lt_dup_witness1 pfdt pfd) = nth_lt l _ (lt_dup_witness2 pfdt pfd). Proof. intros T l. induction l as [|a l ih]; simpl. intros ? h1. contradict (not_dup_nil _ h1). intros h1 h2. destruct in_dec as [h3 | h3]; simpl. erewrite nth_indep. erewrite <- nth_lt_compat. rewrite lind_compat; auto. apply lt_lind. erewrite nth_indep. erewrite <- nth_lt_compat. symmetry. erewrite nth_indep. erewrite <- nth_lt_compat. rewrite ih. reflexivity. apply lt_dup_witness2. apply lt_dup_witness1. Qed. Lemma dup_iff' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), dup l <-> exists m n (pfm:m < length l) (pfn:n < length l), m <> n /\ nth_lt _ _ pfm = nth_lt _ _ pfn. intros T h1 l; split; intro h2. exists _, _, (lt_dup_witness1 h1 h2), (lt_dup_witness2 h1 h2). split. apply dup_witness_neq. apply dup_witness_compat. revert h2. induction l as [|a l ih]; simpl. intro h2. destruct h2 as [? [? [h2]]]. contradict (lt_n_0 _ h2). intro h2. destruct h2 as [m [n [h3 [h4 [h5 h6]]]]]. destruct m as [|m], n as [|n]. contradict h5; auto. apply lt_S_n in h4. erewrite nth_indep in h6. erewrite <- (nth_lt_compat _ _ h4) in h6 at 1. subst. constructor. apply in_nth_lt. erewrite nth_indep in h6. erewrite <- (nth_lt_compat _ _ h4) in h6 at 1. subst. assumption. assumption. apply lt_S_n in h3. erewrite nth_indep in h6. erewrite <- (nth_lt_compat _ _ h3) in h6 at 1. subst. constructor. apply in_nth_lt. assumption. apply lt_S_n in h3. apply lt_S_n in h4. apply dup_intro_cons. apply ih. exists _, _, h3, h4. split. intro; subst. contradict h5; auto. erewrite nth_lt_compat. erewrite nth_indep; auto. rewrite h6 at 1. rewrite nth_lt_compat. apply nth_indep. assumption. Qed. Lemma lind_nth_lt_no_dup : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pf:n < length l) (pfi:In (nth_lt _ _ pf) l), NoDup l -> lind pfdt l (nth_lt l n pf) pfi = n. intros T h1 l n h2 h3 h4. rewrite lind_eq_iff. split. apply nth_lt_functional3. intro h5. rewrite in_firstn_iff in h5. destruct h5 as [m [h5 h6]]. rewrite no_dup_iff in h4. rewrite dup_iff' in h4. contradict h4. exists m, n, (lt_trans _ _ _ h5 h2), h2. split. intro; subst; contradict (lt_irrefl _ h5). rewrite <- h6. apply nth_lt_functional3. assumption. Grab Existential Variables. auto with arith. assumption. Qed. Lemma lind_seg_list : forall i n (pf:In i (seg_list n)), lind nat_eq_dec (seg_list n) i pf = i. intros i n h1. eapply nth_lt_inj. apply no_dup_seg_list. rewrite lind_compat, nth_lt_seg_list; auto. Grab Existential Variables. rewrite in_seg_list_iff in h1. rewrite length_seg_list; auto. Qed. Lemma lind_seq : forall i m n (pf:In i (seq m n)), lind nat_eq_dec (seq m n) i pf = i - m. intros i m n h1. pose proof h1 as h1'. rewrite in_seq_iff in h1'. eapply nth_lt_inj. apply no_dup_seq. rewrite lind_compat, nth_lt_seq; auto. omega. Grab Existential Variables. rewrite seq_length. omega. Qed. Definition seq_lind {T:Type} (l:list T) (m n:nat) (pf:m < length l) := firstn n (skipn m l). Lemma length_seq_lind : forall {T:Type} (l:list T) (m n:nat) (pf:m < length l), n <= length l - m -> length (seq_lind l m n pf) = n. intros T l m n h1 h2. unfold seq_lind. rewrite length_firstn; auto. rewrite length_skipn; auto. Qed. Lemma nth_lt_seq_lind : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m n r:nat) (pfm:m < length l) (pfr:r < length (seq_lind l m n pfm)) (pfmr:m + r < length l), nth_lt (seq_lind l m n pfm) r pfr = nth_lt l (m + r) pfmr. intros T hdt l m n r h1 h2 h3. unfold seq_lind. rewrite nth_lt_firstn_eq. erewrite nth_lt_skipn_eq. apply nth_lt_functional2; auto with arith. assumption. Grab Existential Variables. omega. Qed. Lemma removelast_seq_lind : forall {T:Type} (l:list T) (m n:nat) (pfm:m < length l), m + n <= length l -> removelast (seq_lind l m n pfm) = seq_lind l m (pred n) pfm. intros T l m n h1 hm. unfold seq_lind. destruct n as [|n]; auto. rewrite removelast_firstn. simpl; auto. rewrite length_skipn. omega. Qed. Lemma seq_lind1_pred : forall {T:Type} (l:list T) (pf:1 < length l), seq_lind l 1 (pred (length l)) pf = tl l. intros T l. destruct l as [|a l]. simpl; intros; omega. intros; simpl. unfold seq_lind. simpl. rewrite firstn_length; auto. Qed. Lemma seq_lind1_pred_pred : forall {T:Type} (l:list T) (pf1:1 < length l) (pfn:l <> nil), seq_lind l 1 (pred (pred (length l))) pf1 = removelast (tl l). intros T l h1 h2. pose proof (seq_lind1_pred l h1) as h3. rewrite <- h3. rewrite <- removelast_seq_lind; auto. omega. Qed. (*Substitutes a single element (indicated by a given index [k]) for a given arbitrary type-element [y] in a given list [l] (substitutes the same whether or not [y] is in [l]).*) Definition list_subst {T:Type} (l:list T) (k:nat) (pfk:k < length l) (x:T) : list T := map_seg (seg_fun_subst (nth_lt l) x k pfk). Lemma length_list_subst : forall {T:Type} (l:list T) (k:nat) (pfk:k < length l) (x:T), length (list_subst l k pfk x) = length l. unfold list_subst; intros; rewrite length_map_seg; auto. Qed. Lemma list_subst_cons0 : forall {T:Type} (l:list T) (a b:T) (pf : 0 < length (a::l)), list_subst (a::l) 0 pf b = b::l. intros T l a b h1. unfold list_subst. rewrite map_seg_eq_iff. simpl. exists eq_refl. intros m h2. destruct m; simpl; auto. apply nth_indep; auto with arith. Qed. Lemma list_subst_cons_S : forall {T:Type} (l:list T) (a b:T) (i:nat) (pf:S i < S (length l)), list_subst (b::l) (S i) pf a = b::list_subst l i (lt_S_n _ _ pf) a. intros T l a b i h1. unfold list_subst. rewrite map_seg_eq_iff. ex_goal. simpl; rewrite length_map_seg; auto. exists hex. intros m h2. intro h3. destruct m as [|m]; subst; auto. rewrite nth_lt_cons, nth_lt_map_seg'. simpl. unfold seg_fun_to_S. simpl. f_equal. apply functional_extensionality_dep. intro c. apply functional_extensionality. intro hc. symmetry. rewrite nth_lt_compat. apply nth_indep; auto. apply proof_irrelevance. Qed. Lemma nth_lt_list_subst_same : forall {U:Type} (l:list U) (i:nat) (a:U) (pfi:i < length l) (pfi':i < length (list_subst l i pfi a)), nth_lt (list_subst l i pfi a) i pfi' = a. unfold list_subst; intros U l. induction l as [|a l ih]. simpl. intros; omega. intros i y h2 h3. simpl in h2. pose proof h2 as h4. destruct i as [|i]. rewrite nth_lt_map_seg'. simpl; auto. rewrite nth_lt_map_seg'. apply lt_S_n in h4. specialize (ih i y h4). hfirst ih. rewrite length_map_seg; auto. specialize (ih hf). rewrite nth_lt_map_seg' in ih. simpl. unfold seg_fun_to_S. simpl. rewrite <- ih. f_equal. apply functional_extensionality_dep; intro; apply functional_extensionality; intro. rewrite nth_lt_compat. apply nth_indep; auto. apply proof_irrelevance. apply proof_irrelevance. Qed. Lemma nth_lt_list_subst_neq : forall {U:Type} (l:list U) (i m:nat) (y:U) (pfi:i < length l) (pfm:m < length (list_subst l i pfi y)), i <> m -> nth_lt (list_subst l i pfi y) m pfm = nth_lt l m (lt_eq_trans _ _ _ pfm (length_list_subst _ _ _ _)). unfold list_subst; intros U l. induction l as [|a l ih]. simpl; intros; omega. intros i m y hi. simpl in hi. pose proof hi as h1. destruct i as [|i]. intros h2 h3. apply neq_sym in h3. rewrite neq0_iff in h3. symmetry. gen0. generalize h2. rewrite (S_pred _ _ h3). intros h4 h5. rewrite nth_lt_cons, nth_lt_map_seg'. simpl in h5. simpl. rewrite nth_lt_compat. apply nth_indep. apply lt_S_n in h5; auto. intros h2 h3. pose proof h2 as h4. rewrite length_map_seg in h4. rewrite nth_lt_map_seg'. simpl in h4. destruct m as [|m]. simpl; auto. simpl. apply lt_S_n in h4. apply lt_S_n in h1. specialize (ih i m y h1). hfirst ih. rewrite length_map_seg; auto. specialize (ih hf). hfirst ih. contradict h3; auto. specialize (ih hf0). rewrite nth_lt_map_seg' in ih. rewrite nth_lt_compat in ih. erewrite nth_indep in ih. rewrite <- ih at 1. unfold seg_fun_to_S; simpl. f_equal. apply functional_extensionality_dep; intro; apply functional_extensionality; intro. rewrite nth_lt_compat. apply nth_indep. assumption. apply proof_irrelevance. apply proof_irrelevance. assumption. Qed. Lemma list_subst_seq_pred : forall (s n:nat) (pfl:pred (length (seq s n)) < length (seq s n)), list_subst (seq s n) (pred (length (seq s n))) pfl (pred s) = seq s (pred n) ++ pred s :: nil. unfold list_subst; intros s n h1. rewrite map_seg_eq_iff. ex_goal. rewrite seq_length, app_length, seq_length; simpl. destruct n as [|n]. simpl in h1. apply lt_irrefl in h1; contradiction. simpl. rewrite S_compat; auto. exists hex. intros m h2 h3. pose proof (seg_fun_subst_compat (nth_lt (seq s n)) (pred s) (pred (length (seq s n))) h1 m h2) as h4. simpl in h4. destruct h4 as [h4 h5]. pose proof h2 as h6. rewrite seq_length in h6. symmetry. destruct n as [|n]. contradict (lt_n_0 _ h6). apply lt_dec in h6. simpl in h6. destruct h6 as [h6 | h6]. hfirst h5. rewrite seq_length; simpl; intro; subst. contradict (lt_irrefl _ h6). specialize (h5 hf). rewrite h5. simpl. destruct m as [|m]. erewrite nth_lt_app1. rewrite nth_lt_seq, plus_0_r; auto. erewrite nth_lt_app1. rewrite nth_lt_seq. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite nth_lt_seq; omega. rewrite seq_length; auto with arith. unfold h3. clear h3. subst. hfirst h4. rewrite seq_length; simpl; auto. specialize (h4 hf). rewrite h4. erewrite nth_lt_app2. gen0. rewrite seq_length. intro h6. simpl. rewrite minus_same; auto. Grab Existential Variables. rewrite seq_length; simpl; auto with arith. rewrite seq_length; auto with arith. rewrite seq_length; auto. rewrite seq_length; auto. Qed. Lemma inv_inv_map_in_dep_map_seg_lind_singularize_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T) (pf:In x (map_in_dep_map_seg_lind_singularize pfdt l n pfle)), let pflt := lt_le_trans _ _ _ (lt_inv_inv_map_in_dep_map_seg_lind_singularize pfdt l n pfle x pf) pfle in nth_lt l _ pflt = x. intros T hdt l n hle x h1 h2. rewrite nth_lt_eq_iff. pose proof h1 as h3. unfold map_in_dep_map_seg_lind_singularize in h3. rewrite in_map_in_dep_iff in h3. destruct h3 as [i [h3 h4]]; subst. ex_goal. rewrite list_singularize_in_iff. apply in_nth_lt. exists hex. pose proof h3 as h5. unfold map_seg_lind_singularize in h5. rewrite in_map_seg_iff in h5. exists 0. split. apply gt_lt. rewrite <- (count_occ_In hdt). rewrite list_singularize_in_iff. apply in_nth_lt. unfold inv_inv_map_in_dep_map_seg_lind_singularize, inv_map_in_dep_map_seg_lind_singularize, inv_map_seg_lind_singularize. symmetry. gen0. unfold inv_map_in_dep_map_seg_lind_singularize, map_seg_lind_singularize. rewrite nth_lt_map_seg'. intros h10. pose proof (lind_eq_iff nat_eq_dec (map_seg (fun (i0 : nat) (pfi : i0 < n) => lind_singularize hdt l i0 (lt_le_trans i0 n (length l) pfi hle)))) as hc. match goal with |- lind _ _ ?x' _ = _ => pose x' as x end. gtermr. specialize (hc x h10 y). hfirst hc. unfold y. destruct h5 as [r [hd he]]. generalize hex. generalize (lt_map_seg_lind_singularize hdt l i n hle h3). rewrite <- he. unfold lind_singularize. intro hg. erewrite nth_lt_functional3. rewrite lind_compat. intro hh. fold (lind hdt l (nth_lt l r (lt_le_trans r n (length l ) hd hle)) hh). eapply le_lt_trans. erewrite lind_functional3. apply le_lind_nth_lt. rewrite length_map_seg. assumption. specialize (hc hf). fold x y. rewrite hc at 1. split. rewrite nth_lt_map_seg'. apply f_equal_dep. unfold lind. symmetry. unfold y. change (lind hdt (map_in_dep_map_seg_lind_singularize hdt l n hle) (nth_lt (list_singularize hdt l) i (lt_map_seg_lind_singularize hdt l i n hle h3)) h1 = lind hdt l (nth_lt (list_singularize hdt l) i (lt_map_seg_lind_singularize hdt l i n hle h3)) hex). rewrite lind_eq_iff. split. unfold map_in_dep_map_seg_lind_singularize. rewrite nth_lt_map_in_dep'. apply nth_lt_functional2. unfold map_seg_lind_singularize. rewrite nth_lt_map_seg'. unfold lind_singularize. gen0. erewrite nth_lt_functional3. rewrite lind_compat. intro h11. rewrite lind_nth_lt_no_dup; auto. apply no_dup_list_singularize. intro h11. rewrite in_firstn_iff in h11. destruct h11 as [r [h14 h12]]. unfold map_in_dep_map_seg_lind_singularize in h12. rewrite nth_lt_map_in_dep' in h12. apply nth_lt_inj in h12. unfold map_seg_lind_singularize in h12. rewrite nth_lt_map_seg' in h12. unfold lind_singularize in h12. revert h12. match goal with |- lind _ _ (nth_lt _ _ ?pf) _ = _ -> _ => generalize pf end. intros h15 h16. revert h14. generalize hex. generalize (lt_map_seg_lind_singularize hdt _ _ _ _ h3). rewrite <- h16. intros h17. erewrite nth_lt_functional3. rewrite lind_compat. intros h18 h19. pose proof (le_lind_nth_lt hdt l r h15) as h20. pose proof (proof_irrelevance _ h18 (in_nth_lt _ _ _)) as h21. rewrite <- h21 in h20. omega. apply no_dup_list_singularize. intro h22. rewrite in_firstn_iff in h22. destruct h22 as [r [h23 h24]]. rewrite nth_lt_map_seg' in h24. apply lind_singularize_eq_nth_lt in h24. revert h24. gen0. intros h24 h25. symmetry in h25. revert h25. gen0. intros h25 h26. rewrite (nth_lt_eq_iff hdt) in h26. destruct h26 as [h27 [occ [h28 h29]]]. symmetry in h29. pose proof (lind_eq_iff hdt (map_in_dep_map_seg_lind_singularize hdt l n hle) (nth_lt (list_singularize hdt l) i (lt_map_seg_lind_singularize hdt l i n hle h3)) h1 (lind_occ hdt l occ (nth_lt l r h24) h27)) as ha. hfirst ha. rewrite <- h29. apply lt_lind. specialize (ha hf0). rewrite ha in h29 at 1. destruct h29 as [h29 h30]. unfold map_in_dep_map_seg_lind_singularize in h29. rewrite nth_lt_map_in_dep' in h29. apply nth_lt_inj in h29. unfold map_seg_lind_singularize in h29. rewrite nth_lt_map_seg' in h29. revert h29. gen0. intros h31 h32. unfold lind_singularize in h32. pose proof (lind_eq_iff hdt (list_singularize hdt l) (nth_lt l (lind_occ hdt l occ (nth_lt l r h24) h27) h31) (iff1 (list_singularize_in_iff hdt l (nth_lt l (lind_occ hdt l occ (nth_lt l r h24) h27) h31)) (in_nth_lt l (lind_occ hdt l occ (nth_lt l r h24) h27) h31)) i (lt_map_seg_lind_singularize hdt l i n hle h3)) as hb. rewrite hb in h32 at 1. destruct h32 as [h32 h33]. symmetry in h32. erewrite nth_lt_functional3 in h32. rewrite lind_occ_compat in h32. erewrite nth_lt_functional3 in h33. rewrite lind_occ_compat in h33. rewrite h32 in h33. revert h23. change (r < lind hdt l (nth_lt _ _ (lt_map_seg_lind_singularize hdt l i n hle h3)) hex -> False). generalize hex. rewrite <- h32. intros h34 h35. pose proof (le_lind_nth_lt hdt l r h24) as h36. pose proof (proof_irrelevance _ h34 (in_nth_lt _ _ _)) as h37. rewrite <- h37 in h36. omega. apply no_dup_list_singularize. Grab Existential Variables. unfold y. rewrite length_map_seg. destruct h5 as [r [h5 h6]]. generalize hex. generalize (lt_map_seg_lind_singularize hdt l i n hle h3). rewrite <- h6. intro he. unfold lind_singularize. erewrite nth_lt_functional3. rewrite lind_compat. intro hg. fold (lind hdt l (nth_lt l r (lt_le_trans r n (length l) h5 hle)) hg). eapply le_trans. erewrite lind_functional3. eapply le_lind_nth_lt. apply lt_impl_le; auto. rewrite length_map_in_dep_map_seg_lind_singularize. destruct h5 as [r [h5 h6]]. generalize hex. generalize (lt_map_seg_lind_singularize hdt l i n hle h3). rewrite <- h6. intro he. unfold lind_singularize. erewrite nth_lt_functional3. rewrite lind_compat. intro hg. fold (lind hdt l (nth_lt l r (lt_le_trans r n (length l) h5 hle)) hg). eapply le_trans. erewrite lind_functional3. eapply le_lind_nth_lt. apply lt_impl_le; auto. rewrite length_map_in_dep_map_seg_lind_singularize. destruct h5 as [r [h5 h6]]. generalize hex. generalize (lt_map_seg_lind_singularize hdt l i n hle h3). rewrite <- h6. intro he. unfold lind_singularize. erewrite nth_lt_functional3. rewrite lind_compat. intro hg. fold (lind hdt l (nth_lt l r (lt_le_trans r n (length l) h5 hle)) hg). eapply le_lt_trans. erewrite lind_functional3. eapply le_lind_nth_lt. assumption. Qed. Lemma inv_inv_im_im_lind_singularize_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfle:n <= length l) (x:T) (pf:Inn (im_im_lind_singularize pfdt l n pfle) x), let pflt := lt_le_trans _ _ _ (lt_inv_inv_im_im_lind_singularize pfdt l n pfle x pf) pfle in nth_lt l _ pflt = x. unfold inv_inv_im_im_lind_singularize; intros. erewrite nth_lt_functional3. apply inv_inv_map_in_dep_map_seg_lind_singularize_compat. Qed. Lemma im_list_to_set_seg_list : forall {U:Type} (pfdu:type_eq_dec U) (n:nat) (f:nat->U), Im (list_to_set (seg_list n)) f = list_to_set (map_seg (fun_to_seg f n)). intros; rewrite list_to_set_map_seg_, list_to_set_seg_list_eq; auto. Qed. (*This is the minimal "firstn sub_list" of [l] that contains all [x]s in [l].*) Definition firstn_all_min {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) := firstn (S (lind_occ pfdt l (pred (count_occ pfdt l x)) x pf)) l. (*This is the minimal "skipn sublist" that contains all [x]s in [l].*) Definition skipn_all_min {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) := skipn (lind pfdt l x pf) l. Lemma last_firstn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), last (firstn_all_min pfdt l x pf) x = x. intros T h1 l x h2. unfold firstn_all_min. erewrite firstn_S_eq. rewrite last_app_sing. rewrite lind_occ_compat; auto. Qed. Lemma in_firstn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), In x (firstn_all_min pfdt l x pf). intros T h1 l x h2. unfold firstn_all_min. rewrite (in_firstn_iff' h1). exists h2. pose proof h2 as h2'. rewrite (count_occ_In h1) in h2' . rewrite (S_pred _ _ h2'). simpl. pose proof (lt_lind_S_occ h1 l x h2 (pred (pred (count_occ h1 l x)))) as h3. destruct h3 as [h3|h3]; auto. destruct (zerop (pred (count_occ h1 l x))) as [h4 | h4]. rewrite h4 in h3. simpl in h3. rewrite h4. unfold lind. apply lt_n_Sn. rewrite (S_pred _ _ h4). simpl. omega. rewrite <- h3. simpl. unfold lind. apply lt_n_Sn. Qed. Lemma in_skipn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), In x (skipn_all_min pfdt l x pf). intros T h1 l x h2. unfold skipn_all_min. rewrite (in_skipn_iff h1). exists h2. exists 0. unfold lind. pose proof h2 as h2'. rewrite (count_occ_In h1) in h2'. split; auto with arith. Qed. Lemma hd_skipn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), hd x (skipn_all_min pfdt l x pf) = x. intros T h1 l x h2. pose proof (in_skipn_all_min h1 _ _ h2) as h3. unfold skipn_all_min in h3. unfold skipn_all_min. induction l as [|a l ih]. contradiction. unfold lind in h3. unfold lind. simpl in h3. simpl. destruct (h1 x a) as [|h4]; subst. simpl; auto. simpl in h3. simpl. destruct h2 as [|h2]; subst. pose proof (h4 eq_refl); contradiction. specialize (ih h2). hfirst ih. unfold lind. erewrite lind_occ_functional. apply h3. reflexivity. reflexivity. reflexivity. specialize (ih hf). unfold lind in ih. erewrite lind_occ_functional. rewrite ih at 1. reflexivity. reflexivity. reflexivity. reflexivity. Qed. Lemma count_occ_firstn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), count_occ pfdt (firstn_all_min pfdt l x pf) x = count_occ pfdt l x. intros T h1 l. induction l as [|a l ih]; simpl; auto. intros x h2. rewrite (switch_eq_dec_eq h1 x a). destruct (h1 a x) as [|h3]; subst. f_equal. unfold switch_eq_dec. simpl. destruct (in_dec h1 x l) as [h3 | h3]. pose proof h3 as h3'. rewrite (count_occ_In h1) in h3'. rewrite (S_pred _ _ h3'). simpl. rewrite <- (S_pred _ _ h3'). apply ih. rewrite (count_occ_In h1) in h3. rewrite (not_lt0_iff (count_occ h1 l x)) in h3. rewrite h3. simpl. reflexivity. destruct h2 as [|h2]; subst. contradict h3; auto. apply ih. Qed. Lemma count_occ_skipn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), count_occ pfdt (skipn_all_min pfdt l x pf) x = count_occ pfdt l x. intros T h1 l. induction l as [|a l ih]; simpl; auto. intros; contradiction. intros x h2. destruct h2 as [|h2]; subst. unfold skipn_all_min, lind; simpl. destruct (h1 x x) as [h3 | h3]. rewrite skipn0. simpl. destruct (h1 x x) as [h3' | h3']; auto. contradiction. contradict h3; auto. specialize (ih _ h2). destruct (h1 a x) as [|h3]; subst. rewrite <- ih. unfold skipn_all_min, lind. simpl. destruct (h1 x x) as [h3 | h3]. simpl. destruct (h1 x x) as [h4 | h4]. f_equal; auto. contradiction. contradict h3; auto. unfold skipn_all_min, lind. simpl. destruct (h1 x a) as [|h4]; subst. contradict h3; auto. simpl. unfold skipn_all_min, lind in ih. rewrite <- ih. repeat f_equal; try apply proof_irrelevance. Qed. Lemma firstn_all_min_cons_in_hd : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf':In x (x::l)) (pf:In x l), firstn_all_min pfdt (x::l) x pf' = x::firstn_all_min pfdt l x pf. intros T h1 l x h2 h3; unfold firstn_all_min; simpl. f_equal. destruct (h1 x x) as [|h4]; simpl. pose proof h3 as h3'. rewrite (count_occ_In h1) in h3'. rewrite (S_pred _ _ h3'). destruct in_dec as [h4 | h4]. simpl. rewrite (hd_compat' _ (in_not_nil _ _ h4)) at 1. rewrite (hd_compat' _ (in_not_nil _ _ h4)) at 6. repeat f_equal; try apply proof_irrelevance. contradiction. contradict h4; auto. Qed. Lemma firstn_all_min_cons_nin_hd : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x (x::l)), ~In x l -> firstn_all_min pfdt (x::l) x pf = x::nil. intros T h1 l x h2 h3; unfold firstn_all_min; simpl. f_equal. destruct (h1 x x) as [|h4]; simpl. rewrite (count_occ_In h1) in h3. rewrite (not_lt0_iff (count_occ h1 l x)) in h3 at 1. rewrite h3; simpl; auto. contradict h4; auto. Qed. Lemma skipn_all_min_cons_hd : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf':In x (x::l)), skipn_all_min pfdt (x::l) x pf' = x::l. intros T h1 l x h2. unfold skipn_all_min. erewrite lind_functional. rewrite lind_hd'. simpl; auto. simpl; auto. Grab Existential Variables. apply cons_neq_nil. Qed. Lemma skipn_all_min_cons_in_last : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf':In x (l ++ x::nil)) (pf:In x l), skipn_all_min pfdt (l ++ x::nil) x pf' = (skipn_all_min pfdt l x pf) ++ x::nil. intros T h1 l x h2 h3. unfold skipn_all_min, lind. rewrite skipn_app. f_equal. erewrite lind_occ_functional. rewrite lind_occ_app_lt_in1 at 1. f_equal. f_equal. apply proof_irrelevance. reflexivity. reflexivity. reflexivity. erewrite lind_occ_functional. rewrite lind_occ_app_lt_in1. apply lt_le_weak. apply lt_lind_occ. reflexivity. reflexivity. reflexivity. Grab Existential Variables. rewrite (count_occ_In h1) in h3; auto with arith. rewrite (count_occ_In h1) in h3; auto with arith. Qed. Lemma skipn_all_min_cons_nin_last : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf':In x (l ++ x::nil)) (pf: ~In x l), skipn_all_min pfdt (l ++ x::nil) x pf' = x::nil. intros T h1 l x h2 h3. pose proof h2 as h2'. apply in_app_or in h2'. destruct h2' as [|h4]. contradiction. unfold skipn_all_min. erewrite lind_functional. rewrite lind_app_nin. rewrite skipn_length_app_plus. rewrite lind_cons_same. simpl; auto. assumption. reflexivity. Grab Existential Variables. left; auto. Qed. Lemma lind_occ_hd_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pfn:l <> nil) (occ:nat), let h := hd' l pfn in let pfh := in_hd' l pfn in forall (pfh':In h (tl l)), lind_occ pfdt l (S occ) h pfh = S (lind_occ pfdt (tl l) occ h pfh'). intros T h1 l h2 occ h h3 h4. destruct l as [|a l]; simpl. contradiction. simpl in h, h3, h4. destruct in_dec as [h7 | h7]. unfold h. unfold h in h4. pose (left (a <> a) (eq_refl a)) as h8. pose proof (classic_eq (h1 a a) h8) as h9. rewrite h9. unfold h8. repeat f_equal; try apply proof_irrelevance. contradiction. Qed. Lemma lind_occ_skipn : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (m:nat) (pf:In x (skipn m l)) (occ:nat), m <= length l -> occ < count_occ pfdt (skipn m l) x -> let pf' := in_skipn _ _ _ pf in lind_occ pfdt (skipn m l) occ x pf = lind_occ pfdt l (occ + count_occ pfdt (firstn m l) x) x pf' - m. intros T h1 l. induction l as [|a l ih]; simpl. intros ? ? hi. pose proof hi as hi'. rewrite skipn_nil in hi'. contradiction. intros x m h2 occ h3 h4. destruct m as [|m]. simpl in h2, h4. simpl. destruct h2 as [|h2]; subst. rewrite eq_dec_same at 1 2. rewrite <- plus_n_O. rewrite <- minus_n_O. reflexivity. assumption. rewrite switch_eq_dec_eq in h4. destruct (h1 x a) as [|h5]; subst. simpl in h4. rewrite <- minus_n_O. rewrite <- plus_n_O. reflexivity. simpl in h4. rewrite <- minus_n_O. f_equal. rewrite <- plus_n_O. f_equal. apply proof_irrelevance. simpl in h2, h4. simpl. apply le_S_n in h3. rewrite (switch_eq_dec_eq h1 a x). destruct (h1 x a) as [|h5]; subst. simpl. rewrite plus_S_shift_r. destruct in_dec as [h5 | h5]. rewrite ih at 1; auto. assert (h5 = in_skipn l m a h2). apply proof_irrelevance. subst; omega. contradict h5. eapply in_skipn; apply h2. simpl. rewrite ih at 1; auto. f_equal. f_equal. apply proof_irrelevance. Qed. Lemma count_occ_removelast_firstn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), count_occ pfdt (removelast (firstn_all_min pfdt l x pf)) x = pred (count_occ pfdt l x). intros T h1 l. induction l as [|a l ih]; simpl; auto. intros x h2. rewrite (switch_eq_dec_eq h1 x a). destruct (h1 a x) as [|h3]; subst. unfold switch_eq_dec. simpl. destruct (in_dec h1 x l) as [h3 | h3]. pose proof h3 as h3'. rewrite (count_occ_In h1) in h3'. rewrite (S_pred _ _ h3'). simpl. specialize (ih x h3). rewrite (hd_compat' _ (in_not_nil _ _ h3)) at 1. simpl. destruct (h1 x x) as [|h4]. f_equal. apply ih. contradict h4; auto. rewrite (count_occ_In h1) in h3. rewrite (not_lt0_iff (count_occ h1 l x)) in h3. rewrite h3. simpl. reflexivity. destruct h2 as [|h2]; subst. contradict h3; auto. unfold switch_eq_dec. pose proof h2 as h2'. pose proof (in_firstn_all_min h1 _ _ h2') as h4. rewrite (count_occ_In h1) in h4. simpl. rewrite (hd_compat' _ (in_not_nil _ _ h2)) at 1. simpl. destruct (h1 a x) as [|h5]; subst. contradict h3; auto. apply ih. Qed. Lemma count_occ_tl_skipn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), count_occ pfdt (tl (skipn_all_min pfdt l x pf)) x = pred (count_occ pfdt l x). intros T h1 l. unfold skipn_all_min, lind. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2. destruct h2 as [|h2]; subst. destruct (h1 x x) as [|h2]; simpl; auto. rewrite (switch_eq_dec_eq h1 a x). destruct (h1 a x) as [|h3]; subst. unfold switch_eq_dec. destruct (h1 x x) as [|h3]. simpl; auto. contradict h3; auto. destruct (h1 x a) as [|h4]; subst. contradict h3; auto. unfold switch_eq_dec. simpl. erewrite <- ih. f_equal. Qed. Lemma pred_count_occ_hd : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), (pred (count_occ pfdt (x :: l) x)) = count_occ pfdt l x. intros T h1 l x. simpl. destruct (h1 x x) as [|h2]; subst; simpl; auto. contradict h2; auto. Qed. Lemma lind_occ_firstn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (occ:nat), occ < count_occ pfdt l x -> let pf' := in_firstn_all_min pfdt _ _ pf in lind_occ pfdt (firstn_all_min pfdt l x pf) occ x pf' = lind_occ pfdt l occ x pf. intros T h1 l x h2 occ h3 h4. unfold firstn_all_min. rewrite lind_occ_firstn. f_equal. apply proof_irrelevance. fold (firstn_all_min h1 l x h2). rewrite count_occ_firstn_all_min; auto. Qed. Lemma in_removelast_firstn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), 1 < count_occ pfdt l x -> In x (removelast (firstn_all_min pfdt l x pf)). intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2. rewrite (switch_eq_dec_eq h1 x a). destruct (h1 a x) as [h3 | h3]; subst. unfold switch_eq_dec. simpl. intro h3. apply lt_S_n in h3. rewrite (S_pred _ _ h3). rewrite <- (count_occ_In h1 l x) in h3. destruct in_dec as [h4 | h4]. simpl. rewrite (hd_compat' _ (in_not_nil _ _ h4)) at 1. left; auto. contradiction. intro h4. destruct h2 as [|h2]. contradiction. unfold switch_eq_dec. simpl. rewrite (hd_compat' _ (in_not_nil _ _ h2)) at 1. right. apply ih; auto. Qed. Lemma firstn_all_min_cons_in_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l), firstn_all_min pfdt (x::l) x (in_eq _ _) = x::firstn_all_min pfdt l x pfin. intros T h1 l x h2. unfold firstn_all_min. simpl. f_equal. destruct (h1 x x) as [|h3]. simpl. pose proof h2 as h2'. rewrite (count_occ_In h1) in h2'. rewrite (S_pred _ _ h2'). destruct in_dec as [h3 | h3]. simpl. pose proof (proof_irrelevance _ h2 h3); subst; auto. contradiction. contradict h3; auto. Qed. Lemma lind_occ_lt_length_firstn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l) (occ:nat), occ < count_occ pfdt l x -> lind_occ pfdt l occ x pfin < length (firstn_all_min pfdt l x pfin). intros T h1 l x h2 occ h3. pose proof (in_firstn_all_min h1 _ _ h2) as h4. rewrite <- lind_occ_firstn_all_min. apply lt_lind_occ. assumption. Qed. Lemma lind_removelast_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf':In x (removelast l)), let pf := removelast_in _ _ pf' in lind pfdt (removelast l) x pf' = lind pfdt l x pf. intros T h1 l x h2 h5. pose proof (app_removelast_last x (in_not_nil _ _ (removelast_in l x h2))) as h3. pose proof h5 as h5'. rewrite h3 in h5'. pose proof (lind_functional' h1 _ _ _ _ h5 h5' h3 eq_refl) as h6. symmetry. rewrite h6. erewrite lind_functional. rewrite lind_app_in1. reflexivity. reflexivity. Qed. Lemma lind_occ_removelast_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf':In x (removelast l)) (occ:nat), occ < count_occ pfdt (removelast l) x -> let pf := removelast_in _ _ pf' in lind_occ pfdt (removelast l) occ x pf' = lind_occ pfdt l occ _ pf. intros T h1 l x h2 occ h3 h4. pose proof (app_removelast_last x (in_not_nil _ _ (removelast_in l x h2))) as h5. generalize h4, h2. rewrite h5, removelast_app_sing. intros h4' h2'. pose proof (lind_occ_app_lt_in1 h1 (removelast l) (last l x::nil) _ _ h3) as h6. simpl in h6. symmetry. erewrite lind_occ_functional. rewrite h6 at 1. f_equal. apply proof_irrelevance. reflexivity. reflexivity. reflexivity. Qed. Lemma lind_occ_lt_length_removelast_firstn_all_min : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l) (occ:nat), S occ < count_occ pfdt l x -> lind_occ pfdt l occ x pfin < length (removelast(firstn_all_min pfdt l x pfin)). intros T h1 l x h2 occ h3. pose proof (in_removelast_firstn_all_min h1 l x h2) as h4. hfirst h4. omega. specialize (h4 hf). pose proof (lt_lind_occ h1 _ occ _ h4) as h5. rewrite lind_occ_removelast_in in h5. erewrite lind_occ_functional in h5. rewrite lind_occ_firstn_all_min in h5 at 1. apply h5. omega. reflexivity. reflexivity. reflexivity. rewrite count_occ_removelast_firstn_all_min. omega. Qed. Lemma count_occ_firstn_removelast : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), let lr := removelast l in n <= length lr -> count_occ pfdt (firstn n lr) x = count_occ pfdt (firstn n l) x. simpl; intros; rewrite firstn_removelast; auto. Qed. Lemma count_occ_firstn_lind_occ_S : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (occ:nat), S occ < count_occ pfdt l x -> count_occ pfdt (firstn (lind_occ pfdt l (S occ) x pf) l) x = S (count_occ pfdt (firstn (lind_occ pfdt l occ x pf) (firstn_all_min pfdt l x pf)) x). intros T h1 l. induction l as [|a l ih]; simpl. intros; omega. intros x h2 occ h3. rewrite switch_eq_dec_eq in h3. destruct (h1 x a) as [|h4]; subst. unfold switch_eq_dec in h3. apply lt_S_n in h3. destruct in_dec as [h4 | h4]. destruct occ as [|occ]; simpl. destruct (h1 a a) as [|h5]. simpl. f_equal. fold (lind h1 l a h4). rewrite count_occ_firstn_eq0_iff. intros h4'. pose proof (proof_irrelevance _ h4 h4'); subst; auto. contradict h5; auto. destruct (h1 a a) as [|h5]. f_equal. rewrite ih; auto. f_equal. f_equal. simpl. rewrite (S_pred _ _ h3). destruct in_dec as [h5 | h5]. pose proof (proof_irrelevance _ h4 h5); subst. reflexivity. contradiction. contradict h5; auto. contradict h4. rewrite (count_occ_In h1). omega. unfold switch_eq_dec in h3. destruct h2 as [|h2]; subst. contradict h4; auto. specialize (ih _ h2 _ h3). pose proof (lt_lind_occ h1 l (S occ) x (neq_in_cons l x a h4 (or_intror h2))) as h5. simpl. rewrite (switch_eq_dec_eq h1 x a). destruct (h1 a x) as [|h6]; subst. pose proof (h4 eq_refl); contradiction. unfold switch_eq_dec. pose proof (proof_irrelevance _ h2 (neq_in_cons l x a h4 (or_intror h2))) as h7. rewrite <- h7. rewrite ih. f_equal. f_equal. unfold firstn_all_min. f_equal. f_equal. f_equal. f_equal. apply proof_irrelevance. Qed. Lemma lind_occ_eq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (occ n:nat) (pfin:In x l) (pflt:n < length l), occ < count_occ pfdt l x -> (lind_occ pfdt l occ x pfin = n <-> nth_lt l n pflt = x /\ count_occ pfdt (firstn (S n) l) x = S occ /\ count_occ pfdt (skipn (S n) l) x = count_occ pfdt l x - (S occ)). intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x occ n h2 h3 h4. rewrite switch_eq_dec_eq in h4. destruct h2 as [|h2]; subst. rewrite (eq_dec_same h1) at 1 2 3. rewrite (eq_dec_same h1) in h4 at 1. simpl in h4. destruct occ as [|occ]. split; intro h5; subst; simpl. split; auto with arith. simpl in h5. destruct n as [|n]; auto. destruct h5 as [h5 [h6 h7]]. simpl in h6, h7. destruct l as [|a l]. simpl in h3. omega. simpl in h3. apply lt_S_n in h3. simpl in h4, h5, h6, h7. destruct (h1 a x) as [|h8]. apply S_inj in h6. discriminate. apply S_inj in h6. destruct n as [|n]. contradiction. apply lt_S_n in h3. pose proof (lt_length _ _ h3) as h9. erewrite nth_indep in h5. rewrite <- (nth_lt_compat _ _ h3) in h5 at 1. erewrite nth_lt_functional3 in h5. pose proof (nth_lt_firstn_eq l n (S n)) as h10. hfirst h10. apply lt_length_firstn; auto with arith. specialize (h10 hf). pose proof (in_nth_lt (firstn (S n) l) n hf) as h11. rewrite (count_occ_In h1) in h11. rewrite h10 in h11. erewrite nth_lt_functional3 in h5 at 1. rewrite h5 in h11 at 1. omega. assumption. apply lt_S_n in h4. destruct in_dec as [h5 | h5]. destruct n as [|n]. simpl. intuition. apply lt_S_n in h3. specialize (ih x occ n h5 h3 h4). split; intro h6. apply S_inj in h6. rewrite ih in h6. destruct h6 as [h6 [h7 h8]]. split. erewrite nth_indep. rewrite <- nth_lt_compat at 1. apply h6. assumption. split. f_equal; auto. omega. f_equal. rewrite ih. destruct h6 as [h6 [h7 h8]]. split; auto. erewrite nth_indep in h6 at 1. rewrite <- nth_lt_compat in h6 at 1. apply h6. assumption. apply count_occ_in in h4. contradiction. destruct n as [|n]. destruct (h1 x a) as [|h5]; subst. simpl in h4. simpl. rewrite (eq_dec_same h1) at 1 2; simpl. destruct occ as [|occ]. intuition. apply lt_S_n in h4. destruct in_dec as [h5 | h5]. intuition. contradiction. simpl in h4. assert (he:h1 a x = switch_eq_dec _ _ (right (x = a) h5)). unfold switch_eq_dec. destruct (h1 a x); subst. contradict h5; auto. f_equal; apply proof_irrelevance. rewrite he. unfold switch_eq_dec. simpl. intuition. apply lt_S_n in h3. rewrite (switch_eq_dec_eq h1 a x). destruct (h1 x a) as [|h5]; subst. simpl in h4. rewrite (eq_dec_same h1) at 1 2. destruct occ as [|occ]. split; intro h5. discriminate. destruct h5 as [h5 [h6 h7]]. specialize (ih a 0 n h2 h3). hfirst ih. rewrite <- (count_occ_In h1 l a); auto. specialize (ih hf). apply S_inj in h6. let P := type of ih in match P with ?hyp <-> _ => assert (h10:~ hyp) end. intro h10. rewrite ih in h10. destruct h10; omega. revert h10. generalize h2. rewrite <- h5. intros h10 h11. pose proof (le_lind_nth_lt h1 _ _ h3) as h12. unfold lind in h12. revert h12. generalize (in_nth_lt l n h3). erewrite nth_lt_compat. intros h13 h14. assert (h15: lind_occ h1 l 0 (nth n l a) h10 = lind_occ h1 l 0 (nth n l (nth_lt l n h3)) h13). apply lind_occ_functional; auto. apply nth_indep. assumption. rewrite h15 in h11. assert (h16:lind_occ h1 l 0 (nth n l (nth_lt l n h3)) h13 < n). omega. apply not_lt0_iff in h6. contradict h6. rewrite <- (count_occ_In h1 (firstn (S n) l) a). rewrite in_firstn_iff'. exists h2. unfold lind. eapply lt_trans. erewrite lind_occ_functional. apply h16. reflexivity. reflexivity. rewrite <- h5. apply nth_indep. assumption. auto with arith. apply lt_S_n in h4. destruct in_dec as [h5 | h5]. specialize (ih a occ n h2 h3 h4). split; intro h6. apply S_inj in h6. pose proof (proof_irrelevance _ h2 h5) as h7. rewrite <- h7 in h6. rewrite ih in h6. destruct h6 as [h6 [h8 h9]]. subst. split. rewrite nth_lt_compat. apply nth_indep. assumption. rewrite h8, h9. split; auto. f_equal. pose proof (proof_irrelevance _ h2 h5) as h7. rewrite <- h7. rewrite ih. destruct h6 as [h6 [h8 h9]]. split. rewrite <- h6. rewrite nth_lt_compat. apply nth_indep. assumption. apply S_inj in h8. rewrite h8, h9. intuition. contradiction. simpl in h4. simpl. specialize (ih x occ n h2 h3 h4). split; intro h6. apply S_inj in h6. pose proof (proof_irrelevance _ h2 (neq_in_cons l x a h5 (or_intror h2))) as h8. rewrite h8 in ih. rewrite ih in h6. destruct h6 as [h6 [h9 h10]]. split. rewrite <- h6. rewrite nth_lt_compat. apply nth_indep. assumption. destruct l; tauto. f_equal. pose proof (proof_irrelevance _ h2 (neq_in_cons l x a h5 (or_intror h2))) as h8. rewrite <- h8. rewrite ih. destruct h6 as [h6 [h9 h10]]. split. rewrite nth_lt_compat. rewrite <- h6. apply nth_indep. assumption. destruct l as [|c l]. tauto. simpl; simpl in h9, h10. destruct (h1 c x) as [h11 | h11]. tauto. tauto. Grab Existential Variables. assumption. Qed. Lemma lind_occ_hd'' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (occ:nat) (pfn:l <> nil), let pfin := in_hd' _ pfn in forall (pfin':In (hd' _ pfn) (tl l)), let pfin'' := in_not_nil _ _ pfin' in occ < count_occ pfdt l (hd' _ pfin'') -> occ = 0 \/ lind_occ pfdt l occ _ pfin = S (lind_occ pfdt (tl l) (pred occ) _ pfin'). intros T h1 l occ. revert l. destruct occ as [|occ]; simpl. intros; left; auto. intros l h2 h3 h4. right. pose proof (hd_compat' _ h2) as h5; simpl in h5. pose proof (lind_occ_functional h1 l _ (S occ) (S occ) _ _ (in_hd' _ h2) (in_eq (hd' l h2) (tl l)) h5 eq_refl eq_refl) as h6. rewrite h6; simpl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]. f_equal. f_equal. apply proof_irrelevance. contradiction. contradict hr; auto. Qed. Lemma in_firstn_S_lind_occ : forall {T:Type} (h1:type_eq_dec T) (l:list T) (x:T) (occ:nat) (pf:In x l), In x (firstn (S (lind_occ h1 l occ x pf)) l). intros T h1 l x occ h2. rewrite (in_firstn_iff' h1). exists h2. pose proof (le_lind_occ h1 l x h2 occ) as h4. omega. Qed. Lemma lind_occ_lrep : forall {T:Type} (pfdt:type_eq_dec T) (a x:T) (n:nat) (occ:nat) (pf:In x (lrep a n)), occ < n -> lind_occ pfdt (lrep a n) occ x pf = occ. intros T h1 a x n. revert a x. induction n as [|n ih]; simpl. intros; contradiction. intros a x occ h2 h3. lr_match_goal; fold hlr; destruct hlr as [hl | hr]; subst. destruct occ as [|occ]; auto. apply lt_S_n in h3. lr_match_goal; fold hlr; destruct hlr as [hl | hr]; subst. f_equal; auto. contradict hr. apply in_lrep. intro h4. apply lrep_nil_iff in h4; subst. omega. destruct h2 as [|h2]; subst. contradict hr; auto. pose proof h2 as h2'. apply in_lrep' in h2'. contradiction. Qed. Lemma lind_occ_firstn_pred_count_occ_nth_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pfsm:S m <= length l) (pfin:In (nth_lt l m pfsm) (firstn (S m) l)), lind_occ pfdt (firstn (S m) l) (pred (count_occ pfdt (firstn (S m) l) (nth_lt l m pfsm))) (nth_lt l m pfsm) pfin = m. intros T h1 l m h2 h3. rewrite lind_occ_eq_iff. split. rewrite nth_lt_firstn_eq. apply nth_lt_functional3. split. rewrite (count_occ_In h1) in h3. rewrite <- (S_pred _ _ h3). f_equal. apply firstn_idem. rewrite (count_occ_In h1) in h3. rewrite <- (S_pred _ _ h3). rewrite minus_same. rewrite count_occ_skipn. rewrite firstn_idem. apply minus_same. rewrite (count_occ_In h1) in h3. rewrite (S_pred _ _ h3); auto with arith. Grab Existential Variables. rewrite length_firstn; auto. Qed. Lemma lS_app : forall (l l':list nat), lS (l++l') = lS l ++ lS l'. intros l l'. unfold lS. rewrite map_eq_iff. ex_goal. do 2 rewrite app_length. do 2 rewrite map_length; auto. exists hex. intros n h1. pose proof h1 as h1'. apply lt_length_app_dec in h1'. destruct h1' as [h2 | h2]. rewrite (nth_lt_app1 _ _ _ h2). erewrite nth_lt_app1. erewrite nth_lt_map. reflexivity. destruct (nil_dec' l') as [|h3]; subst. simpl in h2. omega. pose proof h1 as h1'. rewrite app_length in h1'. destruct (lt_le_dec n (length l)) as [h4 | h4]. rewrite (nth_lt_app1 _ _ _ h4). erewrite nth_lt_app1. erewrite nth_lt_map. reflexivity. rewrite (nth_lt_app2 _ _ _ h4). erewrite nth_lt_app2. erewrite nth_lt_map. f_equal. apply nth_lt_functional2. rewrite map_length; auto. Grab Existential Variables. rewrite map_length; auto. rewrite map_length; auto. rewrite map_length; auto. rewrite map_length; auto. Qed. Lemma lS_map_seg : forall {T:Type} (n:nat) (f:seg_fun n nat), lS (map_seg f) = map_seg (fun c pf => S (f c pf)). intros T n. induction n as [|? ih]; simpl; auto. intro. rewrite lS_app, ih; auto. Qed. Lemma faml_lpred_famlS : forall (ll:faml nat), faml_lpred (famlS ll) = ll. unfold faml_lpred, famlS; intros; rewrite map_map. rewrite map_eq_iff. exists eq_refl. intros n h1. rewrite lpred_lS. f_equal. apply proof_irrelevance. Qed. Lemma nin_famlS : forall (ll:faml nat) (ln:list nat), In 0 ln -> ~In ln (famlS ll). intros ll ln h1 h2. rewrite in_famlS_iff in h2. destruct h2 as [ln' [h2 h3]]; subst. rewrite in_lS_iff in h1. destruct h1 as [? [? ?]]; discriminate. Qed. Lemma famlS_eq_iff : forall (ll ll':faml nat), famlS ll = ll' <-> ((forall ln, In ln ll' -> ~In 0 ln) /\ ll = faml_lpred ll'). intros ll ll'. split; intro h1; subst. split. intros ln h1. rewrite in_famlS_iff in h1. destruct h1 as [ln' [h2 h3]]; subst. intro h3. rewrite in_lS_iff in h3. destruct h3 as [? [? ?]]; discriminate. rewrite faml_lpred_famlS; auto. destruct h1 as [h1 h2]; subst. unfold famlS, faml_lpred. rewrite map_map. rewrite map_eq_iff. exists eq_refl. intros n h2. specialize (h1 _ (in_nth_lt ll' n h2)). pose proof (lS_lpred_iff (nth_lt ll' n h2)) as h3. rewrite h3 in h1. rewrite <- h1. f_equal. apply proof_irrelevance. Qed. Lemma cardinal_length_compat : forall {T:Type} (l:list T), NoDup l -> cardinal _ (list_to_set l) (length l). intros T l. induction l as [|a l h1]. intros; simpl. constructor. simpl; intro h2. inversion h2 as [|x l' h4 h5]. clear h2. subst. constructor. apply h1; auto. rewrite <- list_to_set_in_iff. assumption. Qed. Lemma le_card_fun1_length : forall {T:Type} (l:list T), card_fun1 (list_to_set l) <= length l. intros T l. induction l as [|a l h1]; simpl. rewrite card_fun1_empty. auto with arith. destruct (classic_dec (In a l)) as [h2 | h2]. rewrite list_to_set_in_iff in h2. rewrite in_add_eq; auto. rewrite card_add_nin'. apply le_n_S. assumption. apply list_to_set_finite. rewrite <- list_to_set_in_iff. assumption. Qed. Lemma card_fun1_length_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), card_fun1 (list_to_set l) = length (list_singularize pfdt l). intros T hdt l. induction l as [|a l h1]; simpl. rewrite card_fun1_empty. reflexivity. destruct (in_dec hdt a l) as [h2 | h2]. rewrite in_add_eq; auto. rewrite list_singularize_in_cons; auto. rewrite <- list_to_set_in_iff; assumption. rewrite list_singularize_nin; auto. rewrite card_add_nin'; simpl. rewrite h1; auto. apply list_to_set_finite. rewrite <- list_to_set_in_iff; auto. Qed. Lemma list_to_set_no_dup_length : forall {T:Type} (l:list T) (E:Ensemble T), list_to_set l = E -> NoDup l -> exists !n, length l = n. intros T l. induction l as [|a l h1]. simpl. intros. exists 0. red. split; auto. simpl. intros E h2 h3. inversion h3 as [|x l' h4 h5]. subst. clear h3. specialize (h1 _ (eq_refl _ )). specialize (h1 h5). destruct h1 as [n h1]. red in h1. destruct h1 as [h1 h2]. exists (S n). red. split. f_equal. assumption. intros x h6. specialize (h2 _ (eq_refl _)). rewrite <- h2 in h6. assumption. Qed. Lemma list_to_set_maps_eq_length : forall {T U V:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (f:T->V) (g:U->V) (l:list T) (l':list U), NoDup l -> NoDup l' -> Injective f -> Injective g -> list_to_set (map f l) = list_to_set (map g l') -> length l = length l'. intros T U V hdt hdu hdv f g l. induction l as [|a l ih]; simpl. intros ? ? ? ? ? h1. symmetry in h1. apply empty_set_nil in h1. rewrite map_eq_nil_iff in h1; subst; auto. intro l'. destruct l' as [|a' l']. intros ? ? ? ? h1. simpl in h1. apply add_not_empty in h1. contradiction. simpl. intros h1 h2 h3 h4 h5. destruct (hdv (f a) (g a')) as [h6 | h6]. rewrite h6 in h5. apply add_nin_inj_eq in h5. f_equal. apply ih; auto. apply no_dup_cons in h1; auto. apply no_dup_cons in h2; auto. intro h7. rewrite <- h6 in h7. rewrite <- list_to_set_in_iff in h7. rewrite in_map_iff in h7. destruct h7 as [x [h7 h8]]. apply h3 in h7; subst. apply no_dup_cons_nin in h1. contradiction. intro h7. rewrite <- list_to_set_in_iff in h7. rewrite in_map_iff in h7. destruct h7 as [x [h7 h8]]. apply h4 in h7; subst. apply no_dup_cons_nin in h2. contradiction. pose proof (f_equal (fun S => Subtract S (f a)) h5) as h7. simpl in h7. rewrite sub_add_same_nin in h7. pose proof (Add_intro2 _ (list_to_set (map f l)) (f a)) as h8. rewrite h5 in h8. inversion h8 as [? h9 | ? h9]; subst. pose proof h9 as h10. rewrite <- list_to_set_in_iff in h10. rewrite in_map_iff in h10. destruct h10 as [x [h11 h12]]. specialize (ih (remove hdu x (a'::l')) (no_dup_cons _ _ h1)). hfirst ih. apply no_dup_remove. assumption. specialize (ih hf h3 h4). hfirst ih. rewrite h7. rewrite <- h11. symmetry. rewrite list_to_set_map_remove. reflexivity. assumption. right; auto. assumption. apply ih in hf0. rewrite length_remove_no_dup in hf0. simpl in hf0; auto with arith. assumption. right; auto. destruct h9. contradict h6; auto. intro h8. rewrite <- list_to_set_in_iff in h8. rewrite in_map_iff in h8. destruct h8 as [c [h8 h9]]; subst. apply h3 in h8; subst. apply no_dup_cons_nin in h1. contradiction. Qed. Lemma list_to_set_map_in_deps_eq_length : forall {T U V:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (l:list T) (l':list U) (f:fun_in_dep l V) (g:fun_in_dep l' V), NoDup l -> NoDup l' -> inj_dep f -> inj_dep g -> list_to_set (map_in_dep f) = list_to_set (map_in_dep g) -> length l = length l'. intros T U V hdt hdu hdv l. induction l as [|a l ih]; simpl. intros ? ? ? ? ? ? ? h1. symmetry in h1. apply empty_set_nil in h1. rewrite map_in_dep_eq_nil_iff in h1; subst; auto. intro l'. destruct l' as [|a' l']. intros ? ? ? ? ? ? h1. simpl in h1. apply add_not_empty in h1. contradiction. simpl. intros f g h1 h2 h3 h4 h5. destruct (hdv (f a (in_eq _ _)) (g a' (in_eq _ _))) as [h6 | h6]. rewrite h6 in h5. apply add_nin_inj_eq in h5. f_equal. eapply ih; auto. apply no_dup_cons in h1; auto. apply no_dup_cons in h2; auto. eapply inj_from_cons. apply h3. eapply inj_from_cons. apply h4. assumption. intro h7. rewrite <- h6 in h7. rewrite <- list_to_set_in_iff in h7. rewrite in_map_in_dep_iff in h7. destruct h7 as [x [hi h7]]. apply h3 in h7; subst. apply no_dup_cons_nin in h1. contradiction. intro h7. rewrite <- list_to_set_in_iff in h7. rewrite in_map_in_dep_iff in h7. destruct h7 as [x [hi h7]]. apply h4 in h7; subst. apply no_dup_cons_nin in h2. contradiction. pose proof (f_equal (fun S => Subtract S (f a (in_eq _ _))) h5) as h7. simpl in h7. rewrite sub_add_same_nin in h7. pose proof (Add_intro2 _ (list_to_set (map_in_dep (fun_from_cons f))) (f a (in_eq _ _))) as h8. rewrite h5 in h8. inversion h8 as [? h9 | ? h9]; subst. pose proof h9 as h10. rewrite <- list_to_set_in_iff in h10. rewrite in_map_in_dep_iff in h10. destruct h10 as [x [hi h11]]. specialize (ih (remove hdu x (a'::l')) (fun_from_cons f) (fun_in_dep_remove hdu g x) (no_dup_cons _ _ h1)). hfirst ih. apply no_dup_remove. assumption. specialize (ih hf (inj_from_cons _ h3) (inj_dep_remove hdu g x h4)). hfirst ih. rewrite h7. rewrite <- h11. symmetry. erewrite list_to_set_map_in_dep_remove_inj. reflexivity. assumption. apply ih in hf0. rewrite length_remove_no_dup in hf0. simpl in hf0; auto with arith. assumption. right; auto. destruct h9. contradict h6; auto. intro h8. rewrite <- list_to_set_in_iff in h8. rewrite in_map_in_dep_iff in h8. destruct h8 as [c [hi h8]]; subst. apply h3 in h8; subst. apply no_dup_cons_nin in h1. contradiction. Qed. Definition list_fun (T U:Type) : Type := list T -> list U. Definition list_fun1 (T:Type) : Type := list_fun T T. (*Gets an incresaing list of all the indices of a given element in the given list.*) Definition linds_occ' {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) := map_seg (fun (n:nat) (pflt:n < count_occ pfdt l x) => lind_occ pfdt l n x pf). Definition linds_occ {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) : list nat := match (in_dec pfdt x l) with | left pf => linds_occ' pfdt l x pf | right _ => nil end. Lemma linds_occ_functional : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x x':T), x = x' -> l = l' -> linds_occ pfdt l x = linds_occ pfdt l' x'. intros; subst; auto. Qed. Lemma linds_occ_functional' : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x x':T) (pfx:In x l) (pfx':In x' l'), x = x' -> l = l' -> linds_occ' pfdt l x pfx = linds_occ' pfdt l' x' pfx'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma linds_occ_compat' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x :T) (pf:In x l), linds_occ pfdt l x = linds_occ' pfdt l x pf. intros T h1 l x h2. unfold linds_occ, linds_occ'. destruct in_dec as [h3 | h3]. f_equal. apply functional_extensionality_dep. intro a. apply functional_extensionality. intro ha. f_equal. apply proof_irrelevance. contradiction. Qed. Lemma in_linds_occ_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), In n (linds_occ pfdt l x) <-> exists (m:nat) (pf:In x l), m < count_occ pfdt l x /\ lind_occ pfdt l m x pf = n. unfold linds_occ; intros T hdt l x occ. lr_match_goal; fold hlr; destruct hlr as [hl | hr]; simpl. unfold linds_occ'. rewrite in_map_seg_iff. split; intros h1. destruct h1 as [m [h3 h2]]; subst. exists m, hl; auto. destruct h1 as [m h1]. destruct h1 as [h1 [h2 h3]]. pose proof (proof_irrelevance _ h1 hl); subst. exists m, h2; auto. split; intro h1; try destruct h1 as [? [? ?]]; contradiction. Qed. Lemma in_linds_occ_iff' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat) (pf:In x l), (In n (linds_occ' pfdt l x pf) <-> exists (m:nat), m < count_occ pfdt l x /\ lind_occ pfdt l m x pf = n). intros T h1 l x n h2; split; intro h3. rewrite <- linds_occ_compat' in h3. rewrite in_linds_occ_iff in h3. destruct h3 as [m [h3 [h4 h5]]]; subst. exists m. split; auto. f_equal; apply proof_irrelevance. rewrite <- linds_occ_compat'. destruct h3 as [m [h3 h4]]; subst. rewrite in_linds_occ_iff. exists m, h2. split; auto. Qed. Lemma in_lind_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), In (lind pfdt l x pf) (linds_occ pfdt l x). intros T h1 l x h2. rewrite in_linds_occ_iff. exists 0, h2. split. rewrite (count_occ_In h1) in h2. apply gt_lt; auto. reflexivity. Qed. Lemma in_lind_occ_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (m:nat), m < count_occ pfdt l x -> In (lind_occ pfdt l m x pf) (linds_occ pfdt l x). intros T h1 l x h2 m hm. rewrite in_linds_occ_iff. exists m, h2. split; auto. Qed. Lemma in_linds_occ_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat), In n (linds_occ pfdt l x) -> n < length l. intros T h1 l x n h2. rewrite in_linds_occ_iff in h2. destruct h2 as [m [h2 [h3 h4]]]. subst. apply lt_lind_occ. Qed. Lemma linds_occ_nil : forall {T:Type} (pfdt:type_eq_dec T) (x:T), linds_occ pfdt nil x = nil. unfold linds_occ; intros; destruct in_dec; auto. Qed. Lemma linds_occ_nin : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), ~ In a l -> linds_occ pfdt l a = nil. intros T h1 l a h2. unfold linds_occ. destruct in_dec as [|h3]. contradiction. reflexivity. Qed. Lemma linds_occ_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), In a l -> linds_occ pfdt l a <> nil. intros T h1 l a h2. unfold linds_occ. destruct in_dec. unfold linds_occ'. intro h3. rewrite map_seg_eq_iff in h3. destruct h3 as [h3 h4]. simpl in h3. rewrite (count_occ_In h1) in h2. omega. contradiction. Qed. Lemma linds_occ_cons_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), linds_occ pfdt (a::l) a = 0::nil ++ lS (linds_occ pfdt l a). intros T h1 l a; unfold linds_occ, linds_occ'; simpl. rewrite eq_dec_eq at 1; auto. rewrite map_seg_eq_iff. ex_goal. rewrite eq_dec_eq at 1; auto; simpl. f_equal. rewrite length_lS. lr_match_goal'; simpl. rewrite length_map_seg; auto. pose proof (count_occ_In h1 l a) as h2. rewrite h2 in hr. apply not_lt_le in hr; apply le0 in hr; auto. exists hex. intros m h2; simpl. rewrite eq_dec_eq at 1; auto. destruct m as [|m]; auto. rewrite eq_dec_eq in h2 at 1; simpl in h2; auto. apply lt_S_n in h2. lr_match_goal'. gtermr. redterm1 y. pose proof (nth_lt_compat c m) as h3. hfirst h3. unfold c; rewrite length_lS, length_map_seg; auto. specialize (h3 hf); simpl in h3. fold c. erewrite nth_indep. rewrite <- h3 at 1. unfold c. rewrite nth_lt_lS, nth_lt_map_seg'; auto. assumption. pose proof (count_occ_In h1 l a) as h3. rewrite h3 in hr. apply not_lt_le in hr. apply le0 in hr. omega. Qed. Lemma linds_occ_cons_eq' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:In a l), linds_occ' pfdt (a::l) a (in_eq _ _) = 0::nil ++ lS (linds_occ' pfdt l a pf). intros T h1 l a h2. pose proof (linds_occ_cons_eq h1 l a) as h3. unfold linds_occ in h3. destruct in_dec as [h4 | h4], in_dec as [h5 | h5]. simpl in h3. simpl. pose proof (proof_irrelevance _ h2 h5); subst. rewrite <- h3. f_equal. apply proof_irrelevance. contradiction. discriminate. discriminate. Qed. Lemma linds_occ_cons_eq_nin' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:~In a l), linds_occ' pfdt (a::l) a (in_eq _ _) = 0::nil. intros T h1 l. induction l as [|b l ih]; simpl. intros. unfold linds_occ'. rewrite map_seg_eq_iff. simpl. destruct (h1 a a) as [h2 | h2]. exists eq_refl. intros m h3. assert (m = 0). omega. subst. reflexivity. contradict h2; auto. intros a h2. specialize (ih a). unfold linds_occ'. rewrite map_seg_eq_iff. ex_goal. simpl. destruct (h1 a a) as [h3 | h3]. destruct (h1 b a) as [h4 | h4]; subst. contradict h2. left; auto. f_equal. assert (h5:~In a l). contradict h2. right; auto. rewrite (count_occ_In h1) in h5; omega. contradict h3; auto. exists hex. intros m h3 h4. pose proof h3 as h3'. rewrite hex in h3'. simpl in h3'. assert (h5:m = 0). omega. generalize h4. rewrite h5. intros h6. simpl. destruct (h1 a a) as [h7 | h7]; auto. contradict h7; auto. Qed. Lemma linds_occ_cons_neq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a x:T), x <> a -> linds_occ pfdt (a::l) x = lS (linds_occ pfdt l x). intros T h1 l a x h2; unfold linds_occ, linds_occ'; simpl. destruct (h1 a x) as [h3 | h3]; subst. contradict h2; auto. destruct (in_dec h1 x l) as [h4 | h4]. rewrite map_seg_eq_iff. ex_goal. rewrite length_lS, length_map_seg; auto. exists hex. intros m h5; simpl. destruct (h1 x a) as [h6 | h6]; subst. contradict h2; auto. unfold lS. erewrite nth_lt_map. f_equal. rewrite nth_lt_map_seg. apply lind_occ_functional; auto. auto. Grab Existential Variables. assumption. Qed. Lemma linds_occ_cons_neq' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a x:T) (pfx:In x l), let pfx' := in_cons _ _ _ pfx in x <> a -> linds_occ' pfdt (a::l) x pfx' = lS (linds_occ' pfdt l x pfx). intros T h1 l a x h2 h3 h4. pose proof (linds_occ_cons_neq h1 l a x h4) as h5. unfold linds_occ in h5. destruct in_dec as [h6 | h6], in_dec as [h7 | h7]. pose proof (proof_irrelevance _ h2 h7). unfold h3. clear h3. subst. rewrite <- h5. f_equal. apply proof_irrelevance. contradiction. contradict h6; right; auto. contradict h6; right; auto. Qed. Lemma linds_occ_app_nin_l' : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T) (pfxn:~In x l) (pfxin:In x l'), let pfx := in_app_r pfxin l in linds_occ' pfdt (l++l') x pfx = map (plus (length l)) (linds_occ' pfdt l' x pfxin). intros T h1 l l' x h2 h3 h4. unfold linds_occ'. rewrite (count_occ_In h1) in h2. rewrite (not_lt0_iff (count_occ h1 l x)) in h2. rewrite map_seg_eq_iff. ex_goal. rewrite count_occ_app. rewrite map_length. rewrite length_map_seg. rewrite h2; auto with arith. exists hex. intros m h5 h6. pose proof h5 as h5'. rewrite count_occ_app in h5'. rewrite h2 in h5'; simpl in h5'. erewrite nth_lt_functional3. erewrite nth_lt_map. rewrite nth_lt_map_seg. symmetry. pose proof (lind_occ_app_ge_in2 h1 l l' x h3 m) as h7. hfirst h7. rewrite h2; auto with arith. specialize (h7 hf). simpl in h7. rewrite h2, <- minus_n_O in h7. pose proof (proof_irrelevance _ h4 (in_app_r h3 l)); subst; auto. Grab Existential Variables. assumption. erewrite map_length. rewrite length_map_seg. assumption. Qed. Lemma linds_occ_app_nin_l : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T) (pfxn:~In x l) (pfxin:In x l'), let pfx := in_app_r pfxin l in linds_occ pfdt (l++l') x = map (plus (length l)) (linds_occ pfdt l' x). intros; do 2 erewrite linds_occ_compat'; apply linds_occ_app_nin_l'; auto. Grab Existential Variables. assumption. Qed. Lemma linds_occ_app_nin_r' : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T) (pfxin:In x l) (pfxn:~In x l'), let pfx := in_app_l pfxin l' in linds_occ' pfdt (l++l') x pfx = linds_occ' pfdt l x pfxin. intros T h1 l l' x h2 h3 h4. unfold linds_occ'. rewrite (count_occ_In h1) in h3. rewrite (not_lt0_iff (count_occ h1 l' x)) in h3. rewrite map_seg_eq_iff. ex_goal. rewrite count_occ_app. rewrite h3, <- plus_n_O. rewrite length_map_seg; auto. exists hex. intros m h5; simpl. pose proof h5 as h5'. rewrite count_occ_app in h5'. rewrite h3, <- plus_n_O in h5'. erewrite nth_lt_functional3. rewrite nth_lt_map_seg. pose proof (lind_occ_app_lt_in1 h1 l l' x _ h5') as h6. simpl in h6. pose proof (proof_irrelevance _ h4 (in_app_l (count_occ_in h1 l x m h5') l')) as h7. pose proof (proof_irrelevance _ h2 (count_occ_in h1 l x m h5')) as h8. rewrite h7, h8; auto. Grab Existential Variables. assumption. Qed. Lemma linds_occ_app_nin_r : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T) (pfxin:In x l) (pfxn:~In x l'), let pfx := in_app_l pfxin l' in linds_occ pfdt (l++l') x = linds_occ pfdt l x. intros; do 2 erewrite linds_occ_compat'; apply linds_occ_app_nin_r'; auto. Grab Existential Variables. assumption. Qed. Lemma linds_occ_app_in' : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T) (pfx:In x l) (pfx':In x l'), let pfxa := in_app_l pfx l' in linds_occ' pfdt (l++l') x pfxa = linds_occ' pfdt l x pfx ++ map (plus (length l)) (linds_occ' pfdt l' x pfx'). intros T h1 l l' x h2 h3 h4. unfold linds_occ'. rewrite map_seg_eq_iff. ex_goal. rewrite count_occ_app. rewrite app_length. rewrite map_length. do 2 rewrite length_map_seg; auto. exists hex. intros m h5 h6. pose proof h6 as h6'. rewrite app_length in h6'. rewrite map_length in h6'. do 2 rewrite length_map_seg in h6'. destruct (lt_le_dec m (count_occ h1 l x)) as [h7 | h7]. erewrite nth_lt_app1. erewrite nth_lt_functional3. erewrite nth_lt_map_seg. erewrite lind_occ_functional. rewrite lind_occ_app_lt_in1. apply lind_occ_functional; auto. reflexivity. reflexivity. reflexivity. erewrite nth_lt_app2. erewrite nth_lt_map. rewrite nth_lt_map_seg. rewrite length_map_seg. erewrite lind_occ_functional. erewrite lind_occ_app_ge_in2. reflexivity. assumption. reflexivity. reflexivity. reflexivity. Grab Existential Variables. rewrite length_map_seg. omega. rewrite length_map_seg. assumption. assumption. assumption. rewrite length_map_seg; auto. Qed. Lemma linds_occ_app_in : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T) (pfx:In x l) (pfx':In x l'), let pfxa := in_app_l pfx l' in linds_occ pfdt (l++l') x = linds_occ pfdt l x ++ map (plus (length l)) (linds_occ pfdt l' x). intros; do 3 erewrite linds_occ_compat'; apply linds_occ_app_in'; auto. Grab Existential Variables. assumption. assumption. Qed. Lemma length_linds_occ' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfx:In x l), length (linds_occ' pfdt l x pfx) = count_occ pfdt l x. intros T h1 l. induction l as [|a l ih]; simpl; auto. intros x h2. destruct (h1 a x) as [|h3]; subst. destruct (in_dec h1 x l) as [ha | ha]. erewrite linds_occ_functional'. rewrite linds_occ_cons_eq' at 1. simpl. f_equal. rewrite length_lS. apply ih. reflexivity. reflexivity. erewrite linds_occ_functional'. rewrite linds_occ_cons_eq_nin'. simpl. f_equal. rewrite (count_occ_In h1) in ha. omega. apply ha. reflexivity. reflexivity. destruct h2 as [|h2]; subst. contradict h3; auto. specialize (ih _ h2). erewrite linds_occ_functional'. rewrite linds_occ_cons_neq'. rewrite length_lS. apply ih. contradict h3; rewrite h3; auto. reflexivity. reflexivity. Grab Existential Variables. assumption. Qed. Lemma length_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), length (linds_occ pfdt l x) = count_occ pfdt l x. intros T h1 l x. unfold linds_occ. destruct in_dec as [h2 | h2]. apply length_linds_occ'. simpl. rewrite (count_occ_In h1) in h2. omega. Qed. Lemma linds_occ_eq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (ln:list nat), linds_occ pfdt l x = ln <-> (exists pfeq : count_occ pfdt l x = length ln, forall (m : nat) (pflt : m < count_occ pfdt l x), lind_occ pfdt l m x pf = nth_lt ln m (lt_eq_trans m (count_occ pfdt l x) (length ln) pflt pfeq)). intros T h1 l x h2 ln. unfold linds_occ. destruct in_dec as [h3 | h3]. unfold linds_occ'. rewrite map_seg_eq_iff; simpl. pose proof (proof_irrelevance _ h2 h3); subst. tauto. contradiction. Qed. Lemma nin_linds_occ_eq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:~In x l) (ln:list nat), linds_occ pfdt l x = ln <-> ln = nil. intros T h1 l x h2 ln. apply (linds_occ_nin h1) in h2. rewrite h2. intuition. Qed. Lemma in_linds_occ_nth_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (r:nat) (pf:r < length l), In r (linds_occ pfdt l (nth_lt l r pf)). intros T h1 l. induction l as [|a l ih]; simpl. intros; omega. intros r h2. destruct r as [|r]; simpl. rewrite linds_occ_cons_eq. left; auto. apply lt_S_n in h2. destruct (h1 a (nth r l a)) as [h3 | h3]. rewrite <- h3. rewrite linds_occ_cons_eq. right. rewrite app_nil_l. apply in_map. specialize (ih r h2). rewrite nth_lt_compat in ih. erewrite nth_indep in ih. rewrite <- h3 in ih at 1; auto. assumption. specialize (ih r h2). rewrite linds_occ_cons_neq; auto. apply in_map. erewrite nth_lt_compat in ih. erewrite nth_indep in ih. apply ih. assumption. Qed. Lemma in_linds_occ_nth_lt_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), In x l -> forall (i j:nat) (pfii:In i (linds_occ pfdt l x)) (pfij:In j (linds_occ pfdt l x)), let pflti := in_linds_occ_lt pfdt _ _ _ pfii in let pfltj := in_linds_occ_lt pfdt _ _ _ pfij in nth_lt l i pflti = nth_lt l j pfltj. intros T h1 l x h2 i j h3 h4 h5 h6. pose proof h3 as h3'. pose proof h4 as h4'. rewrite in_linds_occ_iff in h3'. rewrite in_linds_occ_iff in h4'. destruct h3' as [m [h9 [h7 h8]]]. destruct h4' as [m' [h9' [h7' h8']]]. gen0. rewrite <- h8. intro h10. symmetry. gen0. rewrite <- h8'. intro h11. erewrite nth_lt_functional3. rewrite lind_occ_compat. erewrite nth_lt_functional3. rewrite lind_occ_compat; auto. Qed. Lemma in_linds_occ_nth_lt_eq' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), In x l -> forall (i j:nat) (pfii:In i (linds_occ pfdt l x)) (pfij:In j (linds_occ pfdt l x)) (pflti:i < length l) (pfltj:j < length l), nth_lt l i pflti = nth_lt l j pfltj. intros T h1 l x h2 i j h3 h4 h5 h6. erewrite nth_lt_functional3. rewrite in_linds_occ_nth_lt_eq at 1. apply nth_lt_functional3. apply h2. Grab Existential Variables. apply h4. assumption. Qed. Lemma in_linds_occ_nth_lt' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (r:nat) (pf:r < length l), In r (linds_occ' pfdt l (nth_lt l r pf) (in_nth_lt _ _ _)). intros T h1 l r h2. pose proof (in_linds_occ_nth_lt h1 l r h2) as h3. unfold linds_occ in h3. destruct in_dec as [h4 | h4]. pose proof (proof_irrelevance _ h4 (in_nth_lt _ _ _)). subst. assumption. red in h3. contradiction. Qed. Lemma nth_lt_linds_occ' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (m:nat) (pfm:m < length (linds_occ' pfdt l x pf)), nth_lt (linds_occ' pfdt l x pf) m pfm = lind_occ pfdt l m x pf. intros T h1 l x h2 m h3. unfold linds_occ'. erewrite nth_lt_functional3. rewrite nth_lt_map_seg; auto. Grab Existential Variables. rewrite length_linds_occ' in h3; auto. Qed. Lemma nth_lt_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (m:nat) (pfm:m < length (linds_occ pfdt l x)), nth_lt (linds_occ pfdt l x) m pfm = lind_occ pfdt l m x pf. intros. gen0. erewrite linds_occ_compat'. intro h1. rewrite nth_lt_linds_occ'. reflexivity. Qed. Lemma hd_prime_linds_occ' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l) (pfn:linds_occ' pfdt l x pfin <> nil), hd' (linds_occ' pfdt l x pfin) pfn = lind pfdt l x pfin. intros T h1 l x h2 h3. rewrite <- nth_lt0. rewrite nth_lt_linds_occ'. unfold lind; auto. Qed. Lemma hd_linds_occ' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l) (pfn:linds_occ' pfdt l x pfin <> nil) (d:nat), hd d (linds_occ' pfdt l x pfin) = lind pfdt l x pfin. intros; rewrite (hd_hd_compat' _ pfn). rewrite hd_prime_linds_occ'; auto. Qed. Lemma hd_prime_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l) (pfn:linds_occ pfdt l x <> nil), hd' (linds_occ pfdt l x) pfn = lind pfdt l x pfin. intros T h1 l x h2. rewrite (linds_occ_compat' h1 l x h2). intro h3. apply hd_prime_linds_occ'. Qed. Lemma hd_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l) (pfn:linds_occ pfdt l x <> nil) (d:nat), hd d (linds_occ pfdt l x) = lind pfdt l x pfin. intros; rewrite (hd_hd_compat' _ pfn). erewrite hd_prime_linds_occ; auto. Qed. Lemma no_dup_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), NoDup (linds_occ pfdt l x). intros T h1 l. induction l as [|a l ih]; simpl. intro x. rewrite linds_occ_nil. constructor. intro x. destruct (h1 a x) as [|h2]; subst. rewrite linds_occ_cons_eq. simpl. constructor. intro h2. rewrite in_lS_iff in h2. destruct h2 as [? [? ?]]; discriminate. apply no_dup_lS; auto. rewrite linds_occ_cons_neq. apply no_dup_lS; auto. contradict h2; auto. Qed. Lemma no_dup_linds_occ': forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), NoDup (linds_occ' pfdt l x pf). intros. rewrite <- linds_occ_compat'. apply no_dup_linds_occ. Qed. Lemma impl_count_occ_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x y:T), list_to_set (linds_occ pfdt l x) = list_to_set (linds_occ pfdt l y) -> count_occ pfdt l x = count_occ pfdt l y. intros T h1 l. induction l as [|a l ih]; simpl; auto. intros x y h2. destruct (h1 a x) as [|h3]; subst. destruct (h1 x y) as [|h4]; subst. reflexivity. rewrite linds_occ_cons_eq, linds_occ_cons_neq in h2. simpl in h2. pose proof (Add_intro2 _ (list_to_set (lS (linds_occ h1 l x))) 0) as h5. rewrite h2 in h5. rewrite <- list_to_set_in_iff in h5. rewrite in_lS_iff in h5. destruct h5 as [? [? ?]]; discriminate. apply neq_sym; auto. destruct (h1 a y) as [|h4]; subst. rewrite linds_occ_cons_eq, linds_occ_cons_neq in h2. simpl in h2. pose proof (Add_intro2 _ (list_to_set (lS (linds_occ h1 l y))) 0) as h5. rewrite <- h2 in h5. rewrite <- list_to_set_in_iff in h5. rewrite in_lS_iff in h5. destruct h5 as [? [? ?]]; discriminate. apply neq_sym; auto. do 2 rewrite linds_occ_cons_neq in h2; auto. apply ih. unfold lS in h2. do 2 rewrite map_im_compat in h2. apply im_inj_inj in h2; auto. red; intros; apply S_inj; auto. Qed. Lemma all_lind_occ_eq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x x':T) (pfx:In x l) (pfx':In x' l), (count_occ pfdt l x = count_occ pfdt l x' /\ (forall (n:nat), n < count_occ pfdt l x -> lind_occ pfdt l n x pfx = lind_occ pfdt l n x' pfx')) <-> x = x'. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. split; intro h2. destruct h2 as [h2 h2']. destruct (h1 a x) as [h3 | h3], (h1 a x') as [h4 | h4]; simpl; subst; auto. destruct (h1 x x) as [h5 | h5], (h1 x' x) as [h6 | h6]; simpl; auto. specialize (h2' 0 (lt_0_Sn _)). simpl in h2'. discriminate. pose proof (h5 eq_refl); contradiction. destruct (h1 x x') as [h4 | h4]; subst; auto. destruct (h1 x' x') as [h5 | h5]. specialize (h2' 0). pose proof (neq_in_cons _ _ _ h4 pfx) as h6. hfirst h2'. pose proof (count_occ_In h1 l x) as h7. rewrite h7 in h6; auto with arith. specialize (h2' hf). simpl in h2'. discriminate. pose proof (h5 eq_refl); contradiction. apply neq_sym in h3. apply neq_sym in h4. pose proof (neq_in_cons _ _ _ h3 pfx) as h5. pose proof (neq_in_cons _ _ _ h4 pfx') as h6. destruct (h1 x a) as [h7 | h7], (h1 x' a) as [h8 | h8]; subst. contradict h3; auto. contradict h3; auto. contradict h4; auto. rewrite <- ih. split; auto. intros n h9. specialize (h2' _ h9). apply S_inj in h2'; auto. apply h2'. subst. destruct (h1 a x') as [h2 | h2], (h1 x' a) as [h3 | h3]; simpl; auto. subst. contradict h3; auto. split; auto. intros; auto. f_equal. f_equal. f_equal. apply proof_irrelevance. Qed. Lemma linds_occ_firstn : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (n:nat) (pfin:In x l) (pfin':In x (firstn n l)), linds_occ pfdt (firstn n l) x = firstn (count_occ pfdt (firstn n l) x) (linds_occ pfdt l x). intros T h1 l x n h2 h3. unfold linds_occ. destruct in_dec as [h4 | h4]. destruct in_dec as [h5 | h5]. unfold linds_occ'. rewrite map_seg_eq_iff. ex_goal. rewrite length_firstn; auto. rewrite length_map_seg. apply le_count_occ_sub_list. apply sub_list_firstn. exists hex. intros m h6 h7. rewrite nth_lt_firstn_eq. erewrite nth_lt_functional3. rewrite nth_lt_map_seg at 1. rewrite lind_occ_firstn. f_equal. apply proof_irrelevance. assumption. contradiction. contradiction. Grab Existential Variables. eapply lt_le_trans. apply h6. apply le_count_occ_sub_list. apply sub_list_firstn. Qed. Lemma linds_occ_firstn_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (m n:nat), m < n -> forall (pfinm:In x (firstn m l)) (pfinn:In x (firstn n l)), linds_occ pfdt (firstn m l) x = firstn (count_occ pfdt (firstn m l) x) (linds_occ pfdt (firstn n l) x). intros T h1 l x m n h2 h3 h4. rewrite (firstn_lt_nest _ _ _ h2) at 1; auto. rewrite linds_occ_firstn at 1; auto. f_equal. f_equal. rewrite <- firstn_lt_nest; auto. rewrite <- firstn_lt_nest; auto. Qed. Lemma in_linds_occ_firstn_le : forall (T:Type) (pfdt:type_eq_dec T) (l:list T) (x:T) (r m n:nat) (pfsm:S m <= length l), m <= n -> r < m -> In r (linds_occ pfdt (firstn m l) x) -> In r (linds_occ pfdt (firstn n l) x). intros T h1 l x r m n h2 hle h3 h4. rewrite in_linds_occ_iff in h4. destruct h4 as [o [h4 [h5 h6]]]; subst. rewrite in_linds_occ_iff. exists o. ex_goal. eapply in_firstn_le. apply hle. assumption. exists hex. split. eapply lt_le_trans. apply h5. apply le_count_occ_sub_list. apply sub_list_firstn_le; auto with arith. rewrite lind_occ_firstn; auto with arith. rewrite lind_occ_firstn; auto with arith. apply lind_occ_functional; auto. eapply lt_le_trans. apply h5. apply le_count_occ_sub_list. apply sub_list_firstn_le; auto with arith. Qed. Lemma in_nin_linds_occ_firstn_S : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (r m:nat) (pfsm:S m <= length l), r < m -> In x (firstn m l) -> ~In (nth_lt l m pfsm) (firstn m l) -> In r (linds_occ pfdt (firstn (S m) l) x) -> In r (linds_occ pfdt (firstn m l) x). intros T h1 l x r m h2 h3 h4 h5 h6. rewrite in_linds_occ_iff in h6. destruct h6 as [o [h6 [h7 h8]]]. rewrite in_linds_occ_iff. subst. rewrite lind_occ_firstn in h3; auto with arith. pose proof h6 as h6'. rewrite (firstn_S_eq _ _ h2) in h6', h7. rewrite count_occ_app in h7. simpl in h7. destruct (h1 (nth_lt l m h2) x) as [h8 | h8]; subst. contradiction. rewrite <- plus_n_O in h7. apply in_app_or in h6'. destruct h6' as [h9 | h9]. exists o, h9. split; auto with arith. rewrite lind_occ_firstn. rewrite lind_occ_firstn. apply lind_occ_functional; auto. eapply lt_le_trans. apply h7. apply le_count_occ_sub_list. apply sub_list_firstn_le; auto with arith. assumption. destruct h9; subst. contradiction. simpl in H. contradiction. Qed. Lemma in_nin_linds_occ_firstn_S' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (m r:nat) (pfsm:S m <= length l), r < S m -> x <> nth_lt l m pfsm -> In r (linds_occ pfdt (firstn (S m) l) x) -> In r (linds_occ pfdt (firstn m l) x). intros T h1 l x m r h2 h3 h4 h5. rewrite in_linds_occ_iff in h5. rewrite in_linds_occ_iff. destruct h5 as [o [h6 [h7 h8]]]. pose proof h6 as h6'. apply (in_firstn_S_dec h1 _ _ h2) in h6'. destruct h6' as [h9 | h9]. exists o. exists h9. rewrite (firstn_S_eq _ _ h2) in h7. rewrite count_occ_app in h7. simpl in h7. lr_if h7; fold hlr in h7; destruct hlr as [hl | hr]; subst. contradict h4; auto. rewrite <- plus_n_O in h7; auto. rewrite lind_occ_firstn; auto. split; auto. rewrite lind_occ_firstn; auto. apply lind_occ_functional; auto. eapply lt_le_trans. apply h7. apply le_count_occ_sub_list. apply sub_list_firstn_le; auto with arith. subst. contradict h4; auto. Qed. Lemma nin_lind_nth_lt_firstn_S : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pfsm:S m <= length l) (pfi:In (nth_lt l m pfsm) (firstn (S m) l)), ~In (nth_lt l m pfsm) (firstn m l) -> lind pfdt (firstn (S m) l) (nth_lt l m pfsm) pfi = m. intros T h1 l m h2 h3 h4. rewrite lind_eq_iff. split. rewrite nth_lt_firstn_eq. apply nth_lt_functional3. erewrite <- firstn_lt_nest; auto with arith. Grab Existential Variables. rewrite length_firstn; auto with arith. Qed. Lemma map_linds_occ_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), map (linds_occ pfdt l) l = map_in_dep (linds_occ' pfdt l). intros T h1 l. rewrite map_eq_iff. ex_goal. rewrite length_map_in_dep; auto. exists hex. intros m h2. erewrite nth_lt_map_in_dep, linds_occ_compat'. reflexivity. Qed. Lemma map_linds_occ_compat' : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (pfi:incl l' l), map (linds_occ pfdt l) l' = map_in_dep (fun (x:T) (pfix:In x l') => (linds_occ' pfdt _ _ (pfi _ pfix))). intros T h1 l l' h2. rewrite map_eq_iff. ex_goal. rewrite length_map_in_dep; auto. exists hex. intros m h3. erewrite nth_lt_map_in_dep, linds_occ_compat'. reflexivity. Qed. Definition remove_sig {T} (l:list T) (a:T) : Type := {x:{t:T | t <> a} | In (proj1_sig x) l}. Definition remove_sig_list {T l a} (pfdt:type_eq_dec T) (pfn:~In a l) : list {x:T | x <> a} := map_in_dep (fun x (pfx:In x l) => match (pfdt x a) with | left pfe => let pfx := in_subst1 pfe pfx in False_rect _ (pfn pfx) | right pfn => exist (fun c => c <> a) _ pfn end). Definition remove_mask {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) : bv := inds_to_bv l (linds_occ pfdt l x). Lemma length_remove_mask : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), length (remove_mask pfdt l x) = length l. intros T h1 l x. unfold remove_mask, inds_to_bv. apply length_map_seg. Qed. Lemma lind_interp_bv_ge_hd : forall {T:Type} (pfdt:type_eq_dec T) (x:T) (l:list T) (v:list bool) (pfe:length l = count_occ bool_eq_dec v false) (pfxnin: ~In x l) (pfxin:In x (interp_bv x v l pfe)) (pfn:bv1s v <> nil), lind pfdt (interp_bv x v l pfe) x pfxin >= hd' (bv1s v) pfn. intros T h1 x l v. revert l x. induction v as [|b v ih]. intros ? ? ? ? ? h5. unfold bv1s in h5; simpl in h5. contradict h5; auto. intros l x h2 h3 h4 h5. destruct b. generalize h5. rewrite bv1s_cons_true; simpl. intros; auto with arith. generalize h5. rewrite bv1s_cons_false; simpl. intro h6. unfold lS. rewrite hd_map'. generalize h4. pose proof h2 as h2'. rewrite count_occ_cons_eq in h2'; auto. apply length_eq_S in h2'. rewrite (interp_bv_false _ _ _ _ h2'). intro h7. destruct h7 as [h7 | h7]; subst. contradict h3. apply in_hd'. rewrite lind_cons_neq. apply le_n_S. apply ih. contradict h3. apply in_tl in h3; auto. intro; subst. contradict h3. apply in_hd'. Qed. Lemma lind_interp_bv : forall {T:Type} (pfdt:type_eq_dec T) (x:T) (l:list T) (v:list bool) (pfe:length l = count_occ bool_eq_dec v false) (pfxnin: ~In x l) (pfxin:In x (interp_bv x v l pfe)) (pfn:bv1s v <> nil), lind pfdt (interp_bv x v l pfe) x pfxin = hd' (bv1s v) pfn. intros T h1 x l v h2 h3 h4 h5. pose proof (in_bv1s_lt v _ (in_hd' _ h5)) as h6. pose proof (lind_eq_iff h1 _ _ h4 (hd' _ h5)) as h7. hfirst h7. rewrite length_interp_bv; auto. specialize (h7 hf). rewrite h7. split. apply nth_lt_interp_bv; auto. apply in_hd'. intro h8. rewrite (in_firstn_iff' h1) in h8. destruct h8 as [h8 h9]. pose proof (proof_irrelevance _ h8 h4); subst. nhl h7. intro h10. rewrite h10 in h9; apply lt_irrefl in h9; contradiction. rewrite h7 in hl. contradict hl. pose proof (lind_interp_bv_ge_hd h1 x l v h2 h3 h4 h5). omega. Qed. Lemma lind_occ_interp_bv : forall {T:Type} (pfdt:type_eq_dec T) (x:T) (v:bv) (l:list T) (occ:nat) (pfeq:length l = count_occ bool_eq_dec v false) (pfnin: ~In x l) (pfin: In x (interp_bv x v l pfeq)) (pfo:occ < length (bv1s v)), lind_occ pfdt (interp_bv x v l pfeq) occ x pfin = nth_lt (bv1s v) occ pfo. intros T pfdt x v. revert x. induction v as [|b v ih]. intros; contradiction. intros x l occ h1 hn h2 h3. destruct b. gen0. rewrite interp_bv_true. intro h5. pose proof h1 as h1'. rewrite count_occ_cons_neq in h1'. destruct (in_dec bool_eq_dec true v) as [h6|h6]. pose proof (in_interp_bv x _ _ h1' h6) as h7. destruct occ as [|occ]. symmetry. gen0. rewrite bv1s_cons_true. intro ha. simpl. rewrite eq_dec_eq at 1; auto. symmetry. gen0. rewrite bv1s_cons_true. intro h8. rewrite nth_lt_S. pose proof h7 as h7'. term0 h5. redterm0 c. redterm0 c0. pose proof (proof_irrelevance _ c1 h1') as he. unfold c1 in he. rewrite <- he in h7'. rewrite (lind_occ_cons_in_eq _ _ _ _ h7' h5). simpl. rewrite nth_lt_lS. f_equal. simpl in h8. pose proof h8 as h8'. apply lt_S_n in h8'. rewrite length_lS in h8'. specialize (ih x l occ _ hn h7' h8'). rewrite ih at 1. apply nth_lt_functional3. auto with arith. destruct (zerop occ) as [|hz]; subst. gterml. redterm3 x0. redterm0 c. fold c0. pose proof (lind_cons_same pfdt c0 x h5) as h7. rewrite h7 at 1. simpl; auto. pose proof (nin_true_bv1s _ h6) as h9. symmetry. gen0. rewrite bv1s_cons_true, h9. intro h10. simpl in h10. assert (h11:occ = 0). omega. gen0. rewrite h11. intro h12. simpl. assert (h13:pfdt x x = left (eq_refl x)). destruct (pfdt x x); f_equal; try apply proof_irrelevance. contradict n; auto. rewrite h13; auto. intro; discriminate. gen0. pose proof h1 as h1'. rewrite count_occ_cons_eq in h1'; auto. pose proof h1' as h4. apply length_eq_S in h4. rewrite (interp_bv_false _ _ _ h1 h4). intro h5. symmetry. assert (hneq: x <> hd' l h4). intro h6; subst; contradict hn; apply in_hd'. gen0. rewrite bv1s_cons_false. intro h6. unfold lS. pose proof h6 as h6'. rewrite length_lS in h6'. rewrite (nth_lt_map _ _ _ h6'). destruct h5 as [h5 | h5]; subst. contradict hneq; auto. destruct occ as [|occ]; simpl. assert (he:pfdt x (hd' l h4) = right hneq). destruct (pfdt x (hd' l h4)) as [h8 | h8]. contradiction. f_equal; apply proof_irrelevance. rewrite he. f_equal. pose proof (lind_interp_bv pfdt x (tl l) v (length_count_occ_false_tl x l v h1)) as h8. hfirst h8. contradict hn. apply in_tl in hn; auto. specialize (h8 hf h5). hfirst h8. apply lt_length in h6'; auto. specialize (h8 hf0). unfold lind in h8. erewrite lind_occ_functional. rewrite h8 at 1. rewrite nth_lt_0_hd'. f_equal. apply proof_irrelevance. reflexivity. reflexivity. reflexivity. assert (h7:pfdt x (hd' l h4) = right hneq). destruct (pfdt x (hd' l h4)). contradiction. f_equal; apply proof_irrelevance. rewrite h7. f_equal. symmetry. apply ih. contradict hn. apply in_tl; auto. Qed. Lemma interp_bv_compat : forall {T:Type} (pfdt:type_eq_dec T) (x:T) (v:bv) (l:list T) (pfe:length l = count_occ bool_eq_dec v false) (i:nat) (pf:i < length v), ~In x l -> bv1s v = linds_occ pfdt (interp_bv x v l pfe) x. intros T h1 x v l h2 i h3 h4. symmetry. destruct (in_dec bool_eq_dec true v) as [h5|h5]. pose proof (in_interp_bv x v l h2 h5) as h6. rewrite (linds_occ_eq_iff _ _ _ h6). ex_goal. apply count_occ_interp_bv; auto. exists hex. intros occ h7. pose proof h7 as h8. rewrite hex in h8. apply lind_occ_interp_bv; auto. rewrite nin_true_interp_bv, nin_true_bv1s; auto. rewrite nin_linds_occ_eq_iff; auto. Qed. (*Returns an analogous list to [l], which has a [true] at values at x, and a [false] at all others.*) Fixpoint list_to_bv {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) : bv := match l with | nil => nil | a::l' => let v := list_to_bv pfdt l' x in if (pfdt x a) then (true :: v) else (false :: v) end. Lemma count_occ_list_to_bv : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), count_occ bool_eq_dec (list_to_bv pfdt l x) false = length (remove pfdt x l). intros T h1 l. induction l as [|a l ih]; simpl; auto. intro x. destruct (h1 x a) as [|h2]; subst. rewrite count_occ_cons_neq. apply ih. intro; discriminate. rewrite count_occ_cons_eq; auto. simpl. f_equal; auto. Qed. Lemma list_to_bv_cons_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), list_to_bv pfdt (x::l) x = true :: list_to_bv pfdt l x. intros T h1 l x. simpl. destruct (h1 x x) as [h3 | h3]; auto. contradict h3; auto. Qed. Lemma list_to_bv_cons_neq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x a:T), x <> a -> list_to_bv pfdt (a::l) x = false :: list_to_bv pfdt l x. intros T h1 l x a h2. simpl. destruct (h1 x a) as [h3 | h3]. contradiction. reflexivity. Qed. Lemma interp_bv_list_to_bv_remove : forall {T:Type} (pfdt:type_eq_dec T) (x:T) (l:list T), let pfeq := eq_sym (count_occ_list_to_bv pfdt l x) in interp_bv x (list_to_bv pfdt l x) (remove pfdt x l) pfeq = l. intros T h1 x l. revert x. induction l as [|a l ih]. intros x. unfold interp_bv; simpl; auto. intros x h2. unfold h2. clear h2. destruct (h1 x a) as [|h3]; subst. gen0. rewrite list_to_bv_cons_eq. rewrite remove_cons_eq. intro h2. rewrite interp_bv_true. f_equal. specialize (ih a). simpl in ih. erewrite f_equal_dep. rewrite ih at 1; auto. reflexivity. gen0. rewrite list_to_bv_cons_neq, remove_cons_neq; auto. intro h4. erewrite interp_bv_false. simpl. f_equal. specialize (ih x). simpl in ih. erewrite f_equal_dep. rewrite ih at 1; auto. reflexivity. Grab Existential Variables. apply cons_neq_nil. Qed. Lemma remove_eq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (x:T), let v := list_to_bv pfdt l x in remove pfdt x l = l' <-> exists (pf:~In x l') (pfe:length l' = count_occ bool_eq_dec v false), l = interp_bv x v l' pfe. Proof. intros T h1 l l' x v; split; intro h3; subst. ex_goal. apply remove_In. exists hex. unfold v. ex_goal. rewrite count_occ_list_to_bv; auto. exists hex0. symmetry. erewrite f_equal_dep. apply interp_bv_list_to_bv_remove; auto. reflexivity. unfold v in h3. clear v. destruct h3 as [h3 h4]. destruct h4 as [h4 h5]. rewrite h5. apply remove_interp_bv; auto. Qed. (*"List Equivalence class" or "List Equal class".*) Definition lec {T:Type} {l:list T} {i:nat} (pfdt:type_eq_dec T) (pf:i < length l) : list nat := linds_occ pfdt l (nth_lt l i pf). (*This increments each element of the list equivalence class of a given representative index.*) Definition lec_S {T:Type} {l:list T} {i:nat} (pfdt:type_eq_dec T)(pf:i < length l) : list nat := lS (lec pfdt pf). Definition lec_cons {T:Type} {l:list T} {i:nat} (pfdt:type_eq_dec T) (pf:i < length l) (a:T) : list nat := let pf' := lt_n_S _ _ pf in lec pfdt pf' (l := a::l). Lemma length_lec : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (pf:i < length l), length (lec pfdt pf) = count_occ pfdt l (nth_lt _ _ pf). intros T h1 i h2 h3. unfold lec. rewrite length_linds_occ; auto. Qed. Lemma lecs_eq : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (m:nat) (pfm:m < length l) (pf:l = l'), let pfe := length_subst l l' pf in let pfm' := lt_eq_trans _ _ _ pfm pfe in lec pfdt pfm = lec pfdt pfm'. simpl; intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma linds_occ_nth_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (pfi:i < length l), linds_occ pfdt l (nth_lt _ _ pfi) = lec pfdt pfi. unfold lec; auto. Qed. Lemma no_dup_lec : forall {T:Type} {l:list T} {i:nat} (pfdt:type_eq_dec T) (pfi:i < length l), NoDup (lec pfdt pfi). unfold lec; intros; apply no_dup_linds_occ. Qed. (*See past [sort] below for a sorted version of [linds] called [sorted_linds].*) (*These functions return a family of lists of the indices of each element in [l].*) Definition linds {T:Type} (pfdt:type_eq_dec T) (l:list T) : faml nat := map (linds_occ pfdt l) (list_singularize pfdt l). Definition linds' {T:Type} (pfdt:type_eq_dec T) (l:list T) : faml nat := let f := fun_in_dep_incl _ _ (linds_occ' pfdt l) (incl_list_singularize pfdt _) in map_in_dep f. Lemma linds_functional : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), l = l' -> linds pfdt l = linds pfdt l'. intros; subst; auto. Qed. Lemma linds_functional' : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), l = l' -> linds' pfdt l = linds' pfdt l'. intros; subst; auto. Qed. Lemma linds_compat' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), linds pfdt l = linds' pfdt l. intros T h1 l. unfold linds, linds'. rewrite (map_eq_iff (list_singularize h1 l) (linds_occ h1 l)). ex_goal. rewrite length_map_in_dep; auto. exists hex. intros n h2. erewrite nth_lt_map_in_dep. unfold fun_in_dep_incl. unfold linds_occ. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. f_equal. apply proof_irrelevance. contradict hr. rewrite list_singularize_in_iff. apply in_nth_lt. Qed. Lemma linds_nil : forall {T:Type} (pfdt:type_eq_dec T), linds pfdt nil = nil. intros; unfold linds; simpl; auto. Qed. Lemma linds_nil' : forall {T:Type} (pfdt:type_eq_dec T), linds' pfdt nil = nil. intros; unfold linds'; simpl; auto. Qed. Lemma linds_eq_nil_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), linds pfdt l = nil <-> l = nil. intros T h1 l. destruct l as [|a l]. rewrite linds_nil; tauto. unfold linds; simpl. destruct (in_dec h1 a l). split; intro h2. apply map_eq_nil in h2. apply list_singularize_nil_iff in h2; subst; auto. discriminate. rewrite list_singularize_nin; auto. simpl. rewrite linds_occ_cons_eq. split; intro; discriminate. Qed. Lemma linds_eq_nil_iff' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), linds' pfdt l = nil <-> l = nil. intros; rewrite <- linds_compat'; apply linds_eq_nil_iff; auto. Qed. Lemma nin_nil_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), ~In nil (linds pfdt l). intros T h1 l h2. unfold linds in h2. rewrite in_map_iff in h2. destruct h2 as [x h2]. destruct h2 as [h2 h3]. unfold linds_occ in h2. destruct in_dec as [h4 | h4]. unfold linds_occ' in h2. apply map_seg_eq_nil in h2. rewrite (count_occ_In h1) in h4. rewrite <- h2 in h4. apply lt_irrefl in h4; auto. contradict h4. rewrite <- list_singularize_in_iff in h3; auto. Qed. Lemma linds_sing : forall {T:Type} (pfdt:type_eq_dec T) (a:T), linds pfdt (a::nil) = (0::nil)::nil. unfold linds; intros; simpl. rewrite linds_occ_cons_eq; simpl; auto. Qed. Lemma inj_dep_linds_occ' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), inj_dep (linds_occ' pfdt l). intros T h1 l. red; intros x y hx hy h2. unfold linds_occ' in h2. rewrite map_seg_eq_iff in h2. destruct h2 as [h2 h5]. pose proof h2 as h2'. rewrite length_map_seg in h2'. erewrite <- all_lind_occ_eq_iff. split. apply h2'. intros n h6. specialize (h5 _ h6). simpl in h5. erewrite nth_lt_functional3 in h5. rewrite nth_lt_map_seg in h5. apply h5. Grab Existential Variables. rewrite <- h2'; auto. Qed. Lemma inj_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x x':T), In x l -> In x' l -> linds_occ pfdt l x = linds_occ pfdt l x' -> x = x'. intros T h1 l x x' h2 h3 h5. pose proof (inj_dep_linds_occ' h1 l) as h4. red in h4. specialize (h4 _ _ h2 h3). unfold linds_occ in h5. lr_match h5; fold hlr in h5; destruct hlr as [hl | hr]; lr_match h5; fold hlr in h5; destruct hlr as [hl' | hr']; try contradiction. pose proof (proof_irrelevance _ h2 hl). pose proof (proof_irrelevance _ h3 hl'); subst; auto. Qed. Lemma in_linds_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat), In ln (linds pfdt l) <-> exists! (x:T), In x l /\ ln = linds_occ pfdt l x. intros T h1 l ln. unfold linds, linds_occ, linds_occ'. rewrite in_map_iff. split; intro h2. destruct h2 as [x [h2 hin]]. rewrite <- list_singularize_in_iff in hin. exists x. red; split. split; auto. intros x' h3. destruct h3 as [h3 h4]; subst. destruct (in_dec h1 x l) as [h5 | h5], (in_dec h1 x' l) as [h6 | h6]. rewrite map_seg_eq_iff in h4. destruct h4 as [h4 h7]. pose proof h4 as h4'. rewrite length_map_seg in h4'. pose proof (all_lind_occ_eq_iff h1 l x x' h5 h6) as h8. rewrite <- h8. split; auto. intros n h9. specialize (h7 n h9). simpl in h7. erewrite nth_lt_functional3 in h7. rewrite nth_lt_map_seg in h7; auto. contradiction. contradiction. contradiction. destruct h2 as [x [h2 hin]]. destruct h2 as [h2 h3]; simpl; subst. exists x. split; auto. rewrite <- list_singularize_in_iff; auto. Grab Existential Variables. rewrite <- h4'. assumption. Qed. Lemma in_linds_iff' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat), In ln (linds' pfdt l) <-> exists! (x:T), exists (pf:In x l), ln = linds_occ' pfdt l x pf. intros T h1 l ln. rewrite <- linds_compat'. rewrite in_linds_iff. split; intro h2. destruct h2 as [x [[h3 h4] h5]]; subst. exists x. red. split. exists h3. rewrite <- linds_occ_compat'; auto. intros x' h4. destruct h4 as [h4 h6]. apply h5. split; auto. rewrite <- linds_occ_compat' in h6; auto. destruct h2 as [x [[h3 h4] h5]]; subst. exists x. red; split. split; auto. rewrite <- linds_occ_compat'; auto. intros x' h4. destruct h4 as [h4 h6]. apply h5. exists h4. do 2 rewrite <- linds_occ_compat'. rewrite <- linds_occ_compat' in h6; auto. Qed. Lemma in_linds_lec_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat), In ln (linds pfdt l) <-> exists (i:nat) (pf:i < length l), ln = lec pfdt pf. intros T h1 l ln. rewrite in_linds_iff. split; intro h2. destruct h2 as [x h2]. red in h2. destruct h2 as [h2 h3]. destruct h2 as [h2]; subst. unfold lec. exists (lind h1 _ _ h2). exists (lt_lind _ _ _ _). rewrite lind_compat; auto. destruct h2 as [i [h2 h3]]; subst. unfold lec. exists (nth_lt l i h2). red; split. split. apply in_nth_lt. reflexivity. intros x h3. destruct h3 as [h3 h4]. apply inj_linds_occ in h4; auto. apply in_nth_lt. Qed. Lemma in_linds_subst : forall {T:Type} {l l':list T} {ln:list nat} (pfdt:type_eq_dec T), In ln (linds pfdt l) -> l = l' -> In ln (linds pfdt l'). intros; subst; auto. Qed. Lemma in_linds_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (i:nat), In i ln -> In ln (linds pfdt l) -> i < length l. intros T h1 l ln i h2 h3. rewrite in_linds_iff in h3. destruct h3 as [x [[h3 h4] h5]]; subst. rewrite in_linds_occ_iff in h2. destruct h2 as [m [h6 [h7 h8]]]; subst. apply lt_lind_occ. Qed. Lemma in_linds_not_nil : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat), In ln (linds pfdt l) -> ln <> nil. intros T h1 l ln h2 h3; subst. rewrite in_linds_iff in h2. destruct h2 as [x [[h3 h4] h5]]. symmetry in h4. pose proof (linds_occ_eq_iff h1 _ _ h3 nil) as h6. rewrite h6 in h4. destruct h4 as [h4]. simpl in h4. pose proof h3 as h3'. pose proof (count_occ_In h1 l x) as h7. rewrite h7 in h3'. rewrite h4 in h3'. apply lt_irrefl in h3'. contradiction. Qed. Lemma in_linds_occ_linds' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), In (linds_occ' pfdt l x pf) (linds' pfdt l). intros T h1 l x h2. unfold linds'. rewrite in_map_in_dep_iff. exists x. ex_goal. rewrite <- list_singularize_in_iff; auto. exists hex. unfold fun_in_dep_incl. f_equal. apply proof_irrelevance. Qed. Lemma in_linds_occ_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), In x l -> In (linds_occ pfdt l x) (linds pfdt l). intros; erewrite linds_compat', linds_occ_compat'. apply in_linds_occ_linds'. Grab Existential Variables. assumption. Qed. Lemma linds_map_inj : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (f:T->U), Injective f -> linds pfdu (map f l) = linds pfdt l. intros T U h1 h2 l f h3. symmetry. unfold linds. rewrite (map_eq_iff (list_singularize h1 l) (linds_occ h1 l)). ex_goal. rewrite map_length. erewrite list_singularize_map_inj. rewrite map_length. reflexivity. assumption. exists hex. intros n h4. erewrite nth_lt_map. rewrite linds_occ_eq_iff. ex_goal. rewrite length_linds_occ. rewrite nth_lt_list_singularize_map. erewrite count_occ_map_inj. reflexivity. rewrite list_singularize_in_iff. apply in_nth_lt. assumption. exists hex0. intros m h5. erewrite nth_lt_linds_occ. symmetry. gen0. erewrite nth_lt_list_singularize_map. intro h6. erewrite lind_occ_map_inj. reflexivity. assumption. Grab Existential Variables. rewrite nth_lt_list_singularize_map. apply in_map. rewrite list_singularize_in_iff. apply in_nth_lt. assumption. rewrite list_singularize_in_iff. apply in_nth_lt. Qed. (*This prints the value given by the indices [ln] in [l]*) Definition in_linds_val {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (pfiln:In ln (linds pfdt l)) : T := let pfn := in_linds_not_nil _ _ _ pfiln in let h := hd' _ pfn in let pfh := in_hd' _ pfn in let pflt := in_linds_lt pfdt _ _ _ pfh pfiln in nth_lt l h pflt. Lemma in_linds_val_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (pfiln:In ln (linds pfdt l)), In (in_linds_val pfdt l ln pfiln) l. unfold in_linds_val; intros; apply in_nth_lt. Qed. Lemma nth_lt_in_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (i:nat) (pfln:In i (linds_occ pfdt l x)), nth_lt l i (in_linds_lt _ _ _ i pfln (in_linds_occ_linds pfdt l x pf)) = x. intros T h1 l x h2 i h3. pose proof h3 as h3'. rewrite in_linds_occ_iff in h3'. destruct h3' as [m [h4 [h5 h6]]]; subst. erewrite nth_lt_functional3. rewrite lind_occ_compat; auto. Qed. Lemma in_linds_length : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (pfi:In ln (linds pfdt l)), length ln = count_occ pfdt l (in_linds_val _ _ _ pfi). intros T h1 l ln h2. pose proof h2 as h2'. rewrite in_linds_iff in h2'. destruct h2' as [x [[h3 h4] h5]]; subst. rewrite length_linds_occ. specialize (h5 (in_linds_val h1 l (linds_occ h1 l x) h2)). hfirst h5. split; auto. apply in_linds_val_in. unfold in_linds_val. pose proof (nth_lt_in_linds_occ h1 l x h3 _ (in_hd' (linds_occ h1 l x) (in_linds_not_nil h1 l (linds_occ h1 l x) h2))) as h6. rewrite <- h6 at 1. repeat f_equal. apply proof_irrelevance. specialize (h5 hf). rewrite h5 at 1; auto. Qed. Lemma count_occ_in_linds_val : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (pfi:In ln (linds pfdt l)), count_occ pfdt l (in_linds_val _ _ _ pfi) = length ln. intros; symmetry; apply in_linds_length. Qed. Lemma nth_lt_eq_in_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln : list nat) (m r:nat) (pfm:m < length l) (pfr:r < length l), In m ln -> nth_lt l m pfm = nth_lt l r pfr -> In ln (linds pfdt l) -> In r ln. intros T h1 l ln m r h2 h3 h4 h5 h6. rewrite in_linds_iff in h6. destruct h6 as [x h6]. red in h6. destruct h6 as [[h6 h7] h8]; subst. rewrite in_linds_occ_iff in h4. destruct h4 as [m' h7]. destruct h7 as [h7 [h9 h10]]; subst. erewrite nth_lt_functional3 in h5. rewrite nth_lt_lind_occ in h5; subst. apply in_linds_occ_nth_lt. Qed. Lemma length_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), length (linds pfdt l) = length (list_singularize pfdt l). intros; unfold linds; rewrite map_length; auto. Qed. Lemma length_linds' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), length (linds' pfdt l) = length (list_singularize pfdt l). intros; unfold linds'; rewrite length_map_in_dep; auto. Qed. Lemma linds_eq_length_singular : forall {T U:Type} {pfdt:type_eq_dec T} {pfdu:type_eq_dec U} (l:list T) (l':list U), linds pfdt l = linds pfdu l' -> length (list_singularize pfdt l) = length (list_singularize pfdu l'). intros; do 2 rewrite <- length_linds. rewrite H; auto. Qed. Lemma linds_eq_length_singular' : forall {T U:Type} {pfdt:type_eq_dec T} {pfdu:type_eq_dec U} (l:list T) (l':list U), linds' pfdt l = linds' pfdu l' -> length (list_singularize pfdt l) = length (list_singularize pfdu l'). intros; do 2 rewrite <- length_linds'. rewrite H; auto. Qed. Lemma no_dup_linds' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), NoDup (linds' pfdt l). intros T h1 l. unfold linds'. apply no_dup_map_in_dep_inj. apply no_dup_list_singularize. apply inj_dep_incl. apply inj_dep_linds_occ'. Qed. Lemma no_dup_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), NoDup (linds pfdt l). intros; rewrite linds_compat'. apply no_dup_linds'. Qed. Lemma no_dup_mem_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), no_dup_mem (linds pfdt l). intros T h1 l; red; intros ln h2. rewrite in_linds_iff in h2. destruct h2 as [x [[h2 h3] h4]]; subst. apply no_dup_linds_occ. Qed. Lemma all_sing_in_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (m:nat), all_sing ln m -> In ln (linds pfdt l) -> ln = m::nil. intros T h1 l ln m h2 h3. rewrite in_linds_iff in h3. destruct h3 as [x [[h3 h4] h5]]; subst. pose proof h3 as h3'. pose proof (no_dup_linds_occ h1 l x) as h6. rewrite (list_singularize_no_dup nat_eq_dec) in h6. rewrite <- h6. erewrite all_sing_list_singularize. reflexivity. intro h7. rewrite linds_occ_eq_iff in h7. destruct h7 as [h7]. rewrite (count_occ_In h1) in h3. simpl in h7; omega. assumption. Grab Existential Variables. assumption. Qed. Lemma nth_lts_eq_length_linds : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l'), (forall (m n : nat) (pfm : m < length l) (pfn : n < length l), let pfm' := lt_eq_trans m (length l) (length l') pfm pfe in let pfn' := lt_eq_trans n (length l) (length l') pfn pfe in nth_lt l m pfm = nth_lt l n pfn <-> nth_lt l' m pfm' = nth_lt l' n pfn') -> length (linds pfdt l) = length (linds pfdu l'). intros; do 2 rewrite length_linds; eapply nth_lts_eq_length_list_singularize. apply H. Qed. Lemma in_lec : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (pf: i < length l), In (lec pfdt pf) (linds pfdt l). intros. apply in_linds_occ_linds. apply in_nth_lt. Qed. Lemma in_lec_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m r:nat) (pflt:m < length l), In r (lec pfdt pflt) -> r < length l. intros T h1 l m r h2 h3. unfold lec in h3. rewrite in_linds_occ_iff in h3. destruct h3 as [? [? [h3 ?]]]; subst. apply lt_lind_occ. Qed. Lemma in_lec_S : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (pf: i < length l), In (lec_S pfdt pf) (famlS (linds pfdt l)). intros T h1 l i h2. unfold lec_S, famlS. apply in_map. apply in_lec. Qed. Lemma nin0_lec_S : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pflt:n < length l), ~In 0 (lec_S pfdt pflt). intros T h1 l n h2 h3. unfold lec_S in h3. rewrite in_lS_iff in h3. destruct h3 as [? [? ?]]; discriminate. Qed. Lemma lec_cons_nth_lt_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pfm:m < length l), let a := nth_lt l m pfm in lec_cons pfdt pfm a = 0 :: lec_S pfdt pfm. intros T h1 l m h2; simpl. unfold lec_cons, lec_S, lec. rewrite linds_occ_eq_iff. ex_goal. simpl. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. rewrite length_lS, length_linds_occ. rewrite nth_lt_compat. do 2 f_equal. apply nth_indep; auto. contradict hr. rewrite nth_lt_compat. apply nth_indep; auto. exists hex. intros occ h3. simpl; simpl in h3. rewrite switch_eq_dec_eq in h3. lr_match_goal; fold hlr in h3; fold hlr; destruct hlr as [hl | hr]; unfold switch_eq_dec in h3. destruct occ as [|occ]; auto. destruct in_dec as [h4 | h4]. erewrite nth_indep. rewrite <- nth_lt_compat. unfold lS. erewrite nth_lt_map. f_equal. symmetry. erewrite nth_lt_functional3. erewrite nth_lt_linds_occ at 1. apply lind_occ_functional; auto. rewrite length_lS, length_linds_occ. apply lt_S_n in h3. erewrite nth_lt_compat. erewrite nth_indep. apply h3. assumption. contradict h4. erewrite nth_indep. rewrite <- nth_lt_compat. apply in_nth_lt. assumption. contradict hr. rewrite nth_lt_compat. f_equal. rewrite nth_lt_compat. apply nth_indep. assumption. Grab Existential Variables. assumption. apply in_nth_lt. rewrite length_linds_occ. apply lt_S_n in h3. erewrite nth_lt_compat; auto. apply lt_S_n in h3. rewrite length_linds_occ, nth_lt_compat; auto. apply lt_S_n in h3. rewrite length_lS, length_linds_occ, nth_lt_compat; auto. left. simpl. apply nth_lt_compat. Qed. Lemma in_lec_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m r:nat) (pfm:m < length l) (pfr:r < length l), In r (lec pfdt pfm) <-> nth_lt l m pfm = nth_lt l r pfr. intros T h1 l m r h2 h3; split; intro h4. unfold lec in h4. rewrite in_linds_occ_iff in h4. destruct h4 as [c [h6 [h4 h5]]]. subst. symmetry. gen0. erewrite lind_occ_functional. intro h5. erewrite nth_lt_functional3. rewrite lind_occ_compat at 1. reflexivity. reflexivity. reflexivity. reflexivity. unfold lec. rewrite in_linds_occ_iff. pose proof (ex_lind_occ_nth_lt h1 _ _ h3) as h5. destruct h5 as [occ [h5 h6]]. rewrite <- h6. exists occ. exists (in_nth_lt _ _ _). split. rewrite h4; auto. gen0. rewrite h4. intro; apply lind_occ_functional; auto. Grab Existential Variables. apply in_nth_lt. Qed. Lemma in_lec_iff' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m r:nat) (pfm:m < length l), In r (lec pfdt pfm) <-> exists pfr:r < length l, nth_lt l m pfm = nth_lt l r pfr. intros T h1 l m r h2; split; intro h3. exists (in_lec_lt _ _ _ _ _ h3). pose proof h3 as h4. rewrite in_lec_iff in h4. apply h4. destruct h3 as [h3 h4]. rewrite in_lec_iff. apply h4. Qed. Lemma in_lec_S_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m r:nat) (pfm:m < length l), In r (lec_S pfdt pfm) <-> 0 < r /\ exists pfpr:pred r < length l, nth_lt l m pfm = nth_lt l (pred r) pfpr. intros T h1 l m r h2. unfold lec_S. rewrite in_lS_iff. split; intro h3. destruct h3 as [m' [h3 h4]]; subst. simpl. split; auto with arith. exists (in_lec_lt _ _ _ _ _ h3). pose proof h3 as h3'. rewrite in_lec_iff' in h3'. destruct h3' as [h4 h5]. rewrite h5. apply nth_lt_functional3. destruct h3 as [hr [h3 h4]]. exists (pred r). split. rewrite in_lec_iff. apply h4. apply (S_pred _ _ hr). Qed. Lemma rep_in_lec : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pf:m < length l), In m (lec pfdt pf). intros T h1 l m h2. rewrite in_lec_iff. reflexivity. Qed. Lemma lec_not_nil : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pf:m < length l), lec pfdt pf <> nil. intros T h1 l m h2 h3. unfold lec in h3. rewrite linds_occ_eq_iff in h3. destruct h3 as [h3 h4]. simpl in h3. pose proof (in_nth_lt l m h2) as h5. rewrite (count_occ_In h1) in h5. rewrite h3 in h5. apply lt_irrefl in h5; auto. Grab Existential Variables. apply in_nth_lt. Qed. Lemma lecs_eq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pfm:m < length l) (n:nat) (pfn:n < length l), lec pfdt pfm = lec pfdt pfn <-> nth_lt _ _ pfm = nth_lt _ _ pfn. intros T h1 l m h2 n h3; split; intro h4. unfold lec in h4. apply inj_linds_occ in h4; try apply in_nth_lt; auto. unfold lec. rewrite h4; auto. Qed. Lemma lec_Ss_eq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pfm:m < length l) (n:nat) (pfn:n < length l), lec_S pfdt pfm = lec_S pfdt pfn <-> nth_lt _ _ pfm = nth_lt _ _ pfn. intros; unfold lec_S; split; intro. apply inj_lS in H. rewrite lecs_eq_iff in H; auto. f_equal. rewrite lecs_eq_iff; auto. Qed. Lemma lecs_eq_count_occ : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (m n:nat) (pfm:m < length l) (pfn:n < length l'), lec pfdt pfm = lec pfdu pfn -> count_occ pfdt l (nth_lt _ _ pfm) = count_occ pfdu l' (nth_lt _ _ pfn). intros T U h1 h2 l l' m n h3 h4 h5. pose proof (f_equal (@length nat) h5) as h6. do 2 rewrite length_lec in h6; auto. Qed. Lemma impl_in_lec_app : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (m:nat) (pfm:m < length l) (pfm':m < length (l++l')) (r:nat), In r (lec pfdt pfm) -> In r (lec pfdt pfm'). intros T h1 l l' m h2 h3 r h4. pose proof (in_lec_lt h1 _ _ _ h2 h4) as h5. rewrite in_lec_iff in h4. rewrite in_lec_iff. rewrite (nth_lt_app1 _ _ _ h5). rewrite (nth_lt_app1 _ _ _ h2). apply h4. Grab Existential Variables. rewrite app_length; auto with arith. Qed. Lemma lec_lrep : forall {T:Type} (pfdt:type_eq_dec T) (a:T) (m n:nat) (pfm:m < length (lrep a n)), lec pfdt pfm = seg_list n. intros T h1 a m n h2. unfold lec. pose proof h2 as h2'. rewrite length_lrep in h2'. rewrite linds_occ_eq_iff. ex_goal. rewrite count_occ_lrep', length_seg_list; auto. apply in_nth_lt. exists hex. intros occ h3. rewrite nth_lt_seg_list. rewrite count_occ_lrep' in h3. apply lind_occ_lrep; auto. apply in_nth_lt. Grab Existential Variables. apply in_nth_lt. Qed. Lemma all_sing_lec : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (m:nat) (pfm:m < length l), all_sing l a -> lec pfdt pfm = seg_list (length l). intros T h1 l a m h2 h3. pose proof h3 as h3'. apply all_sing_lrep in h3'. gen0. rewrite h3'. intro h0. pose proof (length_lrep a (length l)) as h4. pose proof (lec_lrep h1 a m _ h0) as h5. rewrite h5. rewrite h3'. rewrite length_lrep; f_equal; auto. Qed. Lemma all_sing_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pfa:In a l), all_sing l a -> linds pfdt l = seg_list (length l) :: nil. intros T h1 l a h2 h3. unfold linds. rewrite (map_eq_iff (list_singularize h1 l) (linds_occ h1 l)). ex_goal. simpl. erewrite all_sing_list_singularize. simpl; auto. eapply in_not_nil. apply h2. apply h3. apply all_sing_lrep in h3. exists hex. intros n h4. simpl. rewrite linds_occ_eq_iff. ex_goal. rewrite h3 at 1. rewrite count_occ_lrep'; auto. do 3 (destruct n; repeat try rewrite length_seg_list; auto). rewrite <- h3. rewrite list_singularize_in_iff. apply in_nth_lt. exists hex0. intros m h5. destruct n as [|n]. rewrite nth_lt_seg_list. gen0. generalize h4. rewrite h3. intros h4' h7. apply lind_occ_lrep. rewrite h3 in h5 at 1. rewrite count_occ_lrep' in h5; auto. rewrite <- h3. rewrite list_singularize_in_iff. apply in_nth_lt. gen0. generalize dependent h4. rewrite h3. intros h4 h5 h6 i. rewrite lind_occ_lrep. destruct n as [|n]; try rewrite nth_lt_seg_list; auto. rewrite count_occ_lrep' in h6; auto. Grab Existential Variables. rewrite list_singularize_in_iff. apply in_nth_lt. Qed. Lemma neq_nin_0_lec_Sn : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (m m':nat) (pfm:m < length l) (pfsm:S m < length (a::l)), a <> nth_lt l m pfm -> ~In 0 (lec pfdt pfsm). intros T h1 l a m m' h2 h3 h4 h5. rewrite in_lec_iff in h5. simpl in h5. contradict h4. rewrite nth_lt_compat. symmetry. erewrite nth_indep. apply h5. assumption. Grab Existential Variables. simpl. apply lt_0_Sn. Qed. Lemma lec_0_cons_nin_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), ~In a l -> lec pfdt (lt_0_Sn (length l)) (l:=a::l) = 0::nil. intros T h0 l a h1. unfold lec. simpl. rewrite linds_occ_eq_iff. ex_goal. simpl. erewrite eq_dec_same at 1. rewrite (count_occ_In h0) in h1. f_equal. omega. assumption. exists hex. intros m h2. simpl. erewrite eq_dec_same at 1. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. contradiction. do 2 (destruct m; auto). assumption. Grab Existential Variables. left; auto. Qed. Definition linds_remove {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) := remove list_nat_eq_dec (linds_occ pfdt l a) (linds pfdt l). Lemma nin_linds_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), In a l -> ~In (linds_occ pfdt l a) (linds_remove pfdt l a). intros T h1 l a h2 h3. unfold linds_remove in h3. rewrite in_remove_iff in h3. destruct h3 as [? h3]; contradict h3; auto. Qed. Lemma in_linds_nth_lt_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (m n:nat) (pfln:In ln (linds pfdt l)) (pfm:In m ln) (pfn:In n ln), let pfm' := in_linds_lt _ _ _ _ pfm pfln in let pfn' := in_linds_lt _ _ _ _ pfn pfln in nth_lt l m pfm' = nth_lt l n pfn'. intros T h1 l ln m n h2 h3 h4; simpl. pose proof h2 as h2'. rewrite in_linds_iff in h2'. destruct h2' as [x [[h7 h8] h9]]; subst. erewrite nth_lt_functional3. symmetry. erewrite nth_lt_functional3. apply in_linds_occ_nth_lt_eq. apply h7. Grab Existential Variables. apply h4. assumption. Qed. Lemma in_linds_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (m:nat) (pfm:In m ln) (pfln: In ln (linds pfdt l)), let pf := in_linds_lt pfdt l ln _ pfm pfln in In (nth_lt _ _ pf) l. intros; apply in_nth_lt. Qed. Lemma no_dup_faml_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), no_dup_faml (linds pfdt l). intros T h1 l. red. intros l' h2. rewrite in_linds_iff in h2. destruct h2 as [x h3]; red in h3. destruct h3 as [[h3 h4] h5]; subst. apply no_dup_linds_occ. Qed. Lemma disj_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), pairwise_disjointl (linds pfdt l). intros T h1 l. red; intros lm ln h2 h3 h4; do 2 red. apply eq_empty. red; intros r h5. inversion h5 as [? h8 h9]; subst. rewrite <- list_to_set_in_iff in h8, h9. rewrite in_linds_iff in h3, h4. destruct h3 as [x [h3 h10]], h4 as [y [h4 h11]]. destruct h3 as [h3], h4 as [h4]; subst. rewrite in_linds_occ_iff in h8. destruct h8 as [m [h8 [h12 h13]]]; subst. rewrite in_linds_occ_iff in h9. destruct h9 as [m' [h9 [h13 h14]]]; subst. apply lind_occ_inj in h14; subst. contradict h2; auto. Qed. Lemma union_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), FamilyUnion (faml_to_fam (linds pfdt l)) = list_to_set (seg_list (length l)). intros T h1 l. apply Extensionality_Ensembles; red; split; red; intros r h2. inversion h2 as [S h3 h4 h5]; subst. pose proof (in_faml_to_fam_impl _ _ h4) as h6. destruct h6 as [ln [h6 h7]]; subst. rewrite <- list_to_set_in_iff. rewrite in_seg_list_iff. rewrite <- list_to_set_in_iff in h5. rewrite in_linds_iff in h6. destruct h6 as [x [[h7 h8] h9]]; subst. rewrite in_linds_occ_iff in h5. destruct h5 as [m [h10 [h11 h12]]]; subst. apply lt_lind_occ. rewrite <- list_to_set_in_iff in h2. rewrite in_seg_list_iff in h2. pose proof (family_union_intro _ (faml_to_fam (linds h1 l)) (list_to_set (linds_occ h1 l (nth_lt _ _ h2))) r) as h3. hfirst h3. apply in_faml_impl. apply in_linds_occ_linds. apply in_nth_lt. specialize (h3 hf). hfirst h3. rewrite <- list_to_set_in_iff. rewrite in_linds_occ_iff. pose proof (ex_lind_occ_nth_lt h1 _ _ h2) as h4. destruct h4 as [occ h4]. exists occ, (in_nth_lt _ _ _); auto. auto. Qed. Lemma partition_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), partition_faml (linds pfdt l) (seg_list (length l)). intros T h1 l. red; red. split. rewrite <- pairwise_disjointl_compat. apply disj_linds. apply union_linds. Qed. Lemma list_to_set_linds_eq_length_list_singularize : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), list_to_set (linds pfdt l) = list_to_set (linds pfdu l') -> length (list_singularize pfdt l) = length (list_singularize pfdu l'). intros T U h1 h2 l h3 h4. do 2 rewrite linds_compat' in h4. unfold linds' in h4. eapply list_to_set_map_in_deps_eq_length in h4. assumption. assumption. assumption. apply list_nat_eq_dec. apply no_dup_list_singularize. apply no_dup_list_singularize. apply inj_dep_incl. apply inj_dep_linds_occ'. apply inj_dep_incl. apply inj_dep_linds_occ'. Qed. Lemma count_occ_linds_occ' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfi:In x l), count_occ list_nat_eq_dec (linds' pfdt l) (linds_occ' pfdt l x pfi) = 1. intros T h1 l x h2; apply count_occ_no_dup. apply no_dup_linds'. rewrite in_linds_iff'. exists x. red; split. exists h2; auto. intros x' h3. destruct h3 as [h3 h4]. apply inj_dep_linds_occ' in h4; auto. Qed. Lemma count_occ_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfi:In x l), count_occ list_nat_eq_dec (linds pfdt l) (linds_occ pfdt l x) = 1. intros T h1 l x h2. rewrite linds_compat'. rewrite (linds_occ_compat' h1 l _ h2). apply count_occ_linds_occ'. Qed. Lemma count_occ_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat), In ln (linds pfdt l) -> count_occ list_nat_eq_dec (linds pfdt l) ln = 1. intros T h1 l ln h2. apply count_occ_no_dup; auto. apply no_dup_linds. Qed. Lemma count_occ_linds_cons_nin' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (ln:list nat), l <> nil -> ~In a l -> ~In 0 ln -> count_occ list_nat_eq_dec (linds' pfdt (a::l)) ln = count_occ list_nat_eq_dec (linds' pfdt l) (lpred ln). intros T h1 l a ln hn h2 h3. destruct (in_dec list_nat_eq_dec (lpred ln) (linds' h1 l)) as [h4 | h4]. pose proof h4 as h4'. rewrite in_linds_iff' in h4. destruct h4 as [x [[h4 h6] h5]]. rewrite count_occ_no_dup. rewrite h6. rewrite count_occ_linds_occ'; auto. apply no_dup_linds'. rewrite in_linds_iff'. exists x. red; split. exists (in_cons _ _ _ h4). rewrite linds_occ_cons_neq'. rewrite lS_lpred_iff in h3. rewrite h3. f_equal; auto. intro; subst; contradiction. intros x' h7. destruct h7 as [h7 h8]; subst. destruct (h1 x' a) as [|hf]; subst. contradict h3. rewrite in_linds_occ_iff'. exists 0. simpl. destruct (h1 a a) as [he | he]. split; auto with arith. contradict he. reflexivity. destruct h7 as [|h7]; subst. contradict hf; auto. pose proof (proof_irrelevance _ (or_intror h7) (in_cons a x' l h7)) as he. rewrite he in h6. rewrite linds_occ_cons_neq' in h6; auto. rewrite lpred_lS in h6. apply inj_dep_linds_occ' in h6; auto. pose proof h4 as h4'. rewrite (count_occ_In list_nat_eq_dec) in h4. change (~0 < count_occ list_nat_eq_dec (linds' h1 l) (lpred ln)) in h4. rewrite not_lt0_iff in h4. rewrite h4. rewrite <- not_lt0_iff. change (~count_occ list_nat_eq_dec (linds' h1 (a::l)) ln > 0). rewrite <- (count_occ_In list_nat_eq_dec). contradict h4'. rewrite in_linds_iff' in h4'. destruct h4' as [x [[h5 h6] h7]]; subst. destruct (h1 x a) as [|h8]; subst. contradict h3. rewrite in_linds_occ_iff'. exists 0. simpl. destruct (h1 a a) as [he | he]. split; auto with arith. contradict he. reflexivity. destruct h5 as [|h5]; subst. contradict h8; auto. rewrite in_linds_iff'. exists x. red; split. exists h5. pose proof (proof_irrelevance _ (or_intror h5) (in_cons a x l h5)) as he. rewrite he. rewrite linds_occ_cons_neq'. rewrite lpred_lS. reflexivity. assumption. intros x' h9. destruct h9 as [h9 h10]. apply h7; auto. exists (in_cons _ _ _ h9). assert (h11:x' <> a). intro; subst; contradiction. pose proof (proof_irrelevance _ (or_intror h5) (in_cons a x l h5)) as h12. rewrite h12. rewrite linds_occ_cons_neq'; auto. rewrite linds_occ_cons_neq'; auto. f_equal. rewrite h12 in h10. rewrite linds_occ_cons_neq' in h10; auto. rewrite lpred_lS in h10; auto. Qed. Definition coarse_list_in_list_of_lists_seqs_fun_list_nat {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln ln':list nat) (pfcl:coarse_list ln' ln) (pfi:In ln (list_of_lists_seqs (linds pfdt l))) : list T := map_seg (fun i (pfi':i < length ln') => let pfi0 := lt_le_trans _ _ _ pfi' (length_coarse_list _ _ pfcl) in let pfi1 := in_list_of_lists_seqs_lt (linds pfdt l) ln i pfi pfi0 in proj1_sig (constructive_definite_description _ (iff1 (in_linds_iff pfdt l _) (in_nth_lt (linds pfdt l) i pfi1)))). Lemma in_list_of_lists_seqs_mask_to_coarse_list : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (v:bv) (pfeq:length ln = length v), let ll := linds pfdt l in forall (pfi:In ln (list_of_lists_seqs ll)), let ln' := mask_to_coarse_list pfeq in let pfeq' := in_list_of_lists_seqs_length _ _ pfi in let pfeq0 := eq_trans (eq_sym pfeq') pfeq in let ll' := mask_to_coarse_list pfeq0 in In ln' (list_of_lists_seqs ll'). intros; apply coarse_list_in_list_of_lists_seqs. Qed. (*Returns a list with [true]s corresponding to elements whose [count_occ]s are equal to [n], [false] otherwise.*) Definition occ_count_bv {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) : bv := map (fun x => if (nat_eq_dec (count_occ pfdt l x) n) then true else false) (list_singularize pfdt l). Lemma occ_count_bv0 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), occ_count_bv pfdt l 0 = falses (length (list_singularize pfdt l)). intros T h1 l. unfold occ_count_bv, falses. rewrite (map_eq_iff (list_singularize h1 l) (fun x : T => if nat_eq_dec (count_occ h1 l x) 0 then true else false)). ex_goal. rewrite length_lrep; auto. exists hex. intros n h2. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. pose proof (in_nth_lt _ _ h2) as h3. rewrite <- list_singularize_in_iff in h3. rewrite (count_occ_In h1) in h3. omega. rewrite nth_lt_lrep; auto. Qed. Lemma occ_count_bv_length : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat), length (occ_count_bv pfdt l n) = length (list_singularize pfdt l). intros T h1 l n. unfold occ_count_bv. rewrite map_length; auto. Qed. (*A list of all elements that occur [n] times, in the order that they appear in [list_singuliarze l].*) Definition occ_count {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) : list T := mask_to_coarse_list (eq_sym (occ_count_bv_length pfdt l n)). Lemma count_occ_map_count_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), count_occ nat_eq_dec (map (count_occ pfdt l) (list_singularize pfdt l)) (count_occ pfdt l x) = length (inv_iml nat_eq_dec (count_occ pfdt l) (list_singularize pfdt l) (count_occ pfdt l x)). intros; apply count_occ_map. assumption. Qed. Definition linds_prod {T:Type} (pfdt:type_eq_dec T) (l:list T) : faml nat := list_of_lists_seqs (linds pfdt l). Lemma linds_prod_not_nil : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), p_mem_l (fun lx : list nat => lx <> nil) (linds pfdt l). intros T h1 l; red; intros lx h2. rewrite in_linds_iff in h2. destruct h2 as [x h2]. red in h2. destruct h2 as [h2 h3]. destruct h2 as [h2]; subst. intro h4. rewrite (linds_occ_eq_iff _ _ _ h2) in h4. destruct h4 as [h4]. simpl in h4. pose proof h2 as h2'. rewrite (count_occ_In h1) in h2'. omega. Qed. Lemma linds_prod_not_nil' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), l <> nil -> linds_prod pfdt l <> nil. intros T h1 l h2. unfold linds_prod. apply list_of_lists_seqs_not_nil. apply linds_prod_not_nil. Qed. Lemma no_dup_linds_prod : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), NoDup (linds_prod pfdt l). intros T h1 l. unfold linds_prod. apply list_of_lists_seqs_no_dup. apply no_dup_faml_linds. Qed. Lemma in_linds_prod_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat), In ln (linds_prod pfdt l) <-> (length ln = length (list_singularize pfdt l) /\ (forall (i:nat) (pfi:i < length ln), exists pfi':i < length (linds pfdt l), In (nth_lt ln i pfi) (nth_lt _ i pfi'))). intros T h1 l ln. unfold linds_prod. rewrite in_list_of_lists_seqs_iff. rewrite length_linds at 1; tauto. Qed. Lemma in_linds_prod_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (i:nat), In ln (linds_prod pfdt l) -> In i ln -> i < length l. intros T h1 l ln i h2 h3. rewrite in_linds_prod_iff in h2. destruct h2 as [h2 h4]. specialize (h4 (lind nat_eq_dec _ _ h3) (lt_lind _ _ _ _)). destruct h4 as [h4 h5]. rewrite lind_compat in h5. unfold linds in h5. erewrite nth_lt_map in h5. eapply in_linds_occ_lt. apply h5. Grab Existential Variables. eapply lt_eq_trans. apply lt_lind. assumption. Qed. Lemma in_linds_prod_lind_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (pfilnlp:In ln (linds_prod pfdt l)) (i:nat) (pfiln:In i ln), lind nat_eq_dec ln i pfiln < length (linds pfdt l). intros T h1 l ln h2 i h3. rewrite in_linds_prod_iff in h2. destruct h2 as [h2 h4]. specialize (h4 _ (lt_lind nat_eq_dec ln i h3)). destruct h4; auto. Qed. Lemma in_linds_prod_in_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (i:nat) (pfilnlp:In ln (linds_prod pfdt l)) (pfiln:In i ln), let pflt := in_linds_prod_lt pfdt l ln i pfilnlp pfiln in In (linds_occ pfdt l (nth_lt _ _ pflt)) (linds pfdt l). intros; apply in_linds_occ_linds; apply in_nth_lt. Qed. Lemma in_linds_prod_length : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat), In ln (linds_prod pfdt l) -> length ln = length (list_singularize pfdt l). intros T h1 l n h2. rewrite in_linds_prod_iff in h2. destruct h2; auto. Qed. Lemma nil_in_linds_prod : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), In nil (linds_prod pfdt l) -> l = nil. intros T h1 l h2. rewrite in_linds_prod_iff in h2. simpl in h2. destruct h2 as [h2]. symmetry in h2. apply length_eq0 in h2. rewrite list_singularize_nil_iff in h2; auto. Qed. Lemma in_cons_linds_prod : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (i:nat), In (i::ln) (linds_prod pfdt l) -> In ln (list_of_lists_seqs (tl (linds pfdt l))). intros T h1 l ln i h3. rewrite in_linds_prod_iff in h3. simpl in h3. destruct h3 as [h3 h4]. rewrite in_list_of_lists_seqs_iff. split. rewrite <- (length_linds h1 l) in h3. pose proof (lt_length (linds h1 l) 0) as h5. rewrite <- h3 in h5. specialize (h5 (lt_0_Sn _)). rewrite (hd_compat' _ h5) in h3. simpl in h3. apply S_inj in h3; auto. intros j h5. specialize (h4 (S j) (lt_n_S _ _ h5)). simpl in h4. destruct h4 as [h4 h6]. pose proof (lt_length (linds h1 l) (S j) h4) as h7. ex_goal. pose proof h4 as h4'. rewrite (hd_compat' _ h7) in h4'. simpl in h4'. apply lt_S_n in h4'; auto. exists hex. rewrite nth_lt_compat in h6. erewrite <- (nth_indep (linds h1 l) ln) in h6; auto. rewrite (hd_compat' _ h7) in h6. simpl in h6. rewrite nth_lt_compat, (nth_lt_compat (tl (linds h1 l))). erewrite nth_indep, (nth_indep (tl (linds h1 l))); auto. apply h6. Qed. Definition dups {T:Type} (pfdt:type_eq_dec T) (l:list T) : list T := list_singularize pfdt (map_filter_gt nat_eq_dec lt Rlem_dec_lt_nat (count_occ pfdt l) 1 l). Lemma no_dup_dups : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), NoDup l -> dups pfdt l = nil. intros T h1 l. unfold dups. rewrite list_singularize_nil_iff. unfold map_filter_gt. rewrite map_filter_eq_nil_iff. destruct (nil_dec' l) as [|h2]; subst. intros; left; auto. intro h3. right. pose proof (count_occ_no_dup h1 l h3) as h4. intros x h5. specialize (h4 x h5). omega. Qed. Lemma nth_lt_in_or_0 : forall {T:Type} (l:list T) (n:nat) (pf:n < length l), n = 0 \/ (n > 0 /\ In (nth_lt l n pf) l). intros T l n. revert l. induction n as [|n h1]; try tauto. intros l h2. right. split. auto with arith. pose proof (nth_lt_compat l (S n) h2) as h3. simpl in h3. destruct h3; auto. apply in_nth_lt. Qed. Lemma nth_lt_lind_hd : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pfn:l <> nil) (pfin:In (hd' l pfn) l) (pflt:lind pfdt l (hd' l pfn) pfin < length l), nth_lt l (lind pfdt l (hd' l pfn) pfin) pflt = hd' l pfn. intros T h0 l h1 h2 h3. pose proof (lind_compat h0 l (hd' l h1) h2) as h4. erewrite nth_lt_functional2. rewrite h4 at 1; auto. reflexivity. Qed. Definition l_to_inds_aux {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin:In x l) : nat := let pfin := iff1 (list_singularize_in_iff pfdt l x) pfin in lind pfdt _ _ pfin. (*Returns a list with each element of [l] replaced by its index in [list_singularize].*) Definition l_to_inds {T:Type} (pfdt:type_eq_dec T) (l:list T) : list nat := map_in_dep (l_to_inds_aux pfdt l). Lemma l_to_inds_aux_functional3 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfin pfin':In x l), l_to_inds_aux pfdt l x pfin = l_to_inds_aux pfdt l x pfin'. intros; f_equal; apply proof_irrelevance. Qed. Lemma length_l_to_inds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), length (l_to_inds pfdt l) = length l. unfold l_to_inds; intros. apply length_map_in_dep. Qed. Lemma in_l_to_inds_aux : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), In (l_to_inds_aux pfdt l x pf) (l_to_inds pfdt l). unfold l_to_inds_aux, l_to_inds; intros. rewrite in_map_in_dep_iff. exists x, pf. unfold l_to_inds_aux; auto. Qed. Lemma l_to_inds_eq_length : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), l_to_inds pfdt l = l_to_inds pfdu l' -> length l = length l'. unfold l_to_inds; intros T U hdt hdu l l' h1. rewrite map_in_dep_eq_iff in h1. destruct h1 as [h1 h2]. clear h2. rewrite length_map_in_dep in h1; auto. Qed. Lemma l_to_inds_aux_cons_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pfi:In a l) (pfi':In a (a::l)), l_to_inds_aux pfdt (a::l) a pfi' = l_to_inds_aux pfdt l a pfi. intros T h1 l a h2 h3. unfold l_to_inds_aux. simpl. gen0. rewrite list_singularize_in_cons; auto. intro h5. f_equal. apply proof_irrelevance. Qed. Lemma l_to_inds_aux_cons_nin : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pfi:In a (a::l)), ~In a l -> l_to_inds_aux pfdt (a::l) a pfi = 0. intros T h1 l a h2 h3. unfold l_to_inds_aux. simpl. gen0. rewrite list_singularize_nin; auto. intro h5. erewrite lind_functional. rewrite lind_cons_same; auto. reflexivity. Grab Existential Variables. left; auto. Qed. Lemma l_to_inds_aux_cons_in_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a x:T) (pfn:x <> a) (pfi:In a l) (pfi:In x (a::l)), l_to_inds_aux pfdt (a::l) x pfi = l_to_inds_aux pfdt l x (neq_in_cons _ _ _ pfn pfi). intros T h1 l a x h2 h3 h4. destruct h4 as [|h4]; subst. contradict h2; auto. unfold l_to_inds_aux. gen0. rewrite list_singularize_in_cons. intro h5. f_equal; apply proof_irrelevance. assumption. Qed. Lemma l_to_inds_aux_cons_nin_in: forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a x:T) (pfn:x <> a) (pfi:~In a l) (pfi':In x (a::l)), l_to_inds_aux pfdt (a::l) x pfi' = S (l_to_inds_aux pfdt l x (neq_in_cons _ _ _ pfn pfi')). intros T h1 l a x h2 h3 h4. destruct h4 as [|h4]; subst. contradict h2; auto. unfold l_to_inds_aux. gen0. rewrite list_singularize_nin; auto. intro h5. pose proof (lind_cons_neq h1 (list_singularize h1 l) a x) as h6. hfirst h6. rewrite <- list_singularize_in_iff; auto. specialize (h6 hf h2). pose proof (proof_irrelevance _ h5 (or_intror hf)). subst. rewrite h6. repeat f_equal; apply proof_irrelevance. Qed. Lemma fun_from_cons_l_to_inds_aux_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), In a l -> fun_from_cons (l_to_inds_aux pfdt (a::l)) = l_to_inds_aux pfdt l. intros T h1 l a hi. apply functional_extensionality_dep. intro x. apply functional_extensionality. intro h2. unfold fun_from_cons. destruct (h1 x a) as [|h3]; subst. apply l_to_inds_aux_cons_in. rewrite (l_to_inds_aux_cons_in_in h1 _ _ _ h3 hi (in_cons _ _ _ h2)). f_equal; apply proof_irrelevance. Qed. Lemma fun_from_cons_l_to_inds_aux_nin : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), ~In a l -> fun_from_cons (l_to_inds_aux pfdt (a::l)) = fun x pf => S (l_to_inds_aux pfdt l x pf). intros T h1 l a hni. apply functional_extensionality_dep. intro x. apply functional_extensionality. intro h2. assert (h3:x <> a). intro; subst; contradiction. pose proof (l_to_inds_aux_cons_nin_in h1 l a x h3 hni (in_cons a _ _ h2)) as h4. unfold fun_from_cons. erewrite l_to_inds_aux_functional3. rewrite h4 at 1. repeat f_equal. apply proof_irrelevance. Qed. Lemma l_to_inds_cons_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:In a l), l_to_inds pfdt (a::l) = (l_to_inds_aux pfdt l a pf) :: l_to_inds pfdt l. intros T h1 l a h2. unfold l_to_inds; simpl. pose proof (l_to_inds_aux_cons_in h1 l a h2 (in_eq _ _)) as h3. rewrite h3, fun_from_cons_l_to_inds_aux_in; auto. Qed. Lemma l_to_inds_cons_nin : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:~In a l), l_to_inds pfdt (a::l) = 0 :: lS (l_to_inds pfdt l). intros T h1 l a h2. unfold l_to_inds; simpl. pose proof (l_to_inds_aux_cons_nin h1 l a (in_eq _ _) h2) as h3. rewrite h3, fun_from_cons_l_to_inds_aux_nin; auto. repeat f_equal. rewrite lS_map_in_dep; auto. Qed. Lemma l_to_inds_aux_inj_dep : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), inj_dep (l_to_inds_aux pfdt l). intros T h1 l; red. intros x y h2 h3 h4. unfold l_to_inds_aux in h4. rewrite (lind_eq_iff _ _ _ _ _ (lt_lind _ _ _ _)) in h4. destruct h4; subst. rewrite (nth_lt_eq_iff h1). ex_goal. rewrite <- list_singularize_in_iff; auto. exists hex. exists 0. split. apply gt_lt. rewrite <- (count_occ_In h1). rewrite <- list_singularize_in_iff; auto. unfold lind; f_equal; apply proof_irrelevance. Qed. Lemma linds_l_to_inds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), linds nat_eq_dec (l_to_inds pfdt l) = linds pfdt l. intros T h1 l. unfold linds, l_to_inds. rewrite (map_eq_iff (list_singularize nat_eq_dec (map_in_dep (l_to_inds_aux h1 l))) (linds_occ nat_eq_dec (map_in_dep (l_to_inds_aux h1 l)))). ex_goal. rewrite map_length. apply (length_list_singularize_map_in_dep_inj h1). apply l_to_inds_aux_inj_dep. exists hex. intros n h2. pose proof h2 as h2'. rewrite (length_list_singularize_map_in_dep_inj h1) in h2'. rewrite nth_lt_map'. rewrite linds_occ_eq_iff. ex_goal. rewrite length_linds_occ. pose proof (nth_lt_list_singularize_map_in_dep h1 nat_eq_dec (l_to_inds_aux h1 l) n (l_to_inds_aux_inj_dep h1 l) h2') as h3. simpl in h3. erewrite nth_lt_functional3. rewrite h3 at 1. rewrite (count_occ_map_in_dep_inj h1). f_equal. apply nth_lt_functional3. apply l_to_inds_aux_inj_dep. exists hex0. intros m h4. pose proof (nth_lt_linds_occ h1 l (nth_lt (list_singularize h1 l) n (lt_eq_trans n (length (map (linds_occ h1 l) (list_singularize h1 l))) (length (list_singularize h1 l)) (lt_eq_trans n (length (list_singularize nat_eq_dec (map_in_dep (l_to_inds_aux h1 l)))) (length (map (linds_occ h1 l) (list_singularize h1 l))) h2 hex) (map_length (linds_occ h1 l) (list_singularize h1 l))))) as h5. hfirst h5. rewrite list_singularize_in_iff. apply in_nth_lt. specialize (h5 hf m). hfirst h5. rewrite length_linds_occ. pose proof h4 as h4'. erewrite nth_lt_functional3 in h4'. rewrite nth_lt_list_singularize_map_in_dep in h4' at 1. rewrite count_occ_map_in_dep_inj in h4' at 1. apply h4'. apply l_to_inds_aux_inj_dep. specialize (h5 hf0). symmetry. erewrite nth_lt_functional3. rewrite h5 at 1. pose proof (lind_occ_map_in_dep_inj h1 nat_eq_dec (l_to_inds_aux h1 l) (nth_lt _ _ h2')) as h6. hfirst h6. rewrite list_singularize_in_iff. apply in_nth_lt. specialize (h6 hf1). hfirst h6. apply in_map_in_dep. specialize (h6 hf2 m). hfirst h6. apply l_to_inds_aux_inj_dep; auto. specialize (h6 hf3). symmetry. gen0. erewrite nth_lt_functional3. rewrite nth_lt_list_singularize_map_in_dep. intro h7. pose proof (proof_irrelevance _ hf1 (iff2 (list_singularize_in_iff h1 l (nth_lt (list_singularize h1 l) n h2')) (in_nth_lt (list_singularize h1 l) n h2'))). subst. pose proof (proof_irrelevance _ hf2 h7); subst. rewrite h6 at 1. apply lind_occ_functional; auto. apply nth_lt_functional3. apply l_to_inds_aux_inj_dep. Grab Existential Variables. apply l_to_inds_aux_inj_dep. apply l_to_inds_aux_inj_dep. rewrite in_map_in_dep_iff. exists (nth_lt _ _ h2'). ex_goal. rewrite list_singularize_in_iff. apply in_nth_lt. exists hex0. erewrite nth_lt_functional3. rewrite nth_lt_list_singularize_map_in_dep. apply l_to_inds_aux_functional3. Grab Existential Variables. apply l_to_inds_aux_inj_dep. Qed. Lemma linds_list_to_set_injable : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), list_to_set_injable (fun ll => ll = linds pfdt l). intros T h1 l; red; intros; subst; auto. Qed. Fixpoint count_occ_l {T:Type} (pfdt:type_eq_dec T) (l l':list T) : nat := let pfdt' := list_eq_dec pfdt in match l' with | nil => if (pfdt' l nil) then 1 else 0 | a::la => let co := count_occ_l pfdt l la in if (pfdt' l (firstn (length l) l')) then S co else co end. (*Finds the index of the [S occ]th occureence of [l] in [l'] as a sub_list. If [occ >= pred (count_occ_l)] then the value is the same as [pred(count_occ_l)]'s.*) Fixpoint slind_occ {T:Type} (pfdt:type_eq_dec T) (l l':list T) (occ:nat) (pfln:l <> nil) : l' <> nil -> sub_list l l' -> nat := let pfdt' := list_eq_dec pfdt in match l' with | nil => fun pfn _ => False_rect _ (pfn eq_refl) | a::la => fun _ pfsl => match (nil_dec' la) with | left _ => 0 | right pfnla => match (pfdt' l (firstn (length l) (a::la))) with | left _ => match occ with | 0 => 0 | S n => match (sub_list_dec pfdt l la) with | left pfsl' => S (slind_occ pfdt l la n pfln pfnla pfsl') | right _ => 0 end end | right pfn => let pf' := neq_sub_list_cons l la a pfn pfsl in S (slind_occ pfdt l la occ pfln pfnla pf') end end end. (*This is the "right index" (corresponding to the last element in [l]) of the [S occ]s occurence of [l] in [l'].*) Definition slind_occ' {T:Type} (pfdt:type_eq_dec T) (l l':list T) (occ:nat) (pfn:l <> nil) (pfn':l' <> nil) (pfsl:sub_list l l') : nat := slind_occ pfdt l l' occ pfn pfn' pfsl + (length l). Definition slind {T:Type} (pfdt:type_eq_dec T) (l l':list T) (pfn:l <> nil) (pfn':l' <> nil) (pfsl:sub_list l l') := slind_occ pfdt l l' 0 pfn pfn' pfsl. Definition slind' {T:Type} (pfdt:type_eq_dec T) (l l':list T) (pfn:l <> nil) (pfn':l' <> nil) (pfsl:sub_list l l') : nat := slind_occ' pfdt l l' 0 pfn pfn' pfsl. Lemma slind_occ_functional : forall {T:Type} (pfdt:type_eq_dec T) (l l' l0 l0':list T) (occ occ':nat) (pfn:l <> nil) (pfn':l' <> nil) (pfn0:l0 <> nil) (pfn0':l0' <> nil) (pf:sub_list l l0) (pf':sub_list l' l0'), l = l' -> l0 = l0' -> occ = occ' -> slind_occ pfdt l l0 occ pfn pfn0 pf = slind_occ pfdt l' l0' occ' pfn' pfn0' pf'. intros; subst; pose proof (proof_irrelevance _ pf pf'); pose proof (proof_irrelevance _ pfn pfn'); pose proof (proof_irrelevance _ pfn0 pfn0'); subst; auto. Qed. Lemma slind_occ'_functional' : forall {T:Type} (pfdt:type_eq_dec T) (l l' l0 l0':list T) (occ occ':nat) (pfn:l <> nil) (pfn':l' <> nil) (pfn0:l0 <> nil) (pfn0':l0' <> nil) (pf:sub_list l l0) (pf':sub_list l' l0'), l = l' -> l0 = l0' -> occ = occ' -> slind_occ' pfdt l l0 occ pfn pfn0 pf = slind_occ' pfdt l' l0' occ' pfn' pfn0' pf'. intros; subst; pose proof (proof_irrelevance _ pf pf'); pose proof (proof_irrelevance _ pfn pfn'); pose proof (proof_irrelevance _ pfn0 pfn0'); subst; auto. Qed. Lemma slind_functional : forall {T:Type} (pfdt:type_eq_dec T) (l l' l0 l0':list T) (pfn:l <> nil) (pfn':l' <> nil) (pfn0:l0 <> nil) (pfn0':l0' <> nil) (pf:sub_list l l0) (pf':sub_list l' l0'), l = l' -> l0 = l0' -> slind pfdt l l0 pfn pfn0 pf = slind pfdt l' l0' pfn' pfn0' pf'. intros; subst; pose proof (proof_irrelevance _ pf pf'); pose proof (proof_irrelevance _ pfn pfn'); pose proof (proof_irrelevance _ pfn0 pfn0'); subst; auto. Qed. Lemma slind_functional' : forall {T:Type} (pfdt:type_eq_dec T) (l l' l0 l0':list T) (pfn:l <> nil) (pfn':l' <> nil) (pfn0:l0 <> nil) (pfn0':l0' <> nil) (pf:sub_list l l0) (pf':sub_list l' l0'), l = l' -> l0 = l0' -> slind' pfdt l l0 pfn pfn0 pf = slind' pfdt l' l0' pfn' pfn0' pf'. intros; subst; pose proof (proof_irrelevance _ pf pf'); pose proof (proof_irrelevance _ pfn pfn'); pose proof (proof_irrelevance _ pfn0 pfn0'); subst; auto. Qed. Lemma slind_firstn_length : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (pf:l <> nil) (pf':l' <> nil) (pfsl:sub_list l l'), firstn (length l) l' = l -> slind pfdt l l' pf pf' pfsl = 0. intros T h1 l l'. revert l. induction l' as [|a' l' ih]; simpl. intros ? ? h2. contradict h2; auto. intros l h2 h3 h4 h5; unfold slind; simpl. destruct (nil_dec' l') as [h6 | h6]; auto. lr_match_goal; fold hlr; destruct hlr as [hl | hr]; auto. pose proof hr as hr'. rewrite h5 in hr'. contradict hr'; auto. Qed. Lemma slind_occ_lt : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (pf:l <> nil) (pf':l' <> nil) (occ:nat) (pfsl:sub_list l l'), slind_occ pfdt l l' occ pf pf' pfsl < length l'. intros T h1 l l'. revert l. induction l' as [|a' l' ih]; simpl. intros ? ? h2; contradict h2; auto. intros l h2 h3 occ h4. destruct (nil_dec' l') as [|h5]; auto with arith. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. destruct occ as [|occ]; auto with arith. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. apply lt_n_S. apply ih; auto with arith. auto with arith. apply lt_n_S, ih; auto. Qed. (*Returns the pair of surrounding lists of the first occurence of [l] in [l'].*) Definition sub_list_partition0 {T:Type} (pfdt:type_eq_dec T) (l l':list T) (pfn:l <> nil) (pfn':l' <> nil) (pfsl:sub_list l l') : list T * list T := let n := slind pfdt l l' pfn pfn' pfsl in let n' := slind' pfdt l l' pfn pfn' pfsl in (firstn n l', skipn n' l'). Lemma sub_list_partition0_compat : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (pfn:l <> nil) (pfn':l' <> nil) (pfsl:sub_list l l'), let lp := sub_list_partition0 pfdt l l' pfn pfn' pfsl in let lfp := fst lp in let lsp := snd lp in l' = lfp ++ l ++ lsp. intros T h1 l l'. revert l. induction l' as [|a' l' ih]; simpl. intros ? ? h2. contradict h2; auto. intros l h2 h3 h4. unfold slind, slind', slind_occ'. simpl. destruct (nil_dec' l') as [h5 | h5]. subst. simpl. unfold slind', slind_occ'. simpl. apply sub_list_sing_dec in h4. destruct h4; subst. contradict h2; auto. simpl; auto. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. simpl. rewrite hl. rewrite length_firstn. rewrite firstn_skipn; auto. apply length_sub_list; auto. simpl. f_equal. apply ih. Qed. Lemma map_in_dep_decompose : forall {T U:Type} (pfdt:type_eq_dec T) (l l':list T) (pfn: l <> nil) (pfn':l' <> nil) (f:fun_in_dep l' U) (pfsl:sub_list l l'), let (lx, ly) := sub_list_partition0 pfdt l l' pfn pfn' pfsl in let n := slind pfdt l l' pfn pfn' pfsl in let n' := slind' pfdt l l' pfn pfn' pfsl in let pfslx := sub_list_firstn _ n in let pfsly := sub_list_skipn _ n' in let fl := fun_in_dep_decompose1' pfdt l l' pfsl f in let flx := fun_in_dep_decompose1' pfdt _ l' pfslx f in let fly := fun_in_dep_decompose1' pfdt _ l' pfsly f in map_in_dep f = map_in_dep flx ++ map_in_dep fl ++ map_in_dep fly. unfold fun_in_dep_decompose1', fun_in_dep_decompose2', fun_in_dep_decompose1, fun_in_dep_decompose2; intros T U h1 l l' h2 h3 f h4. simpl. pose proof (sub_list_partition0_compat h1 l l' h2 h3 h4) as h6. simpl in h6. rewrite (map_in_dep_functional _ h6). do 2 rewrite map_in_dep_app. f_equal. unfold fun_from_app1. apply map_in_dep_ext. intros x h7. apply functional_extensionality. intro h8. rewrite transfer_fun_in_dep_compat. f_equal. apply proof_irrelevance. f_equal. apply map_in_dep_ext. intros x h7. unfold fun_from_app2, fun_from_app1. apply functional_extensionality. intro h8. rewrite transfer_fun_in_dep_compat. f_equal. apply proof_irrelevance. unfold fun_from_app2. apply map_in_dep_ext. intros x h7. apply functional_extensionality. intro h8. rewrite transfer_fun_in_dep_compat. f_equal. apply proof_irrelevance. Qed. Lemma map_in_dep_firstn_S : forall {T U:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfsn : S n <= length l) (f:fun_in_dep (firstn (S n) l) U), let pfsl := sub_list_firstn_le _ _ _ (le_n_Sn _) in let pflt := lt_le_trans _ _ _ (lt_n_Sn _) pfsn in let f' := fun_in_dep_decompose1' pfdt _ _ pfsl f in map_in_dep f = map_in_dep f' ++ f (nth_lt l n pflt) (in_nth_lt_firstn' _ _ _ pfsn (lt_n_Sn _)) :: nil. intros T U h1 l n h2 f h3 h4. destruct n as [|n]. destruct l as [|a l]. simpl in h2. omega. simpl. repeat f_equal. apply proof_irrelevance. pose proof (map_in_dep_decompose h1 (firstn (S n) l) (firstn (S (S n)) l) (firstn_S_neq_nil _ _ (le_trans _ _ _ (le_n_Sn _) h2)) (firstn_S_neq_nil _ _ h2) f h3) as h5. simpl. simpl. simpl in h5. rewrite h5. unfold fun_in_dep_decompose1', fun_in_dep_decompose1. match goal with |- map_in_dep ?g' ++ _ ++ _ = _ => pose g' as g end. let P := type of g in match P with forall x:T, In x ?l' -> _ => pose l' as l0 end. assert (h8:l0 = nil). unfold l0. rewrite slind_firstn_length. simpl; auto. change (firstn (length (firstn (S n) l)) (firstn (S (S n)) l) = firstn (S n) l). rewrite length_firstn. rewrite <- firstn_le_nest; auto. auto with arith. fold l0 in g. fold l0 g. rewrite (map_in_dep_functional g h8). simpl. f_equal. match goal with |- map_in_dep ?g' = _ => pose g' as h end. let P := type of h in match P with forall x:T, In x ?l' -> _ => pose l' as l1 end. assert (h9:l1 = nth_lt l (S n) h4 :: nil). unfold l1. unfold slind', slind_occ'. match goal with |- skipn (slind_occ h1 ?x2 ?x3 0 ?x4 ?x5 ?x6 + ?x7) _ = _ => change (skipn (slind h1 x2 x3 x4 x5 x6 + x7) (firstn (S (S n)) l) = nth_lt l (S n) h4 ::nil ) end. rewrite slind_firstn_length. simpl. change (skipn (length (firstn (S n) l)) (firstn (S (S n)) l) = nth_lt l (S n) h4::nil). rewrite length_firstn_S. rewrite (firstn_S_eq _ _ h2). rewrite skipn_app. rewrite skipn_firstn_nil. rewrite app_nil_l. repeat f_equal. apply proof_irrelevance. rewrite length_firstn; auto with arith. omega. change (firstn (length (firstn (S n) l)) (firstn (S (S n)) l) = firstn (S n) l). rewrite length_firstn. rewrite <- firstn_le_nest; auto. auto with arith. fold l1 in h. fold l1 h. rewrite (map_in_dep_functional h h9). simpl. rewrite transfer_fun_in_dep_compat. unfold h. repeat f_equal. apply proof_irrelevance. Qed. Lemma firstn_map_in_dep_removelast_S : forall {T U:Type} {l:list T} (f:fun_in_dep l U) (n:nat), S n <= length l -> firstn n (map_in_dep (fun x pf => f x (incl_removelast l _ pf))) = firstn n (map_in_dep f). intros T U l f n h1. pose proof (removelast_eq l) as h2. pose proof (map_in_dep_functional (fun (x : T) (pf : (fun x0 : T => In x0 (removelast l)) x) => f x (incl_removelast l x pf)) h2) as h3. rewrite h3 at 1. pose proof (firstn_map_in_dep_firstn f n (pred (length l))) as h4. hfirst h4. omega. specialize (h4 hf). hfirst h4. auto with arith. specialize (h4 hf0). rewrite <- h4. repeat f_equal. apply functional_extensionality_dep. intro x. apply functional_extensionality. intro h5. rewrite transfer_fun_in_dep_compat. f_equal. apply proof_irrelevance. Qed. Lemma firstn_map_in_dep : forall {T U:Type} (pfdt:type_eq_dec T) (l:list T) (f:fun_in_dep l U) (n:nat) (pfle:n <= length l), firstn n (map_in_dep f) = map_in_dep (fun x (pf:In x (firstn n l)) => f x (incl_firstn _ _ pfle _ pf)). intros T U hdt l f n. revert f. revert l. induction n as [|n ih]; auto. intros l f h0. specialize (ih (removelast l) (fun x pf => f x (incl_removelast l _ pf))). hfirst ih. rewrite length_removelast; omega. specialize (ih hf). erewrite firstn_S_eq. rewrite firstn_map_in_dep_removelast_S in ih at 1. rewrite ih. symmetry. erewrite map_in_dep_firstn_S. pose proof (@map_in_dep_functional). pose proof (firstn_removelast l n) as h1. hfirst h1. rewrite length_removelast; omega. specialize (h1 hf0). pose proof (map_in_dep_functional (fun (x : T) (pf : In x (firstn n (removelast l))) => f x (incl_removelast l x (incl_firstn (removelast l) n hf x pf))) h1) as h2. rewrite h2 at 1. f_equal. f_equal. apply functional_extensionality_dep. intro x. apply functional_extensionality. intro h3. rewrite transfer_fun_in_dep_compat. unfold fun_in_dep_decompose1', fun_in_dep_decompose1. f_equal. apply proof_irrelevance. rewrite nth_lt_map_in_dep'. f_equal. apply f_equal_dep. apply nth_lt_functional3. assumption. Grab Existential Variables. assumption. assumption. rewrite length_map_in_dep; auto with arith. Qed. Lemma nth_lt_lind_occ_l_to_inds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (occ:nat) (x:T) (pfi:In x l) (pf:lind_occ nat_eq_dec (l_to_inds pfdt l) occ (l_to_inds_aux pfdt l x pfi) (in_l_to_inds_aux pfdt l x pfi) < length l), (nth_lt l (lind_occ nat_eq_dec (l_to_inds pfdt l) occ (l_to_inds_aux pfdt l x pfi) (in_l_to_inds_aux pfdt l x pfi)) pf) = x. intros T h1 l occ x h2 h3. pose proof (lind_occ_compat nat_eq_dec (l_to_inds h1 l) occ (l_to_inds_aux h1 l x h2) (in_l_to_inds_aux h1 l x h2)) as h4. unfold l_to_inds in h4. rewrite nth_lt_map_in_dep' in h4. apply l_to_inds_aux_inj_dep in h4. rewrite <- h4. apply nth_lt_functional3. Qed. Lemma count_occ_l_to_inds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), count_occ nat_eq_dec (l_to_inds pfdt l) (l_to_inds_aux pfdt l x pf) = count_occ pfdt l x. intros T h1 l. induction l as [|a l ih]; simpl; auto. intros x h2. destruct h2 as [|h2]; subst. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. rewrite eq_dec_same at 1. f_equal. destruct (in_dec h1 x l) as [h2 | h2]. rewrite fun_from_cons_l_to_inds_aux_in. rewrite (l_to_inds_aux_cons_in _ _ _ h2). apply ih. assumption. pose proof h2 as h2'. rewrite (count_occ_In h1) in h2'. apply not_lt0_iff in h2'. rewrite h2'. rewrite fun_from_cons_l_to_inds_aux_nin. rewrite (l_to_inds_aux_cons_nin _ _ _ _ h2). apply count_occ_gt0. intros n h3. rewrite in_map_in_dep_iff in h3. destruct h3 as [y [h3 h4]]; subst; auto with arith. assumption. assumption. contradict hr. f_equal; apply proof_irrelevance. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. destruct (in_dec h1 a l) as [h3 | h3]. rewrite (l_to_inds_aux_cons_in _ _ _ h3) in hl. destruct (h1 a x) as [h4 | h4]; subst. f_equal. rewrite fun_from_cons_l_to_inds_aux_in; auto. rewrite (l_to_inds_aux_cons_in _ _ _ h2). apply ih. apply neq_sym in h4. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h4) in hl. apply l_to_inds_aux_inj_dep in hl; subst. contradict h4; auto. assumption. rewrite l_to_inds_aux_cons_nin in hl; auto. destruct (h1 a x) as [|h4]; subst. contradiction. apply neq_sym in h4. rewrite (l_to_inds_aux_cons_nin_in _ _ _ _ h4) in hl; auto. discriminate. destruct (h1 a x) as [|h3]; subst. contradict hr. repeat f_equal. apply proof_irrelevance. apply neq_sym in h3. destruct (in_dec h1 a l) as [h4 | h4]. rewrite fun_from_cons_l_to_inds_aux_in; auto. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h3); auto. rewrite (l_to_inds_aux_cons_nin_in _ _ _ _ h3); auto. rewrite fun_from_cons_l_to_inds_aux_nin; auto. rewrite <- (ih _ h2). rewrite <- lS_map_in_dep. rewrite count_occ_lS. simpl. repeat f_equal. apply proof_irrelevance. auto with arith. Qed. Lemma count_occ_firstn_lind_occ_l_to_inds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) (occ:nat), occ < count_occ pfdt l x -> count_occ pfdt (firstn (lind_occ nat_eq_dec (l_to_inds pfdt l) occ (l_to_inds_aux pfdt l x pf) (in_l_to_inds_aux pfdt l x pf)) l) x = occ. intros T h1 l. induction l as [|a l ih]. intros; contradiction. intros x h2 occ h3. destruct (h1 a x) as [|h4]; subst. destruct (in_dec h1 x l) as [h4 | h4]. generalize (in_l_to_inds_aux h1 (x::l) x h2). rewrite (l_to_inds_aux_cons_in _ _ _ h4). rewrite (l_to_inds_cons_in _ _ _ h4). destruct occ as [|occ]. simpl. intro h5. rewrite eq_dec_same at 1. simpl. reflexivity. apply nat_eq_dec. rewrite count_occ_cons_eq in h3; auto. apply lt_S_n in h3. intro h5. erewrite lind_occ_functional3. rewrite lind_occ_hd at 1. simpl. rewrite eq_dec_same at 1. f_equal. specialize (ih _ h4 occ h3). erewrite lind_occ_functional3. rewrite ih at 1; auto. assumption. rewrite count_occ_cons_eq in h3; auto. pose proof h4 as h4'. rewrite (count_occ_In h1) in h4'. apply ngt_nlt in h4'. rewrite not_lt0_iff in h4'. rewrite h4' in h3. assert (occ = 0). omega. subst. simpl. pose proof (proof_irrelevance _ h2 (in_eq _ _)); subst. rewrite eq_dec_same at 1. simpl; auto. apply nat_eq_dec. rewrite count_occ_cons_neq in h3; auto. destruct h2 as [|h2]; subst. contradict h4; auto. destruct (in_dec h1 a l) as [h5 | h5]. generalize (in_l_to_inds_aux h1 (a::l) x (or_intror h2)). apply neq_sym in h4. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h4). rewrite (l_to_inds_cons_in _ _ _ h5). intro h6. simpl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. apply l_to_inds_aux_inj_dep in hl. contradiction. simpl. destruct (h1 a x) as [|h7]; subst. contradict (h4 eq_refl). specialize (ih _ h2 occ h3). erewrite lind_occ_functional3. pose proof (proof_irrelevance _ h2 (neq_in_cons l x a h4 (or_intror h2))) as h8. rewrite h8 in ih. apply ih. assumption. generalize (in_l_to_inds_aux h1 (a::l) x (or_intror h2)). rewrite l_to_inds_cons_nin; auto. apply neq_sym in h4. rewrite (l_to_inds_aux_cons_nin_in _ _ _ _ h4); auto. intro h6. pose proof h6 as h6'. inversion h6'. discriminate. rewrite in_lS_iff in H. destruct H as [m [h7 h8]]. apply S_inj in h8. subst. erewrite lind_occ_functional3. rewrite lind_occ_cons_in_neq. simpl. destruct (h1 a x) as [|h8]. subst. contradict (h4 eq_refl). unfold lS. rewrite lind_occ_map_inj at 1. apply ih. assumption. red; intros; apply S_inj; auto. auto with arith. Grab Existential Variables. rewrite in_lS_iff. exists (l_to_inds_aux h1 l x (neq_in_cons l x a h4 (or_intror h2))). split; auto. rewrite count_occ_l_to_inds; auto. left; auto. Qed. Lemma lind_occ_l_to_inds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (occ:nat) (x:T) (pf:In x l), occ < count_occ pfdt l x -> lind_occ pfdt l occ x pf = lind_occ nat_eq_dec (l_to_inds pfdt l) occ (l_to_inds_aux pfdt l x pf) (in_l_to_inds_aux _ _ _ _). intros T h1 l occ x h2 holt. rewrite lind_occ_eq_iff. split. gen0. rewrite (lind_occ_map_in_dep_inj h1 nat_eq_dec (l_to_inds_aux h1 l)). intro h3. erewrite nth_lt_functional3. rewrite lind_occ_compat; auto. apply l_to_inds_aux_inj_dep. split. erewrite firstn_S_eq. rewrite count_occ_app. simpl. rewrite nth_lt_lind_occ_l_to_inds. rewrite eq_dec_same at 1; auto. rewrite S_compat. f_equal. pose proof (firstn_S_eq _ _ (lt_lind_occ nat_eq_dec (l_to_inds h1 l) occ (l_to_inds_aux h1 l x h2) (in_l_to_inds_aux _ _ _ _))) as h3. unfold l_to_inds in h3 at 1. unfold l_to_inds in h3. rewrite firstn_map_in_dep in h3 at 1. apply count_occ_firstn_lind_occ_l_to_inds. assumption. assumption. rewrite count_occ_skipn. f_equal. erewrite firstn_S_eq. rewrite count_occ_app. simpl. rewrite nth_lt_lind_occ_l_to_inds. rewrite eq_dec_same at 1. rewrite S_compat. f_equal. apply count_occ_firstn_lind_occ_l_to_inds; auto. assumption. assumption. Grab Existential Variables. eapply lt_eq_trans. apply lt_lind_occ. rewrite length_l_to_inds; auto. eapply lt_eq_trans. apply lt_lind_occ. rewrite length_map_in_dep; auto. eapply lt_eq_trans. apply lt_lind_occ. apply length_l_to_inds. eapply lt_eq_trans. apply lt_lind_occ. apply length_l_to_inds. Qed. Lemma linds_occ_l_to_inds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), linds_occ pfdt l x = linds_occ nat_eq_dec (l_to_inds pfdt l) (l_to_inds_aux pfdt l x pf). intros T h1 l x h2. rewrite linds_occ_eq_iff. ex_goal. rewrite length_linds_occ. rewrite count_occ_l_to_inds; auto. exists hex. intros m h3. erewrite nth_lt_linds_occ. apply lind_occ_l_to_inds; auto. Qed. Lemma nth_lts_eq_cons_l_to_inds_aux : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (a:T) (a':U) (pfe:length (a::l) = length (a'::l')) (pfi:In a l) (pfi':In a' l'), (forall (m n : nat) (pfm : m < length (a :: l)) (pfn : n < length (a :: l)), let pfm' := lt_eq_trans m (length (a :: l)) (length (a' :: l')) pfm pfe in let pfn' := lt_eq_trans n (length (a :: l)) (length (a' :: l')) pfn pfe in nth_lt (a :: l) m pfm = nth_lt (a :: l) n pfn <-> nth_lt (a' :: l') m pfm' = nth_lt (a' :: l') n pfn') -> l_to_inds_aux pfdt l a pfi = l_to_inds_aux pfdu l' a' pfi'. intros T U hdt hdu l. induction l as [|a l ih]. intros; contradiction. intro l'. destruct l' as [|a' l']. intros; contradiction. intros b b' h2 h3 h4 h5. destruct h3 as [|h3], h4 as [|h4]; subst. destruct (in_dec hdt b l) as [h6 | h6], (in_dec hdu b' l') as [h7 | h7]. rewrite (l_to_inds_aux_cons_in _ _ _ h6). rewrite (l_to_inds_aux_cons_in _ _ _ h7). eapply ih. intros m n h8 h9 h10 h11. specialize (h5 (S m) (S n)). hfirst h5. simpl; auto with arith. specialize (h5 hf). hfirst h5. simpl; auto with arith. specialize (h5 hf0). simpl in h5. simpl. destruct m as [|m], n as [|n]; auto; try tauto. specialize (h5 (S (S (lind hdt _ _ h6))) 0). hfirst h5. simpl; do 2 apply lt_n_S. apply lt_lind. specialize (h5 hf). hfirst h5. simpl; auto with arith. specialize (h5 hf0). simpl in h5. erewrite nth_indep in h5. erewrite <- nth_lt_compat in h5. rewrite lind_compat in h5. hl h5; auto. rewrite h5 in hl. pose proof (@nth_In _ (lind hdt l b h6) l' b') as h8. hfirst h8. simpl in h2. do 2 apply S_inj in h2. rewrite <- h2. apply lt_lind. specialize (h8 hf1). rewrite hl in h8. contradiction. apply lt_lind. specialize (h5 (S (S (lind hdu _ _ h7))) 0). hfirst h5. simpl; do 2 apply lt_n_S. simpl in h2. pose proof h2 as h2'. do 2 apply S_inj in h2'. rewrite h2'. apply lt_lind. specialize (h5 hf). hfirst h5. simpl; auto with arith. specialize (h5 hf0). simpl in h5. symmetry in h5. erewrite nth_indep in h5. erewrite <- nth_lt_compat in h5. rewrite lind_compat in h5. hl h5; auto. rewrite h5 in hl. pose proof (@nth_In _ (lind hdu l' b' h7) l b) as h8. hfirst h8. simpl in h2. do 2 apply S_inj in h2. rewrite h2. apply lt_lind. specialize (h8 hf1). rewrite hl in h8. contradiction. apply lt_lind. rewrite l_to_inds_aux_cons_nin; auto. rewrite l_to_inds_aux_cons_nin; auto. destruct (hdu b' a') as [|h6]; subst. destruct (in_dec hdt b l) as [h6 | h6]. rewrite (l_to_inds_aux_cons_in _ _ _ h6); auto. rewrite (l_to_inds_aux_cons_in _ _ _ h4); auto. eapply ih. intros m n h7 h8 h9 h10. simpl in h5; simpl. destruct m as [|m], n as [|n]. tauto. specialize (h5 1 (S (S n))). hfirst h5. apply lt_n_S; auto with arith. specialize (h5 hf). hfirst h5. do 2 apply lt_n_S. simpl in h8; auto with arith. specialize (h5 hf0). simpl in h5. assumption. specialize (h5 (S (S m)) 1). hfirst h5. apply lt_n_S; auto with arith. specialize (h5 hf). hfirst h5. auto with arith. specialize (h5 hf0). simpl in h5. assumption. simpl in h7, h8. specialize (h5 (S (S m)) (S (S n))); auto with arith. specialize (h5 0 (S (S (lind hdu _ _ h4)))). simpl in h5. hfirst h5; auto with arith. specialize (h5 hf). hfirst h5. do 2 apply lt_n_S. simpl in h2. do 2 apply S_inj in h2. rewrite h2. apply lt_lind. hfirst h5. do 2 apply lt_n_S. simpl in h2. do 2 apply S_inj in h2. rewrite h2. apply lt_lind. hfirst h5. specialize (h5 hf1). hr h5. erewrite nth_indep. erewrite <- nth_lt_compat. rewrite lind_compat. reflexivity. apply lt_lind. rewrite <- h5 in hr. pose proof (@nth_In _ (lind hdu l' a' h4) l b) as h7. hfirst h7. simpl in h2. do 2 apply S_inj in h2. rewrite h2. apply lt_lind. do 2 apply lt_n_S. simpl in h2. do 2 apply S_inj in h2. rewrite h2. apply lt_lind. specialize (h5 hf2). hr h5. erewrite nth_indep. erewrite <- nth_lt_compat. rewrite lind_compat. reflexivity. apply lt_lind. rewrite <- h5 in hr. pose proof (@nth_In _ (lind hdu l' a' h4) l b) as h7. hfirst h7. simpl in h2. do 2 apply S_inj in h2. rewrite h2. apply lt_lind. specialize (h7 hf3). rewrite <- hr in h7. contradiction. destruct (in_dec hdt b l) as [h7 | h7]. rewrite (l_to_inds_aux_cons_in _ _ _ h7). destruct (in_dec hdu a' l') as [hi | hi]. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h6); auto. eapply ih. intros m n h9 h10 h11 h12. destruct m as [|m], n as [|n]; simpl. tauto. specialize (h5 0 1). hfirst h5. simpl; auto with arith. specialize (h5 hf). hfirst h5. simpl; auto with arith. specialize (h5 hf0). simpl in h5. hl h5; auto. rewrite h5 in hl. contradiction. specialize (h5 0 1). hfirst h5. simpl; auto with arith. specialize (h5 hf). hfirst h5. simpl; auto with arith. specialize (h5 hf0). simpl in h5. hl h5; auto. rewrite h5 in hl. contradiction. specialize (h5 (S (S m)) (S (S n))). simpl in h5. hfirst h5; auto with arith. specialize (h5 0 1). simpl in h5. hfirst h5; auto with arith. specialize (h5 hf). hfirst h5; auto with arith. specialize (h5 hf0). hl h5; auto. rewrite h5 in hl. subst; auto. contradict h6; auto. specialize (h5 0 1). simpl in h5. hfirst h5; auto with arith. specialize (h5 hf). hfirst h5; auto with arith. specialize (h5 hf0). hl h5; auto. rewrite h5 in hl; contradiction. destruct (hdt a b) as [h6 | h6]; subst. destruct (in_dec hdu b' l') as [h7 | h7]. rewrite (l_to_inds_aux_cons_in _ _ _ h3). rewrite (l_to_inds_aux_cons_in _ _ _ h7). eapply ih. intros m n h6 h8 h9. specialize (h5 (S m) (S n)). simpl in h5; simpl. auto with arith. specialize (h5 0 (S (S (lind hdt _ _ h3)))). hfirst h5; simpl; auto with arith. specialize (h5 hf). hfirst h5. simpl; do 2 apply lt_n_S. apply lt_lind. specialize (h5 hf0). simpl in h5. hl h5. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite lind_compat; auto. apply lt_lind. rewrite h5 in hl. pose proof (@nth_In _ (lind hdt l b h3) l' b') as h8. hfirst h8. simpl in h2. do 2 apply S_inj in h2. rewrite <- h2. apply lt_lind. specialize (h8 hf1). rewrite <- hl in h8; contradiction. destruct (in_dec hdt a l) as [h7 | h7], (in_dec hdu b' l') as [h8 | h8]. apply neq_sym in h6. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h6). rewrite (l_to_inds_aux_cons_in _ _ _ h8). eapply ih. intros m n h9 h10 h11 h12. simpl in h5. simpl. destruct m as [|m], n as [|n]. tauto. specialize (h5 0 (S (S n))). auto with arith. simpl in h9. pose proof h9 as h9'. apply lt_S_n in h9'. specialize (h5 (S (S m)) 0); auto with arith. specialize (h5 (S (S m)) (S (S n))); simpl in h5; auto with arith. assumption. specialize (h5 0 (S (S (lind hdt _ _ h3)))). hfirst h5; simpl; auto with arith. specialize (h5 hf). hfirst h5; simpl. simpl in h2. do 2 apply lt_n_S. apply lt_lind. specialize (h5 hf0). simpl in h5. hl h5. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite lind_compat; auto. apply lt_lind. rewrite h5 in hl. pose proof (@nth_In _ (lind hdt l b h3) l' b') as h9. hfirst h9. simpl in h2. pose proof h2 as h2'. do 2 apply S_inj in h2'. rewrite <- h2'. apply lt_lind. specialize (h9 hf1). rewrite <- hl in h9. contradiction. specialize (h5 0 1). hfirst h5. simpl; auto with arith. specialize (h5 hf). hfirst h5. simpl; auto with arith. specialize (h5 hf0). simpl in h5. apply neq_sym in h6. rewrite h5 in h6. contradict h6; auto. specialize (h5 0 1). hfirst h5. simpl; auto with arith. specialize (h5 hf). hfirst h5. simpl; auto with arith. specialize (h5 hf0). simpl in h5. apply neq_sym in h6. rewrite h5 in h6. contradict h6; auto. destruct (hdt a b) as [h7 | h7], (hdu a' b') as [h6 | h6]; subst. rewrite (l_to_inds_aux_cons_in _ _ _ h3). rewrite (l_to_inds_aux_cons_in _ _ _ h4). eapply ih. intros m n h6 h7 h8 h9. specialize (h5 (S m) (S n)). hfirst h5; simpl; auto with arith. specialize (h5 hf). hfirst h5; simpl; auto with arith. specialize (h5 hf0); auto. specialize (h5 0 1). hfirst h5; simpl; auto with arith. specialize (h5 hf). hfirst h5; simpl; auto with arith. specialize (h5 hf0). simpl in h5. hl h5; auto. rewrite h5 in hl; subst. contradict h6; auto. specialize (h5 0 1). hfirst h5; simpl; auto with arith. specialize (h5 hf). hfirst h5; simpl; auto with arith. specialize (h5 hf0). simpl in h5. hr h5; auto. rewrite <- h5 in hr; subst. contradict h7; auto. apply neq_sym in h7. apply neq_sym in h6. destruct (in_dec hdt a l) as [hi | hi], (in_dec hdu a' l') as [hi' | hi']. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h7); auto with arith. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h6); auto with arith. eapply ih. intros m n h8 h9 h10 h11. simpl; destruct m as [|m], n as [|n]. tauto. specialize (h5 0 (S (S n))). hfirst h5; simpl; auto with arith. specialize (h5 hf). hfirst h5; simpl; auto with arith. specialize (h5 hf0); simpl in h5; auto with arith. specialize (h5 (S (S m)) 0). hfirst h5; simpl; auto with arith. specialize (h5 hf). hfirst h5; simpl; auto with arith. specialize (h5 hf0); simpl in h5; auto with arith. specialize (h5 (S (S m)) (S (S n))). hfirst h5; simpl; auto with arith. specialize (h5 hf). hfirst h5; simpl; auto with arith. specialize (h5 hf0); simpl in h5; auto with arith. specialize (h5 1 (S (S (lind hdt _ _ hi)))). hfirst h5; simpl; auto with arith. specialize (h5 hf). hfirst h5; simpl. do 2 apply lt_n_S. apply lt_lind. specialize (h5 hf0). simpl in h5. hl h5. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite lind_compat; auto. apply lt_lind. rewrite h5 in hl. pose proof (@nth_In _ (lind hdt l a hi) l' b') as h8. hfirst h8. simpl in h2, hf0. do 2 apply S_inj in h2. rewrite <- h2. apply lt_lind. specialize (h8 hf1). rewrite <- hl in h8. contradiction. specialize (h5 (S (S (lind hdu _ _ hi'))) 1). hfirst h5; simpl; auto with arith. do 2 apply lt_n_S. simpl in h2. pose proof h2 as h2'. do 2 apply S_inj in h2'. rewrite h2'. apply lt_lind. specialize (h5 hf). hfirst h5; simpl; auto with arith. specialize (h5 hf0). simpl in h5. hr h5. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite lind_compat; auto. apply lt_lind. rewrite <- h5 in hr. pose proof (@nth_In _ (lind hdu l' a' hi') l b) as h8. hfirst h8. simpl in h2, hf0. do 2 apply S_inj in h2. rewrite h2. apply lt_lind. specialize (h8 hf1). rewrite hr in h8. contradiction. rewrite (l_to_inds_aux_cons_nin_in _ _ _ _ h7); auto. rewrite (l_to_inds_aux_cons_nin_in _ _ _ _ h6); auto. f_equal. eapply ih. intros m n h8 h9 h10 h11. simpl. destruct m as [|m], n as [|n]. tauto. specialize (h5 0 (S (S n))). hfirst h5. simpl; auto with arith. specialize (h5 hf). hfirst h5. simpl; auto with arith. specialize (h5 hf0); auto. specialize (h5 (S (S m)) 0). hfirst h5. simpl; auto with arith. specialize (h5 hf). hfirst h5. simpl; auto with arith. specialize (h5 hf0); auto. specialize (h5 (S (S m)) (S (S n))). hfirst h5. simpl; auto with arith. specialize (h5 hf). hfirst h5. simpl; auto with arith. specialize (h5 hf0); auto. Grab Existential Variables. simpl in h2; simpl; auto with arith. simpl in h2; simpl; auto with arith. simpl in h2; simpl; auto with arith. simpl in h2; simpl; auto with arith. simpl in h2; simpl; auto with arith. simpl in h2; simpl; auto with arith. simpl in h2; simpl; auto with arith. simpl in h2; simpl; auto with arith. Qed. Lemma nth_lts_eq_l_to_inds : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l'), (forall (m n:nat) (pfm:m < length l) (pfn:n < length l), let pfm' := lt_eq_trans _ _ _ pfm pfe in let pfn' := lt_eq_trans _ _ _ pfn pfe in nth_lt l m pfm = nth_lt l n pfn <-> nth_lt l' m pfm' = nth_lt l' n pfn') -> l_to_inds pfdt l = l_to_inds pfdu l'. intros T U hdt hdu l. induction l as [|a l ih]. intros l' h1. pose proof h1 as h1'. symmetry in h1'. apply length_eq0 in h1'; subst; auto. intro l'. destruct l' as [|a' l']. simpl. intro; discriminate. intros h1 h2. pose proof h1 as h1'. simpl in h1'. apply S_inj in h1'. destruct (in_dec hdt a l) as [h3 | h3], (in_dec hdu a' l') as [h4 | h4]. rewrite (l_to_inds_cons_in _ _ _ h3). rewrite (l_to_inds_cons_in _ _ _ h4). f_equal. eapply nth_lts_eq_cons_l_to_inds_aux. apply h2. eapply ih. intros m n h5 h6 h7 h8. specialize (h2 (S m) (S n)). hfirst h2. simpl; auto with arith. specialize (h2 hf). hfirst h2. simpl; auto with arith. specialize (h2 hf0). simpl in h2; simpl; auto. rewrite nth_lt_compat. rewrite (nth_lt_compat l n). rewrite (nth_lt_compat l'). rewrite (nth_lt_compat l' n). erewrite nth_indep. erewrite (@nth_indep _ l n). erewrite (@nth_indep _ l'). erewrite (@nth_indep _ l' n). apply h2. assumption. assumption. assumption. assumption. specialize (h2 0 (S (lind hdt _ _ h3))). hfirst h2. simpl; auto with arith. specialize (h2 hf). hfirst h2. simpl; apply lt_n_S. apply lt_lind. specialize (h2 hf0). simpl in h2. hl h2. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite lind_compat; auto. apply lt_lind. rewrite h2 in hl. pose proof (@nth_In _ (lind hdt l a h3) l' a') as h5. hfirst h5. rewrite <- h1'. apply lt_lind. specialize (h5 hf1). rewrite <- hl in h5. contradiction. specialize (h2 (S (lind hdu _ _ h4)) 0). hfirst h2. simpl; auto with arith. simpl in h1. rewrite h1. apply lt_n_S. apply lt_lind. specialize (h2 hf). hfirst h2. simpl; auto with arith. specialize (h2 hf0). simpl in h2. hr h2. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite lind_compat; auto. apply lt_lind. rewrite <- h2 in hr. pose proof (@nth_In _ (lind hdu l' a' h4) l a) as h5. hfirst h5. rewrite h1'. apply lt_lind. specialize (h5 hf1). rewrite hr in h5. contradiction. rewrite (l_to_inds_cons_nin _ _ _ h3). rewrite (l_to_inds_cons_nin _ _ _ h4). f_equal. f_equal. eapply ih. intros m n h5 h6 h7 h8. specialize (h2 (S m) (S n)). hfirst h2. simpl; auto with arith. specialize (h2 hf). hfirst h2. simpl; auto with arith. specialize (h2 hf0). simpl in h2. rewrite nth_lt_compat. rewrite (nth_lt_compat l n). rewrite (nth_lt_compat l'). rewrite (nth_lt_compat l' n). erewrite nth_indep. erewrite (@nth_indep _ l n). erewrite (@nth_indep _ l'). erewrite (@nth_indep _ l' n). apply h2. assumption. assumption. assumption. assumption. Grab Existential Variables. assumption. assumption. Qed. Fixpoint count_occl_aux {T:Type} (pfdt:type_eq_dec T) (l l':list T) : sub_list l l' -> nat := let pfdt' := list_eq_dec pfdt in match l' with | nil => fun _ => 0 | a :: la => fun pf => match pfdt' l (firstn (length l) (a::la)) with | left _ => match sub_list_dec pfdt l la with | left pfsl => S (count_occl_aux pfdt l la pfsl) | right _ => 1 end | right pfn => let pf' := neq_sub_list_cons l la a pfn pf in count_occl_aux pfdt l la pf' end end. Definition count_occl {T:Type} (pfdt:type_eq_dec T) (l l':list T) := let pfdt' := list_eq_dec pfdt in match (sub_list_dec pfdt l l') with | left pfsl => count_occl_aux pfdt l l' pfsl | right _ => 0 end. Lemma count_occl_aux_nil1 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pf:sub_list nil l), count_occl_aux pfdt nil l pf = length l. intros T h1 l. induction l as [|a l ih]; simpl; auto. intro h2. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']; auto with arith. contradict hr'. apply sub_list_nil1. contradict hr; auto. Qed. Lemma count_occl_nil1 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), count_occl pfdt nil l = length l. intros T h1 l. unfold count_occl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. revert hl. intros; apply count_occl_aux_nil1. contradict hr. apply sub_list_nil1. Qed. Lemma count_occl_aux_nil2 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pf:sub_list l nil), count_occl_aux pfdt l nil pf = 0. auto. Qed. Lemma count_occl_nil2 : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), count_occl pfdt l nil = 0. intros T h1 l. unfold count_occl. lr_match_goal; fold hlr; destruct hlr; auto. Qed. Lemma count_occl_cons2_eq : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a':T), l = firstn (length l) (a' :: l') -> count_occl pfdt l (a'::l') = S (count_occl pfdt l l'). intros T h1 l l' a' h3. unfold count_occl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. simpl. lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]. lr_match_goal; fold hlr; destruct hlr as [hl1 | hr1]. repeat f_equal; apply proof_irrelevance. contradiction. contradiction. simpl. lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]. lr_match_goal; fold hlr; destruct hlr as [hl1 | hr1]. contradiction. reflexivity. contradiction. contradict hr. rewrite h3. apply sub_list_firstn. Qed. Lemma count_occl_cons2_neq : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (a':T), l <> firstn (length l) (a' :: l') -> count_occl pfdt l (a'::l') = count_occl pfdt l l'. intros T h1 l l' a' h2. unfold count_occl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. simpl. lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]. contradiction. repeat f_equal; apply proof_irrelevance. pose proof (neq_sub_list_cons _ _ _ h2 hl). contradiction. lr_match_goal; fold hlr; destruct hlr as [hl1 | hr1]. contradict hr. apply sub_list_cons; auto. reflexivity. Qed. Lemma count_occl_sub_list_eq_le : forall {T:Type} (pfdt:type_eq_dec T) (l1 l2 l:list T), l1 <> nil -> l2 <> nil -> l1 = firstn (length l1) l2 -> count_occl pfdt l2 l <= count_occl pfdt l1 l. intros T h1 l1 l2 l. revert l1 l2. induction l as [|a l ih]; simpl. intros; do 2 rewrite count_occl_nil2; auto. intros l1 l2 hn1 hn2 he. pose proof (length_sub_list l1 l2) as h5. hfirst h5. rewrite he. apply sub_list_firstn. specialize (h5 hf). destruct (list_eq_dec h1 l2 (firstn (length l2) (a :: l))) as [h3 | h3]. rewrite count_occl_cons2_eq; auto. destruct (list_eq_dec h1 l1 (firstn (length l1) (a :: l))) as [h4 | h4]. rewrite count_occl_cons2_eq; auto with arith. apply le_lt_eq_dec in h5. destruct h5 as [h5 | h5]. rewrite he in h4 at 1. rewrite h3 in h4. rewrite <- firstn_lt_nest in h4; auto. contradict h4; auto. rewrite h5 in he. rewrite firstn_length in he. rewrite <- he in h3. contradiction. destruct (list_eq_dec h1 l1 (firstn (length l1) (a :: l))) as [h4 | h4]. rewrite count_occl_cons2_neq. rewrite count_occl_cons2_eq. eapply le_trans. apply ih. apply hn1. assumption. assumption. auto with arith. assumption. assumption. rewrite count_occl_cons2_neq; auto. rewrite count_occl_cons2_neq; auto. Qed. Lemma sub_list_count_occl : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), l <> nil -> l' <> nil -> sub_list l l' -> 0 < count_occl pfdt l l'. intros T h1 l l'. revert l. induction l' as [|a' l' ih]; simpl. intros ? ? h2. contradict h2; auto. intros l h2 h3 h4. unfold count_occl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. simpl. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. lr_match_goal; fold hlr; destruct hlr as [hl0 | hr0]; auto with arith. destruct (nil_dec' l') as [|h5]; subst. apply sub_list_sing_dec in h4. destruct h4 as [|h4]; try contradiction. subst. simpl in hr'. contradict hr'; auto. pose proof (neq_sub_list_cons _ _ _ hr' h4) as h6. specialize (ih _ h2 h5 h6). unfold count_occl in ih. lr_match ih; fold hlr in ih; destruct hlr as [hl0 | hr0]. pose proof (proof_irrelevance _ hl0 (neq_sub_list_cons l l' a' hr' hl)); subst; auto. omega. contradiction. Qed. Lemma not_sub_list_count_occl : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), ~sub_list l l' -> count_occl pfdt l l' = 0. intros T h1 l l' h2. unfold count_occl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]; auto. contradiction. Qed. Lemma count_occl_aux_same : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (pf:l <> nil) (pfsl:sub_list l l), count_occl_aux pfdt l l pfsl = 1. intros T h1 l. induction l as [|a l ih]; simpl. intro h2; contradict h2; auto. intros h2 h3. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. lr_match_goal; fold hlr; destruct hlr as [hl' | hr']. pose proof (length_sub_list _ _ hl') as h4. simpl in h4. omega. reflexivity. contradict hr. f_equal. rewrite firstn_length; auto. Qed. Lemma count_occl_same : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), l <> nil -> count_occl pfdt l l = 1. intros T h1 l h2. unfold count_occl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. apply count_occl_aux_same; auto. contradict hr. apply sub_list_refl. Qed. Fixpoint remove1 {T U:Type} (pfdt:type_eq_dec T) (l:list (T*U)) (x:T) := match l with | nil => nil | pr::l' => if (pfdt (fst pr) x) then (remove1 pfdt l' x) else pr::(remove1 pfdt l' x) end. Fixpoint remove2 {T U:Type} (pfdt:type_eq_dec U) (l:list (T*U)) (y:U) := match l with | nil => nil | pr::l' => if (pfdt (snd pr) y) then (remove2 pfdt l' y) else pr::(remove2 pfdt l' y) end. Lemma remove1_nin : forall {T U:Type} (pfdt:type_eq_dec T) (l:list (T*U)) (x:T) (y:U), ~In (x, y) (remove1 pfdt l x). intros T U h0 l. induction l as [|l pr h1]. simpl. auto. simpl. intros x y. destruct (h0 (fst l) x) as [h2 | h3]. apply h1. intro h4. destruct h4 as [h4 | h5]. contradict h3. rewrite h4. simpl. reflexivity. pose proof (h1 x y). contradiction. Qed. Lemma remove2_nin : forall {T U:Type} (pfdt:type_eq_dec U) (l:list (T*U)) (x:T) (y:U), ~In (x, y) (remove2 pfdt l y). intros T U h0 l. induction l as [|l pr h1]. simpl. auto. simpl. intros x y. destruct (h0 (snd l) y) as [h2 | h3]. apply h1. intro h4. destruct h4 as [h4 | h5]. contradict h3. rewrite h4. simpl. reflexivity. pose proof (h1 x y). contradiction. Qed. Lemma remove1_in : forall {T U:Type} (pfdt:type_eq_dec T) (l:list (T*U)) (pr:T*U) (x:T), fst pr <> x -> In pr l -> In pr (remove1 pfdt l x). intros T U h0 l. induction l as [|a l h1]. simpl. auto. simpl. intros pr x h2 h3. destruct (h0 (fst a) x) as [h4 | h5]. subst. destruct h3 as [h3l | h3r]. subst. contradict h2. reflexivity. apply h1; assumption. destruct h3 as [h3l | h3r]. subst. left. reflexivity. right. apply h1; assumption. Qed. Lemma remove2_in : forall {T U:Type} (pfdt:type_eq_dec U) (l:list (T*U)) (pr:T*U) (y:U), snd pr <> y -> In pr l -> In pr (remove2 pfdt l y). intros T U h0 l. induction l as [|a l h1]. simpl. auto. simpl. intros pr y h2 h3. destruct (h0 (snd a) y) as [h4 | h5]. subst. destruct h3 as [h3l | h3r]. subst. contradict h2. reflexivity. apply h1; assumption. destruct h3 as [h3l | h3r]. subst. left. reflexivity. right. apply h1; assumption. Qed. Lemma list_to_set_remove1_inc : forall {T U:Type} (pfdt:type_eq_dec T) (lp:list (T*U)) (x:T), Included (list_to_set (remove1 pfdt lp x)) (list_to_set lp). intros T U h0 lp. induction lp as [|pr lp h1]. simpl. auto with sets. intros x. red. simpl. intros pr' h2. destruct (h0 (fst pr) x) as [h3 | h4]. subst. specialize (h1 (fst pr)). left. auto with sets. simpl in h2. specialize (h1 x). destruct h2 as [pr'' h5 | pr'' h6]. left. auto with sets. destruct h6. right. constructor. Qed. Lemma list_to_set_remove2_inc : forall {T U:Type} (pfdt:type_eq_dec U) (lp:list (T*U)) (y:U), Included (list_to_set (remove2 pfdt lp y)) (list_to_set lp). intros T U h0 lp. induction lp as [|pr lp h1]. simpl. auto with sets. intros y. red. simpl. intros pr' h2. destruct (h0 (snd pr) y) as [h3 | h4]. subst. specialize (h1 (snd pr)). left. auto with sets. simpl in h2. specialize (h1 y). destruct h2 as [pr'' h5 | pr'' h6]. left. auto with sets. destruct h6. right. constructor. Qed. Lemma remove1_inc : forall {T U:Type} (pfdt:type_eq_dec T) (lp:list (T*U)) (x:T) (pr:T*U), In pr (remove1 pfdt lp x) -> In pr lp. intros T U h0 lp x pr h1. rewrite list_to_set_in_iff in h1. rewrite list_to_set_in_iff. pose proof (list_to_set_remove1_inc h0 lp x) as h2. auto with sets. Qed. Lemma remove2_inc : forall {T U:Type} (pfdt:type_eq_dec U) (lp:list (T*U)) (y:U) (pr:T*U), In pr (remove2 pfdt lp y) -> In pr lp. intros T U h0 lp y pr h1. rewrite list_to_set_in_iff in h1. rewrite list_to_set_in_iff. pose proof (list_to_set_remove2_inc h0 lp y) as h2. auto with sets. Qed. Lemma remove1_list_to_set_union : forall {T U:Type} (pfdt:type_eq_dec T) (lab:list (T*U)) (x:T) (A:Ensemble T) (B:Ensemble U), ~ Inn A x -> list_to_set lab = Union (cart_prod (Singleton x) B) (cart_prod A B) -> list_to_set (remove1 pfdt lab x) = cart_prod A B. intros T U hdt lab x A B h0 h2. apply Extensionality_Ensembles. red. split. (* <= *) red. intros pr h4. pose proof (list_to_set_remove1_inc hdt lab x) as h5. rewrite h2 in h5. assert (h6:Inn (Union (cart_prod (Singleton x) B) (cart_prod A B)) pr). auto with sets. destruct h6 as [pr h7 | pr h8]. destruct h7 as [h7]. destruct h7 as [h7l h7r]. rewrite surjective_pairing in h4. destruct h7l. rewrite <- list_to_set_in_iff in h4. pose proof (remove1_nin hdt lab x (snd pr)). contradiction. assumption. (* >= *) red. intros pr h3. rewrite <- list_to_set_in_iff. assert (h4:fst pr <> x). intro h4. destruct h3 as [h3]. destruct h3 as [h3l h3r]. subst. contradiction. apply remove1_in; try assumption. rewrite list_to_set_in_iff. rewrite h2. right. assumption. Qed. Lemma remove1_no_dup : forall {T U:Type} (pfdt:type_eq_dec T) (l:list (T*U)) (x:T), NoDup l -> NoDup (remove1 pfdt l x). intros T U h0 l. induction l as [|a l h1]. simpl. auto. simpl. intros x h2. pose proof (no_dup_cons _ _ h2) as h3. specialize (h1 x h3). destruct (h0 (fst a) x) as [h4 | h5]. assumption. constructor. intro h6. pose proof (remove1_inc _ _ _ _ h6). inversion h2. contradiction. assumption. Qed. Lemma remove2_no_dup : forall {T U:Type} (pfu:type_eq_dec U) (l:list (T*U)) (y:U), NoDup l -> NoDup (remove2 pfu l y). intros T U h0 l. induction l as [|a l h1]. simpl. auto. simpl. intros y h2. pose proof (no_dup_cons _ _ h2) as h3. specialize (h1 y h3). destruct (h0 (snd a) y) as [h4 | h5]. assumption. constructor. intro h6. pose proof (remove2_inc _ _ _ _ h6). inversion h2. contradiction. assumption. Qed. Lemma in_firstn_hd : forall {T:Type} (l:list T) (x:T) (n:nat), 0 < n -> In x (firstn n (x::l)). intros T l x n. revert x l. induction n as [|n h1]; simpl; auto. intros; omega. Qed. Lemma in_skipn_impl : forall {T:Type} (l:list T) (x:T) (n:nat), In x (skipn n l) -> In x l. intros T l x n. revert x l. induction n as [|n h2]; simpl; auto. intros x l. destruct l as [|a l]; auto. intro h1. specialize (h2 _ _ h1). right; auto. Qed. Lemma no_dup_firstn : forall {T:Type} (l:list T) (n:nat), NoDup l -> NoDup (firstn n l). intros T l n h1. revert h1. revert l. induction n as [|n h1]; simpl. intros. constructor. intros l h2. destruct l. constructor. pose proof (no_dup_cons_nin _ _ h2) as h6. pose proof (no_dup_cons _ _ h2) as h4. specialize (h1 _ h4). constructor. contradict h6. eapply in_firstn. apply h6. assumption. Qed. Lemma no_dup_firstn_S : forall {T:Type} (l:list T) (n:nat) (pfsn:S n <= length l), NoDup (firstn n l) -> ~In (nth_lt l n pfsn) (firstn n l) -> NoDup (firstn (S n) l). intros T l n h1 h2 h3. rewrite (firstn_S_eq _ _ h1). apply no_dup_app; auto. constructor. intro; contradiction. constructor. simpl. rewrite add_empty_sing. rewrite list_to_set_in_iff in h3. rewrite nin_iff in h3; auto. Qed. Lemma in_linds_val_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfx:In x l), forall (pf:In (linds_occ pfdt l x) (linds pfdt l)), in_linds_val pfdt l (linds_occ pfdt l x) pf = x. intros T h1 l x h2 h3. unfold in_linds_val. pose proof (nth_lt_in_linds_occ h1 l x h2 _ (in_hd' (linds_occ h1 l x) (in_linds_not_nil h1 l (linds_occ h1 l x) h3))) as h4. rewrite <- h4. apply nth_lt_functional3. Qed. Lemma nth_lt_lind_in_linds_val : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (m:nat) (pfi:In ln (linds pfdt l)) (pfm:In m ln), let pfm' := in_linds_lt pfdt _ _ _ pfm pfi in nth_lt l _ (lt_lind pfdt _ _ (in_linds_val_in _ _ _ pfi)) = nth_lt l m pfm'. intros T h1 l ln m h2 h3 h4. rewrite lind_compat. unfold in_linds_val. apply in_linds_nth_lt_eq. Qed. Lemma in_lind_in_linds_val : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (pf:In ln (linds pfdt l)), In (lind pfdt _ _ (in_linds_val_in _ _ _ pf)) ln. intros T h1 l ln h2. pose proof h2 as h2'. rewrite in_linds_iff in h2'. destruct h2' as [x [[h3 h4] h5]]; subst. rewrite in_linds_occ_iff. exists 0, h3. split. apply lt_gt. rewrite <- (count_occ_In h1 l x); auto. unfold lind. apply lind_occ_functional; auto. rewrite in_linds_val_linds_occ; auto. Qed. Lemma in_lind_occ_in_linds_val : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (pfiln:In ln (linds pfdt l)) (occ:nat) (pfo:occ < count_occ pfdt l (in_linds_val pfdt l ln pfiln)), In (lind_occ pfdt l occ (in_linds_val pfdt l ln pfiln) (in_linds_val_in pfdt l ln pfiln)) ln. intros T h1 l ln h2 occ h3. pose proof h2 as h2'. rewrite in_linds_iff in h2'. destruct h2' as [x [[h4 h5] h6]]; subst. rewrite in_linds_val_linds_occ in h3; auto. rewrite in_linds_occ_iff. exists occ, h4. split; auto. symmetry. gen0. rewrite in_linds_val_linds_occ. intro h7. apply lind_occ_functional; auto. assumption. Qed. Lemma linds_occ_in_linds_val : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat) (pfln:In ln (linds pfdt l)), linds_occ pfdt l (in_linds_val pfdt l ln pfln) = ln. intros T h1 l ln h2. pose proof h2 as h2'. rewrite in_linds_iff in h2'. destruct h2' as [x [[h3 h4] h5]]; subst. apply linds_occ_functional; auto. rewrite in_linds_val_linds_occ; auto. Qed. Lemma lind_occ_in_linds_val : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (pflt:i < length (linds pfdt l)) (ln:list nat) (pfiln:In ln (linds pfdt l)) (occ:nat) (pfo:occ < length ln), lind_occ pfdt l occ (in_linds_val pfdt l ln pfiln) (in_linds_val_in _ _ _ _) = nth_lt ln occ pfo. intros T h1 l i h2 ln h3 occ h4. pose proof h3 as h3'. rewrite in_linds_iff in h3'. destruct h3' as [x [[h5 h6] h7]]; subst. gen0. rewrite in_linds_val_linds_occ. intro h8. erewrite nth_lt_linds_occ. reflexivity. assumption. Qed. Lemma nth_lt_linds_occ_in_linds_val : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (pflt:i < length (linds pfdt l)), nth_lt (linds pfdt l) i pflt = linds_occ pfdt l (in_linds_val pfdt l (nth_lt (linds pfdt l) i pflt) (in_nth_lt (linds pfdt l) i pflt)). intros; rewrite linds_occ_in_linds_val; auto. Qed. Lemma in_lind_occ_nth_lt_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (occ:nat) (pflt:i < length (linds pfdt l)), occ < count_occ pfdt l (in_linds_val pfdt l (nth_lt (linds pfdt l) i pflt) (in_nth_lt (linds pfdt l) i pflt)) -> In (lind_occ pfdt l occ (in_linds_val pfdt l (nth_lt (linds pfdt l) i pflt) (in_nth_lt (linds pfdt l) i pflt)) (in_linds_val_in pfdt l (nth_lt (linds pfdt l) i pflt) (in_nth_lt (linds pfdt l) i pflt))) (nth_lt (linds pfdt l) i pflt). intros; apply in_lind_occ_in_linds_val; auto. Qed. Lemma in_linds_val_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfi:In x l) (ln:list nat) (pfiln:In ln (linds pfdt l)), In (lind pfdt l x pfi) ln -> in_linds_val pfdt l ln pfiln = x. intros T h1 l x h2 ln h3 h4. unfold in_linds_val. pose proof h3 as h3'. rewrite in_linds_iff in h3'. destruct h3' as [x' [[h7 h5] h6]]; subst. pose proof (nth_lt_in_linds_occ h1 _ _ h7 _ h4) as h8. erewrite nth_lt_functional3 in h8. rewrite lind_compat in h8. subst. pose proof (in_linds_val_linds_occ h1 _ _ h7 h3) as h5. rewrite <- h5. unfold in_linds_val. reflexivity. Qed. Lemma nth_lt_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (pfi:i < length (linds pfdt l)), let pfin := in_nth_lt _ _ pfi in let x := in_linds_val pfdt _ _ pfin in nth_lt (linds pfdt l) i pfi = linds_occ pfdt l x. intros; rewrite nth_lt_linds_occ_in_linds_val; auto. Qed. Lemma in_remove_ind_or : forall {T} (l:list T) (x:T) (i:nat) (pfi:i < length l), In x l -> x = nth_lt l i pfi \/ In x (remove_ind l i). intros T l; induction l as [|a l ih]; simpl. intro; contradiction. intros x i; destruct i as [|i]. intros; intuition. intro h1; apply lt_S_n in h1; intros h2; destruct h2 as [|h2]; subst. right; left; auto. specialize (ih x i h1 h2). destruct ih as [|h3]; subst. left. rewrite nth_lt_compat; apply nth_indep; auto. right; right; auto. Qed. Definition rem_nth_lt {T:Type} (pfdt:type_eq_dec T) (l:list T) (i:nat) (pfi:i < length l) : list T := remove pfdt (nth_lt _ _ pfi) l. Lemma interp_remove_nth_lind : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfn:NoDup l) (pfi:In x l), let n := lind pfdt l x pfi in l = interp (remove pfdt x l) (nth n l x) n. intros T h0 l. induction l as [|a l h1]; simpl; auto. intros; contradiction. intros x h2 h3. pose proof (no_dup_cons _ _ h2) as h4. pose proof (no_dup_cons_nin _ _ h2) as h5. destruct (h0 x a) as [h6 | h6]. subst. pose proof (lind_cons_same h0 l a h3) as h7. rewrite h7. rewrite <- (remove_not_in' _ _ _ h5). unfold interp. simpl. reflexivity. destruct h3 as [h3 | h3]. subst. contradict h6; auto. unfold interp; simpl. rewrite lind_cons_neq. specialize (h1 _ h4 h3). simpl in h1. simpl. unfold interp in h1. f_equal. rewrite h1 at 1. reflexivity. assumption. Qed. (*This one allows for duplicates.*) Definition interp_in {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) := let n := lind pfdt l x pf in let l' := firstn n l in let l'' := skipn (S n) l in l' ++ (x::nil) ++ l''. Definition interp_nin {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l) : Prop := let n := lind pfdt l x pf in let l' := firstn n l in let l'' := skipn (S n) l in ~In x l' /\ ~In x l''. Lemma no_dup_interp_nin : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), NoDup l -> interp_nin pfdt l x pf. intros T h1 l. induction l as [|c l ih]; simpl. intros; contradiction. intros x h2 h3. pose proof (no_dup_cons_nin _ _ h3) as h4. pose proof (no_dup_cons _ _ h3) as h5. red; simpl. unfold lind; simpl. destruct (h1 x c) as [h6 | h6]; subst; simpl. split; auto. specialize (ih _ (neq_in_cons _ _ _ h6 h2) h5). red in ih. split. contradict ih. destruct ih as [|ih]; subst. contradict h6; auto. intro h7. destruct h7; contradiction. destruct ih as [h11 h12]. destruct l as [|d l]; simpl; auto. Qed. Lemma interp_in_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfi:In x l), l = interp_in pfdt l x pfi. intros T h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2; simpl. unfold interp_in, lind; simpl. destruct (h1 x a) as [h3 | h3]; subst; simpl; auto. f_equal. destruct h2 as [h2 | h2]; simpl; subst. contradict h3; auto. specialize (ih _ h2). rewrite ih at 1. unfold interp_in. f_equal. unfold lind. f_equal. f_equal. apply proof_irrelevance. simpl. f_equal. destruct l as [|c l]; auto. unfold lind. f_equal. f_equal. apply proof_irrelevance. Qed. Lemma remove_interp : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (n:nat), ~In a l -> NoDup l -> remove pfdt a (interp l a n) = l. intros T h0 l a n. revert l a. induction n as [|n h1]; simpl; auto. intros l a h2 h3. destruct (h0 a a) as [h1 |h1]. rewrite <- (remove_not_in' _ _ _ h2); auto. contradict h1; auto. intro l. destruct l as [| a l]. intros x h2 h3; simpl. destruct (h0 x x) as [h4 | h4]; auto. contradict h4; auto. intros x h3 h4. simpl in h3. apply not_or_and in h3. destruct h3 as [h3 h5]. pose proof (no_dup_cons _ _ h4) as h6. specialize (h1 _ _ h5 h6). simpl. destruct (h0 x a) as [h7 | h7]. clear h1. subst. contradict h3; auto. f_equal. rewrite <- h1 at 3. reflexivity. Qed. Lemma no_dup_interp : forall {T:Type} (l:list T) (x:T) (n:nat), ~In x l -> NoDup l -> n <= length l -> NoDup (interp l x n). intros T l x n. revert x l. unfold interp. induction n as [|n h1]; simpl; auto. intros x l h1 h2 h3. constructor; auto. intros x l h2 h3 h4. destruct l as [|a l]; simpl. constructor; auto. simpl in h2. apply not_or_and in h2. destruct h2 as [h2a h2b]. pose proof (no_dup_cons _ _ h3) as h5. specialize (h1 x l h2b h5). simpl in h4. assert (h7:n <= length l). omega. specialize (h1 h7). simpl in h1. constructor; auto. intro h8. apply in_app_or in h8. pose proof (no_dup_cons_nin _ _ h3) as h9. destruct h8 as [h8 | h8]. contradict h9. eapply in_firstn. apply h8. simpl in h8. destruct h8 as [h8 | h8]; subst. contradict h2a; auto. contradict h9. eapply in_skipn_impl. apply h8. Qed. Lemma list_in_no_dup_remove_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T), ~In a l -> NoDup l -> [l':list T | NoDup l' /\ In a l' /\ remove pfdt a l' = l] = Im (seg_weak (length l)) (interp l a). intros T hdt l a hnin h0. apply Extensionality_Ensembles; red; split; red; intros l' h1. destruct h1 as [h1]. destruct h1 as [h1 [h2 h3]]. subst. pose (lind hdt l' a h2) as n. assert (h3: n <= length (remove hdt a l')). unfold n. rewrite length_remove_no_dup; auto. unfold n. pose proof (lt_lind hdt l' a h2) as h3. omega. apply Im_intro with n. constructor; auto. pose proof (nth_lt_lind_eq hdt l' a h2) as h4. simpl in h4. rewrite <- h4 at 2. pose proof (nth_lt_compat l' (lind hdt l' a h2) (lt_lind hdt l' a h2)) as h5. simpl in h5. rewrite h5. rewrite h4. unfold n. apply (interp_remove_nth_lind hdt _ _ h1 h2). inversion h1 as [n h2 ? h3]. subst. destruct h2 as [h2]. constructor. split. apply no_dup_interp; auto. split. unfold interp. simpl. rewrite in_app_iff. right. left; auto. apply remove_interp; auto. Qed. Lemma finite_list_reps : forall {T:Type} (E:Ensemble T), Finite E -> Finite (list_reps E). intros T E h1. induction h1 as [|S h2 h3 a h4]; simpl. rewrite list_reps_empty. apply Singleton_is_finite. rewrite list_reps_add; auto. apply Finite_Finite_Union; auto. intros U h5. destruct h5 as [l h5]. subst. destruct h5 as [h5]. destruct h5 as [h5 h6]. subst. assert (h7:~In a l). rewrite <- list_to_set_in_iff in h4. assumption. rewrite (list_in_no_dup_remove_eq _ _ _ h7 h6). apply finite_image. apply finite_seg_weak. apply finite_image. assumption. Qed. Lemma destruct_list_cart_prod : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (lab:list (T * U)), lab <> nil -> NoDup lab -> list_to_set lab = cart_prod A B -> exists (la:list T) (lb:list U), list_to_set la = A /\ list_to_set lb = B /\ NoDup la /\ NoDup lb /\ list_to_set lab = list_to_set (list_prod la lb). intros T U A B lab h1 h2 h3. pose proof (list_to_set_finite lab) as h4. assert (h5:list_to_set lab <> Empty_set _). intro h5. pose proof (empty_set_nil _ h5). contradiction. rewrite h3 in h5. rewrite h3 in h4. assert (h6: A <> Empty_set _). intro h6. rewrite h6 in h5. rewrite cart_prod_empty in h5. contradict h5. reflexivity. assert (h7: B <> Empty_set _). intro h7. rewrite h7 in h5. rewrite cart_prod_empty' in h5. contradict h5. reflexivity. pose proof (not_empty_Inhabited _ _ h6) as h8. pose proof (not_empty_Inhabited _ _ h7) as h9. pose proof (cart_prod_fin_rev1 _ _ h4 h9) as h10. pose proof (cart_prod_fin_rev2 _ _ h4 h8) as h11. pose proof (finite_set_list_no_dup _ h10) as h12. pose proof (finite_set_list_no_dup _ h11) as h13. destruct h12 as [la h12]. destruct h13 as [lb h13]. exists la, lb. destruct h12 as [h12l h12r]. destruct h13 as [h13l h13r]. symmetry in h12l. symmetry in h13l. repeat split; try assumption. rewrite h3. symmetry. apply list_prod_cart_prod_compat; assumption. Qed. Lemma in_list_power_add : forall {T U:Type} (la:list T) (lb:list U) (lp:list (T*U)) (a:T) (b:U), In b lb -> In lp (list_power la lb) -> In ((a, b)::lp) (list_power (a::la) lb). intros T U la lb lp a b h2 h3. simpl. rewrite in_flat_map. exists lp. split; auto. rewrite in_map_iff. exists b. split; auto. Qed. Lemma fpl_cons : forall {T U:Type} (pfdt:type_eq_dec T) (pfu:type_eq_dec U) (la:list T) (lb:list U) (lp:list (T*U)) (pr:T*U), In pr lp -> NoDup ((fst pr)::la) -> functionally_paired_l ((fst pr)::la) lb lp -> functionally_paired_l la lb (remove (impl_type_eq_dec_prod pfdt pfu) pr lp). intros T U hdt hdu la lb lp pr h0 h1 h4. constructor. inversion h4 as [h5 h6]. intros x h7. specialize (h5 x (in_cons (fst pr) x la h7)). destruct h5 as [y h5]. exists y. red. red in h5. pose proof NoDup_remove_2. rewrite <- app_nil_l in h1. pose proof (NoDup_remove_2 _ _ _ h1) as h8. rewrite app_nil_l in h8. assert (h9:(fst pr) <> x). intro; subst; contradiction. split. destruct h5 as [h5l h5r]. destruct h5l as [h5a h5b]. split; try assumption. assert (h10:(x,y) <> pr). intro h11. rewrite surjective_pairing in h11. inversion h11. subst. contradict h9. reflexivity. apply remove_a_in_eq; assumption. intros u h10. destruct h5 as [h5l h5r]. destruct h10 as [h10l h10r]. assert (h11:(x,u) <> pr). intro h12. rewrite surjective_pairing in h12. inversion h12. subst. contradict h9. reflexivity. pose proof (in_remove_neq_in_l _ _ _ _ h10r h11) as h12. apply (h5r _ (conj h10l h12)). intros pr' h5. destruct (eq_dec pr pr') as [h6 | h7]. subst. pose proof (remove_In (impl_type_eq_dec_prod hdt hdu) lp pr'). contradiction. apply neq_sym_iff in h7. pose proof (in_remove_neq_in_l _ _ _ _ h5 h7) as h8. pose proof (fpl_in_dom h4 _ h8) as h9. pose proof (fpl_in_ran h4 _ h8) as h10. split. destruct h9 as [h9l | h9r]. rewrite (surjective_pairing pr) in h0. rewrite h9l in h0. rewrite (surjective_pairing pr') in h8. pose proof (fpl_functional h4 _ _ _ h0 h8) as h11. contradict h7. rewrite (surjective_pairing pr). rewrite (surjective_pairing pr'). rewrite h9l. rewrite h11. reflexivity. assumption. assumption. Qed. Lemma fpl_cons_nin : forall {T U:Type} (la:list T) (lb:list U) (lp:list (T*U)) (a:T) (b:U), ~In a la -> In b lb -> functionally_paired_l la lb lp -> functionally_paired_l (a::la) lb ((a, b) :: lp). intros T U la lb lp a b h1 h2 h3. constructor. intros x h4. destruct h4 as [h4l | h4r]. subst. exists b. red. split. split; auto. left. reflexivity. intros b' h4. destruct h4 as [h4l h4r]. destruct h4r as [h5 | h6]. inversion h5; subst. reflexivity. destruct h3 as [h3l h3r]. specialize (h3r _ h6). simpl in h3r. destruct h3r; contradiction. inversion h3 as [h3l h3r]. specialize (h3l _ h4r). destruct h3l as [y h3l]. red in h3l. exists y. red. destruct h3l as [h5 h6]. destruct h5 as [h5l h5r]. split. split; auto. right. auto. intros y' h7. destruct h7 as [h7l h7r]. destruct h7r as [h8 | h9]. inversion h8. subst. contradiction. specialize (h3r _ h9). specialize (h6 _ (conj h7l h9)). assumption. intros pr h4. destruct h4 as [h4l | h4r]. rewrite (surjective_pairing pr) in h4l. inversion h4l. subst. split. left. reflexivity. assumption. destruct h3 as [h3l h3r]. specialize (h3r _ h4r). destruct h3r as [h5 h6]. split. right. auto. assumption. Qed. Inductive synced {T U:Type} (la:list T) (lp:list (T*U)) := | synced_intro : map (fun pr => (fst pr)) lp = la -> synced la lp. Lemma synced_hd : forall {T U:Type} (def:T*U) (la:list T) (lp:list (T*U)), synced la lp -> fst (hd def lp) = hd (fst def) la. intros T U def la lp. induction la as [|a la h1]; induction lp as [|pr lp h2]. simpl. auto. intro h3. destruct h3 as [h3]. simpl in h3. discriminate. simpl. intro h2. destruct h2 as [h2]. simpl in h2. discriminate. intro h3. simpl. destruct h3 as [h3]. simpl in h3. inversion h3. reflexivity. Qed. Lemma synced_nil1 : forall {T U:Type} (lp:list (T*U)), synced nil lp -> lp = nil. intros T U lp. induction lp as [|pr lp h1]. auto. intro h2. destruct h2 as [h2]. simpl in h2. discriminate. Qed. Lemma synced_nil2 : forall {T U:Type} (la:list T), synced la (@nil (T*U)) -> la = nil. intros T U la. induction la as [|a la h1]. auto. intro h2. destruct h2 as [h2]. simpl in h2. discriminate. Qed. Lemma synced_cons : forall {T U:Type} (la:list T) (lp:list (T*U)) (a:T) (pr:(T*U)), synced (a::la) (pr::lp) -> synced la lp. intros T U la lp a pr h1. constructor. destruct h1 as [h1]. simpl in h1. pose proof (cons_inj _ _ _ _ h1) as h2. destruct h2; assumption. Qed. Lemma synced_cons' : forall {T U:Type} (la:list T) (lp:list (T*U)) (a:T), synced (a::la) lp -> synced la (tl lp). intros T U la lp. revert la. induction lp as [|pr lp h1]. simpl. intros la a h1. destruct h1 as [h1]. simpl in h1. discriminate. intros la a h2. simpl. apply (synced_cons _ _ a pr). assumption. Qed. Lemma synced_cons'' : forall {T U:Type} (la:list T) (lp:list (T*U)) (a:T) (b:U), synced la lp -> synced (a::la) ((a, b) :: lp). intros T U la lp a b h1. constructor. destruct h1. simpl. f_equal. assumption. Qed. Lemma remove_synced_cons : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (lp:list (T*U)) (pr:(T*U)), synced ((fst pr)::la) lp -> functionally_paired_l ((fst pr)::la) lb lp -> In pr lp -> NoDup lp -> remove (impl_type_eq_dec_prod pfdt pfdu) pr lp = tl lp. intros T U hdt hdu la lb lp. revert la. induction lp as [|pr' lp h1]. simpl. auto. intros la pr h2 h2' h3 h4. simpl. destruct (impl_type_eq_dec_prod hdt hdu pr pr') as [h5 | h6]. subst. pose proof (no_dup_cons_nin _ _ h4) as h5. symmetry. apply remove_not_in'. assumption. destruct h3 as [h7 | h8]. symmetry in h7. contradiction. pose proof (synced_hd pr _ _ h2) as h7. simpl in h7. pose proof (in_cons pr' pr lp h8) as h9. pose proof (in_eq pr' lp) as h10. rewrite (surjective_pairing pr) in h9. rewrite (surjective_pairing pr') in h10 at 1. rewrite h7 in h10. pose proof (fpl_functional h2' _ _ _ h9 h10) as h11. assert (h12:pr = pr'). rewrite (surjective_pairing pr). rewrite (surjective_pairing pr'). rewrite h7, h11. reflexivity. contradiction. Qed. Lemma fpl_cons_synced : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (lp:list (T*U)) (a:T), synced (a::la) lp -> NoDup (a::la) -> NoDup lp -> functionally_paired_l (a::la) lb lp -> functionally_paired_l la lb (tl lp). intros T U hdt hdu la lb lp a h1 h2 h2' h3. inversion h3 as [h4 h5]. specialize (h4 a (in_eq a la)). destruct h4 as [b h6]. red in h6. destruct h6 as [h6 h7]. destruct h6 as [h6a h6b]. pose proof (fpl_cons hdt hdu _ _ _(a, b) h6b h2 h3) as h8. rewrite (remove_synced_cons hdt hdu la lb) in h8. assumption. simpl. assumption. simpl. assumption. assumption. assumption. Qed. Lemma in_tl_list_power_synced : forall {T U:Type} (la:list T) (lb:list U) (lp:list (T*U)) (a:T), In (tl lp) (list_power la lb) -> synced (a :: la) lp -> functionally_paired_l (a :: la) lb lp -> In lp (list_power (a::la) lb). intros T U la lb lp a h1 h2 h2'. revert h1 h2 h2'. revert a lb la. induction lp as [|pr lp h3]. simpl. intros a lb la h1 h2. destruct h2 as [h2]. simpl in h2. discriminate. simpl. intros a lb la h1 h2 h4. rewrite in_flat_map. exists lp. split. assumption. rewrite in_map_iff. exists (snd pr). destruct h2 as [h2]. simpl in h2. pose proof (cons_inj _ _ _ _ h2) as h5. destruct h5 as [h5l h5r]. rewrite <- h5l. rewrite <- surjective_pairing. split. reflexivity. destruct h4 as [h4 h5]. specialize (h5 _ (in_eq pr lp)). destruct h5; assumption. Qed. Lemma in_list_power_synced_tl: forall {T U:Type} (la:list T) (lb:list U) (lp:list (T*U)) (a:T), synced (a :: la) lp -> functionally_paired_l (a :: la) lb lp -> In lp (list_power (a::la) lb) -> In (tl lp) (list_power la lb). intros T U la lb lp a h1 h2 h2'. revert h1 h2 h2'. revert a lb la. destruct lp as [|pr lp]. simpl. intros a lb la h1. destruct h1 as [h1]. simpl in h1. discriminate. simpl. intros a lb la h1 h2 h4. rewrite in_flat_map in h4. destruct h4 as [lp' h4]. destruct h4 as [h4 h5]. rewrite in_map_iff in h5. destruct h5 as [b h6]. destruct h6 as [h6l h6r]. inversion h6l. subst. assumption. Qed. Lemma in_tl_list_power_synced_iff : forall {T U:Type} (la:list T) (lb:list U) (lp:list (T*U)) (a:T), synced (a :: la) lp -> functionally_paired_l (a :: la) lb lp -> (In (tl lp) (list_power la lb) <-> In lp (list_power (a::la) lb)). intros T U la lb lp a h1 h2. split. intro; apply in_tl_list_power_synced; assumption. intro; apply in_list_power_synced_tl with a; assumption. Qed. Lemma synced_no_dup : forall {T U:Type} (l:list T) (lp:list (T*U)), NoDup l -> synced l lp -> NoDup lp. intros T U l lp h1 h2. revert lp h2. induction h1 as [|a l h3 h4 h5]. intros lp h1. pose proof (synced_nil1 _ h1). subst. constructor. intros lp. induction lp as [|pr lp h2]. intros; constructor. intro h6. pose proof (synced_cons' _ _ _ h6) as h7. simpl in h7. specialize (h5 _ h7). constructor. pose proof (synced_hd pr _ _ h6) as h8. simpl in h8. intro h9. inversion h7 as [h10]. pose proof (in_map (fun pr0:T*U => fst pr0) lp pr h9) as h11. rewrite h10 in h11. rewrite h8 in h11. contradiction. assumption. Qed. Lemma in_list_power_synced : forall {T U:Type} (la:list T) (lb:list U) (lp:list (T*U)), In lp (list_power la lb) -> synced la lp. intros T U la. induction la as [|a la h1]. intros lb lp h1. simpl in h1. destruct h1; try contradiction. subst. constructor. simpl. reflexivity. intros lb lp h2. simpl in h2. rewrite in_flat_map in h2. destruct h2 as [lp' h2]. destruct h2 as [h2l h2r]. constructor. rewrite in_map_iff in h2r. destruct h2r as [b h2r]. destruct h2r as [h2a h2b]. specialize (h1 _ _ h2l). destruct h1 as [h1]. rewrite <- h2a. simpl. rewrite h1. reflexivity. Qed. Lemma synced_length : forall {T U:Type} (la:list T) (lp:list (T*U)), synced la lp -> length la = length lp. intros T U la lp h1. destruct h1 as [h1]. subst. apply map_length. Qed. Lemma fpl_in_list_power : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (lp:list (T*U)), synced la lp -> NoDup la -> NoDup lp -> (functionally_paired_l la lb lp <-> In lp (list_power la lb)). intros T U hdt hdu la lb lp h0 h1 h3. split. (* -> *) revert h0 h1 h3. revert lb lp. induction la as [|a la h1]. intros lb lp h0 h1 h3. simpl. left. symmetry. apply fpl_empty1_s with lb. assumption. intros lb lp h0 h3 h4 h5. inversion h5 as [h6 h7]. specialize (h6 a (in_eq a la)). destruct h6 as [b h6]. red in h6. destruct h6 as [h6a h6b]. destruct h6a as [h8 h9]. pose proof (fpl_cons_synced hdt hdu _ _ _ _ h0 h3 h4 h5) as h10. pose proof (synced_cons' _ _ _ h0) as h11. pose proof (no_dup_cons _ _ h3) as h12. pose proof (no_dup_tl _ h4) as h13. specialize (h1 _ _ h11 h12 h13 h10). apply in_tl_list_power_synced; assumption. (* <- *) intro h4. revert lb lp h0 h1 h3 h4. induction la as [|a la h1]. (* nil *) simpl. intros lb lp h1 h2 h3 h4. destruct h4 as [h4 | h5]. subst. apply fpl_empty1. contradiction. (* cons *) intros lb lp h2 h3 h4 h5. simpl in h5. rewrite in_flat_map in h5. destruct h5 as [lp' h6]. destruct h6 as [h6l h6r]. rewrite in_map_iff in h6r. destruct h6r as [b h7]. destruct h7 as [h7a h7b]. constructor. intros x h8. destruct h8 as [h8 | h9]. subst. exists b. red. split. split; try constructor; auto. intros b' h8. destruct h8 as [h8l h8r]. pose proof (no_dup_cons_nin _ _ h3) as h9. destruct h8r as [h10 | h11]. inversion h10. reflexivity. pose proof (in_list_power1 _ h6l _ h11) as h12. simpl in h12. contradiction. rewrite <- h7a in h2. pose proof (synced_cons _ _ _ _ h2) as h10. rewrite <- h7a in h4. pose proof (no_dup_cons _ _ h3) as h11. pose proof (no_dup_cons _ _ h4) as h12. specialize (h1 _ _ h10 h11 h12 h6l). destruct h1 as [h13 h14]. specialize (h13 _ h9). destruct h13 as [y h13]. exists y. red. red in h13. destruct h13 as [h13a h13b]. split. split. destruct h13a as [h13aa h13ab]. assumption. destruct h13a as [h13aa h13ab]. pose proof (in_cons (a, b) (x, y) lp' h13ab) as h15. rewrite h7a in h15. assumption. intros b' h15. destruct h15 as [h15l h15r]. rewrite <- h7a in h15r. destruct h15r as [h16 | h17]. inversion h16. subst. pose proof (no_dup_cons_nin _ _ h3). contradiction. apply h13b. split; assumption. intros pr h8. split. rewrite <- h7a in h8. destruct h8 as [h8 | h9]. left. rewrite surjective_pairing in h8. inversion h8. reflexivity. pose proof (in_list_power1 _ h6l _ h9). right. assumption. rewrite <- h7a in h8. destruct h8 as [h8l | h8r]. rewrite surjective_pairing in h8l. inversion h8l. subst. assumption. pose proof (in_list_power2 _ h6l _ h8r). assumption. Qed. Lemma in_list_power_fpl : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu: type_eq_dec U) (la:list T) (lb:list U) (lp:list (T*U)), NoDup la -> NoDup lp -> In lp (list_power la lb) -> functionally_paired_l la lb lp. intros T U hdt hdu la lb lp h1 h2 h3. rewrite fpl_in_list_power; auto. apply in_list_power_synced with lb; assumption. Qed. Lemma in_list_of_lists_seqs_map_map_pair : forall {T U V:Type} (p:T->U->V) (li:list T) (lj:list U) (l:list V), In l (list_of_lists_seqs (map (fun i : T => map (p i) lj) li)) -> exists lp:list (T*U), In lp (list_of_lists_seqs (map (fun i:T => map (pair i) lj) li)) /\ l = map (fun pr:(T*U) => p (fst pr) (snd pr)) lp. intros T U V p li lj. induction li as [|i li h1]; induction lj as [|j lj h2]. simpl. intros l h3. exists nil. split. left. reflexivity. simpl. destruct h3 as [h3l | h3r]. subst. reflexivity. contradiction. simpl. intros l h3. exists nil. split. left. reflexivity. simpl. destruct h3; auto. contradiction. simpl. intros; contradiction. intros l h5. simpl in h5. rewrite in_map_iff in h5. destruct h5 as [vl h5]. destruct h5 as [h5l h5r]. destruct (in_app_or _ _ _ h5r) as [h6 | h7]. clear h5r. rewrite in_map_iff in h6. destruct h6 as [lv h6]. destruct h6 as [h6l h6r]. specialize (h1 _ h6r). destruct h1 as [lp h1]. exists ((i, j) :: lp). simpl. rewrite in_map_iff. split. exists (i, j, lp). simpl. split. reflexivity. apply in_or_app. left. rewrite in_map_iff. exists lp. split. reflexivity. destruct h1 as [h1l h1r]. assumption. destruct h1 as [h1l h1r]. subst. simpl. reflexivity. rewrite (surjective_pairing vl) in h7. rewrite in_prod_iff in h7. destruct h7 as [h7l h7r]. specialize (h1 _ h7r). destruct h1 as [lp h1]. rewrite in_map_iff in h7l. destruct h7l as [u h8]. destruct h8 as [h8l h8r]. exists ((i, u)::lp). simpl. rewrite in_map_iff. split. exists ((i, u), lp). simpl. split. reflexivity. apply in_or_app. right. rewrite in_prod_iff. split. rewrite in_map_iff. exists u. split. reflexivity. assumption. destruct h1 as [h1l h1r]. assumption. rewrite <- h5l. f_equal. rewrite h8l. reflexivity. destruct h1 as [h1l h1r]. assumption. Qed. Lemma in_list_of_lists_seqs_map_map_pair_in_list_power : forall {T U V:Type} (p:T->U->V) (li:list T) (lj:list U) (lp:list (T*U)), In lp (list_of_lists_seqs (map (fun i:T => map (pair i) lj) li)) -> In lp (list_power li lj). intros T U V p li lj. induction li as [|i li h1]; induction lj as [|j lj h2]. simpl. auto. simpl. auto. simpl. intros; contradiction. simpl. intros lp h3. rewrite in_flat_map. rewrite in_map_iff in h3. destruct h3 as [tr h4]. destruct h4 as [h4l h4r]. destruct (in_app_or _ _ _ h4r) as [h5 | h6]. clear h4r. rewrite in_map_iff in h5. destruct h5 as [lp' h6]. destruct h6 as [h6l h6r]. specialize (h1 _ h6r). exists lp'. split; auto. simpl. subst. simpl. left. reflexivity. subst. rewrite (surjective_pairing tr) in h6. pose proof (in_prod_iff). rewrite in_prod_iff in h6. destruct h6 as [h6l h6r]. specialize (h1 _ h6r). exists (snd tr). split; auto. simpl. right. rewrite in_map_iff. rewrite in_map_iff in h6l. destruct h6l as [u h6l]. destruct h6l as [h6a h6b]. exists u. rewrite h6a. split; auto. Qed. Lemma map_ext_in : forall {T U:Type} (f g:T->U) (l:list T), (forall x:T, In x l -> f x = g x) -> map f l = map g l. intros T U f g l. induction l as [|a l h1]. simpl. auto. simpl. intro h2. pose proof (h2 a) as h3. rewrite h3. f_equal. apply h1. intros x h4. apply h2. right. assumption. left. reflexivity. Qed. Lemma fpl_impl_map : forall {T U:Type} (li:list T) (lj:list U) (pfi:NoDup li) (lp:list (T*U)) (pffp:functionally_paired_l li lj lp), synced li lp -> let f := fpl_to_fun_in_dep pffp in lp = map_in_dep (fun (i:T) (pfi:In i li) => (i, f i pfi)). intros T U li lj h1 lp h2 h5. simpl. destruct h5 as [h5]. subst. rewrite map_in_dep_map. rewrite <- map_in_dep_id at 1. apply map_in_dep_ext. intros pr h3. apply functional_extensionality. intro h4. rewrite id_in_dep_compat. rewrite (surjective_pairing pr) at 1. f_equal. symmetry. erewrite (fpl_to_fun_in_dep_in_eq _ _ _ _ _ h3). f_equal. apply proof_irrelevance. Qed. Lemma in_list_of_lists_seqs_map_map : forall {T U V:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (p:T->U->V) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (l:list V), In l (list_of_lists_seqs (map (fun i : T => map (p i) lj) li)) -> exists (lp:list (T*U)) (pfin:In lp (list_power li lj)), let pffp := in_list_power_fpl pfdt pfdu li lj lp pfi (list_power_no_dup li lj lp pfi pfj pfin) pfin in let f := fpl_to_fun_in_dep pffp in l = map_in_dep (fun (i:T) (pfi:In i li) => p i (f i pfi)). intros T U V h1 h2 p li lj h3 h4 l h5. simpl. pose proof (in_list_of_lists_seqs_map_map_pair p li lj l h5) as h7. destruct h7 as [lp h7]. destruct h7 as [h7l h7r]. subst. pose proof (in_list_of_lists_seqs_map_map_pair_in_list_power p li lj lp h7l) as h8. exists lp, h8. pose proof (in_list_power_synced _ _ _ h8) as h9. pose proof (list_power_no_dup _ _ _ h3 h4 h8) as h10. pose proof (in_list_power_fpl h1 h2 _ _ _ h3 h10 h8) as h11. pose proof (fpl_impl_map _ _ h3 _ h11 h9) as h12. simpl in h12. rewrite h12 at 1. rewrite map_map_in_dep. apply map_in_dep_ext. intros; apply functional_extensionality; intro. f_equal; simpl; f_equal; apply proof_irrelevance. Qed. Lemma in_list_of_lists_seqs_map_map_pair_rev : forall {T U V:Type} (p:T->U->V) (li:list T) (lj:list U) (lp:list (T*U)), In lp (list_of_lists_seqs (map (fun i:T => map (pair i) lj) li)) -> exists l:list V, In l (list_of_lists_seqs (map (fun i : T => map (p i) lj) li)) /\ l = map (fun pr:(T*U) => p (fst pr) (snd pr)) lp. intros T U V p li lj. induction li as [|i li h1]; induction lj as [|j lj h2]. simpl. intros lp h1. destruct h1 as [h1l | h1r]. subst. simpl. exists nil. split; auto. contradiction. simpl. intros lp h3. destruct h3 as [h3 | h4]. subst. simpl. exists nil. split; auto. contradiction. simpl. intros; contradiction. intros lp h3. simpl. simpl in h3. rewrite in_map_iff in h3. destruct h3 as [tr h3]. destruct h3 as [h3l h3r]. destruct (in_app_or _ _ _ h3r) as [h4 | h5]. clear h3r. rewrite in_map_iff in h4. destruct h4 as [lp' h4]. destruct h4 as [h4l h4r]. rewrite surjective_pairing in h4l. inversion h4l. specialize (h1 _ h4r). destruct h1 as [l h1]. exists ((p i j) :: l). split. rewrite in_map_iff. exists (p i j, l). simpl. split; auto. apply (in_or_app). left. rewrite in_map_iff. exists l. split; auto. destruct h1 as [h1l h1r]. simpl in h1l. assumption. rewrite <- h3l. simpl. f_equal. f_equal. inversion H0. auto. inversion H0. auto. destruct h1 as [h1l h1r]. rewrite H1 in h1r. assumption. rewrite (surjective_pairing tr) in h5. rewrite in_prod_iff in h5. destruct h5 as [h5l h5r]. rewrite in_map_iff in h5l. destruct h5l as [u h5l]. destruct h5l as [h5a h5b]. specialize (h1 _ h5r). destruct h1 as [l h1]. destruct h1 as [h1l h1r]. exists ((p i u) :: l). split. rewrite in_map_iff. exists (p i u, l). simpl. split; auto. apply in_or_app. right. rewrite in_prod_iff. split. rewrite in_map_iff. exists u. split; auto. assumption. rewrite <- h3l. simpl. f_equal. f_equal. inversion h5a. simpl. auto. inversion h5a. simpl. auto. assumption. Qed. Lemma in_list_power_in_list_of_lists_seqs_map_map_pair : forall {T U:Type} (li:list T) (lj:list U) (lp:list (T*U)), In lp (list_power li lj) -> In lp (list_of_lists_seqs (map (fun i:T => map (pair i) lj) li)). intros T U li lj. induction li as [|i li h1]; induction lj as [|j lj h2]. simpl. auto. simpl. auto. simpl. simpl in h1. intros lp h2. rewrite in_flat_map in h2. destruct h2 as [? h2]. destruct h2; contradiction. simpl. intros lp h3. rewrite in_map_iff. rewrite in_flat_map in h3. destruct h3 as [lp' h3]. destruct h3 as [h3l h3r]. simpl in h3r. destruct h3r as [h3a | h3b]. exists (i, j, lp'). simpl. split; auto. apply in_or_app. left. rewrite in_map_iff. exists lp'. split; auto. rewrite in_map_iff in h3b. destruct h3b as [u h3b]. destruct h3b as [h4 h5]. exists (i, u, lp'). simpl. split; auto. apply in_or_app. right. rewrite in_prod_iff. split. rewrite in_map_iff. exists u. split; auto. auto. Qed. Lemma in_map_list_of_lists_seqs_map_in_dep : forall {T U V:Type} (pft:type_eq_dec T) (pfu:type_eq_dec U) (p:T->U->V) (li:list T) (lj:list U) (ls:list (T*U)) (pfi:NoDup li) (pfj:NoDup lj) (pfin:In ls (list_power li lj)), let pffp := in_list_power_fpl pft pfu li lj ls pfi (list_power_no_dup li lj ls pfi pfj pfin) pfin in let f := fpl_to_fun_in_dep pffp in exists l: list V, In l (list_of_lists_seqs (map (fun i : T => map (p i) lj) li)) /\ l = map_in_dep (fun (i : T) (pfi:In i li) => p i (f i pfi)). intros T U V hdt hdu p li lj ls h1 h2 h3. simpl. pose proof (in_list_power_in_list_of_lists_seqs_map_map_pair _ _ _ h3) as h4. pose proof (in_list_of_lists_seqs_map_map_pair_rev p _ _ _ h4) as h5. destruct h5 as [l h5]. destruct h5 as [h5l h5r]. exists l. subst. split; auto. pose proof (in_list_power_synced _ _ _ h3) as h6. pose proof (list_power_no_dup _ _ _ h1 h2 h3) as h7. pose proof (in_list_power_fpl hdt hdu _ _ _ h1 h7 h3 ) as h8. pose proof (fpl_impl_map _ _ h1 _ h8 h6) as h9. rewrite h9 at 1. rewrite map_map_in_dep. apply map_in_dep_ext. simpl. intros x h10. apply functional_extensionality. intro. f_equal. f_equal. apply proof_irrelevance. Qed. Lemma in_list_of_lists_seqs_map_map_pair_in_list_power' : forall {T U:Type} (li:list T) (lj:list U) (lp:list (U*T)), In lp (list_of_lists_seqs (map (fun j : U => map (fun i : T => (j, i)) li) lj)) -> In lp (list_power lj li). intros T U li lj. induction lj as [|j lj h1]; simpl; auto. intros lp h2. rewrite in_map_iff in h2. destruct h2 as [tr h2]. destruct h2 as [h2a h2b]. rewrite (surjective_pairing tr) in h2b. rewrite in_prod_iff in h2b. destruct h2b as [h2l h2r]. specialize (h1 _ h2r). rewrite in_flat_map. exists (snd tr). split; auto. rewrite in_map_iff. rewrite in_map_iff in h2l. destruct h2l as [x h2l]. destruct h2l as [h3 h4]. exists x. split. rewrite <- h2a. f_equal. assumption. assumption. Qed. Lemma in_list_power_in_list_of_lists_seqs_map_map_pair' : forall {T U:Type} (li:list T) (lj:list U) (lp:list (U*T)), In lp (list_power lj li) -> In lp (list_of_lists_seqs (map (fun j : U => map (fun i : T => (j, i)) li) lj)). intros T U li lj. induction lj as [|j lj h1]; simpl; auto. intros lp h2. rewrite in_flat_map in h2. destruct h2 as [lp' h2]. destruct h2 as [h2l h2r]. rewrite in_map_iff. specialize (h1 _ h2l). rewrite in_map_iff in h2r. destruct h2r as [i h2r]. destruct h2r as [h2a h2b]. exists (j,i, lp'). simpl. split. assumption. rewrite in_prod_iff. split. rewrite in_map_iff. exists i. split; auto. assumption. Qed. Lemma in_list_of_lists_seqs_map_map_pair' : forall {T U V:Type} (p:T->U->V) (li:list T) (lj:list U) (l:list V), In l (list_of_lists_seqs (map (fun j : U => map (fun i : T => p i j) li) lj)) -> exists lp:list (U*T), In lp (list_of_lists_seqs (map (fun j : U => map (fun i : T => (j, i)) li) lj)) /\ l = map (fun pr:(U*T) => p (snd pr) (fst pr)) lp. intros T U V p li lj. revert li. induction lj as [|j lj h1]. simpl. intros li l h2. exists nil. simpl. split. left; auto. destruct h2; [subst; auto | contradiction]. intros li l h5. simpl in h5. rewrite in_map_iff in h5. destruct h5 as [pr h5]. destruct h5 as [h5l h5r]. rewrite (surjective_pairing pr) in h5r. rewrite in_prod_iff in h5r. destruct h5r as [h5a h5b]. specialize (h1 _ _ h5b). destruct h1 as [lp h1]. destruct h1 as [h1a h1b]. pose proof (in_list_of_lists_seqs_map_map_pair_in_list_power' _ _ _ h1a) as h7. rewrite in_map_iff in h5a. destruct h5a as [i h5a]. destruct h5a as [h9 h10]. pose proof (in_list_power_add _ _ _ j i h10 h7) as h11. pose proof (in_list_power_in_list_of_lists_seqs_map_map_pair' _ _ _ h11) as h12. exists ((j, i)::lp). split; auto. simpl. subst. f_equal. rewrite h9. reflexivity. assumption. Qed. Lemma in_list_of_lists_seqs_map_map' : forall {T U V:Type} (hdt:type_eq_dec T) (hdu:type_eq_dec U) (p:T->U->V) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (l:list V) (def:T), In l (list_of_lists_seqs (map (fun j : U => map (fun i : T => p i j) li) lj)) -> exists (lp:list (U*T)) (pfin:In lp (list_power lj li)), let pffp := in_list_power_fpl hdu hdt lj li lp pfj (list_power_no_dup lj li lp pfj pfi pfin) pfin in let f := fpl_to_fun_in_dep pffp in l = map_in_dep (fun (j:U) (pfinj:In j lj) => p (f j pfinj) j). intros T U V hdt hdu p li lj h1 h2 l def h3. pose proof (in_list_of_lists_seqs_map_map_pair _ _ _ _ h3) as h4. destruct h4 as [lp h4]. destruct h4 as [h4l h4r]. pose proof (in_list_of_lists_seqs_map_map_pair_in_list_power' li lj lp h4l) as h5. exists lp, h5. pose proof (in_list_power_synced _ _ _ h5) as h6. pose proof (list_power_no_dup _ _ _ h2 h1 h5) as h7. pose proof (in_list_power_fpl hdu hdt _ _ _ h2 h7 h5) as h8. pose proof (fpl_impl_map _ _ h2 _ h8 h6) as h9. rewrite h4r. simpl. simpl in h9. rewrite h9 at 1. rewrite map_map_in_dep. apply map_in_dep_ext. intros. apply functional_extensionality. intro. f_equal; f_equal; simpl. f_equal; apply proof_irrelevance. Qed. Lemma in_list_of_lists_seqs_map_map_pair_rev' : forall {T U V:Type} (p:T->U->V) (li:list T) (lj:list U) (lp:list (U*T)), In lp (list_of_lists_seqs (map (fun j : U => map (fun i : T => (j, i)) li) lj)) -> exists l:list V, In l (list_of_lists_seqs (map (fun j : U => map (fun i : T => p i j) li) lj)) /\ l = map (fun pr:(U*T) => p (snd pr) (fst pr)) lp. intros T U V p li lj. induction lj as [|j lj h1]; destruct li as [|i li]; simpl. intros lp h1. exists nil. split. left. reflexivity. destruct h1 as [h1l | h1r]. subst. simpl. reflexivity. contradiction. intros lp h1. exists nil. split. left. reflexivity. destruct h1 as [h1l | h1r]. subst. simpl. reflexivity. contradiction. intros; contradiction. intros lp h2. rewrite in_map_iff in h2. destruct h2 as [tr [h2 h3]]. destruct (in_app_or _ _ _ h3) as [h4 | h5]. clear h3. rewrite in_map_iff in h4. destruct h4 as [lp' [h4 h5]]. specialize (h1 _ h5). destruct h1 as [l h1]. destruct h1 as [h1l h1r]. exists ((p i j)::l). rewrite in_map_iff. split. exists ((p i j), l). simpl. split. reflexivity. apply in_or_app. left. rewrite in_map_iff. exists l. split; auto. rewrite <- h2. simpl. rewrite <- h4. simpl. f_equal. assumption. rewrite (surjective_pairing tr) in h5. clear h3. rewrite in_prod_iff in h5. destruct h5 as [h5l h5r]. specialize (h1 _ h5r). destruct h1 as [l [h1l h1r]]. rewrite in_map_iff in h5l. destruct h5l as [a h5l]. destruct h5l as [h5a h5b]. exists ((p a j)::l). split. rewrite in_map_iff. exists ((p a j), l). split. simpl. reflexivity. apply in_or_app. right. rewrite in_prod_iff. split. rewrite in_map_iff. exists a. split; auto. assumption. rewrite <- h2. simpl. rewrite <- h5a. simpl. f_equal. assumption. Qed. Lemma map_in_dep_fpl_to_fun_in_dep : forall {T U:Type} {la:list T} {lb:list U} {lp:list (T*U)} (pf:functionally_paired_l la lb lp), synced la lp -> let f := fpl_to_fun_in_dep pf in map_in_dep f = map (snd (B:=U)) lp. intros T U la lb lp h1 h2. simpl. destruct h2 as [h2]. subst. rewrite map_in_dep_map. rewrite map_eq_map_in_dep. apply map_in_dep_ext. intros pl h3. apply functional_extensionality. intro. rewrite (fpl_to_fun_in_dep_in_eq _ _ _ h1 _ h3) at 1; auto. f_equal. apply proof_irrelevance. Qed. Lemma map_in_dep_fpl_to_fun_cons : forall {T U:Type} {la:list T} {lb:list U} {lp:list (T*U)} {a:T} (pf:functionally_paired_l la lb (tl lp)) (pf':functionally_paired_l (a::la) lb lp), NoDup lp -> let f := fpl_to_fun_in_dep pf in let f' := fpl_to_fun_in_dep pf' in map_in_dep f = map_in_dep (fun_from_cons f'). intros T U la lb lp a h1 h2 h3. simpl. apply map_in_dep_ext. intros x h4. apply functional_extensionality. intro h5. rewrite fun_from_cons_compat; auto. pose proof (fpl_to_fun_in_dep_compat h1 x h5) as h6. pose proof (fpl_to_fun_in_dep_compat h2 x (in_cons a x la h5)) as h7. pose proof (in_tl _ _ h6) as h9. apply (fpl_functional h2 _ _ _ h9 h7). Qed. Lemma in_map_in_dep_list_of_lists_seqs_map : forall {T U V:Type} (hdt:type_eq_dec T) (hdu:type_eq_dec U) (p:T->U->V) (li:list T) (lj:list U) (ls: list (U*T)) (pfi:NoDup li) (pfj:NoDup lj) (pfin:In ls (list_power lj li)), let pffp := in_list_power_fpl hdu hdt lj li ls pfj (list_power_no_dup lj li ls pfj pfi pfin) pfin in let f := fpl_to_fun_in_dep pffp in exists l:list V, In l (list_of_lists_seqs (map (fun j : U => map (fun i : T => p i j) li) lj)) /\ l = (map_in_dep (fun (j : U) (pfinj:In j lj) => p (f j pfinj) j)). intros T U V hdt hdu p li lj ls h1 h2 h3. pose proof (in_list_power_in_list_of_lists_seqs_map_map_pair' _ _ _ h3) as h4. pose proof (in_list_of_lists_seqs_map_map_pair_rev' p _ _ _ h4) as h5. destruct h5 as [l h5]. destruct h5 as [h5l h5r]. exists l. subst. split; auto. pose proof (in_list_power_synced _ _ _ h3) as h6. pose proof (list_power_no_dup _ _ _ h2 h1 h3) as h7. pose proof (in_list_power_fpl hdu hdt _ _ _ h2 h7 h3) as h8. pose proof (fpl_impl_map _ _ h2 _ h8 h6) as h9. rewrite h9 at 1. rewrite map_map_in_dep. apply map_in_dep_ext. intros. apply functional_extensionality. intro. f_equal. simpl. f_equal. apply proof_irrelevance. Qed. Lemma in_map_in_dep_list_of_lists_seqs_map' : forall {T U V:Type} (pft:type_eq_dec T) (pfu:type_eq_dec U) (p:T->U->V) (li:list T) (lj:list U) (ls:list (U*T)) (pfi:NoDup li) (pfj:NoDup lj) (pfin:In ls (list_power lj li)), let pffp := in_list_power_fpl pfu pft lj li ls pfj (list_power_no_dup lj li ls pfj pfi pfin) pfin in let f := fpl_to_fun_in_dep pffp in exists l: list V, In l (list_of_lists_seqs (map (fun j : U => map (fun (i:T) => p i j) li) lj)) /\ l = map_in_dep (fun (j : U) (pfj:In j lj) => p (f j pfj) j). intros T U V hdt hdu p li lj ls h1 h2 h3. simpl. pose proof (in_list_power_in_list_of_lists_seqs_map_map_pair _ _ _ h3) as h4. pose proof (in_list_power_in_list_of_lists_seqs_map_map_pair' _ _ _ h3) as h4'. pose proof (in_list_of_lists_seqs_map_map_pair_rev' p _ _ _ h4') as h5. destruct h5 as [l h5]. destruct h5 as [h5l h5r]. exists l. subst. split; auto. pose proof (in_list_power_synced _ _ _ h3) as h6. pose proof (list_power_no_dup _ _ _ h2 h1 h3) as h7. pose proof (in_list_power_fpl hdu hdt _ _ _ h2 h7 h3 ) as h8. pose proof (fpl_impl_map _ _ h2 _ h8 h6) as h9. simpl in h9. rewrite h9 at 1. rewrite map_map_in_dep. apply map_in_dep_ext. intros y h10. apply functional_extensionality. intro h11. simpl. f_equal. f_equal. apply proof_irrelevance. Qed. Lemma sync_fpl : forall {T U:Type} {la:list T} {lb:list U} {lp:list (T*U)}, functionally_paired_l la lb lp -> exists lp':list (T*U), functionally_paired_l la lb lp' /\ synced la lp' /\ list_to_set lp = list_to_set lp'. intros T U la lb lp h1. inversion h1 as [h2 h3]. pose (fun (x:{t:T | In t la}) => proj1_sig (constructive_definite_description _ (h2 (proj1_sig x) (proj2_sig x)))) as f. pose (fun (x:{t:T | In t la}) => (proj1_sig x, f x)) as g. pose (map g (map_sig la)) as lp'. exists lp'. repeat split. intros x h4. exists (f (exist _ x h4)). red. repeat split. unfold f. destruct constructive_definite_description as [b h6]. simpl in h6. simpl. destruct h6 as [h6l h6r]. assumption. unfold lp'. rewrite in_map_iff. unfold g. exists (exist _ x h4). simpl. split. reflexivity. apply map_sig_in. intros b h5. unfold f. destruct constructive_definite_description as [b' h6]. simpl in h6. simpl. unfold lp' in h5. destruct h5 as [h5l h5r]. rewrite in_map_iff in h5r. destruct h5r as [x' h7]. destruct h7 as [h7l h7r]. unfold g in h7l. inversion h7l as [h8]. unfold f in H. unfold f. destruct constructive_definite_description as [a h9]. simpl. simpl in H. subst. destruct h9 as [h9l h9r]. destruct h6 as [h6l h6r]. apply (fpl_functional h1 _ _ _ h6r h9r). unfold lp' in H. unfold g in H. rewrite in_map_iff in H. destruct H as [x h4]. destruct h4 as [h4l h4r]. rewrite surjective_pairing in h4l. inversion h4l. apply proj2_sig. unfold lp' in H. unfold g in H. rewrite in_map_iff in H. destruct H as [x h4]. destruct h4 as [h4l h4r]. rewrite surjective_pairing in h4l. inversion h4l. unfold f. destruct constructive_definite_description as [x' h6]. simpl. destruct h6; assumption. unfold lp'. rewrite map_map. simpl. apply map_sig_compat. apply Extensionality_Ensembles. red. split. red. (* <= *) intros pr h4. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h4. unfold lp'. unfold g. rewrite in_map_iff. rewrite (surjective_pairing pr) in h4. pose proof (fpl_in_dom h1 _ h4) as h5. simpl in h5. exists (exist _ (fst pr) h5). simpl. split. rewrite surjective_pairing. unfold f. destruct constructive_definite_description as [x h6]. simpl. simpl in h6. destruct h6 as [h6l h6r]. pose proof (fpl_functional h1 _ _ _ h4 h6r). subst. reflexivity. apply map_sig_in. red. intros pr h4. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h4. unfold lp' in h4. unfold g in h4. rewrite in_map_iff in h4. destruct h4 as [x h4]. destruct h4 as [h4l h4r]. rewrite <- h4l. unfold f. destruct constructive_definite_description as [b h5]. simpl. destruct h5; assumption. Qed. Section NprodList. Lemma in_nprod_iff' : forall {T n} y (x:T^n), in_nprod y x <-> In y (nprod_to_list _ _ x). intros T n; induction n as [|n ih]; simpl. tauto. intros y x. destruct x as [f s]; simpl. split; intro h1; destruct h1 as [|h1]; auto; right. rewrite ih in h1; auto. rewrite ih; auto. Qed. Lemma nprod_to_list_nprod_of_list_eq : forall {T:Type} (l:list T), nprod_to_list _ (length l) (nprod_of_list _ l) = l. intros T l. induction l as [|a l h1]; simpl. reflexivity. f_equal; auto. Qed. Lemma nprod_of_list_nprod_to_list_eq : forall {T n} (x:T^n), nprod_of_list _ (nprod_to_list _ _ x) = transfer_nprod_r (length_nprod_to_list _) x. intros T n. induction n as [|n ih]; simpl. intro; apply unit_eq. intro x; destruct x as [t x]; simpl. rewrite ih, transfer_nprod_r_S; repeat f_equal; apply proof_irrelevance. Qed. Lemma all_sing_nprod_to_list_iff : forall {T n} (x:T^n) a, all_sing (nprod_to_list _ _ x) a <-> all_sing_nprod x a. intros T n x; split; intro h1; red in h1; red; intros b h2; [rewrite in_nprod_iff' in h2 | rewrite <- in_nprod_iff' in h2]; auto. Qed. Lemma all_sing_nprod_of_list_iff : forall {T} (l:list T) a, all_sing_nprod (nprod_of_list _ l) a <-> all_sing l a. intros T n x; split; intro h1; red in h1; red; intros b h2. specialize (h1 b). hfirst h1. rewrite in_nprod_iff'. rewrite nprod_to_list_nprod_of_list_eq; auto. auto. rewrite in_nprod_iff' in h2. rewrite nprod_to_list_nprod_of_list_eq in h2; auto. Qed. Lemma app_list_nprod_to_list_eq_sing_iff : forall {T n} (x:(list T)^n) (a:T), (forall c, in_nprod c x -> c <> nil) -> (app_list (nprod_to_list _ _ x) = a::nil <-> exists pfe:n = 1, hd_nprod (lt_eq_trans _ _ _ (lt_0_Sn 0) (eq_sym pfe)) x = a::nil). intros T n; induction n as [|n ih]; intros x a h1; split; intro h2. simpl in h2. discriminate. destruct h2; discriminate. destruct x as [c x]; simpl; simpl in h2, h1. destruct c as [|c]; simpl. simpl in h2. destruct n as [|n]; simpl in h2. discriminate. destruct x as [c x]; simpl in h2. rewrite app_eq_sing_iff in h2. destruct h2 as [h2|h2]. destruct h2 as [? h2]; subst. simpl in h1; specialize (h1 nil). hfirst h1; tauto. destruct h2 as [h2]; subst. specialize (h1 nil). hfirst h1; tauto. simpl in h2. inversion h2; subst. inversion h2. apply app_eq_nil in H0. destruct H0; subst. rewrite app_list_eq_nil_iff in H0. destruct H0 as [h3|h3]. rewrite (nprod_to_list_eq_nil_iff x) in h3; subst. exists (eq_refl _). f_equal; auto. simpl in H1. rewrite app_list_eq_nil_iff in H1. destruct H1 as [h4|h4]. rewrite (nprod_to_list_eq_nil_iff x) in h4; subst. exists (eq_refl _). f_equal; auto. inversion h2; clear h2. rewrite app_list_eq_nil_iff in H0. destruct H0 as [h5|h5]. rewrite (nprod_to_list_eq_nil_iff x) in h5; subst. exists (eq_refl _). f_equal; auto. red in h3. specialize (ih x a). destruct (zerop n) as [|h6]; subst. exists (eq_refl _),. specialize (h1 (hd_nprod h6 x)). hfirst h1. right. rewrite hd_nprod_eq. apply in_nth_prod. apply h1 in hf. red in h4. specialize (h4 (hd_nprod h6 x)). hfirst h4. rewrite hd_nprod_eq. rewrite <- in_nprod_iff'. apply in_nth_prod. specialize (h4 hf0). contradiction. destruct h2 as [h2 h3]. pose proof (S_inj _ _ h2); subst. destruct x as [? x]; simpl; simpl in h3. rewrite app_nil_r; auto. Qed. Fixpoint map_tuple_prod_to {T U n} {A:T->Type} (f:forall t:T, A t -> U) : forall (x:T^n) (y:tuple_prod A x), list U := match n with | 0 => fun _ _ => nil | S n' => fun (x:T^S n') (y:tuple_prod A x) => f (fst x) (fst y) :: map_tuple_prod_to f (snd x) (snd y) end. Lemma in_map_tuple_prod_to : forall {T U n} {A:T->Type} (f:forall t:T, A t -> U) (x:T^n) (y:tuple_prod A x) i (pfi:i < n), In (f (nth_prod i pfi x) (nth_tuple_prod y i pfi)) (map_tuple_prod_to f x y). intros T U n; induction n as [|n ih]; simpl. intros; omega. intros A f x y i hi. destruct x as [xf xs]; simpl in y. destruct y as [yf ys]; simpl. destruct i as [|i]; simpl. left; auto. specialize (ih A f xs ys i (lt_S_n _ _ hi)). right; auto. Qed. Lemma in_map_tuple_prod_to_iff : forall {T U n} {A:T->Type} (f:forall t:T, A t -> U) (x:T^n) (y:tuple_prod A x) (u:U), In u (map_tuple_prod_to f x y) <-> exists i (pfi:i < n), u = f (nth_prod i pfi x) (nth_tuple_prod y i pfi). intros T U n. induction n as [|n ih]; simpl. intros; split; intro h1. contradiction. destruct h1 as [? [? ?]]; omega. intros A f x y u. destruct x as [xf xs]; simpl in y. destruct y as [yf ys]; simpl. split; intro h1. destruct h1 as [|h1]; subst. exists 0, (lt_0_Sn _); simpl; auto. rewrite ih in h1. destruct h1 as [i [hi ?]]; subst. exists (S i), (lt_n_S _ _ hi); simpl. gtermr. redterm0 y. redterm1 y. redterm0 c. fold c1. pose proof (proof_irrelevance _ c1 hi) as h2. rewrite h2; auto. destruct h1 as [i [hi ?]]; subst. destruct i as [|i]; simpl. left; auto. right. specialize (ih A f xs ys). gterm1. fold c. specialize (ih c). rewrite ih. unfold c. exists i, (lt_S_n _ _ hi); auto. Qed. Lemma nth_prod_to_list : forall {T:Type} {i n:nat} (pflt:i < n) (x:T^n), let pflt' := lt_eq_trans _ _ _ pflt (eq_sym (length_nprod_to_list _)) in nth_prod i pflt x = nth_lt (nprod_to_list _ _ x) i pflt'. intros T i n. revert i. induction n as [|n ih]; simpl. intros; omega. intros i hlt x. destruct x as [x y], i as [|i]; simpl; auto. rewrite ih. erewrite nth_lt_compat, nth_lt_functional3. apply nth_indep. rewrite length_nprod_to_list; auto with arith. Grab Existential Variables. rewrite length_nprod_to_list; auto with arith. Qed. Lemma nth_prod_of_list : forall {T:Type} (l:list T) (i:nat) (pflt:i < length l), nth_prod i pflt (nprod_of_list _ l) = nth_lt l i pflt. intros T l. induction l as [|a l ih]. simpl; intros; omega. intros i hlt. destruct i as [|i]. rewrite nth_lt_0'. simpl; auto. simpl in hlt. pose proof hlt as h1. apply lt_S_n in h1. rewrite nth_lt_S; simpl. rewrite ih. apply nth_lt_functional3. Qed. Lemma nprod_inj : forall {T:Type} {n:nat} (x1 x2:T^n) (def:T), x1 = x2 -> forall m:nat, m < n -> nth m (nprod_to_list _ _ x1) def = nth m (nprod_to_list _ _ x2) def. intros; subst; auto. Qed. Fixpoint in_n {T:Type} {n:nat} (l:list T) : T^n -> Prop := match n return (T^n-> Prop) with | O => (fun x => True) | S n' => (fun x => In (fst x) l /\ (in_n l (snd x))) end. Lemma in_n_in_nprod_to_list_compat : forall {T:Type} {n:nat} (l:list T) (t:T^n), in_n l t -> (forall x:T, In x (nprod_to_list _ _ t) -> In x l). intros T n. induction n as [|n h1]; simpl. intros l tt. tauto. intros l t. intros h2 ft' h3. destruct t as [ft st]. destruct h3 as [h3l | h3r]. subst. simpl in h2. destruct h2; auto. simpl in h2. destruct h2 as [h2l h2r]. specialize (h1 _ _ h2r _ h3r). assumption. Qed. Lemma in_nprod_to_list_in_n_compat : forall {T:Type} {n:nat} (l:list T) (t:T^n), (forall x:T, In x (nprod_to_list _ _ t) -> In x l) -> in_n l t. intros T n. induction n as [|n h1]; simpl. intros. tauto. intros l t h2. destruct t as [ft st]. simpl. assert (h3:forall x:T, In x (nprod_to_list T n st) -> In x l). intros x h3. pose proof (in_cons ft x _ h3) as h4. apply h2; auto. specialize (h1 _ _ h3). specialize (h2 ft (in_eq _ _)). split; auto. Qed. Lemma in_n_in_nprod_to_list_compat_iff : forall {T:Type} {n:nat} (l:list T) (t:T^n), in_n l t <-> (forall x:T, In x (nprod_to_list _ _ t) -> In x l). intros; split; [apply in_n_in_nprod_to_list_compat | apply in_nprod_to_list_in_n_compat]. Qed. Lemma in_n_cons : forall {T:Type} {n:nat} (l:list T) (a:T) (t:T^n), in_n l t -> in_n (a::l) t. intros T n l a t h1. rewrite in_n_in_nprod_to_list_compat_iff in h1. rewrite in_n_in_nprod_to_list_compat_iff. intros x h2. specialize (h1 _ h2). right. assumption. Qed. Lemma in_n_nprod_to_list : forall {T:Type} {n:nat} (t:T^n), in_n (nprod_to_list _ _ t) t. intros T n. induction n as [|n h1]; simpl. auto. intro t. split. destruct t; simpl. left. auto. destruct t; simpl. apply in_n_cons. apply h1. Qed. Lemma nprod_to_fpl_ex : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U), NoDup la -> let n := length la in forall (t:U^n), in_n lb t -> exists! lp:list (T*U), NoDup lp /\ synced la lp /\ exists (pf:functionally_paired_l la lb lp), let f := fpl_to_fun_in_dep pf in map_in_dep f = nprod_to_list _ _ t. intros T U hdt hdu la. induction la as [|a la h1]; simpl. intros. exists (list_prod nil lb). red. simpl. split. split. constructor. split; try constructor. simpl. reflexivity. constructor. intros; contradiction. intros; contradiction. reflexivity. intros pl h1. destruct h1 as [h1l [h1c h1r]]. destruct h1r as [h1r]. symmetry. apply (fpl_empty1_s _ _ h1r). intros lb h0 t h2. destruct h2 as [h2l h2r]. inversion h0 as [|? ? h0l h0r]. subst. specialize (h1 lb h0r _ h2r). destruct h1 as [lp h1]. red in h1. destruct h1 as [h1l h1r]. destruct h1l as [h3 h4]. destruct h4 as [h4 h5]. exists ((a, (fst t))::lp). red. split. split. constructor. intro hi. destruct h5 as [h5l h5r]. pose proof (fpl_in_dom h5l _ hi) as h0'. simpl in h0'. contradiction. assumption. split. apply synced_cons''. assumption. destruct h5 as [h5 h6]. pose proof (fpl_cons_nin _ _ _ _ _ h0l h2l h5) as h8. exists h8. assert (h10: map_in_dep (fun (x : T) (h1 : In x la) => fpl_to_fun_in_dep h8 x (in_cons _ _ _ h1)) = map_in_dep (fpl_to_fun_in_dep h5)). apply map_in_dep_ext. intros x h11. apply functional_extensionality. intro hin. pose proof (proof_irrelevance _ hin h11); subst. pose proof (fpl_to_fun_in_dep_compat h8 x (in_cons _ _ _ h11)) as h8'. pose proof (fpl_to_fun_in_dep_compat h5 x h11) as h5'. assert (h12:x <> a). intro; subst; contradiction. inversion h8' as [hinv | hinv]. inversion hinv. subst. contradict h12; auto. eapply fpl_functional. apply h5. apply hinv. assumption. unfold nprod. unfold nprod in h10. unfold fun_from_cons. rewrite h10. rewrite h6. pose proof (fpl_to_fun_in_dep_compat h8 a (in_eq a la)) as h11. destruct h11 as [h11l | h11r]. inversion h11l. destruct t. simpl. reflexivity. pose proof (fpl_in_dom h5 _ h11r) as h12. simpl in h12. contradiction. intros lp' h6. destruct h6 as [h6l [h6c h6r]]. destruct h6r as [h7 h8]. pose proof (synced_cons' _ _ _ h6c) as h9. specialize (h1r (tl lp')). assert (h13:hd (a, fpl_to_fun_in_dep h7 a (in_eq a la)) lp' = (a, fpl_to_fun_in_dep h7 a (in_eq a la)) ). inversion h6c as [h12]. apply injective_projections. simpl. pose proof (synced_hd (a, fpl_to_fun_in_dep h7 a (in_eq a la)) _ _ h6c) as h13. simpl in h13. assumption. simpl. assert (h13:lp' <> nil). intro h14. subst. apply synced_nil2 in h6c. discriminate h6c. pose proof (in_hd_not_nil (impl_type_eq_dec_prod hdt hdu) _ (a, fpl_to_fun_in_dep h7 a (in_eq a la)) h13) as h14. pose proof (fpl_to_fun_in_dep_in_eq _ _ _ h7 _ h14) as h15. rewrite h15. f_equal. assert (h13': (hd a (a::la)) = a). simpl. reflexivity. rewrite <- h12 in h13'. assert (h16:a = fst (a, fpl_to_fun_in_dep h7 a (in_eq a la))). auto. rewrite h16 in h13' at 1. rewrite hd_map in h13'. gen0. rewrite h13'. intro. f_equal. apply proof_irrelevance. hfirst h1r. split. apply no_dup_tl; auto. split; auto. pose proof (fpl_to_fun_in_dep_compat h7 a (in_eq a la)) as h10. pose proof (fpl_cons hdt hdu _ _ _ _ h10 h0 h7) as h11. assert (h12: (remove (impl_type_eq_dec_prod hdt hdu) (a, fpl_to_fun_in_dep h7 a (in_eq a la)) lp') = tl lp'). rewrite <- h13. rewrite remove_hd_no_dup; auto. rewrite h12 in h11. exists h11. destruct t. inversion h8 as [h13']. simpl. erewrite map_in_dep_fpl_to_fun_cons. apply H. assumption. specialize (h1r hf). subst. destruct t. simpl. inversion h8. assert (h12:lp' <> nil). intro h13'. subst. simpl in h9. destruct h6c as [h6c]. simpl in h6c. discriminate. pose proof (list_decompose _ (a, u) h12) as h13'. rewrite h13'. f_equal. rewrite surjective_pairing. f_equal. rewrite (synced_hd _ _ _ h6c). simpl. reflexivity. pose proof (in_hd_not_nil (impl_type_eq_dec_prod hdt hdu) _ (a, u) h12 ) as h14. pose proof (fpl_to_fun_in_dep_in_eq _ _ _ h7 _ h14) as h15. rewrite h15. inversion h6c as [h16]. pose proof (in_eq a la) as h17. rewrite <- h16 in h17. pose proof (hd_map (fun pr:T*U=>fst pr) (a, u) lp') as h18. rewrite h16 in h18. simpl in h18. symmetry. gen0. rewrite <- h18. intro. f_equal. apply proof_irrelevance. Qed. Definition nprod_to_fpl {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (pfla:NoDup la) (t:U^(length la)) (pfin:in_n lb t) := proj1_sig (constructive_definite_description _ (nprod_to_fpl_ex pfdt pfdu la lb pfla t pfin)). Lemma nprod_to_fpl_compat : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (pfla:NoDup la), let n := length la in forall (t:U^n) (pfin:in_n lb t), let lp := nprod_to_fpl pfdt pfdu la lb pfla t pfin in NoDup lp /\ synced la lp /\ exists (pf:functionally_paired_l la lb lp), let f := fpl_to_fun_in_dep pf in map_in_dep f = nprod_to_list _ _ t. unfold nprod_to_fpl; intros; destruct constructive_definite_description; auto. Qed. Lemma nprod_to_list_inj : forall {T:Type} (n:nat) (t t':T^n), nprod_to_list _ n t = nprod_to_list _ n t' -> t = t'. intros T n. induction n as [|n h0]. simpl. intros t t' h1. destruct t; destruct t'. reflexivity. simpl. intros t t' h1. destruct t as [ft st]; destruct t' as [ft' st']. inversion h1. subst. clear h1. specialize (h0 st st' H1). apply injective_projections; auto. Qed. Lemma nprod_to_fpl_inj : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (pfla:NoDup la) (t t':U^(length la)) (pfin:in_n lb t) (pfin':in_n lb t'), nprod_to_fpl pfdt pfdu la lb pfla t pfin = nprod_to_fpl pfdt pfdu la lb pfla t' pfin' -> t = t'. intros T U hdt hdu la lb h1 t t' h2 h3 h4. pose proof (nprod_to_fpl_compat hdt hdu la lb h1 t h2) as h5. pose proof (nprod_to_fpl_compat hdt hdu la lb h1 t' h3) as h6. destruct h5 as [h5a [h5b h5c]]. destruct h6 as [h6a [h6b h6c]]. destruct h5c as [h5c h5d]. destruct h6c as [h6c h6d]. generalize dependent h6c. rewrite <- h4. intro h6c. pose proof (proof_irrelevance _ h5c h6c); subst. intro h7. simpl in h7. rewrite h5d in h7. apply (nprod_to_list_inj _ _ _ h7). Qed. Lemma nprod_to_fpl_nil : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (lb:list U) (pfla:NoDup (nil (A:=T))) (t:unit) (pfin:True), nprod_to_fpl pfdt pfdu nil lb pfla t pfin = nil. intros T U hdt hdu lb h1 t h2. pose proof (nprod_to_fpl_compat hdt hdu nil lb h1 t h2) as h3. destruct h3 as [h3a [h3b h3c]]. apply synced_nil1. assumption. Qed. Lemma fpl_in_map_snd_impl_in_lb : forall {T U:Type} (la:list T) (lb:list U) (lp:list (T*U)), functionally_paired_l la lb lp -> forall (y:U), In y (map (snd (B:=U)) lp) -> In y lb. intros T U la lb lp h1 y. intro h2. rewrite in_map_iff in h2. destruct h2 as [pr [h2l h2r]]. destruct h1 as [h1l h1r]. specialize (h1r _ h2r). destruct h1r as [h1a h1b]. subst. assumption. Qed. Lemma map_in_dep_fpl_to_prod : forall {T U:Type} {la:list T} {lb:list U} {lp:list (T*U)} (pffp:functionally_paired_l la lb lp), type_eq_dec T -> type_eq_dec U -> NoDup la -> NoDup lp -> synced la lp -> forall (f:fun_in_dep la U), map_in_dep f = map_in_dep (fun (pr:T*U) (pf:In pr lp) => f (fst pr) (fpl_in_dom pffp pr pf)). intros T U la. induction la as [|a la ih]; simpl. intros ? lp h1 hdt hdu ? f. pose proof (fpl_empty1_s _ _ h1) as h1'. subst. simpl; auto. intros lb lp h1 hdt hdu h7 hnd hs f. pose proof (fpl_to_fun_in_dep_compat h1 _ (in_eq _ _)) as h5. pose proof (fpl_cons hdt hdu la lb lp _ h5 h7 h1) as h6. simpl in h6. pose proof (no_dup_cons_nin _ _ h7) as h8. pose proof (no_dup_cons _ _ h7) as h9. pose proof (no_dup_remove (impl_type_eq_dec_prod hdt hdu) lp (a, fpl_to_fun_in_dep h1 a (in_eq a la)) hnd) as h10. specialize (ih _ _ h6 hdt hdu h9 h10). pose proof (remove_synced_cons hdt hdu la lb lp (a, fpl_to_fun_in_dep h1 a (in_eq a la)) hs h1 h5 hnd) as h11. generalize dependent h6. rewrite h11. intros h6 h12. pose proof (in_not_nil _ _ h5) as h13. pose proof (list_decompose lp (a, fpl_to_fun_in_dep h1 a (in_eq a la)) h13) as h14. pose proof (synced_hd (a, fpl_to_fun_in_dep h1 a (in_eq a la)) _ _ hs) as h15. simpl in h15. rewrite <- h15 in hs. pose proof (remove_synced_cons hdt hdu la lb lp _ hs) as h16. rewrite h15 in h16. specialize (h16 h1). hfirst h16. apply in_hd_not_nil; auto. apply impl_type_eq_dec_prod; auto. specialize (h16 hf hnd). rewrite <- h16 in h14. match goal with |- _ = map_in_dep ?g' => pose g' as g end. pose proof (map_in_dep_functional g h14) as h17. fold g. rewrite h17. rewrite map_in_dep_cons. rewrite transfer_fun_in_dep_compat. unfold g. f_equal. symmetry. gen0. rewrite h15. intro. f_equal. apply proof_irrelevance. rewrite h15 in hs. pose proof (fpl_impl_map _ lb h7 lp h1 hs) as h18. simpl in h18. match goal with |- _ = map_in_dep ?q => pose q as q' end. rewrite h14 in h18 at 1. inversion h18 as [h19]. clear h18. match goal with |- _ = map_in_dep ?k => pose proof (map_in_dep_functional k H) as h22 end. rewrite h22. rewrite map_in_dep_map_in_dep. apply map_in_dep_ext. intros x h23. apply functional_extensionality. intro y. rewrite transfer_fun_in_dep_compat. unfold fun_from_cons. rewrite transfer_fun_in_dep_compat. f_equal. apply proof_irrelevance. Qed. Lemma list_fst_snd_eq : forall {T U:Type} (lp lp':list (T*U)), map (fst (B:=U)) lp = map (fst (B:=U)) lp' -> map (snd (B:=U)) lp = map (snd (B:=U)) lp' -> lp = lp'. intros T U lp. induction lp as [|a lp h1]; auto; simpl. intros lp' h2 h3. symmetry in h3. pose proof (map_eq_nil _ _ h3). subst. reflexivity. intro lp'. induction lp' as [|a' lp' h2]; simpl. intros h2 h3. discriminate. intros h3 h4. inversion h3; inversion h4. specialize (h1 _ H1 H3). f_equal. apply injective_projections; auto. assumption. Qed. Lemma fpl_to_fun_in_dep_eq : forall {T U:Type} {la:list T} {lb:list U} {lp lp':list (T*U)} (pf: functionally_paired_l la lb lp) (pf':functionally_paired_l la lb lp'), lp = lp' -> fpl_to_fun_in_dep pf = fpl_to_fun_in_dep pf'. intros ? ? ? ? ? ? h1 h2 ?; subst. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. Lemma fpl_to_fun_in_dep_eq' : forall {T U:Type} {la:list T} {lb:list U} {lp lp':list (T*U)} (pffp: functionally_paired_l la lb lp) (pffp':functionally_paired_l la lb lp'), lp = lp' -> forall (x:T) (pfa:In x la), fpl_to_fun_in_dep pffp x pfa = fpl_to_fun_in_dep pffp' x pfa. intros ? ? ? ? ? ? h1 h2 ?; subst. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. Fixpoint nprod_of_list_n {T:Type} (l:list T) (n:nat) (def:T) {struct n} : T ^ n := match n return T^n with | O => tt | S n' => ((hd def l), (nprod_of_list_n (tl l) n' def)) end. Lemma nprod_to_fpl_surj : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (lp:list (T*U)) (pf:functionally_paired_l la lb lp) (pfla:NoDup la), NoDup lp -> synced la lp -> let f := fpl_to_fun_in_dep pf in let n := length la in exists (t:U^n) (pfin:in_n lb t), map_in_dep f = nprod_to_list _ n t /\ lp = (nprod_to_fpl pfdt pfdu la lb pfla t pfin). intros T U hdt hdu la lb lp h1 h2 h3 h4 f n. pose proof (synced_length _ _ h4) as h5. pose (nprod_of_list _ (map (@snd _ _) lp)) as t. assert (h6':length (map (snd (B:=U)) lp) = n). unfold n. rewrite h5. apply map_length. subst. pose (transfer_dep h6' t) as t'. exists t'. assert (h8:in_n lb t). rewrite in_n_in_nprod_to_list_compat_iff. intros x h8. unfold t in h8. rewrite nprod_to_list_nprod_of_list_eq in h8. eapply fpl_in_map_snd_impl_in_lb; auto. apply h1; auto. assumption. pose (fun n':(sigT (fun n => (nprod U n))) => (in_n (n:=(projT1 n')) lb (projT2 n'))) as P. pose proof (transfer_dep_prop _ _ h6' t P) as h9. unfold P in h9. simpl in h9. unfold t'. rewrite h9 in h8. exists h8. split. pose proof (fpl_impl_map _ _ h2 _ h1 h4) as h10. pose proof (f_equal (map (@snd T U)) h10) as h11. rewrite map_map_in_dep in h11. simpl in h11. let P := type of h11 in match P with _ = ?x => assert (h12:x = map_in_dep f) end. apply map_in_dep_ext. intros. apply functional_extensionality. intro. unfold f; auto. rewrite <- h12. unfold t. pose proof (transfer_dep_existT _ _ h6' (nprod_of_list U (map (snd(B:=U)) lp))) as h13. dependent rewrite <- h13. rewrite nprod_to_list_nprod_of_list_eq. rewrite (map_in_dep_fpl_to_prod h1); auto. rewrite map_eq_map_in_dep. eapply map_in_dep_ext. intros pr h14. apply functional_extensionality. intro. erewrite fpl_to_fun_in_dep_in_eq. f_equal. pose proof (nprod_to_fpl_compat hdt hdu _ _ h2 _ h8) as h10. destruct h10 as [h10a [h10b h10c]]. destruct h10c as [h10 h11]. unfold t in h11. pose proof (transfer_dep_existT _ _ h6' (nprod_of_list U (map (snd (B:=U)) lp))) as h13. rewrite <- nprod_to_list_nprod_of_list_eq in h11 at 1. pose proof (fpl_impl_map _ _ h2 _ h10 h10b) as h14. pose proof (f_equal (map (@snd T U)) h14) as h15. rewrite h14. pose proof (fpl_impl_map _ _ h2 _ h1 h4) as h10'. pose proof (f_equal (map (@snd T U)) h10') as h11'. rewrite map_map_in_dep in h11'. simpl in h11'. rewrite h10' at 1. f_equal. apply fun_in_dep_ext. intros; auto. f_equal. assert (h16:lp = (nprod_to_fpl hdt hdu la lb h2 (transfer_dep h6' t) h8)). unfold t. unfold nprod_to_fpl. destruct constructive_definite_description as [lp' h18]. simpl. destruct h18 as [h18a [h18b h18c]]. destruct h18c as [h19 h20]. dependent rewrite <- h13 in h20. rewrite nprod_to_list_nprod_of_list_eq in h20. rewrite (map_in_dep_fpl_to_prod h19) in h20; auto. apply list_fst_snd_eq. destruct h18b as [h18b]. pose proof h4 as h4'. clear h4. destruct h4' as [h4']. rewrite <- h18b in h4'. assert (h21:(fun pr:T*U=>fst pr) = (fst (B:=U))). apply functional_extensionality; auto. rewrite h21 in h4'. assumption. symmetry. rewrite <- h20. rewrite map_eq_map_in_dep. apply map_in_dep_ext. intros pr h21. apply functional_extensionality. intro. symmetry. erewrite fpl_to_fun_in_dep_in_eq. f_equal. apply fpl_to_fun_in_dep_eq'; auto. Qed. Lemma in_n_Sn : forall {T:Type} {n:nat} (l:list T) (t:T^n) (x:T), in_n (n:=S n) l (x, t) -> In x l /\ in_n l t. intros T n l t x h1. rewrite in_n_in_nprod_to_list_compat_iff in h1. split. apply h1. simpl. left. reflexivity. assert (h2:forall x0 : T, In x0 (nprod_to_list T n t) -> In x0 l). intros x' h2. specialize (h1 x'). simpl in h1. apply h1. right. assumption. rewrite <- in_n_in_nprod_to_list_compat_iff in h2. assumption. Qed. Lemma im_nprod_new_proj2_ex : forall {T:Type} {l:list T} {n:nat} (t':{t:T^n | in_n l t}), exists! t:{x:T|In x (nprod_to_list _ _ (proj1_sig t'))}^n, proj1_sig t' = im_nprod (@proj1_sig _ _) t. intros T l n t'. induction n as [|n ret]. (* 0 *) simpl. exists tt. red. simpl in t'. destruct t'. simpl. destruct x. split. reflexivity. intros tt' ?. destruct tt'. assumption. (* S n*) destruct t' as [t' h2]. destruct t' as [ft' st']; simpl. apply in_n_Sn in h2. destruct h2 as [h2l h2r]. specialize (ret (exist _ _ h2r)). simpl in ret. pose (exist (fun x=>ft'=x \/ In x (nprod_to_list T n st')) ft' (or_introl (In ft' (nprod_to_list T n st')) (eq_refl _))) as ft. destruct ret as [ret h3]. red in h3. pose (im_nprod (U:={x:T|ft' = x \/ In x (nprod_to_list T n st')}) (fun x=>(exist _ (proj1_sig x) (or_intror (ft' = (proj1_sig x)) (proj2_sig x)))) ret) as st. exists (ft, st). red. simpl. split. unfold st. apply injective_projections; simpl. reflexivity. rewrite im_nprod_im_nprod. simpl. destruct h3 as [h3l h3r]. rewrite h3l at 1. f_equal. intros x h4. destruct x as [fx sx]. simpl in h4. inversion h4 as [h5]. destruct fx as [fx h6]. simpl in h4. simpl in h5. unfold ft. unfold st. apply injective_projections. simpl. destruct h6. apply proj1_sig_injective. simpl. assumption. apply proj1_sig_injective. simpl. assumption. simpl. destruct h3 as [h3l h3r]. rewrite h3l in H at 1. apply im_nprod_inj with (f:=(proj1_sig (P:=fun x:T=>ft' = x \/ In x (nprod_to_list T n st')))). red. apply proj1_sig_injective. rewrite im_nprod_im_nprod. simpl. assumption. Qed. Definition im_nprod_new_proj2 {T:Type} {l:list T} {n:nat} (t':{t:T^n | in_n l t}) : {x:T|In x (nprod_to_list _ _ (proj1_sig t'))}^n := (proj1_sig (constructive_definite_description _ (im_nprod_new_proj2_ex t'))). Lemma im_nprod_new_proj2_compat : forall {T:Type} {l:list T} {n:nat} (t':{t:T^n | in_n l t}), proj1_sig t' = im_nprod (@proj1_sig _ _) (im_nprod_new_proj2 t'). intros T l n t'. unfold im_nprod_new_proj2. destruct constructive_definite_description. simpl. assumption. Qed. Lemma nprod_in_n_to_nprod_in_l_ex : forall {T:Type} {l:list T} {n:nat} (t':{t:T^n|in_n l t}), exists! ret:{x:T|In x l}^n, proj1_sig t' = im_nprod (@proj1_sig _ _) ret. intros T l n t'. pose (im_nprod_new_proj2 t') as t''. pose proof (im_nprod_new_proj2_compat t') as h0. destruct t' as [t' h1]. simpl in h0. pose proof h1 as h1'. rewrite in_n_in_nprod_to_list_compat_iff in h1'. simpl in t''. simpl. exists (im_nprod (fun x=>(exist _ (proj1_sig x) (h1' _ (proj2_sig x)))) t''). red. split. rewrite im_nprod_im_nprod. simpl. unfold t''. assumption. intros x' h2. unfold t''. apply im_nprod_inj with (f:=(proj1_sig (P:=fun x:T => In x l))). red. apply proj1_sig_injective. rewrite im_nprod_im_nprod. simpl. rewrite <- h2. symmetry. assumption. Qed. Definition nprod_in_n_to_nprod_in_l {T:Type} (l:list T) (n:nat) (t':{t:T^n | in_n l t}) : {x:T|In x l}^n := (proj1_sig (constructive_definite_description _ (nprod_in_n_to_nprod_in_l_ex t'))). Lemma nprod_in_n_to_nprod_in_l_compat : forall {T:Type} {l:list T} {n:nat} (t':{t:T^n|in_n l t}), proj1_sig t' = im_nprod (@proj1_sig _ _) (nprod_in_n_to_nprod_in_l _ _ t') . intros T l n t'. unfold nprod_in_n_to_nprod_in_l. destruct constructive_definite_description. simpl. assumption. Qed. Lemma nprod_in_n_to_nprod_in_l_bij : forall {T:Type} (l:list T) (n:nat), bijective (nprod_in_n_to_nprod_in_l l n). intros T l n. red. split. red. intros x1 x2 h1. pose proof (f_equal (fun x => im_nprod (@proj1_sig _ _) x) h1) as h2. simpl in h2. do 2 rewrite <- nprod_in_n_to_nprod_in_l_compat in h2. apply proj1_sig_injective. assumption. red. intro t'. induction n as [|n h1]. simpl in t'. simpl. exists (exist _ tt I). unfold nprod_in_n_to_nprod_in_l. destruct constructive_definite_description. simpl. simpl in x. destruct t'; destruct x. reflexivity. simpl in t'. destruct t' as [ft' st']. specialize (h1 st'). destruct h1 as [t h2]. simpl. destruct t as [t h1]. destruct ft' as [x h3]. exists (exist _ (x, t) (conj h3 h1)). apply injective_projections. simpl. apply proj1_sig_injective. simpl. unfold nprod_in_n_to_nprod_in_l. destruct constructive_definite_description as [y h4]. simpl. simpl in h4. inversion h4. reflexivity. simpl. unfold nprod_in_n_to_nprod_in_l. destruct constructive_definite_description as [y h4]. simpl. simpl in h4. subst. inversion h4. apply im_nprod_inj with (f:=proj1_sig (P:=fun x0:T => In x0 l)). red. apply proj1_sig_injective. subst. rewrite <- nprod_in_n_to_nprod_in_l_compat. simpl. reflexivity. Qed. Lemma in_sig_fun_ex : forall {T:Type} {l:list T} (x:{t:T|In t l}), exists! y:{t:T|Inn (list_to_set l) t}, proj1_sig x = proj1_sig y. intros T l x. destruct x as [x h1]. pose proof h1 as h1'. rewrite list_to_set_in_iff in h1'. exists (exist _ _ h1'). red. simpl. split; auto. intros y h2. destruct y as [y h3]. simpl in h2. subst. pose proof (proof_irrelevance _ h3 h1'); subst; f_equal. Qed. Definition in_sig_fun {T:Type} {l:list T} (x:{t:T|In t l}):= proj1_sig (constructive_definite_description _ (in_sig_fun_ex x)). Lemma in_sig_fun_compat : forall {T:Type} {l:list T} (x:{t:T|In t l}), proj1_sig x = proj1_sig (in_sig_fun x). intros T l x. unfold in_sig_fun. destruct constructive_definite_description as [x' h1]. simpl. assumption. Qed. Lemma in_sig_fun_bij : forall {T:Type} (l:list T), bijective (@in_sig_fun _ l). intros T l. red. split. red. intros x1 x2 h1. pose proof (f_equal (@proj1_sig _ _) h1) as h2. do 2 rewrite <- in_sig_fun_compat in h2. apply proj1_sig_injective; auto. red. intro y. destruct y as [y h1]. pose proof h1 as h1'. rewrite <- list_to_set_in_iff in h1'. exists (exist _ y h1'). apply proj1_sig_injective. simpl. rewrite <- in_sig_fun_compat. simpl. reflexivity. Qed. Lemma finite_list_sig : forall {T:Type} (l:list T), FiniteT {x:T | In x l}. intros T l. pose proof (in_sig_fun_bij l) as h1. apply bijective_impl_invertible in h1. pose proof (list_to_set_finite l) as h2. pose proof (Finite_ens_type _ h2) as h3. apply (bij_finite _ _ (proj1_sig (function_inverse (@in_sig_fun _ l) h1)) h3 (invertible_impl_inv_invertible _ h1)). Qed. Lemma cardinal_full_set_in_l : forall {T:Type} (l:list T), NoDup l -> cardinal _ (Full_set {x:T | In x l}) (length l). intros T l h0. pose (list_to_set l) as A. induction l as [|a l h1]. simpl. rewrite full_false_empty. constructor. simpl. inversion h0. subst. specialize (h1 H2). pose (fun x':{x:T|In x l} => (exist (fun x=>a=x \/ In x l) (proj1_sig x') (or_intror (a=(proj1_sig x')) (proj2_sig x')))) as f. assert (h3:Injective f). red. unfold f. intros x1 x2 h4. destruct x1; destruct x2. simpl in h4. apply proj1_sig_injective. simpl. pose proof (f_equal (@proj1_sig _ _) h4) as h5. simpl in h5. assumption. pose proof (injection_preserves_cardinal _ _ f _ _ h1 h3) as h4. assert (h5:Im (Full_set {x:T | In x l}) f = [x':{x:T | a = x \/ In x l} | In (proj1_sig x') l]). apply Extensionality_Ensembles. red. split. red. intros y h5. destruct h5 as [x h5 y]. subst. constructor. unfold f. simpl. apply proj2_sig. red. intros y h5. destruct h5 as [h5]. destruct y as [y h6]. simpl in h5. apply Im_intro with (exist _ y h5). constructor. unfold f. apply proj1_sig_injective. simpl. reflexivity. rewrite h5 in h4. assert (h6:Full_set {x:T | a = x \/ In x l} = Add [x':{x:T | a = x \/ In x l} | In (proj1_sig x') l] (exist _ a (or_introl (In a l) (eq_refl _)))). apply Extensionality_Ensembles. red. split. red. intros x h6. destruct x as [x h7]. destruct h7 as [h7l | h7r]. right. rewrite h7l. constructor. constructor. constructor. simpl. assumption. red. constructor. rewrite h6. apply card_add. assumption. intro h7. destruct h7 as [h7]. simpl in h7. contradiction. Qed. Lemma cardinal_in_l : forall {T:Type} (l:list T), NoDup l -> FiniteT_card {x:T | In x l} (finite_list_sig l) = length l. intros T l h1. apply FiniteT_card_cond. apply cardinal_full_set_in_l; auto. Qed. Definition l_exp_seg_to_nprod_in_l {T:Type} (l:list T) (n:nat) (f:{m:nat | m < n} -> {x:T|In x l}) : {x:T | In x l} ^ n. induction n as [|n h1]. simpl. refine tt. simpl. pose (fun m':{m:nat | m < n} => f (exist _ (proj1_sig m') (lt_S _ _ (proj2_sig m')))) as f'. refine ((f (exist _ n (lt_n_Sn n)), (h1 f'))). Defined. Lemma l_exp_seg_to_nprod_in_l_compat : forall {T:Type} (l:list T) (n:nat) (f:{m:nat | m < n} -> {x:T|In x l}) (def:{x:T|In x l}) (x:nat) (pf:x < n), nth (n - x -1) (nprod_to_list _ _ (l_exp_seg_to_nprod_in_l _ _ f)) def = f (exist _ x pf). intros T l n f def x. revert f def. induction n as [|n h1]; induction x as [|x h2]. simpl; intros; try omega. intros; omega. intros f def h2. destruct (zerop n) as [h3 | h4]. subst; simpl. pose proof (proof_irrelevance _ h2 (lt_n_Sn _)); subst; auto. simpl. assert (h5:n-0 = n). auto with arith. rewrite h5. destruct n as [|n]. omega. specialize (h1 (fun m' : {m : nat | m < S n} => f (exist (fun m : nat => m < S (S n)) (proj1_sig m') (lt_S (proj1_sig m') (S n) (proj2_sig m')))) def h4). assert (h6:S n - 0 -1 = n). omega. rewrite h6 in h1. rewrite h1. simpl. pose proof (proof_irrelevance _ h2 (lt_S 0 (S n) h4)); f_equal; f_equal; auto. intros f def h3. assert (h4:x < n). omega. assert (h5:S n - S x - 1 = n - x -1). omega. rewrite h5. destruct (zerop (n - x -1)) as [h6 | h7]. rewrite h6. simpl. assert (h7:n = S x). omega. subst. pose proof (proof_irrelevance _ h3 (lt_n_Sn (S x))); subst; auto. destruct (lt_eq_lt_dec x (n-1)) as [[h10' | h11] | h12]. pose proof (S_pred _ _ h7) as h8. rewrite h8. simpl. assert (h9:pred (n - x -1) = n - S x - 1). apply S_inj. rewrite <- (S_pred _ _ h7). assert (h10:1 <= n - S x). pose proof (lt_le_S _ _ h7) as h9. rewrite (minus_plus_simpl_l_reverse (n-x) 1 x) in h9. assert (h11:x <= n). pose proof (lt_S _ _ h4) as h11. unfold lt in h11. apply le_S_n. assumption. rewrite <- (le_plus_minus _ _ h11) in h9. rewrite <- (S_compat x). assumption. rewrite (minus_Sn_m _ _ h10). do 2 rewrite <- pred_of_minus. f_equal. unfold lt in h4. rewrite (minus_Sn_m _ _ h4). auto with arith. rewrite <- h9 in h1. assert (h13 : S x < n). unfold lt in h10'. pose proof (le_lt_n_Sm _ _ h10') as h14. rewrite <- pred_of_minus in h14. rewrite <- (S_pred _ _ h4) in h14. assumption. specialize (h1 (fun m' : {m : nat | m < n} => f (exist (fun m : nat => m < S n) (proj1_sig m') (lt_S (proj1_sig m') n (proj2_sig m')))) def h13). rewrite h1. f_equal. f_equal. apply proof_irrelevance. subst. destruct (zerop (n-1)) as [h8 | h9]. rewrite h8 in h7. rewrite <- (minus_n_O n) in h7. rewrite h8 in h7. contradict h7. auto with arith. rewrite <- pred_of_minus in h9. rewrite <- (pred_of_minus n) in h7. rewrite (n_minus_pred_n _ h9) in h7. assert (h10:1 - 1 = 0). auto with arith. rewrite h10 in h7. contradict h7. auto with arith. unfold lt in h12. rewrite <- pred_of_minus in h12. rewrite <- (S_pred _ _ h4) in h12. clear h5 h7. omega. Qed. Lemma l_exp_seg_to_nprod_in_l_bij : forall {T:Type} (l:list T) (n:nat), (bijective (l_exp_seg_to_nprod_in_l l n)). intros T l n. induction n as [|n h1]. (* 0 *) red. split. red. intros x1 x2 h1. simpl in h1. apply functional_extensionality. intro x. destruct x. pose proof (lt_n_0 x). contradiction. red. intro y. exists (fun x:{m:nat | m < 0} => False_rect {x:T | In x l} (lt_n_0 _ (proj2_sig x))). simpl. simpl in y. destruct y. reflexivity. (* S n *) red. split. (* injective *) red. intros f1 f2 h2. simpl in h2. inversion h2 as [h0]. apply functional_extensionality. intro x. destruct x as [x h3]. destruct (lt_eq_lt_dec x n) as [[h4 | h5] | h6]. pose proof (f_equal (@nprod_to_list {x:T|In x l} n) H) as h5. pose proof (f_equal (fun l => nth (n-x-1) l (f1 (exist _ x h3))) h5) as h6. simpl in h6. do 2 rewrite (l_exp_seg_to_nprod_in_l_compat _ _ _ _ _ h4) in h6. simpl in h6. pose proof (proof_irrelevance _ h3 (lt_S _ _ h4)) as h7; subst; auto. subst. pose proof (proof_irrelevance _ h3 (lt_n_Sn n)); subst; auto. omega. (* surjective *) red. simpl. intro y. red in h1. destruct h1 as [h1l h1r]. red in h1r. destruct y as [fy sy]. specialize (h1r sy). destruct h1r as [f h2]. pose (fun m':{m:nat | m <= n} => let pf := (le_lt_eq_dec (proj1_sig m') n (proj2_sig m')) in match pf with | left pf' => (f (exist (fun m=>m fy end) as f'. pose (fun m':{m:nat | m < S n} => f' (exist _ _ (le_S_n _ _ (proj2_sig m')))) as f''. exists f''. apply injective_projections. simpl. unfold f''. simpl. unfold f'. simpl. destruct le_lt_eq_dec as [h3 | h4]. contradict h3. auto with arith. reflexivity. simpl. unfold f''. simpl. unfold f'. simpl. assert (h3:(fun m' : {m : nat | m < n} => match le_lt_eq_dec (proj1_sig m') n (le_S_n (proj1_sig m') n (lt_S (proj1_sig m') n (proj2_sig m'))) with | left pf' => f (exist (fun m : nat => m < n) (proj1_sig m') pf') | right _ => fy end) = f). apply functional_extensionality. intro x. destruct x as [x h3]. simpl. destruct le_lt_eq_dec as [h4 | h5]. pose proof (proof_irrelevance _ h3 h4); subst; auto. omega. rewrite h3. exact h2. Qed. Lemma finite_l_exp_seg : forall {T:Type} (l:list T) (n:nat), FiniteT ({m:nat | m < n} -> {x:T|In x l}). intros T l n. pose proof (finite_nat_initial_segment n) as h1. pose proof (list_to_set_finitet l) as h2. apply finite_exp; auto. Qed. Lemma finite_nprod_in_l : forall {T:Type} (l:list T) (n:nat), FiniteT ({x:T|In x l}^n). intros T l n. pose proof (finite_nat_initial_segment n) as h1. pose proof (finite_list_sig l) as h2. pose proof (finite_exp _ _ h1 h2) as h3. pose proof (l_exp_seg_to_nprod_in_l_bij l n) as h4. apply bijective_impl_invertible in h4. apply (bij_finite _ _ _ h3 h4). Qed. End NprodList. Section ListSort. Fixpoint maxlR {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) : l <> nil -> T := match l with | nil => fun pf => False_rect T (pf eq_refl) | a::l' => fun pf => match (nil_dec' l') with | left _ => a | right pf' => maxR R pftso a (maxlR R pftso l' pf') end end. Lemma in_maxlR : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (pf:l <> nil), In (maxlR R pftso l pf) l. intros T R h1 l. induction l as [|a l ih]; simpl; auto. intros h2. destruct (nil_dec' l) as [| h3]; subst. left; auto. specialize (ih h3). unfold maxR. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. right; auto. left; auto. Qed. Lemma maxlR_compat : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (pftr:Transitive _ R) (l:list T) (x:T) (pfin:In x l), let pfn := in_not_nil _ _ pfin in Rw R x (maxlR R pftso l pfn). intros T R h1 ht l. induction l as [|a l hi]; simpl. intros; contradiction. intros n h2. red. destruct (nil_dec' l) as [h3 | h3]; subst. right; destruct h2; try contradiction; auto. destruct h2 as [h2l | h2r]. subst. unfold maxR. lr_if_goal. fold hlr. destruct hlr as [hl | hr]. destruct hl; [left; auto | right; auto]. right; auto. unfold maxR. lr_if_goal. fold hlr. destruct hlr as [hl | hr]. specialize (hi _ h2r). simpl in hi. pose proof (proof_irrelevance _ h3 (in_not_nil _ _ h2r)); subst; auto. specialize (hi _ h2r). simpl in hi. red in hi. destruct hi as [hi | hi]. left. red in ht. eapply ht. apply hi. pose proof (proof_irrelevance _ h3 (in_not_nil l n h2r)); subst; auto. rewrite hi. left. pose proof (proof_irrelevance _ h3 (in_not_nil l n h2r)). clear hi. subst; auto. Qed. Lemma maxlR_functional : forall {T:Type} (R:Relation T) (pft:tso_dec R) (l l':list T) (pfl:l <> nil) (pfl':l' <> nil), l = l' -> maxlR R pft l pfl = maxlR R pft l' pfl'. intros T R h1 l l' h2 h3 h4; subst. pose proof (proof_irrelevance _ h2 h3); subst; auto. Qed. Lemma maxlR_functional' : forall {T:Type} (R R':Relation T) (pfso:sort_dec R) (pfso':sort_dec R') (pft:tso_dec R) (l l':list T) (pfl:l <> nil) (pfl':l' <> nil), R = R' -> l = l' -> maxlR R pft l pfl = maxlR R' (sort_dec_tso_dec _ pfso') l' pfl'. intros T R R' h1 h1' ht l l' h2 h2' h3 ?. subst. pose proof (proof_irrelevance _ h2 h2'); subst. f_equal. symmetry. apply sort_dec_tso_dec_eq. Qed. Lemma maxlR_eq_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (m:T) (pfn: l <> nil), let pftso := sort_dec_tso_dec _ pfso in (maxlR R pftso l pfn = m) <-> (In m l /\ forall x:T, In x l -> Rw R x m). intros T R h1 l m h3 htso. split; intro h4. subst. split. apply in_maxlR. intros x h4. erewrite maxlR_functional. apply maxlR_compat; auto. apply sort_dec_trans; auto. reflexivity. destruct h4 as [h4 h5]. specialize (h5 (maxlR R htso l h3) (in_maxlR _ _ _ _)). pose proof (maxlR_compat R htso (sort_dec_trans _ h1) l m h4) as h6. simpl in h6. pose proof (proof_irrelevance _ h3 (in_not_nil l m h4)); subst. apply (antisym_Rw R (sort_dec_so _ h1) _ _ h5 h6). Grab Existential Variables. assumption. Qed. Lemma maxlR_lrep : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (x:T) (n:nat), let pftso := sort_dec_tso_dec _ pfso in maxlR R pftso (lrep x (S n)) (cons_neq_nil _ _) = x. intros T R hso x n. revert x. induction n as [|n ih]; simpl. destruct (nil_dec' nil) as [|h3]; auto. contradict h3; auto. intro x. destruct (nil_dec' (x::lrep x n)) as [h3 | h3]; auto. specialize (ih x). erewrite <- ih at 3. simpl. destruct (nil_dec' (lrep x n)) as [h4 | h4]; auto. apply maxR_idem. unfold maxR. destruct (sort_dec_tso_dec R hso x (maxlR R (sort_dec_tso_dec R hso) (lrep x n) h4)) as [h5 | h5]. destruct h5 as [h5 | h5]. destruct (sort_dec_tso_dec R hso x x) as [h6 | h6]; auto. destruct (sort_dec_tso_dec R hso x (maxlR R (sort_dec_tso_dec R hso) (lrep x n) h4)) as [h7 | h7]; auto. pose proof (sort_dec_trans _ hso _ _ _ h5 h7) as h8. apply (sort_dec_irrefl _ hso) in h8; contradiction. apply (sort_dec_irrefl _ hso) in h6; contradiction. destruct (sort_dec_tso_dec R hso x (maxlR R (sort_dec_tso_dec R hso) (lrep x n) h4)) as [h6 | h6]; auto. destruct (sort_dec_tso_dec R hso x x); auto. Qed. Lemma maxlR_sing : forall {T:Type} (R:Relation T) (pftso:sort_dec R) (l:list T) (x:T) (pfe:Singleton x = list_to_set l), let pft := sort_dec_tso_dec _ pftso in let pf := inh_sing x in let pfn := P_subst pf pfe in let pfn' := inh_not_nil _ pfn in x = maxlR R pft l pfn'. intros T R h1 l x h2; simpl. pose proof h2 as h2'. symmetry in h2'. apply impl_all_sing in h2'. apply all_sing_lrep in h2'. symmetry. gen0. rewrite h2'. intro h3. erewrite maxlR_functional. erewrite <- maxlR_lrep. apply maxlR_functional. reflexivity. simpl. rewrite lrep_nil_iff in h3. rewrite neq0_iff in h3. rewrite (S_pred _ _ h3). simpl. reflexivity. Grab Existential Variables. simpl. apply cons_neq_nil. Qed. Lemma maxlR_list_singularize : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (pfn:l <> nil), let pftso := sort_dec_tso_dec _ pfso in let pfdt := sort_dec_eq_dec _ pfso in let pfn' := iffn2 (list_singularize_nil_iff pfdt l) pfn in maxlR R pftso l pfn = maxlR R pftso (list_singularize pfdt l) pfn'. intros T R h1 l h2 htso hdt hn. unfold htso. rewrite maxlR_eq_iff. split. rewrite list_singularize_in_iff. apply in_maxlR. intros x h3. rewrite (list_singularize_in_iff hdt) in h3. pose proof (maxlR_compat R (sort_dec_tso_dec R h1) (sort_dec_trans _ h1) (list_singularize hdt l) x h3) as h4. simpl in h4. pose proof (proof_irrelevance _ hn (in_not_nil _ _ h3)) as h5. rewrite h5; auto. Qed. Fixpoint minlR {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) : l <> nil -> T := match l with | nil => fun pf => False_rect T (pf eq_refl) | a::l' => fun pf => match (nil_dec' l') with | left _ => a | right pf' => minR R pftso a (minlR R pftso l' pf') end end. Lemma in_minlR : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (pf:l <> nil), In (minlR R pftso l pf) l. intros T R h1 l. induction l as [|a l ih]; simpl; auto. intros h2. destruct (nil_dec' l) as [| h3]; subst. left; auto. specialize (ih h3). unfold minR. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. left; auto. right; auto. Qed. Lemma minlR_compat : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (pftr:Transitive _ R) (l:list T) (x:T) (pfin:In x l), let pfn := in_not_nil _ _ pfin in Rw R (minlR R pftso l pfn) x. intros T R h0 h1 l. induction l as [|a l ih]; simpl. intros; contradiction. intros n h2. red. destruct (nil_dec' l) as [h3 | h3]; subst. right; destruct h2; try contradiction; auto. destruct h2 as [h2l | h2r]. subst. unfold minR. lr_if_goal. fold hlr. destruct hlr as [hl | hr]. right; auto. left; auto. unfold minR. lr_if_goal. fold hlr. destruct hlr as [hl | hr]. destruct hl as [hl | hl]. specialize (ih _ h2r). simpl in ih. red in ih. destruct ih as [ih | ih]. left. red in h1. eapply h1. apply hl. pose proof (proof_irrelevance _ h3 (in_not_nil l n h2r)); subst; auto. left. pose proof (proof_irrelevance _ h3 (in_not_nil l n h2r)) as h4. rewrite <- h4 in ih. clear h4. subst; auto. subst. simpl in ih. specialize (ih _ h2r). pose proof (proof_irrelevance _ h3 (in_not_nil l n h2r)); subst. red in ih; auto. red in ih. specialize (ih _ h2r). pose proof (proof_irrelevance _ h3 (in_not_nil l n h2r)); subst; auto. Qed. Lemma minlR_functional : forall {T:Type} (R:Relation T) (pft:tso_dec R) (l l':list T) (pfl:l <> nil) (pfl':l' <> nil), l = l' -> minlR R pft l pfl = minlR R pft l' pfl'. intros T R h1 l l' h2 h3 h4; subst. pose proof (proof_irrelevance _ h2 h3); subst; auto. Qed. Lemma minlR_functional' : forall {T:Type} (R R':Relation T) (pfso:sort_dec R) (pfso':sort_dec R') (pft:tso_dec R) (l l':list T) (pfl:l <> nil) (pfl':l' <> nil), R = R' -> l = l' -> minlR R pft l pfl = minlR R' (sort_dec_tso_dec _ pfso') l' pfl'. intros T R R' h1 h1' ht l l' h2 h2' h3 ?. subst. pose proof (proof_irrelevance _ h2 h2'); subst. f_equal. symmetry. apply sort_dec_tso_dec_eq. Qed. Lemma minlR_eq_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (m:T) (pfn: l <> nil), let pftso := sort_dec_tso_dec _ pfso in (minlR R pftso l pfn = m) <-> (In m l /\ forall x:T, In x l -> Rw R m x). intros T R h1 l m h3 htso. split; intro h4; subst. split. apply in_minlR. intros x h4. erewrite minlR_functional. apply minlR_compat; auto. apply sort_dec_trans; auto. reflexivity. destruct h4 as [h4 h5]. specialize (h5 (minlR R htso l h3) (in_minlR _ _ _ _)). pose proof (minlR_compat R htso (sort_dec_trans _ h1) l m h4) as h6. simpl in h6. pose proof (proof_irrelevance _ h3 (in_not_nil l m h4)); subst. symmetry. apply (antisym_Rw R (sort_dec_so _ h1) _ _ h5 h6). Grab Existential Variables. assumption. Qed. Lemma minlR_lrep : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (x:T) (n:nat), let pftso := sort_dec_tso_dec _ pfso in minlR R pftso (lrep x (S n)) (cons_neq_nil _ _) = x. intros T R hso x n. revert x. induction n as [|n ih]; simpl. destruct (nil_dec' nil) as [|h3]; auto. contradict h3; auto. intro x. destruct (nil_dec' (x::lrep x n)) as [h3 | h3]; auto. specialize (ih x). erewrite <- ih at 3. simpl. destruct (nil_dec' (lrep x n)) as [h4 | h4]; auto. apply minR_idem. unfold minR. destruct (sort_dec_tso_dec R hso x (minlR R (sort_dec_tso_dec R hso) (lrep x n) h4)) as [h5 | h5]. destruct (sort_dec_tso_dec R hso x x) as [h6 | h6]; auto. destruct (sort_dec_tso_dec R hso x (minlR R (sort_dec_tso_dec R hso) (lrep x n) h4)) as [h6 | h6]; auto. destruct h6 as [h6 | h6]. pose proof (sort_dec_trans _ hso _ _ _ h5 h6) as h7. apply (sort_dec_irrefl _ hso) in h7. contradiction. assumption. Qed. Lemma minlR_sing : forall {T:Type} (R:Relation T) (pftso:sort_dec R) (l:list T) (x:T) (pfe:Singleton x = list_to_set l), let pft := sort_dec_tso_dec _ pftso in let pf := inh_sing x in let pfn := P_subst pf pfe in let pfn' := inh_not_nil _ pfn in x = minlR R pft l pfn'. intros T R h1 l x h2; simpl. pose proof h2 as h2'. symmetry in h2'. apply impl_all_sing in h2'. apply all_sing_lrep in h2'. symmetry. gen0. rewrite h2'. intro h3. erewrite minlR_functional. erewrite <- minlR_lrep. apply minlR_functional. reflexivity. simpl. rewrite lrep_nil_iff in h3. rewrite neq0_iff in h3. rewrite (S_pred _ _ h3). simpl. reflexivity. Grab Existential Variables. simpl. apply cons_neq_nil. Qed. Lemma minlR_list_singularize : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (pfn:l <> nil), let pftso := sort_dec_tso_dec _ pfso in let pfdt := sort_dec_eq_dec _ pfso in let pfn' := iffn2 (list_singularize_nil_iff pfdt l) pfn in minlR R pftso l pfn = minlR R pftso (list_singularize pfdt l) pfn'. intros T R h1 l h2 htso hdt hn. unfold htso. rewrite minlR_eq_iff. split. rewrite list_singularize_in_iff. apply in_minlR. intros x h3. rewrite (list_singularize_in_iff hdt) in h3. pose proof (minlR_compat R (sort_dec_tso_dec R h1) (sort_dec_trans _ h1) (list_singularize hdt l) x h3) as h4. simpl in h4. pose proof (proof_irrelevance _ hn (in_not_nil _ _ h3)) as h5. rewrite h5; auto. Qed. Lemma minlR_not_lt: forall {T:Type} (R:Relation T) (pftso:sort_dec R) (l:list T) (pf:l <> nil) (x:T), In x l -> let pftso := sort_dec_tso_dec _ pftso in ~ R x (minlR R pftso l pf). intros T R h1 l. induction l as [|a l ih]. intros; contradiction. intros h2 x h3 hts h4. destruct h3 as [|h3]; subst. destruct (nil_dec' l) as [| h5]; subst. simpl in h4. lr_match h4; fold hlr in h4; destruct hlr as [hl | hr]. apply (sort_dec_irrefl _ h1) in h4; auto. pose proof (hr eq_refl); auto. simpl in h4. lr_match h4; fold hlr in h4; destruct hlr as [hl | hr]. apply (sort_dec_irrefl _ h1) in h4; auto. unfold minR in h4. lr_if h4; fold hlr in h4; destruct hlr as [hl' | hr']. apply (sort_dec_irrefl _ h1) in h4; auto. pose proof (sort_dec_trans _ h1 _ _ _ hr' h4) as h6. apply (sort_dec_irrefl _ h1) in h6; auto. simpl in h4. lr_match h4; fold hlr in h4; destruct hlr as [hl | hr]. subst. contradict h3. specialize (ih hr x h3). simpl in ih. unfold hts in h4. apply lt_minR in h4. destruct h4; contradiction. apply sort_dec_trans; auto. Qed. Lemma minlR_remove : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T) (pfn:l <> nil), In a l -> let pfdt := sort_dec_eq_dec R pfso in forall (pfn':remove pfdt a l <> nil), let pftso := sort_dec_tso_dec R pfso in minR R pftso a (minlR R pftso (remove pfdt a l) pfn') = minlR R (sort_dec_tso_dec R pfso) l pfn. intros T R h1 l a h2 hnin h3 h4 h5. unfold minR. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. destruct hl as [hl | hl]. pose proof (minlR_compat R h5 (sort_dec_trans _ h1) (remove h3 a l)) as h6. simpl in h6. symmetry. rewrite minlR_eq_iff. split; auto. intros x h7. destruct (sort_dec_eq_dec _ h1 x a) as [| h8]; subst. right; auto. specialize (h6 x). hfirst h6. rewrite in_remove_iff. split; auto. specialize (h6 hf). left. eapply trans_R_Rw. apply sort_dec_trans; auto. apply hl. erewrite minlR_functional. apply h6; auto. reflexivity. pose proof (in_minlR R h5 (remove h3 a l) h4) as h6. rewrite <- hl in h6. apply remove_In in h6. contradiction. symmetry. rewrite minlR_eq_iff. split. pose proof (in_minlR R h5 (remove h3 a l) h4) as h6. rewrite in_remove_iff in h6. destruct h6; auto. intros x h6. pose proof (minlR_compat R h5 (sort_dec_trans _ h1) (remove h3 a l)) as h8. destruct (sort_dec_eq_dec _ h1 x a) as [|h10]; subst. left; auto. specialize (h8 x). hfirst h8. rewrite in_remove_iff. split; auto. specialize (h8 hf); simpl in h8. erewrite minlR_functional. apply h8; auto. reflexivity. Qed. Lemma minlR_list_to_set_eq : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l l':list T) (pf:l <> nil) (pf':l' <> nil), list_to_set l = list_to_set l' -> let pfdt := sort_dec_tso_dec _ pfso in let pftso := sort_dec_tso_dec _ pfso in minlR R pftso l pf = minlR R pftso l' pf'. intros T R h1 l. induction l as [|a l ih]; simpl. intros ? h3. contradict h3; auto. intros l' h3 h4 h5. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. subst. simpl in h5. rewrite add_empty_sing in h5. simpl in ih. symmetry in h5. apply impl_all_sing in h5. apply all_sing_lrep in h5. symmetry. gen0. rewrite h5. intro h6. pose proof h6 as h6'. rewrite lrep_nil_iff in h6'. rewrite neq0_iff in h6'. gen0. rewrite (S_pred _ _ h6'). intro h7. erewrite (minlR_functional R (sort_dec_tso_dec _ h1) (lrep a (S (pred (length l')))) (lrep a (S (pred (length l'))))). rewrite minlR_lrep at 1; auto. reflexivity. destruct (in_dec (sort_dec_eq_dec _ h1) a l) as [h6 | h6]. rewrite in_add_eq in h5. pose proof (f_equal (fun S => Subtract S a) h5) as h7. simpl in h7. specialize (ih l' hr h4 h5). simpl in ih. rewrite <- ih. unfold minR. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. destruct hl' as [h8|]; auto. apply minlR_not_lt in h8; auto. contradiction. reflexivity. rewrite <- list_to_set_in_iff; auto. pose proof (f_equal (fun S => Subtract S a) h5) as h7. simpl in h7. rewrite sub_add_same_nin in h7. rewrite (subtract_remove_compat (sort_dec_eq_dec _ h1)) in h7. destruct (nil_dec' (remove (sort_dec_eq_dec _ h1) a l')) as [h8 | h8]. rewrite remove_eq_nil_iff in h8. destruct h8 as [|h8]; try contradiction. contradict hr. apply all_sing_lrep in h8. rewrite h8 in h7. rewrite remove_lrep in h7. simpl in h7. apply empty_set_nil in h7; auto. specialize (ih _ hr h8 h7). simpl in ih. rewrite ih. erewrite minlR_functional. apply minlR_remove. rewrite list_to_set_in_iff. rewrite <- h5. right; constructor; auto. reflexivity. rewrite <- list_to_set_in_iff; auto. Grab Existential Variables. assumption. Qed. Lemma minlR_incl : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l l':list T) (pf:l <> nil) (pf':l' <> nil), incl l l' -> let pftso := sort_dec_tso_dec _ pfso in Rw R (minlR R pftso l' pf') (minlR R pftso l pf). intros T R h1 l l' h2 h3 h4 h6. pose proof (in_minlR R h6 l h2) as h8. apply h4 in h8. pose proof (minlR_compat R (sort_dec_tso_dec _ h1) (sort_dec_trans _ h1) l' _ h8) as h9. simpl in h9. pose proof (proof_irrelevance _ h3 (in_not_nil l' _ h8)); subst. unfold h6. unfold h6 in h9. assumption. Qed. Lemma lt_minlR_remove : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (pfn:l <> nil), let pfdt := sort_dec_eq_dec _ pfso in let pftso := sort_dec_tso_dec _ pfso in let l' := remove pfdt (minlR R pftso l pfn) l in forall (pfn':l' <> nil), R (minlR R pftso l pfn) (minlR R pftso l' pfn'). intros T R h1 l h2 h3 h4 l' h5. pose proof (in_minlR R h4 l' h5) as h7. unfold l' in h7. apply incl_remove in h7. pose proof (minlR_compat R h4 (sort_dec_trans _ h1) l _ h7) as h8. simpl in h8. pose proof (proof_irrelevance _ h2 (in_not_nil _ _ h7)) as h9. rewrite h9. red in h8. destruct h8 as [| h8]; auto. rewrite <- h9 in h8. pose proof (in_minlR R h4 (remove h3 (minlR R h4 l h2) l) h5) as h10. rewrite <- h8 in h10. apply remove_In in h10. contradiction. Qed. Fixpoint le_listR {T:Type} (R:Relation T) (pf:tso_dec R) (l m:list T) : Prop := match l with | nil => True | a::l' => match m with | nil => False | b::m' => match (pf a b) with | inleft (left _) => True | inleft (right _) => le_listR R pf l' m' | inright P => False end end end. Definition lt_listR {T:Type} (R:Relation T) (pf:tso_dec R) (l m:list T) : Prop := le_listR R pf l m /\ l <> m. Lemma lt_list_impl_leR : forall {T:Type} (R:Relation T) (pf:tso_dec R) (l m:list T), lt_listR R pf l m -> le_listR R pf l m. intros T R h1 l m h2. red in h2. destruct h2; auto. Qed. Lemma not_lt_listR_nil : forall {T:Type} (R:Relation T) (pf:tso_dec R) (l:list T), ~ lt_listR R pf l nil. intros T R h1 l h2. destruct h2 as [h2 h3]. rewrite (hd_compat' _ h3) in h2. simpl in h2; auto. Qed. Lemma le_listR_nil : forall {T:Type} (R:Relation T) (pf:tso_dec R) (l:list T), le_listR R pf nil l. intros; simpl; auto. Qed. Lemma le_listR_dec : forall {T:Type} (R:Relation T) (pfdt:tso_dec R) (l m:list T), {le_listR R pfdt l m} + {~le_listR R pfdt l m}. intros T R h0 l. induction l as [|a l h1]. intro m. left. simpl. auto. intros m. simpl. destruct m as [|b m]. right. intro; contradiction. destruct (h0 a b) as [h3 | h3]. destruct h3 as [h3 | h3]. left. auto. subst. apply h1. right. intro; contradiction. Qed. Lemma refl_le_listR : forall {T:Type} (R:Relation T) (pft:tso_dec R) (pfi:irrefl R) (l:list T), le_listR R pft l l. intros T R ht hi l. induction l as [|a l h1]; simpl. auto. destruct (ht a a) as [h2 | h2]. destruct h2 as [h2 | h2]; auto. apply hi in h2; auto. Qed. Lemma antisym_le_listR : forall {T:Type} (R:Relation T) (pfsd:sort_dec R) (l m:list T), let pft := sort_dec_tso_dec _ pfsd in le_listR R pft l m -> le_listR R pft m l -> l = m. intros T R h0 l. destruct h0 as [hts htr hir]. induction l as [|a l h1]; simpl. intro m. destruct m; auto. intros ? h1. simpl in h1. contradiction. intros m h2. destruct m as [|b m]. contradiction. unfold sort_dec_tso_dec in h2. destruct (hts a b) as [h3 | h3]. destruct h3 as [h3 | h3]. intro h4. simpl in h4. destruct (hts b a) as [h5 | h5]. destruct h5 as [h5|]; subst. pose proof (htr _ _ _ h3 h5) as h6. apply hir in h6; contradiction. apply hir in h3; contradiction. contradiction. subst. intro h3. simpl in h3. destruct (hts b b) as [h4 | h4]. destruct h4 as [h4 | h4]. apply hir in h4. contradiction. specialize (h1 _ h2 h3). subst. reflexivity. contradiction. contradiction. Qed. Lemma antisym_not_le_listR : forall {T:Type} (R:Relation T) (pfdt:sort_dec R) (l m:list T), let pft := sort_dec_tso_dec _ pfdt in ~ le_listR R pft l m -> ~ le_listR R pft m l -> l = m. intros T R h0 l. destruct h0 as [hts htr hir]. induction l as [|a l h1]; simpl; try tauto. intros m h2 h3. destruct m as [|b m]. contradict h3. simpl. auto. destruct (hts a b) as [h4 | h4]. destruct h4 as [h4 | h4]. contradict h2. auto. subst. simpl in h3. destruct (hts b b) as [h4 | h4]. destruct h4 as [h4 | h4]. apply hir in h4. contradiction. specialize (h1 _ h2 h3). subst. reflexivity. apply hir in h4. contradiction. contradict h3. simpl. destruct (hts b a) as [h3 | h3]. destruct h3 as [h3 | h3]; auto. subst. apply hir in h4. contradiction. pose proof (htr _ _ _ h4 h3) as h5. apply hir in h5. contradiction. Qed. Lemma lt_listR_iff : forall {T:Type} (R:Relation T) (pf:sort_dec R) (l m:list T), let pfts := sort_dec_tso_dec _ pf in lt_listR R pfts l m <-> ~ le_listR R pfts m l. intros T R h0 l m. destruct h0 as [hts htr hir]. split; intro h1. red in h1. destruct h1 as [h1 h2]. intro h3. pose proof (antisym_le_listR R (sort_dec_intro _ hts htr hir) _ _ h1 h3). subst. contradict h2; auto. red; split. destruct (le_listR_dec R hts l m) as [h2 | h2]; auto. pose proof (antisym_not_le_listR R (sort_dec_intro _ hts htr hir) _ _ h1 h2). subst. apply refl_le_listR; auto. intro. subst. contradict h1. apply refl_le_listR; auto. Qed. Lemma lt_listR_cons : forall {T:Type} (R:Relation T) (pf:sort_dec R) (l m:list T) (a:T), let pfts := sort_dec_tso_dec _ pf in lt_listR R pfts (a::l) (a::m) -> lt_listR R pfts l m. intros T R h1 l m a hts h2. red in h2. destruct h2 as [h2 h3]. red in h2. destruct (hts a a) as [h4 | h4]. destruct h4 as [h4 | h5]. apply (sort_dec_irrefl _ h1) in h4. contradiction. red. split; auto. contradict h3. subst; auto. contradiction. Qed. Lemma lt_listR_cons_iff : forall {T:Type} (R:Relation T) (pf:sort_dec R) (l l':list T) (a:T) (pfn:l <> nil), let pfts := sort_dec_tso_dec _ pf in lt_listR R pfts l (a::l') <-> R (hd' l pfn) a \/ (hd' l pfn = a /\ lt_listR R pfts (tl l) l'). intros T R h1 l l' a h2; simpl; split; intro h3. rewrite (hd_compat' _ h2) in h3. red in h3. simpl in h3. destruct h3 as [h3 h4]. destruct (sort_dec_tso_dec R h1 (hd' l h2) a) as [h5 | h5]. destruct h5 as [h5 | h5]. left; auto. subst. right. split; auto. red. split; auto. intro; subst. contradict h4; auto. contradiction. destruct h3 as [h3 | h3]. red. split. rewrite (hd_compat' _ h2). simpl. destruct (sort_dec_tso_dec R h1 (hd' l h2) a) as [h5 | h5]. destruct h5 as [h5 | h5]; auto. subst. apply (sort_dec_irrefl _ h1) in h3. contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h3 h5) as h6. apply (sort_dec_irrefl _ h1) in h6; auto. intro; subst. simpl in h3. apply (sort_dec_irrefl _ h1) in h3; auto. destruct h3 as [? h3]; subst. rewrite (hd_compat' _ h2) at 1; simpl. red. split. simpl. destruct (sort_dec_tso_dec R h1 (hd' l h2) (hd' l h2)) as [h4 | h4]. destruct h4 as [h4 | h4]; auto. red in h3. destruct h3; auto. apply (sort_dec_irrefl _ h1) in h4; auto. red in h3. destruct h3 as [h3 h4]. contradict h4. inversion h4; auto. Qed. Lemma le_listR_dec' : forall {T:Type} (R:Relation T) (pf:sort_dec R) (l m:list T), let pfts := sort_dec_tso_dec _ pf in {le_listR R pfts l m} + {lt_listR R pfts m l}. intros T R h0 l m. destruct h0 as [hts htr hir]. destruct (le_listR_dec R hts l m) as [h1 | h1]. left; auto. rewrite <- (lt_listR_iff R (sort_dec_intro _ hts htr hir)) in h1 at 1. right; auto. Qed. Lemma tso_dec_lt_listR : forall {T:Type} (R:Relation T) (pf:sort_dec R), let pfts := sort_dec_tso_dec _ pf in tso_dec (lt_listR R pfts). intros T R h1. destruct h1 as [hts htr hir]. red; intros l m. destruct (le_listR_dec' R (sort_dec_intro _ hts htr hir) l m) as [h1 | h1]. destruct (le_listR_dec' R (sort_dec_intro _ hts htr hir) m l) as [h2 | h2]. pose proof (antisym_le_listR R (sort_dec_intro _ hts htr hir) _ _ h1 h2). subst; auto. left; left; auto. right; auto. Qed. Lemma trans_le_listR : forall {T:Type} (R:Relation T) (pf:sort_dec R), let pfts := sort_dec_tso_dec _ pf in Transitive _ (le_listR R pfts). intros T R h0. destruct h0 as [hts htr hir]. red. intro la. induction la as [|a la h1]; simpl. intros lb lc. destruct lb as [|b lb], lc as [|c lc]; auto. intros lb lc. destruct lb as [|b lb], lc as [|c lc]; auto. intro. contradiction. destruct (hts a b) as [h2 | h2]. destruct h2 as [h2 | h2]. destruct (hts a c) as [h3 | h3]. destruct h3 as [h3 | h3]. auto. subst. intros ? h3. simpl in h3. destruct (hts b c) as [h4 | h4]. destruct h4 as [h4 | h4]. pose proof (htr _ _ _ h4 h2) as ht. apply hir in ht; contradiction. subst. apply hir in h2; contradiction. contradiction. intros ? h4. simpl in h4. destruct (hts b c) as [h5 | h5]. destruct h5 as [h5 | h5]. pose proof (htr _ _ _ h2 h5) as h6. pose proof (htr _ _ _ h6 h3) as h7. apply hir in h7; contradiction. subst. pose proof (htr _ _ _ h3 h2) as h5. apply hir in h5; auto. assumption. subst. intros h2 h3. simpl in h3. destruct (hts b c) as [h4 | h4]. destruct h4 as [h4 | h4]. auto. subst. specialize (h1 _ _ h2 h3). assumption. contradiction. intro. contradiction. Qed. Lemma trans_le_lt_listR : forall {T:Type} (R:Relation T) (pf:sort_dec R) (la lb lc:list T), let pftso := sort_dec_tso_dec _ pf in le_listR R pftso la lb -> lt_listR R pftso lb lc -> lt_listR R pftso la lc. intros T R h1. destruct h1 as [hts htr hir]. intros la lb lc h0 h1 h2. red in h2. destruct h2 as [h2 h3]. pose proof (trans_le_listR R (sort_dec_intro _ hts htr hir) _ _ _ h1 h2) as h4. red. split; auto. intro h5. subst. pose proof (antisym_le_listR R (sort_dec_intro _ hts htr hir) _ _ h2 h1). subst. contradict h3; auto. Qed. Lemma trans_lt_le_listR : forall {T:Type} (R:Relation T) (pf:sort_dec R) (la lb lc:list T), let pftso := sort_dec_tso_dec _ pf in lt_listR R pftso la lb -> le_listR R pftso lb lc -> lt_listR R pftso la lc. intros T R h1. destruct h1 as [hts htr hir]. intros la lb lc h0 h1 h2. red in h1. destruct h1 as [h1 h3]. pose proof (trans_le_listR R (sort_dec_intro _ hts htr hir) _ _ _ h1 h2) as h4. split; auto. intro h5. subst. pose proof (antisym_le_listR R (sort_dec_intro _ hts htr hir) _ _ h2 h1). subst. contradict h3; auto. Qed. Lemma trans_lt_listR : forall {T:Type} (R:Relation T) (pf:sort_dec R), let pftso := sort_dec_tso_dec _ pf in Transitive _ (lt_listR R pftso). intros T R h1. destruct h1 as [hts htr hir]. red. intros la lb lc h1 h2. destruct h1 as [h1 h1'], h2 as [h2 h2']. split. apply (trans_le_listR R (sort_dec_intro _ hts htr hir) _ _ _ h1 h2). intro. subst. pose proof (antisym_le_listR R (sort_dec_intro _ hts htr hir) _ _ h2 h1). subst. contradict h1'; auto. Qed. Lemma irrefl_lt_listR : forall {T:Type} (R:Relation T) (pf:tso_dec R) (l:list T), ~ lt_listR R pf l l. intros T R h0 l h1. red in h1. destruct h1 as [? h1]. contradict h1; auto. Qed. Lemma so_lt_listR : forall {T:Type} (R:Relation T) (pf:sort_dec R), let pftso := sort_dec_tso_dec _ pf in Strict_Order (lt_listR R pftso). intros T R h1. destruct h1 as [hts htr hir]. constructor. apply (trans_lt_listR R (sort_dec_intro _ hts htr hir)). red. apply irrefl_lt_listR. Qed. Lemma le_listR_lt_eq_dec : forall {T:Type} (R:Relation T) (pf:sort_dec R) (l m:list T), let pftso := sort_dec_tso_dec _ pf in le_listR R pftso l m -> {lt_listR R pftso l m} + {l = m}. intros T R h0. destruct h0 as [hts htr hir]. intro l. induction l as [|a l h1]; destruct m as [|b m]; simpl. intros; right; auto. intros; left. red. split. apply le_listR_nil. intro h1. discriminate h1. intros; contradiction. intro h2. destruct (hts a b) as [h3 | h3]. destruct h3 as [h3|]; subst. left. red. simpl. destruct (hts a b) as [h4 | h4]. destruct h4 as [h4 |]; subst. split; auto. intro h5. inversion h5; subst. clear h5. apply hir in h4; auto. apply hir in h3. contradiction. pose proof (htr _ _ _ h3 h4) as h5. apply hir in h5; contradiction. specialize (h1 _ h2). destruct h1 as [h1 | h1]. left. red. split. simpl. destruct (hts b b) as [h3 | h3]. destruct h3 as [h3 | h3]. apply hir in h3; contradiction. assumption. apply hir in h3; auto. intro h3. inversion h3 as [h4]. clear h3. subst. apply irrefl_lt_listR in h1; auto. subst. right; auto. contradiction. Qed. Lemma lt_listR_tso_dec : forall {T:Type} (R:Relation T) (pf:sort_dec R) (l m:list T), let pftso := sort_dec_tso_dec _ pf in {lt_listR R pftso l m} + {l = m} + {lt_listR R pftso m l}. intros T R h0. destruct h0 as [hts htr hir]. intros l m. destruct (le_listR_dec' R (sort_dec_intro _ hts htr hir) l m) as [h0 | h0]. destruct (le_listR_lt_eq_dec R (sort_dec_intro _ hts htr hir) _ _ h0) as [h1 | h1]. left; auto. subst. left. right; auto. right; auto. Qed. Lemma lt_listR_sort_dec : forall {T:Type} (R:Relation T) (pf:sort_dec R), let pftso := sort_dec_tso_dec _ pf in sort_dec (lt_listR R pftso). intros T R h0. destruct h0 as [hts htr hir]. constructor. eapply tso_dec_lt_listR. apply trans_lt_listR. red; apply irrefl_lt_listR. Defined. Lemma lt_listR_sortable : forall {T:Type} (R:Relation T) (pf:sort_dec R), let pftso := sort_dec_tso_dec _ pf in sortable (lt_listR R pftso). intros T R h1; simpl. apply sort_dec_sortable. apply lt_listR_sort_dec. Qed. Lemma minlR_in : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (ll:faml T) (l:list T) (pfin:In l ll), let pft' := sort_dec_tso_dec _ pfso in let pftl := tso_dec_lt_listR _ pfso in minlR (lt_listR R pft') pftl ll (in_not_nil _ _ pfin) = minlR (lt_listR R pft') pftl (l::ll) (cons_neq_nil ll l). intros T R h1 ll l h2 h3 h4. pose proof (minlR_list_singularize (lt_listR R h3) (lt_listR_sort_dec _ _) ll (in_not_nil ll l h2)) as h5. simpl in h5. unfold h4, h3. unfold h3 in h5. pose proof (minlR_list_singularize (lt_listR R h3) (lt_listR_sort_dec _ _) (l::ll) (cons_neq_nil ll l)) as h5'. simpl in h5'. simpl. lr_match_goal; fold hlr; fold hlr in h5'; destruct hlr as [hl | hr]. clear h5'; subst; contradict h2. destruct h1 as [h1 h6 h7]. simpl. simpl in h5. rewrite h5. erewrite sort_dec_tso_dec_eq in h5'. assert (h8:h1 = h3). unfold h3. symmetry. apply sort_dec_tso_dec_eq. rewrite h8. rewrite h5' at 1. simpl; auto. apply minlR_functional. rewrite list_singularize_in_cons; auto. Qed. Lemma maxlR_in : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (ll:faml T) (l:list T) (pfin:In l ll), let pft' := sort_dec_tso_dec _ pfso in let pftl := tso_dec_lt_listR _ pfso in maxlR (lt_listR R pft') pftl ll (in_not_nil _ _ pfin) = maxlR (lt_listR R pft') pftl (l::ll) (cons_neq_nil ll l). intros T R h1 ll l h2 h3 h4. pose proof (maxlR_list_singularize (lt_listR R h3) (lt_listR_sort_dec _ _) ll (in_not_nil ll l h2)) as h5. simpl in h5. unfold h4, h3. unfold h3 in h5. pose proof (maxlR_list_singularize (lt_listR R h3) (lt_listR_sort_dec _ _) (l::ll) (cons_neq_nil ll l)) as h5'. simpl in h5'. simpl. lr_match_goal; fold hlr; fold hlr in h5'; destruct hlr as [hl | hr]. clear h5'; subst; contradict h2. destruct h1 as [h1 h6 h7]. simpl. simpl in h5. rewrite h5. erewrite sort_dec_tso_dec_eq in h5'. assert (h8:h1 = h3). unfold h3. symmetry. apply sort_dec_tso_dec_eq. rewrite h8. rewrite h5' at 1. apply maxlR_functional. rewrite list_singularize_in_cons; auto. Qed. (*upper bound filter for list*) Fixpoint ub_list {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T) : list T := match l with | nil => nil | a :: l' => match (pftso a b) with | inleft (left _) => a::(ub_list R pftso l' b) | inleft (right _) => ub_list R pftso l' b | inright _ => ub_list R pftso l' b end end. (*weak upper bound filter for list*) Fixpoint wub_list {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T) : list T := match l with | nil => nil | a :: l' => match (pftso a b) with | inleft _ => a::(wub_list R pftso l' b) | inright _ => wub_list R pftso l' b end end. (*lower bound filter for list.*) Fixpoint lb_list {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T) : list T := match l with | nil => nil | a :: l' => match (pftso a b) with | inleft _ => lb_list R pftso l' b | inright _ => a::lb_list R pftso l' b end end. (*weak lower bound filter for list.*) Fixpoint wlb_list {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T) : list T := match l with | nil => nil | a :: l' => match (pftso a b) with | inleft (left _) => wlb_list R pftso l' b | inleft (right _) => a::(wlb_list R pftso l' b) | inright _ => a::wlb_list R pftso l' b end end. Lemma ub_list_nil : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (b:T), ub_list R pftso nil b = nil. intros; simpl; auto. Qed. Lemma wub_list_nil : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (b:T), wub_list R pftso nil b = nil. intros; simpl; auto. Qed. Lemma lb_list_nil : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (b:T), lb_list R pftso nil b = nil. intros; simpl; auto. Qed. Lemma wlb_list_nil : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (b:T), wlb_list R pftso nil b = nil. intros; simpl; auto. Qed. Lemma ub_list_cons_same : forall {T:Type} (R:Relation T) (pfir:irrefl R) (pftso:tso_dec R) (l:list T) (b:T), ub_list R pftso (b::l) b = ub_list R pftso l b. intros T R hir h1 l b. simpl. destruct (h1 b b) as [h2 | h2]. destruct h2 as [h2 | h2]; auto. apply hir in h2. contradiction. reflexivity. Qed. Lemma wub_list_cons_same : forall {T:Type} (R:Relation T) (pfir:irrefl R) (pftso:tso_dec R) (l:list T) (b:T), wub_list R pftso (b::l) b = b :: wub_list R pftso l b. intros T R hir h1 l b; simpl. destruct (h1 b b) as [h2 | h2]. destruct h2 as [h2 | h2]. apply hir in h2; contradiction. reflexivity. apply hir in h2; contradiction. Qed. Lemma lb_list_cons_same : forall {T:Type} (R:Relation T) (pfir:irrefl R) (pftso:tso_dec R) (l:list T) (b:T), lb_list R pftso (b::l) b = lb_list R pftso l b. intros T R hir h1 l b. simpl. destruct (h1 b b) as [h2 | h2]. destruct h2 as [h2 | h2]; auto. apply hir in h2. contradiction. Qed. Lemma wlb_list_cons_same : forall {T:Type} (R:Relation T) (pfir:irrefl R) (pftso:tso_dec R) (l:list T) (b:T), wlb_list R pftso (b::l) b = b :: wlb_list R pftso l b. intros T R hir h1 l b; simpl. destruct (h1 b b) as [h2 | h2]. destruct h2 as [h2 | h2]. apply hir in h2; contradiction. reflexivity. apply hir in h2; contradiction. Qed. Lemma ub_list_cons_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), R x b -> let pftso := sort_dec_tso_dec _ pfso in ub_list R pftso (b::l) x = ub_list R pftso l x. intros T R hso l b x h0 htso. simpl. destruct (htso b x) as [h2 | h2]; subst; auto. destruct h2 as [h2|]; subst. pose proof (sort_dec_trans _ hso _ _ _ h0 h2) as h4. apply (sort_dec_irrefl _ hso) in h4; contradiction. apply (sort_dec_irrefl _ hso) in h0; contradiction. Qed. Lemma wub_list_cons_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), R x b -> let pftso := sort_dec_tso_dec _ pfso in wub_list R pftso (b::l) x = wub_list R pftso l x. intros T R hso l b x h0 htso. simpl. destruct (htso b x) as [h2 | h2]; subst; auto. destruct h2 as [h2|]; subst. pose proof (sort_dec_trans _ hso _ _ _ h0 h2) as h4. apply (sort_dec_irrefl _ hso) in h4; contradiction. apply (sort_dec_irrefl _ hso) in h0; contradiction. Qed. Lemma lb_list_cons_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), R x b -> let pftso := sort_dec_tso_dec _ pfso in lb_list R pftso (b::l) x = b :: lb_list R pftso l x. intros T R hso l b x h0 htso. simpl. destruct (htso b x) as [h2 | h2]; subst; auto. destruct h2 as [h2|]; subst. pose proof (sort_dec_trans _ hso _ _ _ h0 h2) as h4. apply (sort_dec_irrefl _ hso) in h4; contradiction. apply (sort_dec_irrefl _ hso) in h0; contradiction. Qed. Lemma wlb_list_cons_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), R x b -> let pftso := sort_dec_tso_dec _ pfso in wlb_list R pftso (b::l) x = b :: wlb_list R pftso l x. intros T R hso l b x h0 htso. simpl. destruct (htso b x) as [h2 | h2]; subst; auto. destruct h2 as [h2|]; subst. pose proof (sort_dec_trans _ hso _ _ _ h0 h2) as h4. apply (sort_dec_irrefl _ hso) in h4; contradiction. apply (sort_dec_irrefl _ hso) in h0; contradiction. Qed. Lemma ub_list_cons_gt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), R b x -> let pftso := sort_dec_tso_dec _ pfso in ub_list R pftso (b::l) x = b :: ub_list R pftso l x. intros T R hso l b x h0 htso. simpl. destruct (htso b x) as [h2 | h2]; subst; auto. destruct h2 as [h2|]; subst; auto. apply (sort_dec_irrefl _ hso) in h0; contradiction. pose proof (sort_dec_trans _ hso _ _ _ h0 h2) as h4. apply (sort_dec_irrefl _ hso) in h4; contradiction. Qed. Lemma wub_list_cons_gt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), R b x -> let pftso := sort_dec_tso_dec _ pfso in wub_list R pftso (b::l) x = b :: wub_list R pftso l x. intros T R hso l b x h0 htso. simpl. destruct (htso b x) as [h2 | h2]; subst; auto. pose proof (sort_dec_trans _ hso _ _ _ h0 h2) as h4. apply (sort_dec_irrefl _ hso) in h4; contradiction. Qed. Lemma lb_list_cons_gt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), R b x -> let pftso := sort_dec_tso_dec _ pfso in lb_list R pftso (b::l) x = lb_list R pftso l x. intros T R hso l b x h0 htso. simpl. destruct (htso b x) as [h2 | h2]; subst; auto. pose proof (sort_dec_trans _ hso _ _ _ h0 h2) as h4. apply (sort_dec_irrefl _ hso) in h4; contradiction. Qed. Lemma wlb_list_cons_gt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), R b x -> let pftso := sort_dec_tso_dec _ pfso in wlb_list R pftso (b::l) x = wlb_list R pftso l x. intros T R hso l b x h0 htso. simpl. destruct (htso b x) as [h2 | h2]; subst; auto. destruct h2 as [h2|]; subst; auto. apply (sort_dec_irrefl _ hso) in h0; contradiction. pose proof (sort_dec_trans _ hso _ _ _ h0 h2) as h4. apply (sort_dec_irrefl _ hso) in h4; contradiction. Qed. Lemma ub_list_trivial : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b:T), (forall x:T, In x l -> R x b) -> let pftso := sort_dec_tso_dec _ pfso in ub_list R pftso l b = l. intros T R h1 l. induction l as [|a l ih]; auto. intros b h2 htso. simpl. destruct (htso a b) as [h3 | h3]. destruct h3 as [h3 | h3]. f_equal; apply ih. intros; apply h2; right; auto. subst. specialize (h2 b (in_eq _ _)). apply (sort_dec_irrefl _ h1) in h2; contradiction. specialize (h2 a (in_eq _ _)). pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. Qed. Lemma wub_list_trivial : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b:T), (forall x:T, In x l -> Rw R x b) -> let pftso := sort_dec_tso_dec _ pfso in wub_list R pftso l b = l. intros T R h1 l. induction l as [|a l ih]; auto. intros b h2 htso. simpl. destruct (htso a b) as [h3 | h3]. destruct h3 as [h3 | h3]. f_equal; apply ih. intros; apply h2; right; auto. subst. f_equal; apply ih. intros; apply h2; right; auto. specialize (h2 a (in_eq _ _)). destruct h2 as [h2|]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. apply (sort_dec_irrefl _ h1) in h3; contradiction. Qed. Lemma lb_list_trivial : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b:T), (forall x:T, In x l -> R b x) -> let pftso := sort_dec_tso_dec _ pfso in lb_list R pftso l b = l. intros T R h1 l. induction l as [|a l ih]; auto. intros b h2 htso. simpl. destruct (htso a b) as [h3 | h3]. destruct h3 as [h3 | h3]. specialize (h2 a (in_eq _ _)). pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. subst. specialize (h2 b (in_eq _ _)). apply (sort_dec_irrefl _ h1) in h2; contradiction. f_equal; apply ih. intros; apply h2; right; auto. Qed. Lemma wlb_list_trivial : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b:T), (forall x:T, In x l -> Rw R b x) -> let pftso := sort_dec_tso_dec _ pfso in wlb_list R pftso l b = l. intros T R h1 l. induction l as [|a l ih]; auto. intros b h2 htso. simpl. destruct (htso a b) as [h3 | h3]. destruct h3 as [h3 | h3]. specialize (h2 a (in_eq _ _)). destruct h2 as [h2|]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. apply (sort_dec_irrefl _ h1) in h3; contradiction. subst. f_equal; apply ih. intros; apply h2; right; auto. f_equal; apply ih. intros; apply h2; right; auto. Qed. Lemma ub_list_app : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l l':list T) (b:T), let pftso := sort_dec_tso_dec _ pfso in ub_list R pftso (l ++ l') b = ub_list R pftso l b ++ ub_list R pftso l' b. intros T R h1 l. induction l as [|a l ih]; auto. intros l' b hts. simpl. destruct (hts a b) as [h2 | h2]. destruct h2 as [h2|]; subst; auto. rewrite <- app_comm_cons. f_equal; auto. auto. Qed. Lemma wub_list_app : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l l':list T) (b:T), let pftso := sort_dec_tso_dec _ pfso in wub_list R pftso (l ++ l') b = wub_list R pftso l b ++ wub_list R pftso l' b. intros T R h1 l. induction l as [|a l ih]; auto. intros l' b hts. simpl. destruct (hts a b) as [h2 | h2]. rewrite <- app_comm_cons. f_equal; auto. auto. Qed. Lemma lb_list_app : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l l':list T) (b:T), let pftso := sort_dec_tso_dec _ pfso in lb_list R pftso (l ++ l') b = lb_list R pftso l b ++ lb_list R pftso l' b. intros T R h1 l. induction l as [|a l ih]; auto. intros l' b hts. simpl. destruct (hts a b) as [h2 | h2]; auto. rewrite <- app_comm_cons. f_equal; auto. Qed. Lemma wlb_list_app : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l l':list T) (b:T), let pftso := sort_dec_tso_dec _ pfso in wlb_list R pftso (l ++ l') b = wlb_list R pftso l b ++ wlb_list R pftso l' b. intros T R h1 l. induction l as [|a l ih]; auto. intros l' b hts. simpl. destruct (hts a b) as [h2 | h2]. destruct h2 as [h2|]; subst; auto. rewrite <- app_comm_cons. f_equal; auto. rewrite <- app_comm_cons. f_equal; auto. Qed. Lemma in_ub_list_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), let pftso := sort_dec_tso_dec _ pfso in In x (ub_list R pftso l b) <-> In x l /\ R x b. intros T R h1 l. induction l as [|a l ih]; simpl. intros; tauto. intros b x. destruct (sort_dec_tso_dec R h1 a b) as [h2 | h2]. destruct h2 as [h2 | h2]. split; intro h3. split. destruct h3 as [|h3]. left; auto. right. rewrite ih in h3. destruct h3; auto. destruct h3 as [|h3]; subst; auto. rewrite ih in h3. destruct h3; auto. destruct h3 as [h3 h4]. destruct h3 as [|h3]; subst. left; auto. pose proof (conj h3 h4) as h5. rewrite <- ih in h5. right; auto. subst. split; intro h2. rewrite ih in h2. destruct h2 as [h2 h3]. split; auto. destruct h2 as [h2 h3]. destruct h2 as [|h2]; subst. apply (sort_dec_irrefl _ h1) in h3; contradiction. rewrite ih. split; auto. rewrite ih. split; intro h3. destruct h3 as [h3 h4]. split; auto. destruct h3 as [h3 h4]. destruct h3 as [|h3]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h2 h4) as h5. apply (sort_dec_irrefl _ h1) in h5. contradiction. split; auto. Qed. Lemma in_wub_list_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), let pftso := sort_dec_tso_dec _ pfso in In x (wub_list R pftso l b) <-> In x l /\ Rw R x b. intros T R h1 l. induction l as [|a l ih]; intros b x h2; split; intro h3. simpl in h3; contradiction. destruct h3; contradiction. fold h2 in ih. simpl in h3; simpl. destruct (h2 a b) as [h4 | h4]. destruct h4 as [h4|]; subst. destruct h3 as [|h3]; subst. split; left; auto. rewrite ih in h3. destruct h3 as [h3 h5]. split; [right|]; auto. destruct h3 as [|h3]; subst. split; [left | right]; auto. rewrite ih in h3. tauto. rewrite ih in h3. tauto. fold h2 in ih. destruct h3 as [h3 h4]. simpl. destruct h3 as [|h3]; subst. destruct (h2 x b) as [h5 | h5]. left; auto. pose proof (trans_Rw_R _ (sort_dec_trans _ h1) _ _ _ h4 h5) as h6. apply (sort_dec_irrefl _ h1) in h6. contradiction. destruct (h2 a b) as [h5 | h5]. destruct h5 as [h5|]; subst. pose proof (conj h3 h4) as h6. rewrite <- ih in h6. right; auto. pose proof (conj h3 h4) as h5. rewrite <- ih in h5. right; auto. pose proof (conj h3 h4) as h6. rewrite <- ih in h6; auto. Qed. Lemma in_lb_list_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), let pftso := sort_dec_tso_dec _ pfso in In x (lb_list R pftso l b) <-> In x l /\ R b x. intros T R h1 l; induction l as [|a l ih]; intros b x h2; split; intro h3. simpl in h3; contradiction. simpl in h3. destruct h3; contradiction. fold h2 in ih. simpl in h3. destruct (h2 a b) as [h4 | h4]. destruct h4 as [h4 |]; subst. rewrite ih in h3. destruct h3. split; try right; auto. rewrite ih in h3. destruct h3 as [h3 h4]. split; try right; auto. destruct h3 as [|h3]; subst. split; try left; auto. rewrite ih in h3. destruct h3; split; try right; auto. simpl. destruct h3 as [h3 h4]. destruct h3 as [|h3]; destruct (h2 x b) as [h5 | h5]; simpl; try destruct h5 as [h5 | h5]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h6. apply (sort_dec_irrefl _ h1) in h6; contradiction. apply (sort_dec_irrefl _ h1) in h4; contradiction. destruct (h2 x b) as [h6 | h6]. destruct h6 as [h6 |]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h5 h6) as h7. apply (sort_dec_irrefl _ h1) in h7; contradiction. apply (sort_dec_irrefl _ h1) in h5; contradiction. left; auto. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. apply (sort_dec_irrefl _ h1) in h7; contradiction. apply (sort_dec_irrefl _ h1) in h4; contradiction. pose proof (conj h3 h4) as h6. fold h2 in ih. rewrite <- ih in h6. destruct (h2 a b); try right; auto. Qed. Lemma in_wlb_list_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b x:T), let pftso := sort_dec_tso_dec _ pfso in In x (wlb_list R pftso l b) <-> In x l /\ Rw R b x. intros T R h1 l. induction l as [|a l ih]; intros b x h2; split; intro h3; simpl in h3; try destruct h3; try contradiction. fold h2 in ih. destruct (h2 a b) as [h4 | h4]. destruct h4 as [|h4]; subst. rewrite ih in h3. destruct h3; split; auto; try right; auto. destruct h3 as [|h3]; subst; split; [left; auto | right; auto | idtac | idtac]. rewrite ih in h3. destruct h3; right; auto. rewrite ih in h3. destruct h3; auto. destruct h3 as [|h3]; subst. split; left; auto. rewrite ih in h3. destruct h3 as [h3 h5]; split; [right | idtac]; auto. destruct H as [|H]; subst. simpl. destruct (h2 x b) as [h3 | h3]. destruct h3 as [h3 | h3]. pose proof (trans_Rw_R _ (sort_dec_trans _ h1) _ _ _ H0 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. subst. left; auto. left; auto. simpl. pose proof (conj H H0) as h7. rewrite <- ih in h7. fold h2 in h7. destruct (h2 a b) as [h3 | h3]. destruct h3 as [|h3]; auto; subst. right; auto. right; auto. Qed. Lemma ub_list_eq_nil_iff : forall {T:Type} (R:Relation T) (pftso:sort_dec R) (l:list T) (b:T), let pftso := sort_dec_tso_dec _ pftso in ub_list R pftso l b = nil <-> l = nil \/ (l <> nil /\ (forall x:T, In x l -> Rw R b x)). intros T R h1 l. destruct l as [|a l]; simpl. tauto. intro b; split; intro h2. destruct (sort_dec_tso_dec R h1 a b) as [h3 | h3]. destruct h3 as [h3 | h3]. discriminate. subst. right. split. intro; discriminate. intros x h3. destruct h3 as [|h3]; subst. right; auto. destruct (sort_dec_tso_dec R h1 b x) as [h4 | h4]. destruct h4 as [h4 | h4]. left; auto. subst. right; auto. pose proof (conj h3 h4) as h5. rewrite <- (in_ub_list_iff R h1) in h5. rewrite h2 in h5. contradict h5. right. split. intro; discriminate. intros x h4. destruct h4 as [|h4]; subst. left; auto. destruct (sort_dec_tso_dec R h1 b x) as [h5 | h5]. destruct h5 as [h5 | h5]. left; auto. subst. right; auto. pose proof (conj h4 h5) as h6. rewrite <- (in_ub_list_iff R h1) in h6. rewrite h2 in h6. contradict h6. destruct h2 as [| h2]; try discriminate. destruct (sort_dec_tso_dec R h1 a b) as [h3 | h3]. destruct h3 as [h3 | h3]. destruct h2 as [hn h2]. specialize (h2 a (or_introl eq_refl)). pose proof (trans_Rw_R _ (sort_dec_trans _ h1) _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. subst. revert h2. revert b. induction l as [|a l ih]; simpl. intros; auto. intros b h2. destruct (sort_dec_tso_dec R h1 a b) as [h3 | h3]. destruct h3 as [h3 | h3]. destruct h2 as [hn h2]. specialize (h2 a (or_intror (or_introl eq_refl))). pose proof (trans_Rw_R _ (sort_dec_trans _ h1) _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. subst. destruct h2 as [hn h2]. apply ih; split; auto. intro; discriminate. apply ih. split. intro; discriminate. intros x h4. destruct h4 as [|h4]; subst. right; auto. apply h2. right. right; auto. revert h3 h2. induction l as [|c l ih]; simpl. intros; auto. intros h2 h3. destruct (sort_dec_tso_dec R h1 c b) as [h4 | h4]. destruct h4 as [h4 | h4]. destruct h3 as [hn h3]. specialize (h3 c (or_intror (or_introl eq_refl))). pose proof (trans_Rw_R _ (sort_dec_trans _ h1) _ _ _ h3 h4) as h5. apply (sort_dec_irrefl _ h1) in h5; contradiction. subst. apply ih; auto. split. intro; discriminate. intros x h4. apply h3. destruct h4. left; auto. right; right; auto. apply ih; auto. split. intro; discriminate. intros x h5. apply h3. destruct h5. left; auto. right; right; auto. Qed. Lemma wub_list_eq_nil_iff : forall {T:Type} (R:Relation T) (pftso:sort_dec R) (l:list T) (b:T), let pftso := sort_dec_tso_dec _ pftso in wub_list R pftso l b = nil <-> l = nil \/ (l <> nil /\ (forall x:T, In x l -> R b x)). intros T R h1 l; induction l as [|a l ih]; intros b h2; split; intro h3; simpl; try fold h2 in ih. left; auto. reflexivity. right. simpl in h3. destruct (h2 a b) as [h4 | h4]. discriminate. rewrite ih in h3. destruct h3 as [|h3]; subst. split. intro; discriminate. intros x h5. destruct h5 as [|h5]; subst; auto. contradict h5. split. intro; discriminate. intros x h5. destruct h5 as [|h5]; subst; auto. destruct h3 as [hn h3]. apply h3; auto. destruct h3 as [h3 | h3]. discriminate. destruct h3 as [hn h3]. pose proof (h3 _ (in_eq _ _)) as h3'. assert (hnin : ~In b l). intro h4. specialize (h3 _ (in_cons _ _ _ h4)). apply (sort_dec_irrefl _ h1) in h3. contradiction. destruct (h2 a b) as [h4 | h4]. destruct h4 as [h4|]; subst. specialize (h3 _ (in_eq _ _)). pose proof ((sort_dec_trans _ h1) _ _ _ h3 h4) as h5. apply (sort_dec_irrefl _ h1) in h5. contradiction. apply (sort_dec_irrefl _ h1) in h3'. contradiction. destruct (nil_dec' l) as [|h5]; subst; auto. rewrite ih. right; split; auto. intros; apply h3; right; auto. Qed. Lemma lb_list_eq_nil_iff : forall {T:Type} (R:Relation T) (pftso:sort_dec R) (l:list T) (b:T), let pftso := sort_dec_tso_dec _ pftso in lb_list R pftso l b = nil <-> l = nil \/ (l <> nil /\ (forall x:T, In x l -> Rw R x b)). intros T R h1 l; induction l as [|a l ih]; intros b h2; split; intro h3; simpl; try fold h2 in ih; auto. simpl in h3. destruct (h2 a b) as [h4 | h4]. destruct h4 as [h4|]; subst. right. split. intro; discriminate. intros x h5. destruct h5 as [|h5]; subst. left; auto. rewrite ih in h3. destruct h3 as [|h3]; subst. contradict h5. apply h3; auto. right. split. intro; discriminate. intros x h4. destruct h4 as [|h4]; subst. right; auto. rewrite ih in h3. destruct h3 as [|h3]; subst. contradict h4. apply h3; auto. discriminate. destruct h3 as [|h3]. discriminate. destruct (h2 a b) as [h4 | h4]. destruct h4 as [h4|]; subst. destruct (nil_dec' l) as [|h5]; subst; auto. rewrite ih. right. split; intros; try apply h3; try right; auto. destruct (nil_dec' l) as [|h5]; subst; auto. rewrite ih. right; intros; try apply h3; try right; auto. split; auto. destruct h3 as [hn h3]. intros; apply h3; auto. right; auto. destruct h3 as [hn h3]. specialize (h3 _ (in_eq _ _)). pose proof (trans_Rw_R _ (sort_dec_trans _ h1) _ _ _ h3 h4) as h5. apply (sort_dec_irrefl _ h1) in h5; contradiction. Qed. Lemma wlb_list_eq_nil_iff : forall {T:Type} (R:Relation T) (pftso:sort_dec R) (l:list T) (b:T), let pftso := sort_dec_tso_dec _ pftso in wlb_list R pftso l b = nil <-> l = nil \/ (l <> nil /\ (forall x:T, In x l -> R x b)). intros T R h1 l; induction l as [|a l ih]; intros b h2; split; intro h3; try fold h2 in ih; auto. simpl in h3. destruct (h2 a b) as [h4 | h4]; try destruct h4 as [h4|]; subst. rewrite ih in h3. right. destruct h3 as [|h3]; subst. split. intros; discriminate. intros x h5. destruct h5 as [|h5]; subst; auto. contradict h5. split. intro; discriminate. intros x h5. destruct h3 as [hn h3]. destruct h5 as [|h5]; subst; auto. discriminate. discriminate. destruct h3 as [|h3]. discriminate. simpl. destruct (h2 a b) as [h4 | h4]. destruct h4 as [h4|]; subst. destruct (nil_dec' l) as [|hnin]; subst; auto. rewrite ih. right. split; intros; try apply h3; auto. right; auto. destruct h3 as [hn h3]. specialize (h3 _ (in_eq _ _)). apply (sort_dec_irrefl _ h1) in h3. contradiction. destruct h3 as [hn h3]. specialize (h3 _ (in_eq _ _)). pose proof (sort_dec_trans _ h1 _ _ _ h3 h4) as h5. apply (sort_dec_irrefl _ h1) in h5; contradiction. Qed. Lemma ub_list_lt_length : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T), length (ub_list R pftso l b) < S (length l). intros T R h1 l. induction l as [|a l ih]; simpl. intros; auto with arith. intro b. destruct (h1 a b) as [h2 | h2]; simpl. destruct h2 as [h2 |]; subst; simpl. apply lt_n_S. apply ih. eapply lt_trans. apply ih. auto with arith. eapply lt_trans. apply ih. auto with arith. Qed. Lemma wub_list_lt_length : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T), length (wub_list R pftso l b) < S (length l). intros T R h1 l. induction l as [|a l ih]; simpl. intros; auto with arith. intro b. destruct (h1 a b) as [h2 | h2]; simpl. destruct h2 as [h2 |]; subst; simpl. apply lt_n_S. apply ih. apply lt_n_S. apply ih. eapply lt_trans. apply ih. auto with arith. Qed. Lemma lb_list_lt_length : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T), length (lb_list R pftso l b) < S (length l). intros T R h1 l. induction l as [|a l ih]; simpl. intros; auto with arith. intro b. destruct (h1 a b) as [h2 | h2]; simpl. destruct h2 as [h2 |]; subst; simpl. eapply lt_trans. apply ih. auto with arith. eapply lt_trans. apply ih. auto with arith. apply lt_n_S. apply ih. Qed. Lemma wlb_list_lt_length : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T), length (wlb_list R pftso l b) < S (length l). intros T R h1 l. induction l as [|a l ih]; simpl. intros; auto with arith. intro b. destruct (h1 a b) as [h2 | h2]; simpl. destruct h2 as [h2 |]; subst; simpl. eapply lt_trans. apply ih. auto with arith. apply lt_n_S. apply ih. apply lt_n_S. apply ih. Qed. Lemma ub_list_le : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a b:T), R a b -> let pftso := sort_dec_tso_dec _ pfso in length (ub_list R pftso l a) <= length (ub_list R pftso l b). intros T R hs l. induction l as [|c l ih]; auto with arith. intros a b h2 ht. simpl. destruct (ht c a) as [h3 | h3]; simpl. destruct h3 as [h3 | h3]; simpl. destruct (ht c b) as [h4 | h4]; simpl. destruct h4 as [h4 | h5]; simpl. simpl. apply le_n_S. apply ih; auto. subst. pose proof (sort_dec_trans _ hs _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ hs) in h4. contradiction. pose proof (sort_dec_trans _ hs _ _ _ h2 h4) as h5. pose proof (sort_dec_trans _ hs _ _ _ h3 h5) as h6. apply (sort_dec_irrefl _ hs) in h6. contradiction. subst. destruct (ht a b) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. simpl. eapply le_trans. apply ih. apply h3. fold ht. auto with arith. apply (sort_dec_irrefl _ hs) in h2. contradiction. pose proof (sort_dec_trans _ hs _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ hs) in h4. contradiction. destruct (ht c b) as [h4 | h4]. destruct h4 as [h4 | h4]; subst. simpl. eapply le_trans. apply ih. apply h2. fold ht. auto with arith. apply ih; auto. apply ih; auto. Qed. Lemma wub_list_le : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a b:T), R a b -> let pftso := sort_dec_tso_dec _ pfso in length (wub_list R pftso l a) <= length (wub_list R pftso l b). intros T R hs l. induction l as [|c l ih]; auto with arith. intros a b h2 ht. simpl. destruct (ht c a) as [h3 | h3]; simpl. destruct h3 as [h3 | h3]; simpl. destruct (ht c b) as [h4 | h4]; simpl. destruct h4 as [h4 | h5]; simpl. simpl. apply le_n_S. apply ih; auto. subst. pose proof (sort_dec_trans _ hs _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ hs) in h4. contradiction. pose proof (sort_dec_trans _ hs _ _ _ h2 h4) as h5. pose proof (sort_dec_trans _ hs _ _ _ h3 h5) as h6. apply (sort_dec_irrefl _ hs) in h6. contradiction. subst. destruct (ht a b) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. simpl. auto with arith. simpl. auto with arith. pose proof (sort_dec_trans _ hs _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ hs) in h4. contradiction. destruct (ht c b) as [h4 | h4]. destruct h4 as [h4 | h4]; subst. simpl. eapply le_trans. apply ih. apply h2. fold ht. auto with arith. simpl. auto with arith. apply ih; auto. Qed. Lemma lb_list_le : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a b:T), R a b -> let pftso := sort_dec_tso_dec _ pfso in length (lb_list R pftso l b) <= length (lb_list R pftso l a). intros T R hs l. induction l as [|c l ih]; auto with arith. intros a b h2 ht. simpl. destruct (ht c a) as [h3 | h3]; simpl. destruct h3 as [h3 | h3]; simpl. destruct (ht c b) as [h4 | h4]; simpl. destruct h4 as [h4 | h5]; simpl. simpl. apply ih; auto. subst. pose proof (sort_dec_trans _ hs _ _ _ h2 h3) as h5. apply (sort_dec_irrefl _ hs) in h5. contradiction. pose proof (sort_dec_trans _ hs _ _ _ h2 h4) as h5. pose proof (sort_dec_trans _ hs _ _ _ h3 h5) as h6. apply (sort_dec_irrefl _ hs) in h6. contradiction. subst. destruct (ht a b) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. simpl. apply ih; auto. apply (sort_dec_irrefl _ hs) in h2. contradiction. pose proof (sort_dec_trans _ hs _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ hs) in h4. contradiction. destruct (ht c b) as [h4 | h4]. destruct h4 as [h4 | h4]; subst. simpl. eapply le_trans. apply ih. apply h2. fold ht. auto with arith. eapply le_trans. apply ih. apply h3. fold ht. auto with arith. simpl. apply le_n_S. apply ih; auto. Qed. Lemma wlb_list_le : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a b:T), R a b -> let pftso := sort_dec_tso_dec _ pfso in length (wlb_list R pftso l b) <= length (wlb_list R pftso l a). intros T R hs l. induction l as [|c l ih]; auto with arith. intros a b h2 ht. simpl. destruct (ht c a) as [h3 | h3]; simpl. destruct h3 as [h3 | h3]; simpl. destruct (ht c b) as [h4 | h4]; simpl. destruct h4 as [h4 | h5]; simpl. simpl. apply ih; auto. subst. pose proof (sort_dec_trans _ hs _ _ _ h2 h3) as h5. apply (sort_dec_irrefl _ hs) in h5. contradiction. pose proof (sort_dec_trans _ hs _ _ _ h2 h4) as h5. pose proof (sort_dec_trans _ hs _ _ _ h3 h5) as h6. apply (sort_dec_irrefl _ hs) in h6. contradiction. subst. destruct (ht a b) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. simpl. eapply le_trans. apply ih. apply h3. fold ht. auto with arith. apply (sort_dec_irrefl _ hs) in h2. contradiction. simpl. apply le_n_S. apply ih. assumption. destruct (ht c b) as [h4 | h4]. destruct h4 as [h4 | h4]; subst. auto with arith. simpl. apply le_n_S. apply ih. assumption. simpl. apply le_n_S. apply ih; auto. Qed. Lemma coarse_list_ub_list : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T), coarse_list (ub_list R pftso l b) l. intros T R h1 l. induction l as [|a l h2]; simpl. intros; constructor. intro b. destruct (h1 a b) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. apply coarse_list_cons2. apply h2. apply coarse_list_cons1. apply h2. apply coarse_list_cons1. apply h2. Qed. Lemma no_dup_ub_list : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (a:T), NoDup l -> NoDup (ub_list R pftso l a). intros T R h1 l a h2. eapply no_dup_coarse_list. apply coarse_list_ub_list. assumption. Qed. Lemma coarse_list_ub_list_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a b:T), R a b -> let pftso := sort_dec_tso_dec _ pfso in coarse_list (ub_list R pftso l a) (ub_list R pftso l b). intros T R hs l. induction l as [|c l ih]. intros; constructor. intros a b h2 htso. simpl. destruct (htso c a) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. pose proof (sort_dec_trans _ hs _ _ _ h3 h2) as h4. destruct (htso c b) as [h5 | h5]. destruct h5 as [h5|]; subst. apply coarse_list_cons2; auto. apply (sort_dec_irrefl _ hs) in h4; contradiction. pose proof (sort_dec_trans _ hs _ _ _ h4 h5) as h6. apply (sort_dec_irrefl _ hs) in h6; contradiction. destruct (htso a b) as [h3 | h3]. destruct h3 as [h3 | h3]. apply coarse_list_cons1; auto. subst. apply (sort_dec_irrefl _ hs) in h2; contradiction. pose proof (sort_dec_trans _ hs _ _ _ h2 h3) as h6. apply (sort_dec_irrefl _ hs) in h6; contradiction. destruct (htso c b) as [h4 | h4]. destruct h4 as [h4 | h4]; subst. apply coarse_list_cons1; auto. auto. auto. Qed. Lemma coarse_list_wub_list_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a b:T), R a b -> let pftso := sort_dec_tso_dec _ pfso in coarse_list (wub_list R pftso l a) (wub_list R pftso l b). intros T R hs l. induction l as [|c l ih]. intros; constructor. intros a b h2 htso. simpl. destruct (htso c a) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. pose proof (sort_dec_trans _ hs _ _ _ h3 h2) as h4. destruct (htso c b) as [h5 | h5]. destruct h5 as [h5|]; subst. apply coarse_list_cons2; auto. apply (sort_dec_irrefl _ hs) in h4; contradiction. pose proof (sort_dec_trans _ hs _ _ _ h4 h5) as h6. apply (sort_dec_irrefl _ hs) in h6; contradiction. destruct (htso a b) as [h3 | h3]. destruct h3 as [h3 | h3]. apply coarse_list_cons2; auto. subst. apply (sort_dec_irrefl _ hs) in h2; contradiction. pose proof (sort_dec_trans _ hs _ _ _ h2 h3) as h6. apply (sort_dec_irrefl _ hs) in h6; contradiction. destruct (htso c b) as [h4 | h4]. destruct h4 as [h4 | h4]; subst. apply coarse_list_cons1; auto. apply coarse_list_cons1; auto. auto. Qed. Lemma coarse_list_lb_list_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a b:T), R a b -> let pftso := sort_dec_tso_dec _ pfso in coarse_list (lb_list R pftso l b) (lb_list R pftso l a). intros T R hs l. induction l as [|c l ih]. intros; constructor. intros a b h2 htso. simpl. destruct (htso c a) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. pose proof (sort_dec_trans _ hs _ _ _ h3 h2) as h4. destruct (htso c b) as [h5 | h5]. destruct h5 as [h5|]; subst; auto. pose proof (sort_dec_trans _ hs _ _ _ h4 h5) as h6. apply (sort_dec_irrefl _ hs) in h6; contradiction. destruct (htso a b) as [h3 | h3]. destruct h3 as [h3 | h3]; auto. pose proof (sort_dec_trans _ hs _ _ _ h2 h3) as h6. apply (sort_dec_irrefl _ hs) in h6; contradiction. destruct (htso c b) as [h4 | h4]. destruct h4 as [h4 | h4]; subst. apply coarse_list_cons1; auto. apply coarse_list_cons1; auto. apply coarse_list_cons2; auto. Qed. Lemma coarse_list_wlb_list_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a b:T), R a b -> let pftso := sort_dec_tso_dec _ pfso in coarse_list (wlb_list R pftso l b) (wlb_list R pftso l a). intros T R hs l. induction l as [|c l ih]. intros; constructor. intros a b h2 htso. simpl. destruct (htso c a) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. pose proof (sort_dec_trans _ hs _ _ _ h3 h2) as h4. destruct (htso c b) as [h5 | h5]. destruct h5 as [h5|]; subst; auto. apply (sort_dec_irrefl _ hs) in h4; contradiction. pose proof (sort_dec_trans _ hs _ _ _ h4 h5) as h6. apply (sort_dec_irrefl _ hs) in h6; contradiction. destruct (htso a b) as [h3 | h3]. destruct h3 as [h3 | h3]; auto. apply coarse_list_cons1; auto. subst; auto. apply (sort_dec_irrefl _ hs) in h2; contradiction. pose proof (sort_dec_trans _ hs _ _ _ h2 h3) as h6. apply (sort_dec_irrefl _ hs) in h6; contradiction. destruct (htso c b) as [h4 | h4]. destruct h4 as [h4 | h4]; subst. apply coarse_list_cons1; auto. apply coarse_list_cons2; auto. apply coarse_list_cons2; auto. Qed. Lemma in_ub_wlb_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a x y:T), let pftso := sort_dec_tso_dec _ pfso in In x (ub_list R pftso l a) -> In y (wlb_list R pftso l a) -> R x y. intros T R h1 l a x y h2 h3 h4. unfold h2 in h3, h4. rewrite in_ub_list_iff in h3; destruct h3 as [h3 h5]. rewrite in_wlb_list_iff in h4; destruct h4 as [h4 h6]. eapply trans_R_Rw. eapply sort_dec_trans. assumption. apply h5. assumption. Qed. Lemma in_lb_wub_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a x y:T), let pftso := sort_dec_tso_dec _ pfso in In x (wub_list R pftso l a) -> In y (lb_list R pftso l a) -> R x y. intros T R h1 l a x y h2 h3 h4. unfold h2 in h3, h4. rewrite in_wub_list_iff in h3; destruct h3 as [h3 h5]. rewrite in_lb_list_iff in h4; destruct h4 as [h4 h6]. eapply trans_Rw_R. eapply sort_dec_trans. assumption. apply h5. assumption. Qed. Lemma ub_list_nest : forall {T:Type} (R:Relation T) (pfs:sort_dec R) (l:list T) (a b:T), let pftso := sort_dec_tso_dec R pfs in ub_list R pftso (ub_list R pftso l a) b = ub_list R pftso l (minR R pftso a b). intros T R h1 l a b h3. unfold minR. destruct (h3 a b) as [h4 | h4]. destruct h4 as [h4 |]; subst. revert h4. revert a b. induction l as [|c l ih]; simpl; auto. intros a b h2. destruct (h3 c a) as [h4 | h4]. destruct h4 as [h4|]; subst; simpl. pose proof (sort_dec_trans _ h1 _ _ _ h4 h2) as h5. destruct (h3 c b) as [h6 | h6]. destruct h6 as [h6 | h6]. f_equal; apply ih; auto. subst. pose proof (sort_dec_irrefl _ h1 _ h5); contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h6 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. auto. auto. induction l as [|a l ih]; simpl; auto. destruct (h3 a b) as [h4 | h4]. destruct h4 as [h4 | h4]; simpl. destruct (h3 a b) as [h5 | h5]. destruct h5 as [h5 | h5]. f_equal; auto. subst. pose proof (sort_dec_irrefl _ h1 _ h4); contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. subst; auto. apply ih. induction l as [|c l ih]; simpl; auto. destruct (h3 c a) as [h5 | h5]. destruct h5 as [h5|]; subst; auto. simpl. destruct (h3 c b) as [h5' | h5']. destruct h5' as [h5' | h5']. f_equal; auto. subst; auto. auto. destruct (h3 a b) as [h5|h5]; subst. destruct h5 as [h5 | h5]. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. subst; auto. auto. destruct (h3 c b) as [h6 | h6]. destruct h6 as [h6 | h6]. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_trans _ h1 _ _ _ h7 h6) as h8. pose proof (sort_dec_irrefl _ h1 _ h8); contradiction. subst. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h8. pose proof (sort_dec_irrefl _ h1 _ h8); contradiction. auto. Qed. Lemma wub_list_nest : forall {T:Type} (R:Relation T) (pfs:sort_dec R) (l:list T) (a b:T), let pftso := sort_dec_tso_dec R pfs in wub_list R pftso (wub_list R pftso l a) b = wub_list R pftso l (minR R pftso a b). intros T R h1 l a b h3. unfold minR. destruct (h3 a b) as [h4 | h4]. destruct h4 as [h4 |]; subst. revert h4. revert a b. induction l as [|c l ih]; simpl; auto. intros a b h2. destruct (h3 c a) as [h4 | h4]. destruct h4 as [h4|]; subst; simpl. pose proof (sort_dec_trans _ h1 _ _ _ h4 h2) as h5. destruct (h3 c b) as [h6 | h6]. destruct h6 as [h6 | h6]. f_equal; apply ih; auto. subst. pose proof (sort_dec_irrefl _ h1 _ h5); contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h6 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. destruct (h3 a b) as [h4 | h4]. destruct h4 as [h4 |]; subst. f_equal; auto. pose proof (sort_dec_irrefl _ h1 _ h2); contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h2 h4) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. auto. induction l as [|a l ih]; simpl; auto. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. destruct hl as [hl | hl]; simpl. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. f_equal; auto. pose proof (sort_dec_trans _ h1 _ _ _ hl hr') as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. subst. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']. f_equal; auto. pose proof (sort_dec_irrefl _ h1 _ hr'); contradiction. auto. induction l as [|c l ih]; simpl; auto. destruct (h3 c a) as [h5 | h5]. destruct h5 as [h5|]; subst; auto. simpl. destruct (h3 c b) as [h5' | h5']. f_equal; auto. auto. destruct (h3 a b) as [h5|h5]; subst. destruct h5 as [h5 | h5]. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. subst. pose proof (sort_dec_irrefl _ h1 _ h4); contradiction. simpl. destruct (h3 a b) as [h6 | h6]. destruct h6 as [h6 | h6]. pose proof (sort_dec_trans _ h1 _ _ _ h5 h6) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. subst. pose proof (sort_dec_irrefl _ h1 _ h5); contradiction. auto. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h8. destruct (h3 c b) as [h9 | h9]. destruct h9 as [h9|]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h8 h9) as h10. pose proof (sort_dec_irrefl _ h1 _ h10); contradiction. pose proof (sort_dec_irrefl _ h1 _ h8); contradiction. assumption. Qed. Lemma lb_list_nest : forall {T:Type} (R:Relation T) (pfs:sort_dec R) (l:list T) (a b:T), let pftso := sort_dec_tso_dec R pfs in lb_list R pftso (lb_list R pftso l a) b = lb_list R pftso l (maxR R pftso a b). intros T R h1 l a b h3. unfold maxR. destruct (h3 a b) as [h4 | h4]. destruct h4 as [h4 |]; subst. revert h4. revert a b. induction l as [|c l ih]; simpl; auto. intros a b h2. destruct (h3 c a) as [h4 | h4]. destruct h4 as [h4|]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h4 h2) as h5. destruct (h3 c b) as [h6 | h6]. destruct h6 as [h6 | h6]. apply ih; auto. subst. pose proof (sort_dec_irrefl _ h1 _ h5); contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h6 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. destruct (h3 a b) as [h4 | h4]. destruct h4 as [h4 |]; subst. simpl. destruct (h3 a b) as [h5 | h5]. destruct h5 as [h5 | h5]. apply ih; auto. subst. pose proof (sort_dec_irrefl _ h1 _ h4); contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. pose proof (sort_dec_irrefl _ h1 _ h2); contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h2 h4) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. simpl. destruct (h3 c b) as [h5 | h5]. destruct h5 as [h5|]; subst. apply ih; auto. apply ih; auto. f_equal; apply ih; auto. induction l as [|a l ih]; simpl; auto. destruct (h3 a b) as [h4 | h4]; subst. destruct h4 as [h4|]; subst. apply ih; auto. simpl. destruct (h3 b b) as [h4 |]. destruct h4 as [h4 | h5]. pose proof (sort_dec_irrefl _ h1 _ h4); contradiction. apply ih. apply ih. simpl. destruct (h3 a b) as [h5 | h5]. destruct h5 as [h5|]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. pose proof (sort_dec_irrefl _ h1 _ h4); contradiction. f_equal; apply ih; auto. induction l as [|c l ih]; simpl; auto. destruct (h3 c a) as [h5 | h5]. destruct h5 as [h5|]; subst; auto. simpl. destruct (h3 c b) as [h5' | h5']. destruct h5' as [h5'|]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_trans _ h1 _ _ _ h7 h5') as h8. pose proof (sort_dec_irrefl _ h1 _ h8); contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. f_equal; auto. Qed. Lemma wlb_list_nest : forall {T:Type} (R:Relation T) (pfs:sort_dec R) (l:list T) (a b:T), let pftso := sort_dec_tso_dec R pfs in wlb_list R pftso (wlb_list R pftso l a) b = wlb_list R pftso l (maxR R pftso a b). intros T R h1 l a b h3. unfold maxR. destruct (h3 a b) as [h4 | h4]. destruct h4 as [h4 |]; subst. revert h4. revert a b. induction l as [|c l ih]; simpl; auto. intros a b h2. destruct (h3 c a) as [h4 | h4]. destruct h4 as [h4|]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h4 h2) as h5. destruct (h3 c b) as [h6 | h6]. destruct h6 as [h6 | h6]. apply ih; auto. subst. pose proof (sort_dec_irrefl _ h1 _ h5); contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h6 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. destruct (h3 a b) as [h4 | h4]. destruct h4 as [h4 |]; subst. simpl. destruct (h3 a b) as [h5 | h5]. destruct h5 as [h5 | h5]. apply ih; auto. subst. pose proof (sort_dec_irrefl _ h1 _ h4); contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. pose proof (sort_dec_irrefl _ h1 _ h2); contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h2 h4) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. simpl. destruct (h3 c b) as [h5 | h5]. destruct h5 as [h5|]; subst. apply ih; auto. f_equal. apply ih; auto. f_equal. apply ih; auto. induction l as [|a l ih]; simpl; auto. destruct (h3 a b) as [h4 | h4]; subst. destruct h4 as [h4|]; subst. apply ih; auto. simpl. destruct (h3 b b) as [h4 |]. destruct h4 as [h4 | h5]. pose proof (sort_dec_irrefl _ h1 _ h4); contradiction. f_equal; apply ih; auto. f_equal; apply ih; auto. simpl. destruct (h3 a b) as [h5 | h5]. destruct h5 as [h5|]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. pose proof (sort_dec_irrefl _ h1 _ h4); contradiction. f_equal; apply ih; auto. induction l as [|c l ih]; simpl; auto. destruct (h3 c a) as [h5 | h5]. destruct h5 as [h5|]; subst; auto. simpl. destruct (h3 a b) as [h5 | h5]. destruct h5 as [h5|]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_irrefl _ h1 _ h7); contradiction. f_equal; auto. f_equal; auto. simpl. destruct (h3 c b) as [h6 | h6]. destruct h6 as [h6|]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. pose proof (sort_dec_trans _ h1 _ _ _ h6 h7) as h8. pose proof (sort_dec_irrefl _ h1 _ h8); contradiction. f_equal. apply ih. f_equal. apply ih. Qed. Lemma ub_list_wlb_list_comm : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a b:T), R a b -> let pftso := sort_dec_tso_dec _ pfso in ub_list R pftso (wlb_list R pftso l a) b = wlb_list R pftso (ub_list R pftso l b) a. intros T R h1 l. induction l as [|c l ih]; auto. intros a b h2 hts. simpl. destruct (hts c b) as [h4 | h4]. destruct h4 as [h4|]; subst. simpl. destruct (hts c a) as [h3 | h3]. destruct h3 as [h3|]; subst; auto. simpl. destruct (hts a b) as [h5 | h5]. destruct h5 as [h5|]; subst. f_equal; auto. apply (sort_dec_irrefl _ h1) in h4; contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h6. apply (sort_dec_irrefl _ h1) in h6; contradiction. simpl. destruct (hts c b) as [h6 | h6]. destruct h6 as [h7 | h7]. f_equal; auto. subst. apply (sort_dec_irrefl _ h1) in h4; contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h4 h6) as h7. apply (sort_dec_irrefl _ h1) in h7; contradiction. destruct (hts b a) as [h4 | h4]. destruct h4 as [h4|]; subst; auto. apply (sort_dec_irrefl _ h1) in h2; contradiction. simpl. destruct (hts b b) as [h5 | h5]. destruct h5 as [h5 | h5]. apply (sort_dec_irrefl _ h1) in h5; contradiction. auto. apply (sort_dec_irrefl _ h1) in h5; contradiction. destruct (hts c a) as [h3 | h3]. destruct h3 as [h3|]; subst; auto. simpl. pose proof (sort_dec_trans _ h1 _ _ _ h2 h4) as h7. apply (sort_dec_irrefl _ h1) in h7; contradiction. simpl. destruct (hts c b) as [h5 | h5]. destruct h5 as [h5 | h5]. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h7. apply (sort_dec_irrefl _ h1) in h7; contradiction. subst. apply (sort_dec_irrefl _ h1) in h4; contradiction. auto. Qed. Lemma in_or_ub_wlb : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), In a l -> let pftso := sort_dec_tso_dec _ pfso in In a (ub_list R pftso l a) \/ In a (wlb_list R pftso l a). intros T R h1 l. destruct l as [|b l]; auto. intros a h2 hts. destruct h2 as [|h2]; subst. destruct (hts a a) as [h2 | h2]. destruct h2 as [h2 | h2]. apply (sort_dec_irrefl _ h1) in h2; contradiction. rewrite wlb_list_cons_same. right; left; auto. apply sort_dec_irrefl; auto. apply (sort_dec_irrefl _ h1) in h2; contradiction. simpl. destruct (hts b a) as [h3 | h3]. destruct h3 as [h3 | h3]. right. unfold hts. rewrite in_wlb_list_iff. split; auto; right; auto. subst. right; left; auto. right. simpl. right. unfold hts. rewrite in_wlb_list_iff. split; auto; right; auto. Qed. Lemma in_or_lb_wub : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), In a l -> let pftso := sort_dec_tso_dec _ pfso in In a (lb_list R pftso l a) \/ In a (wub_list R pftso l a). intros T R h1 l. destruct l as [|b l]; auto. intros a h2 hts. destruct h2 as [|h2]; subst. destruct (hts a a) as [h2 | h2]. destruct h2 as [h2 | h2]. apply (sort_dec_irrefl _ h1) in h2; contradiction. rewrite wub_list_cons_same. right; left; auto. apply sort_dec_irrefl; auto. apply (sort_dec_irrefl _ h1) in h2; contradiction. simpl. destruct (hts b a) as [h3 | h3]. destruct h3 as [h3 | h3]. right. unfold hts. right. rewrite in_wub_list_iff. split; auto; right; auto. subst. right; left; auto. right. simpl. unfold hts. rewrite in_wub_list_iff. split; auto; right; auto. Qed. Lemma in_ub_list_iffn : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a x:T), In x l -> let pftso := sort_dec_tso_dec _ pfso in In x (ub_list R pftso l a) <-> ~In x (wlb_list R pftso l a). intros T R h1 l a x h2 h3; split; intro h4. intro h5. unfold h3 in h4, h5. rewrite in_ub_list_iff in h4. rewrite in_wlb_list_iff in h5. destruct h4 as [? h4], h5 as [? h5]. rewrite R_iffn in h4; auto. unfold h3 in h4. rewrite in_wlb_list_iff in h4. unfold h3. rewrite in_ub_list_iff. assert (h5:~ Rw R a x). tauto. split; auto. rewrite R_iffn; auto. Qed. Lemma in_wub_list_iffn : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a x:T), In x l -> let pftso := sort_dec_tso_dec _ pfso in In x (wub_list R pftso l a) <-> ~In x (lb_list R pftso l a). intros T R h1 l a x h2 h3; split; intro h4. intro h5. unfold h3 in h4, h5. rewrite in_wub_list_iff in h4. rewrite in_lb_list_iff in h5. destruct h4 as [? h4], h5 as [? h5]. rewrite R_iffn in h5; auto. unfold h3 in h4. rewrite in_lb_list_iff in h4. unfold h3. rewrite in_wub_list_iff. assert (h5:~ R a x). tauto. split; auto. rewrite Rw_iffn; auto. Qed. Lemma in_lb_list_iffn : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a x:T), In x l -> let pftso := sort_dec_tso_dec _ pfso in In x (lb_list R pftso l a) <-> ~In x (wub_list R pftso l a). intros T R h1 l a x h2 h3; split; intro h4. intro h5. unfold h3 in h4, h5. rewrite in_lb_list_iff in h4. rewrite in_wub_list_iff in h5. destruct h4 as [? h4], h5 as [? h5]. rewrite R_iffn in h4; auto. unfold h3 in h4. rewrite in_wub_list_iff in h4. unfold h3. rewrite in_lb_list_iff. assert (h5:~ Rw R x a). tauto. split; auto. rewrite R_iffn; auto. Qed. Lemma in_wlb_list_iffn : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a x:T), In x l -> let pftso := sort_dec_tso_dec _ pfso in In x (wlb_list R pftso l a) <-> ~In x (ub_list R pftso l a). intros T R h1 l a x h2 h3; split; intro h4. intro h5. unfold h3 in h4, h5. rewrite in_wlb_list_iff in h4. rewrite in_ub_list_iff in h5. destruct h4 as [? h4], h5 as [? h5]. rewrite R_iffn in h5; auto. unfold h3 in h4. rewrite in_ub_list_iff in h4. unfold h3. rewrite in_wlb_list_iff. assert (h5:~ R x a). tauto. split; auto. rewrite Rw_iffn; auto. Qed. Lemma length_decompose_ub_wlb : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (a:T), length l = length (ub_list R pftso l a) + length (wlb_list R pftso l a). intros T R h1 l. induction l as [|c l ih]; simpl; auto. intro a. destruct (h1 c a) as [h2 | h2]; simpl. destruct h2 as [h2 |]; subst; simpl. f_equal; auto. rewrite plus_S_shift_r. f_equal; auto. rewrite plus_S_shift_r. f_equal; auto. Qed. Lemma length_decompose_lb_wub : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (a:T), length l = length (lb_list R pftso l a) + length (wub_list R pftso l a). intros T R h1 l. induction l as [|c l ih]; simpl; auto. intro a. destruct (h1 c a) as [h2 | h2]; simpl. destruct h2 as [h2 |]; subst; simpl. rewrite plus_S_shift_r. f_equal; auto. rewrite plus_S_shift_r. f_equal; auto. f_equal; auto. Qed. Lemma length_wlb_list_decompose : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a b:T), R a b -> let pftso := sort_dec_tso_dec _ pfso in length (wlb_list R pftso l a) = length (ub_list R pftso (wlb_list R pftso l a) b) + length (wlb_list R pftso l b). intros T R h1 l. induction l as [|c l ih]; auto. intros a b h2 htso. simpl. destruct (htso c a) as [h3 | h3]. destruct h3 as [h3|]; subst. pose proof (sort_dec_trans _ h1 _ _ _ h3 h2) as h4. destruct (htso c b) as [h5 | h5]. destruct h5 as [h5|]; subst; auto. apply (sort_dec_irrefl _ h1) in h4; contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h4 h5) as h6. apply (sort_dec_irrefl _ h1) in h6; contradiction. simpl. destruct (htso a b) as [h3 | h3]. destruct h3 as [h3|]; subst. simpl; f_equal; auto. apply (sort_dec_irrefl _ h1) in h2; contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h6. apply (sort_dec_irrefl _ h1) in h6; contradiction. simpl. destruct (htso c b) as [h5 | h5]. destruct h5 as [h5|]; subst. simpl; f_equal; auto. simpl. rewrite plus_S_shift_r. f_equal; auto. simpl. rewrite plus_S_shift_r. f_equal; auto. Qed. Lemma count_occ_ub_wlb_decompose : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (pftso:tso_dec R) (l:list T) (a x:T), count_occ pfdt l x = count_occ pfdt (ub_list R pftso l a) x + count_occ pfdt (wlb_list R pftso l a) x. intros T h1 R h2 l. induction l as [|c l ih]; simpl; auto. intros a x. destruct (h1 c x) as [|h3]; subst. destruct (h2 x a) as [h4|h4]; subst. destruct h4 as [h4 | h4]. simpl. destruct (h1 x x) as [h5 | h5]. rewrite plus_S_shift_l. f_equal; auto. contradict h5; auto. subst. simpl. destruct (h1 a a) as [h3 | h3]. rewrite plus_S_shift_r. f_equal; auto. contradict h3; auto. simpl. destruct (h1 x x) as [h5 | h5]. rewrite plus_S_shift_r. f_equal; auto. contradict h5; auto. destruct (h2 c a) as [h4 | h4]. destruct h4 as [h4 | h4]. simpl. destruct (h1 c x) as [h5 | h5]. contradiction. apply ih. subst. simpl. destruct (h1 a x) as [h4 | h4]. contradiction. apply ih. simpl. destruct (h1 c x) as [h5 | h5]. contradiction. apply ih. Qed. Lemma coarse_list_wub_list : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T), coarse_list (wub_list R pftso l b) l. intros T R h1 l. induction l as [|a l h2]; simpl. intros; constructor. intro b. destruct (h1 a b) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. apply coarse_list_cons2. apply h2. apply coarse_list_cons2. apply h2. apply coarse_list_cons1. apply h2. Qed. Lemma no_dup_wub_list : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (a:T), NoDup l -> NoDup (wub_list R pftso l a). intros T R h1 l a h2. eapply no_dup_coarse_list. apply coarse_list_wub_list. assumption. Qed. Lemma coarse_list_lb_list : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T), coarse_list (lb_list R pftso l b) l. intros T R h1 l. induction l as [|a l h2]; simpl. intros; constructor. intro b. destruct (h1 a b) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. apply coarse_list_cons1. apply h2. apply coarse_list_cons1. apply h2. apply coarse_list_cons2. apply h2. Qed. Lemma no_dup_lb_list : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (a:T), NoDup l -> NoDup (lb_list R pftso l a). intros T R h1 l a h2. eapply no_dup_coarse_list. apply coarse_list_lb_list. assumption. Qed. Lemma coarse_list_wlb_list : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (b:T), coarse_list (wlb_list R pftso l b) l. intros T R h1 l. induction l as [|a l h2]; simpl. intros; constructor. intro b. destruct (h1 a b) as [h3 | h3]. destruct h3 as [h3 | h3]; subst. apply coarse_list_cons1. apply h2. apply coarse_list_cons2. apply h2. apply coarse_list_cons2. apply h2. Qed. Lemma no_dup_wlb_list : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (a:T), NoDup l -> NoDup (wlb_list R pftso l a). intros T R h1 l a h2. eapply no_dup_coarse_list. apply coarse_list_wlb_list. assumption. Qed. Lemma ub_list_lt_length_step : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b:T), let pftso := sort_dec_tso_dec _ pfso in match ub_list R pftso l b with | nil => True | a :: l' => length (ub_list R pftso l' a) < S (length l') end. intros T R h1 l. induction l as [|a l ih]; auto. simpl; auto. intros b ht. simpl. destruct (ht a b) as [h3 | h3]. destruct h3 as [h3 | h3]. pose proof (coarse_list_ub_list R ht (ub_list R ht l b) a) as h6. pose proof (length_coarse_list _ _ h6) as h7. red. apply le_n_S. assumption. subst. destruct (nil_dec' (ub_list R ht l b)) as [h4 | h4]. rewrite h4; simpl. auto with arith. rewrite (hd_compat' _ h4). pose proof (coarse_list_ub_list R ht (tl (ub_list R ht l b)) (hd' (ub_list R ht l b) h4)) as h6. pose proof (length_coarse_list _ _ h6) as h7. red. apply le_n_S; auto. destruct (nil_dec' (ub_list R ht l b)) as [h4 | h4]. rewrite h4; simpl; auto. rewrite (hd_compat' _ h4). pose proof (coarse_list_ub_list R ht (tl (ub_list R ht l b)) (hd' (ub_list R ht l b) h4)) as h6. pose proof (length_coarse_list _ _ h6) as h7. red. apply le_n_S; auto. Qed. Lemma wub_list_lt_length_step : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b:T), let pftso := sort_dec_tso_dec _ pfso in match wub_list R pftso l b with | nil => True | a :: l' => length (wub_list R pftso l' a) < S (length l') end. intros T R h1 l. induction l as [|a l ih]; auto. simpl; auto. intros b ht. simpl. destruct (ht a b) as [h3 | h3]. destruct h3 as [h3 | h3]. pose proof (coarse_list_wub_list R ht (wub_list R ht l b) a) as h6. pose proof (length_coarse_list _ _ h6) as h7. red. apply le_n_S; auto. subst. pose proof (coarse_list_wub_list R ht (wub_list R ht l b) b) as h6. pose proof (length_coarse_list _ _ h6) as h7. red. apply le_n_S; auto. destruct (nil_dec' (wub_list R ht l b)) as [h4 | h4]. rewrite h4; auto. rewrite (hd_compat' _ h4). pose proof (coarse_list_wub_list R ht (tl (wub_list R ht l b)) (hd' (wub_list R ht l b) h4)) as h6. pose proof (length_coarse_list _ _ h6) as h7. apply le_lt_n_Sm; auto. Qed. Lemma lb_list_lt_length_step : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b:T), let pftso := sort_dec_tso_dec _ pfso in match lb_list R pftso l b with | nil => True | a :: l' => length (lb_list R pftso l' a) < S (length l') end. intros T R h1 l. induction l as [|a l ih]; auto. intros b ht. simpl; auto. intros b ht. simpl. destruct (ht a b) as [h3 | h3]. destruct h3 as [h3 | h3]. apply ih. subst. destruct (nil_dec' (lb_list R ht l b)) as [h4 | h4]. rewrite h4; simpl; auto. rewrite (hd_compat' _ h4). pose proof (coarse_list_lb_list R ht (tl (lb_list R ht l b)) (hd' (lb_list R ht l b) h4)) as h6. pose proof (length_coarse_list _ _ h6) as h7. red. apply le_n_S; auto. pose proof (lb_list_le R h1 (lb_list R ht l b) b a h3) as h4. simpl in h4. fold ht in h4. apply le_lt_n_Sm; auto. eapply le_trans. apply h4. pose proof (coarse_list_lb_list R ht (lb_list R ht l b) b) as h6. pose proof (length_coarse_list _ _ h6) as h7. assumption. Qed. Lemma wlb_list_lt_length_step : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b:T), let pftso := sort_dec_tso_dec _ pfso in match wlb_list R pftso l b with | nil => True | a :: l' => length (wlb_list R pftso l' a) < S (length l') end. intros T R h1 l. induction l as [|a l ih]; auto. intros b ht. simpl; auto. intros b ht. simpl. destruct (ht a b) as [h3 | h3]. destruct h3 as [h3 | h3]. destruct (nil_dec' (wlb_list R ht l b)) as [h4 | h4]. rewrite h4; simpl; auto. rewrite (hd_compat' _ h4). pose proof (coarse_list_wlb_list R ht (tl (wlb_list R ht l b)) (hd' (wlb_list R ht l b) h4)) as h6. pose proof (length_coarse_list _ _ h6) as h7. red. apply le_n_S; auto. subst. pose proof (coarse_list_wlb_list R ht (wlb_list R ht l b) b) as h6. pose proof (length_coarse_list _ _ h6) as h7. red. apply le_n_S; auto. pose proof (coarse_list_wlb_list R ht (wlb_list R ht l b) a) as h6. pose proof (length_coarse_list _ _ h6) as h7. apply le_lt_n_Sm; auto. Qed. Lemma ub_wlb_list_lt_length_step : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b:T), let pftso := sort_dec_tso_dec _ pfso in match ub_list R pftso l b with | nil => True | a :: l' => length (wlb_list R (sort_dec_tso_dec R pfso) l' a) < S (length l') end. intros T R h1 l. induction l as [|a l ih]; simpl; auto. intro b. destruct (sort_dec_tso_dec R h1 a b) as [h4 | h4]. destruct h4 as [h4 | h4]; simpl. pose proof (coarse_list_wlb_list R (sort_dec_tso_dec R h1) (ub_list R (sort_dec_tso_dec R h1) l b) a) as h5. pose proof (length_coarse_list _ _ h5) as h6. apply le_lt_n_Sm; auto. subst. apply ih. apply ih. Qed. Lemma wlb_ub_list_lt_length_step : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (b:T), let pftso := sort_dec_tso_dec _ pfso in match wlb_list R pftso l b with | nil => True | a :: l' => length (ub_list R (sort_dec_tso_dec R pfso) l' a) < S (length l') end. intros T R h1 l. induction l as [|a l ih]; simpl; auto. intro b. destruct (sort_dec_tso_dec R h1 a b) as [h4 | h4]. destruct h4 as [h4 | h4]; simpl. destruct (nil_dec' (wlb_list R (sort_dec_tso_dec R h1) l b)) as [h5 | h5]. rewrite h5; auto. rewrite (hd_compat' _ h5). pose proof (coarse_list_ub_list R (sort_dec_tso_dec R h1) (tl (wlb_list R (sort_dec_tso_dec R h1) l b)) (hd' (wlb_list R (sort_dec_tso_dec R h1) l b) h5) ) as h6. pose proof (length_coarse_list _ _ h6) as h7. apply le_lt_n_Sm; auto. subst. pose proof (coarse_list_ub_list R (sort_dec_tso_dec R h1) (wlb_list R (sort_dec_tso_dec R h1) l b) b) as h6. pose proof (length_coarse_list _ _ h6) as h7. apply le_lt_n_Sm; auto. pose proof (coarse_list_ub_list R (sort_dec_tso_dec R h1) (wlb_list R (sort_dec_tso_dec R h1) l b) a) as h6. pose proof (length_coarse_list _ _ h6) as h7. apply le_lt_n_Sm; auto. Qed. Fixpoint increasingR {T:Type} (R:Relation T) (l:list T) : Prop := match l with | nil => True | a::l' => match l' with | nil => True | a'::l'' => R a a' /\ increasingR R l' end end. Fixpoint weak_increasingR {T:Type} (R:Relation T) (l:list T) : Prop := match l with | nil => True | a::l' => match l' with | nil => True | a'::l'' => Rw R a a' /\ weak_increasingR R l' end end. Lemma increasingR_cons : forall {T:Type} (R:Relation T) (l:list T) (a:T), increasingR R (a::l) -> increasingR R l. intros T R l a h1. simpl in h1. revert a h1. destruct l as [|a l']; simpl; auto. intros a' h1. destruct h1; auto. Qed. Lemma weak_increasingR_cons : forall {T:Type} (R:Relation T) (l:list T) (a:T), weak_increasingR R (a::l) -> weak_increasingR R l. intros T R l a h1. simpl in h1. revert a h1. destruct l as [|a l']; simpl; auto. intros a' h1. destruct h1; auto. Qed. Lemma increasingR_impl_weak : forall {T:Type} (R:Relation T) (l:list T), increasingR R l -> weak_increasingR R l. intros T R l. induction l as [|a l ih]; simpl; auto. intro h1. destruct l as [|c l]; simpl; auto. destruct h1 as [h1 h2]. specialize (ih h2). split; auto; red. left; auto. Qed. Lemma lt_increasingR_cons : forall {T:Type} (R:Relation T) (l:list T) (a m:T), Transitive _ R -> In a l -> increasingR R (m::l)-> R m a. intros T R l. induction l as [|a l ih]; simpl. intros; contradiction. intros a' m ht h1 h2. destruct h2 as [h2 h3]. destruct h1 as [h1 | h1]; subst; auto. specialize (ih _ _ ht h1 h3). eapply ht. apply h2. assumption. Qed. Lemma le_weak_increasingR_cons : forall {T:Type} (R:Relation T) (l:list T) (a m:T), Transitive _ R -> In a l -> weak_increasingR R (m::l)-> Rw R m a. intros T R l. induction l as [|a l ih]; simpl. intros; contradiction. intros a' m ht h1 h2. destruct h2 as [h2 h3]. destruct h1 as [h1 | h1]; subst; auto. specialize (ih _ _ ht h1 h3). eapply trans_Rw; auto. apply h2. assumption. Qed. Lemma increasingR_cons_impl : forall {T:Type} (R:Relation T) (l:list T) (a:T), increasingR R (a::l) -> increasingR R l. intros T R l a; destruct l; simpl. intros; auto. intros h1. destruct h1; auto. Qed. Lemma weak_increasingR_cons_impl : forall {T:Type} (R:Relation T) (l:list T) (a:T), weak_increasingR R (a::l) -> weak_increasingR R l. intros T R l; destruct l; simpl. intros; auto. intros ? h1. destruct h1; auto. Qed. Lemma increasingR_cons_impl_cons : forall {T:Type} (R:Relation T) (l:list T) (a d:T), increasingR R (a::l) -> R d a -> increasingR R (d::a::l). intros l a d h1 h2. simpl. split; auto. Qed. Lemma weak_increasingR_cons_impl_cons : forall {T:Type} (R:Relation T) (l:list T) (a d:T), weak_increasingR R (a::l) -> Rw R d a -> weak_increasingR R (d::a::l). intros T R l a d h1 h2. simpl. split; auto. Qed. Lemma increasingR_impl_cons : forall {T:Type} (R:Relation T) (l:list T) (pf:l <> nil) (d:T), increasingR R l -> R d (hd' l pf) -> increasingR R (d::l). intros T R l. destruct l as [|h l]; simpl. intros; auto. intros h1 d h2 h3. split; auto. Qed. Lemma weak_increasingR_impl_cons : forall {T:Type} (R:Relation T) (l:list T) (pf:l <> nil) (d:T), weak_increasingR R l -> Rw R d (hd' l pf) -> weak_increasingR R (d::l). intros T R l. destruct l as [|h l]; simpl. intros; auto. intros h1 d h2 h3. split; auto. Qed. Lemma increasingR_impl_cons' : forall {T:Type} (R:Relation T) (l:list T) (d:T), increasingR R l -> (forall c:T, In c l -> R d c) -> increasingR R (d::l). intros T R l. induction l as [|c l ih]; simpl. intros; auto. intros d h1 h2. split; auto. Qed. Lemma weak_increasingR_impl_cons' : forall {T:Type} (R:Relation T) (l:list T) (d:T), weak_increasingR R l -> (forall c:T, In c l -> Rw R d c) -> weak_increasingR R (d::l). intros T R l. induction l as [|c l ih]; simpl. intros; auto. intros d h1 h2. split; auto. Qed. Lemma increasingR_tl : forall {T:Type} (R:Relation T) (l:list T), increasingR R l -> increasingR R (tl l). intros T R l. destruct l as [|a l]; simpl; auto. intro h1. eapply increasingR_cons_impl. apply h1. Qed. Lemma weak_increasingR_tl : forall {T:Type} (R:Relation T) (l:list T), weak_increasingR R l -> weak_increasingR R (tl l). intros T R l. destruct l as [|a l]; simpl; auto. intro h1. eapply weak_increasingR_cons_impl. apply h1. Qed. Lemma increasingR_remove : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (l:list T) (a:T), Transitive _ R -> increasingR R l -> increasingR R (remove pfdt a l). intros T hdt R l. induction l as [|c l ih]; simpl; auto. intros a ht h1. destruct (hdt a c) as [h2 | h2]; subst. apply ih; auto. eapply increasingR_cons_impl; auto. apply h1. simpl. pose proof (increasingR_cons_impl R _ _ h1)as h3. apply increasingR_impl_cons'. apply ih; auto. intros d h4. pose proof (remove_In hdt l a) as h5. assert (h6:d <> a). intro; subst; contradiction. pose proof (in_remove_neq_in_l hdt l d a h4 h6) as h7. eapply lt_increasingR_cons. assumption. apply h7. assumption. Qed. Lemma weak_increasingR_remove : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (l:list T) (a:T), Transitive _ R -> weak_increasingR R l -> weak_increasingR R (remove pfdt a l). intros T hdt R l. induction l as [|c l ih]; simpl; auto. intros a ht h1. destruct (hdt a c) as [h2 | h2]; subst. apply ih; auto. eapply weak_increasingR_cons_impl; auto. apply h1. simpl. pose proof (weak_increasingR_cons_impl R _ _ h1)as h3. apply weak_increasingR_impl_cons'. apply ih; auto. intros d h4. pose proof (remove_In hdt l a) as h5. assert (h6:d <> a). intro; subst; contradiction. pose proof (in_remove_neq_in_l hdt l d a h4 h6) as h7. eapply le_weak_increasingR_cons. assumption. apply h7. assumption. Qed. Lemma incrR_list_minus : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (l l':list T), Transitive _ R -> increasingR R l -> let lm := list_minus pfdt l l' in increasingR R lm. intros T hdt R l l'. revert l. induction l' as [|a l ih]; simpl; auto. intros. rewrite list_minus_nil2; auto. intros l' ht h2. rewrite list_minus_cons_remove. apply increasingR_remove. assumption. apply ih; auto. Qed. Lemma weak_incrR_list_minus : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (l l':list T), Transitive _ R -> weak_increasingR R l -> let lm := list_minus pfdt l l' in weak_increasingR R lm. intros T hdt R l l'. revert l. induction l' as [|a l ih]; simpl; auto. intros. rewrite list_minus_nil2; auto. intros l' ht h2. rewrite list_minus_cons_remove. apply weak_increasingR_remove. assumption. apply ih; auto. Qed. Lemma increasingR_cons_lt_subst : forall {T:Type} (R:Relation T) (l:list T) (a b:T), Transitive _ R -> increasingR R (a::l) -> R b a -> increasingR R (b::l). intros T R l a b ht h1 h2. destruct l as [|c l]; simpl; auto. simpl in h1. destruct h1 as [h1 h3]. split; auto. eapply ht. apply h2. assumption. Qed. Lemma weak_increasingR_cons_le_subst : forall {T:Type} (R:Relation T) (l:list T) (a b:T), Transitive _ R -> weak_increasingR R (a::l) -> Rw R b a -> weak_increasingR R (b::l). intros T R l a b ht h1 h2. destruct l as [|c l]; simpl; auto. simpl in h1. destruct h1 as [h1 h3]. split; auto. eapply trans_Rw; auto. apply h2. assumption. Qed. Lemma incrR_dec : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), {increasingR R l} + {~increasingR R l}. intros T R h1 l. induction l as [|a l ih]; simpl. left; auto. destruct ih as [ih | ih]. destruct l as [|b l]. left; auto. destruct (sort_dec_tso_dec _ h1 a b) as [h2 | h2]. destruct h2 as [h2 |]; subst. left; split; auto. right; intro h2. destruct h2 as [h2]. apply (sort_dec_irrefl _ h1) in h2; auto. right. intro h3. destruct h3 as [h3 h4]. pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h5. apply (sort_dec_irrefl _ h1) in h5; auto. destruct l as [|b l]. left; auto. right. intro h2. destruct h2; contradiction. Qed. Lemma weak_incrR_dec : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), {weak_increasingR R l} + {~weak_increasingR R l}. intros T R h1 l. induction l as [|a l ih]; simpl. left; auto. destruct ih as [ih | ih]. destruct l as [|b l]. left; auto. destruct (sort_dec_tso_dec _ h1 a b) as [h2 | h2]. destruct h2 as [h2 |]; subst. left; split; auto. left; auto. left; split; auto. right; auto. right. intro h3. destruct h3 as [h3 h4]. pose proof (trans_Rw_R R (sort_dec_trans _ h1) _ _ _ h3 h2) as h5. apply (sort_dec_irrefl _ h1) in h5; auto. destruct l as [|b l]. left; auto. right. intro h2. destruct h2; contradiction. Qed. Lemma increasingR_cons_no_dup : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (l:list T) (x:T), Strict_Order R -> increasingR R (x::l) -> NoDup l. intros T hdt R l x hso. revert x. inversion hso as [ht hi]. induction l as [|a l ih]; simpl. intros; constructor. intros x h1. destruct h1 as [h1 h2]. specialize (ih _ h2). constructor; auto. intro h3. pose proof (interp_in_eq hdt _ _ h3) as h5. unfold interp_in in h5; simpl in h5. destruct l as [|c l]. simpl in h3; auto. simpl in h5. unfold lind in h5; simpl in h5. destruct (hdt a c) as [h6 | h6]; subst. simpl in h5. destruct h2 as [h2 h6]. apply hi in h2; auto. simpl in h5. inversion h5 as [h8]. assert (h9:In a l). rewrite h8. apply in_app_r. left; auto. simpl in h2. destruct h2 as [h2 h4]. destruct h3 as [h3 | h3]. rewrite h3 in h2. apply hi in h2; auto. pose proof (lt_increasingR_cons R _ _ _ ht h9 h4) as h10. pose proof (ht _ _ _ h2 h10) as h11. apply hi in h11; auto. Qed. Lemma increasingR_cons_nin : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (l:list T) (a:T), Strict_Order R -> increasingR R (a::l) -> ~In a l. intros T hdt R l a hso. revert a. inversion hso as [ht hi]. induction l as [|d l ih]; simpl. intros; auto. intros a h1. destruct h1 as [h1 h3]. intro h2. destruct h2 as [h2 | h2]; subst. apply hi in h1; auto. pose proof (interp_in_eq hdt _ _ h2) as h4. unfold interp_in in h4; simpl in h4. destruct l as [|c l]. simpl in h2; auto. simpl in h4. unfold lind in h4; simpl in h4. destruct (hdt a c) as [h6 | h6]; subst. simpl in h4. destruct h3 as [h3 h6]. pose proof (ht _ _ _ h1 h3) as h7. apply hi in h7; auto. simpl in h4. inversion h4 as [h8]. clear h4. assert (h9:In a l). rewrite h8. apply in_app_r. left; auto. simpl in h2, h3. destruct h3 as [h3 h4]. destruct h2 as [h2 | h2]. clear h8. subst. contradict h6; auto. pose proof (lt_increasingR_cons R _ _ _ ht h9 h4) as h10. pose proof (ht _ _ _ h3 h10) as h11. pose proof (ht _ _ _ h1 h11) as h12. apply hi in h12; auto. Qed. Lemma increasingR_no_dup : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (l:list T), Strict_Order R -> increasingR R l -> NoDup l. intros T hdt R l hso. destruct l as [|a l]; simpl. intros; constructor. intro h1. constructor. eapply increasingR_cons_nin; auto. apply hso. apply h1. eapply increasingR_cons_no_dup; auto. apply hso. apply h1. Qed. Lemma increasingR_firstn : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (l:list T) (m:nat), Transitive _ R -> increasingR R l -> increasingR R (firstn m l). intros T hdt R l. induction l as [|c l ih]; simpl. intros. rewrite firstn_nil. red; auto. intros m ht h1; simpl. destruct m as [|m]; auto. simpl; auto. pose proof (increasingR_cons_impl R _ _ h1) as h2. specialize (ih m ht h2). simpl. apply increasingR_impl_cons'; auto. intros a h3. pose proof (in_firstn _ _ _ h3) as h4. eapply lt_increasingR_cons; auto. apply h4. assumption. Qed. Lemma weak_increasingR_firstn : forall {T:Type} (R:Relation T) (l:list T) (m:nat), Transitive _ R -> weak_increasingR R l -> weak_increasingR R (firstn m l). intros T R l m ht. revert m. induction l as [|c l ih]; simpl. intros. rewrite firstn_nil. red; auto. intros m h1; simpl. destruct m as [|m]; auto. simpl; auto. pose proof (weak_increasingR_cons_impl R _ _ h1) as h2. specialize (ih m h2). simpl. apply weak_increasingR_impl_cons'; auto. intros a h3. pose proof (in_firstn _ _ _ h3) as h4. eapply le_weak_increasingR_cons. assumption. apply h4. assumption. Qed. Lemma increasingR_skipn : forall {T:Type} (R:Relation T) (l:list T) (m:nat), increasingR R l -> increasingR R (skipn m l). intros T R l. induction l as [|c l ih]; simpl. intros. rewrite skipn_nil. red; auto. intros m h1; simpl. destruct m as [|m]; auto. simpl; auto. pose proof (increasingR_cons_impl R _ _ h1) as h2. specialize (ih m h2). simpl. assumption. Qed. Lemma weak_increasingR_skipn : forall {T:Type} (R:Relation T) (l:list T) (m:nat), weak_increasingR R l -> weak_increasingR R (skipn m l). intros T R l. induction l as [|c l ih]; simpl. intros. rewrite skipn_nil. red; auto. intros m h1; simpl. destruct m as [|m]; auto. simpl; auto. pose proof (weak_increasingR_cons_impl R _ _ h1) as h2. specialize (ih m h2). simpl. assumption. Qed. Lemma increasingR_app_iff : forall {T:Type} (R:Relation T), Transitive _ R -> forall (l l':list T) (pfl:l <> nil) (pfl':l' <> nil), (increasingR R l /\ increasingR R l' /\ R (lst l pfl) (hd' l' pfl')) <-> increasingR R (l++l'). intros T R ht l. induction l as [|a l ih]. intros l h1. contradict h1; auto. intros l' h1 h2. split. intros h345. destruct h345 as [h3 [h4 h5]]. destruct l as [|b l]; simpl. simpl in h5. destruct l' as [|c l']; simpl; auto. simpl in h3, h5. split. destruct h3; auto. destruct h3 as [h3 h6]. specialize (ih l' (cons_neq_nil _ _) h2). let P := type of ih in match P with ?pf <-> _ => assert (h7:pf) end. repeat split; auto. simpl. destruct l as [|c l]; auto. erewrite last_indep in h5. apply h5. apply cons_neq_nil. rewrite ih in h7. destruct l as [|c l]; simpl. destruct l' as [|d l']; simpl; auto. split. destruct h6; auto. rewrite <- app_comm_cons in h7. apply increasingR_cons in h7. assumption. intro h3. pose proof (increasingR_cons _ _ _ h3) as h4. destruct l as [|b l]; simpl. simpl in ih, h3. destruct l' as [|c l']. contradict h2; auto. split; auto. split; auto. simpl. destruct h3; auto. specialize (ih l' (cons_neq_nil _ _) h2). rewrite <- ih in h4. destruct h4 as [h4 [h5 h6]]. simpl in h6. split; auto. split. pose proof (lt_increasingR_cons _ _ _ _ ht (in_eq _ _) h3); auto. assumption. split; auto. destruct l as [|c l]; auto. erewrite last_indep. apply h6. apply cons_neq_nil. Qed. Lemma increasingR_last : forall {T:Type} (R:Relation T) (l:list T), Transitive _ R -> increasingR R l -> forall x:T, In x l -> Rw R x (last l x). intros T R l. induction l as [|a l ih]. intros; contradiction. intros htr h1 x h2. erewrite last_indep; try apply cons_neq_nil. rewrite last_cons. pose proof (increasingR_cons R _ _ h1) as h3. specialize (ih htr h3). destruct h2 as [|h2]; subst. simpl in h1. destruct l as [|c l]. simpl. right; auto. destruct h1 as [h1 h2]. left. eapply trans_R_Rw; auto. apply h1. erewrite last_indep. apply ih. left; auto. apply cons_neq_nil. erewrite last_indep. apply ih; auto. eapply in_not_nil. apply h2. Qed. Lemma increasingR_last' : forall {T:Type} (R:Relation T) (pftr:Transitive _ R) (l:list T), increasingR R l -> forall (x:T), In x l -> x <> last l x -> R x (last l x). intros T R htr l h2 x h3 h4. pose proof (increasingR_last R l htr h2 _ h3) as h5. destruct h5 as [h5 | h5]; auto; try contradiction. Qed. Lemma weak_increasingR_last : forall {T:Type} (R:Relation T) (l:list T), Transitive _ R -> weak_increasingR R l -> forall (x:T) (pf:In x l), Rw R x (last l x). intros T R l. induction l as [|a l ih]. intros; contradiction. intros htr h1 x h2. erewrite last_indep; try apply cons_neq_nil. rewrite last_cons. destruct h2 as [|h2]; subst. pose proof (weak_increasingR_cons R _ _ h1) as h3. specialize (ih htr h3). simpl in h1. destruct l as [|c l]. simpl. right; auto. destruct h1 as [h1 h2]. destruct h1 as [h1 | h1]. specialize (ih _ (in_eq _ _)). left. erewrite last_indep. apply (trans_R_Rw _ htr _ _ _ h1 ih). intro; discriminate. subst. apply ih; left; auto. erewrite last_indep. apply ih; auto. eapply weak_increasingR_cons. apply h1. eapply in_not_nil. apply h2. Qed. Lemma weak_increasingR_app_iff : forall {T:Type} (R:Relation T), Transitive _ R -> forall (l l':list T) (pfl:l <> nil) (pfl':l' <> nil), (weak_increasingR R l /\ weak_increasingR R l' /\ Rw R (lst l pfl) (hd' l' pfl')) <-> weak_increasingR R (l++l'). intros T R h0 l. induction l as [|a l ih]. intros l h1. contradict h1; auto. intros l' h1 h2. split; intro h3. destruct h3 as [h3 [h4 h5]]. destruct l as [|b l]; simpl. simpl in h5. destruct l' as [|c l']; simpl; auto. simpl in h3, h5. split. destruct h3; auto. destruct h3 as [h3 h6]. specialize (ih l' (cons_neq_nil _ _) h2). let P := type of ih in match P with ?pf <-> _ => assert (h7:pf) end. simpl; repeat split; auto. destruct l as [|l c]; auto. erewrite last_indep. apply h5. apply cons_neq_nil. rewrite ih in h7. simpl in h7. assumption. pose proof (weak_increasingR_cons _ _ _ h3) as h4. destruct l as [|b l]. simpl; split; auto. split; auto. simpl in h3. pose proof (hd_compat' _ h2) as h5. simpl in h5. rewrite h5 in h3. destruct h3; auto. specialize (ih l' (cons_neq_nil _ _) h2). rewrite <- ih in h4. destruct h4 as [h4 [h5 h6]]. pose proof (le_weak_increasingR_cons _ _ _ _ h0 (in_eq _ _) h3) as h7. simpl. simpl in h6. repeat split; auto. erewrite (last_indep (b::l) (cons_neq_nil _ _) a) at 1. apply h6. Qed. Lemma weak_incrR_no_dup_impl : forall {T:Type} (R:Relation T) (l:list T), weak_increasingR R l -> NoDup l -> increasingR R l. intros T R l. induction l as [|a l ih]; simpl; auto. destruct l as [|b l]; simpl; auto. intro h2. simpl in ih. destruct h2 as [h2 h3]. specialize (ih h3); intro h4; split; auto. red in h2. destruct h2; subst; auto. apply no_dup_cons_nin in h4. contradict h4; left; auto. apply ih. apply no_dup_cons in h4; auto. Qed. Lemma increasingR_map : forall {T U:Type} (R:Relation T) (R':Relation U) (f:T->U) (l:list T), incrR_fun R R' f -> increasingR R l -> increasingR R' (map f l). intros T U R R' f l. revert f. induction l as [|a l ih]; simpl; auto. intros f h1 h2. specialize (ih f h1). hfirst ih. eapply increasingR_cons_impl. apply h2. specialize (ih hf). destruct l as [|c l]. simpl; auto. destruct h2 as [h2 h3]. simpl. split. apply h1; auto. simpl in ih; auto. Qed. Lemma weak_increasingR_map : forall {T U:Type} (R:Relation T) (R':Relation U) (f:T->U) (l:list T), weak_incrR_fun R R' f -> weak_increasingR R l -> weak_increasingR R' (map f l). intros T U R R' f l. revert f. induction l as [|a l ih]; simpl; auto. intros f h1 h2. specialize (ih f h1). hfirst ih. eapply weak_increasingR_cons_impl. apply h2. specialize (ih hf). destruct l as [|c l]. simpl; auto. destruct h2 as [h2 h3]. red in h2. destruct h2 as [h2 | h2]; subst. simpl. split. apply h1; auto. simpl in ih; auto. simpl. split; auto. right; auto. Qed. Lemma weak_incrR_lrep : forall {T:Type} (R:Relation T) (a:T) (n:nat), weak_increasingR R (lrep a n). intros ? ? ? n. induction n as [|n ih]; simpl; auto. destruct n as [|n]; simpl; auto. split. right; auto. assumption. Qed. Lemma incrR_all_sing : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (pf:Strict_Order R) (l:list T) (x:T), all_sing l x -> increasingR R l -> l = nil \/ l = x::nil. intros T hdt R hso l. induction l as [|c l ih]; simpl. intros; left; auto. intros m h1 h2. right. pose proof (h1 _ (in_eq _ _)); subst. f_equal. destruct (nil_dec l) as [|h3]; auto. destruct h3 as [a h3]. specialize (h1 _ (in_cons _ _ _ h3)); subst. apply increasingR_cons_nin in h2; auto. contradiction. Qed. Lemma incrR_no_dup_minimal : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (pft:sort_dec R) (l l':list T), increasingR R l -> NoDup l' -> list_to_set l = list_to_set l' -> le_listR R (sort_dec_tso_dec _ pft) l' l -> l = l'. intros T hdt R ht l l'. revert l. pose proof (sort_dec_so _ ht) as hso. destruct ht as [hts htr hir]. induction l' as [|a l' ih]; simpl. intros ? ? ? h1. apply empty_set_nil in h1; auto. intros l h1 h0 h2 h3. destruct l as [|c l]. contradiction. pose proof (increasingR_no_dup hdt _ _ hso h1) as hn. pose proof (no_dup_cons_nin _ _ hn) as hnin. pose proof (no_dup_cons_nin _ _ h0) as hnin'. pose proof (no_dup_cons _ _ h0) as hnd. unfold sort_dec_tso_dec in h3; simpl in h3. destruct (hts a c) as [h4 | h4]. destruct h4 as [h4 | h4]. assert (h5:In a l). pose proof (Add_intro2 _ (list_to_set l') a) as h5. rewrite <- h2 in h5. rewrite <- list_to_set_in_iff in h5. destruct h5 as [h5 | h5]; subst. apply hir in h4; auto. contradiction. assumption. pose proof (lt_increasingR_cons R _ _ _ htr h5 h1) as h6. pose proof (htr _ _ _ h6 h4) as hi. apply hir in hi; contradiction. subst. simpl in h2. apply add_nin_inj_eq in h2. f_equal. apply ih; auto. eapply increasingR_cons. apply h1. contradict hnin. rewrite list_to_set_in_iff; auto. contradict hnin'. rewrite list_to_set_in_iff; auto. contradiction. Qed. Lemma weak_incrR_length_count_occ_eq : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (pfso:sort_dec R) (l l':list T), weak_increasingR R l -> weak_increasingR R l' -> length l = length l' -> (forall x:T, count_occ pfdt l x = count_occ pfdt l' x) -> l = l'. intros T h1 R h2 l. induction l as [|a l ih]; simpl. intros l ? h3 h4 h5. symmetry in h4. apply length_eq0 in h4; auto. intros l' h3 h4 h5 h6. destruct l' as [|a' l']. simpl in h5. discriminate. simpl in h5. apply S_inj in h5. destruct (h1 a a') as [h7 | h7]. subst. f_equal. apply ih. eapply weak_increasingR_cons. apply h3. eapply weak_increasingR_cons. apply h4. assumption. intro x. specialize (h6 x). destruct (h1 a' x) as [h7 | h7]; subst. simpl in h6. destruct (h1 x x) as [h7 | h7]; auto with arith. contradict h7; auto. simpl in h6. destruct (h1 a' x) as [h8 | h8]. contradiction. assumption. destruct (in_dec h1 a' l) as [h8 | h8]. pose proof (le_weak_increasingR_cons _ _ _ _ (sort_dec_trans _ h2) h8 h3) as h9. destruct (in_dec h1 a l') as [h10 | h10]. pose proof (le_weak_increasingR_cons _ _ _ _ (sort_dec_trans _ h2) h10 h4) as h11. pose proof (antisym_Rw R (sort_dec_so _ h2) _ _ h9 h11) as h12. contradiction. specialize (h6 a). destruct (h1 a a) as [h11 | h11]. simpl in h6. destruct (h1 a' a) as [h12 | h12]. subst; contradict h7; auto. rewrite (count_occ_In h1) in h10. omega. contradict h11; auto. specialize (h6 a'). simpl in h6. destruct (h1 a a') as [h9 | h9]. contradiction. destruct (h1 a' a') as [h10 | h10]. rewrite (count_occ_In h1) in h8. omega. contradict h10; auto. Qed. Lemma weak_incrR_coarse_list_cons_in : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (pft:Transitive _ R) (l l':list T) (a':T), weak_increasingR R (a'::l') -> coarse_list l (a'::l') -> forall (x:T), In x l -> Rw R a' x. intros T hdt R ht l l' a' h1 h2. revert h1. dependent induction h2; simpl; auto. intros; contradiction. intros h3 x h4. destruct (hdt x a') as [|he]; subst. right; auto. destruct l' as [|b l']. rewrite coarse_list_nil2_iff in h2; subst. simpl in h4. contradiction. destruct h3 as [h3 h5]. destruct (hdt x b) as [|he']; subst. assumption. pose proof (in_coarse_list l (b::l') _ h4 h2) as h7. destruct h7 as [|h7]; subst. contradict he'; auto. eapply le_weak_increasingR_cons in h5. eapply trans_Rw; auto. apply h3. apply h5. assumption. assumption. intros h1 x h3. destruct h3 as [|h3]; subst. right; auto. pose proof (in_coarse_list l l' _ h3 h2) as h5. pose proof h5 as h5'. rewrite (hd_compat' _ (in_not_nil _ _ h5)) in h1, h5'. destruct h1 as [h1 h4]. destruct (hdt x (hd' l' (in_not_nil _ _ h5))) as [he | he]. rewrite he; auto. destruct h5' as [h8 | h8]. rewrite h8 in he. contradict h3; auto. pose proof (le_weak_increasingR_cons R _ _ _ ht h8 h4) as h6. eapply trans_Rw. apply ht. apply h1. assumption. Qed. Lemma weak_incrR_coarse_list : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (pft:Transitive _ R) (l l':list T), weak_increasingR R l' -> coarse_list l l' -> weak_increasingR R l. intros T hdt R ht l l' h1 h2. revert h1. dependent induction h2; simpl; auto. intro h3; apply IHh2. destruct l' as [|a' l']; simpl; auto. destruct h3; auto. intro h3. destruct l' as [|a' l']; simpl; auto. destruct l as [|b l]; auto. apply not_coarse_list_cons_nil in h2; contradiction. destruct h3 as [h3 h4]. destruct (nil_dec' l) as [|h5]; subst; auto. rewrite (hd_compat' _ h5). split. eapply trans_Rw; auto. apply h3. eapply weak_incrR_coarse_list_cons_in; auto. apply h4. apply h2. apply in_hd'. rewrite <- hd_compat'. apply IHh2; auto. Qed. Lemma incrR_le_nth_lt : forall {T:Type} (R:Relation T) (pftr:Transitive _ R) (l:list T), increasingR R l -> forall (i j:nat) (pfij:i < j) (pfi:i < length l) (pfj:j < length l), R (nth_lt l i pfi) (nth_lt l j pfj). intros T R h1 l. induction l as [|a l ih]; simpl. intros; omega. intros h0 i j h2 h3 h4. destruct (nil_dec' l) as [|h5]; subst. simpl in h3, h4. omega. rewrite (hd_compat' _ h5) in h0. destruct h0 as [h0 h6]. rewrite <- (hd_compat' _ h5) in h6. destruct i as [|i], j as [|j]. omega. apply lt_S_n in h4. destruct j as [|j]. erewrite nth_indep. erewrite <- nth_lt_compat. rewrite nth_lt_0_hd'. erewrite hd_functional'. apply h0. assumption. specialize (ih h6 0 (S j) (lt_0_Sn _)). hfirst ih. omega. specialize (ih hf h4). rewrite nth_lt_0_hd' in ih. erewrite hd_functional' in h0. rewrite nth_lt_compat in ih. erewrite nth_indep in ih. eapply h1. apply h0. apply ih. assumption. omega. apply lt_S_n in h2. apply lt_S_n in h3. apply lt_S_n in h4. erewrite (nth_indep l a _ h3). erewrite (nth_indep l a _ h4). erewrite <- nth_lt_compat. erewrite <- nth_lt_compat. apply ih; auto with arith. Grab Existential Variables. assumption. assumption. assumption. Qed. Lemma weak_incrR_le_nth_lt : forall {T:Type} (R:Relation T) (pftr:Transitive _ R) (l:list T), weak_increasingR R l -> forall (i j:nat) (pfij:i < j) (pfi:i < length l) (pfj:j < length l), Rw R (nth_lt l i pfi) (nth_lt l j pfj). intros T R h1 l. induction l as [|a l ih]; simpl. intros; omega. intros h0 i j h2 h3 h4. destruct (nil_dec' l) as [|h5]; subst. simpl in h3, h4. omega. rewrite (hd_compat' _ h5) in h0. destruct h0 as [h0 h6]. rewrite <- (hd_compat' _ h5) in h6. destruct i as [|i], j as [|j]. right; auto. destruct h0 as [h0 | h0]. apply lt_S_n in h4. destruct j as [|j]. left. erewrite nth_indep. erewrite <- nth_lt_compat. rewrite nth_lt_0_hd'. erewrite hd_functional'. apply h0. assumption. specialize (ih h6 0 _ (lt_0_Sn _ ) (not_nil_lt _ h5) h4). rewrite nth_lt_0_hd' in ih. erewrite hd_functional' in h0. pose proof (trans_R_Rw _ h1 _ _ _ h0 ih) as h7. left; auto. rewrite nth_lt_compat in h7. erewrite nth_indep. apply h7. assumption. apply lt_S_n in h4. subst. destruct j as [|j]. right; auto. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite nth_lt_0_hd'. apply hd_functional'. assumption. erewrite nth_indep. erewrite <- nth_lt_compat . erewrite hd_functional'. rewrite <- (nth_lt_0_hd' l). apply ih; auto with arith. assumption. omega. apply lt_S_n in h3. apply lt_S_n in h4. erewrite (nth_indep l a _ h3). erewrite (nth_indep l a _ h4). erewrite <- nth_lt_compat. erewrite <- nth_lt_compat. apply ih; auto with arith. Grab Existential Variables. assumption. assumption. apply not_nil_lt; auto. assumption. assumption. assumption. Qed. Lemma weak_incrR_ub_list : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), weak_increasingR R l -> let pftso := sort_dec_tso_dec _ pfso in weak_increasingR R (ub_list R pftso l a). intros T R h1 l a h2 h3; eapply weak_incrR_coarse_list. eapply sort_dec_eq_dec. apply h1. apply sort_dec_trans; auto. apply h2. apply coarse_list_ub_list. Qed. Lemma weak_incrR_wub_list : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), weak_increasingR R l -> let pftso := sort_dec_tso_dec _ pfso in weak_increasingR R (wub_list R pftso l a). intros T R h1 l a h2 h3; eapply weak_incrR_coarse_list. eapply sort_dec_eq_dec. apply h1. apply sort_dec_trans; auto. apply h2. apply coarse_list_wub_list. Qed. Lemma weak_incrR_lb_list : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), weak_increasingR R l -> let pftso := sort_dec_tso_dec _ pfso in weak_increasingR R (lb_list R pftso l a). intros T R h1 l a h2 h3; eapply weak_incrR_coarse_list. eapply sort_dec_eq_dec. apply h1. apply sort_dec_trans; auto. apply h2. apply coarse_list_lb_list. Qed. Lemma weak_incrR_wlb_list : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), weak_increasingR R l -> let pftso := sort_dec_tso_dec _ pfso in weak_increasingR R (wlb_list R pftso l a). intros T R h1 l a h2 h3; eapply weak_incrR_coarse_list. eapply sort_dec_eq_dec. apply h1. apply sort_dec_trans; auto. apply h2. apply coarse_list_wlb_list. Qed. Lemma weak_incrR_app_sing_ub_list_eq_nin : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), ~In a l -> let pftso := sort_dec_tso_dec _ pfso in weak_increasingR R (l++a::nil) -> ub_list R pftso l a = l. intros T R h1 l. induction l as [|c l ih]; auto. intros a hnin htso h2. simpl. destruct (htso c a) as [h3 | h3]. destruct h3 as [h3 |]; subst; auto. f_equal. apply ih. contradict hnin; right; auto. rewrite <- app_comm_cons in h2. apply weak_increasingR_cons_impl in h2; auto. contradict hnin; left; auto. rewrite <- app_comm_cons in h2. pose proof (@le_weak_increasingR_cons). pose proof (le_weak_increasingR_cons R (l++a::nil) a c (sort_dec_trans _ h1) (in_app_r (in_eq _ _) _) h2) as h4. pose proof (trans_Rw_R _ (sort_dec_trans _ h1) _ _ _ h4 h3) as h5. apply (sort_dec_irrefl _ h1) in h5; contradiction. Qed. Lemma weak_incrR_app_sing_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), let pfdt := sort_dec_eq_dec _ pfso in weak_increasingR R (l++a::nil) <-> (weak_increasingR R l /\ (forall x:T, In x l -> Rw R x a)). intros T R h1 l a hdt. destruct l as [|b l]; simpl. intuition. split; intro h2. split. destruct l as [|c l]; auto. rewrite <- app_comm_cons in h2. destruct h2 as [h2 h3]. split; auto. rewrite app_comm_cons in h3. rewrite <- weak_increasingR_app_iff in h3. destruct h3; auto. apply sort_dec_trans; auto. intros x h3. generalize dependent l. intro l. revert x a b. induction l as [|c l ih]; simpl. intuition. subst; auto. intros x a b h3 h2. destruct h3 as [h3 h4]. pose proof (ih x a c h4) as ih'. destruct h2 as [|h2]; subst. pose proof (ih c a c) as h5. eapply (trans_Rw _ (sort_dec_trans _ h1)). apply h3. apply h5; auto. apply ih'; auto. destruct h2 as [h2 h3]. destruct l as [|c l]; simpl. split; auto. destruct h2 as [h2 h4]. split; auto . change (weak_increasingR R (c::l++a::nil)). rewrite app_comm_cons. rewrite <- weak_increasingR_app_iff. simpl; repeat split; auto. apply h3. right. destruct l as [|d l]. left; auto. right. apply in_last. apply cons_neq_nil. apply sort_dec_trans; auto. Grab Existential Variables. apply cons_neq_nil. apply cons_neq_nil. apply cons_neq_nil. apply cons_neq_nil. Qed. Lemma weak_incrR_app_sing_iff_in : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T) (pfin:In a l), let pfdt := sort_dec_eq_dec _ pfso in weak_increasingR R (l++a::nil) <-> (weak_increasingR R (firstn (lind pfdt l a pfin) l) /\ (let i := lind pfdt l a pfin in i > 0 -> Rw R (last (firstn i l) a) a) /\ skipn_all_min pfdt l a pfin = lrep a (count_occ pfdt l a)). intros T R h1 l a h2 h3. rewrite weak_incrR_app_sing_iff; auto. split; intro h4. destruct h4 as [h4 h5]. split. apply weak_increasingR_firstn; auto. apply sort_dec_trans; auto. split. intros i h6. apply h5. pose proof (in_last (firstn i l) a) as h7. hfirst h7. intro h8. apply firstn_eq_nil in h8. destruct h8 as [h8 | h8]; try omega. pose proof h2 as h2'. rewrite h8 in h2'; contradiction. specialize (h7 hf). rewrite in_firstn_iff' in h7. destruct h7; auto. symmetry. rewrite lrep_eq_iff. match goal with |- ?hyp1 /\ ?hyp2 => assert (h6:hyp1) end. intros x h6. unfold skipn_all_min in h6. rewrite (in_skipn_iff (sort_dec_eq_dec _ h1)) in h6. destruct h6 as [h6 [occ h7]]. destruct h7 as [h7 h8]. apply le_lt_eq_dec in h8. destruct h8 as [h8 | h8]. pose proof (weak_incrR_le_nth_lt R (sort_dec_trans _ h1) _ h4 _ _ h8 (lt_lind _ _ _ _) (lt_lind_occ _ _ _ _ _)) as h9. rewrite lind_compat, lind_occ_compat in h9. pose proof (h5 _ h6). eapply antisym_Rw. apply sort_dec_so. apply h1. assumption. assumption. rewrite (lind_eq_iff h3 _ _ h2 _ (lt_lind_occ _ _ _ _ _)) in h8. destruct h8 as [h8 h9]. rewrite lind_occ_compat in h8; auto. split; auto. pose proof (all_sing_length (sort_dec_eq_dec _ h1) _ _ h6) as h7. rewrite count_occ_skipn_all_min in h7; auto. destruct h4 as [h4 [h5 h6]]. split; auto. rewrite <- (firstn_skipn (lind h3 l a h2)). destruct (zerop (lind h3 l a h2)) as [h7 | h7]. rewrite h7. simpl. unfold skipn_all_min in h6. rewrite h7 in h6. simpl in h6. rewrite h6. apply weak_incrR_lrep. specialize (h5 h7). rewrite <- weak_increasingR_app_iff. split; auto. split. unfold skipn_all_min in h6. rewrite h6. apply weak_incrR_lrep. fold (skipn_all_min h3 l a h2). erewrite <- hd_hd_compat'. rewrite hd_skipn_all_min at 1. rewrite lst_compat'. erewrite last_indep. apply h5. intro h8. apply firstn_eq_nil in h8. destruct h8 as [h8 | h8]. omega. subst. contradiction. apply sort_dec_trans; auto. intros x h7. destruct (zerop (lind h3 l a h2)) as [h8 | h8]. unfold skipn_all_min in h6. rewrite h8 in h6. simpl in h6. rewrite h6 in h7. apply in_lrep' in h7. rewrite h7. right; auto. specialize (h5 h8). rewrite <- (firstn_skipn (lind h3 l a h2)) in h7. apply in_app_or in h7. destruct h7 as [h7 | h7]. pose proof (weak_increasingR_last _ _ (sort_dec_trans _ h1) h4 _ h7) as h9. eapply trans_Rw. apply sort_dec_trans; auto. apply h9. erewrite last_indep. apply h5. intro h10. apply firstn_eq_nil in h10. destruct h10 as [h10 | h10]. omega. subst. contradiction. unfold skipn_all_min in h6. rewrite h6 in h7. apply in_lrep' in h7; subst; right; auto. Grab Existential Variables. intro h8. rewrite skipn_eq_nil_iff in h8. destruct h8 as [|h8]; subst. contradiction. pose proof (lt_lind h3 l a h2). omega. intro h8. apply firstn_eq_nil in h8. destruct h8 as [h8|]; subst. omega. contradiction. eapply sort_dec_eq_dec. apply h1. Qed. Definition weak_incrR_add {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (a:T) : list T := ub_list R pftso l a ++ a :: wlb_list R pftso l a. Lemma weak_incrR_add_nil : forall {T:Type} (R:Relation T)(pftso:tso_dec R) (a:T), weak_incrR_add R pftso nil a = a :: nil. intros; unfold weak_incrR_add. rewrite ub_list_nil, wlb_list_nil; auto. Qed. Lemma weak_incrR_add_sing : forall {T:Type} (R:Relation T)(pftso:sort_dec R) (a b:T), let pftso := sort_dec_tso_dec _ pftso in weak_incrR_add R pftso (a::nil) b = (minR R pftso a b) :: (maxR R pftso a b) :: nil. intros T R h1 a b htso. unfold weak_incrR_add, minR, maxR. destruct (htso a b) as [h2 | h2]. destruct h2 as [h2|]; subst. unfold htso. rewrite ub_list_cons_gt, ub_list_nil; auto. repeat rewrite app_nil_l; simpl. f_equal. f_equal. fold htso. destruct (htso a b) as [h3 |h3]; auto. destruct h3 as [h3 | h3]; auto. subst. apply (sort_dec_irrefl _ h1) in h2; contradiction. pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. simpl. destruct (htso b b) as [h2 | h2]. destruct h2 as [h2 | h2]; auto. apply (sort_dec_irrefl _ h1) in h2; contradiction. unfold htso; rewrite ub_list_cons_lt, ub_list_nil, app_nil_l; simpl; auto. fold htso. destruct (htso a b) as [h3 |h3]; auto. f_equal. destruct h3 as [h3 | h3]. pose proof (sort_dec_trans _ h1 _ _ _ h2 h3) as h4. apply (sort_dec_irrefl _ h1) in h4; contradiction. reflexivity. Qed. Lemma in_weak_incrR_add_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a x:T), let pftso := sort_dec_tso_dec _ pfso in In x (weak_incrR_add R pftso l a) <-> (x = a \/ In x l). unfold weak_incrR_add; simpl; intros T R h1 l a x; split; intro h2. apply in_app_or in h2. destruct h2 as [h2 | h2]. rewrite in_ub_list_iff in h2. destruct h2 as [h2 h3]. right; auto. destruct h2 as [|h2]; subst. left; auto. rewrite in_wlb_list_iff in h2. destruct h2; right; auto. destruct h2 as [|h2]; subst. apply in_app_r. left; auto. destruct (sort_dec_tso_dec _ h1 x a) as [h3 | h3]. destruct h3 as [h3 | h3]. apply in_app_l. rewrite in_ub_list_iff. split; auto. subst. apply in_app_r. left; auto. apply in_app_r. right. rewrite in_wlb_list_iff. split; auto. left; auto. Qed. Lemma weak_incrR_add_cons_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a c:T), R a c -> let pftso := sort_dec_tso_dec _ pfso in weak_incrR_add R pftso (a::l) c = a :: weak_incrR_add R pftso l c. intros T R h1 l a c hac h2. unfold weak_incrR_add, minR. simpl. destruct (h2 a c) as [h3 | h3]. destruct h3 as [h3 | h3]. rewrite <- app_comm_cons; auto. subst. apply (sort_dec_irrefl _ h1) in hac. contradiction. pose proof (sort_dec_trans _ h1 _ _ _ hac h3) as h4. apply (sort_dec_irrefl _ h1) in h4. contradiction. Qed. Lemma weak_incrR_add_cons_same : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), let pftso := sort_dec_tso_dec _ pfso in weak_incrR_add R pftso (a::l) a = ub_list R pftso l a ++ a::a::wlb_list R pftso l a. intros T R h1 l a h2. unfold weak_incrR_add. rewrite ub_list_cons_same, wlb_list_cons_same; auto. apply sort_dec_irrefl; auto. apply sort_dec_irrefl; auto. Qed. Lemma weak_incrR_add_cons_gt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a c:T), R c a -> let pftso := sort_dec_tso_dec _ pfso in weak_incrR_add R pftso (a::l) c = ub_list R pftso l c ++ (c :: a :: wlb_list R pftso l c). intros T R h1 l a c h2 h3. unfold weak_incrR_add. unfold h3. rewrite ub_list_cons_lt, wlb_list_cons_lt; auto. Qed. Lemma weak_incrR_add_in : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), let pftso := sort_dec_tso_dec _ pfso in In a (weak_incrR_add R pftso l a). intros; unfold weak_incrR_add. apply in_app_r. left; auto. Qed. Lemma weak_incrR_add_in_l : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a x:T), In x l -> let pftso := sort_dec_tso_dec _ pfso in In x (weak_incrR_add R pftso l a). intros T R hso l a x h2 h3. unfold weak_incrR_add. destruct (h3 a x) as [h4 | h4]. destruct h4 as [h4 | h4]. apply in_app_r. right. unfold h3. rewrite in_wlb_list_iff. split; auto. left; auto. subst. apply in_app_r. left; auto. apply in_app_l. unfold h3. rewrite in_ub_list_iff. split; auto. Qed. Lemma length_weak_incrR_add : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), let pftso := sort_dec_tso_dec _ pfso in length (weak_incrR_add R pftso l a) = S (length l). intros; unfold weak_incrR_add; rewrite app_length. simpl. rewrite plus_S_shift_r. f_equal. symmetry. apply length_decompose_ub_wlb. Qed. Lemma count_occ_weak_incrR_add_eq : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), let pfdt := sort_dec_eq_dec _ pfso in let pftso := sort_dec_tso_dec _ pfso in count_occ pfdt (weak_incrR_add R pftso l a) a = S (count_occ pfdt l a). intros T R h1 l a h2 h3. unfold weak_incrR_add. rewrite count_occ_app. simpl. destruct (h2 a a) as [h4 | h4]. rewrite plus_S_shift_r. f_equal. symmetry; apply count_occ_ub_wlb_decompose. contradict h4; auto. Qed. Lemma count_occ_weak_incrR_add_neq : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a x:T), x <> a -> let pfdt := sort_dec_eq_dec _ pfso in let pftso := sort_dec_tso_dec _ pfso in count_occ pfdt (weak_incrR_add R pftso l a) x = count_occ pfdt l x. intros T R h1 l a x h2 h3 h4. unfold weak_incrR_add. rewrite count_occ_app. simpl. destruct (h3 a x) as [h5 | h5]. subst; contradict h2; auto. symmetry; apply count_occ_ub_wlb_decompose. Qed. Lemma coarse_list_weak_incrR_add_cons : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a c:T), let pftso := sort_dec_tso_dec _ pfso in coarse_list (weak_incrR_add R pftso l c) (weak_incrR_add R pftso (a::l) c). intros T R h1 l a c h2. unfold weak_incrR_add. simpl. destruct (h2 a c) as [h3 | h3]. destruct h3 as [h3 | h3]. rewrite <- app_comm_cons. apply coarse_list_cons1. apply coarse_list_refl. subst. rewrite <- (app_sing (c::wlb_list R h2 l c)). apply (coarse_list_app_interp_compat (sort_dec_eq_dec _ h1) (ub_list R h2 l c ++ c :: wlb_list R h2 l c) (c::wlb_list R h2 l c) (ub_list R h2 l c) (c::nil)). apply coarse_list_refl. rewrite <- (app_sing (a::wlb_list R h2 l c)). pose proof (app_assoc (ub_list R h2 l c) (c::nil) (wlb_list R h2 l c)) as h4. simpl in h4; simpl. rewrite h4. rewrite <- (app_sing (a::wlb_list R h2 l c)) at 1. pose proof (app_assoc (ub_list R h2 l c) (c::nil) (a::wlb_list R h2 l c)) as h5. simpl in h5; simpl. rewrite h5. rewrite <- (app_sing (wlb_list R h2 l c)) at 1. apply (coarse_list_app_interp_compat (sort_dec_eq_dec _ h1) ((ub_list R h2 l c ++ c :: nil) ++ wlb_list R h2 l c) (wlb_list R h2 l c) (ub_list R h2 l c ++ c::nil) (a::nil)). apply coarse_list_refl. Qed. Lemma weak_incrR_add_compat : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), weak_increasingR R l -> let pftso := sort_dec_tso_dec _ pfso in weak_increasingR R (weak_incrR_add R pftso l a). intros T R h1 l a h2 h3. unfold weak_incrR_add. destruct (nil_dec' (ub_list R h3 l a)) as [h4 | h4]. rewrite h4. rewrite app_nil_l. destruct (nil_dec' (wlb_list R h3 l a)) as [h5 | h5]. rewrite h5; simpl; auto. apply (weak_increasingR_impl_cons _ _ h5). eapply weak_incrR_coarse_list. eapply sort_dec_eq_dec. apply h1. apply sort_dec_trans; auto. apply h2. apply coarse_list_wlb_list. pose proof (in_hd' (wlb_list R h3 l a) h5) as h6. unfold h3 in h6. rewrite in_wlb_list_iff in h6. unfold h3. destruct h6; auto. destruct (nil_dec' (wlb_list R h3 l a)) as [h5 | h5]. rewrite h5. rewrite <- (weak_increasingR_app_iff R (sort_dec_trans _ h1) _ _ h4 (cons_neq_nil _ _)). repeat split. apply weak_incrR_ub_list; auto. simpl. pose proof (in_lst (ub_list R h3 l a) h4) as h6. unfold h3 in h6; rewrite in_ub_list_iff in h6. left. unfold h3. destruct h6; auto. rewrite <- (weak_increasingR_app_iff R (sort_dec_trans _ h1) _ _ h4 (cons_neq_nil _ _)). repeat split. apply weak_incrR_ub_list; auto. simpl. rewrite (hd_compat' _ h5). split. pose proof (in_hd' (wlb_list R h3 l a) h5) as h6. unfold h3 in h6; rewrite in_wlb_list_iff in h6. unfold h3. destruct h6; auto. rewrite <- hd_compat'. apply weak_incrR_wlb_list; auto. simpl. pose proof (in_lst (ub_list R h3 l a) h4) as h6. unfold h3 in h6; rewrite in_ub_list_iff in h6. left. unfold h3. destruct h6; auto. Qed. Fixpoint sort_list_aux {T:Type} (R:Relation T) (pftso:tso_dec R) (acc l:list T) : list T := match l with | nil => acc | a::l' => sort_list_aux R pftso (weak_incrR_add R pftso acc a) l' end. Definition sort_list {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) : list T := sort_list_aux R pftso nil l. Lemma sort_list_functional : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (pfso:sort_dec R) (l l':list T) (pfeq:l = l'), sort_list R pftso l = sort_list R (sort_dec_tso_dec _ pfso) l'. intros; subst. f_equal. symmetry; apply sort_dec_tso_dec_eq. Qed. Lemma sort_list_aux_nil : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (acc:list T), sort_list_aux R pftso acc nil = acc. auto. Qed. Lemma sort_list_nil : forall {T:Type} (R:Relation T) (pftso:tso_dec R), sort_list R pftso nil = nil. auto. Qed. (*See less trivial primed version.*) Lemma sort_list_cons : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T) (a:T), sort_list R pftso (a::l) = sort_list_aux R pftso (a::nil) l. auto. Qed. Lemma sort_list_sing : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (a:T), sort_list R pftso (a::nil) = a::nil. intros; unfold sort_list; simpl. rewrite weak_incrR_add_nil; auto. Qed. Lemma in_acc_sort_list_aux : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (acc l:list T) (x:T), In x acc -> let pftso := sort_dec_tso_dec _ pfso in In x (sort_list_aux R pftso acc l). intros T R h1 acc l. revert acc. induction l as [|b l ih]; simpl; auto. intros acc x h2. apply ih. eapply weak_incrR_add_in_l; auto. Qed. Lemma in_sort_list_aux_app_l : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (acc l l':list T) (x:T), let pftso := sort_dec_tso_dec _ pfso in In x l -> In x (sort_list_aux R pftso acc (l++l')). intros T R h1 acc l. revert acc. induction l as [|a l ih]; simpl. intros; contradiction. intros acc l' x h2. destruct h2 as [|h2]; subst. apply in_acc_sort_list_aux. apply weak_incrR_add_in. apply ih; auto. Qed. Lemma in_sort_list_aux : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (acc l:list T) (x:T), let pftso := sort_dec_tso_dec _ pfso in In x l -> In x (sort_list_aux R pftso acc l). intros; rewrite <- (app_nil_r l). apply in_sort_list_aux_app_l; auto. Qed. Lemma in_sort_list_aux_app_r : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (acc l l':list T) (x:T), let pftso := sort_dec_tso_dec _ pfso in In x l' -> In x (sort_list_aux R pftso acc (l++l')). intros T R h1 acc l. revert acc. induction l as [|a l ih]; simpl. intros; apply in_sort_list_aux; auto. intros; auto. Qed. Lemma in_sort_list_aux_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (acc l:list T) (x:T), let pftso := sort_dec_tso_dec _ pfso in In x (sort_list_aux R pftso acc l) <-> In x l \/ In x acc. intros T R h1 acc l x; split; intro h2. revert h2. revert x acc. induction l as [|a l ih]; simpl. intros; right; auto. intros x acc h2. apply ih in h2. destruct h2 as [h2 | h2]. left; right; auto. unfold weak_incrR_add in h2. apply in_app_or in h2. destruct h2 as [h2 | h2]. right. rewrite in_ub_list_iff in h2. destruct h2; auto. destruct h2 as [|h2]; subst. left; left; auto. right. rewrite in_wlb_list_iff in h2. destruct h2; auto. destruct h2 as [h2 | h2]. apply in_sort_list_aux; auto. apply in_acc_sort_list_aux; auto. Qed. Lemma in_sort_list_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (x:T), let pftso := sort_dec_tso_dec _ pfso in In x (sort_list R pftso l) <-> In x l. intros; unfold sort_list, pftso; rewrite in_sort_list_aux_iff; intuition. Qed. Lemma sort_list_aux_eq_nil_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (acc l:list T), let pftso := sort_dec_tso_dec _ pfso in sort_list_aux R pftso acc l = nil <-> l = nil /\ acc = nil. intros T R h1 acc l; split; intro h2. destruct (nil_dec' l) as [|h3]; subst. rewrite sort_list_aux_nil in h2. split; auto. destruct l as [|a l]; auto. rewrite (hd_compat' _ h3) in h2. simpl in h2. pose proof (weak_incrR_add_in R h1 acc a) as h4. simpl in h4. pose proof (in_acc_sort_list_aux R h1 _ l _ h4) as h5. simpl in h5. rewrite h2 in h5. contradiction. destruct h2; subst; auto. Qed. Lemma sort_list_eq_nil_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), let pftso := sort_dec_tso_dec _ pfso in sort_list R pftso l = nil <-> l = nil. simpl; intros; unfold sort_list. rewrite sort_list_aux_eq_nil_iff at 1. split; auto. intro h1; destruct h1; auto. Qed. Lemma weak_incrR_acc_sort_list_aux : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (acc l:list T), weak_increasingR R acc -> let pftso := sort_dec_tso_dec _ pfso in weak_increasingR R (sort_list_aux R pftso acc l). intros T R h1 acc l. revert acc. induction l as [|a l ih]; auto. intros acc h2 htso. apply ih. specialize (ih _ h2). unfold weak_incrR_add. pose proof (weak_incrR_coarse_list (sort_dec_eq_dec _ h1) _ (sort_dec_trans _ h1) _ _ h2 (coarse_list_ub_list R htso acc a)) as h3. pose proof (weak_incrR_coarse_list (sort_dec_eq_dec _ h1) _ (sort_dec_trans _ h1) _ _ h2 (coarse_list_wlb_list R htso acc a)) as h4. destruct (nil_dec' (ub_list R htso acc a)) as [h5 | h5]. rewrite h5. rewrite app_nil_l. unfold htso in h5; rewrite ub_list_eq_nil_iff in h5. destruct h5 as [|h5]; subst. simpl; auto. destruct h5 as [hn h5]. destruct (nil_dec' (wlb_list R htso acc a)) as [h6 | h6]. rewrite h6; simpl; auto. simpl. rewrite (hd_compat' _ h6). pose proof (in_hd' (wlb_list R htso acc a) h6) as h7. unfold htso in h7. rewrite in_wlb_list_iff in h7. unfold htso; destruct h7; auto. split; auto. rewrite <- hd_compat'. fold htso; auto. destruct (nil_dec' (wlb_list R htso acc a)) as [h6 | h6]. rewrite h6; simpl; auto. unfold htso in h6; rewrite wlb_list_eq_nil_iff in h6. destruct h6 as [|h6]; subst. simpl in h5. contradict h5; auto. destruct h6 as [hn h6]. erewrite <- (weak_increasingR_app_iff _ _ _ _ h5 (cons_neq_nil _ _)); simpl. split; auto. split; auto. left. apply h6. pose proof (in_lst (ub_list R htso acc a) h5) as h7. unfold htso in h7. rewrite in_ub_list_iff in h7. unfold htso; destruct h7; auto. rewrite <- weak_increasingR_app_iff. repeat split; auto. pose proof (in_hd' _ h6) as h7. unfold htso in h7. rewrite in_wlb_list_iff in h7. destruct h7 as [h7 h8]. simpl. rewrite (hd_compat' _ h6); simpl. split; auto. rewrite (hd_compat' _ h6) in h4. simpl in h4; auto. simpl. pose proof (in_lst _ h5) as h7. unfold htso in h7. rewrite in_ub_list_iff in h7. destruct h7 as [h7 h8]. left. apply h8. apply sort_dec_trans; auto. Grab Existential Variables. intro; discriminate. apply sort_dec_trans; auto. Qed. Lemma weak_incrR_sort_list : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), let pftso := sort_dec_tso_dec _ pfso in weak_increasingR R (sort_list R pftso l). intros T R h1 l. induction l as [|a l ih]; auto. unfold sort_list; simpl; auto. simpl; simpl in ih. apply weak_incrR_acc_sort_list_aux. red; auto. Qed. Lemma length_sort_list_aux : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (acc l:list T), length (sort_list_aux R pftso acc l) = length acc + length l. intros T R htso acc l. revert acc. induction l as [|a l ih]; simpl; auto with arith. intro acc. rewrite ih. unfold weak_incrR_add. rewrite app_length; simpl. rewrite (length_decompose_ub_wlb _ htso acc a). omega. Qed. Lemma length_sort_list : forall {T:Type} (R:Relation T) (pftso:tso_dec R) (l:list T), length (sort_list R pftso l) = length l. unfold sort_list; intros; rewrite length_sort_list_aux; auto. Qed. Lemma list_to_set_sort_list : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), let pftso := sort_dec_tso_dec _ pfso in list_to_set l = list_to_set (sort_list R pftso l). intros T R h1 l; apply Extensionality_Ensembles; red; split; red; intros x h2. rewrite <- list_to_set_in_iff in h2; rewrite <- list_to_set_in_iff. rewrite in_sort_list_iff; auto. rewrite <- list_to_set_in_iff in h2; rewrite <- list_to_set_in_iff. rewrite in_sort_list_iff in h2; auto. Qed. Lemma faml_to_fam_sort_list : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:faml T), let pfso' := lt_listR_sort_dec _ pfso in let pftso' := sort_dec_tso_dec _ pfso' in faml_to_fam (sort_list _ pftso' l) = faml_to_fam l. intros T R h1 ll; apply Extensionality_Ensembles; red; split; red; intros x h2; apply in_faml_to_fam_impl in h2; destruct h2 as [l' [h2 h3]]; subst. apply in_faml_impl. rewrite in_sort_list_iff in h2; auto. apply in_faml_impl. rewrite in_sort_list_iff; auto. Qed. Lemma count_occ_sort_list_aux : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (acc l:list T) (x:T), let pfdt := sort_dec_eq_dec _ pfso in let pftso := sort_dec_tso_dec _ pfso in count_occ pfdt (sort_list_aux R pftso acc l) x = count_occ pfdt l x + count_occ pfdt acc x. intros T R hso acc l. revert acc. induction l as [|a l ih]. intros; simpl; auto. intros acc x hdt htso. simpl. destruct (hdt a x) as [|h2]; subst. rewrite ih. unfold weak_incrR_add. rewrite count_occ_app. simpl. fold hdt. destruct (hdt x x) as [h2 | h2]. do 2 rewrite plus_S_shift_r. f_equal. f_equal. symmetry. apply count_occ_ub_wlb_decompose. contradict h2; auto. rewrite ih. unfold weak_incrR_add. rewrite count_occ_app; simpl. fold hdt. destruct (hdt a x) as [|h3]. contradiction. f_equal. symmetry. apply count_occ_ub_wlb_decompose. Qed. Lemma count_occ_sort_list : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (x:T), let pfdt := sort_dec_eq_dec _ pfso in let pftso := sort_dec_tso_dec _ pfso in count_occ pfdt (sort_list R pftso l) x = count_occ pfdt l x. simpl; intros; simpl; unfold sort_list; rewrite count_occ_sort_list_aux; simpl; auto with arith. Qed. Lemma weak_incrR_sort_list_eq : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), let pftso := sort_dec_tso_dec _ pfso in weak_increasingR R l -> sort_list R pftso l = l. intros T R h1 l htso h2. eapply weak_incrR_length_count_occ_eq. apply h1. apply weak_incrR_sort_list. assumption. apply length_sort_list. apply count_occ_sort_list. Qed. (*And the much more useful version!*) Lemma sort_list_cons' : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), let pftso := sort_dec_tso_dec _ pfso in sort_list R pftso (a::l) = weak_incrR_add R pftso (sort_list R pftso l) a. intros T R h1 l a htso. eapply (weak_incrR_length_count_occ_eq (sort_dec_eq_dec _ h1)). apply h1. apply weak_incrR_sort_list. apply weak_incrR_add_compat. apply weak_incrR_sort_list. rewrite length_sort_list. unfold htso. rewrite length_weak_incrR_add. rewrite length_sort_list; auto. simpl. intro x. unfold htso. rewrite count_occ_sort_list. destruct (sort_dec_eq_dec _ h1 x a) as [|h2]; subst. rewrite count_occ_weak_incrR_add_eq. rewrite count_occ_sort_list; simpl. destruct (sort_dec_eq_dec R h1 a a) as [|h2]; auto. contradict h2; auto. rewrite count_occ_weak_incrR_add_neq; auto. rewrite count_occ_sort_list; simpl. destruct (sort_dec_eq_dec R h1 a x) as [|h3]; auto. subst. contradict h2; auto. Qed. Lemma sort_list_decompose_ub_wlb : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T), let pftso := sort_dec_tso_dec _ pfso in sort_list R pftso l = (ub_list R pftso (sort_list R pftso l) a) ++ (wlb_list R pftso (sort_list R pftso l) a). intros T R h1 l a hts. eapply (weak_incrR_length_count_occ_eq (sort_dec_eq_dec _ h1)). apply h1. apply weak_incrR_sort_list. destruct (nil_dec' (ub_list R hts (sort_list R hts l) a)) as [h2 | h2]. rewrite h2. rewrite app_nil_l. eapply weak_incrR_coarse_list. eapply sort_dec_eq_dec. apply h1. apply sort_dec_trans; auto. apply weak_incrR_sort_list. apply coarse_list_wlb_list. destruct (nil_dec' (wlb_list R hts (sort_list R hts l) a)) as [h3 | h3]. rewrite h3. rewrite app_nil_r. eapply weak_incrR_coarse_list. eapply sort_dec_eq_dec. apply h1. apply sort_dec_trans; auto. apply weak_incrR_sort_list. apply coarse_list_ub_list. rewrite <- (weak_increasingR_app_iff _ (sort_dec_trans _ h1) _ _ h2 h3). split. eapply weak_incrR_coarse_list. eapply sort_dec_eq_dec. apply h1. apply sort_dec_trans; auto. apply weak_incrR_sort_list. apply coarse_list_ub_list. split. eapply weak_incrR_coarse_list. eapply sort_dec_eq_dec. apply h1. apply sort_dec_trans; auto. apply weak_incrR_sort_list. apply coarse_list_wlb_list. left. eapply in_ub_wlb_lt. apply in_lst. apply in_hd'. rewrite app_length. apply length_decompose_ub_wlb. intro x. rewrite count_occ_app. apply count_occ_ub_wlb_decompose. Qed. Lemma no_dup_sort_list : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), NoDup l -> let pftso := sort_dec_tso_dec _ pfso in NoDup (sort_list R pftso l). intros T R h1 l. induction l as [|a l ih]; simpl. rewrite sort_list_nil; auto. intro h2. pose proof (no_dup_cons_nin _ _ h2) as h3. pose proof (no_dup_cons _ _ h2) as h4. rewrite sort_list_cons'. unfold weak_incrR_add. apply no_dup_app. apply no_dup_ub_list. auto. constructor. contradict h3. rewrite in_wlb_list_iff in h3. destruct h3 as [h3]. rewrite in_sort_list_iff in h3; auto. apply no_dup_wlb_list. auto. apply not_inhabited_empty. intro h5. destruct h5 as [x h5]. destruct h5 as [x h5 h6]. rewrite <- list_to_set_in_iff in h5, h6. pose proof h5 as h5'. rewrite in_ub_list_iffn in h5'. destruct h6 as [h6 | h6]. subst. rewrite in_ub_list_iff in h5. destruct h5 as [? h5]. apply (sort_dec_irrefl _ h1) in h5; auto. contradiction. rewrite in_ub_list_iff in h5. destruct h5; auto. Qed. Lemma lt_length_sort_list_cons : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (a:T) (i:nat) (pf:i < length (sort_list R (sort_dec_tso_dec _ pfso) (a::l))), a <> nth_lt _ _ pf -> pred i < length (sort_list R (sort_dec_tso_dec _ pfso) l). intros T R h1 l a i h2 h3. destruct i as [|i]. revert h3. generalize h2. rewrite sort_list_cons'. intros h2' h3. destruct (nil_dec' l) as [|h4]; subst. simpl in h3. contradict h3; auto. simpl. rewrite length_sort_list. apply not_nil_lt in h4; auto. simpl. pose proof h2 as h2'. rewrite sort_list_cons' in h2'. rewrite length_weak_incrR_add in h2'. apply lt_S_n in h2'; auto. Qed. Fixpoint sort_list_perm {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (i:nat) : i < length (sort_list R (sort_dec_tso_dec _ pfso) l) -> nat := match l with | nil => fun pf => let pfn := length_functional _ _ (sort_list_nil R (sort_dec_tso_dec _ pfso)) in let pflt := lt_eq_trans _ _ _ pf pfn in False_rect nat (lt_n_0 _ pflt) | a::l' => fun pf:i < length (sort_list R (sort_dec_tso_dec _ pfso) (a::l')) => match (sort_dec_tso_dec _ pfso a (nth_lt _ _ pf)) with | inleft (left pflt) => let pfn := lt_irrefl_neq _ _ _ (sort_dec_irrefl _ pfso) pflt in let pflt' := lt_length_sort_list_cons R pfso l' a i pf pfn in S (sort_list_perm R pfso l' (pred i) pflt') | inleft (right pfe) => 0 | inright pfgt => let pfn := gt_irrefl_neq _ _ _ (sort_dec_irrefl _ pfso) pfgt in let pfgt' := lt_length_sort_list_cons R pfso l' a i pf pfn in (sort_list_perm R pfso l' (pred i) pfgt') end end. Lemma sort_list_perm_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (i:nat) (pf:i < length (sort_list R (sort_dec_tso_dec _ pfso) l)), sort_list_perm R pfso l i pf < length l. intros T R h1 l. induction l as [|a l ih]; simpl. intros ? h2. contradict (lt_n_0 _ h2). intros i h2. destruct (sort_dec_tso_dec _ h1 a (nth_lt _ _ h2)) as [h3|]; auto with arith. destruct h3 as [h3 | h3]; auto with arith. Qed. (*Returns a permutation from [sort_list l] to [l] Assumes [NoDup], but not in the definition -- only behavior*) Definition sort_list_perm_no_dup {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (i:nat) (pflt:i < length (sort_list R (sort_dec_tso_dec _ pfso) l)) : nat := let pfin := in_nth_lt _ _ pflt in let pfin' := iff1 (in_sort_list_iff _ pfso _ _) pfin in lind (sort_dec_eq_dec _ pfso) _ _ pfin'. (*Assumes [NoDup] semnatically, but not overtly.*) Definition list_perm_sort_no_dup {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (i:nat) (pflt:i < length l) : nat := let pfin := in_nth_lt _ _ pflt in let pfin' := iff2 (in_sort_list_iff _ pfso _ _) pfin in let pftso := sort_dec_tso_dec _ pfso in lind (sort_dec_eq_dec _ pfso) _ _ pfin'. (*Assumes [NoDup] semnatically, but not overtly.*) Lemma sort_list_perm_no_dup_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (i:nat) (pflt:i < length (sort_list R (sort_dec_tso_dec _ pfso) l)), sort_list_perm_no_dup R pfso l i pflt < length l. unfold sort_list_perm_no_dup; intros; eapply lt_eq_trans. apply lt_lind. reflexivity. Qed. (*Assumes [NoDup] semantically, but not overtly.*) Lemma list_perm_sort_no_dup_lt : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (i:nat) (pflt:i < length l), list_perm_sort_no_dup R pfso l i pflt < length (sort_list R (sort_dec_tso_dec _ pfso) l). unfold list_perm_sort_no_dup; intros; eapply lt_eq_trans. apply lt_lind. reflexivity. Qed. (*Assumes [NoDup] semantically, but not overtly.*) Lemma sort_list_perm_no_dup_compat : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (i:nat) (pflt:i < length (sort_list R (sort_dec_tso_dec _ pfso) l)), let pf' := sort_list_perm_no_dup_lt R pfso l i pflt in nth_lt (sort_list R (sort_dec_tso_dec _ pfso) l) i pflt = nth_lt l (sort_list_perm_no_dup R pfso l i pflt) (sort_list_perm_no_dup_lt _ _ _ _ _). unfold sort_list_perm_no_dup; intros. symmetry. erewrite nth_lt_functional3. rewrite lind_compat; auto. Qed. (*Assumes [NoDup] semantically, but not overtly.*) Lemma list_perm_sort_no_dup_compat : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T) (i:nat) (pflt:i < length l), let pf' := list_perm_sort_no_dup_lt R pfso l i pflt in nth_lt l i pflt = nth_lt (sort_list R (sort_dec_tso_dec _ pfso) l) (list_perm_sort_no_dup R pfso l i pflt) (list_perm_sort_no_dup_lt _ _ _ _ _). unfold list_perm_sort_no_dup; intros. symmetry. erewrite nth_lt_functional3. rewrite lind_compat; auto. Qed. Lemma list_to_set_injable_sort_list_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), let pftso := sort_dec_tso_dec _ pfso in list_to_set_injable (fun lx => lx = l) <-> list_to_set_injable (fun lx => lx = sort_list _ pftso l). intros; split; intro h3; red in h3; red; intros; subst; auto. Qed. Definition sort_set {T:Type} (R:Relation T) (pfso:sort_dec R) (E:Ensemble T) (pf:Finite E) : list T := let pfso' := lt_listR_sortable _ pfso in let pfts := sort_dec_tso_dec _ pfso in min_set_sortable (list_reps E) (finite_list_reps E pf) (inh_list_reps E pf) _ pfso'. Definition is_sorted_set {T:Type} (l:list T) : Prop := exists (R:Relation T) (pfso:sort_dec R), l = sort_set R pfso (list_to_set l) (list_to_set_finite l). Lemma sort_set_functional : forall {T:Type} (R:Relation T) (pfso pfso':sort_dec R) (E E':Ensemble T) (pf:Finite E) (pf':Finite E'), E = E' -> sort_set R pfso E pf = sort_set R pfso' E' pf'. intros; subst. f_equal. apply sort_dec_eq. apply proof_irrelevance. Qed. Lemma in_list_reps_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (E:Ensemble T) (pf:Finite E), Inn (list_reps E) (sort_set R pfso E pf). intros T R h0 E h1. unfold sort_set. match goal with |- Inn _ (min_set_sortable ?x1 ?x2 ?x3 ?x4 ?x5) => pose proof (min_set_sortable_compat x1 x2 x3 x4 x5) as h2 end. simpl in h2. destruct h2; auto. Qed. Lemma in_sort_set_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (E:Ensemble T) (pf:Finite E) (x:T), Inn E x <-> In x (sort_set R pfso E pf). intros T R hs0 E h0 x. split. intro h2. unfold sort_set. match goal with |- In _ (min_set_sortable ?x1 ?x2 ?x3 ?x4 ?x5) => pose proof (min_set_sortable_compat x1 x2 x3 x4 x5) as h1 end. destruct h1 as [h1 h3]. destruct h1 as [h1]. destruct h1 as [h1 h4]. rewrite h1 in h2. rewrite <- list_to_set_in_iff in h2. assumption. intro h2. unfold sort_set in h2. let P := type of h2 in match P with In _ (min_set_sortable ?x1 ?x2 ?x3 ?x4 ?x5) => pose proof (min_set_sortable_compat x1 x2 x3 x4 x5) as h1 end. simpl in h1. destruct h1 as [h1 h3]. eapply in_list_reps. apply h1. assumption. Qed. Lemma inj_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A B:Ensemble T) (pfa:Finite A) (pfb:Finite B), sort_set R pfso A pfa = sort_set R pfso B pfb -> A = B. intros T R hs A B h1 h2 h3. apply Extensionality_Ensembles; red; split; red; intros x h4. rewrite in_sort_set_iff in h4. rewrite h3 in h4 at 1. rewrite <- in_sort_set_iff in h4; auto. rewrite in_sort_set_iff in h4. rewrite <- h3 in h4 at 1. rewrite <- in_sort_set_iff in h4; auto. Qed. (*Search below for [inj_sort_set']*) Lemma no_dup_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (E:Ensemble T) (pf:Finite E), NoDup (sort_set R pfso E pf). intros T R hs E h1. unfold sort_set. match goal with |- NoDup (min_set_sortable ?pf1 ?pf2 ?pf3 ?pf4 ?pf5) => pose proof (min_set_sortable_compat pf1 pf2 pf3 pf4 pf5) as h2 end. simpl in h2. destruct h2 as [h2 h3]. inversion h2 as [h4]. destruct h4; auto. Qed. Lemma sort_set_compat : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (E:Ensemble T) (pf:Finite E), list_to_set (sort_set R pfso E pf) = E. intros T R hs E h0. apply Extensionality_Ensembles; red; split; red; intros x h2. rewrite <- list_to_set_in_iff in h2. eapply in_list_reps. apply in_list_reps_sort_set. apply h2. rewrite <- list_to_set_in_iff. rewrite <- in_sort_set_iff; auto. Qed. Lemma sort_set_empty : forall {T:Type} (R:Relation T) (pfso:sort_dec R), sort_set R pfso (Empty_set _) (Empty_is_finite _ ) = nil. unfold sort_set. intros T R hs. match goal with |- min_set_sortable ?pf1 ?pf2 ?pf3 ?pf4 ?pf5 = _ => pose proof (min_set_sortable_compat pf1 pf2 pf3 pf4 pf5) as h2 end. simpl in h2. destruct h2 as [h2]. rewrite list_reps_empty in h2 at 1. inversion h2; auto. Qed. Lemma sort_set_nil : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (pf:Finite A), sort_set R pfso A pf = nil -> A = Empty_set _. intros T R hs A h1 h2. apply Extensionality_Ensembles; red; split; auto with sets. red. intros x h3. erewrite in_sort_set_iff in h3. rewrite h2 in h3 at 1. red in h3; contradiction. Qed. Lemma inh_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (pf:Finite A), Inhabited A -> sort_set R pfso A pf <> nil. intros T R hs A h1 h2. destruct h2 as [a h2]. rewrite (in_sort_set_iff R hs A h1) in h2. eapply in_not_nil; apply h2. Qed. Lemma sort_set_sing : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (x:T) (pf:Finite (Singleton x)), sort_set R pfso (Singleton x) pf = x::nil. intros T R hs n h1. unfold sort_set. match goal with |- min_set_sortable _ ?pf1 ?pf2 _ _ = _ => generalize pf1 pf2 end. rewrite list_reps_sing. intros h2 h3. rewrite min_set_sortable_sing at 1; auto. Qed. Lemma nin_hd_tl_sort_set' : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (pff:Finite A) (pfia:Inhabited A), let l := sort_set R pfso A pff in let pfic := inh_sort_set R pfso A pff pfia in ~In (hd' l pfic) (tl l). intros T R hs A h1 h2. simpl. apply nin_hd_tl_no_dup'. apply no_dup_sort_set. Qed. Lemma hd_sort_set' : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (pff:Finite A) (pfi:Inhabited A), let pfso' := sort_dec_sortable _ pfso in hd' (sort_set R pfso A pff) (inh_sort_set R pfso A pff pfi) = min_set_sortable A pff pfi R pfso'. intros T R hs A h1 h2; simpl. unfold sort_set, min_set_sortable. pose proof h2 as h2'. destruct h2' as [a h2']. rewrite <- (hd_hd_compat' _ _ a). unfold sort_set, min_set_sortable. destruct constructive_definite_description as [l h3]. simpl. destruct h3 as [h3 h4]. destruct constructive_definite_description as [a' h5]; simpl. destruct h5 as [h5 h6]. pose proof (h6 (hd a l)) as h6'. pose proof h3 as h3'. destruct h3' as [h3']. destruct h3' as [? h3']; subst. destruct (sort_dec_tso_dec _ hs (hd a l) a') as [h7 | h7]; auto. destruct h7 as [h7 | h7]. hfirst h6'. constructor. pose proof (inh_not_nil_in_list_reps _ _ h2 h3) as h8. rewrite (hd_hd_compat' _ h8). rewrite <- list_to_set_in_iff. apply in_hd'. rewrite in_sing_iff. intro; subst. apply (sort_dec_irrefl _ hs) in h7; auto. specialize (h6' hf). pose proof (sort_dec_trans _ hs _ _ _ h6' h7) as hi. apply (sort_dec_irrefl _ hs) in hi. contradiction. assumption. specialize (h4 (a'::remove (sort_dec_eq_dec _ hs) a' l)). hfirst h4. constructor. constructor. split. simpl. rewrite <- subtract_remove_compat. rewrite add_sub_compat_in. destruct h3 as [h3]. destruct h3; auto. assumption. constructor. rewrite in_remove_iff. intro h8. destruct h8 as [h8 h9]. contradict h9; auto. apply no_dup_remove. assumption. rewrite in_sing_iff. intro h8. rewrite (hd_compat' l (inh_not_nil_in_list_reps _ _ h2 h3)) in h8. inversion h8. rewrite <- (hd_hd_compat' l _ a) in h8. inversion h8. lr_if H3. fold hlr in H3. destruct hlr as [hl | hr]. subst. rewrite <- hd_hd_compat' in h7 at 1. apply (sort_dec_irrefl _ hs) in h7; auto. contradict hr; auto. specialize (h4 hf). rewrite (hd_compat' l (inh_not_nil_in_list_reps _ _ h2 h3)) in h4 at 1. red in h4. destruct h4 as [h4 h9]. simpl in h4. destruct ( sort_dec_tso_dec R hs (hd' l (inh_not_nil_in_list_reps (list_to_set l) l h2 h3)) a') as [h10 | h10]. destruct h10 as [h10 | h10]. pose proof (sort_dec_trans _ hs _ _ _ h10 h7) as h11. rewrite <- hd_hd_compat' in h11 at 1. apply (sort_dec_irrefl _ hs) in h11. contradiction. subst. apply hd_hd_compat'. contradiction. Qed. Lemma in_min : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (pff:Finite A) (pfi:Inhabited A) (x:T), Inn A x -> let pfso' := sort_dec_sortable _ pfso in x = min_set_sortable A pff pfi R pfso' \/ In x (tl (sort_set R pfso A pff)). intros T R hso A h1 h2 x h3. pose (inh_sort_set R hso A h1 h2) as h4. rewrite (in_sort_set_iff R hso A h1) in h3. simpl. rewrite (hd_compat' _ h4) in h3. unfold h4 in h3. rewrite hd_sort_set' in h3. destruct h3; [left | right]; auto. Qed. Lemma tl_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (pff:Finite A) (pfi:Inhabited A), let pfso' := sort_dec_sortable _ pfso in let pft := sort_dec_eq_dec _ pfso in tl (sort_set R pfso A pff) = remove pft (min_set_sortable A pff pfi R pfso') (sort_set R pfso A pff). intros T R h0 A hf hi. simpl. pose proof (finite_set_list_no_dup _ hf) as h2. destruct h2 as [l h2]. destruct h2 as [h2 h3]; subst. rewrite <- hd_sort_set'. rewrite remove_hd_no_dup'; auto. apply no_dup_sort_set. Qed. Lemma incrR_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (C:Ensemble T) (pf:Finite C), increasingR R (sort_set R pfso C pf). intros T R hso C h1. unfold sort_set, min_set_sortable. destruct constructive_definite_description as [l h2]. simpl. destruct h2 as [h2 h3]. revert h3 h2 h1. revert C. induction l as [|a l ih]; simpl. intros ? ? h1. intros; auto. intros C h1 h2 h3. destruct h2 as [h2]. destruct h2 as [h2 hn]; subst. pose proof (no_dup_cons _ _ hn) as hn'. pose proof (no_dup_cons_nin _ _ hn) as h7. specialize (ih (Subtract (Add (list_to_set l) a) a)). hfirst ih. intros l' h4. destruct h4 as [h4 h5]. destruct h4 as [h4]. destruct h4 as [h4 h6]; subst. simpl in h4. rewrite sub_add_same_nin in h4. specialize (h1 (a::l')). rewrite in_sing_iff in h5. hfirst h1. constructor. constructor. split. simpl. f_equal; auto. constructor. intro h8. rewrite list_to_set_in_iff in h8. rewrite <- h4 in h8. rewrite <- list_to_set_in_iff in h8. contradiction. assumption. intro h8. rewrite in_sing_iff in h8. inversion h8; subst. contradict h5; auto. specialize (h1 hf). red in h1. destruct h1 as [h1 h9]. simpl in h1. destruct (sort_dec_tso_dec R hso a a) as [h10 | h10]. destruct h10 as [h10 | h10]. apply (sort_dec_irrefl _ hso) in h10. contradiction. red. split; auto. contradiction. contradict h7. rewrite <- list_to_set_in_iff in h7. assumption. specialize (ih hf). hfirst ih. constructor. split. rewrite sub_add_same_nin; auto. contradict h7. rewrite list_to_set_in_iff; auto. assumption. specialize (ih hf0 (subtract_preserves_finite _ _ h3)). destruct l as [|a' l]. simpl; auto. simpl. simpl in ih. split; auto. assert (h8:a <> a'). intro; subst. contradict h7; left; auto. specialize (h1 (a'::a::l)). hfirst h1. constructor. constructor. split. simpl. apply add_comm. constructor. intro h9. destruct h9 as [h9 | h9]. contradiction. pose proof (no_dup_cons_nin _ _ hn'). contradiction. constructor. contradict h7. right; auto. pose proof (no_dup_cons _ _ hn'); auto. rewrite in_sing_iff. intro h9. inversion h9; subst. contradict h8; auto. specialize (h1 hf1). red in h1. simpl in h1. destruct h1 as [h1 h2]. destruct (sort_dec_tso_dec R hso a a') as [h9 | h9]. destruct h9 as [h9 | h9]; auto. subst. contradict h8; auto. contradiction. Qed. Lemma incrR_impl_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), increasingR R l -> l = sort_set R pfso (list_to_set l) (list_to_set_finite _). intros T R hso l h1. pose proof (increasingR_no_dup (sort_dec_eq_dec _ hso) R l (sort_dec_so _ hso )h1) as hn. unfold sort_set, min_set_sortable. destruct constructive_definite_description as [l' h2]; simpl. destruct h2 as [h2 h3]. destruct h2 as [h2]. destruct h2 as [h2 h4]. destruct (list_eq_dec (sort_dec_eq_dec _ hso) l l') as [h5 | h5]; auto. specialize (h3 l). hfirst h3. constructor. constructor. split; auto. rewrite in_sing_iff; auto. specialize (h3 hf). eapply incrR_no_dup_minimal; auto. red in h3. destruct h3; auto. apply (sort_dec_eq_dec _ hso). apply h1. apply lt_list_impl_leR. apply h3. Qed. Lemma is_sorted_set_cons : forall {T:Type} (l:list T) (a:T), is_sorted_set (a::l) -> is_sorted_set l. intros T l m h1. red in h1. red. destruct h1 as [R [hso h2]]. pose proof (incrR_sort_set R hso (list_to_set (m::l)) (list_to_set_finite _ )) as h3. rewrite <- h2 in h3. pose proof (increasingR_cons R l _ h3) as h4. exists R, hso. apply incrR_impl_sort_set; auto. Qed. Lemma sort_set_cons : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (pf:Finite A) (ln:list T) (a:T), sort_set R pfso A pf = a::ln -> ln = sort_set R pfso (Subtract A a) (subtract_preserves_finite _ _ pf). intros T R hso A h1 ln a h2. pose proof (f_equal list_to_set h2) as h3. rewrite sort_set_compat in h3; subst. simpl. pose proof (incrR_sort_set R hso (list_to_set (a::ln)) h1) as h3. rewrite h2 in h3. pose proof (increasingR_cons_nin (sort_dec_eq_dec _ hso) R _ _ (sort_dec_so _ hso) h3) as h4. symmetry. gen0. rewrite sub_add_same_nin. intro h5. symmetry. erewrite sort_set_functional. apply incrR_impl_sort_set. eapply increasingR_cons. apply h3. reflexivity. rewrite <- list_to_set_in_iff; auto. Grab Existential Variables. assumption. Qed. Lemma lst_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (pf:Finite A) (pfi:Inhabited A), let pfso' := sort_dec_sortable _ pfso in let pfn := inh_sort_set R pfso _ _ pfi in lst (sort_set R pfso A pf) pfn = max_set_sortable A pf pfi R pfso'. intros T R h1 A h2 h3. pose proof (incrR_sort_set R h1 A h2) as h4. simpl. match goal with |- _ = max_set_sortable ?x1 ?x2 ?x3 ?x4 ?x5 => pose proof (max_set_sortable_compat x1 x2 x3 x4 x5) as h6 end. simpl in h6. destruct h6 as [h6 h7]. rewrite (in_sort_set_iff R h1 A h2) in h6. rewrite lst_compat'. pose proof (increasingR_last _ _ (sort_dec_trans _ h1) h4 (max_set_sortable A h2 h3 R (sort_dec_sortable R h1)) h6) as h5. let P := type of h5 in match P with Rw R ?x ?y => destruct (sort_dec_eq_dec _ h1 x y) as [h8|h8] end. rewrite h8. apply last_indep. apply inh_sort_set; auto. let P := type of h5 in match P with Rw R _ ?x => specialize (h7 x) end. hfirst h7. constructor. rewrite in_sort_set_iff. apply in_last. apply inh_sort_set; auto. rewrite in_sing_iff; auto. specialize (h7 hf). pose proof (trans_R_Rw R (sort_dec_trans _ h1) _ _ _ h7 h5) as h9. apply (sort_dec_irrefl _ h1) in h9. contradiction. Qed. Lemma max_set_sortable_add_lst_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (x:T) (pf:Finite A) (pfi:Inhabited A), ~Inn A x -> let pf' := Add_preserves_Finite _ A x pf in let pfi' := inh_add A x in let pfx := inh_sort_set R pfso _ _ pfi in let pfso' := sort_dec_sortable _ pfso in x = max_set_sortable (Add A x) pf' pfi' R pfso' -> R (lst (sort_set R pfso A pf) pfx) x. intros T R hso A x h0 h1 h2 h3 h4 h5 h6 h20. rewrite h20. pose proof (max_set_sortable_compat (Add A x) h3 h4 R h6) as h7. simpl in h7. destruct h7 as [h7 h8]. apply h8. constructor. left. rewrite in_sort_set_iff. apply in_lst. rewrite in_sing_iff. intro h9. unfold h6 in h9, h20. rewrite <- h20 in h9 at 1. contradict h2. rewrite in_sort_set_iff. rewrite h9. apply in_lst. Qed. Lemma sort_set_app_add : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A A':Ensemble T) (pfa:Finite A) (pfa':Finite A') (pfi:Inhabited A') (a:T), ~Inn A a -> A' = Add A a -> a = max_set_sortable A' pfa' pfi R (sort_dec_sortable _ pfso) -> sort_set R pfso A' pfa' = sort_set R pfso A pfa ++ a::nil. intros T R hso A A' h1 h2 h3 a h0 h4 h5. pose proof (finite_inh_or_empty _ h1) as hf. destruct hf as [hf |hf]. pose proof (finite_set_list_no_dup _ h1) as h6. destruct h6 as [l h6]. destruct h6 as [h6 h7]. rewrite h6 in h4. gen0. rewrite h4. intro h8. pose proof (list_to_set_app l (a::nil)) as h9. simpl in h9. rewrite add_empty_sing in h9. gen0. unfold Add. rewrite <- h9. intro h10. erewrite sort_set_functional. symmetry. eapply incrR_impl_sort_set. erewrite <- increasingR_app_iff. split. apply incrR_sort_set. split. red; auto. apply max_set_sortable_add_lst_sort_set. assumption. rewrite h5. apply max_set_sortable_functional; auto. rewrite <- h5. rewrite h6. assumption. apply sort_dec_trans; auto. rewrite h9. rewrite list_to_set_app. rewrite sort_set_compat. simpl. f_equal; auto. rewrite <- add_empty_sing; auto. rewrite hf in h4. rewrite add_empty_sing in h4. clear h5. subst. rewrite sort_set_sing. erewrite sort_set_functional. rewrite sort_set_empty at 1. rewrite app_nil_l. reflexivity. reflexivity. Grab Existential Variables. assumption. assumption. apply cons_neq_nil. assumption. Qed. Lemma tl_sort_set_add : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (c:T) (pf:Finite (Add A c)), list_to_set (tl (sort_set R pfso (Add A c) pf)) = Subtract (Add A c) (min_set_sortable _ pf (inh_add _ _) R (sort_dec_sortable _ pfso)). intros T R hso A c h1. pose proof (no_dup_sort_set R hso (Add A c) h1) as hn. apply Extensionality_Ensembles; red; split; red; intros m h3. rewrite <- list_to_set_in_iff in h3. pose proof (min_set_sortable_compat (Add A c) h1 (inh_add _ _) R (sort_dec_sortable _ hso)) as h4. simpl in h4. destruct h4 as [h4 h5]. pose proof h4 as h4'. rewrite (in_sort_set_iff R hso _ h1 _) in h4'. pose proof (nin_hd_tl_sort_set' R hso _ h1 (inh_add _ _)) as h6. simpl in h6. rewrite hd_sort_set' in h6. pose proof (in_tl _ _ h3) as h8. rewrite list_to_set_in_iff in h8. rewrite sort_set_compat in h8. constructor; auto. rewrite in_sing_iff. intro; subst. contradiction. destruct h3 as [h3 h4]. rewrite in_sing_iff in h4. rewrite <- list_to_set_in_iff. pose proof (hd_compat' (sort_set R hso (Add A c) h1) (inh_sort_set R hso _ _ (inh_add _ _))) as h5. simpl in h5. rewrite hd_sort_set' in h5. rewrite (in_sort_set_iff R hso _ h1) in h3. rewrite h5 in h3. destruct h3; try contradiction; auto. Qed. Lemma in_tl_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (pf:Finite A) (x y:T), Inn A x -> Inn A y -> R x y -> In y (tl (sort_set R pfso A pf)). intros T R hso A h1. pose proof (finite_set_list_no_dup _ h1) as h2. destruct h2 as [l [h2 h3]]; subst. revert h1 h3. induction l as [|c l ih]; simpl. intros ? ? ? ? h1. contradict h1. intros h1 h2 x y h3 h4 h5. pose proof (no_dup_cons_nin _ _ h2) as h10. hfirst ih. apply list_to_set_finite. pose proof (no_dup_cons _ _ h2) as h6. specialize (ih hf h6). rewrite list_to_set_in_iff. rewrite tl_sort_set_add. destruct h4 as [y h4 | y h4], h3 as [x h3 | x h3]; subst. assert (hx:c <> x). intro; subst. contradict h10. rewrite list_to_set_in_iff; auto. assert (hy:c <> y). intro; subst. contradict h10. rewrite list_to_set_in_iff; auto. specialize (ih x y h3 h4 h5). constructor. left; auto. rewrite in_sing_iff. intro; subst. pose proof (min_set_sortable_compat (Add (list_to_set l) c) h1 (inh_add (list_to_set l) c) R (sort_dec_sortable _ hso)) as h7. simpl in h7. destruct h7 as [h7 h8]. specialize (h8 x). hfirst h8. constructor. left; auto. rewrite in_sing_iff. intro; subst. apply (sort_dec_irrefl _ hso) in h5. contradiction. specialize (h8 hf0). pose proof (sort_dec_trans _ hso _ _ _ h8 h5) as hi. apply (sort_dec_irrefl _ hso) in hi; contradiction. rewrite in_sing_iff in h3; subst. constructor. left; auto. rewrite in_sing_iff. intro; subst. pose proof (min_set_sortable_compat (Add (list_to_set l) x) h1 (inh_add (list_to_set l) x) R (sort_dec_sortable _ hso)) as h7. simpl in h7. destruct h7 as [h7 h8]. specialize (h8 x). hfirst h8. constructor. right. constructor. rewrite in_sing_iff. intro he. rewrite he in h5. apply (sort_dec_sortable _ hso) in h5; contradiction. specialize (h8 hf0). pose proof (sort_dec_trans _ hso _ _ _ h8 h5) as hi. apply (sort_dec_irrefl _ hso) in hi. contradiction. destruct h4. constructor. right. constructor. rewrite in_sing_iff. intro h7. pose proof (min_set_sortable_compat (Add (list_to_set l) c) h1 (inh_add (list_to_set l) c) R (sort_dec_sortable _ hso)) as h11. simpl in h11. destruct h11 as [h11 h12]. specialize (h12 x). hfirst h12. constructor. left; auto. rewrite in_sing_iff. intro h13. rewrite <- h13 in h3. rewrite h7 in h3 at 1. contradict h10. rewrite list_to_set_in_iff; auto. specialize (h12 hf0). rewrite h7 in h12. pose proof (sort_dec_trans _ hso _ _ _ h12 h5) as hi. apply (sort_dec_irrefl _ hso) in hi. contradiction. destruct h3, h4. apply (sort_dec_irrefl _ hso) in h5. contradiction. Qed. Lemma list_minus_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A B:Ensemble T) (pfa:Finite A) (pfb:Finite B) (pfab:Finite (Setminus A B)), list_minus (sort_dec_eq_dec _ pfso) (sort_set R pfso A pfa) (sort_set R pfso B pfb) = sort_set R pfso (Setminus A B) pfab. intros T R hso A B h1 h2 h3. pose proof (list_to_set_list_minus (sort_dec_eq_dec _ hso) (sort_set R hso A h1) (sort_set R hso B h2)) as h4. do 2 rewrite sort_set_compat in h4. symmetry. gen0. rewrite <- h4. intro h5. symmetry. pose proof (proof_irrelevance _ h5 (list_to_set_finite _)); subst. apply incrR_impl_sort_set. apply incrR_list_minus. apply sort_dec_trans; auto. apply incrR_sort_set. Qed. Definition fam_to_ens_sort_set {T:Type} (R:Relation T) (pfso:sort_dec R) (F:Family T) (pf:finite_mem F) : Ensemble (list T) := Im (full_sig F) (fun S => sort_set R pfso _ (pf _ (proj2_sig S))). Lemma fam_to_ens_sort_set_empty_eq : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (pf:finite_mem (Empty_set (Ensemble T))), fam_to_ens_sort_set R pfso _ pf = Empty_set _. intros. unfold fam_to_ens_sort_set. rewrite full_empty_sig_empty. apply image_empty. Qed. Lemma fam_to_ens_sort_set_sing_eq : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A:Ensemble T) (pf:Finite A) (pffm:finite_mem (Singleton A)), fam_to_ens_sort_set R pfso (Singleton A) pffm = Singleton (sort_set R pfso A pf). intros. unfold fam_to_ens_sort_set. rewrite full_sig_sing. rewrite im_singleton. simpl. f_equal. f_equal. apply proof_irrelevance. Qed. Lemma finite_fam_to_ens_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (F:Family T) (pf:finite_mem F), Finite F -> Finite (fam_to_ens_sort_set R pfso F pf). intros T R hso F h1 h2. unfold fam_to_ens_sort_set. apply finite_image. rewrite <- finite_full_sig_iff; auto. Qed. Lemma inh_fam_to_ens_sort_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (F:Family T) (pf:finite_mem F), Inhabited F -> Inhabited (fam_to_ens_sort_set R pfso F pf). intros T R hso F h1 h2. destruct h2 as [x h2]. apply Inhabited_intro with (sort_set R hso _ (h1 _ h2)). apply Im_intro with (exist _ _ h2). constructor. simpl. reflexivity. Qed. Lemma fam_to_ens_sort_set_functional : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (F F':Family T) (pf:finite_mem F) (pf':finite_mem F'), F = F' -> fam_to_ens_sort_set R pfso F pf = fam_to_ens_sort_set R pfso F' pf'. intros T R hso F F' h1 h2 h3. subst. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. Lemma fam_to_ens_sort_set_subtract : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (F:Family T) (pff:Finite F) (pffm:finite_mem F) (l:list T), Inn F (list_to_set l) -> increasingR R l -> Subtract (fam_to_ens_sort_set R pfso F pffm) l = fam_to_ens_sort_set R pfso (Subtract F (list_to_set l)) (finite_mem_subtract _ _ pffm). intros T R hso F h1 h0 l h2 hi. unfold fam_to_ens_sort_set. match goal with |- Subtract (Im _ ?f) _ = _ => pose proof (im_full_sig_sub_decompose_inj (@ens_eq_dec T) F (list_to_set l) f h2) as h3 end. simpl in h3. rewrite h3. pose proof (incrR_impl_sort_set R hso l hi) as h4. rewrite h4 at 7. pose proof (proof_irrelevance _ (h0 (list_to_set l) h2) (list_to_set_finite _)) as h6. rewrite h6. rewrite sub_add_eq at 1. rewrite subtract_nin. apply im_ext_in. intros x h7. unfold restriction_sig. simpl. apply sort_set_functional; auto. intro h7. clear h4. inversion h7 as [x h8 q h9]. subst. pose proof (incrR_impl_sort_set R hso _ hi) as h4. rewrite <- h4 in h9. unfold restriction_sig in h9. simpl in h9. destruct x as [x h10]. simpl in h9. destruct h10 as [h10 h11]. pose proof h11 as h11'. rewrite in_sing_iff in h11'. pose proof (f_equal (@list_to_set _) h9) as h12. rewrite sort_set_compat in h12. contradiction. Qed. Lemma im_fam_to_ens_sort_set_list_to_set : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (F:Family T) (pf:finite_mem F), Im (fam_to_ens_sort_set R pfso F pf) list_to_set = F. intros T R hso F h1. unfold fam_to_ens_sort_set. rewrite im_im. rewrite im_full_sig_proj1_sig. apply im_ext_in. intros x h2. clear h2. destruct x as [S h2]. simpl. apply sort_set_compat. Qed. Lemma in_fam_to_ens_sort_set_impl : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (F:Family T) (pf:finite_mem F) (l:list T), Inn (fam_to_ens_sort_set R pfso F pf) l -> Inn F (list_to_set l). intros T R hso F h1 l h2. destruct h2 as [l h2]. subst. clear h2. rewrite sort_set_compat. apply proj2_sig. Qed. Lemma in_fam_impl : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (F:Family T) (pffm:finite_mem F) (A:Ensemble T) (pfin:Inn F A), Inn (fam_to_ens_sort_set R pfso F pffm) (sort_set R pfso A (pffm _ pfin)). intros T R hso F h1 A h2. unfold fam_to_ens_sort_set. apply Im_intro with (exist _ _ h2). constructor. simpl. reflexivity. Qed. Lemma in_fam_to_ens_sort_set_incr : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (F:Family T) (pffm:finite_mem F) (l:list T), Inn (fam_to_ens_sort_set R pfso F pffm) l -> increasingR R l. intros T R hso F h1 l h2. destruct h2 as [l h2]; subst. clear h2. destruct l as [l h2]. simpl. apply incrR_sort_set. Qed. Lemma setminus_sort_set_fam : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (A B:Family T) (pfa:finite_mem A) (pfb:finite_mem B), let f := fun S : sig_set A => sort_set R pfso (proj1_sig S) (pfa (proj1_sig S) (proj2_sig S)) in let g := fun S : sig_set B => sort_set R pfso (proj1_sig S) (pfb (proj1_sig S) (proj2_sig S)) in let f' := restriction_sig f _ (setminus_inc _ _) in Setminus (Im (full_sig A) f) (Im (full_sig B) g) = Im (full_sig (Setminus A B)) f'. intros T R hso A B h1 h2 f g f'. apply Extensionality_Ensembles; red; split; red; intros C h3. destruct h3 as [h3 h4]. destruct h3 as [C h3]; subst. clear h3. destruct C as [C h5]. unfold f'. assert (h6:~Inn B C). contradict h4. unfold f, g. apply Im_intro with (exist _ _ h4). constructor. simpl. f_equal. apply proof_irrelevance. apply Im_intro with (exist (fun x => Inn (Setminus A B) x) C (conj h5 h6)). constructor. unfold restriction_sig. f_equal. apply proj1_sig_injective; auto. destruct h3 as [C h3]; subst. clear h3. destruct C as [C h4]. destruct h4 as [h4 h5]. constructor. econstructor. constructor. unfold f', restriction_sig. reflexivity. intro h6. inversion h6 as [x h8 q h9 h10]; subst. unfold f', g, restriction_sig, f in h9. simpl in h9. apply inj_sort_set in h9. subst. destruct x as [x h9]. simpl in h5. contradiction. Qed. Definition min_listR {T:Type} (R:Relation T) (pft:tso_dec R) (l m:list T) : list T := if (le_listR_dec R pft l m) then l else m. Definition max_listR {T:Type} (R:Relation T) (pft:tso_dec R) (l m:list T) : list T := if (le_listR_dec R pft l m) then m else l. Lemma min_listR_compat1 : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l m:list T), le_listR R (sort_dec_tso_dec _ pfso) (min_listR R (sort_dec_tso_dec _ pfso) l m) l. intros T R hso l m. unfold min_listR. destruct (le_listR_dec R (sort_dec_tso_dec _ hso) l m) as [h1 | h1]. apply refl_le_listR; auto. apply (sort_dec_irrefl _ hso). rewrite <- lt_listR_iff in h1. destruct h1; auto. Qed. Lemma min_listR_compat2 : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l m:list T), le_listR R (sort_dec_tso_dec _ pfso) (min_listR R (sort_dec_tso_dec _ pfso) l m) m. intros T R hso l m. unfold min_listR. lr_if_goal; fold hlr; destruct hlr as [h1 | h1]; auto. rewrite <- lt_listR_iff in h1. destruct h1; auto. apply refl_le_listR. apply sort_dec_irrefl; auto. Qed. Lemma min_listR_eq_compat1 : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (la lb lc:list T), let pft := sort_dec_tso_dec _ pfso in le_listR R pft lb lc -> le_listR R pft (min_listR R pft la lb) (min_listR R pft la lc). intros T R hso la lb lc ht h1. unfold min_listR. lr_if_goal; fold hlr; destruct hlr as [h2 | h2]. lr_if_goal; fold hlr; destruct hlr as [h3 | h3]; auto. apply refl_le_listR. apply sort_dec_irrefl. assumption. pose proof (trans_le_listR R hso _ _ _ h2 h1). contradiction. lr_if_goal; fold hlr; destruct hlr as [h3 | h3]; auto. unfold ht in h2. rewrite <- lt_listR_iff in h2. destruct h2; auto. Qed. Lemma min_listR_eq_compat2 : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (la lb lc:list T), let pft := sort_dec_tso_dec _ pfso in le_listR R pft la lb -> le_listR R pft (min_listR R pft la lc) (min_listR R pft lb lc). intros T R hso la lb lc ht h1. unfold min_listR. lr_if_goal; fold hlr; destruct hlr as [h2 | h2]; auto. lr_if_goal; fold hlr; destruct hlr as [h3 | h3]; auto. lr_if_goal; fold hlr; destruct hlr as [h3 | h3]; auto. pose proof (trans_le_listR R hso _ _ _ h1 h3). contradiction. apply refl_le_listR. apply sort_dec_irrefl; auto. Qed. Lemma min_listR_nil1 : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), let pft := sort_dec_tso_dec _ pfso in min_listR R pft nil l = nil. intros. unfold min_listR. lr_if_goal; fold hlr; destruct hlr as [h2 | h2]; auto. contradict h2. apply le_listR_nil. Qed. Lemma min_listR_nil2 : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), let pft := sort_dec_tso_dec _ pfso in min_listR R pft l nil = nil. intros. unfold min_listR. lr_if_goal; fold hlr; destruct hlr as [h2 | h2]; auto. destruct l; auto. contradict h2. Qed. Lemma min_listR_eq_cons : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l m:list T) (b:T), let pft := sort_dec_tso_dec _ pfso in min_listR R pft (b :: l) (b :: m) = b:: min_listR R pft l m. intros T R hso l m b. unfold min_listR. lr_if_goal; fold hlr; destruct hlr as [h1 | h1]; auto. lr_if_goal; fold hlr; destruct hlr as [h2 | h2]; auto. do 2 red in h1. destruct hso as [hts htr hir]. destruct (hts b b) as [h3 | h3]. destruct h3 as [h3 | h3]. apply hir in h3. contradiction. contradiction. contradiction. lr_if_goal; fold hlr; destruct hlr as [h2 | h2]; auto. unfold le_listR in h1. simpl in h1. destruct (sort_dec_tso_dec R hso b b) as [h3 | h3]. destruct h3 as [h3 | h3]. apply (sort_dec_irrefl _ hso) in h3. contradiction. contradiction. apply (sort_dec_irrefl _ hso) in h3. contradiction. Qed. Lemma min_listR_comm : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l m:list T), let pft := sort_dec_tso_dec _ pfso in min_listR R pft l m = min_listR R pft m l. intros T R hso l m ht. unfold min_listR. lr_if_goal; fold hlr; destruct hlr as [h1 | h1]; auto. lr_if_goal; fold hlr; destruct hlr as [h2 | h2]; auto. eapply antisym_le_listR. apply h1. assumption. lr_if_goal; fold hlr; destruct hlr as [h2 | h2]; auto. eapply antisym_not_le_listR; auto. apply h2. assumption. Qed. Lemma min_listR_assoc : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (la lb lc:list T), let pft := sort_dec_tso_dec _ pfso in min_listR R pft la (min_listR R pft lb lc) = min_listR R pft (min_listR R pft la lb) lc. intros T R hso la lb lc. unfold min_listR. destruct (le_listR_dec R) as [h1 | h1]; auto. destruct (le_listR_dec R) as [h2 | h2]; auto. destruct (le_listR_dec R) as [h3 | h3]; auto. destruct (le_listR_dec R) as [h4 | h4]; auto. contradiction. destruct (le_listR_dec R) as [h4 | h4]; auto. pose proof (trans_le_listR R hso _ _ _ h4 h2). contradiction. contradiction. destruct (le_listR_dec R) as [h3 | h3]; auto. destruct (le_listR_dec R) as [h4 | h4]; auto. contradiction. destruct (le_listR_dec R) as [h4 | h4]; auto. contradiction. rewrite <- lt_listR_iff in h4, h3. pose proof (trans_lt_listR R hso _ _ _ h3 h4) as h5. destruct h5 as [h5 h6]. pose proof (antisym_le_listR R hso _ _ h1 h5). subst. contradict h6; auto. destruct (le_listR_dec R) as [h2 | h2]; auto. rewrite <- lt_listR_iff in h1. destruct (le_listR_dec R) as [h3 | h3]; auto. destruct (le_listR_dec R) as [h4 | h4]; auto. pose proof (trans_le_lt_listR R hso _ _ _ h4 h1) as h5. destruct h5 as [h5 h6]. contradict h6; auto. destruct (le_listR_dec R) as [h4 | h4]; auto. rewrite <- lt_listR_iff in h3. pose proof (trans_lt_le_listR R hso _ _ _ h3 h4) as h5. destruct h5 as [h5 h6]. pose proof (antisym_le_listR R hso _ _ h5 h2). subst. reflexivity. contradiction. destruct (le_listR_dec R) as [h3 | h3]; auto. destruct (le_listR_dec R) as [h4 | h4]; auto. contradiction. contradiction. Qed. Lemma min_listR_idem : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l:list T), let pft := sort_dec_tso_dec _ pfso in min_listR R pft l l = l. intros ? R. intros. unfold min_listR. destruct (le_listR_dec R); auto. Qed. Lemma min_listR_mono : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l la lb:list T), let pft := sort_dec_tso_dec _ pfso in le_listR R pft la lb -> le_listR R pft (min_listR R (sort_dec_tso_dec _ pfso) l la) (min_listR R (sort_dec_tso_dec _ pfso) l lb). intros T R hso l la lb ht h1. unfold min_listR. destruct (le_listR_dec R) as [h2 | h2]; auto. destruct (le_listR_dec R) as [h3 | h3]; auto. apply refl_le_listR. apply sort_dec_irrefl; auto. pose proof (trans_le_listR R hso _ _ _ h2 h1). contradiction. destruct (le_listR_dec R) as [h3 | h3]; auto. rewrite <- lt_listR_iff in h2. destruct h2; auto. Qed. Lemma Rw_lt_listR_iff : forall {T:Type} (R:Relation T) (pfs:sort_dec R) (l l':list T), let pft := sort_dec_tso_dec _ pfs in Rw (lt_listR R pft) l l' <-> le_listR R pft l l'. intros T R h1 l l'; split; intro h2. red in h2. destruct h2 as [h2 |]; subst. apply lt_list_impl_leR; auto. apply refl_le_listR; auto. apply sort_dec_irrefl; auto. apply le_listR_lt_eq_dec in h2. destruct h2 as [h2 | h2]. left; auto. right; auto. Qed. Lemma lt_listR_map : forall {T U:Type} (R:Relation T) (R':Relation U) (pftso:tso_dec R) (pfir:irrefl R) (pfso':sort_dec R') (f:T->U), incrR_fun R R' f -> forall (l l':list T), let pftso' := sort_dec_tso_dec _ pfso' in lt_listR R pftso l l' -> lt_listR R' pftso' (map f l) (map f l'). intros T U R R' htso hir hso' f h1 l l' htso'. pose proof (incrR_fun_inj R R' f htso (sort_dec_irrefl _ hso') h1) as h1'. revert h1' h1. revert l f. induction l' as [|a' l' ih]; simpl. intros ? ? h1' h1 h2. apply not_lt_listR_nil in h2; contradiction. intros l f h1' h1 h2. destruct l as [|a l]. simpl. red; simpl; split; auto. intro; discriminate. simpl; simpl in h2. red in h2; red. destruct h2 as [h2 h3]. simpl in h2. simpl. destruct (htso a a') as [h4 | h4]. destruct h4 as [h4 | h4]. destruct (htso' (f a) (f a')) as [h5 | h5]. destruct h5 as [h5 | h5]; auto. split; auto. intro h6. inversion h6 as [h6']. apply map_inj in H ; subst; contradict h3; auto. apply h1' in h6'; subst; auto. apply map_inj in H ; subst; auto. rewrite h6' in h5. apply (sort_dec_irrefl _ hso') in h5; contradiction. apply h1' in h5; subst; auto. contradict h4; auto. apply h1 in h4; subst; contradict h3; auto. pose proof (sort_dec_trans _ hso' _ _ _ h4 h5) as ht. apply (sort_dec_irrefl _ hso') in ht. contradiction. split. destruct (htso' (f a) (f a')) as [h6 | h6]. destruct h6; auto. apply h1' in e; subst. apply lt_list_impl_leR. apply ih; auto. split; auto. intro; subst; contradict h3; auto. subst. apply (sort_dec_irrefl _ hso') in h6; auto. intro h5. inversion h5. apply h1' in H0; subst. apply map_inj in H1; subst; contradict h3; auto. apply map_inj in H1; subst; auto. contradiction. Qed. Definition incrR_list_fun {T U:Type} (R:Relation T) (R':Relation U) (pftso:tso_dec R) (pftso':tso_dec R') (f:list_fun T U) : Prop := incrR_fun (lt_listR R pftso) (lt_listR R' pftso') f. Definition weak_incrR_list_fun {T U:Type} (R:Relation T) (R':Relation U) (pftso:tso_dec R) (pftso':tso_dec R') (f:list_fun T U) : Prop := weak_incrR_fun (lt_listR R pftso) (lt_listR R' pftso') f. Lemma incrR_list_fun_impl_weak : forall {T U:Type} (R:Relation T) (R':Relation U) (pftso:tso_dec R) (pftso':tso_dec R') (f:list_fun T U), incrR_list_fun R R' pftso pftso' f -> weak_incrR_list_fun R R' pftso pftso' f. unfold incrR_list_fun, weak_incrR_list_fun; intros; eapply incrR_fun_impl_weak; auto. Qed. End ListSort. Section ListNat. Definition increasing (l:list nat) : Prop := increasingR lt l. Definition weak_increasing (l:list nat) : Prop := weak_increasingR lt l. Lemma increasing_cons : forall (l:list nat) (a:nat), increasing (a::l) -> increasing l. unfold increasing; apply increasingR_cons. Qed. Lemma weak_increasing_cons : forall (l:list nat) (a:nat), weak_increasing (a::l) -> weak_increasing l. unfold weak_increasing; apply weak_increasingR_cons. Qed. Lemma increasing_impl_weak : forall (l:list nat), increasing l -> weak_increasing l. unfold increasing, weak_increasing; apply increasingR_impl_weak. Qed. Lemma lt_increasing_cons : forall (l:list nat) (a m:nat), In a l -> increasing (m::l)-> m < a. unfold increasing; intros l a m h1 h2; apply (lt_increasingR_cons lt l a m lt_trans h1 h2). Qed. Lemma le_weak_increasing_cons : forall (l:list nat) (a m:nat), In a l -> weak_increasing (m::l)-> m <= a. unfold weak_increasing; intros l a m h1 h2; rewrite <- Rw_le_iff; apply (le_weak_increasingR_cons lt l a m lt_trans h1 h2). Qed. Lemma increasing_cons_impl : forall (l:list nat) (a:nat), increasing (a::l) -> increasing l. unfold increasing; apply increasingR_cons_impl. Qed. Lemma weak_increasing_cons_impl : forall (l:list nat) (a:nat), weak_increasing (a::l) -> weak_increasing l. unfold weak_increasing; apply weak_increasingR_cons_impl. Qed. Lemma increasing_cons_impl_cons : forall (l:list nat) (a d:nat), increasing (a::l) -> d < a -> increasing (d::a::l). unfold increasing; apply increasingR_cons_impl_cons. Qed. Lemma weak_increasing_cons_impl_cons : forall (l:list nat) (a d:nat), weak_increasing (a::l) -> d <= a -> weak_increasing (d::a::l). unfold weak_increasing. intros l a d h1 h2. rewrite <- Rw_le_iff in h2. apply (weak_increasingR_cons_impl_cons lt l a d h1 h2). Qed. Lemma increasing_impl_cons : forall (l:list nat) (pf:l <> nil) (d:nat), increasing l -> d < hd' l pf -> increasing (d::l). unfold increasing; apply increasingR_impl_cons. Qed. Lemma weak_increasing_impl_cons : forall (l:list nat) (pf:l <> nil) (d:nat), weak_increasing l -> d <= hd' l pf -> weak_increasing (d::l). unfold weak_increasing; intros ? ? ? ? h1; rewrite <- Rw_le_iff in h1; eapply weak_increasingR_impl_cons; auto. apply h1. Qed. Lemma increasing_impl_cons' : forall (l:list nat) (d:nat), increasing l -> (forall c:nat, In c l -> d < c) -> increasing (d::l). unfold increasing; apply increasingR_impl_cons'. Qed. Lemma weak_increasing_impl_cons' : forall (l:list nat) (d:nat), weak_increasing l -> (forall c:nat, In c l -> d <= c) -> weak_increasing (d::l). unfold weak_increasing. intros l d h1 h2. apply (weak_increasingR_impl_cons' lt l d h1). intros; rewrite Rw_le_iff; auto. Qed. Lemma increasing_tl : forall (l:list nat), increasing l -> increasing (tl l). unfold increasing; apply increasingR_tl. Qed. Lemma weak_increasing_tl : forall (l:list nat), weak_increasing l -> weak_increasing (tl l). unfold weak_increasing; apply weak_increasingR_tl. Qed. Lemma increasing_remove : forall (l:list nat) (a:nat), increasing l -> increasing (remove nat_eq_dec a l). unfold increasing. intros; apply (increasingR_remove nat_eq_dec lt); auto. red; intro; eapply lt_trans. Qed. Lemma weak_increasing_remove : forall (l:list nat) (a:nat), weak_increasing l -> weak_increasing (remove nat_eq_dec a l). unfold weak_increasing. intros l a h1. apply (weak_increasingR_remove nat_eq_dec lt); auto. red; intro; eapply lt_trans. Qed. Lemma incr_list_minus : forall (l l':list nat), increasing l -> let lm := list_minus nat_eq_dec l l' in increasing lm. unfold increasing; intros l l'. apply (incrR_list_minus nat_eq_dec lt l l' lt_trans); auto. Qed. Lemma weak_incr_list_minus : forall (l l':list nat), weak_increasing l -> let lm := list_minus nat_eq_dec l l' in weak_increasing lm. unfold weak_increasing; intros l l'. apply (weak_incrR_list_minus nat_eq_dec lt l l' lt_trans); auto. Qed. Lemma increasing_cons_lt_subst : forall (l:list nat) (a b:nat), increasing (a::l) -> b < a -> increasing (b::l). unfold increasing; intros l a b h1 h2; apply (increasingR_cons_lt_subst lt l a b lt_trans h1 h2). Qed. Lemma weak_increasing_cons_le_subst : forall (l:list nat) (a b:nat), weak_increasing (a::l) -> b <= a -> weak_increasing (b::l). unfold weak_increasing; intros l a b h1 h2; apply (weak_increasingR_cons_le_subst lt l a b lt_trans h1). rewrite Rw_le_iff; auto. Qed. Lemma increasing_cons_no_dup : forall (l:list nat) (n:nat), increasing (n::l) -> NoDup l. unfold increasing; intros l n; apply (increasingR_cons_no_dup nat_eq_dec lt l n so_lt_nat). Qed. Lemma increasing_cons_nin : forall (l:list nat) (a:nat), increasing (a::l) -> ~In a l. unfold increasing; intros l a; apply (increasingR_cons_nin nat_eq_dec lt l a so_lt_nat). Qed. Lemma increasing_no_dup : forall (l:list nat) , increasing l -> NoDup l. unfold increasing; intro l; apply (increasingR_no_dup nat_eq_dec lt l so_lt_nat). Qed. Lemma increasing_firstn : forall (l:list nat) (m:nat), increasing l -> increasing (firstn m l). unfold increasing; intros l m; apply (increasingR_firstn nat_eq_dec lt l m). red; intro; apply lt_trans. Qed. Lemma weak_increasing_firstn : forall (l:list nat) (m:nat), weak_increasing l -> weak_increasing (firstn m l). unfold weak_increasing; intros l m h1; apply (weak_increasingR_firstn lt l m lt_trans h1). Qed. Lemma increasing_skipn : forall (l:list nat) (m:nat), increasing l -> increasing (skipn m l). unfold increasing; intros l m h1; apply (increasingR_skipn lt l m h1). Qed. Lemma weak_increasing_skipn : forall (l:list nat) (m:nat), weak_increasing l -> weak_increasing (skipn m l). unfold weak_increasing; intros l m h1; apply (weak_increasingR_skipn lt l m h1). Qed. Lemma increasing_app : forall (l l':list nat) (pfl:l <> nil) (pfl':l' <> nil), increasing l -> increasing l' -> lst l pfl < hd' l' pfl' -> increasing (l++l'). unfold increasing; intros; rewrite <- increasingR_app_iff; repeat split; auto. apply H1. apply lt_trans'. Qed. Lemma weak_increasing_app : forall (l l':list nat) (pfl:l <> nil) (pfl':l' <> nil), weak_increasing l -> weak_increasing l' -> lst l pfl <= hd' l' pfl' -> weak_increasing (l++l'). unfold weak_increasing; intros ? ? ? ? ? ? h1; rewrite <- Rw_le_iff in h1; erewrite <- weak_increasingR_app_iff; repeat split; auto. apply h1. apply lt_trans'. Qed. Lemma weak_incr_no_dup_impl : forall (l:list nat), weak_increasing l -> NoDup l -> increasing l. unfold weak_increasing, increasing; intros; apply weak_incrR_no_dup_impl; auto. Qed. Lemma increasing_map : forall (f:nat->nat) (l:list nat), incr_fun f -> increasing l -> increasing (map f l). intros ? ? h1; rewrite incr_fun_iff in h1; unfold incr_fun' in h1; unfold increasing; apply increasingR_map; auto. Qed. Lemma weak_increasing_map : forall (f:nat->nat) (l:list nat), weak_incr_fun f -> weak_increasing l -> weak_increasing (map f l). intros ? ? h1; rewrite weak_incr_fun_iff in h1; unfold incr_fun' in h1; unfold weak_increasing; apply weak_increasingR_map; auto. Qed. Lemma incr_seg_list : forall (n:nat), increasing (seg_list n). intro n. induction n as [|n ih]. unfold seg_list, seq; red; simpl; auto. rewrite seg_list_app; auto with arith. destruct (zerop n) as [hz | hz]; subst. simpl; auto. simpl. pose proof (in_lst _ (seg_list_not_nil _ hz)) as h1. rewrite in_seg_list_iff in h1. rewrite <- minus_n_O. eapply increasing_app; auto. do 2 red; auto. simpl. apply h1. Grab Existential Variables. apply cons_neq_nil. Qed. Lemma incr_seq : forall (m n:nat), increasing (seq m n). intros m n. rewrite seq_eq_seq_0. apply increasing_map. red. intros; omega. apply incr_seg_list. Qed. Lemma incr_all_sing : forall(l:list nat) (m:nat), all_sing l m -> increasing l -> l = nil \/ l = m::nil. unfold increasing; apply incrR_all_sing. apply nat_eq_dec. apply so_lt_nat. Qed. Lemma incr_lS : forall (l:list nat), increasing l -> increasing (lS l). unfold lS; intros; apply increasing_map; auto. apply incr_fun_S. Qed. Lemma weak_incr_lS : forall (l:list nat), weak_increasing l -> weak_increasing (lS l). unfold lS; intros; apply weak_increasing_map; auto. apply weak_incr_fun_S. Qed. Lemma incr_lpred_nin0 : forall (l:list nat) , ~ In 0 l -> increasing l -> increasing (lpred l). intro l. induction l as [|a l ih]; simpl. intros; auto. intros h1 h2. pose proof (increasing_cons_nin l a h2) as h3. assert (h4:a <> 0). intro; subst. contradict h1. left; auto. assert (h5:~In 0 l). intro h5. contradict h1. right; auto. specialize (ih h5). hfirst ih. eapply increasing_cons_impl. apply h2. specialize (ih hf). destruct l as [|c l]; simpl; auto. pose proof (lt_increasing_cons _ _ _ (in_eq _ _) h2) as h76. split. apply mono_pred_lt; auto. omega. assumption. Qed. Lemma incr_lpred_remove_in0_in1 : forall (l:list nat), In 0 l -> In 1 l -> increasing l -> increasing (lpred (remove nat_eq_dec 0 l)). intro l. destruct l as [|a l]. intros; auto. intros h1 h2 h3. pose proof (increasing_cons_nin l a h3) as h4. assert (h5:a = 0). destruct h1 as [|h1]; auto. pose proof (lt_increasing_cons _ _ _ h1 h3) as h5. omega. subst. destruct h2 as [|h2]. discriminate. rewrite remove_cons_eq. rewrite nin_remove_eq; auto. red in h3. pose proof (increasing_cons_impl _ _ h3) as h5. apply incr_lpred_nin0; auto. Qed. Lemma incr_lpred_in0_nin1 : forall (l:list nat) , In 0 l -> ~ In 1 l -> increasing l -> increasing (lpred l). intro l. induction l as [|a l ih]; simpl. intros; contradiction. intros h1 h2 h3. pose proof (increasing_cons_nin _ _ h3) as h6. destruct a as [|a]. simpl. destruct l as [|c l]; simpl; auto. pose proof (lt_increasing_cons _ _ _ (in_eq _ _) h3) as h7. assert (h8: 1 < c). destruct c as [|c]. omega. assert (1 <> S c). contradict h2. right. left; auto. omega. split. omega. pose proof (increasing_cons_impl (c::l) _ h3) as h9. pose proof (incr_lpred_nin0 _ h6 h9) as h10. assumption. simpl. destruct h1 as [h1 | h1]; try discriminate. assert (h7:a <> 0). intro. subst. contradict h2; left; auto. specialize (ih h1). hfirst ih. contradict h2. right; auto. specialize (ih hf). hfirst ih. eapply increasing_cons_impl. apply h3. specialize (ih hf0). eapply increasing_impl_cons; auto. destruct l as [|c l]. simpl. contradict h1. simpl. simpl in h3. destruct h3; omega. Grab Existential Variables. apply neq_nil_lpred. intro; subst. contradict h1. Qed. Lemma weak_incr_lpred : forall (l:list nat), weak_increasing l -> weak_increasing (lpred l). intro l. induction l as [|c l ih]; simpl; auto. destruct l as [|a l]; simpl; auto. intro h1. red in h1. simpl in h1. destruct h1 as [h1 h2]. rewrite Rw_le_iff in h1. red. split. rewrite Rw_le_iff. omega. apply ih; auto. Qed. Lemma increasing_linds_occ : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T), increasing (linds_occ pfdt l x). intros T h1 l. induction l as [|a l ih]; simpl. intros. rewrite linds_occ_nil; constructor. intro x. destruct (h1 x a) as [h3 | h3]; subst. rewrite linds_occ_cons_eq. simpl. destruct (nil_dec' (linds_occ h1 l a)) as [h3 | h3]. rewrite h3; simpl; auto. do 2 red; auto. apply not_nil_cons in h3. destruct h3 as [n [ln h3]]. pose proof (in_eq n ln) as h4. rewrite <- h3 in h4. rewrite in_linds_occ_iff in h4. destruct h4 as [m [h4 h5]]. destruct h5 as [h5 h6]; subst. rewrite h3. simpl. split. apply lt_0_Sn. specialize (ih a). rewrite h3 in ih. apply (incr_lS (lind_occ h1 l m a h4 :: ln)); auto. specialize (ih x). rewrite linds_occ_cons_neq; auto. apply incr_lS; auto. Qed. Lemma increasing_linds_occ' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pf:In x l), increasing (linds_occ' pfdt l x pf). intros; rewrite <- linds_occ_compat'; apply increasing_linds_occ; auto. Qed. Lemma in_linds_incr : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat), In ln (linds pfdt l) -> increasing ln. intros T h1 l ln h2. rewrite in_linds_iff in h2. destruct h2 as [x [[h2 h3] h4]]; subst. apply increasing_linds_occ. Qed. Lemma list_to_set_injable_increasingR : forall {T:Type} (R:Relation T), sort_dec R -> list_to_set_injable (increasingR R). intros T R hso. pose proof (sort_dec_eq_dec _ hso) as hdt. pose proof (sort_dec_trans _ hso) as htr. pose proof (sort_dec_irrefl _ hso) as hir. red; intro l. induction l as [|a l ih]; simpl. intros ? ? ? h1. symmetry in h1; apply empty_set_nil in h1; subst; auto. intros l' h1 h2 h3. destruct l' as [|a' l']. simpl in h3. apply add_not_empty in h3. contradiction. simpl in h3. simpl in h2. destruct l as [|b l]; simpl. simpl in h3. rewrite add_empty_sing in h3. apply sing_eq_add in h3. destruct h3 as [? h3]. subst. destruct h3 as [h3 | h3]. apply empty_set_nil in h3; subst; auto. f_equal. destruct l' as [|a']; auto. simpl in h3. symmetry in h3. apply sing_eq_add in h3. destruct h3 as [h3]; subst. destruct h2 as [h2]. apply hir in h2; contradiction. destruct l' as [|b' l']. simpl in h3. rewrite add_empty_sing in h3. symmetry in h3. apply sing_eq_add in h3. destruct h3 as [h3 h4]. subst. destruct h4 as [h4 | h4]. apply add_not_empty in h4; contradiction. symmetry in h4. apply sing_eq_add in h4. destruct h4 as [? h4]; subst. destruct h1 as [h1]. apply hir in h1; contradiction. simpl in h3. destruct h1 as [h1 h4], h2 as [h2 h5]. destruct (hdt a a') as [|h6]; subst. f_equal. apply ih; auto. pose proof (add_nin_inj_eq (Add (list_to_set l) b) (Add (list_to_set l') b') a') as h6. hfirst h6. intro h8. destruct h8 as [a' h8|a' h8]; subst. rewrite <- list_to_set_in_iff in h8. pose proof (lt_increasingR_cons R l _ _ htr h8 h4) as h9. pose proof (htr _ _ _ h1 h9) as h10. apply hir in h10; contradiction. destruct h8; subst. apply hir in h1; contradiction. specialize (h6 hf). hfirst h6. intro h8. destruct h8 as [a' h8|a' h8]; subst. rewrite <- list_to_set_in_iff in h8. pose proof (lt_increasingR_cons R l' _ _ htr h8 h5) as h9. pose proof (htr _ _ _ h2 h9) as h10. apply hir in h10; contradiction. destruct h8; subst. apply hir in h2; contradiction. specialize (h6 hf0). apply h6; auto. pose proof (Add_intro2 _ (Add (list_to_set l') b') a') as h7. rewrite <- h3 in h7. pose proof (Add_intro2 _ (Add (list_to_set l) b) a) as h7'. rewrite h3 in h7'. destruct h7 as [a' h7 | a' h7], h7' as [a h7' | a h7']. destruct h7 as [a' h7 | a' h7], h7' as [a h7' | a h7']. rewrite <- list_to_set_in_iff in h7, h7'. pose proof (lt_increasingR_cons R l _ _ htr h7 h4) as h9. pose proof (lt_increasingR_cons R l' _ _ htr h7' h5) as h9'. pose proof (htr _ _ _ h1 h9) as h10. pose proof (htr _ _ _ h2 h9') as h11. pose proof (htr _ _ _ h10 h11) as h12. apply hir in h12; contradiction. destruct h7'. rewrite <- list_to_set_in_iff in h7. pose proof (lt_increasingR_cons R l _ _ htr h7 h4) as h8. pose proof (htr _ _ _ h8 h2) as h9. pose proof (htr _ _ _ h9 h1) as h10. apply hir in h10; contradiction. destruct h7. rewrite <- list_to_set_in_iff in h7'. pose proof (lt_increasingR_cons R l' _ _ htr h7' h5) as h8. pose proof (htr _ _ _ h2 h8) as h9. pose proof (htr _ _ _ h9 h1) as h10. apply hir in h10; contradiction. destruct h7, h7'. pose proof (htr _ _ _ h1 h2) as h7. apply hir in h7; contradiction. destruct h7'. destruct h7 as [a' h7 |a' h7]; subst. rewrite <- list_to_set_in_iff in h7. pose proof (lt_increasingR_cons R l _ _ htr h7 h4) as h8. pose proof (htr _ _ _ h8 h1) as h9. apply hir in h9; contradiction. destruct h7; contradict h6; auto. destruct h7; contradict h6; auto. destruct h7; contradict h6; auto. Qed. Lemma list_to_set_injable_increasing : list_to_set_injable increasing. red; intros l l' h1 h2 h3. eapply list_to_set_injable_increasingR; auto. apply sort_dec_lt_nat. assumption. assumption. Qed. Lemma incrR_remove_inj : forall {T:Type} (pfdt:type_eq_dec T) (R:Relation T) (pfso:sort_dec R) (l l':list T) (x:T), In x l -> In x l' -> increasingR R l -> increasingR R l' -> remove pfdt x l = remove pfdt x l' -> l = l'. intros T hdt R hso l l' x hi hi' h1 h2 h3. eapply list_to_set_injable_increasingR. apply hso. assumption. assumption. apply Extensionality_Ensembles; red; split; red; intros c h4; rewrite <- list_to_set_in_iff in h4; rewrite <- list_to_set_in_iff; destruct (hdt x c) as [|h5]; subst; auto. pose proof (in_remove_iff hdt l c x) as h6. hr h6; auto. rewrite <- h6 in hr. rewrite h3 in hr. rewrite in_remove_iff in hr. destruct hr; auto. pose proof (in_remove_iff hdt l' c x) as h6. hr h6; auto. rewrite <- h6 in hr. rewrite <- h3 in hr. rewrite in_remove_iff in hr. destruct hr; auto. Qed. Lemma in_linds_list_to_set_injable : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), list_to_set_injable (fun xl => In xl (linds pfdt l)). intros T h1 l; red; intros lm ln h2 h3 h4. rewrite in_linds_iff in h2, h3. destruct h2 as [x [[h2 h5] h6]], h3 as [y [[h3 h8] h7]]; subst. pose proof (increasing_linds_occ h1 l x) as h8. pose proof (increasing_linds_occ h1 l y) as h9. eapply incrR_no_dup_minimal. apply nat_eq_dec. apply h8. apply no_dup_linds_occ. assumption. pose proof list_to_set_injable_increasing as h10. red in h10. specialize (h10 _ _ h8 h9 h4). rewrite h10. apply refl_le_listR. red; intro; apply lt_irrefl. Grab Existential Variables. apply sort_dec_lt_nat. Qed. Lemma lec_list_to_set_injable : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l :list T) (l':list U) (i j:nat) (pfi:i < length l) (pfj:j < length l'), list_to_set (linds pfdt l) = list_to_set (linds pfdu l') -> list_to_set (lec pfdt pfi) = list_to_set (lec pfdu pfj) -> lec pfdt pfi = lec pfdu pfj. intros T U hdt hdu l l' i j hi hj he he'. pose proof (in_lec hdt _ _ hi) as h1. pose proof (in_lec hdu _ _ hj) as h2. rewrite list_to_set_in_iff in h1. rewrite he in h1. rewrite <- list_to_set_in_iff in h1. eapply in_linds_list_to_set_injable. apply h1. assumption. assumption. Qed. Lemma list_to_set_linds_eq_length : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), list_to_set (linds pfdt l) = list_to_set (linds pfdu l') -> length l = length l'. intros T U h1 h2 l l' h3. pose proof (partition_linds h1 l) as h4. pose proof (partition_linds h2 l') as h5. eapply no_dup_partition_faml_nat_seg_list_to_set_eq. eapply no_dup_linds. eapply no_dup_linds. eapply no_dup_mem_linds. eapply no_dup_mem_linds. eapply in_linds_list_to_set_injable. eapply in_linds_list_to_set_injable. apply h3. apply partition_linds. apply partition_linds. Qed. Lemma linds_eq_length : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), linds pfdt l = linds pfdu l' -> length l = length l'. intros; eapply list_to_set_linds_eq_length. rewrite H at 1. reflexivity. Qed. Definition lastlinds {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) := last (linds_occ pfdt l x) 0. Definition lastlindss {T:Type} (pfdt:type_eq_dec T) (l:list T) : list nat := map_in_dep (fun ln (pf:In ln (linds pfdt l)) => last ln 0). Lemma in_lastlindss_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat), In n (lastlindss pfdt l) <-> exists ln:list nat, In ln (linds pfdt l) /\ n = last ln 0. intros T h1 l n; split; intro h2. unfold lastlindss in h2. rewrite in_map_in_dep_iff in h2. destruct h2 as [ln [h2 h3]]; subst. exists ln; tauto. unfold lastlindss. rewrite in_map_in_dep_iff. destruct h2 as [ln [h2 h3]]; subst. exists ln, h2; auto. Qed. Lemma in_lastlindss_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat), In n (lastlindss pfdt l) -> n < length l. intros T h1 l n h2. rewrite in_lastlindss_iff in h2. destruct h2 as [ln [h2 h3]]; subst. rewrite in_linds_iff in h2. destruct h2 as [x [h2 h3]]. destruct h2 as [h2]; subst. eapply in_linds_occ_lt. apply in_last. apply linds_occ_in; auto. Qed. Lemma length_lastlindss : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), length (lastlindss pfdt l) = length (list_singularize pfdt l). intros T h1 l. unfold lastlindss. rewrite length_map_in_dep. apply length_linds. Qed. Definition maxl (l:list nat) : l <> nil -> nat := maxlR lt (tso_dec_lt_nat) l. Lemma maxl_compat : forall (l:list nat) (n:nat) (pf:In n l), let pfn := in_not_nil l n pf in n <= maxl l pfn. unfold maxl; intros; rewrite <- Rw_le_iff. apply maxlR_compat. red; apply lt_trans. Qed. Definition minl (l:list nat) : l <> nil -> nat := minlR lt (tso_dec_lt_nat) l. Lemma in_minl : forall (l:list nat) (pf:l <> nil), In (minl l pf) l. unfold minl; apply in_minlR. Qed. Lemma minl_compat : forall (l:list nat) (n:nat) (pf:In n l), let pfn := in_not_nil l n pf in In n l -> minl l pfn <= n. unfold minl; intros; rewrite <- Rw_le_iff. apply minlR_compat. red; apply lt_trans. Qed. Definition two_lists_to_rel {T U:Type} (l:list T) (l':list U) (pfe:length l = length l') : Ensemble (T * U) := list_to_set (map_seg (fun m (pfm:m < length l) => (nth_lt l m pfm, nth_lt l' m (lt_eq_trans _ _ _ pfm pfe)))). Lemma two_lists_to_rel_nil : forall (T U:Type), two_lists_to_rel (@nil T) (@nil U) (eq_refl _) = Empty_set _. intros; unfold two_lists_to_rel. apply Extensionality_Ensembles; red; split; red; intros x h1. rewrite <- list_to_set_in_iff in h1. rewrite in_map_seg_iff in h1. simpl in h1. destruct h1 as [? [? ?]]; omega. contradiction. Qed. Lemma in_two_lists_to_rel : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l') (x:T) (y:U), Inn (two_lists_to_rel l l' pfe) (x, y) -> exists (pfx:In x l) (pfy:In y l') (ox oy:nat) (pfox:ox < count_occ pfdt l x) (pfofy:oy < count_occ pfdu l' y), lind_occ pfdt l ox x pfx = lind_occ pfdu l' oy y pfy. intros T U h1 h2 l l' h3 x y h4. unfold two_lists_to_rel in h4. rewrite <- list_to_set_in_iff in h4. rewrite in_map_seg_iff in h4. destruct h4 as [m [h4 h5]]. inversion h5. subst. exists (in_nth_lt _ _ _), (in_nth_lt _ _ _). pose proof (ex_lind_occ_nth_lt h1 _ _ h4) as h6. pose proof (ex_lind_occ_nth_lt h2 _ _ (lt_eq_trans _ _ _ h4 h3)) as h7. destruct h6 as [ox [h6 h8]], h7 as [oy [h7 h9]]. exists ox, oy. exists h6, h7. congruence. Qed. Lemma incl_graph_seg_two_lists_to_rel_in : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l') (x:T) (pfx:In x l) (f:sig_set (list_to_set l) -> U), Included (graph_seg f) (two_lists_to_rel l l' pfe) -> In (f (exist _ _ (iff1 (list_to_set_in_iff l x) pfx))) l'. intros T U hdt hdu l l' h1 x h2 f h3. red in h3. specialize (h3 (x, (f (exist (Inn (list_to_set l)) x (iff1 (list_to_set_in_iff l x) h2))))). hfirst h3. unfold graph_seg. apply Im_intro with (exist (Inn (list_to_set l)) x (iff1 (list_to_set_in_iff l x) h2)). constructor. simpl; auto. specialize (h3 hf). apply (in_two_lists_to_rel hdt hdu) in h3. destruct h3 as [? [? ?]]; auto. Qed. Lemma graph_seg_two_lists_to_rel_compat : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l') (x:T) (pfx:In x l) (f:sig_set (list_to_set l) -> U) (pfinc:Included (graph_seg f) (two_lists_to_rel l l' pfe)), let y := (f (exist _ _ (iff1 (list_to_set_in_iff l x) pfx))) in exists (ox oy:nat) (pfox:ox < count_occ pfdt l x) (pfofy:oy < count_occ pfdu l' y), let pfy := incl_graph_seg_two_lists_to_rel_in pfdt pfdu l l' pfe x pfx f pfinc in lind_occ pfdt l ox x pfx = lind_occ pfdu l' oy y pfy. Proof. intros T U hdt hdu l l' h1 x h2 f h3 y. pose proof (in_two_lists_to_rel hdt hdu l l' h1 x y) as h4. hfirst h4. apply h3. unfold graph_seg. apply Im_intro with (exist (Inn (list_to_set l)) x (iff1 (list_to_set_in_iff l x) h2)). constructor. simpl; auto. specialize (h4 hf). destruct h4 as [h4 [h5 [ox [oy [h6 [h7 h8]]]]]]. exists ox, oy, h6, h7; simpl; auto. pose proof (proof_irrelevance _ h2 h4) as h9. rewrite h9 at 1. unfold y in h8. unfold y. rewrite h8. repeat f_equal. apply proof_irrelevance. Qed. Lemma incl_graph_seg_lists_to_rel_nth_lt_compat : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l') (i:nat) (pfi:i < length l) (f:sig_set (list_to_set l) -> U) (pfinc:Included (graph_seg f) (two_lists_to_rel l l' pfe)), exists (i':nat) (pfi':i' < length l), nth_lt l i pfi = nth_lt l i' pfi' /\ let pfn := in_nth_lt l i' pfi' in let pfn' := iff1 (list_to_set_in_iff l _) pfn in f (exist _ _ pfn') = nth_lt l' i' (lt_eq_trans _ _ _ pfi' pfe). intros T U h1 h2 l l' h3 i h4 f h5. red in h5. specialize (h5 ((nth_lt l i h4), (f (exist (Inn (list_to_set l)) (nth_lt l i h4) (iff1 (list_to_set_in_iff l _) (in_nth_lt _ _ _)))))). hfirst h5. unfold graph_seg. apply Im_intro with (exist (Inn (list_to_set l)) (nth_lt l i h4) (iff1 (list_to_set_in_iff l (nth_lt l i h4)) (in_nth_lt l i h4))). constructor. simpl; auto. specialize (h5 hf). apply (in_two_lists_to_rel h1 h2) in h5. destruct h5 as [h5 [h8 [ox [oy [h9 [h10 h11]]]]]]. exists (lind_occ h1 l ox (nth_lt l i h4) h5), (lt_lind_occ _ _ _ _ _). split. rewrite lind_occ_compat; auto. simpl. pose proof (nth_lt_functional2 _ _ _ h11 (lt_eq_trans _ _ _ (lt_lind_occ _ _ _ _ _) h3)) (lt_lind_occ _ _ _ _ _) as h12. rewrite lind_occ_compat in h12. rewrite h12. f_equal. apply proj1_sig_injective; simpl. rewrite lind_occ_compat; auto. Qed. Definition two_lists_to_rel_im {T U:Type} (pfdt:type_eq_dec T) (l:list T) (l':list U) (pfe:length l = length l') (x:T) : Ensemble U := list_to_set (map_in_dep (fun (n:nat) (pfi:In n (linds_occ pfdt l x)) => let pflt := in_linds_occ_lt _ _ _ _ pfi in nth_lt l' n (lt_eq_trans _ _ _ pflt pfe))). Definition two_lists_to_rel_inv_im {T U:Type} (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l') (y:U) : Ensemble T := list_to_set (map_in_dep (fun (n:nat) (pfi:In n (linds_occ pfdu l' y)) => let pflt := in_linds_occ_lt _ _ _ _ pfi in nth_lt l n (lt_eq_trans _ _ _ pflt (eq_sym pfe)))). Lemma inh_two_lists_to_rel_im : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l') (x:T), In x l -> Inhabited (two_lists_to_rel_im pfdt l l' pfe x). intros T U hdt hdu l l' h1 x h2. econstructor. unfold two_lists_to_rel_im. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. pose proof (ex_lind_occ_nth_lt hdu _ _ (lt_eq_trans _ _ _ (lt_lind hdt _ _ h2) h1)) as h3. destruct h3 as [occ' [h3 h4]]. pose proof (ex_lind_occ_nth_lt hdt _ _ (lt_lind hdt _ _ h2)) as h5. destruct h5 as [occ [h5 h6]]. terml h4. terml h6. exists x1. ex_goal. rewrite in_linds_occ_iff. exists occ, h2. rewrite lind_compat in h5. split; auto. unfold x1. apply lind_occ_functional; auto. rewrite lind_compat; auto. exists hex. unfold x1. rewrite <- h6 in h4. gen0. rewrite <- h4. intro h7. erewrite nth_lt_functional3. rewrite lind_occ_compat. reflexivity. Qed. Lemma inh_two_lists_to_rel_inv_im : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l') (y:U), In y l' -> Inhabited (two_lists_to_rel_inv_im pfdu l l' pfe y). intros T U hdt hdu l l' h1 x h2. econstructor. unfold two_lists_to_rel_inv_im. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. pose proof (ex_lind_occ_nth_lt hdt _ _ (lt_eq_trans _ _ _ (lt_lind hdu _ _ h2) (eq_sym h1))) as h3. destruct h3 as [occ' [h3 h4]]. pose proof (ex_lind_occ_nth_lt hdu _ _ (lt_lind hdu _ _ h2)) as h5. destruct h5 as [occ [h5 h6]]. terml h4. terml h6. exists x1. ex_goal. rewrite in_linds_occ_iff. exists occ, h2. rewrite lind_compat in h5. split; auto. unfold x1. apply lind_occ_functional; auto. rewrite lind_compat; auto. exists hex. unfold x1. rewrite <- h6 in h4. gen0. rewrite <- h4. intro h7. erewrite nth_lt_functional3. rewrite lind_occ_compat. reflexivity. Qed. Lemma in_two_lists_to_rel_im : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l') (x:T) (y:U), Inn (two_lists_to_rel_im pfdt l l' pfe x) y -> exists (pfx:In x l) (pfy:In y l') (occ:nat) (occ':nat), occ < count_occ pfdt l x /\ occ' < count_occ pfdu l' y /\ lind_occ pfdt l occ x pfx = lind_occ pfdu l' occ' y pfy. intros T U h1 h2 l l' h3 x y h5. unfold two_lists_to_rel_im in h5. rewrite <- list_to_set_in_iff in h5. rewrite in_map_in_dep_iff in h5. destruct h5 as [x' [h6 h7]]. subst. pose proof h6 as h6'. rewrite in_linds_occ_iff in h6'. destruct h6' as [occ [h7 [h8 h9]]]. subst. exists h7, (in_nth_lt _ _ _), occ. pose proof (ex_lind_occ_nth_lt h2 _ _ (lt_eq_trans (lind_occ h1 l occ x h7) (length l) (length l') (in_linds_occ_lt h1 l x (lind_occ h1 l occ x h7) h6) h3)) as h9. destruct h9 as [occ' [h9 h10]]. exists occ'. repeat split; auto. Qed. Lemma in_two_lists_to_rel_inv_im : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l') (x:T) (y:U), Inn (two_lists_to_rel_inv_im pfdu l l' pfe y) x -> exists (pfy:In y l') (pfx:In x l) (occ:nat) (occ':nat), occ < count_occ pfdt l x /\ occ' < count_occ pfdu l' y /\ lind_occ pfdt l occ x pfx = lind_occ pfdu l' occ' y pfy. intros T U h1 h2 l l' h3 x y h5. unfold two_lists_to_rel_inv_im in h5. rewrite <- list_to_set_in_iff in h5. rewrite in_map_in_dep_iff in h5. destruct h5 as [x' [h6 h7]]. subst. pose proof h6 as h6'. rewrite in_linds_occ_iff in h6'. destruct h6' as [occ [h7 [h8 h9]]]. subst. pose proof (ex_lind_occ_nth_lt h1 _ _ (lt_eq_trans (lind_occ h2 l' occ y h7) (length l') (length l) (in_linds_occ_lt h2 l' y (lind_occ h2 l' occ y h7) h6) (eq_sym h3))) as h9. destruct h9 as [occ' [h9 h10]]. exists h7, (in_nth_lt _ _ _), occ', occ. repeat split; auto. Qed. Lemma in_two_lists_to_rel1 : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l') (x:T), In x l -> exists y:U, Inn (two_lists_to_rel_im pfdt l l' pfe x) y /\ Inn (two_lists_to_rel l l' pfe) (x, y). intros T U hdt hdu l l' h2 x h3. apply (inh_two_lists_to_rel_im hdt hdu l l' h2) in h3. destruct h3 as [y h4]. exists y. split; auto. apply (in_two_lists_to_rel_im hdt hdu l l' h2 x y) in h4. destruct h4 as [h4 [h5 [occ [occ' [h6 [h7 h8]]]]]]. unfold two_lists_to_rel. rewrite <- list_to_set_in_iff. rewrite in_map_seg_iff. exists (lind_occ hdt l occ x h4), (lt_lind_occ _ _ _ _ _). rewrite lind_occ_compat. f_equal. gen0. rewrite h8. intro h9. erewrite nth_lt_functional3. rewrite lind_occ_compat; auto. Qed. Lemma in_two_lists_to_rel2 : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l') (y:U), In y l' -> exists x:T, Inn (two_lists_to_rel_inv_im pfdu l l' pfe y) x /\ Inn (two_lists_to_rel l l' pfe) (x, y). intros T U hdt hdu l l' h2 y h3. apply (inh_two_lists_to_rel_inv_im hdt hdu l l' h2) in h3. destruct h3 as [x h4]. exists x. split; auto. apply (in_two_lists_to_rel_inv_im hdt hdu l l' h2 x y) in h4. destruct h4 as [h4 [h5 [occ [occ' [h6 [h7 h8]]]]]]. unfold two_lists_to_rel. rewrite <- list_to_set_in_iff. rewrite in_map_seg_iff. exists (lind_occ hdu l' occ' y h4), (lt_eq_trans _ _ _ (lt_lind_occ _ _ _ _ _) (eq_sym h2)). f_equal. gen0. rewrite <- h8. intro h9. erewrite nth_lt_functional3. rewrite lind_occ_compat; auto. erewrite nth_lt_functional3. rewrite lind_occ_compat; auto. Qed. Definition two_lists_to_fun_in_dep {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (pfl:linds pfdt la = linds pfdu lb) : fun_in_dep la U := fun x (pfx:In x la) => let pfl' := linds_eq_length pfdt pfdu la lb pfl in nth_lt lb (lind pfdt _ _ pfx) (lt_eq_trans _ _ _ (lt_lind pfdt _ _ pfx) pfl'). Lemma in_two_lists_to_fun_in_dep : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (pfl:linds pfdt la = linds pfdu lb) (x:T) (pf:In x la), In (two_lists_to_fun_in_dep pfdt pfdu la lb pfl x pf) lb. unfold two_lists_to_fun_in_dep; intros; apply in_nth_lt. Qed. Definition two_lists_to_sig_fun {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (pfl:linds pfdt la = linds pfdu lb) : sig_set (list_to_set la) -> sig_set (list_to_set lb) := fun x => let pfx' := iff2 (list_to_set_in_iff la (proj1_sig x)) (proj2_sig x) in let pfl' := linds_eq_length pfdt pfdu la lb pfl in let pfi := in_two_lists_to_fun_in_dep pfdt pfdu la lb pfl in exist _ _ (iff1 (list_to_set_in_iff lb _) (in_nth_lt lb (lind pfdt _ _ pfx') (lt_eq_trans _ _ _ (lt_lind pfdt _ _ pfx') pfl'))). Definition fun_linds_eq_nat {U V:Type} (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (f:nat->U) (g:nat->V) : Prop := forall n:nat, linds pfdu (map_seg (fun_to_seg f n)) = linds pfdv (map_seg (fun_to_seg g n)). (*See below [linds_eq_lecs] for compatibility lemmas.*) Definition list_to_segT_fun {T:Type} (l:list T) : length l ---> T. induction l as [|a l h1]; simpl. intro n. destruct n as [n h1]. apply lt_n_0 in h1. contradiction. intro n. destruct n as [n h2]. destruct n. refine a. apply lt_S_n in h2. refine (h1 (exist (fun a => a < length l) _ h2)). Defined. Lemma list_to_segT_fun_compat : forall {T:Type} (l:list T), let f := list_to_segT_fun l in forall (n:(length l) @), In (f n) l /\ f n = nth (proj1_sig n) l (hd' l (lt_length l _ (proj2_sig n))). intro T. induction l as [|a l h1]. intros f n. simpl in n. destruct n; omega. intros f n. simpl. simpl in n. unfold f. simpl. destruct n as [n h2]. destruct n as [|n]. split. left. reflexivity. simpl. reflexivity. pose proof h2 as h2'. apply lt_S_n in h2'. simpl in h1. specialize (h1 (exist (fun a => a < length l) _ h2')). destruct h1 as [h1 h3]. split. right. pose proof (proof_irrelevance _ h2' (lt_S_n _ _ h2)); subst. simpl in h3. assumption. simpl. simpl in h3. pose proof (@nth_indep _ l n (hd' l (lt_length l n h2')) (hd' (a::l) (lt_length (a::l) (S n) h2)) h2') as h4. simpl in h4. rewrite <- h4. rewrite <- h3. f_equal. apply proj1_sig_injective; auto. Qed. Lemma list_to_segT_fun_no_dup_inj : forall {T:Type} (l:list T), NoDup l -> Injective (list_to_segT_fun l). intros T l h1. red. intros m n h2. destruct m as [m h3], n as [n h4]. apply proj1_sig_injective; simpl. pose proof (list_to_segT_fun_compat l (exist (fun a => a < length l) _ h3)) as h5. pose proof (list_to_segT_fun_compat l (exist (fun a => a < length l) _ h4)) as h6. simpl in h5, h6. destruct h5 as [h5 h5']. destruct h6 as [h6 h6']. rewrite h2 in h5'. rewrite h6' in h5'. apply nth_inj in h5'; auto. Qed. Definition list_to_segT_fun' {T:Type} (l:list T) : length l ---> {x:T | In x l}. intro x. pose proof (list_to_segT_fun_compat l x) as h1. destruct h1 as [h1 h2]. refine (exist (fun a=> In a l) _ h1). Defined. Lemma list_to_segT_fun_prime_compat : forall {T:Type} (l:list T) (x:(length l) @), let f := list_to_segT_fun l in let f' := list_to_segT_fun' l in f x = proj1_sig (f' x). intros T l x f f'. unfold f', f, list_to_segT_fun'. destruct (list_to_segT_fun_compat l x) as [h1 h2]. simpl. reflexivity. Qed. Lemma list_to_segT_fun_compat' : forall {T:Type} (l:list T), let f := list_to_segT_fun' l in forall (n:(length l) @), proj1_sig (f n) = nth (proj1_sig n) l (hd' l (lt_length l (proj1_sig n) (proj2_sig n))). intros T l f n. unfold f, list_to_segT_fun'. destruct (list_to_segT_fun_compat l n); auto. Qed. Lemma list_to_segT_fun_no_dup_inj' : forall {T:Type} (l:list T), NoDup l -> Injective (list_to_segT_fun' l). intros T l h1. pose proof (list_to_segT_fun_no_dup_inj l h1) as h2. red in h2. red. intros x y h3. destruct x as [x h5], y as [y h6]. pose proof (f_equal (@proj1_sig _ _) h3) as h4. do 2 rewrite <- list_to_segT_fun_prime_compat in h4. specialize (h2 (exist (fun a => a < length l) _ h5) (exist (fun a => a < length l) _ h6) h4); auto. Qed. Lemma list_to_segT_fun_surj' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), surjective (list_to_segT_fun' l). intros T hdt l. red. intro y. destruct y as [y h1]. pose (exist (fun c => c < length l) (lind hdt _ _ h1) (lt_lind hdt _ _ _)) as x. exists x. pose proof (list_to_segT_fun_compat' l x) as h2. apply proj1_sig_injective; simpl. rewrite h2. pose proof (nth_lt_lind_eq hdt _ _ h1) as h3. simpl in h3. rewrite <- h3. simpl. rewrite nth_lt_compat. apply nth_indep. apply lt_lind. Qed. Definition list_to_segT_fun'' {T:Type} (l:list T) : length l ---> {x:T | Inn (list_to_set l) x}. intro x. pose (list_to_segT_fun' l x) as x'. pose (proj2_sig x') as h1. simpl in h1. rewrite list_to_set_in_iff in h1. refine (exist _ _ h1). Defined. Lemma list_to_segT_fun_compat'' : forall {T:Type} (l:list T), let f := list_to_segT_fun'' l in forall (n:(length l) @), proj1_sig (f n) = nth (proj1_sig n) l (hd' l (lt_length l (proj1_sig n) (proj2_sig n))). intros T l f n. pose proof (list_to_segT_fun_compat' l n) as h2. simpl in h2. unfold f, list_to_segT_fun''. simpl. assumption. Qed. Lemma list_to_segT_fun_prime_compat' : forall {T:Type} (l:list T) (x:(length l) @), let f' := list_to_segT_fun' l in let f'' := list_to_segT_fun'' l in proj1_sig (f'' x) = proj1_sig (f' x). intros T l x f' f''. unfold f', f''. unfold list_to_segT_fun''. simpl. reflexivity. Qed. Lemma list_to_segT_fun_no_dup_inj'' : forall {T:Type} (l:list T), NoDup l -> Injective (list_to_segT_fun'' l). intros T l h1. pose proof (list_to_segT_fun_no_dup_inj' l h1) as h2. red in h2. red. intros x y h3. specialize (h2 x y). apply h2. apply proj1_sig_injective. do 2 rewrite <- list_to_segT_fun_prime_compat'. f_equal; auto. Qed. Lemma list_to_segT_fun_surj'' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), surjective (list_to_segT_fun'' l). intros T hdt l. red. intro y. pose proof (list_to_segT_fun_surj' hdt l) as h1. red in h1. destruct y as [y h2]. pose proof h2 as h2'. rewrite <- list_to_set_in_iff in h2'. specialize (h1 (exist (fun a => In a l) _ h2')). destruct h1 as [x h1]. exists x. apply proj1_sig_injective. simpl. rewrite h1. simpl. reflexivity. Qed. Definition list_to_seg_one_t_fun {T:Type} (l:list T) : seg_one_t (length l) -> T. induction l as [|a l h1]; simpl. intro x. destruct x as [x h1]. destruct h1 as [h1 h2]. assert (h3:1 <= 0). eapply le_trans. apply h1. apply h2. apply le_Sn_0 in h3. contradiction. intro x. destruct (seg_one_t_S_dec x) as [h2 | h2]. refine a. refine (h1 (exist (fun a => 1 <= a <= length l) _ h2)). Defined. Lemma list_to_seg_one_t_fun_compat : forall {T:Type} (l:list T), let f := list_to_seg_one_t_fun l in forall (n:nat) (pf:n < length l), let n' := exist _ _ (iff1 (lt_sot_iff _ _) pf) in In (f n') l /\ f n' = nth n l (hd' l (lt_length l n pf)). intros T l. induction l as [|a l h1]. intros f n h2. unfold f, list_to_seg_one_t_fun. simpl in h2. pose proof (lt_n_0 _ h2). contradiction. intros f n h2. simpl. simpl in h1. destruct n as [|n]. unfold f. simpl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. intuition. simpl in hr. omega. simpl in h2. pose proof (lt_S_n _ _ h2) as h3. specialize (h1 _ h3). unfold f. simpl. lr_match_goal; fold hlr; destruct hlr as [hl | hr]. simpl in hl. omega. simpl in hr. pose proof (proof_irrelevance _ hr (iff1 (lt_sot_iff _ _) h3)); subst. destruct h1 as [h1 h1']. split. right; auto. rewrite h1'. apply nth_indep; auto. Qed. Lemma list_to_seg_one_t_fun_no_dup_inj : forall {T:Type} (l:list T), NoDup l -> Injective (list_to_seg_one_t_fun l). intros T l h1. red. intros m n h2. destruct m as [m h3], n as [n h4]. apply proj1_sig_injective; simpl. pose proof (list_to_seg_one_t_fun_compat l (pred m)) as h5. pose proof (list_to_seg_one_t_fun_compat l (pred n)) as h6. assert (h7:pred m < length l). omega. assert (h8:pred n < length l). omega. specialize (h5 h7). specialize (h6 h8). assert (h9:S (pred n) = n). omega. assert (h10:S (pred m) = m). omega. assert (h11:list_to_seg_one_t_fun l (exist (fun a : nat => 1 <= a <= length l) (S (pred m)) (iff1 (lt_sot_iff _ _) h7)) = list_to_seg_one_t_fun l (exist (fun m0 : nat => 1 <= m0 <= length l) m h3)). f_equal. apply proj1_sig_injective; auto. assert (h12: list_to_seg_one_t_fun l (exist (fun a : nat => 1 <= a <= length l) (S (pred n)) (iff1 (lt_sot_iff _ _) h8)) = list_to_seg_one_t_fun l (exist (fun m0 : nat => 1 <= m0 <= length l) n h4)). f_equal. apply proj1_sig_injective; auto. simpl in h5. destruct h5 as [h5 h5']. simpl in h6. destruct h6 as [h6 h6']. rewrite h12 in h6'. rewrite h11 in h5'. rewrite h2 in h5'. rewrite h6' in h5'. apply nth_inj in h5'; auto. omega. Qed. Lemma im_list_to_segT_fun : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), Im (Full_set _) (list_to_segT_fun l) = list_to_set l. intros T hdt l. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]. subst. clear h1. rewrite <- list_to_set_in_iff. pose proof (list_to_segT_fun_compat l x) as h1. destruct h1 as [h1 h2]. rewrite h2. apply nth_In. apply proj2_sig. rewrite <- list_to_set_in_iff in h1. pose (exist (fun c => c < length l) (lind hdt _ _ h1) (lt_lind hdt _ _ h1)) as e. apply Im_intro with e. constructor. pose proof (list_to_segT_fun_compat l e) as h2. pose proof (nth_lt_lind_eq hdt _ _ h1) as h3. simpl in h3. destruct h2 as [h2a h2b]. rewrite h2b. rewrite <- h3 at 1. rewrite nth_lt_compat. apply nth_indep. apply proj2_sig. Qed. Definition list_to_seg_one_t_fun' {T:Type} (l:list T) : seg_one_t (length l) -> {x:T | In x l}. induction l as [|a l h1]; simpl. intro x. destruct x as [x h1]. destruct h1 as [h1 h2]. assert (h3:1 <= 0). eapply le_trans. apply h1. apply h2. apply le_Sn_0 in h3. contradiction. intro x. destruct (seg_one_t_S_dec x) as [h2 | h2]. refine (exist _ a _). left; auto. specialize (h1 (exist _ _ h2)). destruct h1 as [x' h1]. refine (exist _ x' _). right; auto. Defined. Lemma list_to_seg_one_t_fun_prime_compat : forall {T:Type} (l:list T) (x:seg_one_t (length l)), let f := list_to_seg_one_t_fun l in let f' := list_to_seg_one_t_fun' l in f x = proj1_sig (f' x). intros T l. induction l as [|a l h1]. intros x f f'. unfold f, f'. destruct x as [m h1]. simpl in h1. omega. intros x f f'. unfold f, f'. simpl. destruct (seg_one_t_S_dec x) as [h2 | h2]. simpl. reflexivity. rewrite h1. destruct ( @list_to_seg_one_t_fun' T l (@exist nat (fun m : nat => and (le (S O) m) (le m (@length T l))) (pred (@proj1_sig nat (fun m : nat => and (le (S O) m) (le m (S (@length T l)))) x)) h2)). simpl. reflexivity. Qed. Lemma list_to_seg_one_t_fun_compat' : forall {T:Type} (l:list T), let f := list_to_seg_one_t_fun' l in forall (n:nat) (pf:n < length l), let n' := exist _ _ (iff1 (lt_sot_iff _ _) pf) in proj1_sig (f n') = nth n l (hd' l (lt_length l n pf)). intros T l f n h1. simpl. pose proof (list_to_seg_one_t_fun_compat l n h1) as h2. simpl in h2. destruct h2 as [h2 h3]. unfold f. rewrite <- list_to_seg_one_t_fun_prime_compat. rewrite h3. reflexivity. Qed. Lemma list_to_seg_one_t_fun_no_dup_inj' : forall {T:Type} (l:list T), NoDup l -> Injective (list_to_seg_one_t_fun' l). intros T l h1. pose proof (list_to_seg_one_t_fun_no_dup_inj l h1) as h2. red in h2. red. intros x y h3. destruct x as [x h5], y as [y h6]. pose proof (f_equal (@proj1_sig _ _) h3) as h4. do 2 rewrite <- list_to_seg_one_t_fun_prime_compat in h4. specialize (h2 (exist _ _ h5) (exist _ _ h6) h4); auto. Qed. Lemma im_list_to_segT_fun' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), im_proj1_sig (Im (Full_set _) (list_to_segT_fun' l)) = list_to_set l. intros T hdt l. unfold im_proj1_sig. rewrite im_im. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]. subst. clear h1. rewrite <- list_to_set_in_iff. pose proof (list_to_segT_fun_compat' l x) as h1. rewrite h1 at 1. apply nth_In. apply proj2_sig. rewrite <- list_to_set_in_iff in h1. pose (exist (fun c=> c < length l) (lind hdt _ _ h1) (lt_lind hdt _ _ h1)) as e. apply Im_intro with e. constructor. pose proof (list_to_segT_fun_compat' l e) as h2. pose proof (nth_lt_lind_eq hdt l x h1) as h3. simpl in h3. rewrite h2 at 1. rewrite <- h3 at 1. rewrite nth_lt_compat. apply nth_indep. apply proj2_sig. Qed. Definition list_to_seg_one_t_fun'' {T:Type} (l:list T) : seg_one_t (length l) -> {x:T | Inn (list_to_set l) x}. intro x. pose (list_to_seg_one_t_fun' l x) as x'. pose (proj2_sig x') as h1. simpl in h1. rewrite list_to_set_in_iff in h1. refine (exist _ _ h1). Defined. Lemma list_to_seg_one_t_fun_compat'' : forall {T:Type} (l:list T), let f := list_to_seg_one_t_fun'' l in forall (n:nat) (pf:n < length l), let n' := exist _ _ (iff1 (lt_sot_iff _ _) pf) in proj1_sig (f n') = nth n l (hd' l (lt_length l n pf)). intros T l f n h1 n'. pose proof (list_to_seg_one_t_fun_compat' l n) as h2. simpl in h2. specialize (h2 h1). unfold f, list_to_seg_one_t_fun''. simpl. unfold n'. assumption. Qed. Lemma list_to_seg_one_t_fun_prime_compat' : forall {T:Type} (l:list T) (x:seg_one_t (length l)), let f' := list_to_seg_one_t_fun' l in let f'' := list_to_seg_one_t_fun'' l in proj1_sig (f'' x) = proj1_sig (f' x). intros T l x f' f''. unfold f', f''. unfold list_to_seg_one_t_fun''. simpl. reflexivity. Qed. Lemma list_to_seg_one_t_fun_no_dup_inj'' : forall {T:Type} (l:list T), NoDup l -> Injective (list_to_seg_one_t_fun'' l). intros T l h1. pose proof (list_to_seg_one_t_fun_no_dup_inj' l h1) as h2. red in h2. red. intros x y h3. specialize (h2 x y). apply h2. apply proj1_sig_injective. do 2 rewrite <- list_to_seg_one_t_fun_prime_compat'. f_equal; auto. Qed. Lemma im_list_to_segT_fun'' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), im_proj1_sig (Im (Full_set _) (list_to_segT_fun'' l)) = list_to_set l. intros T hdt l. unfold im_proj1_sig. rewrite im_im. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1]. subst. clear h1. rewrite <- list_to_set_in_iff. pose proof (list_to_segT_fun_compat'' l x) as h1. rewrite h1. apply nth_In. apply proj2_sig. rewrite <- list_to_set_in_iff in h1. pose (exist (fun c=> c < length l) (lind hdt _ _ h1) (lt_lind hdt _ _ h1)) as e. apply Im_intro with e. constructor. pose proof (list_to_segT_fun_compat'' l e) as h2. pose proof (nth_lt_lind_eq hdt l x h1) as h3. simpl in h3. rewrite h2. rewrite <- h3 at 1. rewrite nth_lt_compat. apply nth_indep. apply proj2_sig. Qed. Definition list_to_segT_card_fun {T:Type} (pfdt:type_eq_dec T) (l:list T) : card_fun1 (list_to_set l) ---> sig_set (list_to_set l) := fun x => let pf0 := lt_eq_trans _ _ _ (proj2_sig x) (card_fun1_length_eq pfdt l) in let pf1 := in_nth_lt _ _ pf0 in let pf2 := iff2 (list_singularize_in_iff pfdt l _) pf1 in let pf3 := iff1 (list_to_set_in_iff _ (nth_lt _ _ pf0)) pf2 in exist (fun c => Inn (list_to_set l) c) _ pf3. Lemma list_to_segT_card_fun_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:(card_fun1 (list_to_set l)) @), let f := list_to_segT_card_fun pfdt l in let pf0 := lt_eq_trans _ _ _ (proj2_sig x) (card_fun1_length_eq pfdt l) in proj1_sig (f x) = nth_lt _ _ pf0. intros T h1 l x f h2. unfold f. unfold list_to_segT_card_fun; simpl. apply nth_lt_functional3. Qed. Lemma list_to_segT_card_fun_inj : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), let f := list_to_segT_card_fun pfdt l in Injective f. intros T h1 l f. red; intros x y h2. unfold f, list_to_segT_card_fun in h2. apply exist_injective in h2. apply nth_lt_inj in h2. apply proj1_sig_injective in h2; auto. apply no_dup_list_singularize. Qed. Lemma list_to_segT_card_fun_surj : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), let f := list_to_segT_card_fun pfdt l in surjective f. intros T h1 l f. red; intro x. destruct x as [x h2]. pose proof h2 as h2'. rewrite <- list_to_set_in_iff in h2'. rewrite (list_singularize_in_iff h1) in h2'. pose proof (lt_lind h1 _ _ h2') as h3. rewrite <- card_fun1_length_eq in h3. exists (exist (fun c => c < card_fun1 (list_to_set l)) _ h3). unfold f, list_to_segT_card_fun. simpl. apply proj1_sig_injective; simpl. gen0. erewrite lind_functional. intro h4. erewrite nth_lt_functional3. rewrite lind_compat. reflexivity. reflexivity. Grab Existential Variables. assumption. Qed. Lemma list_to_segT_card_fun_bij : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), let f := list_to_segT_card_fun pfdt l in bijective f. intros; red; split; [apply list_to_segT_card_fun_inj; auto | apply list_to_segT_card_fun_surj; auto]. Qed. Lemma im_list_to_segT_card_fun : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), im_proj1_sig (Im (Full_set _) (list_to_segT_card_fun pfdt l)) = list_to_set l. intros T h1 l. unfold im_proj1_sig. rewrite im_im. unfold list_to_segT_card_fun. apply Extensionality_Ensembles; red; split; red; intros x h2. destruct h2 as [x h2]; subst. rewrite <- list_to_set_in_iff; simpl. rewrite (list_singularize_in_iff h1). apply in_nth_lt. simpl. rewrite <- list_to_set_in_iff in h2. rewrite (list_singularize_in_iff h1) in h2. pose proof (lt_lind h1 _ _ h2) as h3. rewrite <- (card_fun1_length_eq h1 l) in h3. apply Im_intro with (exist (fun c => c < card_fun1 (list_to_set l)) _ h3). constructor. simpl. symmetry. gen0. erewrite lind_functional. intro. erewrite nth_lt_functional3. rewrite lind_compat. reflexivity. reflexivity. Grab Existential Variables. assumption. Qed. (*This defines a lexicographic ordering relation on the list of naturals, which is useful in the next definition.*) (*The extra verbase [tso_dec] condition is included for compatibility with essential lemmas.*) Definition le_list_nat (l m:list nat) := le_listR lt (sort_dec_tso_dec _ sort_dec_lt_nat) l m. Definition lt_list_nat (l m:list nat) : Prop := lt_listR lt (sort_dec_tso_dec _ sort_dec_lt_nat) l m. End ListNat. Lemma length_list_of_lists_seqs : forall {T:Type} (ll:faml T), length (list_of_lists_seqs ll) = timesl (map (@length T) ll). intros T ll. induction ll as [|al ll ih]; simpl; auto. rewrite map_length. rewrite prod_length. f_equal; auto. Qed. Notation "l /= m" := (le_list_nat l m) (at level 20) : list_scope. Notation "l /_ m" := (lt_list_nat l m) (at level 20) : list_scope. Section ListNat'. Lemma lt_list_nat_impl_le : forall (l m:list nat), l /_ m -> l /= m. unfold lt_list_nat, le_list_nat; apply lt_list_impl_leR. Qed. Lemma not_lt_list_nat_nil : forall (l:list nat), ~ l /_ nil. unfold lt_list_nat; apply not_lt_listR_nil. Qed. Lemma le_list_nat_nil : forall (l:list nat), nil /= l. unfold lt_list_nat, le_list_nat; apply le_listR_nil. Qed. Lemma le_list_nat_dec : forall (l m:list nat), {l /= m} + {~l /= m}. unfold le_list_nat; apply le_listR_dec. Qed. Lemma refl_le_list_nat : forall (l:list nat), l /= l. unfold le_list_nat; apply refl_le_listR. red; apply lt_irrefl. Qed. Lemma antisym_le_list_nat : forall (l m:list nat), l /= m -> m /= l -> l = m. unfold le_list_nat; intros; eapply antisym_le_listR. apply H. assumption. Qed. Lemma antisym_not_le_list_nat : forall (l m:list nat), ~ l /= m -> ~ m /= l -> l = m. unfold le_list_nat; intros; eapply antisym_not_le_listR. apply H. apply H0. Qed. Lemma lt_list_nat_iff : forall (l m:list nat), l /_ m <-> ~ m /= l. unfold lt_list_nat, le_list_nat; intros; eapply lt_listR_iff. Qed. Lemma lt_list_nat_cons : forall (l m:list nat) (a:nat), (a::l) /_ (a::m) -> l /_ m. unfold lt_list_nat; intros; eapply lt_listR_cons. apply H. Qed. Lemma lt_list_nat_cons_iff : forall (l l':list nat) (a:nat) (pfn:l <> nil), l /_ (a::l') <-> (hd' l pfn) < a \/ (hd' l pfn = a /\ (tl l) /_ l'). unfold lt_list_nat; intros; apply lt_listR_cons_iff. Qed. Lemma le_list_nat_dec' : forall (l m:list nat), {l /= m} + {m /_ l}. unfold lt_list_nat, le_list_nat; eapply le_listR_dec'. Qed. Lemma lt_list_nat_tso_dec : tso_dec lt_list_nat. unfold lt_list_nat; intros; apply tso_dec_lt_listR. Qed. Lemma trans_le_list_nat : Transitive _ le_list_nat. unfold le_list_nat; intros; apply trans_le_listR. Qed. Lemma trans_le_lt_list_nat : forall (la lb lc:list nat), la /= lb -> lb /_ lc -> la /_ lc. unfold le_list_nat, lt_list_nat; intros; eapply trans_le_lt_listR. apply H. assumption. Qed. Lemma trans_lt_le_list_nat : forall (la lb lc:list nat), la /_ lb -> lb /= lc -> la /_ lc. unfold le_list_nat, lt_list_nat; intros; eapply trans_lt_le_listR. apply H. assumption. Qed. Lemma trans_lt_list_nat : Transitive _ lt_list_nat. unfold lt_list_nat; intros; apply trans_lt_listR. Qed. Lemma irrefl_lt_list_nat : forall (l:list nat), ~ l /_ l. unfold lt_list_nat; intros; apply irrefl_lt_listR. Qed. Lemma so_lt_list_nat : Strict_Order lt_list_nat. unfold lt_list_nat; intros; apply so_lt_listR. Qed. Lemma le_list_nat_lt_eq_dec : forall (l m:list nat), l /= m -> {l /_ m} + {l = m}. unfold lt_list_nat, le_list_nat; intros; apply le_listR_lt_eq_dec; auto. Qed. Lemma lt_list_nat_sort_dec : sort_dec lt_list_nat. apply lt_listR_sort_dec. Qed. Lemma lt_list_nat_irrefl : irrefl lt_list_nat. red; apply irrefl_lt_listR. Qed. Lemma lt_list_nat_sortable : sortable lt_list_nat. apply sort_dec_sortable; apply lt_list_nat_sort_dec. Qed. Lemma lt_list_nat_tso_dec_eq : lt_list_nat_tso_dec = sort_dec_tso_dec lt_list_nat lt_list_nat_sort_dec. apply functional_extensionality_dep. intro x. apply functional_extensionality_dep. intro y. destruct (lt_list_nat_tso_dec x y) as [h1 | h1]. destruct h1 as [h1 | h1]. destruct ( sort_dec_tso_dec lt_list_nat lt_list_nat_sort_dec x y) as [h2 | h2]. destruct h2 as [h2 | h2]. repeat f_equal. apply proof_irrelevance. pose proof h1 as h1'. rewrite h2 in h1'. apply lt_list_nat_irrefl in h1'. contradiction. pose proof (trans_lt_list_nat _ _ _ h1 h2) as h3. apply lt_list_nat_irrefl in h3. contradiction. subst. destruct (sort_dec_tso_dec lt_list_nat lt_list_nat_sort_dec y y) as [h1 | h1]. destruct h1 as [h1 | h1]. contradict (lt_list_nat_irrefl _ h1). repeat f_equal. apply proof_irrelevance. contradict (lt_list_nat_irrefl _ h1). destruct (sort_dec_tso_dec lt_list_nat lt_list_nat_sort_dec x y) as [h2 | h2]. destruct h2 as [h2 | h2]. pose proof (trans_lt_list_nat _ _ _ h1 h2) as h3. apply lt_list_nat_irrefl in h3. contradiction. subst. contradict (lt_list_nat_irrefl _ h1). f_equal. apply proof_irrelevance. Qed. Lemma list_nat_eq_dec_eq' : list_nat_eq_dec = sort_dec_eq_dec _ (lt_list_nat_sort_dec). apply functional_extensionality_dep. intro x. apply functional_extensionality_dep. intro y. destruct (list_nat_eq_dec x y) as [h1 | h1], (sort_dec_eq_dec lt_list_nat lt_list_nat_sort_dec x y) as [h2 | h2]; subst; f_equal; try apply proof_irrelevance. contradict (h2 eq_refl). contradict (h1 eq_refl). Qed. Lemma Rw_lt_list_nat_iff : forall l l':list nat, Rw lt_list_nat l l'<-> le_list_nat l l'. intros l l'. split; intro h1. red in h1. destruct h1 as [h1 |]; subst. apply lt_list_nat_impl_le; auto. apply refl_le_list_nat. apply le_list_nat_lt_eq_dec in h1. destruct h1 as [h1|]; subst. left; auto. right; auto. Qed. Lemma le0_nil : forall (l:list nat), l <> nil -> (0::nil) /= l. intro l. induction l as [|a l ih]. intro h1; contradict h1; auto. intro. unfold le_list_nat. simpl. destruct (sort_dec_tso_dec lt sort_dec_lt_nat 0 a) as [h1 | h1]. destruct h1; auto. omega. Qed. Lemma lt_lS : forall (l l':list nat), l' <> nil -> (0::l) /_ lS l'. intros l l'. destruct l' as [|a l']. intro h1; contradict h1; auto. simpl. split. simpl. destruct (sort_dec_tso_dec lt sort_dec_lt_nat 0 (S a)) as [h1 | h1]. destruct h1; auto. discriminate. omega. intro h1. inversion h1. Qed. Lemma lt_hd0_famlS : forall (l:list nat) (ll:faml nat) (pf:famlS ll <> nil), ~In nil ll -> (0::l) /_ (hd' (famlS ll) pf). intros l ll h1 hnin. pose proof (in_hd' (famlS ll) h1) as hi. pose proof h1 as h1'. rewrite famlS_eq_nil_iff in h1'. generalize h1. rewrite (hd_compat' _ h1'). rewrite famlS_cons. intro h0. simpl. apply lt_lS. intro h2. rewrite <- h2 in hnin. contradict hnin. apply in_hd'. Qed. Lemma no_dup_famlS : forall (ll:faml nat), NoDup ll -> NoDup (famlS ll). intro ll. induction ll as [|al ll ih]; simpl; auto. intro h1. pose proof (no_dup_cons_nin _ _ h1) as h2. pose proof (no_dup_cons _ _ h1) as h3. apply ih in h3. constructor. contradict h2. rewrite in_famlS_iff in h2. destruct h2 as [ln' [h4 h5]]. apply inj_lS in h5; subst; auto. assumption. Qed. Definition seg_faml (n:nat) := map (fun x => (x::nil)) (seg_list n) . Lemma sing_unq_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), sing_unq (linds pfdt l). intros T h1 l; red. intros lm ln h2 h3 h4. rewrite in_linds_iff in h2. destruct h2 as [x [[h2 h5] h8]]. rewrite in_linds_iff in h3. destruct h3 as [x' [[h3 h6] h7]]; subst. destruct (lt_list_nat_tso_dec (linds_occ h1 l x) (linds_occ h1 l x')) as [[h9 | h9] | h9]. symmetry. eapply incrR_no_dup_minimal. apply nat_eq_dec. apply increasing_linds_occ. apply no_dup_linds_occ. symmetry; auto. apply lt_list_nat_impl_le. assumption. assumption. eapply incrR_no_dup_minimal. apply nat_eq_dec. apply increasing_linds_occ. apply no_dup_linds_occ. assumption. apply lt_list_nat_impl_le. assumption. Qed. Lemma no_dup_in_linds_prod : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat), In ln (linds_prod pfdt l) -> NoDup ln. intros T h1 l ln h2. rewrite in_linds_prod_iff in h2. destruct h2 as [ha hb]. rewrite no_dup_iff. rewrite dup_iff'. intro h2. destruct h2 as [m [n [h2 [h3 [h4 h5]]]]]. pose proof h2 as h2'. pose proof h3 as h3'. rewrite ha in h2', h3'. pose proof (hb _ h2) as h6. pose proof (hb _ h3) as h7. destruct h6 as [h6 h8], h7 as [h7 h9]. rewrite h5 in h8. pose proof (partition_linds h1 l) as h12. pose proof (partition_faml_dec (linds h1 l) _ _ _ h12 (in_nth_lt _ _ h6) (in_nth_lt _ _ h7)) as h13. destruct h13 as [h13 | h13]. pose proof (in_linds_list_to_set_injable h1 l) as h14. red in h14. specialize (h14 (nth_lt (linds h1 l) m h6) (nth_lt (linds h1 l) n h7) (in_nth_lt _ _ _) (in_nth_lt _ _ _) h13). apply nth_lt_inj in h14. contradiction. apply no_dup_linds. do 2 red in h13. rewrite list_to_set_in_iff in h8, h9. pose proof (Intersection_intro _ _ _ _ h8 h9) as h14. rewrite h13 in h14. contradiction. apply nat_eq_dec. Qed. (*Maps corresponding parts between the partitions [linds] and [linds (remove).*) Definition linds_remove_fun {T:Type} (pfdt:type_eq_dec T) (l:list T) (r:T) (ln:list nat) (pfln:In ln (linds pfdt l)) : list nat := if (list_nat_eq_dec ln (linds_occ pfdt l r)) then nil else let v := in_linds_val _ _ _ pfln in linds_occ pfdt (remove pfdt r l) v. Lemma in_linds_remove_fun : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (r:T) (pf:In r l) (ln:list nat) (pfln:In ln (linds pfdt l)), ln <> linds_occ pfdt l r -> In (linds_remove_fun pfdt l r ln pfln) (linds pfdt (remove pfdt r l)). intros T h1 l r hr ln h2 h3. unfold linds_remove_fun. lr_if_goal; fold hlr; destruct hlr as [hl | hr']. contradiction. rewrite in_linds_iff. exists (in_linds_val h1 l ln h2). red. split. split; auto. rewrite in_remove_iff. split. apply in_linds_val_in. intro; subst. rewrite linds_occ_in_linds_val in hr'. contradict hr'; auto. intros x h4. destruct h4 as [h4 h5]. apply inj_linds_occ in h5; auto. rewrite in_remove_iff. split. apply in_linds_val_in. intro; subst. rewrite linds_occ_in_linds_val in hr'; contradict hr'; auto. Qed. Lemma length_linds_remove_fun' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (r:T), In r l -> forall (ln:list nat) (pfln:In ln (linds pfdt l)), ln <> linds_occ pfdt l r -> length (linds_remove_fun pfdt l r ln pfln) = length (linds_occ pfdt l (in_linds_val pfdt l ln pfln)). intros T h1 l r hir ln h0 h2. unfold linds_remove_fun. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. contradiction. unfold linds_occ. destruct in_dec as [h3 | h3]. destruct in_dec as [h4 | h4]. unfold linds_occ'. do 2 rewrite length_map_seg. rewrite count_occ_remove_neq; auto. intro; subst. rewrite linds_occ_in_linds_val in h2. contradict h2; auto. pose proof h3 as h3'. rewrite in_remove_iff in h3'. destruct h3'; contradiction. contradict h3. rewrite in_remove_iff. split. apply in_linds_val_in. intro; subst. rewrite linds_occ_in_linds_val in h2. contradict h2; auto. Qed. Lemma length_linds_remove_fun : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (r:T), In r l -> forall (ln:list nat) (pfln:In ln (linds pfdt l)), ln <> linds_occ pfdt l r -> length (linds_remove_fun pfdt l r ln pfln) = length ln. intros T h1 l r h2 ln h3 h4. unfold linds_remove_fun. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. contradiction. rewrite length_linds_occ; auto. rewrite count_occ_remove_neq; auto. rewrite count_occ_in_linds_val; auto. intro; subst. rewrite linds_occ_in_linds_val in hr. contradict hr; auto. Qed. Lemma count_occ_remove_in_linds_val : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (r:T), In r l -> forall (ln:list nat) (pfln:In ln (linds pfdt l)), ln <> linds_occ pfdt l r -> count_occ pfdt (remove pfdt r l) (in_linds_val pfdt l ln pfln) = length ln. intros T h1 l r h2 ln h3 h4. rewrite count_occ_remove_neq; auto. rewrite count_occ_in_linds_val; auto. intro; subst. rewrite linds_occ_in_linds_val in h4. contradict h4; auto. Qed. (*For each [ln] in [linds_prod] and an [x] in [l], calculates the element of [ln] which is the index of [x] in [l].*) Definition in_l_domain_linds_prod_ind {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfx:In x l) (ln:list nat) (pfln:In ln (linds_prod pfdt l)) : nat := let pfx' := iff1 (list_singularize_in_iff pfdt l x) pfx in let i := lind pfdt (list_singularize pfdt l) x pfx' in let pfeq := eq_sym (in_linds_prod_length pfdt l ln pfln) in let pflt := lt_eq_trans _ _ _ (lt_lind pfdt _ _ pfx') pfeq in nth_lt ln i pflt. Lemma in_l_domain_linds_prod_ind_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfx:In x l) (ln:list nat) (pfln:In ln (linds_prod pfdt l)), In (in_l_domain_linds_prod_ind pfdt l x pfx ln pfln) ln. unfold in_l_domain_linds_prod_ind; intros; apply in_nth_lt. Qed. Lemma lind_in_l_domain_linds_prod_ind_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfx:In x l) (ln:list nat) (pfln:In ln (linds_prod pfdt l)), let pfin := in_l_domain_linds_prod_ind_in pfdt l x pfx ln pfln in lind nat_eq_dec _ _ pfin < length (list_singularize pfdt l). intros T h1 l x h2 ln h3 h4. eapply lt_eq_trans. apply lt_lind. erewrite in_linds_prod_length. reflexivity. assumption. Qed. Lemma in_l_domain_linds_prod_ind_lt : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfx:In x l) (ln:list nat) (pfln:In ln (linds_prod pfdt l)), in_l_domain_linds_prod_ind pfdt l x pfx ln pfln < length l. intros T h1 l x h2 ln h3. pose proof (in_l_domain_linds_prod_ind_in h1 l x h2 ln h3) as h4. eapply in_linds_prod_lt. apply h3. assumption. Qed. Lemma lind_list_singularize : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfx:In x l) (ln:list nat) (pfln:In ln (linds_prod pfdt l)), let pfi := iff1 (list_singularize_in_iff pfdt _ _) pfx in lind pfdt _ x pfi = lind nat_eq_dec _ _ (in_l_domain_linds_prod_ind_in pfdt l x pfx ln pfln). intros T h1 l x h2 ln h3 h4. unfold in_l_domain_linds_prod_ind. symmetry. erewrite lind_functional. rewrite lind_nth_lt_no_dup. reflexivity. eapply no_dup_in_linds_prod. apply h3. apply nth_lt_functional3. Grab Existential Variables. erewrite in_linds_prod_length. apply lt_lind. assumption. apply in_nth_lt. Qed. Lemma in_l_domain_linds_prod_ind_compat : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfx:In x l) (ln:list nat) (pfln:In ln (linds_prod pfdt l)), nth_lt l _ (in_l_domain_linds_prod_ind_lt pfdt l x pfx ln pfln) = x. intros T h1 l x h2 ln h3. rewrite (nth_lt_eq_iff h1). exists h2. pose proof (lind_list_singularize h1 l x h2 ln h3) as h5. simpl in h5. pose proof h3 as h3'. rewrite in_linds_prod_iff in h3'. destruct h3' as [h6 h7]. pose proof (lind_in_l_domain_linds_prod_ind_lt h1 l x h2 ln h3) as h8. simpl in h8. rewrite <- h6 in h8. specialize (h7 _ h8). destruct h7 as [h7 h9]. unfold linds in h9. pose proof (nth_lt_map (list_singularize h1 l) (linds_occ h1 l) (lind nat_eq_dec ln (in_l_domain_linds_prod_ind h1 l x h2 ln h3) (in_l_domain_linds_prod_ind_in h1 l x h2 ln h3))) as h10. hfirst h10. rewrite <- h6. apply lt_lind. specialize (h10 hf h7). rewrite h10 in h9 at 1. rewrite in_linds_occ_iff in h9. destruct h9 as [occ [h12 [h9 h11]]]. symmetry in h11. erewrite nth_lt_functional3 in h11. erewrite lind_compat in h11. exists occ. simpl in h5. revert h9. generalize hf. rewrite <- h5. intro h13. erewrite nth_lt_functional3. rewrite lind_compat. intro h14. split; auto. rewrite h11. apply lind_occ_functional; auto. symmetry. gen0. rewrite <- h5. intro h15. erewrite nth_lt_functional3. rewrite lind_compat. reflexivity. Qed. Lemma in_l_domain_linds_prod_ind_compat' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (x:T) (pfx:In x l) (ln:list nat) (pfln:In ln (linds_prod pfdt l)), let pfx' := iff1 (list_singularize_in_iff pfdt _ _) pfx in let pfi := lind_in_l_domain_linds_prod_ind_lt pfdt l x pfx ln pfln in nth_lt (list_singularize pfdt l) _ pfi = x. Proof. intros T h1 l x h2 ln h3 h4 h5. pose proof (lind_list_singularize h1 l x h2 ln h3) as h6. simpl in h6. gen0. simpl. rewrite <- h6. intro h7. erewrite nth_lt_functional3. apply lind_compat. Qed. Definition two_lists_linds_prod_fun {T U:Type} {la:list T} {lb:list U} (pfdt:type_eq_dec T) (pfe:length la = length lb) (ln:list nat) (pfln:In ln (linds_prod pfdt la)) : fun_in_dep la U := (fun x pf => nth_lt _ _ (lt_eq_trans _ _ _ (in_l_domain_linds_prod_ind_lt pfdt la x pf ln pfln) pfe)). Definition two_lists_linds_sig_fun {T U:Type} {la:list T} {lb:list U} (pfdt:type_eq_dec T) (pfe:length la = length lb) (ln:list nat) (pfln:In ln (linds_prod pfdt la)) : sig_set (list_to_set la) -> U := (fun x => nth_lt _ _ (lt_eq_trans _ _ _ (in_l_domain_linds_prod_ind_lt pfdt la (proj1_sig x) (iff2 (list_to_set_in_iff _ _) (proj2_sig x)) ln pfln) pfe)). Lemma two_lists_linds_sig_fun_compat : forall {T U:Type} {la:list T} {lb:list U} (pfdt:type_eq_dec T) (pfe:length la = length lb) (ln:list nat) (pfln:In ln (linds_prod pfdt la)) (x:T) (pf:In x la), two_lists_linds_prod_fun pfdt pfe ln pfln x pf = two_lists_linds_sig_fun pfdt pfe ln pfln (exist _ _ (iff1 (list_to_set_in_iff _ _) pf)). unfold two_lists_linds_prod_fun, two_lists_linds_sig_fun; intros; simpl. apply nth_lt_functional2. f_equal. apply proof_irrelevance. Qed. Definition sort_set_nat (E:Ensemble nat) (pf:Finite E) : list nat := sort_set lt sort_dec_lt_nat E pf. Definition is_sorted_list_nat (l:list nat) : Prop := l = sort_set_nat (list_to_set l) (list_to_set_finite l). Lemma sort_set_nat_functional : forall (E E':Ensemble nat) (pf:Finite E) (pf':Finite E'), E = E' -> sort_set_nat E pf = sort_set_nat E' pf'. unfold sort_set_nat; intros; apply sort_set_functional; auto. Qed. Lemma in_list_reps_sort_set_nat : forall (E:Ensemble nat) (pf:Finite E), Inn (list_reps E) (sort_set_nat E pf). unfold sort_set_nat; intros; apply in_list_reps_sort_set; auto. Qed. Lemma in_sort_set_nat_iff : forall (E:Ensemble nat) (pf:Finite E) (m:nat), Inn E m <-> In m (sort_set_nat E pf). unfold sort_set_nat; intros; apply in_sort_set_iff; auto. Qed. Lemma inj_sort_set_nat : forall (A B:Ensemble nat) (pfa:Finite A) (pfb:Finite B), sort_set_nat A pfa = sort_set_nat B pfb -> A = B. unfold sort_set_nat; intros; eapply inj_sort_set; auto. apply H. Qed. (*Search below for [inj_sort_set_nat']*) Lemma no_dup_sort_set_nat : forall (E:Ensemble nat) (pf:Finite E), NoDup (sort_set_nat E pf). unfold sort_set_nat; intros; eapply no_dup_sort_set; auto. Qed. Lemma sort_set_nat_compat : forall (E:Ensemble nat) (pf:Finite E), list_to_set (sort_set_nat E pf) = E. unfold sort_set_nat; intros; eapply sort_set_compat; auto. Qed. Lemma sort_set_nat_empty : sort_set_nat (Empty_set nat) (Empty_is_finite _ ) = nil. unfold sort_set_nat; intros; eapply sort_set_empty; auto. Qed. Lemma sort_set_nat_nil : forall (A:Ensemble nat) (pf:Finite A), sort_set_nat A pf = nil -> A = Empty_set _. unfold sort_set_nat; intros; eapply sort_set_nil; auto. apply H. Qed. Lemma inh_sort_set_nat : forall (A:Ensemble nat) (pf:Finite A), Inhabited A -> sort_set_nat A pf <> nil. unfold sort_set_nat; intros; eapply inh_sort_set; auto. Defined. Lemma sort_set_nat_sing : forall (n:nat) (pf:Finite (Singleton n)), sort_set_nat (Singleton n) pf = n::nil. unfold sort_set_nat; intros; eapply sort_set_sing; auto. Qed. Lemma nin_hd_tl_sort_set_nat' : forall (A:Ensemble nat) (pff:Finite A) (pfia:Inhabited A), let l := sort_set_nat A pff in let pfic := inh_sort_set_nat A pff pfia in ~In (hd' l pfic) (tl l). simpl; unfold sort_set_nat, inh_sort_set_nat; intros; apply nin_hd_tl_sort_set'. Qed. Lemma hd_sort_set_nat' : forall (A:Ensemble nat) (pff:Finite A) (pfi:Inhabited A), hd' (sort_set_nat A pff) (inh_sort_set_nat A pff pfi) = min_set_nat A pff pfi. unfold sort_set_nat, min_set_nat; intros; erewrite hd_functional'; rewrite hd_sort_set' at 1. apply f_equal_prop_dep; auto. Qed. Lemma in_nat : forall (A:Ensemble nat) (pff:Finite A) (pfi:Inhabited A) (m:nat), Inn A m -> m = min_set_nat A pff pfi \/ In m (tl (sort_set_nat A pff)). unfold sort_set_nat, min_set_nat, lt_sortable; intros. pose proof (in_min lt sort_dec_lt_nat A pff pfi m H) as h1. simpl in h1. destruct h1 as [h1 | h1]. left; rewrite h1. apply f_equal_prop_dep; auto. right; auto. Qed. Lemma tl_sort_set_nat : forall (A:Ensemble nat) (pff:Finite A) (pfi:Inhabited A), tl (sort_set_nat A pff) = remove nat_eq_dec (min_set_nat A pff pfi) (sort_set_nat A pff). unfold sort_set_nat, min_set_nat, lt_sortable; intros. erewrite min_set_sortable_functional, sort_set_functional. erewrite tl_sort_set. f_equal. apply type_eq_dec_eq. reflexivity. reflexivity. Grab Existential Variables. apply sort_dec_lt_nat. assumption. assumption. Qed. Lemma in0_sort_set_nat : forall (A:Ensemble nat) (pf:Finite A) (pfi:Inn A 0), hd' (sort_set_nat A pf) (inh_sort_set_nat A pf (Inhabited_intro _ _ _ pfi)) = 0. intros A h1 h2. rewrite <- (hd_hd_compat' _ _ 0). unfold sort_set_nat, sort_set, min_set_sortable. destruct constructive_definite_description as [l h3]. simpl. destruct h3 as [h3 h4]. destruct l as [|a l]; simpl; auto. inversion h3 as [h5]. destruct h5 as [h5 h6]. clear h3. subst. pose proof (no_dup_cons_nin _ _ h6) as h5. rewrite <- list_to_set_in_iff in h2. destruct h2 as [|h2]; auto. specialize (h4 (0::a::remove nat_eq_dec 0 l)). hfirst h4. constructor. constructor. split. simpl. rewrite <- subtract_remove_compat. rewrite <- add_comm. f_equal. rewrite add_sub_compat_in; auto. rewrite <- list_to_set_in_iff. assumption. constructor. intro h7. destruct h7 as [| h7]; subst. contradiction. apply remove_In in h7 ;auto. constructor. contradict h5. rewrite in_remove_iff in h5. destruct h5; auto. pose proof (no_dup_cons _ _ h6) as h7. apply no_dup_remove; auto. rewrite in_sing_iff. intro h7. inversion h7 as [h8]. clear h7 H. subst. contradiction. specialize (h4 hf). red in h4. destruct h4 as [h4 h7]. simpl in h4. destruct (sort_dec_tso_dec lt sort_dec_lt_nat a 0) as [h8 | h8]. destruct h8 as [h8 | h8]. omega. assumption. omega. Qed. Lemma incr_sort_set_nat : forall (C:Ensemble nat) (pf:Finite C), increasing (sort_set_nat C pf). unfold sort_set_nat; intros; apply incrR_sort_set. Qed. Lemma incr_no_dup_minimal : forall (l l':list nat), increasing l -> NoDup l' -> list_to_set l = list_to_set l' -> l' /= l -> l = l'. unfold le_list_nat, increasing; intros; eapply incrR_no_dup_minimal. apply nat_eq_dec. apply H. assumption. assumption. apply H2. Qed. Lemma incr_impl_sort_set_nat : forall (l:list nat), increasing l -> l = sort_set_nat (list_to_set l) (list_to_set_finite _). unfold increasing, sort_set_nat; intros; eapply incrR_impl_sort_set; auto. Qed. Lemma sort_set_nat_cons : forall (A:Ensemble nat) (pf:Finite A) (ln:list nat) (a:nat), sort_set_nat A pf = a::ln -> ln = sort_set_nat (Subtract A a) (subtract_preserves_finite _ _ pf). unfold sort_set_nat; intros; eapply sort_set_cons; auto. Qed. Lemma max_set_nat_add_lst_sort_set: forall (A:Ensemble nat) (m:nat) (pf:Finite A) (pfi:Inhabited A), ~Inn A m -> let pf' := Add_preserves_Finite _ A m pf in let pfi' := inh_add A m in let pfn := inh_sort_set_nat _ _ pfi in m = max_set_nat (Add A m) pf' pfi' -> lst (sort_set_nat A pf) pfn < m. simpl; unfold sort_set_nat, max_set_nat; intros; eapply max_set_sortable_add_lst_sort_set; auto. erewrite max_set_sortable_functional. apply H0. reflexivity. Qed. Lemma sort_set_nat_app_add : forall (A A':Ensemble nat) (pfa:Finite A) (pfa':Finite A') (pfi:Inhabited A') (a:nat), ~Inn A a -> A' = Add A a -> a = max_set_nat A' pfa' pfi -> sort_set_nat A' pfa' = sort_set_nat A pfa ++ a::nil. unfold sort_set_nat, max_set_nat; intros; eapply sort_set_app_add; auto. rewrite H1. apply max_set_sortable_functional; auto. Grab Existential Variables. assumption. Qed. Lemma sort_set_nat_seg : forall (n:nat), sort_set_nat (seg n) (finite_seg n) = seg_list n. intro n. erewrite sort_set_nat_functional. erewrite incrR_impl_sort_set. apply sort_set_nat_functional. reflexivity. apply incr_seg_list. symmetry. apply list_to_set_seg_list_eq. Grab Existential Variables. rewrite list_to_set_seg_list_eq. apply finite_seg. Qed. Lemma tl_sort_set_nat_add : forall (A:Ensemble nat) (c:nat) (pf:Finite (Add A c)), list_to_set (tl (sort_set_nat (Add A c) pf)) = Subtract (Add A c) (min_set_nat _ pf (inh_add _ _)). unfold sort_set_nat, min_set_nat; intros. erewrite min_set_sortable_functional; auto; eapply tl_sort_set_add. Qed. Lemma in_tl_sort_set_nat : forall (A:Ensemble nat) (pf:Finite A) (m n:nat), Inn A m -> Inn A n -> m < n -> In n (tl (sort_set_nat A pf)). unfold sort_set_nat; intros; eapply in_tl_sort_set. apply H. assumption. assumption. Qed. Lemma list_to_set_lpred_sort_set_nat : forall (A:Ensemble nat) (pf:Finite A), list_to_set (lpred (sort_set_nat A pf)) = ens_pred A. intros A h1. apply Extensionality_Ensembles; red; split; red; intros x h4. rewrite <- list_to_set_in_iff in h4. unfold lpred in h4. rewrite in_map_iff in h4. destruct h4 as [m h4]. destruct h4 as [h4 h5]; subst. rewrite <- in_sort_set_nat_iff in h5. econstructor. apply h5. reflexivity. destruct h4 as [x h4]; subst. rewrite <- list_to_set_in_iff. unfold lpred. rewrite in_map_iff. exists x. split; auto. rewrite <- in_sort_set_nat_iff; auto. Qed. Lemma list_to_set_tl_lpred_sort_set_nat : forall (A:Ensemble nat) (pf:Finite A), Inn A 0 -> Inn A 1 -> list_to_set (tl (lpred (sort_set_nat A pf))) = ens_pred A. intros A h1 h2 h3. rewrite <- (list_to_set_lpred_sort_set_nat _ h1). apply Extensionality_Ensembles; red; split; red; intros m h4. rewrite <- list_to_set_in_iff in h4. rewrite <- list_to_set_in_iff. apply in_tl; auto. rewrite <- list_to_set_in_iff in h4. rewrite <- list_to_set_in_iff. rewrite tl_lpred. rewrite in_lpred_iff in h4. rewrite in_lpred_iff. destruct h4 as [m' h4]. destruct h4 as [h4 h5]; subst. rewrite <- in_sort_set_nat_iff in h4. apply (in_nat _ h1 (Inhabited_intro _ _ _ h2)) in h4. rewrite min_set_nat_0 in h4; auto. destruct h4 as [h4 | h4]; subst. exists 1. split. rewrite (tl_sort_set_nat _ _ (Inhabited_intro _ _ _ h2)). rewrite in_remove_iff. split. rewrite <- in_sort_set_nat_iff; auto. rewrite min_set_nat_0; auto. reflexivity. exists m'. split; auto. Qed. Lemma increasing_tl_lpred_sort_set_nat : forall (A:Ensemble nat) (pf:Finite A), increasing (tl (lpred (sort_set_nat A pf))). intros A h1. pose proof (finite_set_list_no_dup _ h1) as h2. destruct h2 as [l [h2 h3]]; subst. destruct (in_dec nat_eq_dec 0 l) as [h2 | h2], (in_dec nat_eq_dec 1 l) as [h4 | h4]. rewrite tl_lpred. rewrite list_to_set_in_iff in h2. pose proof h2 as h2'. rewrite (in_sort_set_nat_iff _ h1) in h2. pose proof (no_dup_sort_set_nat _ h1) as h5. rewrite <- (remove_hd_no_dup' nat_eq_dec _ (in_not_nil _ _ h2) h5). pose proof (in_inh _ _ h4) as h6. erewrite hd_functional'. rewrite (hd_sort_set_nat' _ _ h6) at 1. rewrite min_set_nat_0; auto. apply incr_lpred_remove_in0_in1. assumption. rewrite <- in_sort_set_nat_iff. rewrite <- list_to_set_in_iff; auto. apply incr_sort_set_nat. apply increasing_tl. apply incr_lpred_in0_nin1; auto. rewrite <- in_sort_set_nat_iff. rewrite <- list_to_set_in_iff; auto. rewrite <- in_sort_set_nat_iff. rewrite <- list_to_set_in_iff; auto. apply incr_sort_set_nat. apply increasing_tl. apply incr_lpred_nin0. rewrite <- in_sort_set_nat_iff. rewrite <- list_to_set_in_iff; auto. apply incr_sort_set_nat. apply increasing_tl. apply incr_lpred_nin0. rewrite <- in_sort_set_nat_iff. rewrite <- list_to_set_in_iff; auto. apply incr_sort_set_nat. Qed. Lemma lpred_sort_set_nat_in0_in1 : forall (A:Ensemble nat) (pf:Finite A), Inn A 0 -> Inn A 1 -> lpred (sort_set_nat A pf) = 0::sort_set_nat (ens_pred A) (finite_ens_pred _ pf). intros A h1 h2 h3. pose proof (in0_sort_set_nat A h1 h2) as h4. pose proof (hd_lpred' (sort_set_nat A h1)) as h5. hfirst h5. apply neq_nil_lpred. intro h6. pose proof h2 as h2'. rewrite (in_sort_set_nat_iff A h1) in h2'. rewrite h6 in h2'. contradict h2'. specialize (h5 hf). rewrite (hd_compat' _ hf). symmetry in h5. erewrite hd_functional' in h5. rewrite h4 in h5 at 1. simpl in h5. rewrite <- h5. f_equal. pose proof (incr_lpred_remove_in0_in1 (sort_set_nat A h1)) as h6. hfirst h6. rewrite <- in_sort_set_nat_iff; auto. specialize (h6 hf0). hfirst h6. rewrite <- in_sort_set_nat_iff; auto. specialize (h6 hf1). hfirst h6. apply incr_sort_set_nat. specialize (h6 hf2). pose proof (increasing_tl _ hf2) as h7. pose proof (list_to_set_tl_lpred_sort_set_nat A h1 h2 h3) as h8. symmetry. gen0. rewrite <- h8. intro h9. symmetry. erewrite sort_set_nat_functional. apply incr_impl_sort_set_nat. apply increasing_tl_lpred_sort_set_nat. reflexivity. Qed. Lemma lpred_sort_set_nat_nin0 : forall (A:Ensemble nat) (pf:Finite A), ~Inn A 0 -> lpred (sort_set_nat A pf) = sort_set_nat (ens_pred A) (finite_ens_pred _ pf). intros A h1 h2. pose proof (incr_lpred_nin0 (sort_set_nat A h1)) as h4. hfirst h4. contradict h2. rewrite <- in_sort_set_nat_iff in h2; auto. specialize (h4 hf). hfirst h4. apply incr_sort_set_nat. specialize (h4 hf0). symmetry. erewrite sort_set_nat_functional. symmetry. apply incr_impl_sort_set_nat. assumption. rewrite list_to_set_lpred_sort_set_nat; auto. Qed. Lemma lpred_sort_set_nat_in0_nin1 : forall (A:Ensemble nat) (pf:Finite A), Inn A 0 -> ~Inn A 1 -> lpred (sort_set_nat A pf) = sort_set_nat (ens_pred A) (finite_ens_pred _ pf). intros A h1 h2 h3. pose proof (incr_lpred_in0_nin1 (sort_set_nat A h1)) as h4. hfirst h4. rewrite <- in_sort_set_nat_iff; auto. specialize (h4 hf). hfirst h4. contradict h3. rewrite <- in_sort_set_nat_iff in h3; auto. specialize (h4 hf0). hfirst h4. apply incr_sort_set_nat. specialize (h4 hf1). symmetry. erewrite sort_set_nat_functional. symmetry. apply incr_impl_sort_set_nat. assumption. rewrite list_to_set_lpred_sort_set_nat; auto. Qed. Lemma list_minus_sort_set_nat : forall (A B:Ensemble nat) (pfa:Finite A) (pfb:Finite B) (pfab:Finite (Setminus A B)), list_minus nat_eq_dec (sort_set_nat A pfa) (sort_set_nat B pfb) = sort_set_nat (Setminus A B) pfab. intros A B h1 h2 h3. pose proof (list_to_set_list_minus nat_eq_dec (sort_set_nat A h1) (sort_set_nat B h2)) as h4. do 2 rewrite sort_set_nat_compat in h4. symmetry. gen0. rewrite <- h4. intro h5. symmetry. pose proof (proof_irrelevance _ h5 (list_to_set_finite _)); subst. apply incr_impl_sort_set_nat. apply incr_list_minus. apply incr_sort_set_nat. Qed. Definition fam_nat_to_ens_sort_set (F:Family nat) (pf:finite_mem F) : Ensemble (list nat) := fam_to_ens_sort_set lt sort_dec_lt_nat F pf. Lemma fam_nat_to_ens_sort_set_empty_eq : forall (pf:finite_mem (Empty_set (Ensemble nat))), fam_nat_to_ens_sort_set _ pf = Empty_set _. intro; unfold fam_nat_to_ens_sort_set; apply fam_to_ens_sort_set_empty_eq. Qed. Lemma fam_nat_to_ens_sort_set_sing_eq : forall (A:Ensemble nat) (pf:Finite A) (pffm:finite_mem (Singleton A)), fam_nat_to_ens_sort_set (Singleton A) pffm = Singleton (sort_set_nat A pf). unfold fam_nat_to_ens_sort_set; apply fam_to_ens_sort_set_sing_eq. Qed. Lemma finite_fam_nat_to_ens_sort_set : forall (F:Family nat) (pf:finite_mem F), Finite F -> Finite (fam_nat_to_ens_sort_set F pf). unfold fam_nat_to_ens_sort_set; apply finite_fam_to_ens_sort_set. Qed. Lemma inh_fam_nat_to_ens_sort_set : forall (F:Family nat) (pf:finite_mem F), Inhabited F -> Inhabited (fam_nat_to_ens_sort_set F pf). unfold fam_nat_to_ens_sort_set; apply inh_fam_to_ens_sort_set. Qed. Lemma fam_nat_to_ens_sort_set_functional : forall (F F':Family nat) (pf:finite_mem F) (pf':finite_mem F'), F = F' -> fam_nat_to_ens_sort_set F pf = fam_nat_to_ens_sort_set F' pf'. unfold fam_nat_to_ens_sort_set; apply fam_to_ens_sort_set_functional. Qed. Lemma fam_nat_to_ens_sort_set_subtract : forall (F:Family nat) (pff:Finite F) (pffm:finite_mem F) (l:list nat), Inn F (list_to_set l) -> increasing l -> Subtract (fam_nat_to_ens_sort_set F pffm) l = fam_nat_to_ens_sort_set (Subtract F (list_to_set l)) (finite_mem_subtract _ _ pffm ). unfold fam_nat_to_ens_sort_set; apply fam_to_ens_sort_set_subtract. Qed. Lemma im_fam_nat_to_ens_sort_set_list_to_set : forall (F:Family nat) (pf:finite_mem F), Im (fam_nat_to_ens_sort_set F pf) list_to_set = F. unfold fam_nat_to_ens_sort_set; apply im_fam_to_ens_sort_set_list_to_set. Qed. Lemma in_fam_nat_to_ens_sort_set_impl : forall (F:Family nat) (pf:finite_mem F) (l:list nat), Inn (fam_nat_to_ens_sort_set F pf) l -> Inn F (list_to_set l). unfold fam_nat_to_ens_sort_set; apply in_fam_to_ens_sort_set_impl. Qed. Lemma in_fam_nat_impl : forall (F:Family nat) (pffm:finite_mem F) (A:Ensemble nat) (pfin:Inn F A), Inn (fam_nat_to_ens_sort_set F pffm) (sort_set_nat A (pffm _ pfin)). unfold fam_nat_to_ens_sort_set, sort_set_nat; apply in_fam_impl. Qed. Lemma in_fam_nat_to_ens_sort_set_incr : forall (F:Family nat) (pffm:finite_mem F) (l:list nat), Inn (fam_nat_to_ens_sort_set F pffm) l -> increasing l. unfold fam_nat_to_ens_sort_set; apply in_fam_to_ens_sort_set_incr. Qed. Lemma setminus_sort_set_nat_fam : forall (A B:Family nat) (pfa:finite_mem A) (pfb:finite_mem B), let f := fun S : sig_set A => sort_set_nat (proj1_sig S) (pfa (proj1_sig S) (proj2_sig S)) in let g := fun S : sig_set B => sort_set_nat (proj1_sig S) (pfb (proj1_sig S) (proj2_sig S)) in let f' := restriction_sig f _ (setminus_inc _ _) in Setminus (Im (full_sig A) f) (Im (full_sig B) g) = Im (full_sig (Setminus A B)) f'. simpl; unfold sort_set_nat; apply setminus_sort_set_fam. Qed. (*Once again, the verbose [tso_dec] is for lemma compatibility.*) Definition min_list_nat (l m:list nat) : list nat := min_listR lt (sort_dec_tso_dec _ sort_dec_lt_nat) l m. Definition max_list_nat (l m:list nat) : list nat := max_listR lt (sort_dec_tso_dec _ sort_dec_lt_nat) l m. Lemma min_list_nat_compat1 : forall (l m:list nat), min_list_nat l m /= l. unfold min_list_nat, le_list_nat; eapply min_listR_compat1. Qed. Lemma min_list_nat_compat2 : forall (l m:list nat), min_list_nat l m /= m. unfold min_list_nat, le_list_nat; eapply min_listR_compat2. Qed. Lemma min_list_nat_eq_compat1 : forall (la lb lc:list nat), lb /= lc -> min_list_nat la lb /= min_list_nat la lc. unfold min_list_nat, le_list_nat; eapply min_listR_eq_compat1. Qed. Lemma min_list_nat_eq_compat2 : forall (la lb lc:list nat), la /= lb -> min_list_nat la lc /= min_list_nat lb lc. unfold min_list_nat, le_list_nat; eapply min_listR_eq_compat2. Qed. Lemma min_list_nat_nil1 : forall (l:list nat), min_list_nat nil l = nil. unfold min_list_nat; eapply min_listR_nil1. Qed. Lemma min_list_nat_nil2 : forall (l:list nat), min_list_nat l nil = nil. unfold min_list_nat; eapply min_listR_nil2. Qed. Lemma min_list_nat_eq_cons : forall (l m:list nat) (b:nat), min_list_nat (b :: l) (b :: m) = b:: min_list_nat l m. unfold min_list_nat; eapply min_listR_eq_cons. Qed. Lemma min_list_nat_comm : forall (l m:list nat), min_list_nat l m = min_list_nat m l. unfold min_list_nat; apply min_listR_comm. Qed. Lemma min_list_nat_assoc : forall (la lb lc:list nat), min_list_nat la (min_list_nat lb lc) = min_list_nat (min_list_nat la lb) lc. unfold min_list_nat; apply min_listR_assoc. Qed. Lemma min_list_nat_idem : forall (l:list nat), min_list_nat l l = l. unfold min_list_nat; apply min_listR_idem. Qed. Lemma min_list_nat_mono : forall (l la lb:list nat), la /= lb -> min_list_nat l la /= min_list_nat l lb. unfold min_list_nat, le_list_nat; apply min_listR_mono. Qed. Definition list_nat_fun : Type := list_fun nat nat. Definition incr_list_nat_fun (f:list_nat_fun) : Prop := incrR_list_fun lt lt (sort_dec_tso_dec _ sort_dec_lt_nat) (sort_dec_tso_dec _ sort_dec_lt_nat) f. Definition weak_incr_list_nat_fun (f:list_nat_fun) : Prop := weak_incrR_list_fun lt lt (sort_dec_tso_dec _ sort_dec_lt_nat) (sort_dec_tso_dec _ sort_dec_lt_nat) f. Lemma incr_list_nat_fun_impl_weak : forall (f:list_nat_fun), incr_list_nat_fun f -> weak_incr_list_nat_fun f. intros; do 3 red; intros. left. apply H in H0; auto. Qed. Lemma lt_map : forall (f:nat->nat), incr_fun f -> forall (l l':list nat), l /_ l' -> (map f l) /_ (map f l'). unfold lt_list_nat; intros; eapply lt_listR_map. red. apply lt_irrefl. rewrite incr_fun_iff in H. red in H; auto. apply H0. Qed. Lemma incr_list_nat_fun_lS : incr_list_nat_fun lS. do 3 red; intros; unfold lS; eapply lt_listR_map. red; apply lt_irrefl. change (incr_fun' S). rewrite <- incr_fun_iff. apply incr_fun_S; auto. apply H. Qed. Lemma weak_incr_list_nat_fun_lS : weak_incr_list_nat_fun lS. apply incr_list_nat_fun_impl_weak; apply incr_list_nat_fun_lS. Qed. Lemma incr_in0 : forall (l:list nat), increasing l -> In 0 l -> l = 0::tl l. intros l h1 h2. destruct l as [|a l]; simpl. contradiction. destruct h2 as [|h2]; subst; auto. pose proof (lt_increasing_cons _ _ a h2 h1). omega. Qed. Lemma linds_occ_nth_lt_eq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (n:nat) (pfn:n < length l) (ln:list nat), ln = linds_occ pfdt l (nth_lt _ _ pfn) <-> (increasing ln /\ (forall i:nat, (In i ln <-> exists pfm:i < length l, nth_lt l i pfm = nth_lt l n pfn))). intros T h1 l n h2 ln; split; intro h3; subst. split. apply increasing_linds_occ. intros i. rewrite in_linds_occ_iff. split; intro h3. destruct h3 as [m [h5 [h3 h4]]]. subst. exists (lt_lind_occ _ _ _ _ _). rewrite lind_occ_compat; auto. destruct h3 as [h3 h4]. pose proof (ex_lind_occ_nth_lt h1 _ _ h3) as h5. destruct h5 as [occ [h6 h7]]. exists occ, (in_nth_lt _ _ _). rewrite <- h4 at 1. split; auto. gen0. rewrite <- h4; auto. intro. rewrite <- h7. apply lind_occ_functional; auto. destruct h3 as [h3 h4]. pose proof list_to_set_injable_increasing as h5. red in h5. apply h5; auto. apply increasing_linds_occ. apply Extensionality_Ensembles; red; split; intros i h6. rewrite <- list_to_set_in_iff in h6. rewrite h4 in h6. destruct h6 as [h6 h7]. rewrite <- list_to_set_in_iff. rewrite in_linds_occ_iff. pose proof (ex_lind_occ_nth_lt h1 _ _ h6) as h8. destruct h8 as [occ [h8 h9]]. exists occ, (in_nth_lt _ _ _). split. rewrite h7 in h8; auto. gen0. rewrite <- h7. intro h10. rewrite <- h9. apply lind_occ_functional; auto. rewrite <- list_to_set_in_iff. rewrite h4. rewrite <- list_to_set_in_iff in h6. rewrite in_linds_occ_iff in h6. destruct h6 as [m [h8 [h6 h7]]]; subst. exists (lt_lind_occ _ _ _ _ _). rewrite lind_occ_compat; auto. Qed. Lemma in_linds_iff_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat), In ln (linds pfdt l) <-> (increasing ln /\ exists (n:nat) (pfn:n < length l), (forall i:nat, (In i ln <-> exists pfm:i < length l, nth_lt l i pfm = nth_lt l n pfn))). intros T h1 l ln. pose proof (linds_occ_nth_lt_eq_iff h1 l) as h0. split; intro h2. rewrite in_linds_iff in h2. destruct h2 as [x h2]. red in h2. destruct h2 as [[h2 h3] h4]; subst. specialize (h0 _ (lt_lind h1 _ _ h2) (linds_occ h1 l x)). hl h0. symmetry. f_equal. rewrite lind_compat; auto. rewrite h0 in hl. destruct hl as [h5 h6]. split; auto. exists _, (lt_lind h1 _ _ h2); auto. destruct h2 as [h2 h3]. destruct h3 as [i [h4 h5]]. pose proof (iff2 (h0 i h4 ln) (conj h2 h5)); subst. apply in_linds_occ_linds. apply in_nth_lt. Qed. Lemma incr_lec : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pfm:m < length l), increasing (lec pfdt pfm). unfold lec; intros; apply increasing_linds_occ. Qed. Lemma incr_lec_S : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pfm:m < length l), increasing (lec_S pfdt pfm). unfold lec_S; intros; apply incr_lS; apply increasing_linds_occ. Qed. Lemma incr_lec_cons : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (m:nat) (pfm:m < length l), increasing (lec_cons pfdt pfm a). unfold lec_cons; intros; apply incr_lec. Qed. Lemma lec_eq_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (m:nat) (pfm:m < length l) (ln:list nat), lec pfdt pfm = ln <-> (increasing ln /\ forall (r:nat), In r ln <-> exists (pfr:r < length l), nth_lt l r pfr = nth_lt l m pfm). intros T h1 l m h2 ln. pose proof (in_linds_iff_iff h1 l ln) as h3. split; intro h4; subst. pose proof (in_lec h1 l _ h2) as h4. rewrite h3 in h4. destruct h4 as [h4 h5]. split; auto. destruct h5 as [i [h5 h6]]. intros r; split; intro h7. pose proof h7 as h7'. rewrite h6 in h7'. destruct h7' as [ha hb]. exists ha; auto. specialize (h6 m). pose proof (rep_in_lec h1 l _ h2) as h10. rewrite h6 in h10. destruct h10 as [h10 h11]. pose proof (proof_irrelevance _ h10 h2); subst; auto. congruence. destruct h7 as [h7 h8]. rewrite in_lec_iff'. exists h7; auto. destruct h4 as [h4 h5]. apply list_to_set_injable_increasing. apply incr_lec. assumption. apply Extensionality_Ensembles; red; split; red; intros x h6. rewrite <- list_to_set_in_iff in h6. rewrite <- list_to_set_in_iff. pose proof (in_lec_lt h1 l m x h2 h6) as h7. rewrite (in_lec_iff _ _ _ _ _ h7) in h6. specialize (h5 x). hr h5. exists h7; auto. rewrite <- h5 in hr; auto. rewrite <- list_to_set_in_iff in h6. rewrite <- list_to_set_in_iff. rewrite in_lec_iff'. rewrite h5 in h6. destruct h6 as [h6 h7]. exists h6; auto. Qed. Lemma lec_Sn_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (m m':nat) (pfm:m < length l) (pfsm:S m < length (a::l)), a <> nth_lt l m pfm -> lec pfdt pfsm = lS (lec pfdt pfm). intros T h1 l a m m' h2 h3 h4. rewrite lec_eq_iff. split. apply incr_lS. apply incr_lec. intros r. split; intro h5. rewrite in_lS_iff in h5. destruct h5 as [r' [h5 h6]]; subst. simpl. rewrite in_lec_iff' in h5. destruct h5 as [h5 h6]. exists (lt_n_S _ _ h5). rewrite nth_lt_compat in h6. erewrite nth_indep in h6. rewrite h6 at 1. rewrite nth_lt_compat. apply nth_indep. assumption. assumption. destruct h5 as [h5 h6]. rewrite in_lS_iff. destruct r as [|r]. simpl in h6. contradict h4. rewrite nth_lt_compat. erewrite nth_indep. apply h6. assumption. simpl in h6. exists r. split; auto. rewrite in_lec_iff'. simpl in h5. exists (lt_S_n _ _ h5). erewrite nth_lt_compat. erewrite nth_indep. rewrite <- h6 at 1. rewrite nth_lt_compat. apply nth_indep. apply lt_S_n in h5; auto. assumption. Qed. Lemma nin_lec_S_cons_eq : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (n:nat) (pf:S n < length (a::l)), ~In a l -> lec_S pfdt (lt_S_n _ _ pf) = lec pfdt pf. intros T h1 l a n h3 h4. symmetry. rewrite lec_eq_iff. split. apply incr_lec_S. intro r. destruct r as [|r]. split; intro h5. rewrite in_lec_S_iff in h5. destruct h5; omega. destruct h5 as [h5 h6]. simpl in h6. contradict h4. rewrite h6. erewrite nth_indep. erewrite <- nth_lt_compat. apply in_nth_lt. simpl in h3. apply lt_S_n in h3; auto. unfold lec_S. rewrite in_lS_iff. split; intro h5. destruct h5 as [m h5]. rewrite in_lec_iff' in h5. destruct h5 as [[h5 h6] h7]. apply S_inj in h7; subst. exists (lt_n_S _ _ h5). simpl. rewrite nth_lt_compat in h6. erewrite nth_indep in h6. rewrite h6 at 1. rewrite nth_lt_compat. apply nth_indep. assumption. let P := type of h6 in match P with nth _ _ (nth_lt _ _ ?pf) = _ => apply pf end. destruct h5 as [h5 h6]. exists r. rewrite in_lec_iff. split; auto. simpl in h6. erewrite nth_lt_compat. erewrite nth_indep. rewrite <- h6 at 1. rewrite nth_lt_compat. apply nth_indep. simpl in h5; apply lt_S_n in h5; auto. simpl in h3; apply lt_S_n in h3; auto. Grab Existential Variables. simpl in h5; pose proof (lt_S_n _ _ h5); auto. simpl in h3; apply lt_S_n in h3; auto. Qed. Lemma lec_cons_neq_eq : forall {T:Type} (pfdt:type_eq_dec T) {l:list T} (m:nat) (pfm:m < length l) (a:T), a <> nth_lt l m pfm -> lec_cons pfdt pfm a = lec_S pfdt pfm. intros T h1 l m h2 a h3. symmetry. apply list_to_set_injable_increasing. apply incr_lec_S. apply incr_lec_cons. apply Extensionality_Ensembles; red; split; red; intros x h4; rewrite <- list_to_set_in_iff; rewrite <- list_to_set_in_iff in h4. rewrite in_lec_S_iff in h4. destruct h4 as [h4 [h5 h6]]. unfold lec_cons. rewrite in_lec_iff. simpl. destruct x as [|x]. contradict (lt_irrefl _ h4). simpl in h6. erewrite nth_lt_compat, nth_indep in h6. rewrite h6 at 1. rewrite nth_lt_compat. apply nth_indep. simpl in h5; auto. assumption. unfold lec_cons in h4. pose proof (in_lec_lt _ _ _ _ _ h4) as hlt. rewrite in_lec_iff in h4. simpl in h4. rewrite in_lec_S_iff. destruct x as [|x]. contradict h3. erewrite nth_lt_compat, nth_indep. rewrite h4 at 1; auto. assumption. split; auto with arith. simpl. simpl in hlt. apply lt_S_n in hlt. exists hlt. erewrite nth_lt_compat, nth_indep. rewrite h4 at 1. rewrite nth_lt_compat. apply nth_indep. assumption. assumption. Grab Existential Variables. assumption. simpl. omega. Qed. End ListNat'. Lemma plusl_map_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (f:T->nat) (x:T), NoDup l -> In x l -> plusl (map f l) = (f x) + (plusl (map f (remove pfdt x l))). intros T h0 l. induction l as [|a l h1]; simpl. intros; contradiction. intros f x h2 h3. destruct h3 as [h3l | h3r]. subst. destruct (h0 x x) as [h3 | h4]. pose proof (no_dup_cons_nin _ _ h2) as h4. rewrite <- remove_not_in'; auto. contradict h4. reflexivity. pose proof (no_dup_cons _ _ h2) as h4. specialize (h1 f x h4 h3r). rewrite h1. destruct (h0 x a) as [h5 | h6]. subst. pose proof (no_dup_cons_nin _ _ h2). contradiction. simpl. rewrite plus_assoc. rewrite (plus_comm (f a) (f x)). rewrite plus_assoc. reflexivity. Qed. Lemma plusl_nat_valued_fun_functional : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), NoDup l -> NoDup l' -> list_to_set l = list_to_set l' -> forall f:T->nat, plusl (map f l) = plusl (map f l'). intros T h0 l. induction l as [|a l h1];simpl; auto. intros l' h1 h2 h3. symmetry in h3. pose proof (empty_set_nil _ h3). subst. simpl. reflexivity. intros l' h2 h3 h4. pose proof (subtract_remove_compat h0 l' a) as h5. rewrite <- h4 in h5. pose proof (no_dup_cons_nin _ _ h2) as h6. rewrite list_to_set_in_iff in h6. rewrite sub_add_same_nin in h5; auto. pose proof (no_dup_remove h0 _ a h3) as h7. pose proof (no_dup_cons _ _ h2) as h8. specialize (h1 _ h8 h7 h5). intro f. rewrite h1. rewrite (plusl_map_remove h0 _ _ a h3). reflexivity. pose proof (Add_intro2 _ (list_to_set l) a) as h9. rewrite h4 in h9. rewrite list_to_set_in_iff. assumption. Qed. Lemma plusl_map_remove_no_dup : forall {T:Type} (pfdt:type_eq_dec T) (f:T->nat) (l:list T) (x:T) (pf:In x l), NoDup l -> plusl (map f l) = f x + plusl (map f (remove pfdt x l)). intros T h1 f l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h2 h3. destruct (h1 x a) as [|h4]; subst. f_equal. rewrite <- remove_not_in'; auto. apply no_dup_cons_nin in h3; auto. simpl. rewrite plus_assoc. rewrite (plus_comm (f x) (f a)). rewrite <- plus_assoc. f_equal. apply ih. destruct h2 as [|h2]; subst. contradict h4; auto. assumption. apply no_dup_cons in h3; auto. Qed. Lemma length_eq_plusl : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), length l = plusl (map (count_occ pfdt l) (list_singularize pfdt l)). intros T hdt l. pose proof (partition_linds hdt l) as h2. pose proof (partition_faml_length nat_eq_dec _ _ (no_dup_seg_list (length l)) (no_dup_linds hdt l) (no_dup_mem_linds _ _) (in_linds_list_to_set_injable hdt l) h2) as h3. rewrite length_seg_list in h3. rewrite <- h3. f_equal. unfold linds. rewrite map_map. apply map_ext_in. intros. rewrite length_linds_occ; auto. Qed. Lemma plusl_sort_list : forall (ln:list nat), plusl (sort_list _ tso_dec_lt_nat ln) = plusl ln. intro ln. induction ln as [|n ln ih]; simpl; auto. pose proof (sort_list_cons' _ sort_dec_lt_nat ln n) as h1. simpl in h1. pose proof (sort_dec_tso_dec_eq _ sort_dec_lt_nat) as h2. rewrite <- h2 at 1. rewrite h1. unfold weak_incrR_add. rewrite plusl_app; simpl. rewrite plus_assoc. rewrite (plus_comm _ n). rewrite <- plus_assoc. f_equal. rewrite <- ih. rewrite <- plusl_app. f_equal. symmetry. erewrite <- sort_list_decompose_ub_wlb. f_equal; auto. Qed. Lemma count_occs_eq_inv_iml : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T) (m:nat), (forall x:T, count_occ pfdt l x = count_occ pfdt l' x) -> list_to_set (inv_iml nat_eq_dec (count_occ pfdt l) (list_singularize pfdt l) m) = list_to_set (inv_iml nat_eq_dec (count_occ pfdt l') (list_singularize pfdt l') m). intros T h1 l l' m h2. pose proof (count_occs_eq_list_to_set _ _ _ h2) as h3. apply Extensionality_Ensembles; red; split; red; intros x h5; rewrite <- list_to_set_in_iff in h5; rewrite <- list_to_set_in_iff; rewrite in_inv_iml_iff in h5; rewrite in_inv_iml_iff; destruct h5 as [h5 h6]; subst. rewrite <- list_singularize_in_iff in h5; split; try rewrite <- list_singularize_in_iff; rewrite list_to_set_in_iff in h5; try rewrite list_to_set_in_iff. rewrite h3 in h5; auto. symmetry. apply h2; auto. rewrite <- list_singularize_in_iff in h5. split. rewrite <- list_singularize_in_iff. rewrite list_to_set_in_iff in h5. rewrite list_to_set_in_iff. rewrite <- h3 in h5; auto. apply h2; auto. Qed. Lemma count_occs_eq_sort_list_map : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (l':list T), (forall x:T, count_occ pfdt l x = count_occ pfdt l' x) -> sort_list lt tso_dec_lt_nat (map (count_occ pfdt l') (list_singularize pfdt l')) = sort_list lt tso_dec_lt_nat (map (count_occ pfdt l) (list_singularize pfdt l)). intros T hdt l l' h2. pose proof (count_occs_eq_impl _ _ _ h2) as h3. pose proof (count_occs_eq_list_to_set _ _ _ h2) as h4. eapply weak_incrR_length_count_occ_eq. apply sort_dec_lt_nat. erewrite <- sort_dec_tso_dec_eq at 1. apply weak_incrR_sort_list. erewrite <- sort_dec_tso_dec_eq at 1. apply weak_incrR_sort_list. rewrite length_sort_list. rewrite length_sort_list. do 2 rewrite map_length; auto. intro m. erewrite <- sort_dec_tso_dec_eq at 1. erewrite count_occ_sort_list. symmetry. erewrite <- sort_dec_tso_dec_eq at 1. erewrite count_occ_sort_list. pose proof (type_eq_dec_eq nat_eq_dec) as ht. rewrite <- ht at 1. symmetry. rewrite <- ht at 1. rewrite count_occ_map; auto. rewrite count_occ_map; auto. apply no_dup_list_to_set_eq_length; auto. eapply no_dup_coarse_list. apply coarse_list_inv_iml. apply no_dup_list_singularize. eapply no_dup_coarse_list. apply coarse_list_inv_iml. apply no_dup_list_singularize. apply count_occs_eq_inv_iml; auto. Grab Existential Variables. apply sort_dec_lt_nat. apply sort_dec_lt_nat. apply sort_dec_lt_nat. Qed. Lemma count_occs_eq_plusl_map : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), (forall x:T, count_occ pfdt l x = count_occ pfdt l' x) -> plusl (map (count_occ pfdt l) (list_singularize pfdt l)) = plusl (map (count_occ pfdt l') (list_singularize pfdt l')). intros T hdt l l' h1. rewrite <- plusl_sort_list. symmetry. rewrite <- plusl_sort_list. rewrite (count_occs_eq_sort_list_map hdt l l'); auto. Qed. Lemma count_occs_eq_length : forall {T:Type} (pfdt:type_eq_dec T) (l l':list T), (forall x:T, count_occ pfdt l x = count_occ pfdt l' x) -> length l = length l'. intros T hdt l l' h2. rewrite (length_eq_plusl hdt). rewrite (length_eq_plusl hdt). apply count_occs_eq_plusl_map; auto. Qed. Lemma sort_lists_eq_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l l':list T), let pfdt := sort_dec_eq_dec _ pfso in sort_list R (sort_dec_tso_dec _ pfso) l = sort_list R (sort_dec_tso_dec _ pfso) l' <-> (forall x, count_occ pfdt l x = count_occ pfdt l' x). intros T R hso l l' hdt; split; intro h1. intro x. pose proof (count_occ_sort_list R hso l x) as h2. pose proof (count_occ_sort_list R hso l' x) as h3. simpl in h2, h3. unfold hdt. rewrite <- h2, <- h3. rewrite h1; auto. pose proof h1 as hc. apply count_occs_eq_list_to_set in h1; auto. rewrite (list_to_set_sort_list R hso) in h1. symmetry in h1. rewrite (list_to_set_sort_list R hso) in h1. eapply weak_incrR_length_count_occ_eq. apply hso. apply weak_incrR_sort_list. apply weak_incrR_sort_list. do 2 rewrite length_sort_list. eapply list_to_sets_eq_length in h1. apply (count_occs_eq_length hdt); auto. intro x. do 2 rewrite count_occ_sort_list. apply hc. Grab Existential Variables. assumption. Qed. Lemma list_to_sets_eq_iff : forall {T:Type} (R:Relation T) (pfso:sort_dec R) (l l':list T), NoDup l -> NoDup l' -> (list_to_set l = list_to_set l' <-> sort_list R (sort_dec_tso_dec _ pfso) l = sort_list R (sort_dec_tso_dec _ pfso) l'). intros T R h1 l l' h2 h4; split; intro h5. rewrite sort_lists_eq_iff. intro x. destruct (in_dec (sort_dec_eq_dec _ h1) x l) as [h6 | h6], (in_dec (sort_dec_eq_dec _ h1) x l') as [h7 | h7]. rewrite count_occ_no_dup; auto. rewrite count_occ_no_dup; auto. rewrite list_to_set_in_iff in h6. rewrite h5 in h6. rewrite <- list_to_set_in_iff in h6. contradiction. rewrite list_to_set_in_iff in h7. rewrite <- h5 in h7. rewrite <- list_to_set_in_iff in h7. contradiction. rewrite (count_occ_In (sort_dec_eq_dec _ h1)) in h6. rewrite (count_occ_In (sort_dec_eq_dec _ h1)) in h7. omega. apply Extensionality_Ensembles; red; split; red; intros x h6; rewrite <- list_to_set_in_iff in h6; rewrite <- list_to_set_in_iff; rewrite <- in_sort_list_iff in h6; rewrite <- in_sort_list_iff. rewrite h5 in h6 at 1. apply h6. rewrite <- h5 in h6 at 1. apply h6. Qed. Section FamlNat. Lemma faml_nat_eq_dec : type_eq_dec (faml nat). apply list_eq_dec. apply list_nat_eq_dec. Qed. Definition minf (l:faml nat) (pf:l <> nil) : list nat := minlR lt_list_nat lt_list_nat_tso_dec l pf. Lemma in_minf : forall (l:faml nat) (pf:l <> nil), In (minf l pf) l. unfold minf; apply in_minlR. Qed. Lemma minf_compat : forall (ll:faml nat) (l:list nat) (pf:In l ll), minf ll (in_not_nil _ _ pf) /= l. unfold minf; intros; rewrite <- Rw_lt_list_nat_iff. apply minlR_compat. red; apply trans_lt_list_nat. Qed. Lemma minf_sing : forall (ll:faml nat) (l:list nat) (pfe:Singleton l = list_to_set ll), let pf := inh_sing l in let pfn := P_subst pf pfe in let pfn' := inh_not_nil _ pfn in l = minf ll pfn'. unfold minf; intros. erewrite minlR_functional'. apply minlR_sing. apply lt_list_nat_sort_dec. reflexivity. reflexivity. Grab Existential Variables. assumption. apply lt_list_nat_sort_dec. Qed. Lemma minf_in : forall (ll:faml nat) (l:list nat) (pf:In l ll), minf ll (in_not_nil _ _ pf) = minf (l::ll) (cons_neq_nil _ _). unfold minf, lt_list_nat. intros ll l h1. pose proof (minlR_in lt sort_dec_lt_nat ll l h1) as h2. assert (h3: (fun l0 m : list nat => lt_listR lt (sort_dec_tso_dec lt sort_dec_lt_nat) l0 m) = (lt_listR lt (sort_dec_tso_dec lt sort_dec_lt_nat))). apply functional_extensionality; intro al. apply functional_extensionality; intro bl. reflexivity. pose proof (minlR_functional' (lt_listR lt (sort_dec_tso_dec lt sort_dec_lt_nat)) (fun l0 m : list nat => lt_listR lt (sort_dec_tso_dec lt sort_dec_lt_nat) l0 m) (lt_listR_sort_dec _ _)) as h4. hfirst h4. rewrite h3. apply lt_listR_sort_dec. specialize (h4 hf (tso_dec_lt_listR _ _ ) ll ll (in_not_nil ll l h1) (in_not_nil ll l h1) h3 eq_refl). pose proof (sort_dec_tso_dec_eq (fun l0 m : list nat => lt_listR lt (sort_dec_tso_dec lt sort_dec_lt_nat) l0 m) hf lt_list_nat_tso_dec) as h8. gen0. rewrite <- h8. intro hn. pose proof (proof_irrelevance _ hn (in_not_nil ll l h1)); subst. rewrite <- h4. rewrite h2. eapply minlR_functional'. apply lt_listR_sort_dec. reflexivity. reflexivity. Qed. Definition le_faml_nat (ll lm:faml nat) : Prop := le_listR lt_list_nat (sort_dec_tso_dec _ lt_list_nat_sort_dec) ll lm. Definition lt_faml_nat (l m:faml nat) : Prop := le_faml_nat l m /\ l <> m. Notation "l //= m" := (le_faml_nat l m) (at level 20) : list_scope. Notation "l //_ m" := (lt_faml_nat l m) (at level 20) : list_scope. Lemma lt_faml_nat_impl_le : forall (l m:faml nat), l //_ m -> l //= m. unfold lt_faml_nat, le_faml_nat; apply lt_list_impl_leR. Qed. Lemma le_faml_nat_nil : forall (l:faml nat), nil //= l. intro; unfold le_list_nat; apply le_listR_nil. Qed. Lemma le_faml_nat_dec : forall (l m:faml nat), {l //= m} + {~l //= m}. unfold le_faml_nat; intros; apply le_listR_dec. Qed. Lemma refl_le_faml_nat : forall (l:faml nat), l //= l. unfold le_faml_nat; apply refl_le_listR. apply lt_list_nat_irrefl. Qed. Lemma antisym_le_faml_nat : forall (l m:faml nat), l //= m -> m //= l -> l = m. unfold le_faml_nat; apply antisym_le_listR. Qed. Lemma antisym_not_le_faml_nat : forall (l m:faml nat), ~ l //= m -> ~ m //= l -> l = m. unfold le_faml_nat; apply antisym_not_le_listR. Qed. Lemma lt_faml_nat_iff : forall (l m:faml nat), l //_ m <-> ~ m //= l. unfold lt_faml_nat; intros; apply lt_listR_iff. Qed. Lemma lt_faml_nat_cons : forall (l m:faml nat) (al:list nat), (al :: l) //_ (al :: m)-> l //_ m. unfold lt_faml_nat; apply lt_listR_cons. Qed. Lemma le_faml_nat_dec' : forall (l m:faml nat), {l //= m} + {m //_ l}. unfold le_faml_nat; apply le_listR_dec'. Qed. Lemma tso_lt_faml_nat : total_strict_order lt_faml_nat. unfold lt_faml_nat; apply tso_dec_total_strict_order; apply tso_dec_lt_listR. Qed. Lemma trans_le_faml_nat : Transitive _ le_faml_nat. unfold le_faml_nat; apply trans_le_listR. Qed. Lemma trans_le_lt_faml_nat : forall (la lb lc:faml nat), la //= lb -> lb //_ lc -> la //_ lc. unfold le_faml_nat, lt_faml_nat; eapply trans_le_lt_listR. Qed. Lemma trans_lt_le_faml_nat : forall (la lb lc:faml nat), la //_ lb -> lb //= lc -> la //_ lc. unfold lt_faml_nat, le_faml_nat; apply trans_lt_le_listR. Qed. Lemma trans_lt_faml_nat : Transitive _ lt_faml_nat. unfold lt_faml_nat; apply trans_lt_listR. Qed. Lemma irrefl_lt_faml_nat : forall (l:faml nat), ~ l //_ l. unfold lt_listR; apply irrefl_lt_listR. Qed. Lemma so_lt_faml_nat : Strict_Order lt_faml_nat. unfold lt_faml_nat; apply so_lt_listR. Qed. Lemma lt_faml_nat_sortable : sortable lt_faml_nat. unfold lt_faml_nat; apply lt_listR_sortable. Qed. Lemma count_occ_famlS_nin0 : forall (ll:faml nat) (l:list nat), ~In 0 l -> count_occ list_nat_eq_dec (famlS ll) l = count_occ list_nat_eq_dec ll (lpred l). intro ll. induction ll as [|al ll ih]; simpl; auto. intros l h0. destruct (list_nat_eq_dec (lS al) l) as [h1 | h1]. subst. destruct (list_nat_eq_dec al (lpred (lS al))) as [h2 | h2]. f_equal; auto. rewrite lpred_lS in h2; contradict h2; auto. destruct (list_nat_eq_dec al (lpred l)) as [h2 | h2]. rewrite h2 in h1. apply neq_sym in h1. rewrite <- lS_lpred_iff in h1. contradict h1; auto. auto. Qed. Lemma count_occ_famlS_in0 : forall (ll:faml nat) (l:list nat), In 0 l -> count_occ list_nat_eq_dec (famlS ll) l = 0. intros ll l h1. assert (h2:~In l (famlS ll)). intro h2. rewrite in_famlS_iff in h2. destruct h2 as [ln [h2 h3]]; subst. rewrite in_lS_iff in h1. destruct h1 as [? [? ?]]; discriminate. rewrite (count_occ_In list_nat_eq_dec) in h2; omega. Qed. Definition sort_family (F:Family nat) (pff:Finite F) (pffm:finite_mem F) : faml nat := match (finite_inh_or_empty_dec F pff) with | left pfi => sort_set lt_list_nat lt_list_nat_sort_dec (fam_nat_to_ens_sort_set F pffm) (finite_fam_nat_to_ens_sort_set _ _ pff) | right _ => nil end. Lemma sort_family_compat : forall (F:Family nat) (pf:Finite F) (pffm:finite_mem F), list_to_set (sort_family F pf pffm) = fam_nat_to_ens_sort_set F pffm. intros F h1 h2. unfold sort_family. apply Extensionality_Ensembles; red; split; red; intros x h3. rewrite <- list_to_set_in_iff in h3. lr_if h3; fold hlr in h3; destruct hlr as [hl | hr]. rewrite <- in_sort_set_iff in h3; auto. contradict h3. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. rewrite <- list_to_set_in_iff. rewrite <- in_sort_set_iff; auto. subst. rewrite fam_nat_to_ens_sort_set_empty_eq in h3. contradict h3. Qed. Definition is_sorted_family (ll:faml nat) : Prop := ll = sort_family (faml_to_fam ll) (finite_faml_to_fam _) (finite_mem_faml _). Lemma sort_family_functional : forall (F F':Family nat) (pff:Finite F) (pff':Finite F') (pfm:finite_mem F) (pfm':finite_mem F'), F = F' -> sort_family F pff pfm = sort_family F' pff' pfm'. intros F F' h1 h2 h3 h4 h5. subst. pose proof (proof_irrelevance _ h1 h2); pose proof (proof_irrelevance _ h3 h4); subst; auto. Qed. Lemma no_dup_sort_family : forall (F:Family nat) (pff:Finite F) (pffm:finite_mem F), NoDup (sort_family F pff pffm). intros E h1 h2. unfold sort_family. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. apply no_dup_sort_set. constructor. Qed. Lemma in_list_reps_sort_family : forall (F:Family nat) (pf:Finite F) (pffm:finite_mem F), Inn (list_reps (fam_nat_to_ens_sort_set F pffm)) (sort_family F pf pffm ). intros F h0 ha. unfold sort_family. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. constructor. split. rewrite sort_set_compat; auto. apply no_dup_sort_set. subst. rewrite fam_nat_to_ens_sort_set_empty_eq; auto. rewrite list_reps_empty. constructor. Qed. Lemma impl_in_sort_family : forall (F:Family nat) (pf:Finite F) (pffm:finite_mem F) (A:Ensemble nat) (pfa:Finite A), Inn F A -> In (sort_set_nat A pfa) (sort_family F pf pffm). intros F h1 h2 A h3 h4. unfold sort_family. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. rewrite <- in_sort_set_iff. erewrite sort_set_nat_functional. apply in_fam_nat_impl. reflexivity. subst. destruct h4. Grab Existential Variables. assumption. Qed. (*This is kind of weak and loses information, use the next.*) Lemma in_sort_family_impl : forall (F:Family nat) (pf:Finite F) (pffm:finite_mem F) (l:list nat), In l (sort_family F pf pffm) -> Inn F (list_to_set l). intros F h1 h2 l h5. unfold sort_family in h5. lr_if h5; fold hlr in h5; destruct hlr as [hl | hr]. rewrite <- in_sort_set_iff in h5. apply in_fam_nat_to_ens_sort_set_impl in h5; auto. contradict h5. Qed. (*This is stronger than the original.*) Lemma in_sort_family_impl' : forall (F:Family nat) (pf:Finite F) (pffm:finite_mem F) (l:list nat), In l (sort_family F pf pffm) -> exists (S:Ensemble nat) (pfi:Inn F S), l = sort_set_nat S (pffm _ pfi). intros F h1 h2 l h4. unfold sort_family in h4. lr_if h4; fold hlr in h4; destruct hlr as [hl | hr]. rewrite <- in_sort_set_iff in h4. destruct h4 as [l h4]; subst. clear h4. destruct l as [l h3]. exists _, h3. unfold sort_set_nat. apply sort_set_functional; simpl; auto. contradict h4. Qed. Lemma in_sort_family_unq : forall (F:Family nat) (pf:Finite F) (pffm:finite_mem F) (l l':list nat), list_to_set l = list_to_set l' -> In l (sort_family F pf pffm) -> In l' (sort_family F pf pffm) -> l = l'. intros F h1 h2 l l' h4 h5 h6. apply in_sort_family_impl' in h5. apply in_sort_family_impl' in h6. destruct h5 as [S [h5 h5']], h6 as [S' [h6 h6']]. subst. do 2 rewrite sort_set_nat_compat in h4. subst. apply sort_set_nat_functional; auto. Qed. Lemma sort_family_compat' : forall (F:Family nat) (pf:Finite F) (pffm:finite_mem F), faml_to_fam (sort_family F pf pffm) = F. intros F h1 h2. rewrite faml_to_fam_eq. rewrite map_im_compat. rewrite sort_family_compat. apply im_fam_nat_to_ens_sort_set_list_to_set. Qed. Lemma sort_family_sing : forall (A:Ensemble nat) (pff:Finite A) (pff':Finite (Singleton A)) (pffm:finite_mem (Singleton A)), sort_family (Singleton A) pff' pffm = (sort_set_nat A pff)::nil. intros ? h1. intros. unfold sort_family. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. simpl. gen0. erewrite fam_nat_to_ens_sort_set_sing_eq. intro. erewrite sort_set_sing. reflexivity. apply sing_not_empty in hr. contradiction. Qed. Lemma incl_sort_family_incl : forall (F F':Family nat) (pff:Finite F) (pff':Finite F') (pffm:finite_mem F) (pffm': finite_mem F'), Included F F' -> incl (sort_family F pff pffm) (sort_family F' pff' pffm'). intros F F' h1 h1' h2 h2' h4. rewrite incl_iff. do 2 rewrite sort_family_compat. red; intros x h5. destruct h5 as [x h5]. subst. clear h5. destruct x as [x h5]. simpl. pose proof (h4 _ h5) as h6. apply Im_intro with (exist _ _ h6). constructor; simpl. apply sort_set_functional; auto. Qed. Lemma in_is_sorted_family_ : forall (ll:faml nat) (l:list nat) (pf:In l ll), is_sorted_family ll -> is_sorted_list_nat l. intros ll l h1 h2. red in h2. red. simpl in h2. destruct ll as [|al ll]. red in h1. contradiction. pose proof h1 as h1'. rewrite list_to_set_in_iff in h1'. rewrite h2 in h1'. rewrite sort_family_compat in h1'. destruct h1' as [l h1']; subst. apply sort_set_nat_functional. rewrite sort_set_nat_compat. reflexivity. Qed. Lemma incr_impl_in_sort_family : forall (F:Family nat) (pf:Finite F) (pffm:finite_mem F) (l:list nat), increasing l -> Inn F (list_to_set l) -> In l (sort_family F pf pffm). intros F h1 h2 l h3 h4. pose proof (sort_family_compat F h1 h2) as h5. simpl in h5. destruct h5 as [h5 h6]. pose proof (finite_inh_or_empty _ h1) as h7. pose proof (incrR_impl_sort_set _ sort_dec_lt_nat l h3) as h10. destruct h7 as [h7 | h7]. rewrite h10. apply impl_in_sort_family. assumption. clear h10. subst; contradiction. Qed. Lemma inj_sort_family : forall (A B:Family nat) (pffa:Finite A) (pffma:finite_mem A) (pffb:Finite B) (pffmb:finite_mem B), sort_family A pffa pffma = sort_family B pffb pffmb -> A = B. intros A B h1 h1' h2 h2' h3. apply Extensionality_Ensembles; red; split; red; intros C h4. pose proof (impl_in_sort_family _ h1 h1' _ (h1' _ h4) h4) as hi. rewrite h3 in hi at 1. apply in_sort_family_impl in hi. rewrite sort_set_nat_compat in hi; auto. pose proof (impl_in_sort_family _ h2 h2' _ (h2' _ h4) h4) as hi. rewrite <- h3 in hi. apply in_sort_family_impl in hi. rewrite sort_set_nat_compat in hi; auto. Qed. Lemma sort_family_eq_nil : forall (F:Family nat) (pf:Finite F) (pffm:finite_mem F), sort_family F pf pffm = nil -> F = Empty_set _. intros F h1 h2 h3. do 2 red in h2. apply eq_empty. red; intros A h4. pose proof (impl_in_sort_family _ h1 h2 _ (h2 _ h4) h4) as h5. rewrite h3 in h5. contradiction. Qed. Definition increasing_faml (ll:faml nat) : Prop := increasingR lt_list_nat ll. Definition weak_increasing_faml (ll:faml nat) : Prop := weak_increasingR lt_list_nat ll. Lemma increasing_cons_faml : forall (ll:faml nat) (l:list nat), increasing_faml (l::ll) -> increasing_faml ll. unfold increasing_faml; apply increasingR_cons. Qed. Lemma weak_increasing_cons_faml : forall (ll:faml nat) (l:list nat), weak_increasing_faml (l::ll) -> weak_increasing_faml ll. unfold weak_increasing_faml; apply weak_increasingR_cons. Qed. Lemma increasing_impl_weak_faml : forall (ll:faml nat), increasing_faml ll -> weak_increasing_faml ll. unfold weak_increasing_faml; apply increasingR_impl_weak. Qed. Lemma lt_increasing_faml_cons : forall (ll:faml nat) (l l':list nat), In l ll -> increasing_faml (l'::ll) -> l' /_ l. unfold increasing_faml, lt_list_nat; intros; eapply lt_increasingR_cons. apply sort_dec_trans. apply lt_listR_sort_dec. apply H. assumption. Qed. Lemma le_weak_increasing_faml_cons : forall (ll:faml nat) (l l':list nat), In l ll -> weak_increasing_faml (l'::ll) -> l' /= l. unfold weak_increasing_faml, le_list_nat; intros. erewrite <- Rw_lt_listR_iff. eapply le_weak_increasingR_cons. apply trans_lt_list_nat. apply H. assumption. Qed. Lemma increasing_faml_cons_impl : forall (ll:faml nat) (l:list nat), increasing_faml (l::ll) -> increasing_faml ll. unfold increasing_faml. apply increasingR_cons_impl. Qed. Lemma weak_increasing_faml_cons_impl : forall (ll:faml nat) (l:list nat), weak_increasing_faml (l::ll) -> weak_increasing_faml ll. unfold weak_increasing_faml; apply weak_increasingR_cons_impl. Qed. Lemma increasing_faml_cons_impl_cons : forall (ll:faml nat) (l l':list nat), increasing_faml (l::ll) -> l' /_ l -> increasing_faml (l'::l::ll). unfold increasing_faml, lt_list_nat; apply increasingR_cons_impl_cons. Qed. Lemma waek_increasing_faml_impl_aux_cons : forall (ll:faml nat) (l l':list nat), weak_increasing_faml (l::ll) -> l' /= l -> weak_increasing_faml (l'::l::ll). unfold weak_increasing_faml, lt_list_nat; intros. rewrite <- Rw_lt_list_nat_iff in H0. apply weak_increasingR_cons_impl_cons; auto. Qed. Lemma increasing_faml_impl_cons : forall (ll:faml nat) (pf:ll <> nil) (l:list nat), increasing_faml ll -> l /_ hd' ll pf -> increasing_faml (l::ll). unfold increasing_faml, lt_list_nat; apply increasingR_impl_cons. Qed. Lemma weak_increasing_faml_impl_cons : forall (ll:faml nat) (pf:ll <> nil) (l:list nat), weak_increasing_faml ll -> l /= hd' ll pf -> weak_increasing_faml (l::ll). unfold weak_increasing_faml; intros. rewrite <- Rw_lt_list_nat_iff in H0. eapply weak_increasingR_impl_cons; auto. apply H0. Qed. Lemma increasing_faml_impl_cons' : forall (ll:faml nat) (l:list nat), increasing_faml ll -> (forall l':list nat, In l' ll -> l /_ l') -> increasing_faml (l::ll). unfold increasing; apply increasingR_impl_cons'. Qed. Lemma weak_increasing_faml_impl_cons' : forall (ll:faml nat) (l:list nat), weak_increasing_faml ll -> (forall l':list nat, In l' ll -> l /= l') -> weak_increasing_faml (l::ll). unfold weak_increasing_faml, le_list_nat; intros. eapply weak_increasingR_impl_cons'; auto. intros. rewrite Rw_lt_list_nat_iff. apply H0; auto. Qed. Lemma increasing_faml_tl : forall (ll:faml nat), increasing_faml ll -> increasing_faml (tl ll). unfold increasing_faml; apply increasingR_tl. Qed. Lemma weak_increasing_faml_tl : forall (ll:faml nat), weak_increasing_faml ll -> weak_increasing_faml (tl ll). unfold weak_increasing_faml; apply weak_increasingR_tl. Qed. Lemma increasing_faml_remove : forall (ll:faml nat) (l:list nat), increasing_faml ll -> increasing_faml (remove list_nat_eq_dec l ll). unfold increasing_faml; intros; eapply increasingR_remove; auto. apply trans_lt_list_nat. Qed. Lemma weak_increasing_faml_remove : forall (ll:faml nat) (l:list nat), weak_increasing_faml ll -> weak_increasing_faml (remove list_nat_eq_dec l ll). unfold weak_increasing_faml; intros; eapply weak_increasingR_remove; auto. apply trans_lt_list_nat. Qed. Lemma incr_faml_list_minus : forall (ll ll':faml nat), increasing_faml ll -> let lm := list_minus list_nat_eq_dec ll ll' in increasing_faml lm. unfold increasing_faml; intros; eapply incrR_list_minus; auto. apply trans_lt_list_nat. Qed. Lemma weak_incr_faml_list_minus : forall (ll ll':faml nat), weak_increasing_faml ll -> let lm := list_minus list_nat_eq_dec ll ll' in weak_increasing_faml lm. unfold weak_increasing_faml; intros; eapply weak_incrR_list_minus; auto. apply trans_lt_list_nat. Qed. Lemma increasing_faml_cons_lt_subst : forall (ll:faml nat) (l l':list nat), increasing_faml (l::ll) -> l' /_ l -> increasing_faml (l'::ll). unfold increasing_faml, lt_list_nat; intros; eapply increasingR_cons_lt_subst. apply trans_lt_list_nat. apply H. assumption. Qed. Lemma weak_increasing_faml_cons_le_subst : forall (ll:faml nat) (l l':list nat), weak_increasing_faml (l::ll) -> l' /= l -> weak_increasing_faml (l'::ll). unfold weak_increasing_faml, le_list_nat; intros; eapply weak_increasingR_cons_le_subst. apply trans_lt_list_nat. apply H. rewrite Rw_lt_list_nat_iff. assumption. Qed. Lemma increasing_faml_cons_no_dup : forall (ll:faml nat) (l:list nat), increasing_faml (l::ll) -> NoDup ll. unfold increasing_faml; intros; eapply increasingR_cons_no_dup. apply list_nat_eq_dec. apply so_lt_list_nat. apply H. Qed. Lemma increasing_faml_cons_nin : forall (ll:faml nat) (l:list nat), increasing_faml (l::ll) -> ~In l ll. unfold increasing_faml; intros; eapply increasingR_cons_nin. apply list_nat_eq_dec. apply so_lt_list_nat. apply H. Qed. Lemma increasing_faml_no_dup : forall (ll:faml nat) , increasing_faml ll -> NoDup ll. unfold increasing_faml; intros; eapply increasingR_no_dup. apply list_nat_eq_dec. apply so_lt_list_nat. apply H. Qed. Lemma increasing_faml_app : forall (ll ll':faml nat) (pfl:ll <> nil) (pfl':ll' <> nil), increasing_faml ll -> increasing_faml ll' -> lst ll pfl /_ hd' ll' pfl' -> increasing_faml (ll++ll'). unfold increasing_faml, lt_list_nat. intros; rewrite <- increasingR_app_iff. split. apply H. split; auto. apply H1. apply trans_lt_list_nat. Qed. Lemma weak_increasing_faml_app : forall (ll ll':faml nat) (pfl:ll <> nil) (pfl':ll' <> nil), weak_increasing_faml ll -> weak_increasing_faml ll' -> lst ll pfl /= hd' ll' pfl' -> weak_increasing_faml (ll++ll'). unfold weak_increasing_faml; intros; rewrite <- weak_increasingR_app_iff; auto. split; auto. split; auto. rewrite Rw_lt_list_nat_iff. apply H1. apply trans_lt_list_nat. Qed. Lemma no_dup_min_increasing_faml : forall (lm:faml nat) (C:Family nat) (pfc:Finite C) (pfm:finite_mem C), fam_nat_to_ens_sort_set C pfm = list_to_set lm -> NoDup lm -> (forall x : list (list nat), Inn (Subtract (list_reps (fam_nat_to_ens_sort_set C pfm)) lm) x -> lm //_ x) -> increasing_faml lm. intro lm. induction lm as [|al ll ih]; simpl; auto. intros; constructor. revert ih. destruct ll as [|bl ll]. intros. do 2 red; auto. intros h1 C h0 h2 h3 h4 h5. pose proof (h5 (bl::al::ll)) as h6. hfirst h6. constructor. constructor. simpl. rewrite add_comm. split; auto. apply no_dup_comm_cons; auto. rewrite in_sing_iff. intro h7. inversion h7; subst. pose proof (no_dup_cons_nin _ _ h4) as h8. contradict h8. left; auto. pose proof (no_dup_cons _ _ h4) as h20. pose proof (no_dup_cons2 _ _ _ h4) as h20'. pose proof (no_dup_cons_nin _ _ h4) as h21. pose proof (no_dup_cons2_nin _ _ _ h4) as h21'. split. specialize (h6 hf). red in h6. destruct h6 as [h6 h7]. simpl in h6. do 2 red in h6. destruct (sort_dec_tso_dec lt_list_nat lt_list_nat_sort_dec al bl) as [h9 | h9]. destruct h9 as [h9 | h9]. split; auto. pose proof (f_equal (fun D => Subtract D al) h3) as h10. simpl in h10. rewrite sub_add_same_nin in h10. red in h9. destruct h9; auto. intro h11. destruct h11 as [al h11 | al h11]. contradict h21'. rewrite <- list_to_set_in_iff in h11. assumption. destruct h11. contradict h21. left; auto. intro. subst. contradict h21. left; auto. subst. contradict h21. left; auto. contradiction. pose proof (f_equal (fun C => Subtract C al) h3) as h22. simpl in h22. assert (h10:al <> bl). intro; subst; contradict h21; left; auto. pose proof (Add_intro2 _ (Add (list_to_set ll) bl) al) as hc. simpl in h3. rewrite <- h3 in hc. pose proof (in_fam_nat_to_ens_sort_set_impl _ _ _ hc) as hc'. rewrite sub_add_same_nin in h22. pose proof (Add_intro2 _ (Add (list_to_set ll) bl) al) as h11. rewrite <- h3 in h11. pose proof (in_fam_nat_to_ens_sort_set_incr _ _ _ h11) as h12. rewrite fam_nat_to_ens_sort_set_subtract in h22; auto. specialize (h1 _ (subtract_preserves_finite _ _ h0) (finite_mem_subtract _ _ h2) h22 h20). hfirst h1. intros x h13. destruct h13 as [h13 h14]. rewrite in_sing_iff in h14. destruct h13 as [h13]. destruct h13 as [h13 h15]. specialize (h5 (al::x)). hfirst h5. constructor. constructor. simpl. rewrite <- h13. rewrite <- fam_nat_to_ens_sort_set_subtract. split. rewrite add_sub_eq. rewrite in_add_eq; auto. apply list_nat_eq_dec. constructor; auto. intro h16. rewrite list_to_set_in_iff in h16. rewrite <- h13 in h16. inversion h16 as [d h17 q h18]. destruct d as [d hd]. pose proof hd as hd'. rewrite h18 in hd' at 1. simpl in hd'. destruct hd' as [h19 h23]. rewrite in_sing_iff in h23; simpl in h23. rewrite sort_set_nat_compat in h23. contradict h23; auto. assumption. assumption. assumption. rewrite in_sing_iff. intro h17. inversion h17. contradiction. eapply lt_faml_nat_cons. apply h5; auto. specialize (h1 hf0); auto. intro h23. destruct h23 as [al h23 | al h23]. contradict h21'. rewrite list_to_set_in_iff; auto. destruct h23. contradict h10; auto. Qed. Lemma incr_sort_family : forall (C:Family nat) (pff:Finite C) (pffm:finite_mem C), increasing_faml (sort_family C pff pffm). unfold increasing_faml, sort_family; intros. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. apply incrR_sort_set. red; auto. Qed. Definition increasing_mem (ll:faml nat) : Prop := p_mem increasing (list_to_set ll). Lemma increasing_mem_sort_family : forall (A:Family nat) (pff:Finite A) (pffm:finite_mem A), increasing_mem (sort_family A pff pffm). intros A h1 h2. do 2 red. intros x h3. rewrite <- list_to_set_in_iff in h3. pose proof (in_sort_family_impl' _ h1 h2 _ h3) as h4. destruct h4 as [S [h4 h5]]. subst. apply incr_sort_set_nat. Qed. Lemma increasing_mem_list_minus : forall (ll ll':faml nat), increasing_mem ll -> increasing_mem (list_minus list_nat_eq_dec ll ll'). intros ll ll' h1. do 2 red in h1; do 2 red. intros x h2; apply h1. rewrite list_to_set_list_minus in h2. destruct h2; auto. Qed. Lemma increasing_mem_remove : forall (ll:faml nat) (l:list nat), increasing_mem ll -> increasing_mem (remove list_nat_eq_dec l ll). intros ll ll' h1. do 2 red in h1; do 2 red. intros x h2; apply h1. rewrite <- subtract_remove_compat in h2. destruct h2; auto. Qed. Lemma fam_nat_to_ens_list_faml_to_fam : forall (ll:faml nat), increasing_faml ll -> increasing_mem ll -> fam_nat_to_ens_sort_set (faml_to_fam ll) (finite_mem_faml ll) = list_to_set ll. intro ll. induction ll as [|al ll ih]; simpl. rewrite fam_nat_to_ens_sort_set_empty_eq; auto. unfold fam_nat_to_ens_sort_set, fam_to_ens_sort_set. rewrite im_full_sig_add. simpl. intros h0 hp. pose proof (increasing_faml_cons_impl _ _ h0) as h1. specialize (ih h1). rewrite <- ih. unfold fam_nat_to_ens_sort_set. f_equal. apply im_ext_in. intros x h2. unfold restriction_sig. simpl. apply sort_set_nat_functional; auto. symmetry. pose proof (hp _ (Add_intro2 _ _ _)) as h2. erewrite sort_set_nat_functional. apply incr_impl_sort_set_nat; auto. reflexivity. eapply p_mem_downward_closed. apply incl_add. apply hp. Qed. Lemma incr_faml_no_dup_minimal : forall (ll ll':faml nat), increasing_faml ll -> NoDup ll' -> list_to_set ll = list_to_set ll' -> ll' //= ll -> ll = ll'. unfold increasing_faml, le_list_nat. apply incrR_no_dup_minimal. apply list_nat_eq_dec. Qed. Lemma incr_impl_is_sorted_family : forall (ll:faml nat), increasing_faml ll -> increasing_mem ll -> is_sorted_family ll. intros l h1 h0; red. pose proof (increasing_faml_no_dup l h1) as hn. unfold sort_family. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. symmetry. unfold sort_set, min_set_sortable. destruct constructive_definite_description as [l' h2]; simpl. destruct h2 as [h2 h3]. destruct (tso_dec_lt_listR _ lt_list_nat_sort_dec l l') as [h4 | h4]. destruct h4 as [h4|]; auto. specialize (h3 l). hfirst h3. constructor; auto. constructor. split; auto. rewrite fam_nat_to_ens_list_faml_to_fam; auto. intro h5; destruct h5. apply (irrefl_lt_listR _ _ _ h4). specialize (h3 hf). pose proof (trans_lt_listR _ _ _ _ _ h3 h4) as h5. pose proof (irrefl_lt_listR _ _ _ h5); contradiction. destruct h2 as [h2], h2 as [h2 h5]. rewrite fam_nat_to_ens_list_faml_to_fam in h2. symmetry. eapply incr_faml_no_dup_minimal; auto. unfold le_faml_nat. apply lt_list_impl_leR; auto. assumption. assumption. rewrite faml_to_fam_eq in hr. apply empty_set_nil in hr. apply map_eq_nil in hr; auto. Qed. Lemma weak_incr_faml_no_dup_impl : forall (l:faml nat), weak_increasing_faml l -> NoDup l -> increasing_faml l. unfold weak_increasing_faml, increasing_faml; intros; apply weak_incrR_no_dup_impl; auto. Qed. Lemma weak_increasing_faml_map : forall (f:list_nat_fun) (ll:faml nat), weak_incr_list_nat_fun f -> weak_increasing_faml ll -> weak_increasing_faml (map f ll). unfold weak_increasing_faml; apply weak_increasingR_map. Qed. Lemma weak_incr_famlS : forall (ll:faml nat), weak_increasing_faml ll -> weak_increasing_faml (famlS ll). intros ll h1. unfold famlS. apply weak_increasing_faml_map; auto. apply weak_incr_list_nat_fun_lS; auto. Qed. (*Note that [linds] and [linds'] are unsorted since [list_singularize] filters backwards, so [list_singularize [a b c b a]] would be [c b a]. Since we wish to compare [linds]s, we need something canonical, i.e. sorted, to account for [list_singularize]'s permuting in friction against [sort]'s permuting. So when you want to compute or sort something like [linds (a::l)], you will need a canonical sorted list instead.*) Definition sorted_linds {T:Type} (pfdt:type_eq_dec T) (l:list T) : faml nat := sort_list lt_list_nat (sort_dec_tso_dec _ lt_list_nat_sort_dec) (linds pfdt l). Definition sorted_linds' {T:Type} (pfdt:type_eq_dec T) (l:list T) : faml nat := sort_list lt_list_nat (sort_dec_tso_dec _ lt_list_nat_sort_dec) (linds' pfdt l). Lemma sorted_linds_compat' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), sorted_linds pfdt l = sorted_linds' pfdt l. unfold sorted_linds', sorted_linds; intros; rewrite linds_compat'; auto. Qed. Lemma sorted_linds_nil' : forall {T:Type} (pfdt:type_eq_dec T), sorted_linds' pfdt nil = nil. unfold sorted_linds'; intros. rewrite linds_nil', sort_list_nil; auto. Qed. Lemma sorted_linds_nil : forall {T:Type} (pfdt:type_eq_dec T), sorted_linds pfdt nil = nil. intros; rewrite sorted_linds_compat'; apply sorted_linds_nil'. Qed. Lemma nin_nil_sorted_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), ~In nil (sorted_linds pfdt l). intros; intro h1. unfold sorted_linds in h1. rewrite in_sort_list_iff in h1. rewrite in_linds_iff in h1. destruct h1 as [x h2]; red in h2. destruct h2 as [h2 h3]. destruct h2 as [h2 h4]. symmetry in h4. unfold linds_occ in h4. destruct in_dec as [h5 | h5]. unfold linds_occ' in h4. apply map_seg_eq_nil in h4. rewrite (count_occ_In pfdt) in h2. omega. contradiction. Qed. Lemma nin_nil_sorted_linds' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), ~In nil (sorted_linds' pfdt l). intros; rewrite <- sorted_linds_compat'; apply nin_nil_sorted_linds; auto. Qed. Lemma sorted_linds_eq_nil_iff' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), sorted_linds' pfdt l = nil <-> l = nil. intros T h1 l; split; intro h2. unfold sorted_linds' in h2. rewrite (sort_list_eq_nil_iff _ lt_list_nat_sort_dec) in h2. rewrite linds_eq_nil_iff' in h2; auto. subst. apply sorted_linds_nil'. Qed. Lemma sorted_linds_eq_nil_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), sorted_linds pfdt l = nil <-> l = nil. intros; rewrite sorted_linds_compat'; apply sorted_linds_eq_nil_iff'. Qed. Lemma in_sorted_linds_iff : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (ln:list nat), In ln (sorted_linds pfdt l) <-> In ln (linds pfdt l). unfold sorted_linds; intros; split; intro h1. rewrite in_sort_list_iff in h1; auto. rewrite in_sort_list_iff; auto. Qed. Lemma length_sorted_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), length (sorted_linds pfdt l) = length (list_singularize pfdt l). unfold sorted_linds; intros; rewrite length_sort_list. unfold linds. apply map_length; auto. Qed. Lemma incr_sorted_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), increasing_faml (sorted_linds pfdt l). intros T h1 l. unfold increasing_faml, sorted_linds. apply weak_incrR_no_dup_impl. apply weak_incrR_sort_list. apply no_dup_sort_list. apply no_dup_linds. Qed. Lemma length_sorted_linds' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), length (sorted_linds' pfdt l) = length (list_singularize pfdt l). intros; rewrite <- sorted_linds_compat'; apply length_sorted_linds; auto. Qed. Lemma sorted_linds_eq_length : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), sorted_linds pfdt l = sorted_linds pfdu l' -> length (list_singularize pfdt l) = length (list_singularize pfdu l'). intros T U h1 h2 l l' h3. pose proof (length_sorted_linds h1 l) as h4. rewrite h3 in h4. rewrite length_sorted_linds in h4; auto. Qed. Lemma sorted_linds_eq_length' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), sorted_linds' pfdt l = sorted_linds' pfdu l' -> length (list_singularize pfdt l) = length (list_singularize pfdu l'). intros T U h1 h2 l l' h3. pose proof (length_sorted_linds' h1 l) as h4. rewrite h3 in h4. rewrite length_sorted_linds' in h4; auto. Qed. Lemma no_dup_sorted_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), NoDup (sorted_linds pfdt l). intros T h1 l. unfold sorted_linds. apply no_dup_sort_list. apply no_dup_linds. Qed. Lemma no_dup_sorted_linds' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), NoDup (sorted_linds' pfdt l). intros; rewrite <- sorted_linds_compat'; apply no_dup_sorted_linds; auto. Qed. Lemma incl_sorted_linds' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), incl (sorted_linds' pfdt l) (linds' pfdt l). intros T h1 l; red; intros l' h2. unfold sorted_linds' in h2. rewrite in_sort_list_iff in h2; auto. Qed. Lemma incl_sorted_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), incl (sorted_linds pfdt l) (linds pfdt l). intros; rewrite sorted_linds_compat', linds_compat'; apply incl_sorted_linds'. Qed. Lemma incl_linds_sorted : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), incl (linds' pfdt l) (sorted_linds' pfdt l). intros T h1 l; red; intros l' h2. unfold sorted_linds'. rewrite in_sort_list_iff ; auto. Qed. Lemma incl_linds_sorted' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), incl (linds pfdt l) (sorted_linds pfdt l). intros; rewrite sorted_linds_compat', linds_compat'; apply incl_linds_sorted. Qed. Lemma faml_to_fam_sorted_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), faml_to_fam (sorted_linds pfdt l) = faml_to_fam (linds pfdt l). intros T h1 l. apply Extensionality_Ensembles; red; split; red; intros ln h2; apply in_faml_to_fam_impl in h2; destruct h2 as [ln' [h2 h3]]; subst; unfold sorted_linds in h2. rewrite in_sort_list_iff in h2. apply in_faml_impl; auto. apply in_faml_impl. unfold sorted_linds. rewrite in_sort_list_iff; auto. Qed. (*We need to used [sorted_linds'] (rather than [linds']) to account for [list_singularize]'s permuting in friction against [sort]'s.*) Lemma sorted_linds_cons_in' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:In a l), let ln := linds_occ' pfdt l a pf in let ls := sorted_linds' pfdt l in sorted_linds' pfdt (a::l) = (0::(lS ln))::famlS (remove list_nat_eq_dec ln ls). intros T h1 l a h2; simpl. eapply (weak_incrR_length_count_occ_eq list_nat_eq_dec). apply lt_list_nat_sort_dec. apply weak_incrR_sort_list. simpl. destruct (nil_dec' (famlS (remove list_nat_eq_dec (linds_occ' h1 l a h2) (sorted_linds' h1 l)))) as [h3 | h3]. unfold famlS in h3. apply map_eq_nil in h3. rewrite remove_eq_nil_iff in h3. destruct h3 as [h3 | h3]. unfold sorted_linds' in h3. rewrite sort_list_eq_nil_iff in h3. rewrite linds_eq_nil_iff' in h3; subst. contradiction. apply (all_sing_remove' list_nat_eq_dec) in h3. rewrite h3; simpl; auto. rewrite (hd_compat' _ h3). split. left. apply lt_hd0_famlS. intro h4. rewrite in_remove_iff in h4. destruct h4 as [h4 h5]. apply nin_nil_sorted_linds' in h4; auto. rewrite <- hd_compat'. apply weak_incr_famlS. apply weak_increasingR_remove. apply trans_lt_list_nat. unfold sorted_linds. apply weak_incrR_sort_list. unfold sorted_linds' at 1. rewrite length_sort_list. rewrite length_linds'. simpl. rewrite list_singularize_in_cons; auto. rewrite length_famlS. rewrite length_remove_no_dup. rewrite length_sorted_linds'. rewrite list_singularize_in_iff in h2. rewrite <- (S_pred _ _ (not_nil_lt _ (in_not_nil _ _ h2))); auto. apply no_dup_sorted_linds'. unfold sorted_linds'. rewrite in_sort_list_iff. unfold linds'. rewrite in_map_in_dep_iff. exists a. ex_goal. rewrite <- list_singularize_in_iff; auto. exists hex. rewrite fun_in_dep_incl_compat. f_equal. apply proof_irrelevance. intro l'. unfold sorted_linds'. assert (h3:list_nat_eq_dec = sort_dec_eq_dec _ lt_list_nat_sort_dec). apply type_eq_dec_eq. rewrite h3. rewrite count_occ_sort_list. unfold linds'. destruct (list_nat_eq_dec l' (0::lS (linds_occ' h1 l a h2))) as [he | he]; subst. rewrite count_occ_cons_eq. rewrite count_occ_map_in_dep. rewrite <- h3. rewrite count_occ_famlS_in0. rewrite length_inv_im_in_dep_inj; auto. apply no_dup_list_singularize. apply inj_dep_incl. apply inj_dep_linds_occ'. eapply in_not_nil. rewrite in_inv_im_in_dep_iff. pose proof (in_eq a l) as h4. rewrite (list_singularize_in_iff h1) in h4. exists h4. unfold fun_in_dep_incl. erewrite linds_occ_functional'. rewrite linds_occ_cons_eq' at 1. simpl. reflexivity. reflexivity. reflexivity. left; auto. assumption. reflexivity. pose proof (no_dup_linds_occ' h1 (a :: l) _ (in_eq _ _)) as h4. pose proof (no_dup_map_in_dep_inj (fun_in_dep_incl (a :: l) (list_singularize h1 (a :: l)) (linds_occ' h1 (a :: l)) (incl_list_singularize h1 (a :: l))) (no_dup_list_singularize _ _)) as h5. hfirst h5. apply inj_dep_incl; auto. apply inj_dep_linds_occ'; auto. specialize (h5 hf). let P := type of h5 in match P with NoDup ?l' => pose l' as ml end. destruct (in_dec list_nat_eq_dec l' ml) as [h6 | h6]. rewrite count_occ_no_dup; auto. rewrite count_occ_no_dup; auto. constructor. intro h7. rewrite in_famlS_iff in h7. destruct h7 as [al [h7 h8]]. pose proof (in_eq 0 (lS (linds_occ' h1 l a h2))) as h9. rewrite h8 in h9. rewrite in_lS_iff in h9. destruct h9 as [h9 [h10 h11]]. discriminate. apply no_dup_map_inj. apply no_dup_remove. apply no_dup_sort_list. apply no_dup_map_in_dep_inj. apply no_dup_list_singularize. apply inj_dep_incl; auto. apply inj_dep_linds_occ'. apply inj_lS. right. unfold ml in h6. rewrite in_map_in_dep_iff in h6. destruct h6 as [x [h6 h7]]. unfold fun_in_dep_incl in h7. pose proof h6 as h6'. destruct (h1 x a) as [h8 | h8]; subst. contradict he. erewrite linds_occ_functional'. rewrite linds_occ_cons_eq' at 1. reflexivity. reflexivity. reflexivity. rewrite <- list_singularize_in_iff in h6'. destruct h6' as [|h6']; subst. contradict h8; auto. rewrite in_famlS_iff. exists (linds_occ' h1 l x h6'). split. rewrite in_remove_iff. split. rewrite in_sort_list_iff. rewrite in_map_in_dep_iff. exists x. ex_goal. rewrite <- list_singularize_in_iff; auto. exists hex. unfold fun_in_dep_incl. f_equal. apply proof_irrelevance. intro h9. apply inj_dep_linds_occ' in h9. contradiction. erewrite linds_occ_functional'. rewrite linds_occ_cons_neq'. reflexivity. apply h8. reflexivity. reflexivity. rewrite (count_occ_In (sort_dec_eq_dec _ lt_list_nat_sort_dec)) in h6. rewrite (not_lt0_iff (count_occ (sort_dec_eq_dec _ lt_list_nat_sort_dec) ml l')) in h6. fold ml. rewrite h6. simpl. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. rewrite hl in he. contradict he; auto. symmetry. rewrite <- not_lt0_iff. rewrite <- (count_occ_In (sort_dec_eq_dec _ lt_list_nat_sort_dec) (famlS (remove (sort_dec_eq_dec lt_list_nat lt_list_nat_sort_dec) (linds_occ' h1 l a h2) (sort_list lt_list_nat (sort_dec_tso_dec lt_list_nat lt_list_nat_sort_dec) (map_in_dep (fun_in_dep_incl l (list_singularize h1 l) (linds_occ' h1 l) (incl_list_singularize h1 l)))))) l'). intro h7. rewrite in_famlS_iff in h7. destruct h7 as [ln [h7 h8]]. rewrite in_remove_iff in h7. destruct h7 as [h7 h9]. rewrite in_sort_list_iff in h7. rewrite in_map_in_dep_iff in h7. destruct h7 as [x [h7 h10]]. destruct (h1 x a) as [|h11]; subst. unfold fun_in_dep_incl in h9. contradict h9. f_equal; apply proof_irrelevance. rewrite <- (not_lt0_iff (count_occ (sort_dec_eq_dec lt_list_nat lt_list_nat_sort_dec) ml (lS (fun_in_dep_incl l (list_singularize h1 l) (linds_occ' h1 l) (incl_list_singularize h1 l) x h7)))) in h6. rewrite <- (count_occ_In (sort_dec_eq_dec _ lt_list_nat_sort_dec) ml (lS (fun_in_dep_incl l (list_singularize h1 l) (linds_occ' h1 l) (incl_list_singularize h1 l) x h7))) in h6. contradict h6. unfold ml. rewrite in_map_in_dep_iff. exists x. ex_goal. rewrite <- list_singularize_in_iff. right; auto. pose proof (iff2 (list_singularize_in_iff h1 _ _) h7); auto. exists hex. unfold fun_in_dep_incl. erewrite linds_occ_functional'. rewrite linds_occ_cons_neq'. reflexivity. apply h11. reflexivity. reflexivity. Qed. Lemma sorted_linds_cons_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:In a l), let ln := linds_occ pfdt l a in let ls := sorted_linds pfdt l in sorted_linds pfdt (a::l) = (0::(lS ln))::famlS (remove list_nat_eq_dec ln ls). simpl; intros ? ? ? ? h2; do 2 rewrite sorted_linds_compat'; rewrite (linds_occ_compat' _ _ _ h2); apply sorted_linds_cons_in'. Qed. Lemma sorted_linds_cons_nin' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:~In a l), sorted_linds' pfdt (a::l) = (0::nil)::(famlS (sorted_linds' pfdt l)). intros T h1 l a h2. destruct (nil_dec' l) as [|hn]; subst. unfold sorted_linds', linds', fun_in_dep_incl; simpl. erewrite linds_occ_functional'. rewrite linds_occ_cons_eq_nin' at 1. rewrite sort_list_sing; auto. apply h2. reflexivity. reflexivity. eapply (weak_incrR_length_count_occ_eq list_nat_eq_dec). apply lt_list_nat_sort_dec. apply weak_incrR_sort_list. eapply weak_increasingR_impl_cons. apply weak_incr_famlS. apply weak_incrR_sort_list. rewrite Rw_lt_list_nat_iff. pose proof (lt_hd0_famlS nil (sorted_linds' h1 l)) as ha. hfirst ha. contradict hn. rewrite famlS_eq_nil_iff in hn. unfold sorted_linds' in hn. rewrite (sort_list_eq_nil_iff _ _ (linds' h1 l))in hn at 1. rewrite linds_eq_nil_iff' in hn; auto. specialize (ha hf). hfirst ha. intro hb. unfold sorted_linds' in hb. rewrite in_sort_list_iff in hb. unfold linds' in hb. rewrite in_map_in_dep_iff in hb. destruct hb as [x [hc hd]]. unfold fun_in_dep_incl in hd. unfold linds_occ' in hd. rewrite map_seg_eq_iff in hd. destruct hd as [hd]. simpl in hd. pose proof hc as hc'. rewrite <- list_singularize_in_iff in hc'. rewrite (count_occ_In h1) in hc'. omega. specialize (ha hf0). apply lt_list_nat_impl_le; auto. simpl. erewrite hd_functional'. apply ha. unfold sorted_linds'. unfold linds'. simpl. rewrite length_sort_list. rewrite length_map_in_dep. destruct (in_dec h1 a l) as [|h3]. contradiction. simpl. f_equal. rewrite list_singularize_nin; auto. rewrite length_famlS. rewrite length_sort_list. rewrite length_map_in_dep. reflexivity. intro ln. unfold sorted_linds'. assert (h3:list_nat_eq_dec = sort_dec_eq_dec _ lt_list_nat_sort_dec). apply type_eq_dec_eq. rewrite h3. rewrite count_occ_sort_list. simpl. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. subst. unfold linds'. rewrite count_occ_map_in_dep. rewrite length_inv_im_in_dep_inj. f_equal. symmetry. rewrite <- not_lt0_iff. intro h4. rewrite <- (count_occ_In (sort_dec_eq_dec lt_list_nat lt_list_nat_sort_dec) (famlS (sort_list lt_list_nat (sort_dec_tso_dec lt_list_nat lt_list_nat_sort_dec) (map_in_dep (fun_in_dep_incl l (list_singularize h1 l) (linds_occ' h1 l) (incl_list_singularize h1 l))))) (0 :: nil)) in h4 at 1. rewrite in_famlS_iff in h4. destruct h4 as [? [? h4]]. pose proof (in_eq 0 nil) as h5. rewrite h4 in h5. rewrite in_lS_iff in h5. destruct h5 as [? [? ?]]; discriminate. assumption. simpl. rewrite list_singularize_nin. constructor. rewrite <- list_singularize_in_iff; auto. apply no_dup_list_singularize. assumption. apply inj_dep_incl. apply inj_dep_linds_occ'. eapply (in_not_nil _ a). rewrite in_inv_im_in_dep_iff. ex_goal. rewrite <- list_singularize_in_iff. left; auto. exists hex. unfold fun_in_dep_incl. unfold linds_occ'. simpl. destruct (h1 a a) as [h4 | h4]. rewrite map_seg_eq_iff. simpl. ex_goal. f_equal. rewrite (count_occ_In h1) in h2; omega. exists hex0. intros m h5. destruct in_dec as [|h6]. contradiction. assert (h7:m = 0). omega. subst. reflexivity. contradict h4; auto. assumption. destruct (in_dec nat_eq_dec 0 ln) as [h4 | h4]. rewrite <- h3. rewrite count_occ_famlS_in0. rewrite <- not_lt0_iff. intro h5. rewrite <- (count_occ_In list_nat_eq_dec (linds' h1 (a::l)) ln) in h5. rewrite <- linds_compat' in h5. rewrite in_linds_iff in h5. destruct h5 as [x [[h5 h6] h7]]; subst. rewrite in_linds_occ_iff in h4. destruct h4 as [m [h6 [h8 h9]]]. destruct (h1 x a) as [|h10]; subst. simpl in h8, h9. destruct (h1 a a) as [h10 | h10]. pose proof h2 as h2'. rewrite (count_occ_In h1) in h2'. assert (m = 0). omega. subst. contradict hr. rewrite linds_occ_cons_eq. simpl; f_equal. symmetry. rewrite linds_occ_nin; auto. pose proof (h10 eq_refl); contradiction. simpl in h9. destruct (h1 x a) as [h11 | h11]. contradiction. discriminate. assumption. rewrite <- h3. rewrite count_occ_famlS_nin0; auto. rewrite h3. rewrite count_occ_sort_list. rewrite <- h3. apply count_occ_linds_cons_nin'; auto. Grab Existential Variables. intro h3. rewrite famlS_eq_nil_iff in h3. unfold sorted_linds' in h3. rewrite (sort_list_eq_nil_iff lt_list_nat lt_list_nat_sort_dec (linds' h1 l)) in h3. rewrite linds_eq_nil_iff' in h3. contradiction. Qed. Lemma sorted_linds_cons_nin : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:~In a l), sorted_linds pfdt (a::l) = (0::nil)::(famlS (sorted_linds pfdt l)). simpl; intros; do 2 rewrite sorted_linds_compat'; apply sorted_linds_cons_nin'; auto. Qed. Lemma sorted_linds_app_sing_in' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:In a l), let ll := sorted_linds' pfdt l in let ln := linds_occ' pfdt l a pf in sorted_linds' pfdt (l++a::nil) = sort_list _ lt_list_nat_tso_dec ((ln++length l::nil) :: (remove list_nat_eq_dec ln ll)). intros T h1 l a h2; simpl. eapply (weak_incrR_length_count_occ_eq list_nat_eq_dec). apply lt_list_nat_sort_dec. apply weak_incrR_sort_list. unfold sorted_linds'. rewrite lt_list_nat_tso_dec_eq. apply weak_incrR_sort_list. rewrite length_sorted_linds'. rewrite length_sort_list. simpl. rewrite list_singularize_rev_cons. rewrite app_length. simpl. rewrite S_compat. f_equal. rewrite length_remove_no_dup. rewrite length_remove_no_dup. rewrite length_sorted_linds'; auto. apply no_dup_sorted_linds'. unfold sorted_linds'. rewrite in_sort_list_iff. apply in_linds_occ_linds'. apply no_dup_list_singularize. rewrite <- list_singularize_in_iff. assumption. intro l'. rewrite list_nat_eq_dec_eq'. unfold sorted_linds'. rewrite lt_list_nat_tso_dec_eq. do 2 rewrite count_occ_sort_list. simpl. lr_if_goal; fold hlr; destruct hlr as [hl | hr]; subst. rewrite count_occ_no_dup. match goal with |- _ = S ?r' => assert (h3:r' = 0) end. rewrite <- not_lt0_iff. apply ngt_nlt. rewrite <- count_occ_In. intro h3. rewrite in_remove_iff in h3. destruct h3 as [h3 h4]. rewrite in_sort_list_iff in h3. pose proof (in_app_r (in_eq (length l) nil) (linds_occ' h1 l a h2)) as h5. rewrite in_linds_iff' in h3. destruct h3 as [x [h3 h7]]. destruct h3 as [h3 h6]. rewrite h6 in h5. rewrite <- linds_occ_compat' in h5. apply in_linds_occ_lt in h5. apply lt_irrefl in h5; auto. rewrite h3; auto. apply no_dup_linds'. rewrite in_linds_iff'. exists a. red. split. exists (in_app_r (in_eq a nil) _). symmetry. pose proof (proof_irrelevance _ (in_app_r (in_eq a nil) l) (in_app_l h2 (a::nil))) as h3. rewrite h3. rewrite (linds_occ_app_in' _ _ _ _ _ (in_eq _ _)) at 1. f_equal. unfold linds_occ'. simpl. destruct (h1 a a); simpl. rewrite plus_0_r; auto. contradict n; auto. intros x h3. destruct h3 as [h3 h4]. destruct (h1 a x) as [h5 | h5]; auto. pose proof h3 as h3'. apply in_app_or in h3'. destruct h3' as [h7 | h7]. pose proof (proof_irrelevance _ h3 (in_app_l h7 (a::nil))). subst. rewrite (linds_occ_app_nin_r' h1 _ _ _ h7) in h4. pose proof (in_app_r (in_eq (length l) nil) (linds_occ' h1 l a h2)) as h8. rewrite h4 in h8. rewrite <- linds_occ_compat' in h8. apply in_linds_occ_lt in h8. apply lt_irrefl in h8. contradiction. intro h8. destruct h8; subst. contradict h5; auto. contradiction. destruct h7; auto. contradiction. destruct (in_dec list_nat_eq_dec l' (linds' h1 (l++a::nil))) as [h3 | h3]. pose proof h3 as h3'. rewrite in_linds_iff' in h3'. destruct h3' as [x h5]. red in h5. destruct h5 as [h5 h7]. destruct h5 as [h5 h8]; subst. pose proof h5 as h5'. apply in_app_or in h5'. destruct h5' as [h8 | h8]. assert (h9:x <> a). intro; subst. contradict hr. pose proof (proof_irrelevance _ h5 (in_app_l h2 _)); subst. rewrite (linds_occ_app_in' _ _ _ _ _ (in_eq _ _)). f_equal. unfold linds_occ'. simpl. destruct (h1 a a) as [h9 | h9]; simpl. rewrite plus_0_r; auto. contradict h9; auto. pose proof (proof_irrelevance _ h5 (in_app_l h8 _)). subst. rewrite linds_occ_app_nin_r' in hr. rewrite count_occ_no_dup; auto. rewrite count_occ_no_dup; auto. apply no_dup_remove. apply no_dup_sort_list. apply no_dup_linds'. rewrite in_remove_iff. split. rewrite in_sort_list_iff. rewrite in_linds_iff'. exists x. red. split. exists h8. rewrite linds_occ_app_nin_r'; auto. intro h10. inversion h10; subst. contradict h9; auto. contradiction. intros x' h10. destruct h10 as [h10 h11]. rewrite linds_occ_app_nin_r' in h11. apply inj_dep_linds_occ' in h11; auto. intro h12. inversion h12; subst. contradict h9; auto. contradiction. rewrite linds_occ_app_nin_r'. intro h10. apply inj_dep_linds_occ' in h10. contradiction. intro h10. inversion h10. subst. contradict h9; auto. contradiction. apply no_dup_linds'. intro h10. inversion h10; subst. contradict h9; auto. contradiction. inversion h8; subst. pose proof (proof_irrelevance _ h5 (in_app_l h2 _)); subst. contradict hr. rewrite (linds_occ_app_in' _ _ _ _ _ (in_eq _ _)). f_equal. unfold linds_occ'. simpl. destruct (h1 x x) as [h10 | h10]. simpl. rewrite plus_0_r; auto. contradict h10; auto. contradiction. pose proof h3 as h3'. rewrite (count_occ_In (sort_dec_eq_dec lt_list_nat lt_list_nat_sort_dec)) in h3'. apply ngt_nlt in h3'. rewrite not_lt0_iff in h3'. rewrite h3'. symmetry. rewrite <- not_lt0_iff. apply ngt_nlt. rewrite <- (count_occ_In _). intro h4. rewrite in_remove_iff in h4. destruct h4 as [h4 h5]. rewrite in_sort_list_iff in h4. rewrite in_linds_iff' in h4. destruct h4 as [x h4]. red in h4. destruct h4 as [h4 h7]. destruct h4 as [h4 h8]. subst. contradict h3. rewrite in_linds_iff'. exists x. red. split. exists (in_app_l h4 (a::nil)). rewrite linds_occ_app_nin_r'; auto. intro h8. destruct h8; subst. contradict h5; f_equal; apply proof_irrelevance. contradiction. intros x' h8. destruct h8 as [h8 h9]. destruct (h1 x' a) as [h10 | h10]; subst. contradict hr; auto. rewrite h9. pose proof (proof_irrelevance _ h8 (in_app_l h2 (a::nil))); subst. rewrite (linds_occ_app_in' _ _ _ _ _ (in_eq _ _)). f_equal. unfold linds_occ'; simpl. destruct (h1 a a) as [h10 | h10]. simpl; rewrite plus_0_r; auto. contradict h10; auto. pose proof h8 as h11. apply in_app_or in h11. destruct h11 as [h11 | h12]. pose proof (proof_irrelevance _ h8 (in_app_l h11 (a::nil))); subst. rewrite linds_occ_app_nin_r' in h9. apply inj_dep_linds_occ' in h9; auto. contradict h10. destruct h10; auto. contradiction. destruct h12; subst. contradict h10; auto. contradiction. Qed. Lemma sorted_linds_app_sing_in : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:In a l), let ll := sorted_linds pfdt l in let ln := linds_occ pfdt l a in sorted_linds pfdt (l++a::nil) = sort_list _ lt_list_nat_tso_dec ((ln++length l::nil) :: (remove list_nat_eq_dec ln ll)). intros T h1 l a h2; simpl. do 2 rewrite sorted_linds_compat'. rewrite (linds_occ_compat' _ _ _ h2). apply sorted_linds_app_sing_in'. Qed. Lemma sorted_linds_app_sing_nin' : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:~In a l), sorted_linds' pfdt (l++a::nil) = sorted_linds' pfdt l ++ (length l :: nil) :: nil. intros T h1 l a h2. destruct (nil_dec' l) as [|h3]; subst; simpl. unfold sorted_linds'. rewrite <- linds_compat'. rewrite linds_sing. rewrite sort_list_sing; auto. eapply (weak_incrR_length_count_occ_eq list_nat_eq_dec). apply lt_list_nat_sort_dec. apply weak_incrR_sort_list. pose proof h3 as h3'. rewrite <- (sorted_linds_eq_nil_iff' h1) in h3. pose proof (in_lst _ h3) as h4. assert (h5:lst (sorted_linds' h1 l) h3 <> nil). intro h5. rewrite h5 in h4. pose proof (nin_nil_sorted_linds' h1 l); contradiction. pose proof (lt_list_nat_cons_iff (lst (sorted_linds' h1 l) h3) nil (length l) h5) as h6. pose proof (weak_increasingR_app_iff lt_list_nat trans_lt_list_nat (sorted_linds' h1 l) ((length l::nil)::nil) h3) (sing_not_nil _) as h7. rewrite <- h7. split. apply weak_incrR_sort_list. simpl. split; auto. left. rewrite h6. left. pose proof (in_linds_lt h1 l (lst (sorted_linds' h1 l) h3) _ (in_hd' _ h5)) as h8. hfirst h8. rewrite linds_compat'. apply incl_sorted_linds'. apply in_lst. eapply in_linds_lt. apply in_hd'. rewrite linds_compat'. eapply incl_sorted_linds'. apply in_lst. rewrite app_length. do 2 rewrite length_sorted_linds'. rewrite list_singularize_nin_app_nil. rewrite app_length; simpl; auto. assumption. intro l'. destruct (in_dec list_nat_eq_dec l' (sorted_linds' h1 (l++a::nil))) as [h4 | h4]. rewrite count_occ_no_dup; auto. pose proof h4 as h4'. unfold sorted_linds' in h4. rewrite in_sort_list_iff in h4. rewrite in_linds_iff' in h4. destruct h4 as [x [[h4 h5] h6]]. subst. rewrite count_occ_app. simpl. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. rewrite S_compat. f_equal. symmetry. rewrite <- not_lt0_iff. apply ngt_nlt. rewrite <- count_occ_In. intro h7. unfold sorted_linds' in h7. rewrite in_sort_list_iff in h7. rewrite in_linds_iff' in h7. destruct h7 as [x' [h7 h8]]. destruct h7 as [h7 h9]. rewrite <- hl in h9. pose proof (in_linds_occ_linds' h1 l x' h7) as h10. rewrite <- h9 in h10. rewrite <- linds_compat' in h10. pose proof (in_linds_lt h1 l _ _ (in_eq _ _) h10) as h11. apply lt_irrefl in h11; contradiction. assert (h7:x <> a). intro; subst. pose proof (proof_irrelevance _ h4 (in_app_r (in_eq _ _) l)); subst. rewrite linds_occ_app_nin_l' in hr. unfold linds_occ' in hr. simpl in hr. destruct (h1 a a) as [h8 | h8]. simpl in hr. rewrite plus_0_r in hr. contradict hr; auto. contradict (h8 eq_refl). assumption. pose proof h4 as h8. apply in_app_or in h8. destruct h8 as [h8 | h8]. rewrite count_occ_no_dup; auto with arith. apply no_dup_sorted_linds'. unfold sorted_linds' in h4'. rewrite in_sort_list_iff in h4'. rewrite in_linds_iff' in h4'. destruct h4' as [x' h10]. red in h10. destruct h10 as [h10 h11]. destruct h10 as [h12 h13]. apply inj_dep_linds_occ' in h13. subst. unfold sorted_linds'. rewrite in_sort_list_iff. rewrite in_linds_iff'. exists x'. red. split. exists h8. pose proof (proof_irrelevance _ h4 (in_app_l h8 _)); subst. rewrite linds_occ_app_nin_r'; auto. contradict h7. destruct h7; auto. contradiction. intros x h9. destruct h9 as [h9 h10]. pose proof (proof_irrelevance _ h4 (in_app_l h8 _)). subst. rewrite linds_occ_app_nin_r' in h10; auto. apply inj_dep_linds_occ' in h10; auto. contradict h7. destruct h7; auto. contradiction. destruct h8; subst. contradict h7; auto. contradiction. apply no_dup_sorted_linds'. pose proof h4 as h4'. rewrite (count_occ_In list_nat_eq_dec) in h4. apply ngt_nlt in h4. rewrite not_lt0_iff in h4. rewrite h4. symmetry. rewrite <- not_lt0_iff. apply ngt_nlt. rewrite <- count_occ_In. intro h5. unfold sorted_linds' in h5. apply in_app_or in h5. destruct h5 as [h5 | h5]. rewrite in_sort_list_iff in h5. contradict h4'. unfold sorted_linds'. rewrite in_sort_list_iff. rewrite in_linds_iff'. rewrite in_linds_iff' in h5. destruct h5 as [x h5]. red in h5. destruct h5 as [h5 h6]. destruct h5 as [h5 h7]; subst. exists x. red. split. ex_goal. eapply in_app_l; auto. exists hex. pose proof (proof_irrelevance _ hex (in_app_l h5 (a::nil))); subst. rewrite linds_occ_app_nin_r'; auto. intro h7; destruct h7; subst. contradiction. contradiction. intros x' h7. destruct h7 as [h7 h8]. assert (h9:x <> a). intro; subst; contradiction. pose proof @linds_occ_app_in'. pose proof (linds_occ_app_nin_r' h1 l (a::nil) x h5) as h10. hfirst h10. contradict h9. destruct h9; auto. contradiction. specialize (h10 hf). simpl in h10. rewrite h8 in h10. apply inj_dep_linds_occ' in h10; auto. destruct h5; subst. contradict h4'. unfold sorted_linds'. rewrite in_sort_list_iff. rewrite in_linds_iff'. exists a. red. split. exists (in_app_r (in_eq _ _) _). rewrite linds_occ_app_nin_l'. unfold linds_occ'. simpl. destruct (h1 a a) as [h5 | h5]. simpl. rewrite plus_0_r; auto. contradict h5; auto. assumption. intros x h5. destruct h5 as [h5 h6]. rewrite h6 in h4. rewrite <- not_lt0_iff in h4. apply nlt_ngt in h4. rewrite <- count_occ_In in h4. contradict h4. unfold sorted_linds'. rewrite in_sort_list_iff. apply in_linds_occ_linds'. contradiction. Qed. Lemma sorted_linds_app_sing_nin : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:~In a l), sorted_linds pfdt (l++a::nil) = sorted_linds pfdt l ++ (length l :: nil) :: nil. intros; do 2 rewrite sorted_linds_compat'; apply sorted_linds_app_sing_nin'; auto. Qed. Lemma list_to_set_linds_occ_eq : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (a:T) (a':U), list_to_set (linds pfdt (a::l)) = list_to_set (linds pfdu (a'::l')) -> linds_occ pfdt l a = linds_occ pfdu l' a'. intros T U h1 h2 l l' a a' h5. destruct (in_dec h1 a l) as [h3 | h3], (in_dec h2 a' l') as [h4 | h4]. pose proof (in_linds_occ_linds h1 (a::l) _ (in_eq _ _)) as h6. rewrite linds_occ_cons_eq in h6. simpl in h6. rewrite list_to_set_in_iff in h6. rewrite h5 in h6. rewrite <- list_to_set_in_iff in h6. rewrite in_linds_iff in h6. destruct h6 as [y h6]. red in h6. destruct h6 as [h6 h7]. destruct h6 as [h6 h8]. destruct h6 as [|h6]; subst. rewrite linds_occ_cons_eq in h8. simpl in h8. inversion h8. apply inj_lS in H0; auto. destruct (h2 y a') as [|h9]; subst. rewrite linds_occ_cons_eq in h8. simpl in h8. inversion h8. apply inj_lS in H0; auto. rewrite linds_occ_cons_neq in h8; auto. pose proof (in_eq 0 (lS (linds_occ h1 l a))) as h10. rewrite h8 in h10. apply nin0_lS in h10. contradiction. pose proof (in_linds_occ_linds h1 (a::l) _ (in_eq _ _)) as h6. rewrite list_to_set_in_iff in h6. rewrite h5 in h6. rewrite <- list_to_set_in_iff in h6. rewrite linds_occ_cons_eq in h6. simpl in h6. rewrite in_linds_iff in h6. destruct h6 as [y h6]. red in h6. destruct h6 as [h6 h7]. destruct h6 as [h6 h8]. destruct h6 as [|h6]; subst. rewrite linds_occ_cons_eq in h8. simpl in h8. inversion h8. apply inj_lS in H0; auto. assert (h9:y <> a'). intro; subst; contradiction. rewrite (linds_occ_cons_neq h2 _ _ _ h9) in h8. pose proof (in_eq 0 (lS (linds_occ h1 l a))) as h10. rewrite h8 in h10. apply nin0_lS in h10; contradiction. pose proof (in_linds_occ_linds h2 (a'::l') _ (in_eq _ _)) as h6. rewrite list_to_set_in_iff in h6. rewrite <- h5 in h6. rewrite <- list_to_set_in_iff in h6. rewrite linds_occ_cons_eq in h6. simpl in h6. rewrite in_linds_iff in h6. destruct h6 as [x [h6 h7]]. destruct h6 as [h6 h8]. destruct h6 as [|h6]; subst. rewrite linds_occ_cons_eq in h8. simpl in h8. inversion h8. apply inj_lS in H0; auto. assert (h9:x <> a). intro; subst; contradiction. rewrite linds_occ_cons_neq in h8; auto. pose proof (in_eq 0 (lS (linds_occ h2 l' a'))) as h10. rewrite h8 in h10. apply nin0_lS in h10. contradiction. do 2 (rewrite nin_linds_occ_eq_iff; auto). Qed. Lemma list_to_set_linds_from_cons : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (a:T) (a':U), list_to_set (linds pfdt (a::l)) = list_to_set (linds pfdu (a'::l')) -> list_to_set (linds pfdt l) = list_to_set (linds pfdu l'). intros T U h1 h2 l l' a a' h5. destruct (in_dec h1 a l) as [h3 | h3], (in_dec h2 a' l') as [h4 | h4]. pose proof (list_to_set_linds_occ_eq _ _ _ _ _ _ h5) as hl. erewrite list_to_set_sort_list in h5. symmetry in h5. erewrite list_to_set_sort_list in h5. pose proof (sorted_linds_cons_in h2 _ _ h4) as h6; simpl in h6. pose proof (sorted_linds_cons_in h1 _ _ h3) as h7; simpl in h7. unfold sorted_linds in h6, h7. rewrite h6, h7 in h5 at 1. simpl in h5. rewrite hl in h5. pose proof (f_equal (fun S => Subtract S (0::lS (linds_occ h2 l' a'))) h5) as ha. simpl in ha. rewrite sub_add_same_nin in ha. rewrite sub_add_same_nin in ha. unfold famlS in ha. rewrite map_im_compat in ha. rewrite map_im_compat in ha. rewrite <- subtract_remove_compat in ha. rewrite <- subtract_remove_compat in ha. rewrite <- list_to_set_sort_list in ha. rewrite <- list_to_set_sort_list in ha. rewrite im_subtract_inj in ha. rewrite im_subtract_inj in ha. apply in_sub_inj in ha. apply im_inj_inj in ha; auto. apply inj_lS. apply Im_intro with (linds_occ h2 l' a'). rewrite <- list_to_set_in_iff. apply in_linds_occ_linds. assumption. reflexivity. apply Im_intro with (linds_occ h2 l' a'). rewrite <- list_to_set_in_iff. rewrite <- hl. apply in_linds_occ_linds. assumption. reflexivity. apply inj_lS. apply inj_lS. intro hb. rewrite <- list_to_set_in_iff in hb. apply nin_famlS in hb; auto. left; auto. intro hb. rewrite <- list_to_set_in_iff in hb. apply nin_famlS in hb; auto. left; auto. erewrite list_to_set_sort_list in h5. symmetry in h5. erewrite list_to_set_sort_list in h5. pose proof (sorted_linds_cons_nin h2 _ _ h4) as h6; simpl in h6. pose proof (sorted_linds_cons_in h1 _ _ h3) as h7; simpl in h7. unfold sorted_linds in h6, h7. rewrite h6, h7 in h5 at 1. simpl in h5. pose proof (Add_intro2 _ (list_to_set (famlS (sort_list lt_list_nat (sort_dec_tso_dec lt_list_nat lt_list_nat_sort_dec) (linds h2 l')))) (0 :: nil)) as h8. rewrite h5 in h8. inversion h8; subst. rewrite <- list_to_set_in_iff in H. apply nin_famlS in H. contradiction. left; auto. inversion H. rewrite lS_eq_nil_iff in H1. pose proof (in_lind_linds_occ h1 l a h3) as h9. rewrite H1 in h9. contradiction. erewrite list_to_set_sort_list in h5. symmetry in h5. erewrite list_to_set_sort_list in h5. pose proof (sorted_linds_cons_in h2 _ _ h4) as h6; simpl in h6. pose proof (sorted_linds_cons_nin h1 _ _ h3) as h7; simpl in h7. unfold sorted_linds in h6, h7. rewrite h6, h7 in h5 at 1. simpl in h5. pose proof (Add_intro2 _ (list_to_set (famlS (sort_list lt_list_nat (sort_dec_tso_dec lt_list_nat lt_list_nat_sort_dec) (linds h1 l)))) (0 :: nil)) as h8. rewrite <- h5 in h8. inversion h8; subst. rewrite <- list_to_set_in_iff in H. apply nin_famlS in H. contradiction. left; auto. inversion H. rewrite lS_eq_nil_iff in H1. pose proof (in_lind_linds_occ h2 l' a' h4) as h9. rewrite H1 in h9. contradiction. erewrite list_to_set_sort_list in h5. symmetry in h5. erewrite list_to_set_sort_list in h5. pose proof (sorted_linds_cons_nin h2 _ _ h4) as h6; simpl in h6. pose proof (sorted_linds_cons_nin h1 _ _ h3) as h7; simpl in h7. unfold sorted_linds in h6, h7. rewrite h6, h7 in h5 at 1. simpl in h5. apply add_nin_inj_eq in h5. unfold famlS in h5. do 2 rewrite map_im_compat in h5. do 2 rewrite <- list_to_set_sort_list in h5. apply im_inj_inj in h5; auto. apply inj_lS. intro h8. rewrite <- list_to_set_in_iff in h8. apply nin_famlS in h8; auto. left; auto. intro h8. rewrite <- list_to_set_in_iff in h8. apply nin_famlS in h8; auto. left; auto. Qed. Lemma linds_decompose_remove : forall {T:Type} (pfdt:type_eq_dec T) (l:list T) (a:T) (pf:In a l), sorted_linds pfdt l = sort_list _ lt_list_nat_tso_dec ((linds_occ pfdt l a) :: (linds_remove pfdt l a)). intros T h1 l a h2. unfold sorted_linds. eapply list_to_set_injable_increasingR. apply lt_list_nat_sort_dec. apply weak_incrR_no_dup_impl. apply weak_incrR_sort_list. apply no_dup_sort_list. apply no_dup_linds. apply weak_incrR_no_dup_impl. rewrite lt_list_nat_tso_dec_eq. apply weak_incrR_sort_list. rewrite lt_list_nat_tso_dec_eq. apply no_dup_sort_list. constructor. apply nin_linds_remove; auto. unfold linds_remove. apply no_dup_remove. apply no_dup_linds. rewrite <- list_to_set_sort_list. rewrite lt_list_nat_tso_dec_eq. rewrite <- list_to_set_sort_list. apply Extensionality_Ensembles; red; split; red; intros x h3; rewrite <- list_to_set_in_iff in h3; rewrite <- list_to_set_in_iff. rewrite in_linds_iff in h3. destruct h3 as [x' [h3 h4]]. destruct h3 as [h3]; subst. unfold linds_remove. destruct (h1 x' a) as [|h5]; subst. left; auto. right. rewrite in_remove_iff. split. apply in_linds_occ_linds; auto. contradict h5. apply inj_linds_occ in h5; auto. destruct h3 as [|h3]; subst. apply in_linds_occ_linds; auto. unfold linds_remove in h3. rewrite in_remove_iff in h3. destruct h3; auto. Qed. Lemma partition_sorted_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), partition_faml (sorted_linds pfdt l) (seg_list (length l)). intros T h1 l. pose proof (partition_linds h1 l) as h2. red in h2; red. unfold sorted_linds. pose proof (faml_to_fam_sort_list _ sort_dec_lt_nat (linds h1 l)) as h3. simpl in h3. erewrite sort_list_functional in h3. rewrite h3; auto. reflexivity. Qed. Lemma linds_occ_eq_l_to_inds_aux : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (a:T) (a':U) (pf: In a l) (pf':In a' l'), list_to_set (linds pfdt l) = list_to_set (linds pfdu l') -> linds_occ pfdt l a = linds_occ pfdu l' a' -> l_to_inds_aux pfdt l a pf = l_to_inds_aux pfdu l' a' pf'. intros T U h1 h2 l. induction l as [|a l ih]; simpl. intros; contradiction. intros l'. destruct l' as [|a' l']; simpl. intros; contradiction. intros x x' h3 h4 hls h5. destruct (in_dec h1 a l) as [hi | hi]. destruct (in_dec h2 a' l') as [hi' | hi']. destruct h3 as [|h3]; subst. rewrite linds_occ_cons_eq in h5. simpl in h5. rewrite (l_to_inds_aux_cons_in _ _ _ hi). destruct h4 as [|h4]; subst. rewrite linds_occ_cons_eq in h5. simpl in h5. inversion h5. apply inj_lS in H0. rewrite (l_to_inds_aux_cons_in _ _ _ hi'). apply ih; auto. eapply list_to_set_linds_from_cons. apply hls. destruct (h2 x' a') as [| h7]; subst. rewrite linds_occ_cons_eq in h5. simpl in h5. inversion h5. apply inj_lS in H0. rewrite (l_to_inds_aux_cons_in _ _ _ h4). apply ih; auto. eapply list_to_set_linds_from_cons. apply hls. rewrite linds_occ_cons_neq in h5; auto. pose proof (in_eq 0%nat (lS (linds_occ h1 l x))) as h8. simpl in h5. rewrite h5 in h8. rewrite in_lS_iff in h8. destruct h8 as [? [? ?]]; discriminate. destruct h4 as [|h4]; subst. simpl in h5. rewrite linds_occ_cons_eq in h5. simpl in h5. rewrite (l_to_inds_aux_cons_in _ _ _ hi'). destruct (h1 x a) as [|he]; subst. rewrite linds_occ_cons_eq in h5. simpl in h5. inversion h5. apply inj_lS in H0. rewrite (l_to_inds_aux_cons_in _ _ _ hi). apply ih; auto. eapply list_to_set_linds_from_cons. apply hls. rewrite linds_occ_cons_neq in h5; auto. pose proof (in_eq 0 (lS (linds_occ h2 l' x'))) as h6. rewrite <- h5 in h6. rewrite in_lS_iff in h6. destruct h6 as [? [? ?]]; discriminate. destruct (h1 x a) as [|h6]; subst. rewrite (l_to_inds_aux_cons_in _ _ _ h3). rewrite linds_occ_cons_eq in h5. simpl in h5. destruct (h2 x' a') as [|ha]; subst. rewrite linds_occ_cons_eq in h5. simpl in h5. inversion h5. apply inj_lS in H0. rewrite (l_to_inds_aux_cons_in h2 _ _ h4). apply ih; auto. eapply list_to_set_linds_from_cons. apply hls. rewrite linds_occ_cons_neq in h5. pose proof (in_eq 0 (lS (linds_occ h1 l a))) as h6. rewrite h5 in h6. rewrite in_lS_iff in h6. destruct h6 as [? [? ?]]; discriminate. assumption. rewrite linds_occ_cons_neq in h5; auto. destruct (h2 x' a') as [|he]; subst. rewrite linds_occ_cons_eq in h5. simpl in h5. pose proof (in_eq 0 (lS (linds_occ h2 l' a'))) as h7. rewrite <- h5 in h7. rewrite in_lS_iff in h7. destruct h7 as [? [? ?]]; discriminate. rewrite linds_occ_cons_neq in h5; auto. apply inj_lS in h5. erewrite l_to_inds_aux_cons_in_in; auto. erewrite l_to_inds_aux_cons_in_in; auto. specialize (ih l' x x' h3 h4). hfirst ih. eapply list_to_set_linds_from_cons. apply hls. erewrite l_to_inds_aux_functional3. rewrite ih at 1. apply l_to_inds_aux_functional3. assumption. assumption. pose proof (in_linds_occ_linds h1 (a::l) x h3) as h6. rewrite list_to_set_in_iff in h6. rewrite hls in h6. erewrite list_to_set_sort_list in h6. pose proof (sorted_linds_cons_nin h2 _ _ hi') as h7. unfold sorted_linds in h7. rewrite h7 in h6 at 1. rewrite <- list_to_set_in_iff in h6. destruct h3 as [|h3]; subst. rewrite linds_occ_cons_eq in h6. simpl in h6. destruct h6 as [h6 | h6]. inversion h6. symmetry in H0. rewrite lS_eq_nil_iff in H0. pose proof (in_lind_linds_occ h1 _ _ hi) as h8. apply in_not_nil in h8. contradiction. apply nin_famlS in h6. contradiction. left; auto. destruct h4 as [|h4]; subst. destruct (h1 x a) as [|h8]; subst. rewrite linds_occ_cons_eq in h6. simpl in h6. destruct h6 as [h6 | h6]. inversion h6. symmetry in H0. rewrite lS_eq_nil_iff in H0. pose proof (in_lind_linds_occ h1 _ _ h3) as h8. apply in_not_nil in h8; contradiction. rewrite linds_occ_cons_eq in h5. rewrite linds_occ_cons_eq in h5. simpl in h5. inversion h5. apply inj_lS in H0. pose proof (in_lind_linds_occ h1 l a h3) as h8. rewrite H0 in h8. rewrite in_linds_occ_iff in h8. destruct h8 as [? [? ?]]; contradiction. rewrite linds_occ_cons_neq in h6; auto. destruct h6 as [h6 | h6]. pose proof (in_eq 0 nil) as h9. rewrite h6 in h9. apply nin0_lS in h9. contradiction. rewrite in_famlS_iff in h6. destruct h6 as [ln' [h6 h9]]. apply inj_lS in h9. subst. rewrite linds_occ_cons_eq in h5. rewrite linds_occ_cons_neq in h5; auto. simpl in h5. pose proof (in_eq 0 (lS (linds_occ h2 l' x'))) as h9. rewrite <- h5 in h9. apply nin0_lS in h9. contradiction. assert (h8:x' <> a'). intro; subst; contradiction. destruct (h1 x a) as [|h9]; subst. rewrite linds_occ_cons_eq in h6. simpl in h6. destruct h6 as [h6 | h6]. inversion h6. symmetry in H0. rewrite lS_eq_nil_iff in H0. pose proof (in_lind_linds_occ h1 _ _ h3) as h9. rewrite H0 in h9; contradiction. apply nin_famlS in h6. contradiction. left; auto. rewrite linds_occ_cons_neq in h6; auto. simpl in h6. destruct h6 as [h6 | h6]. pose proof (in_eq 0 nil) as h10. rewrite h6 in h10. apply nin0_lS in h10. contradiction. rewrite in_famlS_iff in h6. destruct h6 as [ln [h6 h10]]. apply inj_lS in h10; subst. rewrite linds_occ_cons_neq in h5; auto. rewrite linds_occ_cons_neq in h5; auto. apply inj_lS in h5. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h9); auto. rewrite (l_to_inds_aux_cons_nin_in _ _ _ _ h8); auto. pose proof hls as h10. erewrite list_to_set_sort_list in h10. symmetry in h10. erewrite list_to_set_sort_list in h10. pose proof (sorted_linds_cons_in h1 l a hi) as h11. simpl in h11. unfold sorted_linds in h11. rewrite h11 in h10 at 1. pose proof (sorted_linds_cons_nin h2 l' a' hi') as h12. unfold sorted_linds in h12. rewrite h12 in h10 at 1. simpl in h10. pose proof (Add_intro2 _ (list_to_set (famlS (sort_list lt_list_nat (sort_dec_tso_dec lt_list_nat lt_list_nat_sort_dec) (linds h2 l')))) (0 :: nil)) as h13. rewrite h10 in h13 at 1. inversion h13; subst. rewrite <- list_to_set_in_iff in H. apply nin_famlS in H. contradiction. left; auto. inversion H. rewrite lS_eq_nil_iff in H1. pose proof (in_lind_linds_occ h1 l a hi) as h14. rewrite H1 in h14. contradiction. destruct (in_dec h2 a' l') as [h6 | h6]. pose proof hls as hls'. erewrite list_to_set_sort_list in hls'. symmetry in hls'. erewrite list_to_set_sort_list in hls'. symmetry in hls'. pose proof (sorted_linds_cons_nin h1 _ _ hi) as h7. simpl in h7. pose proof (sorted_linds_cons_in h2 _ _ h6) as h8. simpl in h8. unfold sorted_linds in h7, h8. rewrite h7, h8 in hls' at 1. simpl in hls'. pose proof (Add_intro2 _ (list_to_set (famlS (sort_list lt_list_nat (sort_dec_tso_dec lt_list_nat lt_list_nat_sort_dec) (linds h1 l)))) (0 :: nil)) as h9. rewrite hls' in h9. inversion h9; subst. rewrite <- list_to_set_in_iff in H. apply nin_famlS in H. contradiction. left; auto. inversion H. rewrite lS_eq_nil_iff in H1. pose proof (in_lind_linds_occ h2 _ _ h6) as h10. rewrite H1 in h10. contradiction. pose proof hls as hls'. erewrite list_to_set_sort_list in hls'. symmetry in hls'. erewrite list_to_set_sort_list in hls'. symmetry in hls'. pose proof (sorted_linds_cons_nin h1 _ _ hi) as h7. simpl in h7. pose proof (sorted_linds_cons_nin h2 _ _ h6) as h8. simpl in h8. unfold sorted_linds in h7, h8. rewrite h7, h8 in hls' at 1. simpl in hls'. apply add_nin_inj_eq in hls'. unfold famlS in hls'. do 2 rewrite map_im_compat in hls'. do 2 rewrite <- list_to_set_sort_list in hls'. apply im_inj_inj in hls'. destruct h3 as [|h3]; subst. rewrite linds_occ_cons_eq in h5. destruct h4 as [|h4]; subst. rewrite linds_occ_cons_eq in h5; auto. simpl in h5. inversion h5. apply inj_lS in H0. rewrite (l_to_inds_aux_cons_nin _ _ _ _ hi); auto. rewrite (l_to_inds_aux_cons_nin _ _ _ _ h6); auto. assert (h9:x' <> a'). intro; subst; contradiction. rewrite (linds_occ_cons_neq h2 _ _ _ h9) in h5. simpl in h5. pose proof (in_eq 0 (lS (linds_occ h1 l x))) as h10. rewrite h5 in h10. rewrite in_lS_iff in h10. destruct h10 as [? [? ?]]; discriminate. destruct h4 as [|h4]; subst. assert (h9:x <> a). intro; subst; contradiction. rewrite linds_occ_cons_eq in h5. rewrite linds_occ_cons_neq in h5; auto. simpl in h5. pose proof (in_eq 0 (lS (linds_occ h2 l' x'))) as h10. rewrite <- h5 in h10. apply nin0_lS in h10; contradiction. assert (h9:x <> a). intro; subst; contradiction. assert (h10:x' <> a'). intro; subst; contradiction. rewrite (l_to_inds_aux_cons_nin_in h1 _ _ _ h9); auto. rewrite (l_to_inds_aux_cons_nin_in h2 _ _ _ h10); auto. f_equal. apply ih; auto. rewrite linds_occ_cons_neq in h5; auto. rewrite linds_occ_cons_neq in h5; auto. apply inj_lS in h5; auto. apply inj_lS. intro h9. rewrite <- list_to_set_in_iff in h9. apply nin_famlS in h9; auto. left; auto. intro h9. rewrite <- list_to_set_in_iff in h9. apply nin_famlS in h9; auto. left; auto. Grab Existential Variables. assumption. assumption. Qed. Lemma list_to_set_linds_eq_l_to_inds : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), list_to_set (linds pfdt l) = list_to_set (linds pfdu l') -> l_to_inds pfdt l = l_to_inds pfdu l'. intros T U h1 h2 l. induction l as [|a l ih]. intros l' h3. simpl in h3. symmetry in h3. apply empty_set_nil in h3. unfold l_to_inds; simpl. symmetry. rewrite linds_eq_nil_iff in h3; subst; auto. intros l'. destruct l' as [|a' l']. intro h3. simpl in h3. rewrite (list_to_set_sort_list _ lt_list_nat_sort_dec) in h3. destruct (in_dec h1 a l) as [h4 | h4]. pose proof (sorted_linds_cons_in h1 l a h4) as h5. simpl in h5. unfold sorted_linds in h5. rewrite h5 in h3. simpl in h3. apply add_not_empty in h3; contradiction. pose proof (sorted_linds_cons_nin h1 l a h4) as h5. unfold sorted_linds in h5. rewrite h5 in h3. simpl in h3. apply add_not_empty in h3; contradiction. intro h3. pose proof (list_to_set_linds_from_cons h1 h2 l l' a a' h3) as h4. pose proof (list_to_set_linds_occ_eq h1 h2 l l' a a' h3) as he. destruct (in_dec h1 a l) as [h5 | h5]. rewrite (l_to_inds_cons_in h1 _ _ h5). destruct (in_dec h2 a' l') as [h6 | h6]. rewrite (l_to_inds_cons_in h2 _ _ h6). pose proof (linds_occ_eq_l_to_inds_aux h1 h2 l l' a a' h5 h6 h4 he) as hl. rewrite hl. f_equal; auto. pose proof (in_lind_linds_occ h1 l a h5) as h7. rewrite he in h7. rewrite in_linds_occ_iff in h7. destruct h7 as [? [? ?]]. contradiction. destruct (in_dec h2 a' l') as [h6 | h6]. pose proof (in_lind_linds_occ h2 _ _ h6) as h7. rewrite <- he in h7. rewrite in_linds_occ_iff in h7. destruct h7 as [? [? ?]]. contradiction. rewrite l_to_inds_cons_nin; auto. rewrite l_to_inds_cons_nin; auto. do 2 f_equal; auto. Qed. Lemma list_to_set_eq_linds : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), list_to_set (linds pfdt l) = list_to_set (linds pfdu l') -> linds pfdt l = linds pfdu l'. intros T U h1 h2 l l' h3. pose proof (list_to_set_linds_eq_l_to_inds h1 h2 _ _ h3) as h4. rewrite <- (@linds_l_to_inds T). rewrite <- (@linds_l_to_inds U). rewrite h4; auto. Qed. Lemma sorted_linds_eq_iff : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), sorted_linds pfdt l = sorted_linds pfdu l' <-> linds pfdt l = linds pfdu l'. intros T U h1 h2 l l'. unfold sorted_linds. pose proof (list_to_sets_eq_iff lt_list_nat lt_list_nat_sort_dec (linds h1 l) (linds h2 l') (no_dup_linds _ _) (no_dup_linds _ _)) as h3. simpl in h3. rewrite <- h3. split; intro h4. apply list_to_set_eq_linds in h4; auto. congruence. Qed. Lemma linds_eq_lecs : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (m:nat) (pfm:m < length l) (pf:linds pfdt l = linds pfdu l'), let pfl := linds_eq_length pfdt pfdu l l' pf in let pfm' := lt_eq_trans _ _ _ pfm pfl in lec pfdt pfm = lec pfdu pfm'. intros T U hdt hdu l l' m hm h1 he hm'. pose proof (in_lec hdt l m hm) as h2. rewrite h1 in h2. pose proof (in_lec hdu l' m hm') as h3. pose proof (in_linds_list_to_set_injable hdu l') as h4. red in h4. apply h4; auto. eapply partition_faml_eq. apply (partition_linds hdu). apply h2. apply h3. apply rep_in_lec. apply rep_in_lec. Qed. Lemma linds_eq_counts : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (i:nat) (pfi: i < length l) (pf:linds pfdt l = linds pfdu l'), let pfi' := lt_eq_trans _ _ _ pfi (linds_eq_length pfdt pfdu _ _ pf) in count_occ pfdt l (nth_lt l i pfi) = count_occ pfdu l' (nth_lt l' i pfi'). intros T U hdt hdu l l' i h1 h2 h3. erewrite <- length_linds_occ. erewrite <- length_linds_occ. do 2 rewrite linds_occ_nth_lt. pose proof (linds_eq_lecs hdt hdu _ _ i h1 h2) as h4. simpl in h4. rewrite h4. repeat f_equal. Qed. Lemma list_to_set_linds_eq_l_to_inds_aux : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (i:nat) (pflt:i < length l) (pflt':i < length l'), list_to_set (linds pfdt l) = list_to_set (linds pfdu l') -> l_to_inds_aux pfdt l (nth_lt l i pflt) (in_nth_lt _ _ _) = l_to_inds_aux pfdu l' (nth_lt l' i pflt') (in_nth_lt _ _ _). intros T U hdt hdu l. induction l as [|a l ih]. simpl; intros; omega. intro l'. destruct l' as [|a' l']. intros i h1 h2 h3. simpl in h2. omega. intros i h1 h2 h3. simpl in h2. pose proof h3 as h3'. apply (list_to_set_eq_linds hdt hdu (a::l) (a'::l')) in h3. rewrite <- sorted_linds_eq_iff in h3. simpl in h1. destruct (in_dec hdt a l) as [h4 | h4]. destruct (in_dec hdu a' l') as [h5 | h5]. destruct (hdt a (nth_lt (a ::l) i h1)) as [h6 | h6], (hdu a' (nth_lt (a'::l') i h2)) as [h7 | h7]. gen0. generalize (in_nth_lt (a'::l')i h2). rewrite h6, h7. intros h8 h9. simpl. destruct i as [|i]. rewrite (l_to_inds_aux_cons_in _ _ _ h4); auto. rewrite (l_to_inds_aux_cons_in _ _ _ h5); auto. gen0. generalize h5. rewrite h6, h7. intros h10 h11. simpl. pose proof (sorted_linds_cons_in hdt l a h4) as h12. simpl in h12. pose proof (sorted_linds_cons_in hdu l' a' h5) as h13. simpl in h13. unfold sorted_linds in h12, h13. unfold sorted_linds in h3. rewrite h12, h13 in h3. inversion h3. apply inj_famlS in H1. apply inj_lS in H0. rewrite H0 in H1. eapply incrR_remove_inj in H1. pose proof (sorted_linds_eq_iff hdt hdu l l') as h14. unfold sorted_linds in h14. rewrite h14 in H1. apply linds_occ_eq_l_to_inds_aux. rewrite H1; auto. assumption. apply lt_list_nat_sort_dec. rewrite in_sort_list_iff. rewrite <- H0. apply in_linds_occ_linds. assumption. rewrite in_sort_list_iff. apply in_linds_occ_linds. assumption. apply incr_sorted_linds. apply incr_sorted_linds. simpl in h8, h9. pose proof (@nth_indep _ l i a (nth i l a)) as h10. pose proof h1 as h1'. pose proof h2 as h2'. apply lt_S_n in h1'. apply lt_S_n in h2'. specialize (h10 h1'). gen0. rewrite <- h10. intro. pose proof (@nth_indep _ l' i a' (nth i l' a')) as h10'. specialize (h10' h2'). symmetry. gen0. rewrite <- h10'. intro h11. rewrite (l_to_inds_aux_cons_in _ _ _ (nth_In l' a' h2')). rewrite (l_to_inds_aux_cons_in _ _ _ (nth_In l a h1')). symmetry. rewrite sorted_linds_cons_in in h3; auto. rewrite sorted_linds_cons_in in h3; auto. inversion h3. apply inj_famlS in H1. apply inj_lS in H0. rewrite H0 in H1. eapply incrR_remove_inj in H1. rewrite sorted_linds_eq_iff in H1. specialize (ih l' i h1' h2'). hfirst ih. rewrite H1; auto. specialize (ih hf). revert ih. intro. gen0. erewrite nth_indep. intro h12. symmetry. gen0. erewrite nth_indep. intro h13. symmetry. erewrite f_equal_dep. rewrite ih at 1. apply f_equal_dep. rewrite nth_lt_compat. reflexivity. rewrite nth_lt_compat. reflexivity. assumption. assumption. apply lt_list_nat_sort_dec. rewrite <- H0. unfold sorted_linds. rewrite in_sort_list_iff. apply in_linds_occ_linds. assumption. unfold sorted_linds. rewrite in_sort_list_iff. apply in_linds_occ_linds. assumption. apply incr_sorted_linds. apply incr_sorted_linds. destruct i as [|i]. simpl in h7. contradict h7; auto. simpl in h6, h7. simpl. gen0. generalize (in_nth_lt (a' :: l') (S i) h2). simpl. rewrite <- h6. intros h8 h9. rewrite (l_to_inds_aux_cons_in _ _ _ h4). apply neq_sym in h7. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h7). gen0. rewrite h6. intro h10. rewrite sorted_linds_cons_in in h3; auto. rewrite sorted_linds_cons_in in h3; auto. inversion h3. apply inj_famlS in H1. apply inj_lS in H0. rewrite H0 in H1. eapply incrR_remove_inj in H1. rewrite sorted_linds_eq_iff in H1. pose proof h1 as h1'. pose proof h2 as h2'. apply lt_S_n in h1'. apply lt_S_n in h2'. specialize (ih l' i h1' h2'). hfirst ih. rewrite H1; auto. specialize (ih hf). revert ih. intro. gen0. erewrite nth_indep. intro h12. symmetry. gen0. erewrite nth_indep. intro h13. symmetry. erewrite f_equal_dep. rewrite ih at 1. apply f_equal_dep. rewrite nth_lt_compat. reflexivity. rewrite nth_lt_compat. reflexivity. assumption. assumption. apply lt_list_nat_sort_dec. rewrite <- H0. unfold sorted_linds. rewrite in_sort_list_iff. apply in_linds_occ_linds. assumption. unfold sorted_linds. rewrite in_sort_list_iff. apply in_linds_occ_linds. assumption. apply incr_sorted_linds. apply incr_sorted_linds. assumption. destruct i as [|i]. simpl in h6. contradict h6; auto. simpl in h6, h7. simpl. symmetry. gen0. generalize (in_nth_lt (a :: l) (S i) h1). simpl. rewrite <- h7. intros h8 h9. rewrite (l_to_inds_aux_cons_in _ _ _ h5). apply neq_sym in h6. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h6). gen0. rewrite h7. intro h10. rewrite sorted_linds_cons_in in h3; auto. rewrite sorted_linds_cons_in in h3; auto. inversion h3. apply inj_famlS in H1. apply inj_lS in H0. rewrite H0 in H1. eapply incrR_remove_inj in H1. rewrite sorted_linds_eq_iff in H1. pose proof h1 as h1'. pose proof h2 as h2'. apply lt_S_n in h1'. apply lt_S_n in h2'. specialize (ih l' i h1' h2'). hfirst ih. rewrite H1; auto. specialize (ih hf). revert ih. intro. gen0. erewrite nth_indep. intro h12. symmetry. gen0. erewrite nth_indep. intro h13. erewrite f_equal_dep. rewrite ih at 1. apply f_equal_dep. rewrite nth_lt_compat. reflexivity. rewrite nth_lt_compat. reflexivity. assumption. assumption. apply lt_list_nat_sort_dec. rewrite <- H0. unfold sorted_linds. rewrite in_sort_list_iff. apply in_linds_occ_linds. assumption. unfold sorted_linds. rewrite in_sort_list_iff. apply in_linds_occ_linds. assumption. apply incr_sorted_linds. apply incr_sorted_linds. assumption. destruct i as [|i]. simpl in h6. contradict h6; auto. simpl. apply neq_sym in h6. apply neq_sym in h7. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h6); auto. rewrite (l_to_inds_aux_cons_in_in _ _ _ _ h7); auto. simpl. pose proof h1 as h1'. pose proof h2 as h2'. apply lt_S_n in h1'. apply lt_S_n in h2'. specialize (ih l' i h1' h2'). rewrite sorted_linds_cons_in in h3; auto. rewrite sorted_linds_cons_in in h3; auto. inversion h3. apply inj_lS in H0. rewrite H0 in H1. apply inj_famlS in H1. eapply incrR_remove_inj in H1. rewrite sorted_linds_eq_iff in H1. hfirst ih. rewrite H1; auto. specialize (ih hf). revert ih. Gen0. rewrite nth_lt_compat. intro h8. intro h9. symmetry in h9. revert h9. Gen0. rewrite nth_lt_compat. intros h10 h11. erewrite f_equal_dep. rewrite <- h11 at 1. apply f_equal_dep. apply nth_indep. assumption. apply nth_indep. assumption. apply lt_list_nat_sort_dec. unfold sorted_linds. rewrite in_sort_list_iff. rewrite <- H0. apply in_linds_occ_linds. assumption. unfold sorted_linds. rewrite in_sort_list_iff. apply in_linds_occ_linds. assumption. apply incr_sorted_linds. apply incr_sorted_linds. rewrite sorted_linds_cons_in in h3. rewrite sorted_linds_cons_nin in h3. inversion h3. rewrite lS_eq_nil_iff in H0. pose proof (in_lind_linds_occ hdt _ _ h4) as h6. rewrite H0 in h6. contradiction. assumption. assumption. destruct (in_dec hdu a' l') as [h5 | h5]. rewrite sorted_linds_cons_nin in h3; auto. rewrite sorted_linds_cons_in in h3; auto. inversion h3. apply inj_famlS in H1. symmetry in H0. rewrite lS_eq_nil_iff in H0. pose proof (in_lind_linds_occ hdu l' a' h5) as h6. rewrite H0 in h6. contradiction. simpl. destruct i as [|i]. rewrite l_to_inds_aux_cons_nin; auto. rewrite l_to_inds_aux_cons_nin; auto. assert (h6:a <> nth i l a). intro h6. rewrite h6 in h4. contradict h4. apply nth_In; auto with arith. assert (h7:a' <> nth i l' a'). intro h7. rewrite h7 in h5. contradict h5. apply nth_In; auto with arith. apply neq_sym in h6. apply neq_sym in h7. rewrite (l_to_inds_aux_cons_nin_in _ _ _ _ h6 h4). rewrite (l_to_inds_aux_cons_nin_in _ _ _ _ h7 h5). f_equal. rewrite sorted_linds_cons_nin in h3; auto. rewrite sorted_linds_cons_nin in h3; auto. inversion h3. apply inj_famlS in H0. rewrite sorted_linds_eq_iff in H0. pose proof h1 as h1'. pose proof h2 as h2'. apply lt_S_n in h1'. apply lt_S_n in h2'. specialize (ih l' i h1' h2'). hfirst ih. rewrite H0; auto. specialize (ih hf). revert ih. Gen0. rewrite nth_lt_compat. intros h8 h9. symmetry in h9. revert h9. Gen0. rewrite nth_lt_compat. intros h9 h10. erewrite f_equal_dep. rewrite <- h10 at 1. apply f_equal_dep. apply nth_indep. assumption. apply nth_indep. assumption. Qed. Lemma linds_eq_l_to_inds_aux : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (i:nat) (pflt:i < length l) (pflt':i < length l'), linds pfdt l = linds pfdu l' -> l_to_inds_aux pfdt l (nth_lt l i pflt) (in_nth_lt _ _ _) = l_to_inds_aux pfdu l' (nth_lt l' i pflt') (in_nth_lt _ _ _). intros; apply list_to_set_linds_eq_l_to_inds_aux. rewrite H; auto. Qed. Lemma list_to_set_linds_map_seg_eq_l_to_inds_aux : forall {U V:Type} (hdu:type_eq_dec U) (hdv:type_eq_dec V) (f:nat->U) (g:nat->V) (i n:nat) (pflt:i < n), list_to_set (linds hdu (map_seg (fun_to_seg f n))) = list_to_set (linds hdv (map_seg (fun_to_seg g n))) -> forall (pfi:In (fun_to_seg f n i pflt) (map_seg (fun_to_seg f n))) (pfi':In (fun_to_seg g n i pflt) (map_seg (fun_to_seg g n))), l_to_inds_aux hdu (map_seg (fun_to_seg f n)) (f i) pfi = l_to_inds_aux hdv (map_seg (fun_to_seg g n)) (g i) pfi'. intros. pose proof (list_to_set_linds_eq_l_to_inds_aux hdu hdv (map_seg (fun_to_seg f n)) (map_seg (fun_to_seg g n)) i) as h1. hfirst h1. rewrite length_map_seg; auto. specialize (h1 hf). hfirst h1. rewrite length_map_seg; auto. specialize (h1 hf0 H). erewrite f_equal_dep. rewrite h1 at 1. apply f_equal_dep. erewrite nth_lt_functional3. rewrite nth_lt_map_seg. reflexivity. erewrite nth_lt_functional3. rewrite nth_lt_map_seg. reflexivity. Grab Existential Variables. assumption. assumption. Qed. Lemma linds_map_seg_eq_l_to_inds_aux : forall {U V:Type} (hdu:type_eq_dec U) (hdv:type_eq_dec V) (f:nat->U) (g:nat->V) (i n:nat) (pflt:i < n), linds hdu (map_seg (fun_to_seg f n)) = linds hdv (map_seg (fun_to_seg g n)) -> forall (pfi:In (fun_to_seg f n i pflt) (map_seg (fun_to_seg f n))) (pfi':In (fun_to_seg g n i pflt) (map_seg (fun_to_seg g n))), l_to_inds_aux hdu (map_seg (fun_to_seg f n)) (f i) pfi = l_to_inds_aux hdv (map_seg (fun_to_seg g n)) (g i) pfi'. intros; apply list_to_set_linds_map_seg_eq_l_to_inds_aux; auto. rewrite H; auto. Qed. Lemma linds_eq_lind_occ : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (i occ:nat) (pfi: i < length l) (pfi':i < length l'), linds pfdt l = linds pfdu l' -> occ < count_occ pfdt l (nth_lt _ _ pfi) -> lind_occ pfdt l occ (nth_lt _ _ pfi) (in_nth_lt _ _ _) = lind_occ pfdu l' occ (nth_lt _ _ pfi') (in_nth_lt _ _ _). intros T U h1 h2 l l' i occ h3 h4 h5 h6. pose proof (linds_eq_l_to_inds_aux h1 h2 _ _ i h3 h4 h5) as h11. pose proof (lind_occ_l_to_inds h1 l occ (nth_lt l i h3) (in_nth_lt _ _ _) h6) as h8. pose proof h6 as h6'. pose proof (linds_eq_counts h1 h2 l l' i h3 h5) as h9. simpl in h9. rewrite h9 in h6'. erewrite nth_lt_functional3 in h6'. pose proof (lind_occ_l_to_inds h2 l' occ (nth_lt l' i h4) (in_nth_lt _ _ _) h6') as h10. rewrite h8, h10. gen0. rewrite h11. intro h12. apply lind_occ_functional. apply list_to_set_linds_eq_l_to_inds. rewrite h5; auto. reflexivity. apply list_to_set_linds_eq_l_to_inds_aux. reflexivity. Qed. Lemma linds_map_seg_eq_lind_occ : forall {U V:Type} (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (f:nat->U) (g:nat->V) (i occ n:nat) (pflt:i < n), linds pfdu (map_seg (fun_to_seg f n)) = linds pfdv (map_seg (fun_to_seg g n)) -> occ < count_occ pfdu (map_seg (fun_to_seg f n)) (f i) -> forall (pfi:In (fun_to_seg f n i pflt) (map_seg (fun_to_seg f n))) (pfi':In (fun_to_seg g n i pflt) (map_seg (fun_to_seg g n))), lind_occ pfdu (map_seg (fun_to_seg f n)) occ (f i) pfi = lind_occ pfdv (map_seg (fun_to_seg g n)) occ (g i) pfi'. intros U V hdu hdv f g i occ n h1 h2 h3 h4 h5. pose proof (linds_eq_lind_occ hdu hdv (map_seg (fun_to_seg f n)) (map_seg (fun_to_seg g n)) i occ) as h6. hfirst h6. rewrite length_map_seg; auto. specialize (h6 hf). hfirst h6. rewrite length_map_seg; auto. specialize (h6 hf0 h2). hfirst h6. erewrite nth_lt_functional3. rewrite nth_lt_map_seg. rewrite fun_to_seg_compat; auto. specialize (h6 hf1). revert h6. Gen0. generalize (in_nth_lt (map_seg (fun_to_seg g n)) i hf0). erewrite nth_lt_functional3. rewrite nth_lt_map_seg. erewrite (nth_lt_functional3 (map_seg (fun_to_seg f n))). rewrite nth_lt_map_seg. do 2 rewrite fun_to_seg_compat. intros h6 h7 h8. pose proof (proof_irrelevance _ h7 h4). pose proof (proof_irrelevance _ h5 h6). subst; auto. Grab Existential Variables. assumption. assumption. assumption. Qed. Lemma linds_eq_linds_occ : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (i:nat) (pfi: i < length l) (pfi':i < length l'), linds pfdt l = linds pfdu l' -> linds_occ pfdt l (nth_lt _ _ pfi) = linds_occ pfdu l' (nth_lt _ _ pfi'). intros T U h1 h2 l l' i h3 h4 h5. pose proof (@list_to_set_linds_eq_l_to_inds). pose proof (@linds_occ_l_to_inds). erewrite linds_occ_l_to_inds. symmetry. erewrite linds_occ_l_to_inds. erewrite list_to_set_linds_eq_l_to_inds. repeat f_equal. apply linds_eq_l_to_inds_aux; auto. rewrite h5; auto. Qed. Lemma linds_map_seg_eq_linds_occ : forall {U V:Type} (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (f:nat->U) (g:nat->V) (i n:nat) (pf:i < n), linds pfdu (map_seg (fun_to_seg f n)) = linds pfdv (map_seg (fun_to_seg g n)) -> linds_occ pfdu (map_seg (fun_to_seg f n)) (f i) = linds_occ pfdv (map_seg (fun_to_seg g n)) (g i). intros U V hdu hdv f g i n h1 h2. pose proof (linds_eq_linds_occ hdu hdv (map_seg (fun_to_seg f n)) (map_seg (fun_to_seg g n)) i) as h3. hfirst h3. rewrite length_map_seg; auto. specialize (h3 hf). hfirst h3. rewrite length_map_seg; auto. specialize (h3 hf0 h2). erewrite nth_lt_functional3 in h3. rewrite nth_lt_map_seg in h3. erewrite nth_lt_functional3 in h3. rewrite nth_lt_map_seg in h3. do 2 rewrite fun_to_seg_compat in h3; auto. Grab Existential Variables. assumption. assumption. Qed. Lemma fun_linds_eq_nat_linds_occ : forall {U V:Type} (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (f:nat->U) (g:nat->V), fun_linds_eq_nat pfdu pfdv f g -> forall (i n:nat), i < n -> linds_occ pfdu (map_seg (fun_to_seg f n)) (f i) = linds_occ pfdv (map_seg (fun_to_seg g n)) (g i). intros U V hdu hdv f g h1 i n h2. red in h1. specialize (h1 n). apply linds_map_seg_eq_linds_occ; auto. Qed. Lemma fun_linds_eq_nat_lind_occ : forall {U V:Type} (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (f:nat->U) (g:nat->V), fun_linds_eq_nat pfdu pfdv f g -> forall (i n occ:nat) (pflt:i < n), occ < count_occ pfdu (map_seg (fun_to_seg f n)) (f i) -> forall (pfi:In (fun_to_seg f n i pflt) (map_seg (fun_to_seg f n))) (pfi':In (fun_to_seg g n i pflt) (map_seg (fun_to_seg g n))), lind_occ pfdu (map_seg (fun_to_seg f n)) occ (f i) pfi = lind_occ pfdv (map_seg (fun_to_seg g n)) occ (g i) pfi'. intros U V hdu hdv f g h1 i n occ h2 h3 h4 h5. red in h1. specialize (h1 n). apply linds_map_seg_eq_lind_occ; auto. Qed. Lemma linds_nth_lt_compat : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l' :list U) (m n:nat) (pf:linds pfdt l = linds pfdu l') (pfm:m < length l) (pfn:n < length l), let pfl := linds_eq_length pfdt pfdu l l' pf in let pfm' := lt_eq_trans _ _ _ pfm pfl in let pfn' := lt_eq_trans _ _ _ pfn pfl in nth_lt l m pfm = nth_lt l n pfn -> nth_lt l' m pfm' = nth_lt l' n pfn'. intros T U hdt hdu l l' m n h1 hm hn ha hm' hn' h2. rewrite <- (lecs_eq_iff hdt) in h2. rewrite <- (lecs_eq_iff hdu). pose proof (in_lec hdt l m hm) as h3. pose proof (in_lec hdt l n hn) as h4. pose proof (in_lec hdu l' m hm') as h3'. pose proof (in_lec hdu l' n hn') as h4'. rewrite <- h1 in h3'. rewrite <- h1 in h4'. rewrite in_linds_iff in h3', h4'. destruct h3' as [m' h5], h4' as [n' h6]. red in h5, h6. destruct h5 as [h5 h7], h6 as [h6 h8]. destruct h5 as [h5 h9], h6 as [h6 h10]. erewrite f_equal_dep; auto. rewrite <- (linds_eq_lecs hdt hdu l l') at 1; auto. symmetry. erewrite f_equal_dep; auto. rewrite <- (linds_eq_lecs hdt hdu l l') at 1. erewrite f_equal_dep. symmetry. erewrite f_equal_dep. apply h2. reflexivity. reflexivity. Grab Existential Variables. assumption. assumption. assumption. assumption. Qed. Lemma linds_nth_lt_compat_iff : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l' :list U) (m n:nat) (pf:linds pfdt l = linds pfdu l') (pfm:m < length l) (pfn:n < length l), let pfl := linds_eq_length pfdt pfdu l l' pf in let pfm' := lt_eq_trans _ _ _ pfm pfl in let pfn' := lt_eq_trans _ _ _ pfn pfl in nth_lt l m pfm = nth_lt l n pfn <-> nth_lt l' m pfm' = nth_lt l' n pfn'. intros T U hdt hdu l l' m n h1 hm hn hl hm' hn'. split. apply linds_nth_lt_compat; auto. intro h2. erewrite nth_lt_functional3. rewrite (linds_nth_lt_compat hdu hdt l' l m n (eq_sym h1) hm' hn' h2) at 1. apply nth_lt_functional3. Qed. Lemma linds_eq_iff : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l' :list U), linds pfdt l = linds pfdu l' <-> l_to_inds pfdt l = l_to_inds pfdu l'. intros T U hdt hdu l l'; rewrite <- sorted_linds_eq_iff. unfold sorted_linds. pose proof (list_to_sets_eq_iff _ lt_list_nat_sort_dec (linds hdt l) (linds hdu l') (no_dup_linds _ _) (no_dup_linds _ _)) as h1. rewrite <- h1. split; intro h2. apply list_to_set_linds_eq_l_to_inds in h2; auto. erewrite <- linds_l_to_inds. rewrite h2. symmetry. erewrite <- linds_l_to_inds. reflexivity. Qed. Lemma linds_eq_iff' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U) (pfe:length l = length l'), linds pfdt l = linds pfdu l' <-> (forall (m n:nat) (pfm:m < length l) (pfn:n < length l), let pfm' := lt_eq_trans _ _ _ pfm pfe in let pfn' := lt_eq_trans _ _ _ pfn pfe in nth_lt l m pfm = nth_lt l n pfn <-> nth_lt l' m pfm' = nth_lt l' n pfn'). intros T U hdt hdu l l' he; split; intro h1. intros m n h2 h3 h4 h5. split; intro h6. erewrite nth_lt_functional3. erewrite (nth_lt_functional3 l' n). eapply linds_nth_lt_compat. apply h6. erewrite nth_lt_functional3. erewrite (nth_lt_functional3 l n). eapply linds_nth_lt_compat. apply h6. rewrite linds_eq_iff. eapply nth_lts_eq_l_to_inds. apply h1. Grab Existential Variables. symmetry. apply h1. apply h1. Qed. Lemma linds_eq_firstn : forall {T U:Type} {pfdt:type_eq_dec T} {pfdu:type_eq_dec U} (l:list T) (l':list U) (m:nat), linds pfdt l = linds pfdu l' -> linds pfdt (firstn m l) = linds pfdu (firstn m l'). intros T U hdt hdu l l' m h1. pose proof (linds_eq_length _ _ _ _ h1) as he. rewrite linds_eq_iff'. rewrite linds_eq_iff' in h1. intros r s h2 h3 h4 h5. do 4 rewrite nth_lt_firstn_eq. erewrite nth_lt_functional3. erewrite (nth_lt_functional3 l s). erewrite (nth_lt_functional3 l' r). erewrite (nth_lt_functional3 l' s). apply h1. Grab Existential Variables. eapply lt_le_trans. apply h3. apply length_firstn_le. eapply lt_le_trans. apply h2. apply length_firstn_le. apply length_eq_firstn; auto. assumption. Qed. Lemma two_lists_to_fun_in_dep_compat : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (pfl:linds pfdt la = linds pfdu lb) (i:nat) (pfi:i < length la), let pfe := linds_eq_length pfdt pfdu la lb pfl in let pfi' := lt_eq_trans _ _ _ pfi pfe in two_lists_to_fun_in_dep pfdt pfdu _ _ pfl _ (in_nth_lt _ _ pfi) = nth_lt _ _ pfi'. simpl; unfold two_lists_to_fun_in_dep; intros T U h1 h2 la lb h3 i h4. pose proof (in_lind_linds_occ h1 _ _ (in_nth_lt la i h4)) as h6. pose proof (in_linds_occ_linds h1 _ _ (in_nth_lt la i h4)) as h7. rewrite h3 in h7. rewrite in_linds_iff in h7. destruct h7 as [y h8]. red in h8. destruct h8 as [h8 h9]. destruct h8 as [h8 h10]. rewrite h10 in h6. rewrite in_linds_occ_iff in h6. destruct h6 as [m [h6 h7]]. destruct h7 as [h7 h11]. gen0. rewrite <- h11. intro h12. erewrite nth_lt_functional3. rewrite lind_occ_compat. apply h9. split. apply in_nth_lt. pose proof (in_linds_occ_linds h1 la (nth_lt la i h4) (in_nth_lt _ _ _)) as h13. rewrite h3 in h13. rewrite in_linds_iff in h13. destruct h13 as [y' [h13 h14]]. destruct h13 as [h13 h15]. pose proof h15 as h15'. rewrite h10 in h15'. apply inj_linds_occ in h15'; auto; subst. do 2 rewrite linds_occ_nth_lt. apply linds_eq_lecs; auto. Qed. Lemma two_lists_to_sig_fun_compat : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (la:list T) (lb:list U) (pfl:linds pfdt la = linds pfdu lb) (i:nat) (pfi:i < length la), let pfe := linds_eq_length pfdt pfdu la lb pfl in let pfi' := lt_eq_trans _ _ _ pfi pfe in two_lists_to_sig_fun pfdt pfdu _ _ pfl (exist _ _ (iff1 (list_to_set_in_iff _ _) (in_nth_lt _ _ pfi))) = exist _ _ (iff1 (list_to_set_in_iff _ _) (in_nth_lt _ _ pfi')). simpl; unfold two_lists_to_sig_fun; intros T U h1 h2 la lb h3 i h4. apply proj1_sig_injective; simpl. pose proof (in_lind_linds_occ h1 _ _ (in_nth_lt la i h4)) as h6. pose proof (in_linds_occ_linds h1 _ _ (in_nth_lt la i h4)) as h7. rewrite h3 in h7. rewrite in_linds_iff in h7. destruct h7 as [y h8]. red in h8. destruct h8 as [h8 h9]. destruct h8 as [h8 h10]. rewrite h10 in h6. rewrite in_linds_occ_iff in h6. destruct h6 as [m [h6 h7]]. destruct h7 as [h7 h11]. gen0. erewrite lind_functional3 in h11. rewrite <- h11. intro h12. erewrite nth_lt_functional3. rewrite lind_occ_compat. apply h9. split. apply in_nth_lt. pose proof (in_linds_occ_linds h1 la (nth_lt la i h4) (in_nth_lt _ _ _)) as h13. rewrite h3 in h13. rewrite in_linds_iff in h13. destruct h13 as [y' [h13 h14]]. destruct h13 as [h13 h15]. pose proof h15 as h15'. rewrite h10 in h15'. apply inj_linds_occ in h15'; auto; subst. do 2 rewrite linds_occ_nth_lt. apply linds_eq_lecs; auto. Qed. Lemma faml_to_fam_list_minus_sort_family : forall (A B:Family nat) (pffa:Finite A) (pffma:finite_mem A) (pffb:Finite B) (pffmb:finite_mem B), faml_to_fam (list_minus list_nat_eq_dec (sort_family A pffa pffma) (sort_family B pffb pffmb)) = Setminus A B. intros A B h1 h2 h3 h4. pose proof (increasing_mem_sort_family A h1 h2) as hia. pose proof (increasing_mem_sort_family B h3 h4) as hib. apply Extensionality_Ensembles; red; split; red; intros C h5. pose proof (in_faml_to_fam_impl _ _ h5) as h6. destruct h6 as [l h6]; subst. destruct h6 as [h6 h6']. apply (in_list_minus_impl list_nat_eq_dec) in h6. destruct h6 as [h6 h10]. subst. pose proof (in_sort_family_impl' _ _ _ _ h6) as h7. destruct h7 as [S [h7 h8]]; subst. rewrite sort_set_nat_compat. constructor; auto. contradict h10. eapply impl_in_sort_family; auto. destruct h5 as [h5 h6]. pose proof (h2 _ h5) as h7. pose proof (sort_set_nat_compat C h7) as h8. rewrite <- h8. eapply in_faml_impl. apply impl_in_list_minus. eapply impl_in_sort_family; auto. contradict h6. apply in_faml_impl in h6. rewrite sort_family_compat' in h6. rewrite sort_set_nat_compat in h6. assumption. Qed. Lemma list_minus_sort_family : forall (A B:Family nat) (pffa:Finite A) (pffma:finite_mem A) (pffb:Finite B) (pffmb:finite_mem B) (pfab:Finite (Setminus A B)) (pffmab:finite_mem (Setminus A B)), list_minus list_nat_eq_dec (sort_family A pffa pffma) (sort_family B pffb pffmb) = sort_family (Setminus A B) pfab pffmab. intros A B h1 h1' h2 h2' h3 h3'. pose proof (faml_to_fam_list_minus_sort_family A B h1 h1' h2 h2') as he. gterml. pose proof (sort_family_functional _ _ (finite_faml_to_fam x) h3 (finite_mem_faml x) h3' he) as h4. rewrite <- h4. eapply incr_impl_is_sorted_family. unfold x. apply incr_faml_list_minus. apply incr_sort_family. unfold x. apply increasing_mem_list_minus. apply increasing_mem_sort_family. Qed. Lemma faml_to_fam_remove_can_faml_nat : forall (A:Family nat) (B:Ensemble nat) (pffa:Finite A) (pffma:finite_mem A) (pffb:Finite B) , faml_to_fam (remove list_nat_eq_dec (sort_set_nat B pffb) (sort_family A pffa pffma)) = Subtract A B. intros A B h1 h2 h3. pose proof (faml_to_fam_list_minus_sort_family A (Singleton B) h1 h2 (Singleton_is_finite _ _)) as h4. hfirst h4. do 2 red; intros X h5. destruct h5; auto. specialize (h4 hf). pose proof (sort_family_sing B h3 (Singleton_is_finite (Ensemble nat) B) hf) as h5. unfold Subtract. rewrite <- h4. f_equal. rewrite h5. rewrite list_minus_sing. reflexivity. Qed. Lemma remove_sort_family : forall (A:Family nat) (B:Ensemble nat) (pffa:Finite A) (pffma:finite_mem A) (pffb:Finite B) (pfab:Finite (Subtract A B)) (pffmab:finite_mem (Subtract A B)), remove list_nat_eq_dec (sort_set_nat B pffb) (sort_family A pffa pffma) = sort_family (Subtract A B) pfab pffmab. intros A B h1 h1' h2 h2' h3. pose proof (faml_to_fam_remove_can_faml_nat A B h1 h1' h2) as he. gterml. pose proof (sort_family_functional _ _ (finite_faml_to_fam x) h2' (finite_mem_faml x) h3 he) as h4. rewrite <- h4. eapply incr_impl_is_sorted_family. unfold x. apply increasing_faml_remove. apply incr_sort_family. unfold x. apply increasing_mem_remove. apply increasing_mem_sort_family. Qed. End FamlNat. Lemma card_in_l : forall {T:Type} (l:list T), NoDup l -> FiniteT_card _ (list_to_set_finitet l) = length l. intros T l h0. induction l as [|a l h1]. simpl. rewrite <- FiniteT_card_False. pose (fun x:{_:T | False} => False_rect False (proj2_sig x)) as f. assert (h1:invertible f). apply bijective_impl_invertible. red; split; red; intros; contradiction. pose proof (FiniteT_card_bijection _ _ (list_to_set_finitet nil) f h1) as h2. simpl in h2. rewrite <- h2. pose proof (proof_irrelevance _ (bij_finite {_ : T | False} False f (list_to_set_finitet nil) h1) empty_finite) as h3. rewrite h3. reflexivity. pose proof (no_dup_cons_nin _ _ h0) as h0'. simpl. pose proof (FiniteT_card_option _ (list_to_set_finitet l)) as h3. pose (fun x:(option {x : T | In x l}) => let f := (fun x => a = x \/ In x l) in match x with | Some e => exist f (proj1_sig e) (or_intror (proj2_sig e)) | None => exist f a (or_introl (eq_refl _)) end) as g. assert (h4:invertible g). apply bijective_impl_invertible. red; split; red. intros o p h4. destruct o; destruct p. simpl in h4. apply exist_injective in h4. f_equal. apply proj1_sig_injective. assumption. simpl in h4. apply exist_injective in h4. contradict h0'. destruct s. simpl in h4. subst. assumption. simpl in h4. apply exist_injective in h4. contradict h0'. destruct s. simpl in h4. subst. assumption. reflexivity. intro y. destruct y as [y h4]. destruct h4 as [h4l | h4r]. exists None. simpl. apply proj1_sig_injective. simpl. assumption. exists (Some (exist _ y h4r)). simpl. reflexivity. rewrite <- (FiniteT_card_bijection _ _ (add_finite _ (list_to_set_finitet l)) g h4) in h3. pose proof (no_dup_cons _ _ h0) as h5. specialize (h1 h5). rewrite h1 in h3. rewrite <- h3. f_equal. apply proof_irrelevance. Qed. Section Transfer. Lemma fun_map_sig_functional : forall {T U:Type} (l l':list T) (f:list {x:T |In x l}->U) (pf:l = l'), let f' := (fun c:list {x:T |In x l'} => f (transfer_dep_r (U := fun lx:list T => list {x:T |In x lx}) pf c)) in f (map_sig l) = f' (map_sig l'). intros. unfold f'. clear f'. subst. rewrite transfer_dep_r_eq_refl. reflexivity. Qed. Lemma map_sig_transfer_dep_r_compat : forall {T U:Type} (l l':list T) (f:{x:T | In x l}->U) (pf:l = l'), let f' := (fun c:{x:T |In x l'} => f (transfer_dep_r (U := fun xl => {x:T |In x xl}) pf c)) in map f (transfer_dep_r (U := fun xl => list {x:T |In x xl}) pf (map_sig l')) = map f' (map_sig l'). intros. unfold f'. clear f'. subst. rewrite transfer_dep_r_eq_refl. apply map_ext_in. intros; rewrite transfer_dep_r_eq_refl. reflexivity. Qed. Lemma map_sig_transfer_dep_compat : forall {T U:Type} (l l':list T) (f:{x:T | In x l'}->U) (pf:l = l'), let f' := (fun c:{x:T |In x l} => f (transfer_dep (U := fun xl => {x:T |In x xl}) pf c)) in map f (transfer_dep (U := fun xl => list {x:T |In x xl}) pf (map_sig l)) = map f' (map_sig l). intros. unfold f'. clear f'. subst. rewrite transfer_dep_eq_refl. apply map_ext_in. intros; rewrite transfer_dep_eq_refl. reflexivity. Qed. End Transfer. bool2-v0-3/AtomlessBooleanAlgebras.v0000644000175000017500000002567313575574055020276 0ustar dwyckoff76dwyckoff76(*Copyright (C) 2016, Daniel Wyckoff*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import FunctionProperties. Require Import SetUtilities. Require Import SetUtilities2. Require Import LogicUtilities. Require Import BoolAlgBasics. Require Import FiniteOperations. Require Import FunctionProperties. Require Import TypeUtilities. Require Import Homomorphisms. Require Import FieldsOfSets. Require Import ArithUtilities. Require Import Even. Require Import Subalgebras. Require Import DecidableDec. Require Import ExtensionsOfHomomorphisms. Require Import FiniteMaps. Require Import ProofIrrelevance. Require Import Description. Require Import SetUtilities2_5. Definition atomless (B:Bool_Alg) : Prop := forall x:bt B, ~ atom x. Definition atomless_p {T:Type} (Bp:Bool_Alg_p T) : Prop := forall x:btp Bp, ~atom_p x. Lemma atomless_p_iff : forall {T:Type} (Bp:Bool_Alg_p T), atomless_p Bp <-> atomless (ba_conv Bp). intros T Bp. split; intro h1; red in h1; red. intros x h2. specialize (h1 _ h2). assumption. intros x h2. rewrite atom_p_iff in h2. specialize (h1 _ h2). assumption. Qed. Lemma degen_atomless : forall (B:Bool_Alg), degen B -> atomless B. intros B h1. red. intros x h2. red in h2. destruct h2 as [h2 h3]. pose proof (degen_eq _ h1 x 0). subst. contradict h2. reflexivity. Qed. Lemma atomless_impl_infinite : forall (B:Bool_Alg), ~degen B -> atomless B -> ~finite_ba B. intros B h0 h1 h2. assert (h3:forall x:bt B, x <> 0 -> exists y:bt B, lt 0 y /\ lt y x). intros x h3. red in h1. specialize (h1 x). apply NNPP. intro h4. contradict h1. pose proof (not_ex_all_not _ (fun y => lt 0 y /\ lt y x) h4) as h5. simpl in h5. red. split. assumption. intros b h6. specialize (h5 b). apply not_and_or in h5. destruct h5 as [h5 | h7]. destruct (zero_lt_or _ b) as [h7 | h8]. right. assumption. contradiction. left. unfold lt in h7. apply not_and_or in h7. destruct h7 as [h8 | h9]. contradiction. apply NNPP. intro h10. apply neq_sym in h10. contradiction. assert (h3':forall x:bt_pos B, exists y:bt_pos B, lt (proj1_sig y) (proj1_sig x)). intro x. destruct x as [x h4]. pose proof (iff2 (neq_zero_lt_iff _ x) h4) as h5. apply h3 in h5. destruct h5 as [y h5]. destruct h5 as [h5 h6]. exists (exist _ _ h5). simpl. assumption. assert (h3'':forall x:sig_set (ba_pos B), exists y:sig_set (ba_pos B), lt (proj1_sig y) (proj1_sig x)). intro x. destruct x as [x h4]. destruct h4 as [h4]. simpl. pose proof (h3' (exist _ _ h4)) as h5. destruct h5 as [y h5]. destruct y as [y h6]. assert (h7:In (ba_pos B) y). constructor; auto. exists (exist _ _ h7). simpl. simpl in h5. assumption. red in h2. pose proof (finite_ba_pos B h2) as hf. assert (hf':Finite (full_sig (ba_pos B))). rewrite <- finite_full_sig_iff. assumption. pose proof (bij_seg _ hf) as h4. destruct h4 as [g h4]. assert (h5:Inhabited (ba_pos B)). apply Inhabited_intro with 1. constructor. red. split. apply zero_min. contradict h0. red. assumption. assert (h5':Inhabited (full_sig (ba_pos B))). destruct h5 as [x h5]. apply Inhabited_intro with (exist _ _ h5). constructor. pose (fun x:sig_set (ba_pos B) => finite_constructive_definite_description _ hf g h4 (fun y =>lt (proj1_sig y) (proj1_sig x)) (h3'' x)) as psi. pose (fun x:sig_set (ba_pos B) => proj1_sig (psi x)) as psi'. assert (h6:Strict_Order (exist_rel (lt (B:=B)) (fun a : bt B => In (ba_pos B) a))). apply strict_order_exist_rel. apply strict_order_lt. assert (h7:(forall x : sig_set (ba_pos B), exist_rel (lt (B:=B)) (fun a : bt B => In (ba_pos B) a) (psi' x) x)). intro x. red. unfold psi', psi. pose proof (proj2_sig (finite_constructive_definite_description (ba_pos B) hf g h4 (fun y : sig_set (ba_pos B) => lt (proj1_sig y) (proj1_sig x)) (h3'' x))) as h7. simpl in h7. assumption. pose proof (strict_order_decreasing_uo_infinite h5' psi' (exist_rel (@lt B) (fun a=>In (ba_pos B) a)) h6 h7) as h8. contradict h8. rewrite <- (finite_full_sig_iff (ba_pos B)). assumption. Qed. Lemma atomless_impl_infinite_p : forall {T:Type} (Bp:Bool_Alg_p T), ~degen_p Bp -> atomless_p Bp -> ~finite_ba_p Bp. intros T Bp h1 h2. rewrite degen_p_iff in h1. rewrite atomless_p_iff in h2. pose proof (atomless_impl_infinite _ h1 h2) as h3. rewrite finite_ba_p_iff. assumption. Qed. Lemma below_non_zero_atomless_ex : forall (B:Bool_Alg), atomless B -> forall (p:bt B), p <> 0 -> exists a : bt B, a <> 0 /\ lt a p. intros B h1 x h3. red in h1. specialize (h1 x). apply NNPP. intro h4. contradict h1. pose proof (not_ex_all_not _ (fun y => y <> 0 /\ lt y x) h4) as h5. simpl in h5. red. split. assumption. intros b h6. specialize (h5 b). apply not_and_or in h5. destruct h5 as [h5 | h7]. destruct (zero_lt_or _ b) as [h7 | h8]. right. assumption. left. assert (h9 : b = 0). tauto. subst. red in h8. destruct h8 as [h8 h9]. contradict h9. reflexivity. unfold lt in h7. apply not_and_or in h7. destruct h7 as [h8 | h9]. contradiction. apply NNPP. intro h10. apply not_or_and in h10. destruct h10 as [h10 h11]. apply neq_sym in h10. contradiction. Qed. Lemma below_non_zero_atomless_ex_p : forall {T:Type} (Bp:Bool_Alg_p T), atomless_p Bp -> forall (p:btp Bp), p <> %0 -> exists a : btp Bp, a <> %0 /\ lt_p a p. intros T Bp h1 p h2. rewrite atomless_p_iff in h1. rewrite ba_conv_zero in h2. pose proof (below_non_zero_atomless_ex _ h1 p h2) as h3. destruct h3 as [a h3]. exists a; auto. Qed. Definition countable_below_non_zero_atomless {B:Bool_Alg} (S:Ensemble nat) (pfinf: ~Finite S) (f:sig_set S -> bt B) (pfb:bijective f) (pfa:atomless B) (p:bt B) (pfnz:p <> 0) : bt B. assert (h1 : exists x : sig_set (Full_set (bt B)), (proj1_sig x) <> 0 /\ lt (proj1_sig x) p). pose proof (below_non_zero_atomless_ex B pfa p pfnz) as h1. destruct h1 as [x h2]. exists (exist _ _ (Full_intro _ x)). simpl. assumption. refine (proj1_sig (proj1_sig (countable_set_constructive_definite_description (full_wrap_fun2 f) _ pfinf (iff1 (bij_full_wrap_fun2_iff f) pfb) h1))). Defined. Definition countable_below_non_zero_atomless_p {T:Type} {Bp:Bool_Alg_p T} (S:Ensemble nat) (pfinf: ~Finite S) (f:sig_set S -> btp Bp) (pfb:bijective f) (pfa:atomless_p Bp) (p:btp Bp) (pfnz:p <> %0) : btp Bp := countable_below_non_zero_atomless S pfinf (ba_conv_fun2 f) pfb pfa p pfnz. Lemma countable_below_non_zero_atomless_compat : forall {B:Bool_Alg} (S:Ensemble nat) (pfinf: ~Finite S) (f:sig_set S -> bt B) (pfb:bijective f) (pfa:atomless B) (p:bt B) (pfnz:p <> 0), let a := countable_below_non_zero_atomless S pfinf f pfb pfa p pfnz in a <> 0 /\ lt a p. unfold countable_below_non_zero_atomless. intros B S pfinf f pfb pfa p pfnz. pose proof (proj2_sig (countable_set_constructive_definite_description (full_wrap_fun2 f) (fun x : sig_set (Full_set (bt B)) => proj1_sig x <> 0 /\ lt (proj1_sig x) p) pfinf (iff1 (bij_full_wrap_fun2_iff f) pfb) match below_non_zero_atomless_ex B pfa p pfnz with | ex_intro x h2 => ex_intro (fun x0 : sig_set (Full_set (bt B)) => proj1_sig x0 <> 0 /\ lt (proj1_sig x0) p) (exist (In (Full_set (bt B))) x (Full_intro (bt B) x)) h2 end)) as h1. destruct h1; split; auto. Qed. Lemma countable_below_non_zero_atomless_compat_p : forall {T:Type} {Bp:Bool_Alg_p T} (S:Ensemble nat) (pfinf: ~Finite S) (f:sig_set S -> btp Bp) (pfb:bijective f) (pfa:atomless_p Bp) (p:btp Bp) (pfnz:p <> %0), let a := countable_below_non_zero_atomless_p S pfinf f pfb pfa p pfnz in a <> %0 /\ lt_p a p. unfold countable_below_non_zero_atomless_p, countable_below_non_zero_atomless. intros B S pfinf f pfb pfa pfnz p h1. pose proof (proj2_sig (countable_set_constructive_definite_description (full_wrap_fun2 (ba_conv_fun2 pfb)) (fun x : sig_set (Full_set (bt (ba_conv S))) => proj1_sig x <> 0 /\ lt (proj1_sig x) p) f (iff1 (bij_full_wrap_fun2_iff (ba_conv_fun2 pfb)) pfa) match below_non_zero_atomless_ex (ba_conv S) pfnz p h1 with | ex_intro x h2 => ex_intro (fun x0 : sig_set (Full_set (bt (ba_conv S))) => proj1_sig x0 <> 0 /\ lt (proj1_sig x0) p) (exist (In (Full_set (bt (ba_conv S)))) x (Full_intro (bt (ba_conv S)) x)) h2 end)) as h2. destruct h2; split; auto. Qed. Lemma atomless_int_alg_ba : atomless int_alg_ba. red. intros x h1. red in h1. destruct h1 as [h1 h2]. destruct x as [S h3]. assert (h4:S <> Empty_set _). intro; subst. contradict h1. apply proj1_sig_injective; simpl. reflexivity. apply not_empty_Inhabited in h4. pose proof (union_of_ho_invls_ex_strict_included _ h4 h3) as h5. destruct h5 as [S' [h5 [h6 h7]]]. pose proof (h2 (exist _ _ h6)) as h2'. assert (h8:Strict_Included (proj1_sig (exist (fun S0 : Ensemble QArith_base.Q => union_of_ho_invls S0) S' h6)) (proj1_sig (exist (fun S0 : Ensemble QArith_base.Q => union_of_ho_invls S0) S h3))). simpl. assumption. rewrite <- int_alg_ba_lt_iff in h8. destruct h8 as [h8 h9]. apply h2 in h8. destruct h8 as [h8 | h8]. pose proof (f_equal (@proj1_sig _ _) h8) as h10. simpl in h10. subst. contradict h9. apply proj1_sig_injective; auto. simpl in h8. pose proof (f_equal (@proj1_sig _ _) h8) as h10. simpl in h10. subst. assert (h10: le (B:=int_alg_ba) (exist (fun S0 : Ensemble QArith_base.Q => union_of_ho_invls S0) (Empty_set QArith_base.Q) h6) (exist (fun S0 : Ensemble QArith_base.Q => union_of_ho_invls S0) S h3)). pose proof (zero_min (exist _ _ h3) (B:=int_alg_ba)) as h11. simpl in h11. pose proof (proof_irrelevance _ h6 (union_of_ho_invls_empty)); subst; auto. specialize (h2 (exist _ _ h6) h10). destruct h2 as [h2 | h2]. pose proof (f_equal (@proj1_sig _ _) h2) as h11. simpl in h11. subst. contradict h9. apply proj1_sig_injective; auto. apply Inhabited_not_empty in h7. contradict h7. reflexivity. Qed. bool2-v0-3/SetUtilities3.v0000644000175000017500000002600513575574055016246 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2015 Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import FiniteMaps. Require Import FiniteMaps2. Require Import NPeano. Require Import DecidableDec. Require Import SetUtilities. Require Import TypeUtilities. Require Import EqDec. Require Import SetUtilities1_5. Require Import LogicUtilities. Require Import ProofIrrelevance. (*Version 0.3 bool2 -- certified -- Version 8.4 pl4 Coq.*) (*This uses [ens_eq_dec], so this might be reworkable.*) Lemma card_power_set : forall {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T), Finite A -> card_fun1 (power_set A) = 2 ^ (card_fun1 A). intros T hdt A h1. pose proof (fin_ens_eq_dec _ hdt) as heq. pose proof (finite_inh_or_empty_dec _ h1) as h2. destruct h2 as [hne | he]. destruct hne as [a hina]. pose proof (cardinal_couple true false) as hcc. assert (h:true <> false). intro h. discriminate. specialize (hcc h). clear h. pose proof bool_finite as hfb. rewrite <- Finite_FiniteT_iff in hfb. pose ([pr:(Ensemble T)*(Fin_map A (Full_set bool) false) | Included (fst pr) A /\ forall x:T, Inn A x -> (Inn (fst pr) x -> (snd pr) |-> x = true) /\ (~Inn (fst pr) x -> (snd pr) |-> x = false)]) as S. assert (hs:fun_well_defined S). red. constructor. intros X h2. destruct h2 as [h2]. destruct h2 as [F h2]. exists F. red. split. split; auto; try constructor. exists X; auto. intros F' h3. destruct h3 as [h3l h3r]. destruct h3l as [h3l]. destruct h3l as [X' h3l]. unfold S in h3l. unfold S in h3r. inversion h3l as [h3a]. clear h3l. simpl in h3a. destruct h3a as [h3a h3b]. inversion h3r as [h4]. clear h3r. simpl in h4. destruct h4 as [h4l h4r]. inversion h2 as [h2a]. simpl in h2a. destruct h2a as [h2l h2r]. apply fin_map_ext. intros x h5. specialize (h3b _ h5). specialize (h4r _ h5). specialize (h2r _ h5). destruct h4r as [h4a h4b]. destruct h2r as [h2a h2b]. destruct (classic (Inn X x)) as [h6 | h7]. simpl. pose proof (fin_map_app_compat F x). specialize (h2a h6). specialize (h4a h6). congruence. specialize (h2b h7). specialize (h4b h7). congruence. intros pr h2. assert (hf:Finite (fst pr)). destruct h2 as [h2]. destruct h2 as [h2]. eapply Finite_downward_closed. apply h1. assumption. destruct (heq A (fst pr) h1 hf) as [h3 | h4]. unfold S in h2. destruct h2 as [h2]. destruct h2 as [h2l h2r]. split. constructor. pose (fun_to_fin_map hdt A false h1 (fun x=>true)) as F. assert (hinc:Included (Im A (fun _ : T => true)) (Full_set bool)). red. intros; constructor. pose (fin_map_new_ran F hfb hinc) as F'. exists F'. constructor. simpl. split. assumption. intros x h5. split. intro h6. specialize (h2r _ h5). destruct h2r as [h2a h2b]. specialize (h2a h6). unfold F'. pose proof (fin_map_new_ran_compat F hfb hinc x) as hc. rewrite <- hc. unfold F. rewrite fun_to_fin_map_compat. reflexivity. assumption. assumption. intro h6. specialize (h2r _ h5). destruct h2r as [h2a h2b]. rewrite h3 in h5. contradiction. constructor. exists (fst pr). constructor. simpl. split; auto. destruct h2 as [h2]. destruct h2 as [h2l h2r]. apply neq_sym in h4. pose proof (Strict_super_set_contains_new_element _ _ _ h2l h4) as hinh. destruct hinh as [x hinh]. destruct hinh as [hinx hnin]. split. constructor. pose (fun_to_fin_map hdt _ false h1 (fun x=>(if (classic_dec (In A x)) then (if (classic_dec (In (fst pr) x)) then true else false) else false))) as F. assert (h6: Included (Im A (fun x : T => if classic_dec (In A x) then if classic_dec (In (fst pr) x) then true else false else false)) (Full_set bool)). red. intros; constructor. pose (fin_map_new_ran F hfb h6) as F'. exists F'. constructor. simpl. split; auto. unfold F'. intros t h9. rewrite <- fin_map_new_ran_compat. unfold F. rewrite fun_to_fin_map_compat. destruct (classic_dec (In A t)) as [h10 | h11]. destruct (classic_dec (In (fst pr) t)) as [h12 | h13]. split. auto. intro. contradiction. split. intro h14. contradiction. auto. contradiction. assumption. assumption. constructor. pose proof (h2r _ hinx) as hinx'. destruct hinx' as [h2a h2b]. pose proof (h2b hnin) as hnin'. exists (fst pr). constructor. simpl. split; auto. red in hs. assert (h2:dom_rel S = power_set A). apply Extensionality_Ensembles. red. split. red. intros X h2. destruct h2 as [h2]. destruct h2 as [f h2]. destruct h2 as [h2]. simpl in h2. destruct h2 as [h2l]. constructor. assumption. red. intros C h2. destruct h2 as [h2]. constructor. pose (fun_to_fin_map hdt _ false h1 (fun x=>(if (classic_dec (In A x)) then (if (classic_dec (In C x)) then true else false) else false))) as F. assert (h3:Included (Im A (fun x : T => if classic_dec (In A x) then if classic_dec (In C x) then true else false else false)) (Full_set bool)). red; intros; constructor. pose (fin_map_new_ran F hfb h3) as F'. exists F'. constructor. simpl. split. assumption. intros x h4. unfold F'. rewrite <- fin_map_new_ran_compat. unfold F. rewrite fun_to_fin_map_compat. split. destruct (classic_dec) as [h5 | h6]. intro h6. destruct (classic_dec) as [h7 | h8]. reflexivity. contradiction. contradiction. intro h5. destruct (classic_dec) as [h6 | h7]. destruct (classic_dec) as [h8 | h9]. contradiction. reflexivity. contradiction. assumption. assumption. rewrite h2 in hs. assert (h3:ran_rel S = Full_set (Fin_map A (Full_set bool) false)). apply Extensionality_Ensembles. red. split. red. intros; constructor. red. intros F ?. constructor. exists [x:T | In A x /\ F |-> x = true]. constructor. simpl. split. red. intros x h3. destruct h3 as [h3]. destruct h3; auto. intros x h3. split. intro h4. destruct h4 as [h4]. destruct h4; auto. intro h4. assert (h5:~ (In A x /\ F |-> x = true)). intro h5. contradict h4. constructor. assumption. apply not_and_or in h5. destruct h5 as [h5l | h5r]. contradiction. destruct (F |-> x). contradict h5r. reflexivity. reflexivity. rewrite h3 in hs. pose proof (finite_fin_maps A (Full_set bool) false h1 hfb) as h4. pose proof (power_set_finite _ h1) as h5. pose (fun_to_fin_map hdt _ false h1 (fun x=>false)) as D. assert (h6:Included (Im A (fun _:T=>false)) (Full_set bool)). red. intros; constructor. pose (fin_map_new_ran D hfb h6) as D'. pose (fin_map_intro _ _ D' (ens_eq_dec T) h5 h4 _ hs) as F. assert (h7:card_fun1 (Full_set bool) = 2). pose proof (card_fun1_compat (Full_set bool)) as h7. destruct h7 as [h7l h7r]. specialize (h7l hfb). rewrite bool_couple in h7l. eapply cardinal_is_functional. rewrite bool_couple. apply h7l. apply hcc. reflexivity. rewrite <- h7. pose proof (card_fin_maps hdt _ _ false h1 hfb) as hcf. rewrite <- hcf. rewrite FiniteT_card_fun1_compat. apply bij_dom_ran_card_eq' with F. red. split. red. intros C E h8 h9 h10. pose proof (fin_map_app_fin_map_to_fps_compat F (C, F|->C) h8 (eq_refl _)) as h11. pose proof (fin_map_app_fin_map_to_fps_compat F (E, F|->E) h9 (eq_refl _)) as h12. rewrite h10 in h11. unfold F in h11. unfold F in h12. rewrite <- fin_map_to_fps_s_compat in h11. rewrite <- fin_map_to_fps_s_compat in h12. inversion h11 as [h11']. simpl in h11'. destruct h11' as [h11a h11b]. inversion h12 as [h12']. simpl in h12'. destruct h12' as [h12a h12b]. apply Extensionality_Ensembles. red. split. red. intros x h13. unfold D' in h11b. assert (h15:In A x). auto with sets. specialize (h11b _ h15). destruct h11b as [h11c h11d]. specialize (h11c h13). rewrite fin_map_app_compat in h12. inversion h12 as [h16]. clear h12. simpl in h16. destruct h16 as [h17 h18]. specialize (h18 _ h15). destruct h18 as [h18l h18r]. apply NNPP. intro h16. specialize (h18r h16). pose proof (fin_map_to_fps_s_compat (ens_eq_dec T) _ _ D' h5 h4 _ hs) as h19. pose proof (subsetT_eq_compat _ _ _ _ hs (fp_fin_map_to_fps (fin_map_intro (power_set A) (Full_set (Fin_map A (Full_set bool) false)) D' (ens_eq_dec T) h5 h4 S hs)) h19) as h20. dependent rewrite <- h20 in h18r. unfold D' in h18r. pose proof (type_eq_dec_eq (fin_map_type_eq_dec (fin_map_intro (power_set A) (Full_set (Fin_map A (Full_set bool) false)) (fin_map_new_ran D hfb h6) (ens_eq_dec T) h5 h4 S hs)) (ens_eq_dec T)) as he. terml h11c. terml h18r. assert (he1:x0 = x1). unfold x0, x1. pose proof @fin_fps_to_f_functional. erewrite fin_fps_to_f_functional. reflexivity; auto. reflexivity. fold x0 x1 in h11c, h18r. rewrite <- he1 in h18r. rewrite h11c in h18r at 1. discriminate. red. intros x h13. unfold D' in h11b. assert (h15:In A x). auto with sets. specialize (h12b _ h15). destruct h12b as [h12c h12d]. specialize (h12c h13). rewrite fin_map_app_compat in h11. inversion h11 as [h16]. clear h12. simpl in h16. destruct h16 as [h17 h18]. specialize (h18 _ h15). destruct h18 as [h18l h18r]. apply NNPP. intro h16. specialize (h18r h16). pose proof (fin_map_to_fps_s_compat (ens_eq_dec T) _ _ D' h5 h4 _ hs) as h19. pose proof (subsetT_eq_compat _ _ _ _ hs (fp_fin_map_to_fps (fin_map_intro (power_set A) (Full_set (Fin_map A (Full_set bool) false)) D' (ens_eq_dec T) h5 h4 S hs)) h19) as h20. dependent rewrite <- h20 in h18r. terml h12c. terml h18r. assert (he1:x0 = x1). unfold x0, x1. pose proof @fin_fps_to_f_functional. erewrite fin_fps_to_f_functional. reflexivity; auto. reflexivity. fold x0 x1 in h12c, h18r. rewrite <- he1 in h18r. rewrite h12c in h18r. discriminate. red. intros G ?. assert (h8:In (power_set A) [x:T|In A x /\ G |->x = true]). constructor. red. intros x h8. destruct h8 as [h8]. destruct h8; assumption. exists [x:T|In A x /\ G |-> x = true]. split; auto. pose proof (in_fin_map_to_fps F _ h8) as h9. unfold F in h9. rewrite <- fin_map_to_fps_s_compat in h9. inversion h9 as [h10]. clear h9. simpl in h10. destruct h10 as [h10l h10r]. apply fin_map_ext. intros x h11. specialize (h10r _ h11). destruct h10r as [h10a h10b]. destruct (eq_dec (G|->x) true) as [h12|h13]. assert (h13:In [x:T | In A x /\ G |-> x =true] x). constructor. split; auto. specialize (h10a h13). simpl. congruence. assert (h14:~In [x:T | In A x /\ G |->x = true] x). intro h15. destruct h15 as [h15]. destruct h15 as [? h15]. rewrite h15 in h13. contradict h13. reflexivity. specialize (h10b h14). simpl. rewrite h10b. destruct (G |-> x). contradict h13. reflexivity. reflexivity. subst. rewrite power_set_empty. rewrite card_fun1_sing. rewrite card_fun1_empty. simpl. reflexivity. Qed. bool2-v0-3/FunctionProperties2.v0000644000175000017500000033353113575574055017465 0ustar dwyckoff76dwyckoff76(*Copyright (C) 2019, Daniel Wyckoff*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import ProofIrrelevance. Require Import FunctionalExtensionality. Require Import NaryFunctions. Require Import Equality. Require Import TypeUtilities. Require Import EqDec. Require Import ArithUtilities. Require Import FunctionProperties. Require Import LogicUtilities. Require Import DecidableDec. Require Import ListUtilities. Require Import SetUtilities. Require Import Description. (*Version 0.3 bool2 -- Certified -- Coq 8.4 pl4.*) (*Is true when given list enumerates the elements in a FiniteT, with respect to the structural recursion (and indexing).*) Inductive FiniteT_list : forall {T}, FiniteT T -> list T -> Prop := | FiniteT_to_list_empty : FiniteT_list empty_finite nil | FiniteT_to_list_add : forall {T} (pff:FiniteT T) (l:list T), FiniteT_list pff l -> FiniteT_list (add_finite T pff) (list_to_option l) | FiniteT_to_list_bij : forall {T U} (pff:FiniteT T) (l:list T) (g:T->U) (pfinv:invertible g), FiniteT_list (bij_finite T U g pff pfinv) (map g l). (*Finds minimum point at which f and g differ (with respect to an indexing..*) Inductive disagrees_at_min {T U} {f g:T->U} (pff:FiniteT T) (h:T->(FiniteT_card _ pff) @) (pffi:FiniteT_index pff h) : disagrees f g -> T -> Prop := | disagrees_at_min_intro : forall c (pfn:f c <> g c), (forall c' (pfn':f c' <> g c'), proj1_sig (h c) <= proj1_sig (h c')) -> disagrees_at_min pff h pffi (disagrees_intro c f g pfn) c. (*Do a sort on [nprod T] and a [FintieT_index] on it also.*) Inductive lt_fun {T U} (pfft:FiniteT T) (ht:T->(FiniteT_card _ pfft) @) (pffu:FiniteT U) (hu:U->(FiniteT_card _ pffu) @) (pffit:FiniteT_index pfft ht) (pffiu:FiniteT_index pffu hu) : Relation (T->U) := | lt_fun_intro : forall c (f g:T->U) (pfd:disagrees f g), disagrees_at_min pfft ht pffit pfd c -> proj1_sig (hu (f c)) < proj1_sig (hu (g c)) -> lt_fun pfft ht pffu hu pffit pffiu f g. Definition transfer_sig_fun1 {T U} {P Q:T->Prop} (pf:impl_pred P Q) (f:{x:T | P x} -> U) : {x:T | Q x} -> U := fun y => f (transfer_sig pf y). Definition transfer_sig_fun1' {T U} {P Q:T->Prop} (pf:iff_pred P Q) (f:{x:T | P x} -> U) : {x:T | Q x} -> U := transfer_sig_fun1 (iff_pred2 _ _ pf) f. Definition transfer_sig_fun2 {T U} {P Q:U->Prop} (pf:pred_impl P Q) (f:T -> {y:U | P y}) : T -> {y:U | Q y} := fun x => transfer_sig pf (f x). Definition transfer_sig_fun2' {T U} {P Q:U->Prop} (pf:iff_pred P Q) (f:T -> {y:U | P y}) : T -> {y:U | Q y} := transfer_sig_fun2 (iff_pred1 _ _ pf) f. Definition transfer_sig_fun_dep {T U} {P:T->Prop} {Q R:U->Prop} (pf:pred_impl Q R) (f:prop_dep_fun P {y:U | Q y}) : prop_dep_fun P {y:U | R y} := fun y (pfy:P y) => let (x, pfx) := (f y pfy) in transfer_sig pf (exist Q x pfx). Definition transfer_sig_fun_dep' {T U} {P:T->Prop} {Q R:U->Prop} (pf:iff_pred Q R) (f:prop_dep_fun P {y:U | Q y}) : prop_dep_fun P {y:U | R y} := transfer_sig_fun_dep (iff_pred1 _ _ pf) f. Lemma inj_transfer_sig_fun1 : forall {T U} {P Q:T->Prop} (pf:impl_pred P Q) (f:{x:T | P x} -> U), Injective f -> Injective (transfer_sig_fun1 pf f). unfold transfer_sig_fun1, transfer_sig; intros ? ? ? ? ? ? h2; red in h2; red; intros ? ? h3; apply h2 in h3; apply exist_injective in h3; apply proj1_sig_injective in h3; auto. Qed. Lemma inj_transfer_sig_fun1' : forall {T U} {P Q:T->Prop} (pf:iff_pred P Q) (f:{x:T | P x} -> U), Injective f -> Injective (transfer_sig_fun1' pf f). intros; apply inj_transfer_sig_fun1; auto. Qed. Lemma surj_transfer_sig_fun1' : forall {T U} {P Q:T->Prop} (pf:iff_pred P Q) (f:{x:T | P x} -> U), surjective f -> surjective (transfer_sig_fun1' pf f). unfold transfer_sig_fun1', transfer_sig_fun1, transfer_sig. intros ? ? ? ? he f h1; red in h1; red; intro y; specialize (h1 y); destruct h1 as [x h2]; subst; destruct x as [x hx]. red in he; destruct he as [he he']. red in he; do 2 red in he'. pose proof (he x hx) as h1. exists (exist Q x h1). repeat f_equal; apply proof_irrelevance. Qed. Lemma bij_transfer_sig_fun1' : forall {T U} {P Q:T->Prop} (pf:iff_pred P Q) (f:{x:T | P x} -> U), bijective f -> bijective (transfer_sig_fun1' pf f). intros ? ? ? ? ? f h1; red in h1; red; destruct h1;split; [apply inj_transfer_sig_fun1'; auto | apply surj_transfer_sig_fun1'; auto]. Qed. Lemma inv_transfer_sig_fun1' : forall {T U} {P Q:T->Prop} (pf:iff_pred P Q) (f:{x:T | P x} -> U), invertible f -> invertible (transfer_sig_fun1' pf f). intros ? ? ? ? ? ? h1; apply invertible_impl_bijective in h1; apply bijective_impl_invertible; apply bij_transfer_sig_fun1'; auto. Qed. Definition inj_transfer_sig_fun2 : forall {T U} {P Q:U->Prop} (pf:pred_impl P Q) (f:T -> {y:U | P y}), Injective f -> Injective (transfer_sig_fun2 pf f). unfold transfer_sig_fun2, transfer_sig; intros ? ? ? ? ? ? h1; red in h1; red; intros ? ? h2; apply exist_injective in h2; apply proj1_sig_injective in h2; apply h1 in h2; auto. Qed. Definition inj_transfer_sig_fun2' : forall {T U} {P Q:U->Prop} (pf:iff_pred P Q) (f:T -> {y:U | P y}), Injective f -> Injective (transfer_sig_fun2' pf f). intros; apply inj_transfer_sig_fun2; auto. Qed. Lemma surj_transfer_sig_fun2' : forall {T U} {P Q:U->Prop} (pf:iff_pred P Q) (f:T -> {y:U | P y}), surjective f -> surjective (transfer_sig_fun2' pf f). unfold transfer_sig_fun2', transfer_sig_fun2, transfer_sig; intros T U P Q h1 f h2; red in h2; red; intros x; destruct x as [x hx]; destruct h1 as [h1 h3]; do 2 red in h3. pose proof (h3 x hx) as h4. specialize (h2 (exist _ _ h4)); destruct h2 as [y h5]. exists y. apply proj1_sig_injective; auto; simpl. rewrite h5; auto. Qed. Lemma bij_transfer_sig_fun2' : forall {T U} {P Q:U->Prop} (pf:iff_pred P Q) (f:T -> {y:U | P y}), bijective f -> bijective (transfer_sig_fun2' pf f). intros ? ? ? ? ? ? h1; destruct h1; split; [apply inj_transfer_sig_fun2'; auto | apply surj_transfer_sig_fun2'; auto]. Qed. Lemma invertible_transfer_sig_fun2' : forall {T U} {P Q:U->Prop} (pf:iff_pred P Q) (f:T -> {y:U | P y}), invertible f -> invertible (transfer_sig_fun2' pf f). intros ? ? ? ? ? ? h1; apply invertible_impl_bijective in h1; apply bijective_impl_invertible; apply bij_transfer_sig_fun2'; auto. Qed. Lemma inj_transfer_sig_fun_dep : forall {T U} {P:T->Prop} {Q R:U->Prop} (pf:pred_impl Q R) (f:prop_dep_fun P {y:U | Q y}), inj_dep f -> inj_dep (transfer_sig_fun_dep pf f). unfold transfer_sig_fun_dep, transfer_sig; intros ? ? ? ? ? h1 f h2; red in h2; red; intros x y hx hy h5. specialize (h2 x y hx hy). destruct (f x hx), (f y hy). apply exist_injective in h5; apply proj1_sig_injective in h5; apply exist_injective in h5; subst. apply h2. f_equal. apply proof_irrelevance. Qed. Lemma inj_transfer_sig_fun_dep' : forall {T U} {P:T->Prop} {Q R:U->Prop} (pf:iff_pred Q R) (f:prop_dep_fun P {y:U | Q y}), inj_dep f -> inj_dep (transfer_sig_fun_dep' pf f). intros; apply inj_transfer_sig_fun_dep; auto. Qed. Lemma surj_transfer_sig_fun_dep' : forall {T U} {P:T->Prop} {Q R:U->Prop} (pf:iff_pred Q R) (f:prop_dep_fun P {y:U | Q y}), surj_dep f -> surj_dep (transfer_sig_fun_dep' pf f). unfold transfer_sig_fun_dep', transfer_sig_fun_dep, transfer_sig; intros ? ? ? ? ? he f h1; red in h1; red; intro y; destruct y as [y hy]. destruct he as [h3 h4]; do 2 red in h4; pose proof (h4 _ hy) as h5. specialize (h1 (exist _ _ h5)); destruct h1 as [x [h1 h2]]. exists x, h1. destruct (f x h1); apply exist_injective in h2; subst. f_equal; apply proof_irrelevance. Qed. Lemma bij_transfer_sig_fun_dep' : forall {T U} {P:T->Prop} {Q R:U->Prop} (pf:iff_pred Q R) (f:prop_dep_fun P {y:U | Q y}), bij_dep f -> bij_dep (transfer_sig_fun_dep' pf f). intros ? ? ? ? ? ? ? h1; red in h1; red; destruct h1 as [h1 h2]; split; [apply inj_transfer_sig_fun_dep'; auto | apply surj_transfer_sig_fun_dep'; auto]. Qed. Lemma inv_id_prop_dep_sig' : forall {T:Type} (P:T->Prop), inv_dep (id_prop_dep_sig' P). unfold id_prop_dep_sig'; intros T P. gterm0. fold c. apply (inv_dep_intro c (@proj1_sig T P) (@proj2_sig T P)); unfold c; simpl; auto. intro y; destruct y; auto. Qed. Lemma inv_id_in_dep_sig' : forall {T:Type} (l:list T), inv_dep (id_in_dep_sig' l). intros; apply inv_id_prop_dep_sig'. Qed. Definition fun_subst {T U} (pfdt:type_eq_dec T) (f:T->U) (x:T) (y:U) : T -> U := fun t => if (pfdt t x) then y else f x. Definition seg_fun_transfer {U n n'} (pf:n = n') (f:seg_fun n U) : seg_fun n' U := transfer_dep pf f (U := fun i => seg_fun i U). Lemma seg_fun_transfer_eq_refl : forall {U n} (pf:n = n) (f:seg_fun n U), seg_fun_transfer pf f = f. unfold seg_fun_transfer; intros ? ? h1. pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. intro; rewrite transfer_dep_eq_refl; auto. Qed. Lemma seg_fun_transfer_compat : forall {U n n'} (pfe:n = n') (f:seg_fun n U) i (pfi:i < n'), seg_fun_transfer pfe f i pfi = f i (lt_eq_trans _ _ _ pfi (eq_sym pfe)). intros; subst. rewrite seg_fun_transfer_eq_refl. f_equal; apply proof_irrelevance. Qed. Lemma seg_fun_transfer_eq : forall {U n n'} (pfe:n = n') (f:seg_fun n U), seg_fun_transfer pfe f = fun i (pfi:i < n') => f i (lt_eq_trans _ _ _ pfi (eq_sym pfe)). intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro; rewrite seg_fun_transfer_compat; auto. Qed. Lemma in_seg_fun_dec_transfer : forall {U n n'} (pfe:n = n') y (f:seg_fun n U), in_seg_fun_dec y f -> in_seg_fun_dec y (seg_fun_transfer pfe f). intros; subst; rewrite seg_fun_transfer_eq_refl; auto. Qed. Lemma in_seg_fun_dec_transfer' : forall {U n n'} (pfe:n = n') y (f:seg_fun n U), in_seg_fun_dec y (seg_fun_transfer pfe f) -> in_seg_fun_dec y f. intros ? ? ? ? ? ? h1; subst; rewrite seg_fun_transfer_eq_refl in h1; auto. Qed. Definition in_seg_fun_dec_subst {U n y y'} (f:seg_fun n U) (pfi:in_seg_fun_dec y f) (pfe:y = y') : in_seg_fun_dec y' f. intros; subst; auto. Defined. Lemma seg_fun_to_pred_compat : forall {U n} (f:seg_fun n U) i (pfi:i < pred n), seg_fun_to_pred f i pfi = f (S i) (lt_pred_S _ _ pfi). intros ? n; destruct n as [|n]. intros; omega. intros f i. destruct i as [|i]; simpl; unfold seg_fun_to_S; intros; f_equal; apply proof_irrelevance. Qed. Lemma seg_fun_from_pred_compat : forall {U n} (f:seg_fun n U) i (pfi:i < pred n), seg_fun_from_pred f i pfi = f i (lt_pred' _ _ pfi). intros ? n; destruct n as [|n]. intros; omega. intros f i. destruct i as [|i]; simpl; unfold seg_fun_from_S; intros; f_equal; apply proof_irrelevance. Qed. Definition seg_fun_dep_w {T n} (U:T->Type) (w:T^n) : Type := forall i (pfi:i < n), U (nth_prod i pfi w). Definition seg_fun_to_list {U n} := map_seg (U :=U ) (n := n). Definition list_to_seg_fun {U} (l:list U) : seg_fun (length l) U := fun i pf => nth_lt l i pf. Lemma seg_fun_to_list_from_list : forall {U n} (hdu:type_eq_dec U) (f:seg_fun (S (S n)) U), seg_fun_to_list (seg_fun_to_S (seg_fun_from_S f)) = seg_fun_to_list (seg_fun_from_S (seg_fun_to_S f)). intros; rewrite leq_iff; auto. ex_goal. do 2 rewrite length_map_seg; auto. exists hex. intros i hi h1. do 2 rewrite nth_lt_map_seg'. unfold seg_fun_to_S, seg_fun_from_S. f_equal. apply proof_irrelevance. Qed. Lemma seg_fun_to_list_to_S : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun (S n) U), 0 < n -> seg_fun_to_list (seg_fun_to_S f) = tl (seg_fun_to_list f). intros U n hdu; induction n as [|n ih]; simpl; auto. intros f ?. destruct (zerop n) as [|h1]; subst. simpl. unfold seg_fun_to_S; repeat f_equal; apply proof_irrelevance. specialize (ih (seg_fun_from_S f)). rewrite <- seg_fun_to_list_from_list. rewrite ih. simpl. rewrite tl_app. rewrite tl_app. rewrite tl_app. f_equal. unfold seg_fun_to_S; repeat f_equal; apply proof_irrelevance. rewrite map_seg_eq_nil_iff; omega. apply app_not_nil_r. intro; discriminate. rewrite map_seg_eq_nil_iff; omega. assumption. assumption. Qed. Fixpoint count_seg_fun_im {U n} (pfdu:type_eq_dec U) y : seg_fun n U -> nat := match n with | 0 => fun _ => 0 | S n' => fun f => let f' := seg_fun_to_S f in let s := count_seg_fun_im pfdu y f' in if (pfdu y (f 0 (lt_0_Sn _))) then S s else s end. Fixpoint count_seg_fun_im_nat {n} y : seg_fun n nat -> nat := match n with | 0 => fun _ => 0 | S n' => fun f => let f' := seg_fun_to_S f in let s := count_seg_fun_im_nat y f' in if (beq_nat y (f 0 (lt_0_Sn _))) then S s else s end. Lemma count_seg_fun_im_functional : forall {U n} (pfdu:type_eq_dec U) y y' (f f':seg_fun n U), y = y' -> f = f' -> count_seg_fun_im pfdu y f = count_seg_fun_im pfdu y' f'. intros; subst; auto. Qed. Lemma count_seg_fun_im_nat_functional : forall {n} y y' (f f':seg_fun n nat), y = y' -> f = f' -> count_seg_fun_im nat_eq_dec y f = count_seg_fun_im nat_eq_dec y' f'. intros; subst; auto. Qed. Lemma unq_inv_dep_ex : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) (pfinv:inv_dep f), exists! pr:{g:U->T & (forall u, P (g u))}, (forall (x:T) (pfx:P x), (projT1 pr) (f x pfx) = x) /\ (forall (y:U), f ((projT1 pr) y) (projT2 pr y) = y). intros T U P f hinv. induction hinv. exists (existT _ g pfp); red; simpl. split; try intros g' hg'; auto. destruct g' as [g' h2]; simpl in hg'. eapply projT_injective; simpl. apply proof_irrelevance. Grab Existential Variables. simpl. apply functional_extensionality. intro y. destruct hg' as [h3 h4]. specialize (H0 y). specialize (h4 y). rewrite <- h4 at 1. rewrite H; auto. Qed. Fixpoint nth_seg_fun_prod_type {n} : forall {f:seg_fun n Type} (g:seg_fun_prod_type f) i pfi, f i pfi := match n with | 0 => fun _ _ _ h1 => False_rect _ (lt_n_0 _ h1) | S n' => fun f g => let f' := seg_fun_to_S f in let (fg, sg) := g in fun i => match i with | 0 => fun pf0 => transfer_dep_pf (lt_0_Sn _) pf0 fg (P := fun i => i < S n') (U := fun j pfj => f j pfj) (x := 0) | S i' => fun pfi => let pfi' := lt_S_n _ _ pfi in transfer_seg_fun_to_S_app f i' pfi' pfi (@nth_seg_fun_prod_type n' f' sg i' pfi') end end. Inductive in_seg_fun_prod_type_aux {n} (f:seg_fun n Type) (g:seg_fun_prod_type f) i (pfi:i < n) (x:f i pfi) : Prop := | in_seg_fun_prod_type_inro : forall j (pfj:j < n) (pfe:f j pfj = f i pfi), transfer pfe (nth_seg_fun_prod_type g j pfj) = x -> in_seg_fun_prod_type_aux f g i pfi x. Definition in_seg_fun_prod_type {n} {f:seg_fun n Type} {i} {pfi:i < n} (x:f i pfi) (g:seg_fun_prod_type f) : Prop := in_seg_fun_prod_type_aux f g i pfi x. Lemma seg_fun_prod_type_to_ex : forall {n} {f:seg_fun n Type} (g:seg_fun_prod_type f), exists! (k:seg_fun_dep f), g = seg_fun_prod k. intro n. induction n as [|n ih]; simpl. intros f x. exists (fun i (pfi:i < 0) => False_rect _ (lt_n_0 _ pfi)). red. split. apply unit_eq. intros f' h1; subst. apply functional_extensionality_dep. intro. apply functional_extensionality_dep. intro; omega. intros f g. destruct g as [fg sg]. specialize (ih (seg_fun_to_S f) sg). destruct ih as [k ih]. exists (fun i => match i with | 0 => fun (pf0:0 < S n) => transfer_dep_pf (lt_0_Sn n) pf0 fg (x:=0) | S i' => fun (pfi:S i' < S n) => let pfi' := lt_S_pred' _ _ (lt_0_Sn _) pfi in transfer_seg_fun_to_S_app f i' pfi' pfi (k i' pfi') end). red. split. f_equal. unfold transfer_dep_pf. rewrite transfer_dep_eq_refl; auto. red in ih. destruct ih as [h1 h2]; subst. f_equal. apply functional_extensionality_dep. intro i. apply functional_extensionality_dep. intro hi. unfold seg_fun_dep_to_S. pose proof (transfer_seg_fun_to_S_app_subst k i hi) as hki. rewrite <- hki. pose proof (proof_irrelevance _ hi (lt_S_pred' (S i) n (lt_0_Sn i) (lt_n_S i n hi))) as h3. rewrite <- h3; auto. red in ih. destruct ih as [h5 h6]. subst. intros sg' h1. inversion h1; subst. inversion h1. specialize (h6 _ H0); subst. apply functional_extensionality_dep. intro i. apply functional_extensionality_dep. intro hi. destruct i as [|i]. unfold transfer_dep_pf. pose proof (proof_irrelevance _ (lt_0_Sn _) hi) as h2. rewrite <- h2. rewrite eq_eq_refl at 1. rewrite transfer_dep_eq_refl; auto. unfold transfer_seg_fun_to_S_app, eq_rect_r, eq_rect. pose proof (proof_irrelevance _ (lt_n_S i n (lt_S_pred' (S i) n (lt_0_Sn i) hi)) hi) as h2. match goal with |- context [match ?pf with | _ => _ end] => pose pf as h3 end. pose (lt_S_pred' (S i) n (lt_0_Sn i) hi) as hlt. fold hlt in h2. pose proof (lt_S_n _ _ hi) as h4. pose proof (proof_irrelevance _ hlt h4) as h5. fold hlt. rewrite h5. unfold hlt in h2. fold hlt in h2. rewrite h5 in h2. clear h5. clear hlt. subst. rewrite eq_eq_refl at 1. unfold seg_fun_dep_to_S; auto. Qed. Definition seg_fun_prod_type_to {n} {f:seg_fun n Type} (g:seg_fun_prod_type f) : seg_fun_dep f := proj1_sig (constructive_definite_description _ (seg_fun_prod_type_to_ex g)). Lemma seg_fun_prod_type_to_compat : forall {n} {f:seg_fun n Type} (g:seg_fun_prod_type f), let k := seg_fun_prod_type_to g in g = seg_fun_prod k. unfold seg_fun_prod_type_to; intros; destruct constructive_definite_description; auto. Qed. Definition inverse_dep {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) (pfinv:inv_dep f) : {g:U->T & (forall u, P (g u))} := proj1_sig (constructive_definite_description _ (unq_inv_dep_ex f pfinv)). Lemma inverse_dep_compat : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) (pfinv:inv_dep f), let pr := inverse_dep f pfinv in (forall (x:T) (pfx:P x), (projT1 pr) (f x pfx) = x) /\ (forall (y:U), f ((projT1 pr) y) (projT2 pr y) = y). unfold inverse_dep; intros; destruct constructive_definite_description; auto. Qed. Definition inverse_dep' {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) (pfinv:inv_dep f) : U->{x:T | P x} := let pr := inverse_dep f pfinv in fun y => exist P (projT1 pr y) (projT2 pr y). Lemma inverse_dep_compat' : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) (pfinv:inv_dep f), let g := inverse_dep' f pfinv in (forall (x:T) (pfx:P x), g (f x pfx) = exist _ x pfx) /\ (forall (y:U), f (proj1_sig (g y)) (proj2_sig (g y)) = y). unfold inverse_dep'; simpl; intros ? ? ? f hinv; pose proof (inverse_dep_compat f hinv) as h1; destruct h1; split; intro; auto. intro; apply proj1_sig_injective; simpl; auto. Qed. Lemma inverse_dep_compat'' : forall {T U} {P:T->Prop} (f:prop_dep_fun P U) (pfinv:inv_dep f) x, proj1_sig (inverse_dep' f pfinv x) = proj1_sig (inverse_dep f pfinv) x. intros ? ? ? f hinv x; unfold inverse_dep'; simpl. pose proof (inverse_dep_compat' f hinv) as h1; simpl in h1. destruct h1 as [h1 h3]. specialize (h3 x). destruct (inverse_dep f hinv); simpl; auto. Qed. Lemma bij_impl_inv_dep : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U), bij_dep f -> inv_dep f. intros T U P f h1. red in h1. destruct h1 as [h1 h2]. red in h1, h2. assert (g:forall y:U, {x:T | exists pfx:P x, y = f x pfx}). intro. apply constructive_definite_description. specialize (h2 y). destruct h2 as [x [h2 h3]]; subst. exists x. red. split. exists h2; auto. intros x' h3. destruct h3 as [h3 h4]. eapply h1. apply h4. pose proof (inv_dep_intro f (fun y:U => proj1_sig (g y)) (fun y:U => (destruct_exists_pf_set_aux _ _ _ (proj2_sig (g y))))) as h3. apply h3; clear h3. intros x hx. destruct (g (f x hx)); simpl. destruct e as [h3 h4]. eapply h1. apply (eq_sym h4). intro y. destruct (g y); simpl. destruct e as [h3 h4]; subst. f_equal; apply proof_irrelevance. Qed. Lemma inv_impl_bij_dep : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U), inv_dep f -> bij_dep f. intros T U P f h1. pose proof (inverse_dep_compat f h1) as h2; simpl in h2. destruct h2 as [h2 h3]. constructor; red. intros x y hx hy h4. pose proof (h2 _ hx) as h5. pose proof (h2 _ hy) as h6. rewrite <- h4 in h6. congruence. intro y. specialize (h3 y). exists (projT1 (inverse_dep f h1) y), (projT2 (inverse_dep f h1) y); auto. Qed. Lemma inv_impl_inj_dep : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U), inv_dep f -> inj_dep f. intros ? ? ? ? h1; apply inv_impl_bij_dep in h1; destruct h1; auto. Qed. Lemma inv_impl_surj_dep : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U), inv_dep f -> surj_dep f. intros ? ? ? ? h1; apply inv_impl_bij_dep in h1; destruct h1; auto. Qed. Lemma inverse_dep_inv' : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) (pfinv:inv_dep f), invertible (inverse_dep' f pfinv). intros T U P f hinv. pose proof (inverse_dep_compat' f hinv) as h4; simpl in h4; destruct h4 as [h4 h5]. pose proof hinv as h1. apply inv_impl_bij_dep in h1. apply bijective_impl_invertible. red in h1; red; destruct h1 as [h1 h2]; split; red. intros x y h3. pose proof (h2 x) as h6. destruct h6 as [a [h6 ?]]; subst. pose proof (h2 y) as h7. destruct h7 as [b [h7 ?]]; subst. do 2 rewrite h4 in h3. apply exist_injective in h3; subst; f_equal; apply proof_irrelevance. intro y. destruct y as [y hy]. exists (f y hy); auto. Qed. Lemma inverse_dep_bij' : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) (pfinv:inv_dep f), bijective (inverse_dep' f pfinv). intros; apply invertible_impl_bijective; apply inverse_dep_inv'. Qed. Lemma inverse_dep_inj' : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) (pfinv:inv_dep f), Injective (inverse_dep' f pfinv). intros; apply invertible_impl_injective; apply inverse_dep_inv'. Qed. Lemma inverse_dep_surj' : forall {T U:Type} {P:T->Prop} (f:prop_dep_fun P U) (pfinv:inv_dep f), surjective (inverse_dep' f pfinv). intros; apply invertible_impl_surjective; apply inverse_dep_inv'. Qed. Lemma inv_fun_in_dep_fun_compose : forall {T U V:Type} {P:T->Prop} (f:prop_dep_fun P U) (g:U->V), inv_dep f -> invertible g -> inv_dep (fun_in_dep_fun_compose f g). intros T U V P f g h1 h2. apply bij_impl_inv_dep. apply inv_impl_bij_dep in h1. apply invertible_impl_bijective in h2. red in h1, h2. destruct h1 as [h1a h1b], h2 as [h2a h2b]. red; split; red. intros x y hx hy h1. unfold fun_in_dep_fun_compose in h1. apply h2a in h1. apply h1a in h1; auto. intro y. specialize (h2b y). destruct h2b as [x]; subst. specialize (h1b x). destruct h1b as [z [hz ?]]; subst. exists z, hz. unfold fun_in_dep_fun_compose; auto. Qed. Lemma inv_tl : forall {T U} {l:list T} (f:fun_in_dep l U) (pf0:l <> nil) (pfi:In (hd' l pf0) (tl l)), inv_dep f -> inv_dep (tl_in_dep f). intros T U l f h0 hi h1. apply inv_impl_bij_dep in h1. apply bij_impl_inv_dep. eapply bij_tl; auto. apply hi. Qed. Lemma inv_dep_sig : forall {T U} {P:T->Prop} (f:prop_dep_fun P U), inv_dep f-> invertible (prop_dep_fun_sig f). unfold prop_dep_fun_sig; intros T U P f h1. apply bijective_impl_invertible. apply inv_impl_bij_dep in h1. apply bij_dep_sig; auto. Qed. Lemma inv_transfer_sig_fun_dep' : forall {T U} {P:T->Prop} {Q R:U->Prop} (pf:iff_pred Q R) (f:prop_dep_fun P {y:U | Q y}), inv_dep f -> inv_dep (transfer_sig_fun_dep' pf f). unfold transfer_sig_fun_dep', transfer_sig_fun_dep; intros T U P Q R h1 f h2; apply inv_impl_bij_dep in h2; apply bij_impl_inv_dep; auto; apply bij_transfer_sig_fun_dep'; auto. Qed. Lemma false_dep_eq : forall {T U} (f:T->False->U), inv_dep f -> equi U False. intros T U f h1. pose proof (inv_dep_sig _ h1) as h2. pose proof (invertible_impl_inv_invertible _ h2) as h4. eapply equi_trans. econstructor. apply h4. apply equi_sym. apply equi_sig_false. Qed. Lemma equi_false_inv : forall T, equi T False -> invertible (false_fun T). intros T h1. apply bijective_impl_invertible. inversion h1. red; split; red. intros; contradiction. intro y. pose (f y); contradiction. Qed. Lemma surj_false_fun_false : forall U, ~ (equi U False) -> surjective (false_fun (U->False)). intro; red; intros h1 f. pose proof (inv_fun_false f) as h2. contradict h1. econstructor. apply h2. Qed. Lemma bij_false_fun_false : forall U, ~ (equi U False) -> bijective (false_fun (U->False)). intros; red; split; [apply inj_false_fun_false | apply surj_false_fun_false]; auto. Qed. Lemma inv_false_fun_false : forall U, ~ (equi U False) -> invertible (false_fun (U->False)). intros; apply bijective_impl_invertible; apply bij_false_fun_false; auto. Qed. Lemma equi_fun_false_false : forall U, ~ equi U False -> equi (U->False) False. intros U h1. pose proof (inv_false_fun_false _ h1) as h2. apply equi_sym. econstructor. apply h2. Qed. Definition count_seg_fun {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) (i:nat) (pfi:i < n) : nat := count_seg_fun_im pfdu (f i pfi) f. Definition count_seg_fun_nat {n} (f:seg_fun n nat) (i:nat) (pfi:i < n) : nat := count_seg_fun_im_nat (f i pfi) f. Lemma count_seg_fun_im_compat : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) i (pfi:i < n), count_seg_fun_im pfdu (f i pfi) f = count_seg_fun pfdu f i pfi. auto. Qed. Lemma count_seg_fun_im_nat_compat : forall {n} (f:seg_fun n nat) i (pfi:i < n), count_seg_fun_im_nat (f i pfi) f = count_seg_fun_nat f i pfi. auto. Qed. Lemma count_seg_fun_im_nat_compat' : forall {n} (f:seg_fun n nat) y, count_seg_fun_im_nat y f = count_seg_fun_im nat_eq_dec y f. intro n. induction n as [|n ih]; simpl; auto. intros f y. destruct (nat_eq_dec y (f 0 (lt_0_Sn n))) as [|h1]; subst. rewrite beq_nat_same. f_equal; auto. rewrite <- neq_nat_compat in h1. rewrite h1; auto. Qed. Lemma count_seg_fun_nat_compat : forall {n} (f:seg_fun n nat), count_seg_fun_nat f = count_seg_fun nat_eq_dec f. intros n f; apply functional_extensionality_dep; intro i; apply functional_extensionality; intro hi. unfold count_seg_fun, count_seg_fun_nat. apply count_seg_fun_im_nat_compat'. Qed. Lemma count_seg_fun_im_to_S : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun (S n) U), count_seg_fun_im pfdu (f 0 (lt_0_Sn _)) f = S (count_seg_fun_im pfdu (f 0 (lt_0_Sn _)) (seg_fun_to_S f)). intros; simpl; lr_if_goal'; auto. contradict hr; auto. Qed. Lemma count_seg_fun_im_in_iff : forall {n U} (pfdu:type_eq_dec U) (f:seg_fun n U) y, count_seg_fun_im pfdu y f > 0 <-> in_seg_fun y f. intro n. induction n as [|n ih]; simpl. intuition. intros U hdu f y. lr_if_goal; fold hlr; destruct hlr as [hl | hr]; subst. split; intro h1; auto with arith. split; intro h1. rewrite ih in h1. right; auto. destruct h1; subst. contradict hr; auto. rewrite ih; auto. Qed. Lemma count_seg_fun_im_in_nat_iff : forall {n} (f:seg_fun n nat) y, count_seg_fun_im_nat y f > 0 <-> in_seg_fun y f. intros n f y. pose proof (count_seg_fun_im_nat_compat' f y) as h1. rewrite h1. apply count_seg_fun_im_in_iff. Qed. Lemma count_seg_fun_im_eq0_iff : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) y, count_seg_fun_im pfdu y f = 0 <-> ~in_seg_fun y f. intros; rewrite <- count_seg_fun_im_in_iff. split; intro h1; [rewrite h1 at 1 | omega]; auto with arith. Qed. Lemma count_seg_fun_im_nat_eq0_iff : forall {n} (f:seg_fun n nat) y, count_seg_fun_im_nat y f = 0 <-> ~in_seg_fun y f. intros; rewrite count_seg_fun_im_nat_compat'; apply count_seg_fun_im_eq0_iff. Qed. Lemma count_seg_fun_im2 : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) i j (pfi:i < n) (pfj:j < n), i <> j -> f i pfi = f j pfj -> count_seg_fun_im pfdu (f i pfi) f > 1. intros U n. induction n as [|n ih]; simpl. intros; omega . intros hdu f i j hi hj hn he. destruct i as [|i]. erewrite f_equal_dep. rewrite eq_dec_same at 1. apply lt_n_S. apply gt_lt. rewrite count_seg_fun_im_in_iff. erewrite f_equal_dep. rewrite he at 1. rewrite in_seg_fun_iff. exists (pred j). ex_goal. omega. exists hex. unfold seg_fun_to_S. gen0. apply neq_sym in hn. rewrite <- (S_pred _ _ (iff1 (neq0_iff _) hn)). intro; f_equal; apply proof_irrelevance. reflexivity. assumption. reflexivity. destruct j as [|j]. rewrite he. erewrite f_equal_dep. rewrite eq_dec_same at 1. apply lt_n_S. apply gt_lt. rewrite count_seg_fun_im_in_iff. erewrite f_equal_dep. rewrite <- he at 1. erewrite f_equal_dep. erewrite seg_fun_to_S_compat'. apply in_seg_fun_compat. reflexivity. reflexivity. assumption. reflexivity. lr_if_goal'. apply lt_n_S. apply gt_lt. rewrite count_seg_fun_im_in_iff. erewrite f_equal_dep. erewrite seg_fun_to_S_compat' at 1. apply in_seg_fun_compat. reflexivity. specialize (ih hdu (seg_fun_to_S f) i j (lt_S_n _ _ hi) (lt_S_n _ _ hj)). hfirst ih. contradict hn; subst; auto. specialize (ih hf). hfirst ih. unfold seg_fun_to_S. erewrite f_equal_dep. rewrite he at 1. f_equal; apply proof_irrelevance. reflexivity. specialize (ih hf0). unfold seg_fun in ih. erewrite f_equal_dep. apply ih. reflexivity. Grab Existential Variables. auto with arith. auto with arith. Qed. (*Returns a strict least upper bound of a [seg_fun] into [nat]*) Fixpoint seg_fun_lub_weak {n} : seg_fun n nat -> nat := match n with | 0 => fun _ => 0 | S n' => fun f => let f' := seg_fun_to_S f in let b := seg_fun_lub_weak f' in if (ge_nat b (f 0 (lt_0_Sn _))) then b else f 0 (lt_0_Sn _) end. Definition seg_fun_lub {n} (f:seg_fun n nat) : nat := S (seg_fun_lub_weak f). Lemma le_seg_fun_lub_weak0 : forall {n} (pf:0 < n) (f:seg_fun n nat), f 0 pf <= seg_fun_lub_weak f. intro n. induction n as [|n ih]; simpl. intros; omega. intros h1 f. destruct n as [|n]; simpl. destruct (zerop (f 0 (lt_0_Sn _))) as [h2 | h2]. rewrite h2. rewrite <- h2. erewrite f_equal_dep. apply le_refl. reflexivity. rewrite (S_pred _ _ h2). rewrite <- (S_pred _ _ h2). erewrite f_equal_dep. apply le_refl. reflexivity. destruct (lt_le_dec (seg_fun_lub_weak (seg_fun_to_S (seg_fun_to_S f))) (seg_fun_to_S f 0 (lt_0_Sn n))) as [h5 | h5]. pose proof h5 as h5'. apply lt_not_le in h5'. rewrite <- le_nat_compat, le_nat_ge, <- eq_false_iff in h5'. rewrite h5'. destruct (lt_le_dec (seg_fun_to_S f 0 (lt_0_Sn n)) (f 0 (lt_0_Sn (S n)))) as [h6 | h6]. pose proof h6 as h6'. apply lt_not_le in h6'. rewrite <- le_nat_compat, le_nat_ge, <- eq_false_iff in h6'. rewrite h6'. erewrite f_equal_dep. apply le_refl. reflexivity. pose proof h6 as h6'. rewrite <- le_nat_compat, le_nat_ge in h6'. rewrite h6'. erewrite f_equal_dep. apply h6. reflexivity. pose proof h5 as h5'. rewrite <- le_nat_compat, le_nat_ge in h5'. rewrite h5'. destruct (lt_le_dec (seg_fun_lub_weak (seg_fun_to_S (seg_fun_to_S f))) (f 0 (lt_0_Sn (S n)))) as [h4 | h4]. pose proof h4 as h4'. apply lt_not_le in h4'. rewrite <- le_nat_compat, le_nat_ge, <- eq_false_iff in h4'. rewrite h4'. erewrite f_equal_dep. apply le_refl. reflexivity. pose proof h4 as h4'. rewrite <- le_nat_compat, le_nat_ge in h4'. rewrite h4'. specialize (ih (lt_0_Sn _) (seg_fun_to_S f)). eapply le_trans. erewrite f_equal_dep. apply h4. reflexivity. apply le_refl. Qed. Lemma lt_seg_fun_lub0 : forall {n} (pf:0 < n) (f:seg_fun n nat), f 0 pf < seg_fun_lub f. unfold seg_fun_lub; intros; apply le_lt; apply le_seg_fun_lub_weak0. Qed. Lemma seg_fun_lub_compat : forall {n} (f:seg_fun n nat), 0 < n -> (forall i (pfi:i < n), f i pfi < seg_fun_lub f) /\ (forall b, (forall i (pfi:i < n), f i pfi < b) -> seg_fun_lub f <= b). unfold seg_fun_lub; intro n. induction n as [|n ih]; simpl. intros ? h1. contradict (lt_irrefl _ h1). intro f; split. intros i h1. destruct i as [|i]. destruct (lt_le_dec (seg_fun_lub_weak (seg_fun_to_S f)) (f 0 (lt_0_Sn n))) as [h4 | h4]. pose proof h4 as h4'. apply lt_not_le in h4'. rewrite <- le_nat_compat, le_nat_ge, <- eq_false_iff in h4'. rewrite h4'. erewrite f_equal_dep. apply lt_n_Sn. reflexivity. pose proof h4 as h4'. rewrite <- le_nat_compat, le_nat_ge in h4. rewrite h4. apply le_lt in h4'. erewrite f_equal_dep. apply h4'. reflexivity. pose proof (lt_S_n _ _ h1) as h2. specialize (ih (seg_fun_to_S f) (Olt _ _ h2)). destruct ih as [h3 h4]. specialize (h3 i h2). destruct (lt_le_dec (seg_fun_lub_weak (seg_fun_to_S f)) (f 0 (lt_0_Sn n))) as [h5 | h5]. pose proof h5 as h5'. apply lt_gt in h5' . apply lt_not_le in h5'. apply nle_nge in h5'. rewrite <- ge_nat_compat in h5'. rewrite <- eq_false_iff in h5'. rewrite h5'. unfold seg_fun_to_S in h3 at 1. eapply lt_le_trans. erewrite f_equal_dep. apply h3. reflexivity. apply le_n_S. auto with arith. apply le_ge in h5. rewrite <- ge_nat_compat in h5. rewrite h5. unfold seg_fun_to_S in h3 at 1. erewrite f_equal_dep. apply h3. reflexivity. intros b h1. destruct (lt_le_dec (seg_fun_lub_weak (seg_fun_to_S f)) (f 0 (lt_0_Sn n))) as [h2 | h2]. pose proof h2 as h3. apply lt_not_le in h3. apply nle_nge in h3. rewrite <- ge_nat_compat in h3. rewrite <- eq_false_iff in h3. rewrite h3. apply lt_le. apply lt_n_S; auto. apply le_ge in h2. pose proof h2 as h2'. rewrite <- ge_nat_compat in h2. rewrite h2. apply lt_le. apply lt_n_S; auto. destruct (zerop n) as [|h3]; subst; simpl. specialize (h1 0 (lt_0_Sn _)); omega. specialize (ih (seg_fun_to_S f) h3). destruct ih as [h4 h5]. specialize (h5 b). hfirst h5. intros i hi. unfold seg_fun_to_S. specialize (h1 _ (lt_n_S _ _ hi)); auto. specialize (h5 hf). omega. Qed. Lemma seg_fun_lub_weak_eq_iff : forall {n} (f:seg_fun n nat) s, 0 < n -> (seg_fun_lub_weak f = s <-> (forall i (pfi:i < n), f i pfi < S s) /\ (forall b, (forall i (pfi:i < n), f i pfi < b) -> S s <= b)). intros n f s h1. pose proof (seg_fun_lub_compat f h1) as h2. unfold seg_fun_lub in h2. split; intro h3; subst; auto with arith. destruct h2 as [h2a h2b], h3 as [h3a h3b]. specialize (h3b (S (seg_fun_lub_weak f)) h2a). specialize (h2b (S s) h3a). omega. Qed. (*Adds the pair [m |-> y] to an [f:seg_fun], and shifts the old [m -> f m] to [S m -> f m], and likewise for function pairs above.*) Fixpoint add_seg_fun {U n} m (y:U) : seg_fun n U -> seg_fun (S n) U := match n with | 0 => fun _ _ _ => y | S n' => match m with | 0 => fun f i => match i with | 0 => fun _ => y | S i' => fun pfi => let pfi := lt_S_n _ _ pfi in f i' pfi end | S m' => fun f i => let f' := seg_fun_to_S f in match i with | 0 => fun _ => f 0 (lt_0_Sn _) | S i' => fun pfi => let pfi' := lt_S_n _ _ pfi in add_seg_fun m' y f' i' pfi' end end end. Lemma add_seg_fun_lt_compat : forall {U n} m y (f:seg_fun n U) i (pfi: i < n) (pfi': i < S n), i < m -> add_seg_fun m y f i pfi' = f i pfi. intros U n. induction n as [|n ih]; simpl. intros; omega. intro m; destruct m as [|m]. intros; omega. intros y f i; destruct i as [|i]. intros; f_equal; apply proof_irrelevance. intros hi hi' h1. specialize (ih m y (seg_fun_to_S f) i (lt_S_n _ _ hi) (lt_S_n _ _ hi')). hfirst ih; auto with arith. specialize (ih hf). rewrite ih. unfold seg_fun_to_S. f_equal; apply proof_irrelevance. Qed. Lemma add_seg_fun_compat_eq : forall {U n} m y (f:seg_fun n U) (pfm:m < S n), m < n -> add_seg_fun m y f m pfm = y. intros U n. induction n as [|n ih]; simpl; auto. intro m. destruct m as [|m]; auto. intros y f hm h1. rewrite ih; auto with arith. Qed. Lemma add_seg_fun_compat_gt : forall {U n} m y (f:seg_fun n U) i (pfi: i < S n) (pfi': pred i < n), i > m -> add_seg_fun m y f i pfi = f (pred i) pfi'. intros U n. induction n as [|n ih]; simpl. intros; omega. intro m; destruct m as [|m]. intros y f i; destruct i as [|i]. intros; omega. simpl; intros; f_equal; apply proof_irrelevance. intros y f i. destruct i as [|i]. simpl; intros; f_equal; apply proof_irrelevance. intros hi hi' h1. simpl. erewrite ih; auto with arith. unfold seg_fun_to_S. destruct i as [|i]. omega. simpl. f_equal. apply proof_irrelevance. Grab Existential Variables. simpl in hi'. omega. Qed. (*Removes one [pair-value] pair in a [seg_fun] at [m:nat] and shifts down the greater part of the bipartite partition.*) Fixpoint remove_seg_fun {U n} m : seg_fun n U -> seg_fun (pred n) U := match n return (seg_fun n U -> seg_fun (pred n) U ) with | 0 => fun _ _ pfi => False_rect _ (lt_n_0 _ pfi) | S n' => fun f => let f' := seg_fun_to_S f in let f_ := seg_fun_from_S f in match m with | 0 => f' | S m' => fun i => match i with | 0 => fun pfi => f_ 0 pfi | S i' => fun pfi => let pfi' := lt_S_pred _ _ pfi in remove_seg_fun m' f' i' pfi' end end end. Lemma remove_seg_fun_functional : forall {U n} m m' (f:seg_fun n U), m = m' -> remove_seg_fun m f = remove_seg_fun m' f. intros; subst; auto. Qed. Lemma remove_seg_fun_lt : forall {U n} (f:seg_fun n U) m i (pfi:i < pred n), i < m -> let pfi' := lt_pred' _ _ pfi in remove_seg_fun m f i pfi = f i pfi'. intros U n. induction n as [|n ih]. simpl; intros; omega. simpl. intros f m. destruct m as [|m]. intros; omega. intro i. destruct i as [|i]. unfold seg_fun_from_S. intros; f_equal; apply proof_irrelevance. intros hi h1. apply lt_S_n in h1. specialize (ih (seg_fun_to_S f) m i (lt_S_pred _ _ hi) h1). rewrite ih. unfold seg_fun_to_S. f_equal. apply proof_irrelevance. Qed. Lemma remove_seg_fun_ge : forall {U n} (f:seg_fun n U) m i (pfi:i < pred n), i >= m -> let pfi' := lt_pred_S _ _ pfi in remove_seg_fun m f i pfi = f (S i) pfi'. intros U n. induction n as [|n ih]. simpl; intros; omega. intros f m. destruct m as [|m]. intros; simpl. unfold seg_fun_to_S; f_equal; apply proof_irrelevance. simpl. intro i. destruct i as [|i]. intros; omega. intros hi h1. rewrite ih; auto with arith. unfold seg_fun_to_S. f_equal. apply proof_irrelevance. Qed. Lemma remove_seg_fun_ge_pred : forall {U n} m (f:seg_fun n U), m >= pred n -> remove_seg_fun m f = remove_seg_fun (pred n) f. intros U n. induction n as [|n ih]; auto. intro m. destruct m as [|m]. intros ? h1. apply le0 in h1; subst; auto. simpl in h1; subst; simpl; auto. intros f h1. simpl in h1. rewrite (remove_seg_fun_functional _ _ _ (eq_refl n)). simpl. destruct n as [|n]. apply functional_extensionality_dep. intro i. destruct i as [|i]; apply functional_extensionality_dep; intro; omega. apply le_S_n in h1. apply functional_extensionality_dep. intro i. destruct i as [|i]; auto. apply functional_extensionality_dep. intro h2. pose proof (lt_S_n _ _ h2) as h3. specialize (ih _ (seg_fun_to_S f) h1). pose proof (equal_f_dep _ _ ih _ (lt_S_pred _ _ h2)) as h4; auto. Qed. Lemma remove_seg_fun_to_S : forall {U n} m (f:seg_fun (S n) U), 0 < m -> forall (pfn:0 < n), remove_seg_fun (S m) f = seg_fun_transfer (eq_sym (S_pred _ _ pfn)) (seg_fun_cons (f 0 (lt_0_Sn _)) (remove_seg_fun m (seg_fun_to_S f))). intros; simpl. apply functional_extensionality_dep. intro i. destruct i as [|i]. apply functional_extensionality. intro h1. rewrite seg_fun_transfer_compat. simpl. unfold seg_fun_from_S; f_equal; apply proof_irrelevance. apply functional_extensionality_dep. intro h1. pose proof (lt_eq_trans _ _ _ h1 (S_pred _ _ pfn)) as h2. apply lt_S_n in h2. rewrite seg_fun_transfer_compat. destruct (lt_le_dec i m) as [h3 | h3]. rewrite remove_seg_fun_lt; auto. simpl. rewrite remove_seg_fun_lt; auto. repeat f_equal; apply proof_irrelevance. rewrite remove_seg_fun_ge; auto. simpl. rewrite remove_seg_fun_ge; auto. repeat f_equal; apply proof_irrelevance. Qed. Lemma remove_seg_fun_eq : forall {U n} m (f:seg_fun n U), 0 < pred m -> forall (pfn:0 < n) (pfp:0 < pred n), remove_seg_fun m f = seg_fun_transfer (eq_sym (S_pred _ _ pfp)) (seg_fun_cons (f 0 pfn) (remove_seg_fun (pred m) (seg_fun_to_S (seg_fun_transfer (S_pred _ _ pfn) f)))). intros U n. destruct n as [|n]. simpl. intros; omega. intros m f h1 h2 h3. destruct m as [|m]. contradict (lt_irrefl _ h1). erewrite remove_seg_fun_to_S; auto. repeat f_equal. apply proof_irrelevance. apply functional_extensionality_dep. intro. apply functional_extensionality. intro. rewrite seg_fun_transfer_compat. f_equal. apply proof_irrelevance. Qed. Lemma remove_seg_fun_to_S_swap0 : forall {U n} (f:seg_fun (S n) U) (pf:0 < n), remove_seg_fun 0 (seg_fun_to_S f) = seg_fun_to_S (seg_fun_transfer (S_pred _ _ pf) (remove_seg_fun 0 f)). intros; apply functional_extensionality_dep; intro; apply functional_extensionality; intro. rewrite remove_seg_fun_ge; auto with arith. unfold seg_fun_to_S. rewrite seg_fun_transfer_compat. rewrite remove_seg_fun_ge; auto with arith. f_equal; apply proof_irrelevance. Qed. Lemma remove_seg_fun_to_S_swap_pos : forall {U n} (f:seg_fun (S n) U) r (pf:0 < n), 0 < r -> remove_seg_fun r (seg_fun_to_S f) = seg_fun_to_S (seg_fun_transfer (S_pred _ _ pf) (remove_seg_fun (S r) f)). intros U n. destruct n as [|n]. intros; omega. intros f r h0 hr. apply functional_extensionality_dep. intro i. apply functional_extensionality. intro hi. destruct (lt_le_dec i r) as [h2|h2]; subst. rewrite remove_seg_fun_lt; auto with arith. unfold seg_fun_to_S. rewrite seg_fun_transfer_compat. rewrite remove_seg_fun_lt; auto with arith. f_equal; apply proof_irrelevance. unfold seg_fun_to_S; rewrite seg_fun_transfer_compat. rewrite remove_seg_fun_ge; auto with arith. rewrite remove_seg_fun_ge; auto with arith. f_equal; apply proof_irrelevance. Qed. Lemma seg_fun_to_S_remove1 : forall {U n} (f:seg_fun (S (S n)) U), seg_fun_to_S (remove_seg_fun 1 f) = seg_fun_to_S (seg_fun_to_S f). intros; apply functional_extensionality_dep; intro i; apply functional_extensionality; intro hi; unfold seg_fun_to_S; rewrite remove_seg_fun_ge; auto with arith. f_equal; apply proof_irrelevance. Qed. Lemma seg_fun_remove0_to_S : forall {U n} (f:seg_fun (S (S n)) U), remove_seg_fun 0 (seg_fun_to_S f) = seg_fun_to_pred (seg_fun_to_S f). intros; apply functional_extensionality_dep; intro i; apply functional_extensionality; intro hi; rewrite remove_seg_fun_ge; auto with arith; unfold seg_fun_to_pred, seg_fun_to_S; f_equal; try apply proof_irrelevance. Qed. Lemma seg_fun_from_S_remove : forall {U n} (f:seg_fun (S (S n)) U) r, seg_fun_from_S (remove_seg_fun r f) = remove_seg_fun r (seg_fun_from_S f). intros; apply functional_extensionality_dep; intro i; apply functional_extensionality; intro hi; rewrite seg_fun_from_S_compat. destruct (lt_le_dec i r) as [h2|h2]; subst. do 2 (rewrite remove_seg_fun_lt; auto). rewrite seg_fun_from_S_compat; f_equal; apply proof_irrelevance. do 2 (rewrite remove_seg_fun_ge; auto). rewrite seg_fun_from_S_compat; f_equal; apply proof_irrelevance. Qed. Lemma seg_fun_endo_transfer : forall {n n'} (f:seg_fun n nat) r (pfe:n = n'), seg_fun_endo f r -> seg_fun_endo (seg_fun_transfer pfe f) r. intros; subst; rewrite seg_fun_transfer_eq_refl; auto. Qed. Lemma inj_seg_fun_endo_transfer : forall {n n'} (f:seg_fun n nat) r (pfe:n = n'), inj_seg_fun_endo f r -> inj_seg_fun_endo (seg_fun_transfer pfe f) r. intros; subst; rewrite seg_fun_transfer_eq_refl; auto. Qed. Lemma surj_seg_fun_endo_transfer : forall {n n'} (f:seg_fun n nat) r (pfe:n = n'), surj_seg_fun_endo f r -> surj_seg_fun_endo (seg_fun_transfer pfe f) r. intros; subst; rewrite seg_fun_transfer_eq_refl; auto. Qed. Lemma seg_fun_endo_transfer' : forall {n n'} (f:seg_fun n nat) r (pfe:n = n'), seg_fun_endo (seg_fun_transfer pfe f) r -> seg_fun_endo f r. intros ? ? ? ? ? h1; subst. rewrite seg_fun_transfer_eq_refl in h1; auto. Qed. Lemma inj_seg_fun_endo_transfer' : forall {n n'} (f:seg_fun n nat) r (pfe:n = n'), inj_seg_fun_endo (seg_fun_transfer pfe f) r -> inj_seg_fun_endo f r. intros ? ? ? ? ? h1; subst. rewrite seg_fun_transfer_eq_refl in h1; auto. Qed. Lemma surj_seg_fun_endo_transfer' : forall {n n'} (f:seg_fun n nat) r (pfe:n = n'), surj_seg_fun_endo (seg_fun_transfer pfe f) r -> surj_seg_fun_endo f r. intros ? ? ? ? ? h1; subst. rewrite seg_fun_transfer_eq_refl in h1; auto. Qed. Lemma remove_seg_fun0 : forall {U n} (f:seg_fun (S n) U), remove_seg_fun 0 f = seg_fun_to_S f. intros ? n; destruct n; simpl; auto. Qed. Lemma remove_seg_fun0': forall {U n} (f:seg_fun n U), remove_seg_fun 0 f = seg_fun_to_pred f. intros ? n; destruct n; simpl; auto. Qed. Lemma seg_fun_to_S_eq : forall {U n} (f:seg_fun (S n) U), seg_fun_to_S f = remove_seg_fun 0 f. intros ? n; destruct n; simpl; auto. Qed. Lemma seg_fun_from_S_eq : forall {U n} (f:seg_fun (S n) U), seg_fun_from_S f = remove_seg_fun n f. intros U n. destruct n as [|n]. intro f. apply functional_extensionality_dep. intro i. apply functional_extensionality. intro; omega. intro f. apply functional_extensionality_dep. intro i. destruct i as [|i]. apply functional_extensionality. intro; auto. apply functional_extensionality. intro hi. rewrite remove_seg_fun_lt; auto. unfold seg_fun_from_S. f_equal; apply proof_irrelevance. Qed. Lemma seg_fun_to_S_remove_seg_fun1 : forall {n} (f:seg_fun (S (S n)) nat), seg_fun_to_S (remove_seg_fun 1 f) = seg_fun_to_S (seg_fun_to_S f). intros n f; apply functional_extensionality_dep; intro i; apply functional_extensionality; intro hi; unfold remove_seg_fun, seg_fun_to_S; f_equal; try apply proof_irrelevance. Qed. Lemma inv_seg_fun_eq_same_iff : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) i (pfi:i < n), (inv_seg_fun pfdu (f i pfi) f = i <-> forall j (pfj:j < i) (pfj':j < n), f j pfj' <> f i pfi). intros U n. induction n as [|n ih]; simpl. intro hdu; intros f i hi. contradict (lt_n_O _ hi). intros hdu f i hi. destruct i as [|i]. erewrite f_equal_dep. rewrite eq_dec_same at 1; auto. intuition. reflexivity. lr_if_goal; fold hlr; destruct hlr as [hl | hr]. split; intro h1; try discriminate. specialize (h1 0 (lt_0_Sn _) (lt_0_Sn _)). contradiction. pose proof (lt_S_n _ _ hi) as h1. specialize (ih hdu (seg_fun_to_S f) i h1). unfold seg_fun_to_S in ih at 1. split. intros h3 j hj hj'. apply S_inj in h3. destruct j as [|j]. erewrite f_equal_dep. apply hr. reflexivity. erewrite f_equal_dep in h3. rewrite ih in h3 at 1. unfold seg_fun_to_S in h3. erewrite f_equal_dep. apply neq_sym. erewrite f_equal_dep. apply neq_sym. apply h3. apply lt_S_n. apply hj. reflexivity. reflexivity. reflexivity. intros h2. f_equal. erewrite f_equal_dep. rewrite ih at 1. intros j h3 h4. unfold seg_fun_to_S. specialize (h2 (S j) (lt_n_S _ _ h3) (lt_n_S _ _ h4)). contradict h2. rewrite h2. f_equal; apply proof_irrelevance. reflexivity. Grab Existential Variables. apply lt_S_n in hj'; auto. Qed. Lemma seg_fun_endo_to_S' : forall {n} (f:seg_fun (S n) nat) r s, r <= s -> f 0 (lt_0_Sn _) < s -> seg_fun_endo (seg_fun_to_S f) r -> seg_fun_endo f s. intros n f r s hr h1 h2. inversion h2 as [? ? h3 h4]; subst. constructor. intros i hi. destruct i as [|i]. erewrite f_equal_dep. apply h1. reflexivity. specialize (h3 _ (lt_S_n _ _ hi)). unfold seg_fun_to_S in h3. eapply lt_le_trans. erewrite f_equal_dep. apply h3. reflexivity. assumption. Qed. Lemma seg_fun_endo_from_S' : forall {n} (f:seg_fun (S n) nat) r s, r <= s -> f n (lt_n_Sn _) < s -> seg_fun_endo (seg_fun_from_S f) r -> seg_fun_endo f s. intros n f r s hr h1 h2. inversion h2 as [? ? h3 h4]; subst. constructor. intros i hi. pose proof (lt_dec _ _ hi) as h4. simpl in h4. destruct h4 as [h4|]; subst; auto. specialize (h3 i h4). unfold seg_fun_from_S in h3. eapply lt_le_trans. erewrite f_equal_dep. apply h3. reflexivity. assumption. erewrite f_equal_dep. apply h1. reflexivity. Qed. Lemma seg_fun_endo_to_S_bound : forall {n} b (f:seg_fun (S n) nat), seg_fun_endo f b -> ~in_seg_fun (pred b) (seg_fun_to_S f) -> seg_fun_endo (seg_fun_to_S f) (pred b). intros n b f h1 h2. destruct h1 as [? ? h1]. constructor. intros i hi. specialize (h1 (S i) (lt_n_S _ _ hi)). apply lt_dec in h1. destruct h1 as [h1|h1]. unfold seg_fun_to_S; auto. contradict h2. rewrite <- h1. erewrite f_equal_dep. erewrite seg_fun_to_S_compat'. apply in_seg_fun_compat. reflexivity. Grab Existential Variables. auto with arith. Qed. Lemma shift_seg_fun_endo : forall n k, seg_fun_endo (shift_seg_fun n k) n. intros n k; constructor; intros; apply lt_shift_seg_fun. Qed. Lemma inj_seg_fun_endo_shift : forall n k, inj_seg_fun_endo (shift_seg_fun n k) n. intros; constructor; try apply shift_seg_fun_endo; apply inj_shift_seg_fun. Qed. Lemma seg_fun_endo_right_shift : forall n, seg_fun_endo (right_shift_seg_fun n) n. intros; apply shift_seg_fun_endo. Qed. Lemma seg_fun_endo_left_shift : forall n, seg_fun_endo (left_shift_seg_fun n) n. intros; apply shift_seg_fun_endo. Qed. Lemma add_seg_to_identify_endo : forall n, 0 < n -> seg_fun_endo (add_seg_to_identify n) n. intros; constructor; auto; intros. apply lt_add_seg_to_identify; auto. Qed. Lemma tl_seg_for_fictitious_endo : forall n, seg_fun_endo (tl_seg_for_fictitious n) n. intro n. constructor. apply lt_tl_seg_for_fictitious. Qed. Lemma seg_fun_endo_remove : forall {n i m} (f:seg_fun n nat), i < n -> seg_fun_endo f m-> seg_fun_endo (remove_seg_fun i f) m. intros n i m f h1 h2. revert h1. revert i. induction h2 as [r f ih]. intros i h1. econstructor. intros j hj. destruct (le_lt_dec i j) as [h2 | h2]. pose proof (remove_seg_fun_ge f i j hj h2) as h3. simpl in h3. rewrite h3; auto. pose proof (remove_seg_fun_lt f i j hj h2) as h3. simpl in h3. rewrite h3; auto. Qed. Lemma seg_fun_endo_remove' : forall {n r} (f:seg_fun n nat) b c (pfr:r < n) , b <= c -> f r pfr < c -> seg_fun_endo (remove_seg_fun r f) b -> seg_fun_endo f c. intros n r f b c hb hc h1 h2. inversion h2 as [? ? h3 h4]; subst. destruct (zerop (pred n)) as [h4|h4]. apply pred0_dec in h4. destruct h4 as [h4 | h4]. contradict (lt_n_0 _ (lt_eq_trans _ _ _ hb h4)). pose proof (lt_eq_trans _ _ _ hb h4) as h5. apply lt1 in h5. revert h1. revert hb. subst. intros h0 h1. constructor. intros i hi. pose proof (lt1 _ hi); subst. erewrite f_equal_dep. apply h1. reflexivity. constructor. intros i hi. destruct (zerop r) as [|h5]; subst. destruct i as [|i]. erewrite f_equal_dep. apply h1. reflexivity. specialize (h3 i (lt_S_pred _ _ hi)). rewrite remove_seg_fun_ge in h3. eapply lt_le_trans. erewrite f_equal_dep. apply h3. reflexivity. assumption. auto with arith. pose proof (lt_dec _ _ hi) as h6. destruct (lt_le_dec i r) as [h7|h7]; subst. destruct h6 as [h6|]; subst. specialize (h3 i h6). rewrite remove_seg_fun_lt in h3; auto. eapply lt_le_trans. erewrite f_equal_dep. apply h3. reflexivity. assumption. omega. destruct h6 as [h6|]; subst. specialize (h3 (pred i)). hfirst h3. omega. specialize (h3 hf). apply le_dec' in h7. destruct h7 as [h7|]; subst. rewrite remove_seg_fun_ge in h3; auto. revert h3. generalize (lt_pred_S _ _ hf). assert (h8:0 < i). omega. rewrite <- (S_pred _ _ h8). intros h9 h10. erewrite f_equal_dep. eapply lt_le_trans. apply h10. assumption. reflexivity. erewrite f_equal_dep. apply h1. reflexivity. apply le_dec' in h7. destruct h7 as [h7|]; subst. assert (h8:pred (pred n) < pred n). omega. specialize (h3 (pred (pred n)) h8). rewrite remove_seg_fun_ge in h3; auto. revert h3. generalize (lt_pred_S _ _ h8). rewrite <- (S_pred _ _ h4). intros h9 h10. eapply lt_le_trans. erewrite f_equal_dep. apply h10. reflexivity. assumption. erewrite f_equal_dep. apply h1. reflexivity. Qed. Lemma inj_seg_fun_endo_remove : forall {n i m} (f:seg_fun n nat), i < n -> inj_seg_fun_endo f m-> inj_seg_fun_endo (remove_seg_fun i f) m. intros n i m f h1 h2. constructor. apply seg_fun_endo_remove; auto. inversion h2; auto. intros r s hr hs h3. inversion h2; subst. destruct (lt_le_dec r i) as [h4 | h4]. rewrite (remove_seg_fun_lt f i r hr h4) in h3. destruct (lt_le_dec s i) as [h5 | h5]. rewrite (remove_seg_fun_lt f i s hs h5) in h3. apply H0 in h3; auto. rewrite (remove_seg_fun_ge f i s hs h5) in h3. apply H0 in h3; subst. omega. rewrite (remove_seg_fun_ge f i r hr h4) in h3. destruct (lt_le_dec s i) as [h5 | h5]. rewrite (remove_seg_fun_lt f i s hs h5) in h3. apply H0 in h3; subst; omega. rewrite (remove_seg_fun_ge f i s hs h5) in h3. apply H0 in h3. apply S_inj in h3; auto. Qed. Lemma inj_seg_fun_endo_to_S : forall {n} (f:seg_fun (S n) nat) r, inj_seg_fun_endo f r-> inj_seg_fun_endo (seg_fun_to_S f) r. intros; rewrite seg_fun_to_S_eq. eapply inj_seg_fun_endo_remove in H. apply H. auto with arith. Qed. Lemma inj_seg_fun_endo_from_S : forall {n} (f:seg_fun (S n) nat) r, inj_seg_fun_endo f r-> inj_seg_fun_endo (seg_fun_from_S f) r. intros; rewrite seg_fun_from_S_eq. eapply inj_seg_fun_endo_remove in H. apply H. auto with arith. Qed. Definition inj_seg_fun_lub {n} (f:seg_fun n nat) : Prop := inj_seg_fun_endo f (seg_fun_lub f). Definition inv_seg_fun_lub_weak {n} (f:seg_fun n nat) (pf0:0 < n) : nat := inv_seg_fun nat_eq_dec (seg_fun_lub_weak f) f. Definition remove_seg_fun_lub_weak {n} (f:seg_fun n nat) (pf0:0 < n) : seg_fun (pred n) nat := remove_seg_fun (inv_seg_fun_lub_weak f pf0) f. Lemma le_seg_fun_lub_weak_to_S : forall {n} (f:seg_fun (S n) nat), seg_fun_lub_weak (seg_fun_to_S f) <= seg_fun_lub_weak f. intro n. destruct n as [|n]. simpl; auto with arith. intro f. simpl. destruct (gt_ge_dec (seg_fun_to_S f 0 (lt_0_Sn n)) (seg_fun_lub_weak (seg_fun_to_S (seg_fun_to_S f)))) as [h1 | h1]. apply gt_not_ge in h1. rewrite <- ge_nat_compat in h1. rewrite <- eq_false_iff in h1. rewrite h1. destruct (gt_ge_dec (f 0 (lt_0_Sn (S n))) (seg_fun_to_S f 0 (lt_0_Sn n))) as [h2 | h2]. pose proof h2 as h2'. apply gt_not_ge in h2. rewrite <- ge_nat_compat in h2. rewrite <- eq_false_iff in h2. rewrite h2. omega. rewrite <- ge_nat_compat in h2. rewrite h2; auto. rewrite <- ge_nat_compat in h1. rewrite h1. destruct (gt_ge_dec (f 0 (lt_0_Sn (S n))) (seg_fun_lub_weak (seg_fun_to_S (seg_fun_to_S f)))) as [h2 | h2]. pose proof h2 as h2'. apply gt_not_ge in h2. rewrite <- ge_nat_compat in h2. rewrite <- eq_false_iff in h2. rewrite h2. omega. rewrite <- ge_nat_compat in h2. rewrite h2; auto. Qed. Lemma lt_seg_fun_lub_to_S_dec : forall {n} (f:seg_fun (S n) nat) (pf:0 < S n), {f 0 pf = seg_fun_lub_weak f} + {seg_fun_lub_weak f = seg_fun_lub_weak (seg_fun_to_S f)}. intro n. destruct n as [|n]. simpl. intros f h1. destruct (zerop (f 0 (lt_0_Sn 0))) as [h2 | h2]. rewrite h2. right; auto. rewrite (S_pred _ _ h2). rewrite <- (S_pred _ _ h2). left. apply f_equal_dep; auto. intros f h1. simpl. destruct (gt_ge_dec (seg_fun_to_S f 0 (lt_0_Sn n)) (seg_fun_lub_weak (seg_fun_to_S (seg_fun_to_S f)))) as [h2 | h2]. pose proof h2 as h2'. apply gt_not_ge in h2'. rewrite <- ge_nat_compat in h2'. rewrite <- eq_false_iff in h2'. rewrite h2'. destruct (gt_ge_dec (f 0 (lt_0_Sn (S n))) (seg_fun_to_S f 0 (lt_0_Sn n))) as [h3 | h3]. pose proof h3 as h3'. apply gt_not_ge in h3'. rewrite <- ge_nat_compat in h3'. rewrite <- eq_false_iff in h3'. rewrite h3'. left. apply f_equal_dep; auto. pose proof h3 as h3'. rewrite <- ge_nat_compat in h3'. rewrite h3'. right; auto. pose proof h2 as h2'. rewrite <- ge_nat_compat in h2'. rewrite h2'. destruct (gt_ge_dec (f 0 (lt_0_Sn _)) (seg_fun_lub_weak (seg_fun_to_S (seg_fun_to_S f)))) as [h3 | h3]. pose proof h3 as h3'. apply gt_not_ge in h3'. rewrite <- ge_nat_compat in h3'. rewrite <- eq_false_iff in h3'. rewrite h3'. left; f_equal; auto; apply proof_irrelevance. pose proof h3 as h3'. rewrite <- ge_nat_compat in h3'. rewrite h3'. right; auto. Qed. Section PreImage. Definition pre_class {X S} (v:X->S) (s:S) : Ensemble X := [x:X | v x = s]. Lemma in_pre_class : forall {X S} (v:X->S) x, Inn (pre_class v (v x)) x. intros; constructor; auto. Qed. Lemma pre_class_inj : forall {X S} (v:X->S) (s s':S), in_pre s v -> pre_class v s = pre_class v s' -> s = s'. intros X S v s s' h1 h2. destruct h1. pose proof (in_pre_class v x) as h3. rewrite h2 in h3. destruct h3; auto. Qed. Lemma in_pre_inhabited : forall {X S} (v:X->S) (s:S), in_pre s v -> Inhabited (pre_class v s). intros X S v s h1. destruct h1. apply Inhabited_intro with x. constructor; auto. Qed. Lemma inh_in_pre : forall {X S} (v:X->S) (s:S), Inhabited (pre_class v s) -> in_pre s v. intros X S v s h1. inversion h1 as [x h2]. destruct h2 as [h2]; subst. constructor. Qed. Inductive pre_classes {X S} (v:X->S) : Family X := | pre_classes_intro : forall s, in_pre s v -> Inn (pre_classes v) (pre_class v s). Lemma inh_ind_fam_impl : forall {X S} (v:X->S), inh_ind_var v -> Inhabited (pre_classes v). intros X S v h1. destruct h1 as [s h1]. apply Inhabited_intro with (pre_class v s). constructor; auto. Qed. Lemma in_pre_classes : forall {S X} (v:X->S) x, Inn (pre_classes v) (pre_class v (v x)). intros; constructor; constructor. Qed. Lemma in_pre_impl : forall {X S} {v:X->S} {s:S}, in_pre s v -> Inn (pre_classes v) (pre_class v s). intros ? ? ? ? h1; destruct h1; constructor; constructor. Qed. Inductive pre_app_aux {X S} (v:X->S) : forall A:Ensemble X, Inn (pre_classes v) A -> S -> Prop := pre_app_aux_intro : forall s (pfip:in_pre s v), pre_app_aux v (pre_class v s) (pre_classes_intro v s pfip) s. Lemma pre_app_ex : forall {X S} (v:X->S) A (pfi:Inn (pre_classes v) A), exists! s:S, pre_app_aux v A pfi s /\ A = pre_class v s. intros X S v A h1. inversion h1; subst. exists s. red; split. split; auto. pose proof (proof_irrelevance _ h1 (pre_classes_intro v s H)); subst. constructor. intros s' h2. destruct h2 as [h2 h3]. apply pre_class_inj in h3; auto. Qed. Definition pre_app {X S} (v:X->S) A (pfi:Inn (pre_classes v) A) : S := proj1_sig (constructive_definite_description _ (pre_app_ex v A pfi)). Lemma pre_app_compat : forall {X S} (v:X->S) A (pfi:Inn (pre_classes v) A), let s := pre_app v A pfi in pre_app_aux v A pfi s /\ A = pre_class v s. unfold pre_app; intros; destruct constructive_definite_description; auto. Qed. Inductive pre_app_inh_aux {S X} {U:S->Type} (v:X->S) (a:forall x:X, U (v x)) : forall s (pfip:in_pre_aux v s), U s -> Prop := pre_app_inh_aux_intro : forall x, pre_app_inh_aux v a (v x) (in_pre_intro v x) (a x). End PreImage. Section SetSortedSegFun. (*This is a recursive predicate that has some advantages and disadvandages over the simpler [seg_fun]. [seg_fun] is great when you want simplicity and power, and want easy syntax, but bad if you do a lot of parameterizations with indices or inequality proofs, cause then you get a semantically cumbersome syntax. Those invariably necessitate [transfer] functions, which in general are a nuisance, and can be obviated if the terms in the defining equality are in a type in [Set]. Then we can compute the predicates in [bool] (rather than have them be in [Prop]). So, using an inductive predicate will be an important technique to handle the [destructing] of [sseg_funs] ([Set]-seg_funs]. When we destruct terms, the process should be as painless as pattern-matching.*) (*The trade-off of using this over [seg_fun] is less computational and metaphysical purity -- which is why this section only indicates how this would be gone about.*) (*[sseg_fun offset n f offset] assures us that f has only [n] intended values (from 0 to [pred n]); the lingering values in the function are the same as the last value (at [pred n]). It also ensures us in the 3rd constructor that it was constructed from a previous function, with the adding of a value at [offset], and [n] shifted up once.*) Inductive sseg_fun_aux {U} : (nat->U) -> nat -> nat -> Prop := sseg_fun0 : forall (u:U), sseg_fun_aux (fun _ => u) 0 0 | sseg_fun1 : forall (u:U), sseg_fun_aux (fun _ => u) 1 0 | sseg_fun_S : forall (n offset offset':nat) (f:nat->U) (u:U), sseg_fun_aux f n offset -> sseg_fun_aux (fun i => if (lt_nat i (S n)) then if (lt_nat i offset') then f i else if (beq_nat i offset') then u else f (S i) else f n) (S n) (min offset' n). Inductive sseg_fun {U} (f:nat->U) n : Prop := sseg_fun_intro : forall offset, sseg_fun_aux f n offset -> sseg_fun f n. Inductive sseg_fun_remove_aux {U} : forall (f:nat->U) (n:nat) (offset:nat), sseg_fun_aux f (S n) offset -> forall (f':nat->U), sseg_fun f' n -> Prop := seg_fun_remove_aux_intro : forall f n old_offset offset (pfss:sseg_fun_aux f n old_offset) u, let pfss' := sseg_fun_S n old_offset offset f u pfss in let pfss_ := sseg_fun_intro f n _ pfss in sseg_fun_remove_aux _ n (min offset n) pfss' f pfss_. End SetSortedSegFun. Definition l_to_fun {T} (l:list T) (pfl:l <> nil) : nat -> T := fun i => nth i l (lst l pfl). Inductive segf {T} : (nat -> T) -> Prop := | segf_intro : forall (l:list T) (pfl:l <> nil), segf (l_to_fun l pfl). Lemma in_seg_fun_list_iff : forall {U n} (f:seg_fun n U) y, in_seg_fun y f <-> In y (seg_fun_to_list f). intros; split; intro h1. rewrite in_seg_fun_iff in h1. destruct h1 as [? [h1 ?]]; subst. apply in_map_seg. rewrite in_map_seg_iff in h1. destruct h1 as [? [? ?]]; subst. apply in_seg_fun_compat. Qed. Lemma inj_seg_fun_endo_to_S' : forall {n} (f:seg_fun (S n) nat) r s, r <= s -> f 0 (lt_0_Sn _) < s -> inj_seg_fun_endo (seg_fun_to_S f) r -> ~in_seg_fun (f 0 (lt_0_Sn _)) (seg_fun_to_S f) -> inj_seg_fun_endo f s. intros n f r s hr h1 h2 h3. inversion h2 as [? ? h4 h5]; subst. inversion h4 as [? ? h6 h7]; subst. constructor. eapply seg_fun_endo_to_S'. apply hr. assumption. assumption. intros i j hi hj h7. destruct i as [|i], j as [|j]; subst; auto. contradict h3. erewrite f_equal_dep. rewrite h7 at 1. erewrite f_equal_dep. erewrite seg_fun_to_S_compat' at 1. apply in_seg_fun_compat. reflexivity. reflexivity. contradict h3. erewrite f_equal_dep. rewrite <- h7 at 1. erewrite f_equal_dep. erewrite seg_fun_to_S_compat' at 1. apply in_seg_fun_compat. reflexivity. reflexivity. specialize (h5 i j (lt_S_n _ _ hi) (lt_S_n _ _ hj)). unfold seg_fun_to_S in h5. hfirst h5. erewrite f_equal_dep. rewrite h7 at 1. f_equal; apply proof_irrelevance. reflexivity. specialize (h5 hf); auto with arith. Grab Existential Variables. auto with arith. auto with arith. Qed. Lemma inj_seg_fun_endo_to_S_bound : forall {n} b (f:seg_fun (S n) nat), inj_seg_fun_endo f b -> ~in_seg_fun (pred b) (seg_fun_to_S f) -> inj_seg_fun_endo (seg_fun_to_S f) (pred b). intros n b f h1 h2. inversion h1 as [? ? h3 h4]; subst. constructor. apply seg_fun_endo_to_S_bound; auto. intros i j hi hj h5. specialize (h4 (S i) (S j) (lt_n_S _ _ hi) (lt_n_S _ _ hj) h5). apply S_inj in h4; auto. Qed. Lemma inj_seg_fun_endo_remove' : forall {n} (f:seg_fun n nat) r s m (pfm: m < n), r <= s -> f m pfm < s -> inj_seg_fun_endo (remove_seg_fun m f) r -> ~in_seg_fun (f m pfm) (remove_seg_fun m f) -> inj_seg_fun_endo f s. intro n. induction n as [|n ih]. intros; omega. intros f r s m hm. destruct m as [|m]. intros h1 h2 h3 h4. inversion h3 as [? ? h5 h6]; subst. constructor. eapply seg_fun_endo_to_S'. apply h1. erewrite f_equal_dep. apply h2. reflexivity. assumption. intros i j hi hj h7. destruct i as [|i], j as [|j]; auto. contradict h4. erewrite f_equal_dep. rewrite h7 at 1. erewrite f_equal_dep. erewrite seg_fun_to_S_compat'. apply in_seg_fun_compat. reflexivity. reflexivity. contradict h4. erewrite f_equal_dep. rewrite <- h7 at 1. erewrite f_equal_dep. erewrite seg_fun_to_S_compat' at 1. apply in_seg_fun_compat. reflexivity. reflexivity. specialize (h6 i j (lt_S_n _ _ hi) (lt_S_n _ _ hj)). f_equal. apply h6. simpl. unfold seg_fun_to_S. erewrite f_equal_dep. rewrite h7 at 1. f_equal; apply proof_irrelevance. reflexivity. pose proof (lt_S_n _ _ hm) as h0. intros hr h1 h2 h3. inversion h2 as [? ? h5 h6]; subst. inversion h5 as [? ? h7 h8]; subst. constructor. eapply seg_fun_endo_remove' in h5. apply h5. assumption. apply h1. intros i j hi hj hk. destruct i as [|i], j as [|j]; auto. pose proof (lt_S_n _ _ hj) as h8. destruct (lt_le_dec j m) as [h9 | h9]. pose proof h6 as h6'. specialize (h6 0 (S j) (Olt _ _ h8)). pose proof (lt_dec _ _ hj) as h10. simpl in h10. destruct h10 as [h10|]. revert h1. subst. intro h1. specialize (h6 h10). hfirst h6. rewrite remove_seg_fun_lt; auto with arith. rewrite remove_seg_fun_lt; auto. erewrite f_equal_dep. rewrite hk at 1. f_equal. apply proof_irrelevance. reflexivity. auto with arith. specialize (h6 hf); discriminate. revert h1 hk; subst; intros h1 hk. destruct (lt_le_dec j (S m)) as [h9' | h9']. omega. specialize (h6' 0 j). hfirst h6'. simpl; auto with arith. specialize (h6' hf). hfirst h6'. simpl. auto with arith. specialize (h6' hf0). hfirst h6'. rewrite remove_seg_fun_lt; auto with arith. rewrite remove_seg_fun_ge; auto with arith. erewrite f_equal_dep. rewrite hk at 1. apply f_equal; apply proof_irrelevance. reflexivity. specialize (h6' hf1). revert h1; subst; intro h1. omega. apply le_dec' in h9. destruct h9 as [h9 | h9]. pose proof h6 as h6'. specialize (h6 0 j (Olt _ _ h8) (lt_S_n _ _ hj)). hfirst h6. rewrite remove_seg_fun_lt; auto with arith. rewrite remove_seg_fun_ge. erewrite f_equal_dep. rewrite hk at 1. f_equal; apply proof_irrelevance. reflexivity. destruct j as [|j]; subst. apply le0 in h9; subst. contradict h3. rewrite in_seg_fun_iff. exists 0. ex_goal. simpl; auto with arith. exists hex. rewrite remove_seg_fun_lt; auto with arith. erewrite f_equal_dep. rewrite hk at 1. f_equal; apply proof_irrelevance. reflexivity. auto with arith. specialize (h6 hf); subst. contradict h3. rewrite in_seg_fun_iff. exists 0. ex_goal. simpl; auto with arith. exists hex. rewrite remove_seg_fun_lt; auto with arith. apply le0 in h9; subst. erewrite f_equal_dep. rewrite hk at 1. f_equal; apply proof_irrelevance. reflexivity. subst. contradict h3. erewrite f_equal_dep. rewrite <- hk at 1. rewrite in_seg_fun_iff. exists 0, (Olt _ _ h0). rewrite remove_seg_fun_lt; auto with arith. f_equal. apply proof_irrelevance. reflexivity. pose proof (lt_dec _ _ hi) as h8. simpl in h8. destruct h8 as [h8|]; revert h1 hk; subst; intros h1 hk. destruct (lt_le_dec i m) as [h10|h10]. specialize (h6 (S i) 0 h8 (Olt _ _ h0)). hfirst h6. rewrite remove_seg_fun_lt; auto with arith. rewrite remove_seg_fun_lt; auto with arith. erewrite f_equal_dep. rewrite hk at 1. f_equal; apply proof_irrelevance. reflexivity. apply h6 in hf; discriminate. apply le_dec' in h10. destruct h10 as [h10|]; subst. specialize (h6 i 0 (lt_S_n _ _ hi) (Olt _ _ h0)). hfirst h6. rewrite remove_seg_fun_ge. rewrite remove_seg_fun_lt; auto with arith. erewrite f_equal_dep. rewrite hk at 1. f_equal; apply proof_irrelevance. reflexivity. destruct i as [|i]. apply le0 in h10; subst. contradict h3. erewrite f_equal_dep. rewrite hk at 1. rewrite in_seg_fun_iff. exists 0, h0. rewrite remove_seg_fun_lt; f_equal. apply proof_irrelevance. auto with arith. reflexivity. simpl in h10; auto with arith. specialize (h6 hf). subst. contradict h3. apply le0 in h10; subst. rewrite in_seg_fun_iff. exists 0, h0. simpl. rewrite seg_fun_from_S_compat. erewrite f_equal_dep. rewrite <- hk at 1. f_equal. apply proof_irrelevance. reflexivity. contradict h3. rewrite in_seg_fun_iff. exists 0, (Olt _ _ h0). rewrite remove_seg_fun_lt; auto with arith. erewrite f_equal_dep. rewrite <- hk at 1. f_equal; apply proof_irrelevance. reflexivity. red in h0. apply le_S_n in h0. apply le_dec' in h0. destruct h0 as [h0|]; subst. specialize (h6 i 0 (lt_n_Sn _) (lt_0_Sn _)). hfirst h6. rewrite remove_seg_fun_ge. rewrite remove_seg_fun_lt; auto with arith. erewrite f_equal_dep. rewrite hk at 1. f_equal; apply proof_irrelevance. reflexivity. destruct i as [|i]. apply le0 in h0; subst. contradict h3. rewrite in_seg_fun_iff. exists 0, lt_0_1. rewrite remove_seg_fun_lt. erewrite f_equal_dep. rewrite <- hk at 1. f_equal; apply proof_irrelevance. reflexivity. auto with arith. auto with arith. specialize (h6 hf); revert hm h3 h1; subst; intros hm h3 h1. apply le0 in h0; subst. contradict h3. rewrite in_seg_fun_iff. exists 0, lt_0_1. rewrite remove_seg_fun_lt. erewrite f_equal_dep. rewrite <- hk at 1. f_equal; apply proof_irrelevance. reflexivity. auto with arith. contradict h3. rewrite in_seg_fun_iff. exists 0, (lt_0_Sn _). rewrite remove_seg_fun_lt; auto with arith. erewrite f_equal_dep. rewrite <- hk at 1. f_equal; apply proof_irrelevance. reflexivity. specialize (ih (seg_fun_to_S f) r s m h0 hr). hfirst ih. unfold seg_fun_to_S. erewrite f_equal_dep. apply h1. reflexivity. specialize (ih hf). hfirst ih. destruct (zerop m) as [|h8]; subst. destruct n as [|n]. omega. rewrite remove_seg_fun0. apply inj_seg_fun_endo_to_S in h2. rewrite seg_fun_to_S_remove_seg_fun1 in h2; auto. erewrite remove_seg_fun_to_S_swap_pos. apply inj_seg_fun_endo_to_S. apply inj_seg_fun_endo_transfer. assumption. assumption. specialize (ih hf0). hfirst ih. contradict h3. rewrite in_seg_fun_iff in h3. rewrite in_seg_fun_iff. destruct h3 as [k [h10 h11]]. destruct (lt_le_dec k m) as [h12|h12]; subst. rewrite remove_seg_fun_lt in h11; auto. exists (S k). ex_goal. simpl; omega. exists hex. rewrite remove_seg_fun_lt; auto with arith. unfold seg_fun_to_S in h11. erewrite f_equal_dep. rewrite h11 at 1. f_equal; apply proof_irrelevance. reflexivity. rewrite remove_seg_fun_ge in h11; auto with arith. unfold seg_fun_to_S in h11. exists (S k). ex_goal. omega. exists hex. rewrite remove_seg_fun_ge; auto with arith. erewrite f_equal_dep. rewrite h11 at 1. f_equal; apply proof_irrelevance. reflexivity. specialize (ih hf1). inversion ih as [? ? h11 h12]; subst. specialize (h12 i j (lt_S_n _ _ hi) (lt_S_n _ _ hj)). hfirst h12. unfold seg_fun_to_S. erewrite f_equal_dep. rewrite hk at 1. f_equal; apply proof_irrelevance. reflexivity. specialize (h12 hf2); subst; auto. Grab Existential Variables. destruct (zerop n); revert hm h1 h3 hk; subst; auto; intros hm h1 h3 hk. omega. auto with arith. auto with arith. Qed. Lemma inv_seg_fun_lt : forall {U n} (pfdu:type_eq_dec U) (y:U) (f:seg_fun n U), in_seg_fun y f -> inv_seg_fun pfdu y f < n. intros U n. induction n as [|n ih]; simpl; auto with arith. intros; contradiction. intros hdu y; simpl. intro f. destruct (hdu (f 0 (lt_0_Sn _)) y) as [|h2]; subst; auto with arith. intro h1. apply lt_n_S; auto. apply ih. destruct h1; try contradiction; auto. Qed. Lemma inv_seg_fun_dec_lt : forall {U n} (pfdu:type_eq_dec U) (y:U) (f:seg_fun n U), in_seg_fun_dec y f -> inv_seg_fun pfdu y f < n. intros; apply inv_seg_fun_lt. apply in_seg_fun_dec_impl; auto. Qed. Lemma inv_seg_fun_eq_bound_iff : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) y, inv_seg_fun pfdu y f = n <-> ~in_seg_fun y f. intros U n. induction n as [|n ih]; simpl. intuition. intros hdu f y. destruct (hdu (f 0 (lt_0_Sn n)) y) as [h1 | h1]. intuition. split; intro h2. apply S_inj in h2. rewrite ih in h2. intro h3. destruct h3; contradiction. f_equal. rewrite ih. contradict h2. right; auto. Qed. Lemma inv_seg_fun_ge_bound_iff : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) y, inv_seg_fun pfdu y f >= n <-> ~in_seg_fun y f. intros U n hdu f y; split; intros h1. intro h2. apply (inv_seg_fun_lt hdu) in h2. omega. rewrite <- (inv_seg_fun_eq_bound_iff hdu) in h1. rewrite h1; auto with arith. Qed. Lemma inv_seg_fun_compat : forall {U n} (pfdu:type_eq_dec U) (y:U) (f:seg_fun n U) (pf:inv_seg_fun pfdu y f < n), in_seg_fun_dec y f -> f (inv_seg_fun pfdu y f) pf = y. intros U n. induction n as [|n ih]. intros; omega. intros hdu y f. specialize (ih hdu y (seg_fun_to_S f)). intro h1. simpl in h1. simpl. destruct (hdu (f 0 (lt_0_Sn n)) y) as [|h3]; subst. intros; f_equal. apply proof_irrelevance. pose proof (lt_S_n _ _ h1) as h5. intro h6. destruct h6 as [h6|h6]; subst. erewrite <- ih; auto. unfold seg_fun_to_S; f_equal; apply proof_irrelevance. contradict h3; auto. Grab Existential Variables. assumption. Qed. Lemma inv_seg_fun_compat_le : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) i (pfi:i < n), inv_seg_fun pfdu (f i pfi) f <= i. intros U n. induction n as [|n ih]; simpl; auto with arith. intros hdu f i hi. destruct i as [|i]. erewrite f_equal_dep. rewrite eq_dec_same at 1. apply le_refl. assumption. reflexivity. specialize (ih hdu (seg_fun_to_S f) i (lt_S_n _ _ hi)). lr_if_goal; fold hlr; destruct hlr as [hl|hr]; auto with arith. apply le_n_S. unfold seg_fun_to_S in ih. erewrite f_equal_dep in ih. apply ih. reflexivity. Qed. Lemma inv_seg_fun_eq_iff : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) y k, inv_seg_fun pfdu y f = k <-> (k = n /\ ~in_seg_fun y f) \/ (exists (pfk:k < n), f k pfk = y /\ (forall j (pfj:j < n), j < k -> f j pfj <> y)). intros U n hdu f y k. destruct (in_seg_fun_dec' hdu f y) as [h0 | h0]. split; intro h2; subst. destruct (lt_le_dec (inv_seg_fun hdu y f) n) as [h2 | h2]. right. exists h2; split. apply inv_seg_fun_compat; auto. apply in_seg_fun_impl_dec; auto. intros j hj. pose proof (inv_seg_fun_le hdu f j hj) as h4. intro h5. intro; subst. omega. apply le_ge in h2. rewrite inv_seg_fun_ge_bound_iff in h2. contradiction. destruct h2 as [h2|h2]. destruct h2; contradiction. destruct h2 as [h2 [? h4]]; subst. rewrite inv_seg_fun_eq_same_iff; auto. split; intro h1; subst. left. split; auto. rewrite inv_seg_fun_eq_bound_iff; auto. destruct h1 as [h1 | h1]. destruct h1; subst. rewrite inv_seg_fun_eq_bound_iff; auto. destruct h1 as [h1 [? h2]]; subst. rewrite inv_seg_fun_eq_same_iff; auto. Qed. Lemma inv_seg_fun_eq0_iff : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) y (pf0:0 < n), inv_seg_fun pfdu y f = 0 <-> y = f 0 pf0. intros; rewrite inv_seg_fun_eq_iff. split; intro h1. destruct h1 as [h1|h1]. destruct h1; subst. omega. destruct h1 as [? [h1 h2]]; subst. f_equal; apply proof_irrelevance. right. subst. exists pf0. split; auto. intros; omega. Qed. Lemma inv_seg_fun_to_S_eq : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun (S n) U) y, in_seg_fun_dec y (seg_fun_to_S f) -> y <> f 0 (lt_0_Sn _) -> inv_seg_fun pfdu y (seg_fun_to_S f) = pred (inv_seg_fun pfdu y f). intros U n hdu f y h2 hne. pose proof hne as hn. rewrite <- (inv_seg_fun_eq0_iff hdu) in hne. rewrite neq0_iff in hne. rewrite inv_seg_fun_eq_iff. right. ex_goal. pose proof (inv_seg_fun_lt hdu y (seg_fun_to_S f) (in_seg_fun_dec_impl _ _ h2)) as h3. rewrite (S_pred _ _ hne). simpl. lr_if_goal; fold hlr; destruct hlr as [hl | hr]; subst. contradict hn; auto. simpl. apply inv_seg_fun_lt; auto. apply in_seg_fun_dec_impl; auto. exists hex. split. unfold seg_fun_to_S. gen0. rewrite <- (S_pred _ _ hne). intro h3. apply inv_seg_fun_compat. simpl. left; auto. intros j hj h3. unfold seg_fun_to_S. intro; subst. pose proof (inv_seg_fun_weak_bound hdu f (S j) (lt_n_S _ _ hj) (f (S j) (lt_n_S _ _ hj)) (eq_refl _)) as h5. omega. Qed. Lemma in_seg_fun_dec_lub_weak : forall {n} (f:seg_fun n nat), 0 < n -> in_seg_fun_dec (seg_fun_lub_weak f) f. intro n. induction n as [|n ih]. intros; omega. intros f h1. destruct (zerop n) as [|h2]; subst; simpl. right. destruct (zerop (f 0 (lt_0_Sn _))) as [h2 | h2]. rewrite h2; auto. rewrite (S_pred _ _ h2); auto. specialize (ih (seg_fun_to_S f) h2). destruct (lt_le_dec (seg_fun_lub_weak (seg_fun_to_S f)) (f 0 (lt_0_Sn n))) as [h3 | h3]. apply lt_not_le in h3. apply nle_nge in h3. rewrite <- ge_nat_compat in h3. rewrite <- eq_false_iff in h3. rewrite h3; right; auto. apply le_ge in h3. rewrite <- ge_nat_compat in h3. rewrite h3. left; auto. Qed. Lemma in_seg_fun_lub_weak : forall {n} (f:seg_fun n nat), 0 < n -> in_seg_fun (seg_fun_lub_weak f) f. intros; apply in_seg_fun_dec_impl; apply in_seg_fun_dec_lub_weak; auto. Qed. Lemma in_remove_seg_fun_dec : forall {U n} (pfdu:type_eq_dec U) m (f:seg_fun n U) y, in_seg_fun_dec y (remove_seg_fun m f) -> in_seg_fun_dec y f. intros U n hdu m f y h1; apply in_seg_fun_impl_dec; apply in_seg_fun_dec_impl in h1; auto. rewrite in_seg_fun_iff in h1. destruct h1 as [i [h1 ?]]; subst. rewrite in_seg_fun_iff. destruct (lt_le_dec i m) as [h2 | h2]. exists i, (lt_pred' _ _ h1). rewrite remove_seg_fun_lt; auto. exists (S i), (lt_pred_S _ _ h1). rewrite remove_seg_fun_ge; auto. Qed. Lemma in_remove_seg_fun : forall {U n} m (f:seg_fun n U) y, in_seg_fun y (remove_seg_fun m f) -> in_seg_fun y f. intros U n m f y h1. rewrite in_seg_fun_iff in h1. destruct h1 as [i [h1 ?]]; subst. rewrite in_seg_fun_iff. destruct (lt_le_dec i m) as [h2 | h2]. exists i, (lt_pred' _ _ h1). rewrite remove_seg_fun_lt; auto. exists (S i), (lt_pred_S _ _ h1). rewrite remove_seg_fun_ge; auto. Qed. (*Same as [remove_seg_fun], only you specify an element in the image.*) (*Only of course removes the first instance of [y]*) Fixpoint remove_im_seg_fun {U n} (pfdu:type_eq_dec U) (y:U) : forall (f:seg_fun n U), in_seg_fun_dec y f -> seg_fun (pred n) U := match n with | 0 => fun f pfin => False_rect _ pfin | S n' => fun f pfin => let f' := seg_fun_to_S f in match pfin with | inleft pfin' => fun i => match i with | 0 => fun pf0 => if (pfdu (f 0 (lt_0_Sn _)) y) then f' 0 pf0 else f 0 (lt_0_Sn _) | S i' => fun pfi => let pfi' := lt_S_pred _ _ pfi in let pfi'' := lt_n_S _ _ pfi in if (pfdu (f 0 (lt_0_Sn _)) y) then f (S (S i')) pfi'' else remove_im_seg_fun pfdu y f' pfin' i' pfi' end | inright pfe => f' end end. Lemma remove_im_seg_fun_lt : forall {U n} (pfdu:type_eq_dec U) (y:U) (f:seg_fun n U) (pfin:in_seg_fun_dec y f) i (pfi:i < pred n) (pfi':i < n), i < inv_seg_fun pfdu y f -> remove_im_seg_fun pfdu y f pfin i pfi = f i pfi'. intros U n. induction n as [|n ih]; simpl. intros; omega. intros hdu y f hin i hi hi'. lr_if_goal; fold hlr; destruct hlr as [hl | hr]; subst. intros; omega. destruct hin as [h1 | h1]. destruct i as [|i]. intro. f_equal; apply proof_irrelevance. intro h2. apply lt_S_n in h2. specialize (ih hdu y (seg_fun_to_S f) h1 _ (lt_S_pred _ _ hi) (lt_S_n _ _ hi') h2). rewrite ih. unfold seg_fun_to_S. f_equal; apply proof_irrelevance. contradiction. Qed. Lemma remove_im_seg_fun_ge : forall {U n} (pfdu:type_eq_dec U) (y:U) (f:seg_fun n U) (pfin:in_seg_fun_dec y f) i (pfi:i < pred n), i >= inv_seg_fun pfdu y f -> let pfi' := lt_pred_S _ _ pfi in remove_im_seg_fun pfdu y f pfin i pfi = f (S i) pfi'. intros U n. induction n as [|n ih]; simpl. intros; omega. intros hdu y f hin i hi. intro h2. destruct hin as [h1 | h1]. destruct i as [|i]. lr_if h2; fold hlr in h2; destruct hlr as [hl | hr]; subst. rewrite eq_dec_same at 1. unfold seg_fun_to_S. f_equal; apply proof_irrelevance. assumption. omega. lr_if h2; fold hlr in h2; destruct hlr as [hl | hr]; subst. rewrite eq_dec_same at 1. f_equal. apply proof_irrelevance. assumption. apply le_S_n in h2. lr_if_goal; fold hlr; destruct hlr as [hl' | hr']; subst. f_equal. apply proof_irrelevance. rewrite ih; auto. unfold seg_fun_to_S. f_equal; apply proof_irrelevance. unfold seg_fun_to_S. f_equal; apply proof_irrelevance. Qed. Lemma remove_im_seg_fun_eq_same_iff : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun (S n) U) y i (pfi:i < n) (pfin:in_seg_fun_dec y f), remove_im_seg_fun pfdu y f pfin i pfi = y <-> (i >= inv_seg_fun pfdu y f /\ f (S i) (lt_n_S _ _ pfi) = y). intros U n. induction n as [|n ih]; simpl. intros; omega. intros hdu f y i hi. destruct i as [|i]; simpl. intro hin. destruct hin as [[h2|h2]|n2]. lr_if_goal; fold hlr; destruct hlr as [hl|hr]. unfold seg_fun_to_S; intuition. split; intro h3; subst. contradict hr; auto. omega. subst. lr_if_goal; fold hlr; destruct hlr as [hl|hr]. intuition. split; intro;[contradiction|omega]. subst. rewrite eq_dec_same at 1. unfold seg_fun_to_S. intuition. assumption. specialize (ih hdu (seg_fun_to_S f) y i (lt_S_n _ _ hi)). intro hin. destruct hin as [h2|h2]. specialize (ih h2). lr_if_goal; fold hlr; destruct hlr as [hl|hr]. intuition. simpl in ih. destruct h2 as [h2|h2], i as [|i]. lr_if_goal; fold hlr; fold hlr in ih; destruct hlr as [hl'|hr']; subst. split; intro h3. unfold seg_fun_to_S in hr, h3. split; auto with arith. unfold seg_fun_to_S. rewrite <- h3. f_equal; apply proof_irrelevance. destruct h3 as [h3 h4]. unfold seg_fun_to_S in h4; unfold seg_fun_to_S. rewrite <- h4. f_equal; apply proof_irrelevance. split; intro; subst. pose proof (iff1 ih (eq_refl _)); omega. destruct H as [h3 h4]; subst. omega. lr_if_goal; fold hlr; fold hlr in ih; destruct hlr as [hl'|hr']; subst. erewrite f_equal_dep. rewrite ih at 1. split; intro h3; destruct h3 as [h3 h4]; split; auto with arith. rewrite <- h4. unfold seg_fun_to_S; f_equal; apply proof_irrelevance. unfold seg_fun_to_S in h4; unfold seg_fun_to_S. rewrite <- h4; f_equal; apply proof_irrelevance. reflexivity. erewrite f_equal_dep. rewrite ih at 1. split; intro h3; destruct h3 as [h3 h4]; subst. split. apply le_n_S; auto. unfold seg_fun_to_S; f_equal; apply proof_irrelevance. split. eapply lt_S_n in h3; auto. unfold seg_fun_to_S; f_equal; apply proof_irrelevance. reflexivity. subst. rewrite eq_dec_same in ih at 1. rewrite eq_dec_same at 1. simpl in ih. erewrite f_equal_dep. rewrite ih at 1. split; intro h1; destruct h1 as [h1 h2]; unfold seg_fun_to_S in h2; unfold seg_fun_to_S. split; auto with arith. rewrite <- h2; f_equal; apply proof_irrelevance. split; auto. rewrite <- h2; f_equal; apply proof_irrelevance. reflexivity. assumption. assumption. subst. rewrite eq_dec_same in ih at 1. rewrite eq_dec_same at 1; simpl in ih. erewrite f_equal_dep. rewrite ih at 1. split; intro h1; destruct h1 as [h1 h2]; unfold seg_fun_to_S in h2; unfold seg_fun_to_S. split; auto with arith. rewrite <- h2; f_equal; apply proof_irrelevance. split; auto with arith. rewrite <- h2; f_equal; apply proof_irrelevance. reflexivity. assumption. assumption. subst. rewrite eq_dec_same at 1. split; intro h2; destruct h2 as [h2]. split; auto with arith. unfold seg_fun_to_S; auto. assumption. Qed. Lemma remove_im_seg_fun_in_nin : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) y b (pfin:in_seg_fun_dec y f), in_seg_fun b f -> ~in_seg_fun b (remove_im_seg_fun pfdu y f pfin) -> b = y. intros U n hdu f y b hin hin' h2. rewrite in_seg_fun_iff in hin'. destruct hin' as[i [hi ?]]; subst. pose proof (in_seg_fun_dec_impl _ _ hin) as hj. rewrite in_seg_fun_iff in hj. destruct hj as [j [hj ?]]; subst. destruct i as [|i]. destruct (zerop (inv_seg_fun hdu (f j hj) f)) as [h5|h5]. rewrite inv_seg_fun_eq0_iff in h5. rewrite <- h5 at 1. reflexivity. destruct (zerop j) as [|h3]; subst. f_equal; apply proof_irrelevance. destruct (hdu (f 0 hi) (f j hj)) as [h4|h4]; auto. contradict h2. rewrite in_seg_fun_iff. exists 0. ex_goal. omega. exists hex. erewrite remove_im_seg_fun_lt; auto. destruct (zerop (inv_seg_fun hdu (f (S i) hi) f)) as [h5|h5]. rewrite inv_seg_fun_eq0_iff in h5. destruct j as [|j]. rewrite h5. f_equal; apply proof_irrelevance. rewrite h5 in h2. destruct (hdu (f (S i) hi) (f (S j) hj)) as [|h6]; auto. rewrite h5 in h6. contradict h2. rewrite in_seg_fun_iff. exists 0. ex_goal. destruct n as [|n]. omega. simpl; omega. exists hex. erewrite remove_im_seg_fun_lt. reflexivity. destruct (zerop (inv_seg_fun hdu (f (S j) hj) f)) as [h7|]; auto. rewrite inv_seg_fun_eq0_iff in h7. rewrite <- h7 in h6 at 1. contradict h6; auto. destruct (hdu (f (S i) hi) (f j hj)) as [|h6]; auto. contradict h2. rewrite in_seg_fun_iff. destruct (lt_le_dec (S i) (inv_seg_fun hdu (f j hj) f)) as [h7|h7]. exists (S i). ex_goal. pose proof (inv_seg_fun_lt hdu (f j hj) f) as h8. hfirst h8. apply in_seg_fun_compat. specialize (h8 hf). omega. exists hex. erewrite remove_im_seg_fun_lt; auto. apply le_dec' in h7. destruct h7 as [h7|h7]. exists i. ex_goal. omega. exists hex. erewrite remove_im_seg_fun_ge; auto. f_equal; apply proof_irrelevance. contradict h6. gen0. rewrite <- h7. intro. rewrite inv_seg_fun_compat; auto. Grab Existential Variables. omega. Qed. Lemma inj_seg_fun_eq_iff_inv : forall {U n} (pfdu:type_eq_dec U) (f:seg_fun n U) y , inj_seg f -> forall i (pf:i < n), f i pf = y <-> (exists _:in_seg_fun_dec y f, i = inv_seg_fun pfdu y f). intros U n hdu f y h1 i h2; split; intro h4; subst. exists (in_seg_fun_dec_compat _ _ _ ). pose proof (inv_seg_fun_le hdu f i h2) as h3. apply le_lt_eq_dec in h3. destruct h3 as [h3 | h3]; auto. apply lt_neq in h3. rewrite (inv_seg_fun_eq_same_iff _ _) in h3. contradict h3. intros j hj h3 h4. apply h1 in h4; subst; contradict (lt_irrefl _ hj). destruct h4 as [h4]; subst. apply in_seg_fun_dec_impl in h4. rewrite in_seg_fun_iff in h4. destruct h4 as [i [h4 ?]]; subst. rewrite inv_seg_fun_compat; auto. apply in_seg_fun_dec_compat. Qed. Lemma inj_seg_fun_count_compat : forall {n} (f:seg_fun n nat) r, inj_seg_fun_endo f r -> forall y, in_seg_fun_dec y f -> count_seg_fun_im_nat y f = 1. intro n. induction n as [|n ih]; simpl. intros; contradiction. intros f r h1 y h2. destruct (nat_eq_dec y (f 0 (lt_0_Sn n))) as [|h3]; subst. rewrite beq_nat_same. f_equal. apply le0. apply not_lt_le. apply ngt_nlt. rewrite count_seg_fun_im_in_nat_iff. intro h3. rewrite in_seg_fun_iff in h3. destruct h3 as [i [h3 h4]]. unfold seg_fun_to_S in h4. inversion h1 as [? ? h5 h6]; subst. apply h6 in h4. discriminate. pose proof h3 as h3'. erewrite <- neq_nat_compat in h3. rewrite h3. eapply ih. apply inj_seg_fun_endo_to_S in h1. apply h1. destruct h2; subst; auto. contradict h3'; auto. Qed. Lemma lt_inv_seg_fun_lub_weak : forall {n} (f:seg_fun n nat) (pf0:0 < n), inv_seg_fun_lub_weak f pf0 < n. unfold inv_seg_fun_lub_weak; intros; apply inv_seg_fun_dec_lt; apply in_seg_fun_dec_lub_weak; auto. Qed. Lemma inv_seg_fun_lub_weak_compat : forall {n} (f:seg_fun n nat) (pf0:0 < n) (pf:inv_seg_fun_lub_weak f pf0 < n), f (inv_seg_fun_lub_weak f pf0) pf = seg_fun_lub_weak f. unfold inv_seg_fun_lub_weak; intros n f h0 h1. pose proof (inv_seg_fun_compat nat_eq_dec (seg_fun_lub_weak f) f h1) as h2. rewrite <- h2. apply f_equal_dep; f_equal. apply in_seg_fun_dec_lub_weak; auto. Qed. Lemma le_seg_fun_lub_weak_from_S : forall {n} (f:seg_fun (S n) nat), seg_fun_lub_weak (seg_fun_from_S f) <= seg_fun_lub_weak f. intros n f. destruct (zerop n) as [|hn]; subst. simpl; auto with arith. pose proof (seg_fun_lub_compat f (lt_0_Sn _)) as h6. destruct h6 as [h6 h7]. specialize (h6 (inv_seg_fun_lub_weak (seg_fun_from_S f) hn)). hfirst h6. apply lt_S. apply lt_inv_seg_fun_lub_weak. specialize (h6 hf). pose proof (inv_seg_fun_lub_weak_compat (seg_fun_from_S f) hn (lt_inv_seg_fun_lub_weak _ _)) as h8. rewrite <- h8. rewrite seg_fun_from_S_compat. unfold seg_fun_lub in h6. apply lt_dec in h6. destruct h6 as [h6 | h6]. erewrite f_equal_dep. pose proof (lt_eq_trans _ _ _ h6 (eq_refl (seg_fun_lub_weak f))) as h9. eapply lt_le_weak. apply h9. reflexivity. erewrite f_equal_dep. rewrite h6 at 1. auto. reflexivity. Qed. Lemma le_seg_fun_lub_weak_remove : forall {n} (f:seg_fun n nat) i, seg_fun_lub_weak (remove_seg_fun i f) <= seg_fun_lub_weak f. intros n f i. destruct (zerop n) as [|h1]; subst. simpl; auto. pose proof (seg_fun_lub_compat f h1) as h2. destruct h2 as [h2 h3]. destruct (zerop (pred n)) as [h4|h4]. apply pred0_dec in h4. destruct h4; subst; auto with arith. specialize (h2 (inv_seg_fun nat_eq_dec (seg_fun_lub_weak (remove_seg_fun i f)) f)). hfirst h2. apply inv_seg_fun_lt. pose proof (in_seg_fun_lub_weak (remove_seg_fun i f) h4) as h5. apply in_remove_seg_fun in h5; auto. specialize (h2 hf). rewrite inv_seg_fun_compat in h2. unfold seg_fun_lub in h2; auto with arith. pose proof (in_seg_fun_lub_weak (remove_seg_fun i f) h4) as h5. apply in_remove_seg_fun in h5; auto. apply in_seg_fun_impl_dec; auto. apply nat_eq_dec. Qed. Lemma seg_fun_endo_remove_im : forall {n i m} (f:seg_fun n nat), i < n -> seg_fun_endo f m -> forall (pfi:in_seg_fun_dec i f), seg_fun_endo (remove_im_seg_fun nat_eq_dec i f pfi) m. intros n i m f h1 h2 h3. inversion h2; subst. constructor. intros j h4. destruct (lt_le_dec j (inv_seg_fun nat_eq_dec i f)) as [h5 | h5]. erewrite remove_im_seg_fun_lt; auto. erewrite remove_im_seg_fun_ge; auto. Grab Existential Variables. omega. Qed. Lemma seg_fun_endo_remove_im' : forall {n} (f:seg_fun n nat) a b c (pfi:in_seg_fun_dec a f), b <= c -> a < c -> seg_fun_endo (remove_im_seg_fun nat_eq_dec a f pfi) b -> seg_fun_endo f c. intros n f a b c hin h2 h3 h4. inversion h4 as [? ? h5 h6]; subst. destruct (zerop n) as [|hz]; subst. constructor. intros i hi. contradict (lt_n_0 _ hi). constructor. intros i hi. destruct (lt_dec _ _ hi) as [h6 | h6]. destruct (lt_le_dec i (inv_seg_fun nat_eq_dec a f)) as [h7 | h7]. specialize (h5 i h6). erewrite remove_im_seg_fun_lt in h5; auto. eapply lt_le_trans. apply h5. assumption. apply le_lt_eq_dec in h7. destruct h7 as [h7|h7]. destruct i as [|i]. omega. apply lt_le in h7. specialize (h5 i (lt_trans _ _ _ (lt_n_Sn _) h6)). rewrite remove_im_seg_fun_ge in h5; auto. eapply lt_le_trans. erewrite f_equal_dep. apply h5. reflexivity. assumption. subst. rewrite inv_seg_fun_compat; auto. destruct (lt_le_dec i (inv_seg_fun nat_eq_dec a f)) as [h7 | h7]. subst. pose proof (inv_seg_fun_lt nat_eq_dec a f) as h8. hfirst h8. apply in_seg_fun_dec_impl; auto. specialize (h8 hf). omega. subst. destruct (zerop (pred n)) as [h6|h6]. apply pred0_dec in h6; destruct h6; subst. contradict (lt_irrefl _ hi). simpl. simpl in h7. apply le0 in h7. lr_if h7; fold hlr in h7; destruct hlr as [hl | hr]. subst. erewrite f_equal_dep. apply h3. reflexivity. discriminate. specialize (h5 (pred (pred n))). hfirst h5. eapply lt_pred_n. apply h6. apply le_dec' in h7. destruct h7 as [h7 | h7]. specialize (h5 hf). rewrite remove_im_seg_fun_ge in h5; auto. generalize hi. rewrite (S_pred _ _ h6). intro h8. erewrite f_equal_dep. eapply lt_le_trans. apply h5. assumption. reflexivity. rewrite inv_seg_fun_eq_iff in h7. destruct h7 as [h7 | h7]. destruct h7 as [h7]. contradict (lt_irrefl _ (eq_lt_trans _ _ _ (eq_sym h7) hi)). destruct h7 as [h7 [h8 ?]]; subst. erewrite f_equal_dep. apply h3. reflexivity. Qed. Lemma inj_seg_fun_endo_remove_im : forall {n i m} (f:seg_fun n nat), i < n -> inj_seg_fun_endo f m -> forall (pfi:in_seg_fun_dec i f), inj_seg_fun_endo (remove_im_seg_fun nat_eq_dec i f pfi) m. intros n i m f h1 h2 h3. inversion h2; subst. constructor. apply seg_fun_endo_remove_im; auto. intros r s hr hs h4. destruct (lt_le_dec r (inv_seg_fun nat_eq_dec i f)) as [h5 | h5], (lt_le_dec s (inv_seg_fun nat_eq_dec i f)) as [h6 | h6]. pose proof @remove_im_seg_fun_lt. rewrite (remove_im_seg_fun_lt nat_eq_dec i f h3 r hr (lt_pred' _ _ hr)), (remove_im_seg_fun_lt nat_eq_dec i f h3 s hs (lt_pred' _ _ hs)) in h4. apply H0 in h4; auto. rewrite (remove_im_seg_fun_lt nat_eq_dec i f h3 s hs (lt_pred' _ _ hs)) in h4; auto. auto. rewrite (remove_im_seg_fun_lt nat_eq_dec i f h3 r hr (lt_pred' _ _ hr)), (remove_im_seg_fun_ge nat_eq_dec i f h3 s hs) in h4; auto. apply H0 in h4; subst; omega. rewrite (remove_im_seg_fun_ge nat_eq_dec i f h3 r), (remove_im_seg_fun_lt nat_eq_dec i f h3 s hs (lt_pred' _ _ hs)) in h4; auto. apply H0 in h4; subst; omega. rewrite (remove_im_seg_fun_ge nat_eq_dec i f h3 s hs), (remove_im_seg_fun_ge nat_eq_dec i f h3 r hr) in h4; auto. apply H0 in h4. apply S_inj in h4; auto. Qed. Lemma inj_seg_fun_endo_remove_im' : forall {n} (f:seg_fun n nat) a b c (pfin:in_seg_fun_dec a f), ~in_seg_fun a (remove_im_seg_fun nat_eq_dec a f pfin) -> b <= c -> a < c -> inj_seg_fun_endo (remove_im_seg_fun nat_eq_dec a f pfin) b -> inj_seg_fun_endo f c. intros n f a b c hin hnin hb ha h1. inversion h1 as [? ? h2 h3]; subst. constructor. eapply seg_fun_endo_remove_im'. apply hb. apply ha. apply h2. destruct (zerop (pred n)) as [h4|h4]; subst. apply pred0_dec in h4; destruct h4; subst. intros; omega. intros; omega. pose proof (in_seg_fun_dec_impl _ _ hin) as hin'. intros i j hi hj h5. destruct (lt_dec _ _ hi) as [h6|h6], (lt_dec _ _ hj) as [h7|h7]. destruct (lt_le_dec i (inv_seg_fun nat_eq_dec a f)) as [h8|h8], (lt_le_dec j (inv_seg_fun nat_eq_dec a f)) as [h9|h9]. specialize (h3 _ _ h6 h7). apply h3. erewrite remove_im_seg_fun_lt at 1. erewrite remove_im_seg_fun_lt at 1. apply h5. assumption. assumption. destruct (zerop j) as [|hz]; subst. apply le0 in h9. omega. apply le_dec' in h9. destruct h9 as [h9|h9]. specialize (h3 i (pred j) h6). hfirst h3. apply lt_S_n. rewrite <- (S_pred _ _ hz). rewrite <- (S_pred _ _ (Olt _ _ hi)); auto. specialize (h3 hf). hfirst h3. erewrite remove_im_seg_fun_lt; auto. erewrite remove_im_seg_fun_ge; auto. symmetry. gen0. rewrite <- (S_pred _ _ hz). intro h10. erewrite h5 at 1. f_equal; apply proof_irrelevance. apply h3 in hf0; subst. pose proof (lt_le_trans _ _ _ h8 h9) as h10. apply lt_irrefl in h10; contradiction. subst. rewrite inv_seg_fun_compat in h5; auto; subst. contradict hnin. rewrite in_seg_fun_iff. exists i, h6. erewrite remove_im_seg_fun_lt; auto. destruct (zerop i) as [|hz]; subst. apply le0 in h8. omega. apply le_dec' in h8. destruct h8 as [h8|h8]. specialize (h3 (pred i) j). hfirst h3. eapply lt_trans. apply lt_pred. rewrite <- (S_pred _ _ hz). apply h6. eapply lt_pred_n. apply h6. specialize (h3 hf h7). hfirst h3. erewrite remove_im_seg_fun_ge; auto. erewrite remove_im_seg_fun_lt; auto. gen0. rewrite <- (S_pred _ _ hz). intro. erewrite <- h5 at 1. f_equal; apply proof_irrelevance. specialize (h3 hf0); subst. pose proof (lt_le_trans _ _ _ h9 h8) as h10. apply lt_irrefl in h10; contradiction. subst. rewrite inv_seg_fun_compat in h5; subst; auto. pose proof (inv_seg_fun_weak_bound nat_eq_dec f j hj _ (eq_refl _)) as h10. pose proof (lt_le_trans _ _ _ h9 h10) as h11. contradict (lt_irrefl _ h11). destruct (zerop i) as [|h10], (zerop j) as [|h11]; subst; auto. apply le0 in h8. rewrite inv_seg_fun_eq0_iff in h8. rewrite h5 in h8 at 1. subst. rewrite <- h5 in hnin at 1. contradict hnin. rewrite in_seg_fun_iff. apply le_dec' in h9. destruct h9 as [h9|h9]. exists (pred j). ex_goal. omega. exists hex. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ h11). intro. rewrite h5. f_equal; apply proof_irrelevance. rewrite inv_seg_fun_eq_same_iff in h9. specialize (h9 0 h11 hi). rewrite h5 in h9. contradict h9; auto. apply le0 in h9. rewrite inv_seg_fun_eq0_iff in h9. rewrite <- h5 in h9 at 1. subst. apply le_dec' in h8. destruct h8 as [h8|h8]. contradict hnin. rewrite in_seg_fun_iff. exists (pred i). ex_goal. omega. exists hex. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ h10). intros; f_equal. apply proof_irrelevance. rewrite inv_seg_fun_eq_same_iff in h8. specialize (h8 0 h10 hj). rewrite <- h5 in h8. contradict h8; auto. apply le_dec' in h8. destruct h8 as [h8|h8]. apply le_dec' in h9. destruct h9 as [h9|h9]. specialize (h3 (pred i) (pred j)). hfirst h3. eapply lt_trans. eapply lt_pred_n. apply h10. assumption. specialize (h3 hf). hfirst h3. eapply lt_trans. eapply lt_pred_n. apply h11. assumption. specialize (h3 hf0). hfirst h3. rewrite remove_im_seg_fun_ge; auto. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ h10). intros; symmetry. gen0. rewrite <- (S_pred _ _ h11). intros. erewrite f_equal_dep. rewrite <- h5 at 1. f_equal; apply proof_irrelevance. reflexivity. specialize (h3 hf1). rewrite (S_pred _ _ h10). rewrite h3. symmetry. eapply S_pred. apply h11. subst. rewrite inv_seg_fun_compat in h5; subst; auto. contradict hnin. rewrite in_seg_fun_iff. exists (pred i). ex_goal. omega. exists hex. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ h10). intros; f_equal; apply proof_irrelevance. subst. rewrite inv_seg_fun_compat in h5; subst. apply le_dec' in h9. destruct h9 as [h9|]; auto. contradict hnin. rewrite in_seg_fun_iff. exists (pred j). ex_goal. omega. exists hex. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ h11). intro; f_equal; apply proof_irrelevance. assumption. subst. pose proof (inv_seg_fun_lt nat_eq_dec a f hin') as h7. apply lt_le_pred in h7. apply le_dec' in h7. destruct (lt_le_dec i (inv_seg_fun nat_eq_dec a f)) as [h8|h8]. destruct h7 as [h7|h7]. specialize (h3 i (pred (pred n)) h6). hfirst h3. eapply lt_pred_n. apply h6. specialize (h3 hf). hfirst h3. erewrite remove_im_seg_fun_lt; auto. erewrite remove_im_seg_fun_ge; auto. symmetry. gen0. rewrite <- (S_pred _ _ h4). intro. rewrite h5 at 1. f_equal; apply proof_irrelevance. specialize (h3 hf0). subst. pose proof (lt_le_trans _ _ _ h8 h7) as h9. apply lt_irrefl in h9; contradiction. rewrite inv_seg_fun_eq_iff in h7. destruct h7 as [h7 | h7]. destruct h7; contradiction. destruct h7 as [? [h7 h9]]; subst. specialize (h9 i hi h6). rewrite h5 in h9. contradict h9; f_equal; apply proof_irrelevance. destruct (zerop i) as [|hz]; subst. apply le0 in h8. rewrite inv_seg_fun_eq0_iff in h8. rewrite h5 in h8 at 1. subst. contradict hnin. rewrite in_seg_fun_iff. destruct h7 as [h7|h7]. exists (pred (pred n)). ex_goal. eapply lt_pred_n. apply h6. exists hex. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ h6). intro; f_equal; apply proof_irrelevance. rewrite inv_seg_fun_eq_iff in h7. destruct h7 as [h7|h7]. destruct h7; contradiction. destruct h7 as [h7 [h8 h9]]. specialize (h9 0 hi h6). rewrite h5 in h9. contradict h9; auto. apply le_dec' in h8. destruct h7 as [h7|h7]. destruct h8 as [h8|h8]. specialize (h3 (pred i) (pred (pred n))). hfirst h3. eapply lt_trans. apply lt_n_Sn. rewrite <- (S_pred _ _ hz). assumption. specialize (h3 hf). hfirst h3. eapply lt_pred_n. apply h6. specialize (h3 hf0). hfirst h3. rewrite remove_im_seg_fun_ge; auto. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ hz). intro. symmetry. gen0. rewrite <- (S_pred _ _ h6). intro. erewrite f_equal_dep. rewrite <- h5 at 1. f_equal; apply proof_irrelevance. reflexivity. specialize (h3 hf1). rewrite (S_pred _ _ hz). rewrite h3. rewrite <- (S_pred _ _ h6); auto. subst. rewrite inv_seg_fun_compat in h5; subst. rewrite inv_seg_fun_eq_same_iff. intros j hj' h9 h10. destruct (lt_le_dec j (inv_seg_fun nat_eq_dec (f (pred n) hj) f)) as [h11|h11]. contradict hnin. rewrite in_seg_fun_iff. exists j, hj'. erewrite remove_im_seg_fun_lt; auto. rewrite <- h10. reflexivity. specialize (h3 (pred (pred n)) (pred j)). hfirst h3. eapply lt_pred_n. apply hj'. specialize (h3 hf). hfirst h3. pose proof (inv_seg_fun_eq0_iff nat_eq_dec f (f (pred n) hj) (Olt _ _ h9)) as h12. nhl h12. rewrite neq0_iff. assumption. rewrite h12 in hl. destruct (zerop j) as [|h13]; subst; auto. rewrite (S_pred _ _ h13). eapply mono_pred_lt. rewrite <- (S_pred _ _ h13); auto. auto with arith. specialize (h3 hf0). destruct (zerop j) as [|h12]; subst. contradict hnin. rewrite in_seg_fun_iff. exists _, hf. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ hj'). intro; f_equal; apply proof_irrelevance. apply le_dec' in h11. destruct h11 as [h11|h11]. hfirst h3. erewrite remove_im_seg_fun_ge; auto. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ h4). intro; symmetry; gen0; rewrite <- (S_pred _ _ h12). intro. erewrite f_equal_dep. rewrite h10 at 1. f_equal; apply proof_irrelevance. reflexivity. specialize (h3 hf1). pose proof (f_equal S h3) as h13. rewrite <- (S_pred _ _ h4), <- (S_pred _ _ h12) in h13; subst. apply lt_irrefl in hj'; auto. subst. contradict hnin. rewrite in_seg_fun_iff. exists _, (lt_pred_n _ _ hj'). rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ h4). intro; f_equal; apply proof_irrelevance. assumption. pose proof (f_equal_dep f _ _ (inv_seg_fun_lt nat_eq_dec a f hin') hj h7) as h9. rewrite inv_seg_fun_compat in h9; subst. rewrite h7 in h8. destruct h8; omega. assumption. subst. pose proof (inv_seg_fun_lt nat_eq_dec a f hin') as h9. apply lt_le_pred in h9. apply le_dec' in h9. destruct h9 as [h9|h9]. destruct (lt_le_dec j (inv_seg_fun nat_eq_dec a f)) as [h8 | h8]. specialize (h3 j (pred (pred n)) h7). hfirst h3. eapply lt_pred_n. apply h7. specialize (h3 hf). hfirst h3. erewrite remove_im_seg_fun_lt; auto. erewrite remove_im_seg_fun_ge; auto. symmetry. gen0. rewrite <- (S_pred _ _ h7). intro. rewrite <- h5 at 1. f_equal; apply proof_irrelevance. specialize (h3 hf0). subst. pose proof (lt_le_trans _ _ _ h8 h9) as h10. contradict (lt_irrefl _ h10). destruct (zerop j) as [|h10]; subst. apply le0 in h8. rewrite inv_seg_fun_eq0_iff in h8. rewrite <- h5 in h8 at 1; subst. contradict hnin. rewrite in_seg_fun_iff. exists (pred (pred n)). ex_goal. eapply lt_pred_n. apply h7. exists hex. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ h7). intro; f_equal; apply proof_irrelevance. apply le_dec' in h8. destruct h8 as [h8|h8]. specialize (h3 (pred (pred n)) (pred j)). hfirst h3. eapply lt_pred_n. apply h4. specialize (h3 hf). hfirst h3. eapply lt_trans. eapply lt_pred_n. apply h10. assumption. specialize (h3 hf0). hfirst h3. rewrite remove_im_seg_fun_ge; auto. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ h7). intro. symmetry. gen0. rewrite <- (S_pred _ _ h10). intro. erewrite f_equal_dep. rewrite <- h5 at 1. f_equal. apply proof_irrelevance. reflexivity. specialize (h3 hf1). pose proof (f_equal S h3) as h11. rewrite <- (S_pred _ _ h4), <- (S_pred _ _ h10) in h11; auto. subst. rewrite inv_seg_fun_compat in h5; subst. symmetry. rewrite inv_seg_fun_eq_same_iff. intros j h12 h13. intro h14. contradict hnin. rewrite in_seg_fun_iff. exists (pred (pred n)). ex_goal. eapply lt_pred_n. apply h12. exists hex. rewrite remove_im_seg_fun_ge; auto. gen0. rewrite <- (S_pred _ _ h4). intro; f_equal. apply proof_irrelevance. assumption. rewrite inv_seg_fun_eq_iff in h9. destruct h9 as [h9 | h9]. destruct h9; contradiction. destruct h9 as [h9 [h10 h11]]; subst. specialize (h11 j hj h7). rewrite <- h5 in h11. contradict h11; f_equal; apply proof_irrelevance. subst; auto. Qed. bool2-v0-3/EqDec.glob0000644000175000017500000001173413575600573015171 0ustar dwyckoff76dwyckoff76DIGEST 91b6b8e27208538ea7f528017b65842e FBooleanAlgebrasIntro2.EqDec R766:789 Coq.Logic.FunctionalExtensionality <> <> lib R807:822 Coq.Logic.ProofIrrelevance <> <> lib def 837:847 <> type_eq_dec R882:882 BooleanAlgebrasIntro2.EqDec <> T var R886:886 Coq.Init.Specif <> :type_scope:'{'_x_'}'_'+'_'{'_x_'}' not R892:896 Coq.Init.Specif <> :type_scope:'{'_x_'}'_'+'_'{'_x_'}' not R903:903 Coq.Init.Specif <> :type_scope:'{'_x_'}'_'+'_'{'_x_'}' not R888:890 Coq.Init.Logic <> :type_scope:x_'='_x not R887:887 BooleanAlgebrasIntro2.EqDec <> x var R891:891 BooleanAlgebrasIntro2.EqDec <> y var R898:901 Coq.Init.Logic <> :type_scope:x_'<>'_x not R897:897 BooleanAlgebrasIntro2.EqDec <> x var R902:902 BooleanAlgebrasIntro2.EqDec <> y var prf 913:926 <> type_eq_dec_eq R967:977 BooleanAlgebrasIntro2.EqDec <> type_eq_dec def R979:979 BooleanAlgebrasIntro2.EqDec <> T var R989:991 Coq.Init.Logic <> :type_scope:x_'='_x not R987:988 BooleanAlgebrasIntro2.EqDec <> pf var R992:994 BooleanAlgebrasIntro2.EqDec <> pf' var R1158:1174 Coq.Logic.ProofIrrelevance <> proof_irrelevance prfax R1158:1174 Coq.Logic.ProofIrrelevance <> proof_irrelevance prfax R1238:1254 Coq.Logic.ProofIrrelevance <> proof_irrelevance prfax R1238:1254 Coq.Logic.ProofIrrelevance <> proof_irrelevance prfax prf 1269:1289 <> impl_type_eq_dec_prod R1360:1370 BooleanAlgebrasIntro2.EqDec <> type_eq_dec def R1374:1374 Coq.Init.Datatypes <> :type_scope:x_'*'_x not R1373:1373 BooleanAlgebrasIntro2.EqDec <> T var R1375:1375 BooleanAlgebrasIntro2.EqDec <> U var R1339:1349 BooleanAlgebrasIntro2.EqDec <> type_eq_dec def R1351:1351 BooleanAlgebrasIntro2.EqDec <> U var R1318:1328 BooleanAlgebrasIntro2.EqDec <> type_eq_dec def R1330:1330 BooleanAlgebrasIntro2.EqDec <> T var prf 1814:1829 <> match_eq_dec_sym R1857:1867 BooleanAlgebrasIntro2.EqDec <> type_eq_dec def R1869:1869 BooleanAlgebrasIntro2.EqDec <> T var R1877:1877 BooleanAlgebrasIntro2.EqDec <> T var R1896:1896 BooleanAlgebrasIntro2.EqDec <> U var R1974:1980 Coq.Init.Logic <> :type_scope:x_'='_x not R1911:1912 BooleanAlgebrasIntro2.EqDec <> pf var R1916:1916 BooleanAlgebrasIntro2.EqDec <> y var R1914:1914 BooleanAlgebrasIntro2.EqDec <> x var R1932:1935 Coq.Init.Specif <> left constr R1942:1943 BooleanAlgebrasIntro2.EqDec <> bl var R1953:1957 Coq.Init.Specif <> right constr R1964:1965 BooleanAlgebrasIntro2.EqDec <> br var R1988:1989 BooleanAlgebrasIntro2.EqDec <> pf var R1993:1993 BooleanAlgebrasIntro2.EqDec <> x var R1991:1991 BooleanAlgebrasIntro2.EqDec <> y var R2009:2012 Coq.Init.Specif <> left constr R2019:2020 BooleanAlgebrasIntro2.EqDec <> bl var R2030:2034 Coq.Init.Specif <> right constr R2041:2042 BooleanAlgebrasIntro2.EqDec <> br var def 2463:2477 <> type_dep_eq_dec R2491:2491 BooleanAlgebrasIntro2.EqDec <> T var R2524:2524 BooleanAlgebrasIntro2.EqDec <> T var R2560:2560 Coq.Init.Specif <> :type_scope:'{'_x_'}'_'+'_'{'_x_'}' not R2566:2570 Coq.Init.Specif <> :type_scope:'{'_x_'}'_'+'_'{'_x_'}' not R2577:2577 Coq.Init.Specif <> :type_scope:'{'_x_'}'_'+'_'{'_x_'}' not R2562:2564 Coq.Init.Logic <> :type_scope:x_'='_x not R2561:2561 BooleanAlgebrasIntro2.EqDec <> a var R2565:2565 BooleanAlgebrasIntro2.EqDec <> b var R2572:2575 Coq.Init.Logic <> :type_scope:x_'<>'_x not R2571:2571 BooleanAlgebrasIntro2.EqDec <> a var R2576:2576 BooleanAlgebrasIntro2.EqDec <> b var R2535:2535 BooleanAlgebrasIntro2.EqDec <> P var R2537:2537 BooleanAlgebrasIntro2.EqDec <> b var R2528:2528 BooleanAlgebrasIntro2.EqDec <> P var R2530:2530 BooleanAlgebrasIntro2.EqDec <> a var prf 2801:2822 <> type_dep_eq_dec_compat R2847:2847 BooleanAlgebrasIntro2.EqDec <> T var R2878:2892 BooleanAlgebrasIntro2.EqDec <> type_dep_eq_dec def R2894:2894 BooleanAlgebrasIntro2.EqDec <> P var R2861:2871 BooleanAlgebrasIntro2.EqDec <> type_eq_dec def R2873:2873 BooleanAlgebrasIntro2.EqDec <> T var def 2969:2983 <> dep_type_eq_dec R2992:2992 BooleanAlgebrasIntro2.EqDec <> S var R3023:3033 BooleanAlgebrasIntro2.EqDec <> type_eq_dec def R3036:3036 BooleanAlgebrasIntro2.EqDec <> T var R3038:3038 BooleanAlgebrasIntro2.EqDec <> s var prf 3050:3059 <> sig_eq_dec R3085:3095 BooleanAlgebrasIntro2.EqDec <> type_eq_dec def R3097:3097 BooleanAlgebrasIntro2.EqDec <> T var R3103:3103 BooleanAlgebrasIntro2.EqDec <> T var R3117:3127 BooleanAlgebrasIntro2.EqDec <> type_eq_dec def R3129:3129 Coq.Init.Specif <> :type_scope:'{'_x_':'_x_'|'_x_'}' not R3131:3131 Coq.Init.Specif <> :type_scope:'{'_x_':'_x_'|'_x_'}' not R3133:3135 Coq.Init.Specif <> :type_scope:'{'_x_':'_x_'|'_x_'}' not R3139:3139 Coq.Init.Specif <> :type_scope:'{'_x_':'_x_'|'_x_'}' not R3136:3136 BooleanAlgebrasIntro2.EqDec <> P var R3132:3132 BooleanAlgebrasIntro2.EqDec <> T var R3307:3323 Coq.Logic.ProofIrrelevance <> proof_irrelevance prfax R3307:3323 Coq.Logic.ProofIrrelevance <> proof_irrelevance prfax R3354:3360 Coq.Init.Logic <> f_equal thm R3364:3372 Coq.Init.Specif <> proj1_sig def R3354:3360 Coq.Init.Logic <> f_equal thm R3364:3372 Coq.Init.Specif <> proj1_sig def ax 3726:3731 <> eq_dec R3752:3762 BooleanAlgebrasIntro2.EqDec <> type_eq_dec def R3764:3764 BooleanAlgebrasIntro2.EqDec <> T var bool2-v0-3/SetUtilities1_5.v0000644000175000017500000002137713575574055016477 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2017 Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Version 8.4 pl4 Coq.*) Require Import SetUtilities. Require Import ListUtilities. Require Import EqDec. Require Import LogicUtilities. Section TypeEqDec. Section InDec. Fixpoint n_eq_dec_l {T:Type} (l:list T) (x:T) : Set := match l with | nil => unit | a::l' => (({x = a} + {x <> a}) * n_eq_dec_l l' x)%type end. Lemma type_eq_dec_n_eq_dec_l : forall {T:Type}, type_eq_dec T -> forall (l:list T) (x:T), n_eq_dec_l l x. intros T h1 l. induction l as [|a l ih]; simpl. intro; constructor. intro x. red in h1. specialize (h1 x a). apply (h1, (ih x)). Qed. Definition in_ens_dec {T:Type} (A:Ensemble T) (x:T) : Set := {Inn A x} + {~Inn A x}. Lemma list_to_set_in_ens_dec : forall {T:Type}, type_eq_dec T -> forall (A:Ensemble T) (l:list T) (x:T), list_to_set l = A-> in_ens_dec A x. intros T h0 A l. revert A. induction l as [|a l ih]; simpl. intros; subst; simpl. red. right. intro; contradiction. intros A x h2. subst. destruct (h0 x a) as [h1 | h1]; subst. red. left. right. constructor; auto. destruct (in_dec h0 a l) as [h2 | h2]. rewrite in_add_eq; auto. rewrite <- list_to_set_in_iff; auto. red. destruct (in_dec h0 x (a::l)) as [h3 | h3]. rewrite list_to_set_in_iff in h3. simpl in h3. left; auto. rewrite list_to_set_in_iff in h3. simpl in h3. right; auto. Qed. Lemma in_ens_dec_eq : forall {T:Type} (A:Ensemble T) (x:T) (pf pf':in_ens_dec A x), pf = pf'. intros T A x h2 h3. destruct h2 as [h2 | h2]. destruct h3 as [h3 | h3]; auto. pose proof (proof_irrelevance _ h2 h3); subst; auto. contradiction. destruct h3 as [h3 | h3]. contradiction. pose proof (proof_irrelevance _ h2 h3); subst; auto. Qed. Lemma in_fin_ens_dec : forall {T:Type}, type_eq_dec T -> forall (A:Ensemble T) (pf:Finite A) (x:T), in_ens_dec A x. intros T h1 A h0 x. pose proof (list_to_set_in_ens_dec h1 A) as h2. pose proof (fun_list_to_set_eq (U:=in_ens_dec) A h0) as h3. hfirst h3. intros; apply in_ens_dec_eq. specialize (h3 hf). apply h3. intros x' l h4. specialize (h2 l x' h4). assumption. Qed. Section Dependent. Lemma dep_type_n_eq_dec_l : forall {T:Type} (P:T->Prop), type_dep_eq_dec P -> forall (l:list T), p_mem_l P l -> forall (x:T) (pfx:P x), n_eq_dec_l l x. intros T P h1 l. induction l as [|a l ih]; simpl. intro; constructor. intros h0 x hx. hfirst ih. red. red in h0. intros. apply h0. right; auto. specialize (ih hf x hx). red in h0. specialize (h0 _ (in_eq _ _)). red in h1. specialize (h1 _ _ hx h0). destruct h1 as [h1 | h1]; subst. refine (_, ih). left; auto. refine (_, ih). right; auto. Qed. Definition in_ens_dep_dec {T:Type} {P:T->Prop} (A:Ensemble T) (x:T) (pfx:P x): Set := {Inn A x} + {~Inn A x}. Lemma list_to_set_in_ens_dep_dec : forall {T:Type} (P:T->Prop), type_dep_eq_dec P -> forall (A:Ensemble T) (l:list T), p_mem_l P l -> forall (x:T) (pfx:P x), list_to_set l = A-> in_ens_dep_dec A x pfx. intros T P h0 A l. revert A. induction l as [|a l ih]; simpl. intros; subst; simpl. red. right. intro; contradiction. intros A hpm x hx h2. subst. pose proof hpm as h8. red in hpm. specialize (hpm a (in_eq _ _)). pose proof (in_dep_dec h0 (a::l) h8 x hx) as h9. red. destruct h9 as [h9 | h9]. rewrite list_to_set_in_iff in h9. simpl in h9. left; auto. rewrite list_to_set_in_iff in h9. simpl in h9. right; auto. Qed. Lemma type_dep_eq_dec_in_ens_dec : forall {T:Type} (P:T->Prop), type_dep_eq_dec P -> forall (A:Ensemble T) (pff:Finite A), p_mem P A -> forall (x:T) (pfx:P x), in_ens_dep_dec A x pfx. intros T P h1 A h2 h3 x h4. pose proof (list_to_set_in_ens_dep_dec P h1 A) as h5. pose proof (fun_dep_list_to_set_eq (U:=@in_ens_dep_dec _ P) A h2) as h6. apply h6. intros; apply in_ens_dec_eq. intros x' h7 l h8; subst. apply (h5 l). rewrite p_mem_l_iff; auto. reflexivity. Qed. End Dependent. End InDec. Section EnsEqDec. Fixpoint n_fin_in_dec_l {T:Type} (pfdt:type_eq_dec T) (E:Ensemble T) (pff:Finite E) (l:list T) : Set := match l with | nil => unit | a::l' => (in_ens_dec E a * n_fin_in_dec_l pfdt E pff l')%type end. Lemma type_eq_dec_n_fin_in_dec_l : forall {T:Type} (pfdt:type_eq_dec T) (E:Ensemble T) (pff:Finite E) (l:list T), n_fin_in_dec_l pfdt E pff l. intros T hdt E h0 l. revert h0. revert E. induction l as [|a l h0]; simpl. intros; constructor. intros E h1. specialize (h0 E h1). refine (in_fin_ens_dec hdt E h1 a, h0). Qed. Definition type_dep_eq_dec_finite_l_dec {T:Type} (l:list T) := forall (E:Ensemble T) (pf:Finite E), {list_to_set l = E} + {list_to_set l <> E}. Lemma list_to_set_type_dep_eq_dec_finite_dec : forall {T:Type}, type_eq_dec T -> forall (l:list T), type_dep_eq_dec_finite_l_dec l. intros T h1 l. induction l as [|a l ih]; simpl. red; simpl. intros E h2. pose proof (finite_inh_or_empty_dec E h2) as h3. destruct h3 as [h3 | h3]. right. intro. subst. destruct h3; contradiction. subst. left; auto. red; simpl. intros E h2. destruct (in_dec h1 a l) as [h6 | h6]. red in ih. specialize (ih E h2). destruct ih as [h4 | h4]. left. rewrite in_add_eq; auto. rewrite <- list_to_set_in_iff; auto. right. rewrite in_add_eq; auto. rewrite <- list_to_set_in_iff; auto. pose proof (in_fin_ens_dec h1 E h2 a) as h7. red in h7. destruct h7 as [h7 | h7]. red in ih. specialize (ih (Subtract E a) (subtract_preserves_finite _ _ h2)). destruct ih as [h8 | h8]; subst. left. rewrite h8. apply add_sub_compat_in; auto. right. intro; subst. contradict h8. rewrite sub_add_eq. rewrite list_to_set_in_iff in h6. rewrite subtract_nin; auto. red in ih. specialize (ih E h2). destruct ih as [h8 | h8]; subst. right. intro h8. contradict h7. rewrite <- h8. right. constructor. right. intro; subst. contradict h7. right. constructor. Qed. Lemma fin_ens_eq_dec_eq : forall {T:Type} (A B:Ensemble T) (pfa:Finite A) (pfb:Finite B) (pf pf':{A = B} + {A <> B}), pf = pf'. intros T A B h1 h2 h3 h4. destruct h3 as [h3 | h5], h4 as [h4 | h6]; simpl. f_equal. apply proof_irrelevance. contradiction. contradiction. f_equal. apply proof_irrelevance. Qed. Lemma fin_ens_eq_dec : forall (T:Type) (pfd:type_eq_dec T), type_dep_eq_dec (@Finite T). intros T h0. red. intros A B h1 h2. pose proof (list_to_set_type_dep_eq_dec_finite_dec h0). pose proof (@fun_finite_list_to_set_eq T (fun A B (pfa:Finite A) (pfb:Finite B) => {A = B} + {A <> B})) as h3. simpl in h3. apply h3; auto. intros; apply fin_ens_eq_dec_eq; auto. intros C h5 l h6. subst. apply (list_to_set_type_dep_eq_dec_finite_dec h0 l C h5). Qed. Lemma finite_mem_dep_in_dec : forall {T:Type} (pfd:type_eq_dec T) (F:Family T), finite_mem F -> type_dep_eq_dec (Inn F). intros T F h0 h1. red. intros A B h2 h3. do 2 red in h1. pose proof (h1 _ h2) as h4. pose proof (h1 _ h3) as h5. apply fin_ens_eq_dec; auto. Qed. Lemma eq_empty_dec : forall {T} (pfdt:type_eq_dec T) (E:Ensemble T) (pff:Finite E), {E = Empty_set _} + {E <> Empty_set _}. intros; apply fin_ens_eq_dec; auto. apply Empty_is_finite. Qed. Definition in_ens_eq_dec {T:Type} (A:Ensemble T) : Type := forall (x y:T), Inn A x -> Inn A y -> {x = y} + {x <> y}. End EnsEqDec. End TypeEqDec. bool2-v0-3/ListUtilities3.v0000644000175000017500000000561013575574055016425 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2017 Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Import SetUtilities. Require Import SetUtilities2. Require Import SetUtilities2_5. Require Import ListUtilities. Require Import ListUtilities2. Require Import EqDec. Require Import LogicUtilities. Require Import ArithUtilities. Lemma list_constructive_definite_description : forall {T:Type} (pfd:type_eq_dec T) (l:list T) (P:T->Prop), (exists x:T, In x l /\ P x) -> {b:T | In b l /\ P b}. intros T hd l P h1. pose (list_to_seg_one_t_fun' (list_singularize hd l)) as f. assert (h2: exists x:sig_set (list_to_set l), In (proj1_sig x) l /\ P (proj1_sig x)). destruct h1 as [x h1]. destruct h1 as [h1 h2]. rewrite list_to_set_in_iff in h1. exists (exist _ _ h1); simpl. rewrite <- list_to_set_in_iff in h1. split; auto. pose proof (card_fun1_length_eq hd l) as h3. assert (h4:forall x:sig_set (seg (card_fun1 (list_to_set l))), 1 <= S (proj1_sig x) <= length (list_singularize hd l)). intro x. destruct x as [x h5]. simpl. destruct h5 as [h5]. rewrite <- h3. omega. pose (fun x:sig_set (seg (card_fun1 (list_to_set l))) => f (exist _ _ (h4 x))) as f'. assert (h6:forall x:{x0 : T | In x0 (list_singularize hd l)}, Inn (list_to_set l) (proj1_sig x)). intro x. destruct x as [x h5]. simpl. rewrite <- list_singularize_in_iff in h5. rewrite <- list_to_set_in_iff. assumption. pose (fun x:sig_set (seg (card_fun1 (list_to_set l))) => exist _ _ (h6 (f' x))) as f''. pose proof (finite_constructive_definite_description' _ (list_to_set_finite l) f'') as h7. apply h7; auto. apply inj_impl_bij_sig_set_eq_card. apply finite_seg. apply list_to_set_finite. rewrite card_fun1_seg_eq; auto. pose proof (list_to_seg_one_t_fun_no_dup_inj' (list_singularize hd l) (no_dup_list_singularize hd _)) as h8. red in h8. red. intros x y h9. destruct x as [x h10], y as [y h11]. unfold f'' in h9. apply exist_injective in h9. apply proj1_sig_injective in h9. unfold f', f in h9. apply h8 in h9. apply exist_injective in h9. simpl in h9. apply S_inj in h9. subst. apply proj1_sig_injective; auto. Qed. bool2-v0-3/EqDec.v0000644000175000017500000000752013575574055014516 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014, Daniel Wyckoff*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import FunctionalExtensionality. Require Import ProofIrrelevance. Definition type_eq_dec (T:Type) : Type := forall (x y:T), {x = y} + {x <> y}. Lemma type_eq_dec_eq : forall {T:Type} (pf pf' : type_eq_dec T), pf = pf'. intros T h1 h2. red in h1. red in h2. extensionality x. extensionality y. destruct (h1 x y) as [h3 | h3], (h2 x y) as [h4 | h4]; subst. f_equal. apply proof_irrelevance. contradict h4; auto. contradict h3; auto. f_equal. apply proof_irrelevance. Qed. Lemma impl_type_eq_dec_prod : forall {T U:Type}, type_eq_dec T -> type_eq_dec U -> type_eq_dec (T*U). intros T U h1 h2. red in h1, h2. red. intros x y. destruct x as [x1 x2], y as [y1 y2]. specialize (h1 x1 y1). specialize (h2 x2 y2). destruct h1 as [| h1], h2 as [| h2]; subst. left; auto. right. intro h3. inversion h3. contradiction. right. intro h3. inversion h3. contradiction. right. intro h3. inversion h3. contradiction. Qed. (*This doesn't seem to be that useful, but the next one is!*) (*Wyckoff*) Lemma match_eq_dec_sym : forall {T U:Type} (pf:type_eq_dec T) (x y:T) (bl br:U), match (pf x y) with | left _ => bl | right _ => br end = match (pf y x) with | left _ => bl | right _ => br end. intros T U h1 x y ? ?. destruct (h1 x y) as [h2 | h2], (h1 y x) as [h3 | h3]; auto; subst; try contradict h2; try contradict h3; auto. Qed. (* maybe I can twist this somewhere better.*) (* Lemma type_eq_dec_prod_impl : forall {T U:Type}, type_eq_dec (T*U) -> {inhabited T} + {inhabited U} -> (type_eq_dec T)*(type_eq_dec U). intros T U h1 h3. destruct h3 as [h3 | h3]. *) Definition type_dep_eq_dec {T:Type} (P:T->Prop) : Type := forall (a b:T), P a -> P b -> {a = b} + {a <> b}. (*Note so I don't try again : something like "Lemma type_dep_eq_dec_impl : forall {T:Type} (P:T->Prop), type_dep_eq_dec P -> excl_midP P -> type_eq_dec T." can't be done. But the opposite can:*) Lemma type_dep_eq_dec_compat : forall {T:Type} (P:T->Prop), type_eq_dec T -> type_dep_eq_dec P. intros ? ? h1. red in h1; red. intros; apply h1. Qed. Definition dep_type_eq_dec {S} (T:S->Type) : Type := forall s, type_eq_dec (T s). Lemma sig_eq_dec : forall {T:Type} (pf:type_eq_dec T) (P:T->Prop), type_eq_dec {x:T | P x}. intros T pf P; red; red in pf. intros x y. destruct x as [x hx], y as [y hy]. specialize (pf x y). destruct pf as [pf | pf]; subst. left; f_equal; apply proof_irrelevance. right; intro h1. apply (f_equal (@proj1_sig _ _)) in h1; simpl in h1; subst. contradict pf; auto. Qed. (*Hopefully one day I will be able to refactor all usages of this in bool2 to some consequence of Boolean deciable equality, and eliminate this axiom altogether. It's mainly used for decidable equality of arbitrarily infinite sets, which can perhaps be dealt with more ad hoc.*) Axiom eq_dec : forall {T:Type}, type_eq_dec T. (*This axiom is so far only "used" in an Ltac tactic, where you have to separate proof-terms by type.*) (* Axiom type_decidable : type_eq_dec Type.*)bool2-v0-3/FiniteMaps2.v0000644000175000017500000010713613575574055015662 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2015-2019 Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import SetUtilities. Require Import TypeUtilities. Require Import LogicUtilities. Require Import FiniteMaps. Require Import ArithUtilities. Require Import ListUtilities. Require Import FiniteBags. Require Import NPeano. Require Import FunctionProperties. Require Import FunctionalExtensionality. Require Import Description. Require Import TypeUtilities2. Require Import NaryFunctions. Require Import DecidableDec. Require Import SetUtilities2. Require Import ProofIrrelevance. Require Import ListUtilities2. Require Import EqDec. Require Import SetUtilities2_5. Lemma fin_map_to_l_exp_seg_ex : forall {T U:Type} (pfdt:type_eq_dec T) {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) (la:list T) (lb:list U), NoDup la -> list_to_set la = A -> list_to_set lb = B -> exists! f:{m:nat | m < (length la)} -> {y:U|In y lb}, forall (x:T) (pf:In x la), F |-> x = proj1_sig (f (exist (fun c=> c < length la) _ (lt_lind pfdt la _ pf))). intros T U hdt A B def F la lb hn h3 h4. subst. assert (h5:forall i:{m:nat | m < length la}, In (F |-> (nth_lt _ _ (proj2_sig i))) lb). intro i. rewrite list_to_set_in_iff. subst. apply fin_map_app_in. rewrite <- list_to_set_in_iff. pose proof (nth_lt_compat _ _ (proj2_sig i)) as h3. destruct h3 as [h3l h3r]. apply in_nth_lt. exists (fun m=> (exist _ (F |-> (nth_lt _ _ (proj2_sig m))) (h5 m))). red; simpl. split. intros x h6. pose proof (lind_compat hdt _ _ h6) as h7. rewrite h7. reflexivity. intros f h6. apply functional_extensionality. intro i. destruct i as [i h7]. simpl. apply proj1_sig_injective. simpl. rewrite (h6 _ (in_nth_lt _ _ h7)). f_equal. f_equal. apply proj1_sig_injective. simpl. eapply nth_lt_inj. apply hn. rewrite lind_compat. reflexivity. Qed. Definition fin_map_to_l_exp_seg {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (F:Fin_map A B def) (la:list T) (lb:list U) (pfa:NoDup la) (pflsa:list_to_set la = A) (pflsb:list_to_set lb = B) : {m:nat | m < length la}->{y:U|In y lb} := proj1_sig (constructive_definite_description _ (fin_map_to_l_exp_seg_ex pfdt F la lb pfa pflsa pflsb)). Lemma fin_map_to_l_exp_seg_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (F:Fin_map A B def) (la:list T) (lb:list U) (pfa:NoDup la) (pflsa:list_to_set la = A) (pflsb:list_to_set lb = B), let f:=(fin_map_to_l_exp_seg pfdt F la lb pfa pflsa pflsb) in forall (x:T) (pf:In x la), F |-> x = proj1_sig (f ((exist (fun c=> c < length la) _ (lt_lind pfdt la x pf)))). intros T U A B def h0 F la lb h1 h2 h3 f x h4. unfold fin_map_to_l_exp_seg in f. destruct constructive_definite_description as [f' h5]. simpl in f. apply h5. Qed. Lemma fin_map_to_l_exp_bij : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (def:U) (la:list T) (lb:list U) (pfa:NoDup la) (pflsa:list_to_set la = A) (pflsb:list_to_set lb = B), bijective (fun f:Fin_map A B def=> fin_map_to_l_exp_seg pfdt f la lb pfa pflsa pflsb). intros T U hdt A B def la lb h1 h2 h3. subst. red. split. red. intros F F' h2. pose proof (fin_map_to_l_exp_seg_compat hdt F la lb h1 eq_refl eq_refl) as h3. pose proof (fin_map_to_l_exp_seg_compat hdt F' la lb h1 eq_refl eq_refl) as h4. apply fin_map_ext. intros x h5. rewrite <- list_to_set_in_iff in h5. rewrite (h3 _ h5). rewrite (h4 _ h5). f_equal. rewrite h2. reflexivity. red. intro f. pose [pr:T*U | exists pf:In (fst pr) la, snd pr = (proj1_sig (f (exist (fun c=> c < length la) _ (lt_lind hdt _ _ pf))))] as S. assert (h2:fun_well_defined S). red. split. intros x h2. destruct h2 as [h2]. destruct h2 as [y h2]. exists y. red. split. split; auto. constructor. exists x. assumption. intros y' h3. destruct h3 as [h3l h3r]. inversion h2 as [h4]. clear h2. simpl in h4. destruct h4 as [h4 h5]. inversion h3r as [h6]. clear h3r. simpl in h6. destruct h6 as [h6 h7]. pose proof (proof_irrelevance _ h4 h6); subst; auto. intros pr h2. destruct h2 as [h2]. destruct h2 as [h2 h3]. split. constructor. exists (snd pr). constructor. simpl. exists h2. assumption. constructor. exists (fst pr). constructor. simpl. exists h2; auto. red in h2. assert (h3:dom_rel S = (list_to_set la)). apply Extensionality_Ensembles. red. split. red. intros x h3. destruct h3 as [h3]. destruct h3 as [y h3]. inversion h3 as [h4]. clear h3. simpl in h4. destruct h4 as [h4]. rewrite <- list_to_set_in_iff. assumption. red. intros x h3. constructor. rewrite <- list_to_set_in_iff in h3. exists (proj1_sig (f (exist (fun c => c < length la) _ (lt_lind hdt la x h3)))). constructor. simpl. exists h3; auto. rewrite h3 in h2. pose proof (list_to_set_finite la) as h4. assert (h5:Finite (ran_rel S)). pose proof (fp_ran_rel_im' hdt h4 h2 def) as hran. rewrite hran. apply finite_image. apply h4. pose (fin_map_intro _ _ def hdt h4 h5 _ h2) as F. pose proof (list_to_set_finite lb) as h6. assert (h7:Included (ran_rel S) (list_to_set lb)). red. intros y h8. destruct h8 as [h8]. destruct h8 as [x h8]. inversion h8 as [h9]. clear h8. simpl in h9. destruct h9 as [h9 ?]. subst. pose proof (proj2_sig (f (exist (fun c=> c < length la) _ (lt_lind hdt la x h9)))) as h10. simpl in h10. rewrite <- list_to_set_in_iff. assumption. pose (fin_map_new_ran F h6 h7) as F'. exists F'. pose proof (fin_map_to_l_exp_seg_compat hdt F' la lb h1 eq_refl eq_refl) as h8. apply functional_extensionality. intro i. destruct i as [i h9]. apply proj1_sig_injective. specialize (h8 (nth_lt la _ h9)). pose proof (nth_lt_compat la _ h9) as h10. pose proof (in_nth_lt la i h9) as h10'. specialize (h8 h10'). unfold F' in h8. rewrite <- fin_map_new_ran_compat in h8. pose proof h10' as h10''. rewrite list_to_set_in_iff in h10''. pose proof (in_fin_map_to_fps F _ h10'') as h11. rewrite h8 in h11. unfold F in h11. rewrite <- fin_map_to_fps_s_compat in h11. inversion h11 as [h12]. simpl in h12. unfold F'. destruct h12 as [h12 h13]. unfold F. pose proof (lind_compat hdt la (nth_lt la i h9) h10') as h14. apply nth_lt_inj in h14. pose proof (proof_irrelevance _ h10' h12) as h15. assert (h15':exist (fun c=> c < length la) _ (lt_lind hdt _ _ h12) = exist (fun m=> m < length la) i h9). apply proj1_sig_injective. simpl. rewrite <- h15. assumption. rewrite h15' in h13. rewrite <- h13. f_equal. f_equal. rewrite h15. rewrite h15'. reflexivity. assumption. rewrite <- list_to_set_in_iff. assumption. Qed. Lemma card_l_exp_seg : forall {T:Type} (l:list T) (n:nat), NoDup l -> FiniteT_card _ (finite_l_exp_seg l n) = (length l)^n. intros T l n h1. pose proof (l_exp_seg_to_nprod_in_l_bij l n) as h2. apply bijective_impl_invertible in h2. pose proof (FiniteT_card_bijection). pose proof (finite_l_exp_seg l n) as h3. pose proof (FiniteT_card_bijection _ _ h3 _ h2) as h4. pose proof (proof_irrelevance _ (finite_l_exp_seg l n) h3); subst. rewrite <- h4. pose proof (card_nprod_in_l _ h1 n) as h5. pose proof (proof_irrelevance _ (finite_nprod_in_l l n) (bij_finite ({m : nat | m < n} -> {x : T | In x l}) ({x : T | In x l} ^ n) (l_exp_seg_to_nprod_in_l l n) (finite_l_exp_seg l n) h2)) as h6. rewrite <- h6; auto. Qed. Lemma card_fin_maps : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (def:U) (pfa:Finite A) (pfb:Finite B), FiniteT_card (Fin_map A B def) (finitet_fin_maps _ _ def pfa pfb) = (card_fun1 B)^(card_fun1 A). intros T U hdt A B def h1 h2. pose proof (finite_set_list_no_dup _ h1) as h3. destruct h3 as [la h3]. destruct h3 as [h3l h3r]. pose proof (finite_set_list_no_dup _ h2) as h4. destruct h4 as [lb h4]. destruct h4 as [h4l h4r]. subst. pose proof (fin_map_to_l_exp_bij hdt _ _ def _ lb h3r (eq_refl _) (eq_refl _)) as h5. apply bijective_impl_invertible in h5. pose proof (finitet_fin_maps _ _ def h1 h2) as h6. pose proof (FiniteT_card_bijection _ _ h6 _ h5) as h7. pose proof (proof_irrelevance _ h6 (finitet_fin_maps (list_to_set la) (list_to_set lb) def h1 h2)); subst. rewrite <- h7. pose proof (card_l_exp_seg _ (length la) h4r) as h8. pose proof (proof_irrelevance _ (finite_l_exp_seg lb (length la)) (bij_finite (Fin_map (list_to_set la) (list_to_set lb) def) ({m : nat | m < length la} -> {y : U | In y lb}) (fun f : Fin_map (list_to_set la) (list_to_set lb) def => fin_map_to_l_exp_seg hdt f la lb h3r eq_refl eq_refl) (finitet_fin_maps (list_to_set la) (list_to_set lb) def h1 h2) h5)) as h9. rewrite <- h9. clear h9. rewrite h8. pose proof (card_fun1_compat (list_to_set lb)) as h9. pose proof (card_fun1_compat (list_to_set la)) as h10. destruct h9 as [h9l h9r]. destruct h10 as [h10l h10r]. specialize (h9l h2). specialize (h10l h1). pose proof (cardinal_length_compat la h3r) as h11. pose proof (cardinal_length_compat lb h4r) as h12. assert (h13:length la = card_fun1 (list_to_set la)). eapply cardinal_is_functional. apply h11. apply h10l. reflexivity. assert (h14:length lb = card_fun1 (list_to_set lb)). eapply cardinal_is_functional. apply h12. apply h9l. reflexivity. rewrite h13. rewrite h14. reflexivity. Qed. Lemma ex_same_im_subset_fin_map_inj : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (F:Fin_map A B def), exists (A':Ensemble T) (pf:Included A' A), let F':=restriction F pf in im_fin_map F = im_fin_map F' /\ inj F'. intros T U A B def hdt F. pose proof (fin_map_fin_dom F) as h1. pose proof (ex_same_im_subset_sig_inj hdt h1 (fin_map_to_sig_fun F)) as h2. destruct h2 as [A' [h2 [h3 h4]]]. destruct h3 as [h3]. destruct h3 as [h3a h3b]. exists A', h2. simpl. split. rewrite restriction_sig_fin_map_to_sig_fun in h3b. do 2 rewrite im_fin_map_to_sig_fun in h3b. pose proof (proof_irrelevance _ h2 h3a); subst; auto. rewrite restriction_sig_fin_map_to_sig_fun in h4. rewrite <- fin_map_inj_iff. assumption. Qed. Lemma card_im_eq_inj : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (F:Fin_map A B def), card_fun1 A = card_fun1 (im_fin_map F) -> inj F. intros T U A B def hdt F h1. pose proof (ex_same_im_subset_fin_map_inj hdt F) as h5. simpl in h5. destruct h5 as [A' [h5 [h6 h7]]]. rewrite h6 in h1. pose proof (fin_map_fin_dom F) as hf. pose proof (fin_map_card_im_le_dom' (restriction F h5)) as h8. pose proof (incl_card_fun1 _ _ hf h5) as h9. assert (h10:card_fun1 A' = card_fun1 A). omega. apply incl_card_fun1_eq in h10; auto. subst. pose proof (restriction_trivial_eq F h5) as h10. rewrite h10 in h7. assumption. Qed. Lemma dom_ran_card_eq_bij : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (F:Fin_map A B def), card_fun1 A = card_fun1 (im_fin_map F) -> card_fun1 B = card_fun1 (im_fin_map F) -> bij F. intros T U A B def hdt F h1 h0. pose proof (card_im_eq_inj hdt F h1) as h2. symmetry in h0. pose proof (im_ran_card_eq_surj F h0) as h3. split; auto. Qed. (*This is for constructing a finite map from an explicit list of pairs that conform to two easier to formulate properties.*) Lemma fps_to_f_eq_card_dom_rel_ex : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (S:Ensemble (T*U)) (def:U), Finite S -> card_fun1 S = card_fun1 (dom_rel S) -> exists! (F:Fin_map (dom_rel S) (ran_rel S) def), S = [pr:T*U | snd pr = F |-> (fst pr) /\ Inn (dom_rel S) (fst pr)]. intros T U hdt hdu S def h1 h2. pose proof (impl_type_eq_dec_prod hdt hdu) as hdtu. pose proof (finite_inh_or_empty _ h1) as h0. destruct h0 as [h0a | h0b]. destruct h0a as [s h0a]. assert (h3:forall x:sig_set S, Inn (dom_rel S) (fst (proj1_sig x))). intro p. destruct p as [p h4]. simpl. constructor. exists (snd p). rewrite <- surjective_pairing. assumption. pose (fun x:sig_set S=>fst (proj1_sig x)) as f. pose (sig_fun_to_fin_map hdtu f h1 (fst s)) as f'. assert (h4: im_fin_map f' = dom_rel S). apply Extensionality_Ensembles. red. split. intros x h4. destruct h4 as [x h4]. subst. pose proof (h3 (exist _ _ h4)) as h5. simpl in h5. unfold f'. rewrite (sig_fun_to_fin_map_compat hdtu f h1 (fst s) x h4) at 1. unfold f. simpl. assumption. red. intros x h4. destruct h4 as [h4]. destruct h4 as [y h4]. unfold im_fin_map. apply Im_intro with (x, y); auto. unfold f'. rewrite (sig_fun_to_fin_map_compat hdtu f h1 (fst s) _ h4) at 1. unfold f. simpl. reflexivity. rewrite <- h4 in h2. pose proof @card_im_eq_inj. pose proof (card_im_eq_inj hdtu f' h2) as h5. pose (fin_map_left_inverse hdt f' h5 s) as f2. assert (h6:Included (dom_rel S) (im_fin_map f')). rewrite h4. auto with sets. pose (restriction f2 h6) as f3. assert (h7:forall x:T, Inn (dom_rel S) x -> Inn (ran_rel S) (snd (f3 |-> x))). intros x h7. destruct h7 as [h7]. destruct h7 as [y h7]. unfold f3. rewrite restriction_compat. unfold f2. assert (h8:Inn (im_fin_map f') x). rewrite h4. constructor. exists y; auto. destruct h8 as [x h8]. subst. rewrite fin_map_left_inverse_compat. constructor. exists (fst x). rewrite <- surjective_pairing. assumption. assumption. constructor. exists y. assumption. pose (fun x:sig_set (dom_rel S) => (snd (f3 |-> (proj1_sig x)))) as g. pose (sig_fun_to_fin_map_ran hdt g (dom_rel_finite S h1) def (ran_rel S)). assert (h8: Included (Im (full_sig (dom_rel S)) g) (ran_rel S)). red. intros x h9. destruct h9 as [x h9]. subst. clear h9. constructor. unfold g. destruct x as [x h8]. simpl. unfold f3. rewrite restriction_compat. unfold f2. rewrite <- h4 in h8. destruct h8 as [x h8]. subst. rewrite fin_map_left_inverse_compat. exists (fst x). rewrite <- surjective_pairing. assumption. assumption. assumption. pose (sig_fun_to_fin_map_ran hdt g (dom_rel_finite S h1) def (ran_rel S) (ran_rel_finite S h1) h8) as f4. exists f4. red. split. apply Extensionality_Ensembles. red. split. (* <= *) red. intros pr h9. constructor. split. unfold f4. assert (h10:Inn (dom_rel S) (fst pr)). constructor. exists (snd pr). rewrite <- surjective_pairing. assumption. pose proof (sig_fun_to_fin_map_ran_compat hdt g (dom_rel_finite S h1) def (ran_rel S) (ran_rel_finite S h1) h8 _ h10) as h11. rewrite h11. unfold g. simpl. unfold f3. rewrite restriction_compat. unfold f2. assert (h12:fst pr = f' |-> pr). unfold f'. pose proof (sig_fun_to_fin_map_compat hdtu f h1 (fst s) pr h9) as h12. rewrite h12. unfold f. simpl. reflexivity. rewrite h12. rewrite fin_map_left_inverse_compat at 1. reflexivity. assumption. constructor. exists (snd pr). rewrite <- surjective_pairing. assumption. constructor. exists (snd pr). rewrite <- surjective_pairing. assumption. (* >= *) red. intros p h9. destruct h9 as [h9]. destruct h9 as [h9 h10]. destruct p as [x y]. simpl in h9, h10. rewrite <- h4 in h10. destruct h10 as [p h10]. subst. unfold f4. destruct p as [x y]. assert (h13:Inn (dom_rel S) (f' |-> (x, y))). constructor. unfold f'. exists y. rewrite (sig_fun_to_fin_map_compat hdtu f h1 (fst s) _ h10). unfold f. simpl. assumption. pose proof (sig_fun_to_fin_map_ran_compat hdt g (dom_rel_finite S h1) def (ran_rel S) (ran_rel_finite S h1) h8 _ h13) as h14. rewrite h14 at 1. unfold g, f3. simpl. rewrite restriction_compat. unfold f2. rewrite fin_map_left_inverse_compat. simpl. unfold f'. pose proof (sig_fun_to_fin_map_compat hdtu f h1 (fst s) (x, y) h10) as h15. rewrite h15. unfold f. simpl. assumption. assumption. assumption. intros k h9. apply fin_map_ext. intros x h10. unfold f4. pose proof (sig_fun_to_fin_map_ran_compat hdt g (dom_rel_finite S h1) def (ran_rel S) (ran_rel_finite S h1) h8 _ h10) as h11. rewrite h11 at 1. unfold g, f3. simpl. rewrite restriction_compat. unfold f2. pose h10 as h10'. rewrite <- h4 in h10'. destruct h10' as [p h10' ? h12]. rewrite h12. rewrite fin_map_left_inverse_compat. rewrite h9 in h10'. destruct h10' as [h10']. destruct h10' as [h10a h10b]. rewrite h10a. f_equal. unfold f'. destruct p as [x' y']. simpl in h10a, h10b. simpl. assert (h13: Inn S (x', y')). rewrite h9. constructor. simpl. split; auto. pose proof (sig_fun_to_fin_map_compat hdtu f h1 (fst s) (x', y') h13) as h14. rewrite h14. unfold f. simpl. reflexivity. assumption. assumption. subst. rewrite dom_rel_empty, ran_rel_empty. exists (empty_map1 _ _ hdt def (Empty_set U) (Empty_is_finite U)). red. split. apply Extensionality_Ensembles. red; split; auto with sets. red. intros x h3. destruct h3 as [h3]. destruct h3; contradiction. intros f h3. symmetry. pose proof (type_eq_dec_eq hdt (fin_map_type_eq_dec f)). subst. apply empty_map1_compat. Qed. Definition fps_to_f_eq_card_dom_rel {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (S:Ensemble (T*U)) (def:U) (pff:Finite S) (pfeq:card_fun1 S = card_fun1 (dom_rel S)) := proj1_sig (constructive_definite_description _ (fps_to_f_eq_card_dom_rel_ex pfdt pfdu S def pff pfeq)). Lemma fps_to_f_eq_card_dom_rel_compat : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (S:Ensemble (T*U)) (def:U) (pff:Finite S) (pfeq:card_fun1 S = card_fun1 (dom_rel S)), let F := fps_to_f_eq_card_dom_rel pfdt pfdu S def pff pfeq in S = [pr:T*U | snd pr = F |-> (fst pr) /\ Inn (dom_rel S) (fst pr)]. intros T U S hdt hdu def h1 h2 F. unfold F, fps_to_f_eq_card_dom_rel. destruct constructive_definite_description as [F' h3]. simpl. assumption. Qed. Lemma fps_to_f_eq_card_dom_rel_compat' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (S:Ensemble (T*U)) (def:U) (pff:Finite S) (pfeq:card_fun1 S = card_fun1 (dom_rel S)) (pr:T*U), Inn S pr -> fps_to_f_eq_card_dom_rel pfdt pfdu S def pff pfeq |-> (fst pr) = snd pr. intros T U S def hdt hdu h1 h2 pr h3. pose proof (fps_to_f_eq_card_dom_rel_compat S def hdt hdu h1 h2) as h4. simpl in h4. rewrite h4 in h3. destruct h3 as [h3]. destruct h3 as [h3a h3b]. rewrite h3a. reflexivity. Qed. Definition fin_map_compose_list_to_seg {T U:Type} {l:list T} (pfdt:type_eq_dec T) {B:Ensemble U} {def:U} (a:Fin_map (list_to_set l) B def) : Fin_map (seg (length l)) B def. pose proof (segT_injector_compat (list_to_segT_fun l)) as h1. simpl in h1. destruct h1 as [h1 h2]. pose proof (im_list_to_segT_fun pfdt l) as hi. rewrite hi in h1. let P := type of h2 in match P with injective ?f' => pose f' as f end. assert (h3:forall m:sig_set (seg (length l)), Inn (list_to_set l) (list_to_segT_fun l (st_conv_rev m))). intro m. pose proof (list_to_segT_fun_compat l (st_conv_rev m)) as h3. destruct h3 as [h3 h4]. rewrite <- list_to_set_in_iff. assumption. pose (fun m:sig_set (seg (length l)) => a |-> (list_to_segT_fun l (st_conv_rev m))) as fc. refine (sig_fun_to_fin_map_ran nat_eq_dec fc (finite_seg _) def B (fin_map_fin_ran a) _). red. intros x h4. destruct h4 as [x h4]. subst. unfold fc. eapply im_fin_map_ran. apply in_im_fin_map. apply h3. Defined. Lemma fin_map_compose_list_to_seg_compat : forall {T U:Type} {l:list T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (a:Fin_map (list_to_set l) B def), let f :=list_to_segT_fun l in let F := fin_map_compose_list_to_seg pfdt a in forall m:(length l) @, F |-> (proj1_sig m) = a |-> (f m). intros T U l B def hdt a f F m. unfold f. pose proof (list_to_segT_fun_compat l m) as h1. destruct h1 as [h1 h2]. unfold F, fin_map_compose_list_to_seg. destruct (segT_injector_compat (list_to_segT_fun l)) as [h3 h4]. match goal with |- sig_fun_to_fin_map_ran ?C ?D ?E ?F ?G ?H ?I |-> _ = _ => pose (sig_fun_to_fin_map_ran_compat C D E F G H I (proj1_sig m)) as h5 end. simpl in h5. assert (h6 : Inn (seg (length l)) (proj1_sig m)). constructor. apply proj2_sig. specialize (h5 h6). rewrite h5 at 1. f_equal. f_equal. apply proj1_sig_injective; simpl. reflexivity. Qed. (*For use in functionality related to el_prod_lind.*) Definition fin_map_compose_list_to_seg_lind {T U:Type} {l:list T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (a:Fin_map (list_to_set l) B def) (m:(length l) @) : Fin_map (segT_injector (list_to_segT_fun (firstn (proj1_sig m) l))) B def. pose proof (segT_injector_compat (list_to_segT_fun (firstn (proj1_sig m) l))) as h1. simpl in h1. destruct h1 as [h1 h2]. pose proof (im_list_to_segT_fun pfdt (firstn (proj1_sig m) l)) as hi. rewrite hi in h1. let P := type of h2 in match P with injective ?f' => pose f' as f end. assert (h3 : forall c:sig_set (segT_injector (list_to_segT_fun (firstn (proj1_sig m) l))), Inn (list_to_set l) (f c)). intro c. destruct c as [c hc]. destruct c as [c hc']. pose proof hc' as hc''. rewrite length_firstn in hc''. rewrite <- list_to_set_in_iff. unfold f, restriction_fun. simpl. pose proof (list_to_segT_fun_compat (firstn (proj1_sig m) l) (exist (fun m0 : nat => m0 < length (firstn (proj1_sig m) l)) c hc')) as h4. simpl in h4. destruct h4 as [h4 h5]. assert (h6:0 < proj1_sig m). destruct (zerop (proj1_sig m)) as [h6|]; auto. rewrite h6 in hc''. apply lt_n_0 in hc''. contradiction. eapply in_firstn. apply h4. apply lt_le_weak. apply proj2_sig. pose (fun c:sig_set (segT_injector (list_to_segT_fun (firstn (proj1_sig m) l))) => a |-> (f c)) as fc. refine (sig_fun_to_fin_map_ran (A:= segT_injector (list_to_segT_fun (firstn (proj1_sig m) l))) (segT_eq_dec _) fc (finite_segT_injector _) def B (fin_map_fin_ran a) _). red. intros x h4. destruct h4 as [x h4]. subst. unfold fc. eapply im_fin_map_ran. apply in_im_fin_map. apply h3. Defined. Lemma fin_map_compose_list_to_seg_lind_compat : forall {T U:Type} {l:list T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (a:Fin_map (list_to_set l) B def) (m:(length l) @), let f := restriction_fun (list_to_segT_fun (firstn (proj1_sig m) l)) (segT_injector (list_to_segT_fun (firstn (proj1_sig m) l))) in let F := fin_map_compose_list_to_seg_lind pfdt a m in forall c:sig_set (segT_injector (list_to_segT_fun (firstn (proj1_sig m) l))), F |-> (proj1_sig c) = a |-> f c. intros T U l B def hdt a m f F c. unfold F, f, restriction_fun. pose proof (list_to_segT_fun_compat (firstn (proj1_sig m) l) (proj1_sig c)) as h1. destruct h1 as [h1 h2]. unfold fin_map_compose_list_to_seg_lind. destruct (segT_injector_compat (list_to_segT_fun (firstn (proj1_sig m) l))) as [h3 h4]. match goal with |- sig_fun_to_fin_map_ran ?C ?D ?E ?F ?G ?H ?I |-> _ = _ => pose (sig_fun_to_fin_map_ran_compat C D E F G H I (proj1_sig c)) as h5 end. simpl in h5. assert (h6 : Inn (segT_injector (list_to_segT_fun (firstn (proj1_sig m) l))) (proj1_sig c)). apply proj2_sig. specialize (h5 h6). rewrite h5 at 1. f_equal. Qed. (*For use in functionality related to el_prod_l *) Definition fin_map_compose_list_to_seg_l {T U:Type} {l:list T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (a:Fin_map (list_to_set l) B def) : Fin_map (segT_injector (list_to_segT_fun l)) B def. pose proof (segT_injector_compat (list_to_segT_fun l)) as h1. simpl in h1. destruct h1 as [h1 h2]. pose proof (im_list_to_segT_fun pfdt l) as hi. rewrite hi in h1. let P := type of h2 in match P with injective ?f' => pose f' as f end. assert (h3:forall c:sig_set (segT_injector (list_to_segT_fun l)), Inn (list_to_set l) (f c)). intro c. rewrite <- list_to_set_in_iff. unfold f. unfold restriction_fun. pose proof (list_to_segT_fun_compat l (proj1_sig c)) as h3. destruct h3; auto. pose (fun c:sig_set (segT_injector (list_to_segT_fun l)) => a |-> (f c)) as fc. refine (sig_fun_to_fin_map_ran (segT_eq_dec _) fc (finite_segT_injector _) def B (fin_map_fin_ran a) _). red. intros x h4. destruct h4 as [x h4]. subst. unfold fc. eapply im_fin_map_ran. apply in_im_fin_map. apply h3. Defined. Lemma fin_map_compose_list_to_seg_l_compat : forall {T U:Type} {l:list T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (a:Fin_map (list_to_set l) B def), let f := restriction_fun (list_to_segT_fun l) (segT_injector (list_to_segT_fun l)) in let F := fin_map_compose_list_to_seg_l pfdt a in forall c:sig_set (segT_injector (list_to_segT_fun l)), F |-> (proj1_sig c) = a |-> f c. intros T U l B def hdt a f F c. unfold F, f. unfold restriction_fun. pose proof (list_to_segT_fun_compat l (proj1_sig c)) as h1. destruct h1 as [h1 h2]. unfold fin_map_compose_list_to_seg_l. destruct (segT_injector_compat (list_to_segT_fun l)) as [h3 h4]. match goal with |- sig_fun_to_fin_map_ran ?C ?D ?E ?F ?G ?H ?I |-> _ = _ => pose (sig_fun_to_fin_map_ran_compat C D E F G H I (proj1_sig c)) as h5 end. simpl in h5. assert (h6 : Inn (segT_injector (list_to_segT_fun l)) (proj1_sig c)). apply proj2_sig. specialize (h5 h6). rewrite h5 at 1. f_equal. Qed. Lemma fin_map_seg_transfer_ex : forall {U:Type} {def:U} {B:Ensemble U} {m n:nat} (f:Fin_map (seg m) B def) (pf:m = n), exists! F:Fin_map (seg n) B def, forall b:n @, F |-> (proj1_sig b) = f |-> (proj1_sig (segT_transfer_r pf b)). intros U def B m n f h1. subst. exists f. red. split. intro b. rewrite segT_transfer_r_eq_refl. reflexivity. intros f' h1. apply fin_map_ext. intros a h2. destruct h2 as [h2]. specialize (h1 (exist (fun c => c < n) _ h2)). simpl in h1. rewrite h1; auto. Qed. Definition fin_map_seg_transfer {U:Type} {def:U} {B:Ensemble U} {m n:nat} (f:Fin_map (seg m) B def) (pf:m = n) : Fin_map (seg n) B def := proj1_sig (constructive_definite_description _ (fin_map_seg_transfer_ex f pf)). Lemma fin_map_seg_transfer_compat : forall {U:Type} {def:U} {B:Ensemble U} {m n:nat} (f:Fin_map (seg m) B def) (pf:m = n), let F := fin_map_seg_transfer f pf in forall b:n @, F |-> (proj1_sig b) = f |-> proj1_sig (segT_transfer_r pf b). unfold fin_map_seg_transfer; intros; destruct constructive_definite_description; auto. Qed. Lemma fin_map_seg_transfer_compat' : forall {U:Type} {def:U} {B:Ensemble U} {m n:nat} (f:Fin_map (seg m) B def) (pf:m = n) (b:nat), b < n -> fin_map_seg_transfer f pf |-> b = f |-> b. intros U def B m n f h1 b h2. pose proof (fin_map_seg_transfer_compat f h1) (exist (fun c => c < n) _ h2) as h3. simpl in h3. rewrite h3. f_equal. rewrite segT_transfer_r_compat. simpl. reflexivity. Qed. Lemma fin_map_seg_transfer_eq_refl : forall {U:Type} {def:U} {B:Ensemble U} {m:nat} (f:Fin_map (seg m) B def), fin_map_seg_transfer f eq_refl = f. intros; apply fin_map_ext; auto; intros; rewrite fin_map_seg_transfer_compat'; auto. destruct H; auto. Qed. Lemma fin_map_seg_transfer0 : forall {U:Type} {def:U} {B:Ensemble U} {m:nat} (F:Fin_map (seg 0) B def), F = fin_map_dom_subst (eq_sym seg_0_eq) (empty_map1 _ _ nat_eq_dec _ _ (fin_map_fin_ran F)). intros; apply fin_map_ext; intros; destruct H; omega. Qed. Lemma fin_map_seg_one_transfer_ex : forall {U:Type} {def:U} {B:Ensemble U} {m n:nat} (f:Fin_map (Full_set (seg_one_t m)) B def) (pf:m = n), exists! F:Fin_map (Full_set (seg_one_t n)) B def, forall b:seg_one_t n, F |-> b = f |-> sot_transfer_r pf b. intros U def B m n f h1. subst. exists f. red. split. intro b. rewrite sot_transfer_r_eq_refl. reflexivity. intros f' h1. apply fin_map_ext. intros a h2. clear h2. specialize (h1 a). rewrite h1. rewrite sot_transfer_r_eq_refl. reflexivity. Qed. Definition fin_map_seg_one_transfer {U:Type} {def:U} {B:Ensemble U} {m n:nat} (f:Fin_map (Full_set (seg_one_t m)) B def) (pf:m = n) : Fin_map (Full_set (seg_one_t n)) B def := proj1_sig (constructive_definite_description _ (fin_map_seg_one_transfer_ex f pf)). Lemma fin_map_seg_one_transfer_compat : forall {U:Type} {def:U} {B:Ensemble U} {m n:nat} (f:Fin_map (Full_set (seg_one_t m)) B def) (pf:m = n), let F := fin_map_seg_one_transfer f pf in forall b:seg_one_t n, F |-> b = f |-> sot_transfer_r pf b. unfold fin_map_seg_one_transfer; intros; destruct constructive_definite_description; auto. Qed. Lemma fin_map_seg_pred_ex' : forall {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (Full_set ((S n) @)) B def), exists! f' : Fin_map (Full_set (n @)) B def, forall m:n @, f' |-> m = f |-> (segT_S m). intros U def B n f. pose (fun m:sig_set (Full_set (n @)) => f |-> (segT_S (proj1_sig m))) as g. pose (sig_fun_to_fin_map_ran (segT_eq_dec _) g (finite_full_set_segT _) def _ (fin_map_fin_ran f)) as F. let P := type of F in match P with ?hyp -> _ => assert (h1:hyp) end. red. intros x h1. destruct h1 as [x h1]. subst. unfold g. apply fin_map_app_in. constructor. pose (F h1) as F'. exists F'. red. split. unfold F', F. intro m. pose proof (sig_fun_to_fin_map_ran_compat (segT_eq_dec _) g (finite_full_set_segT _) def _ (fin_map_fin_ran f) h1 m (Full_intro _ _)) as h2. rewrite h2. unfold g. simpl. reflexivity. intros G h2. apply fin_map_ext. intros m h3. specialize (h2 m). rewrite h2. unfold F', F. pose proof (sig_fun_to_fin_map_ran_compat (segT_eq_dec _) g (finite_full_set_segT _) def _ (fin_map_fin_ran f) h1 m (Full_intro _ _)) as h4. rewrite h4. unfold g. simpl. reflexivity. Qed. Definition fin_map_seg_pred' {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (Full_set ((S n) @)) B def) := proj1_sig (constructive_definite_description _ (fin_map_seg_pred_ex' f)). Lemma fin_map_seg_pred_compat' : forall {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (Full_set ((S n) @)) B def), let f' := fin_map_seg_pred' f in forall m:n @, f' |-> m = f |-> (segT_S m). unfold fin_map_seg_pred'; intros; destruct constructive_definite_description; auto. Qed. Lemma fin_map_seg_pred_ex : forall {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (seg (S n)) B def), exists! f' : Fin_map (seg n) B def, forall m:nat, m < n -> f' |-> m = f |-> S m. intros U def B n f. pose (fun m:sig_set (seg n) => f |-> S (proj1_sig m)) as g. pose (sig_fun_to_fin_map_ran nat_eq_dec g (finite_seg _) def _ (fin_map_fin_ran f)) as F. let P := type of F in match P with ?hyp -> _ => assert (h1:hyp) end. red. intros x h1. destruct h1 as [x h1]. subst. clear h1. destruct x as [x h1]. unfold g. apply fin_map_app_in. constructor. simpl. apply lt_n_S. destruct h1; auto. pose (F h1) as F'. exists F'. red. split. unfold F', F. intros m hi. assert (hi':Inn (seg n) m). constructor; auto. pose proof (sig_fun_to_fin_map_ran_compat nat_eq_dec g (finite_seg _) def _ (fin_map_fin_ran f) h1 m hi') as h2. rewrite h2. unfold g. simpl. reflexivity. intros G h2. apply fin_map_ext. intros m h3. pose proof h3 as h3'. destruct h3' as [h3']. specialize (h2 m). rewrite h2. unfold F', F. assert (hi':Inn (seg n) m). constructor; auto. pose proof (sig_fun_to_fin_map_ran_compat nat_eq_dec g (finite_seg _) def _ (fin_map_fin_ran f) h1 m h3) as h4. rewrite h4. unfold g. simpl. reflexivity. assumption. Qed. Definition fin_map_seg_pred {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (seg (S n)) B def) := proj1_sig (constructive_definite_description _ (fin_map_seg_pred_ex f)). Lemma fin_map_seg_pred_compat : forall {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (seg (S n)) B def), let f' := fin_map_seg_pred f in forall m:nat, m < n -> f' |-> m = f |-> (S m). unfold fin_map_seg_pred; intros; destruct constructive_definite_description; auto. Qed. (*strong inequality*) Lemma fin_map_seg_rest_ex : forall {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (seg n) B def) (m:nat) (pf:m < n), exists! F:Fin_map (seg m) B def, forall r:nat, r < m -> F |-> r = f |-> r. intros U def B n f m h1. exists (restriction f (incl_lt_seg _ _ h1)). red. split. intros r h2. apply restriction_compat. constructor; auto. intros f' h2. apply fin_map_ext. intros r h3. destruct h3 as [h3]. specialize (h2 _ h3). rewrite h2. apply restriction_compat. constructor; auto. Qed. Definition fin_map_seg_rest {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (seg n) B def) (m:nat) (pf:m < n) : Fin_map (seg m) B def := proj1_sig (constructive_definite_description _ (fin_map_seg_rest_ex f m pf)). Lemma fin_map_seg_rest_compat : forall {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (seg n) B def) (m:nat) (pf:m < n) (r:nat), let F := fin_map_seg_rest f m pf in r < m -> F |-> r = f |-> r. unfold fin_map_seg_rest; intros; destruct constructive_definite_description; auto. Qed. (*weak inequality*) Lemma fin_map_seg_rest_ex' : forall {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (seg n) B def) (m:nat) (pf:m <= n), exists! F:Fin_map (seg m) B def, forall r:nat, r < m -> F |-> r = f |-> r. intros U def B n f m h1. apply le_lt_eq_dec in h1. destruct h1 as [h1|]; subst. exists (fin_map_seg_rest f m h1). red; split. intros; rewrite fin_map_seg_rest_compat; auto. intros f' h2. apply fin_map_ext. intros ? h3; rewrite h2, fin_map_seg_rest_compat; auto. destruct h3; auto. destruct h3; auto. exists f. red; split. intros; auto. intros f' h1. apply fin_map_ext; intros; auto. rewrite h1; auto. destruct H; auto. Qed. Definition fin_map_seg_rest' {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (seg n) B def) (m:nat) (pf:m <= n) : Fin_map (seg m) B def := proj1_sig (constructive_definite_description _ (fin_map_seg_rest_ex' f m pf)). Lemma fin_map_seg_rest_compat' : forall {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (seg n) B def) (m:nat) (pf:m <= n) (r:nat), let F := fin_map_seg_rest' f m pf in r < m -> F |-> r = f |-> r. unfold fin_map_seg_rest'; intros; destruct constructive_definite_description; auto. Qed. Lemma fin_map_seg_rest_compat'' : forall {U:Type} {def:U} {B:Ensemble U} {n:nat} (f:Fin_map (seg n) B def) (m:nat) (pf:m < n) (r:nat), let pf' := lt_n_Sm_le _ _ (lt_trans _ _ _ pf (lt_n_Sn _)) in let F := fin_map_seg_rest f m pf in let F' := fin_map_seg_rest' f m pf' in r < m -> F |-> r = F' |-> r. unfold fin_map_seg_rest', fin_map_seg_rest; intros; do 2 destruct constructive_definite_description; simpl; auto. rewrite e, e0; auto. Qed. bool2-v0-3/WellOrders.v0000644000175000017500000005261013575574055015617 0ustar dwyckoff76dwyckoff76(*Copyright (C) 2016-2019, Daniel Wyckoff, except for the functions labeld "Schepler", which were copied and pasted from Daniel Schepler's "Zorn's Lemma."*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Export Relation_Definitions. Require Import Relation_Definitions_Implicit. Require Import Classical_Wf. Require Import Description. Require Import FunctionalExtensionality. Require Import Classical. Require Import SetUtilities. Require Import Relations. Require Import LogicUtilities. Require Import ZornsLemma. Require Import Arith. Require Import Equality. Require Import ArithUtilities. Require Import DecidableDec. Require Import EqDec. Require Import FunctionProperties. Section WellOrder. Variable T:Type. (*Schepler*) Record well_order (R:relation T) : Prop := { wo_well_founded: well_founded R; wo_total_strict_order: total_strict_order R }. (*Schepler*) Lemma wo_irrefl: forall R:relation T, well_order R -> (forall x:T, ~ R x x). Proof. intuition. assert (forall y:T, Acc R y -> y <> x). intros. induction H1. intuition. rewrite H3 in H2. apply H2 with x. trivial. trivial. pose proof (wo_well_founded R H). unfold well_founded in H2. pose proof (H1 x (H2 x)). auto. Qed. (*Schepler*) Lemma wo_antisym: forall R:relation T, well_order R -> (forall x y:T, R x y -> ~ R y x). Proof. intuition. assert (forall z:T, Acc R z -> z <> x /\ z <> y). intros. induction H2. intuition. rewrite H4 in H3. pose proof (H3 y H1). tauto. rewrite H4 in H3. pose proof (H3 x H0). tauto. pose proof (wo_well_founded R H). unfold well_founded in H3. pose proof (H2 x (H3 x)). tauto. Qed. (*Schepler*) Lemma wo_transitive: forall R:relation T, well_order R -> transitive R. Proof. intros. unfold transitive. intros. case (wo_total_strict_order R H x z). trivial. intro. case H2. intro. rewrite H3 in H0. pose proof (wo_antisym R H y z). contradict H0. auto. intro. assert (forall a:T, Acc R a -> a <> x /\ a <> y /\ a <> z). intros. induction H4. intuition. rewrite H2 in H5. pose proof (H5 z H3). tauto. rewrite H2 in H5. pose proof (H5 x H0). tauto. rewrite H2 in H5. pose proof (H5 y H1). tauto. rewrite H2 in H5. pose proof (H5 z H3). tauto. rewrite H2 in H5. pose proof (H5 x H0). tauto. rewrite H2 in H5. pose proof (H5 y H1). tauto. pose proof (wo_well_founded R H). unfold well_founded in H5. pose proof (H4 x (H5 x)). tauto. Qed. (*Wyckoff*) Definition unq_minimal_element_property (R : relation T) := forall S : Ensemble T, Inhabited S -> exists! x : T, Inn S x /\ (forall y : T, Inn S y -> ~ R y x). (*Wyckoff*) Lemma wo_implies_umep : forall {R:Relation T}, well_order R -> unq_minimal_element_property R. intros R h1. destruct h1 as [h1 h2]. pose proof (WF_implies_MEP _ _ h1) as h3. red in h3. red. intros S h4. specialize (h3 S h4). destruct h3 as [x [h5 h6]]. exists x. red. split. split; auto. intros x' h7. destruct h7 as [h7 h8]. red in h2. specialize (h2 x x'). destruct h2 as [h2 | h2]. specialize (h8 _ h5). contradiction. destruct h2 as [h2 | h2]; auto. specialize (h6 _ h7). contradiction. Qed. (*Wyckoff*) Definition min_set (R:relation T) (pfw:well_order R) (A:Ensemble T) (pfi:Inhabited A) : T := (proj1_sig (constructive_definite_description _ ((wo_implies_umep pfw) A pfi))). (*Wyckoff*) Lemma min_set_compat : forall (R:relation T) (pfw:well_order R) (A:Ensemble T) (pfi:Inhabited A), let m := min_set R pfw A pfi in Inn A m /\ (forall y : T, Inn A y -> ~ R y m). intros R h1 A pfi m. unfold m, min_set. destruct constructive_definite_description. simpl. assumption. Qed. (*Wyckoff*) Lemma min_set_compat' : forall (R:relation T) (pfw:well_order R) (A:Ensemble T) (pfi:Inhabited A) (y:T), R y (min_set R pfw A pfi) -> ~ Inn A y. intros R h1 A h2 y h3. intro h6. pose proof (min_set_compat R h1 A h2) as h4. destruct h4 as [h4 h5]. specialize (h5 _ h6). contradiction. Qed. (*Wyckoff*) Lemma min_set_full_compat : forall (pfdt:type_eq_dec T) (R:relation T) (pfw:well_order R) (pfi:Inhabited (Full_set T)) (y:T), {min_set R pfw _ pfi = y} + {R (min_set R pfw _ pfi) y}. intros h0 R h1 h2 y. pose proof h1 as h1'. destruct (h0 y (min_set R h1 (Full_set T) h2)) as [h3 | h4]. left; auto. destruct h1 as [h1a h1b]. red in h1b. pose proof (h1b y (min_set R {| wo_well_founded := h1a; wo_total_strict_order := h1b |} (Full_set T) h2)) as h5. right. pose proof (wo_antisym). pose proof (min_set_compat R {| wo_well_founded := h1a; wo_total_strict_order := h1b |} (Full_set T) h2) as h6. simpl in h6. destruct h6 as [h6 h7]. specialize (h7 _ (Full_intro _ y)). tauto. Qed. (*Wyckoff*) Lemma min_set_full_compat' : forall (pfdt:type_eq_dec T) (R:relation T) (pfw:well_order R) (pfi:Inhabited (Full_set T)) (y:T), min_set R pfw _ pfi = y \/ R (min_set R pfw _ pfi) y. intros h0 R h1 h2 y. destruct (min_set_full_compat h0 R h1 h2 y). left; auto. right; auto. Qed. (*Wyckoff*) Definition min_elt (R:Relation T) (pfw:well_order R) := min_set R pfw (Full_set T). (*Wyckoff*) Lemma ex_wo_max_set_unq : forall (R:Relation T), well_order R -> forall (S:Ensemble T), ex_max_set R S -> ex_unq_max_set R S. intros R h1 S h2. pose proof h1 as h1'. red in h2. red. destruct h2 as [m h2]. destruct h2 as [h2 h3]. destruct h1 as [h1a h1b]. red in h1b. exists m. red. split. split; auto. intros x h4. destruct h4 as [h4 h5]. specialize (h1b m x). destruct h1b as [h1b | h1b]. specialize (h3 _ h4). destruct h3 as [h3 | h6]. pose proof (wo_antisym _ h1') as h6. specialize (h6 _ _ h1b). contradiction. subst; auto. destruct h1b as [h1b | h1b]; auto. specialize (h5 _ h2). destruct h5 as [h5 | h6]. pose proof (wo_antisym _ h1') as h6. specialize (h6 _ _ h5). contradiction. assumption. Qed. (*Wyckoff*) Definition wo_max_set (R:Relation T) (pfw:well_order R) (S:Ensemble T) (pfe:ex_max_set R S) : T := (proj1_sig (constructive_definite_description _ (ex_wo_max_set_unq R pfw S pfe))). (*Wyckoff*) Lemma wo_max_set_compat : forall (R:Relation T) (pfw:well_order R) (S:Ensemble T) (pfe:ex_max_set R S), let m := wo_max_set R pfw S pfe in Inn S m /\ (forall w : T, Inn S w -> R w m \/ w = m). intros R h1 S h2. simpl. unfold wo_max_set. destruct constructive_definite_description as [m h3]; simpl. assumption. Qed. (*Wyckoff*) Lemma ex_max_set_or_grows : forall (R:Relation T) (pfw:total_strict_order R) (S:Ensemble T), {ex_max_set R S} + {forall x:T, Inn S x -> Inhabited (afterset' R S x)}. intros R h0 S. destruct (classic_dec (ex_max_set R S)) as [h1 | h2]. left; auto. right. intros x h3. unfold ex_max_set in h2. pose proof (not_ex_all_not _ _ h2 x) as h4. clear h2. simpl in h4. apply not_and_or in h4. destruct h4 as [h4 | h5]. contradiction. apply not_all_ex_not in h5. destruct h5 as [y h5]. apply Inhabited_intro with y. constructor. split; auto. apply imply_to_and in h5. destruct h5 as [h5 h6]. apply not_or_and in h6. destruct h6 as [h6 h7]. red in h0. specialize (h0 x y). apply neq_sym in h7. tauto. Qed. (*Wyckoff*) Definition wo_max_elt (R:Relation T) (pfw:well_order R) (pfe:ex_max_set R (Full_set T)) := wo_max_set R pfw (Full_set T) pfe. (*Wyckoff*) Lemma ex_max_set_or_grows_full : forall (R:Relation T) (pfw:total_strict_order R), {ex_max_set R (Full_set T)} + {forall x:T, Inhabited (afterset' R (Full_set T) x)}. intros R h0. destruct (ex_max_set_or_grows R h0 (Full_set _)) as [h1 | h2]. left; auto. right. intro x. apply h2; auto. constructor. Qed. End WellOrder. Implicit Arguments total_strict_order [[T]]. Implicit Arguments well_order [[T]]. Implicit Arguments wo_well_founded [[T] [R]]. Implicit Arguments wo_transitive [[T] [R]]. Implicit Arguments wo_total_strict_order [[T] [R]]. Implicit Arguments wo_irrefl [[T] [R]]. Implicit Arguments wo_antisym [[T] [R]]. Arguments min_elt [T] _ _ _. Section Covers. Variable T:Type. (*Wyckoff*) (* Named without vowels to avoid name-clashing with similar definition in the standard library.*) Definition cvrs {T:Type} (R:Relation T) (x y:T) := R y x /\ ~ (exists z:T, R y z /\ R z x). (*Wyckoff*) Definition ex_cover' {T:Type} (R:Relation T) (x:T) := exists y:T, cvrs R x y. (*Wyckoff*) Definition ex_cover {T:Type} (R:Relation T) (x:T) := exists! y:T, cvrs R x y. (*Wyckoff*) Definition covered_elt {T:Type} (R:Relation T) (x:T) (pf:ex_cover R x) := (proj1_sig ((constructive_definite_description _ pf))). (*Wyckoff*) Lemma covered_elt_compat : forall (R:Relation T) (x:T) (pf:ex_cover R x), let y := covered_elt _ _ pf in cvrs R x y. intros R x pf y. unfold y, covered_elt. destruct constructive_definite_description. simpl. assumption. Qed. (*Wyckoff*) Lemma ex_cover_prime_compat : forall (R:Relation T), total_strict_order R -> forall (x:T), ex_cover R x <-> ex_cover' R x. intros R h0 x. split; intro h1. red in h1. destruct h1 as [y [h1 h2]]. red. exists y. assumption. red in h1. destruct h1 as[y [h1 h2]]. red. exists y. red. split. split; auto. intros y' h3. destruct h3 as [h3 h4]. red in h0. specialize (h0 y y'). destruct h0 as [h0 | h0]. contradict h2. exists y'. split; auto. destruct h0 as [h0 | h0]. subst; auto. contradict h4. exists y. split; auto. Qed. (*Wyckoff*) Lemma covered_elt_lt : forall (R:Relation T) (x:T) (pf:ex_cover R x), R (covered_elt _ _ pf) x. intros R x h1. pose proof (covered_elt_compat _ _ h1) as h2. destruct h2; auto. Qed. (*Wyckoff*) Lemma tso_lt_covered_elt : forall (R:Relation T), total_strict_order R -> forall (a x:T) (pf:ex_cover R x), R a x -> a <> covered_elt _ _ pf -> R a (covered_elt _ _ pf). intros R h0 a x h1 h2 h3. red in h0. specialize (h0 a (covered_elt R x h1)). destruct h0 as [h0 | h0]. assumption. destruct h0 as [h0 | h0]. contradiction. pose proof (covered_elt_compat R x h1) as h4. simpl in h4. destruct h4 as [h4 h5]. pose proof (not_ex_all_not _ (fun z:T => R (covered_elt R x h1) z /\ R z x) h5) as h6. specialize (h6 a). simpl in h6. contradict h6. split; auto. Qed. (*Wyckoff*) Lemma not_ex_cover_min_elt : forall (R:relation T) (pfw:well_order R) (pfi:Inhabited (Full_set T)), ~ex_cover R (min_elt R pfw pfi). intros R h1 h2 h3. red in h3. destruct h3 as [y h3]. red in h3. destruct h3 as [h3 h4]. red in h3. destruct h3 as [h3 h5]. pose proof (min_set_compat _ R h1 (Full_set T) h2) as h6. simpl in h6. destruct h6 as [h6 h7]. specialize (h7 _ (Full_intro _ y)). contradiction. Qed. (*Wyckoff*) Inductive ex_cover_n (R:Relation T) : T-> nat -> Prop := | ex_cover_n_zero : forall (x:T), ex_cover R x -> ex_cover_n R x 0 | ex_cover_n_S : forall (x:T) (n:nat) (pf:ex_cover R x), ex_cover_n R x n -> ex_cover_n R (covered_elt _ _ pf) n -> ex_cover_n R x (S n). (*Wyckoff*) Definition ex_cover_nat (R:Relation T) (x:T) := forall (n:nat), ex_cover_n R x n. (*Wyckoff*) Definition exists_final_seg (R:Relation T) (x:T) : Prop := exists (n:nat), ex_cover_n R x n /\ ~ex_cover_n R x (S n). (*Wyckoff*) Lemma ex_cover_n_bound : forall (R:Relation T) (x:T) (n:nat), ex_cover_n R x n -> forall m:nat, m < n -> ex_cover_n R x m. intros R x n h1. induction h1 as [x h2 | x n h3 h4 h5]. intros; omega. intros m h6. destruct (zerop n) as [h7 | h8]. subst. assert (h7:m = 0). omega. subst. assumption. destruct (zerop m) as [h9 | h10]. subst. constructor. assumption. pose proof (S_pred _ _ h10) as h11. rewrite h11 in h6. apply lt_S_n in h6. pose proof (h5 _ h6) as h12. pose proof (IHh1_1 _ h6) as h13. pose proof (ex_cover_n_S _ _ _ _ h12 h13) as h14. rewrite <- (S_pred _ _ h10) in h14. assumption. Qed. (*Wyckoff*) Definition isolated_from (R:Relation T) (x:T) := [n:nat | ~ex_cover_n R x n]. (*Wyckoff*) Lemma isolated_from_inhabited : forall (R:Relation T) (x:T), ~ex_cover_nat R x -> Inhabited (isolated_from R x). intros R x h1. unfold ex_cover_nat in h1. apply not_all_ex_not in h1. destruct h1 as [n]. apply Inhabited_intro with n. constructor. assumption. Qed. (*Wyckoff*) Definition nat_translatable_isolated_from_nat : forall (R:Relation T) (x:T), Inhabited (isolated_from R x) -> nat_translatable (isolated_from R x). intros R x h1. rewrite nat_translatable_iff'; auto. intros n h2. destruct h2 as [h2]. constructor. intro h3. assert (h4:n < S n). auto with arith. pose proof (ex_cover_n_bound _ _ _ h3 _ h4). contradiction. Qed. (*Wyckoff*) Definition nat_translatable_isolated_from_nat' : forall (R:Relation T) (x:T), Inhabited (isolated_from R x) -> nat_translatable' (isolated_from R x). intros. rewrite <- nat_translatable_prime_compat. apply nat_translatable_isolated_from_nat; auto. Qed. (*Wyckoff*) Lemma max_ex_cover_n_ex : forall (R:Relation T) (x:T), ~ex_cover_nat R x -> ex_cover R x -> exists! (n:nat), ex_cover_n R x n /\ ~ex_cover_n R x (S n). intros R x h2 ha. pose proof (isolated_from_inhabited R x h2) as h3. pose proof (nat_translatable_isolated_from_nat _ _ h3) as h4. apply (nat_translatable_boundary' (ex_cover_n R x)); auto. constructor; auto. Qed. (*Wyckoff*) Definition max_ex_cover_n (R:Relation T) (x:T) (pfnat:~ex_cover_nat R x) (pf:ex_cover R x) := proj1_sig (constructive_definite_description _ (max_ex_cover_n_ex R x pfnat pf)). (*Wyckoff*) Lemma max_ex_cover_n_compat : forall (R:Relation T) (x:T) (pfnat:~ex_cover_nat R x) (pf:ex_cover R x), let n := max_ex_cover_n R x pfnat pf in ex_cover_n R x n /\ ~ex_cover_n R x (S n). intros R x h1 h2. simpl. unfold max_ex_cover_n. destruct constructive_definite_description as [n h3]. simpl. assumption. Qed. End Covers. Section WellOrderedInduction. Definition well_ordered_induction_fun_aux_step {T U:Type} (R:T->T->Prop) (f_S:T->U->U) (f_sup:{x:T & {y:T | R y x}->U}->U) : (forall x : T, (forall y : T, R y x -> U) -> U). intros x h5. destruct (classic_dec (ex_cover R x)) as [h6 | h7]. pose proof (covered_elt_lt _ R x h6) as h8. specialize (h5 _ h8). refine (f_S x h5). refine (f_sup (existT _ x (fun y':{y:T | R y x} => (h5 _ (proj2_sig y'))))). Defined. End WellOrderedInduction. Section WellOrder'. (*Wyckoff*) Lemma total_strict_order_exist_rel : forall {T:Type} (R:relation T), total_strict_order R -> forall (P:T->Prop), total_strict_order (exist_rel R P). intros T R h1 P. red. intros x y. destruct x as [x h2], y as [y h3]. unfold exist_rel; simpl. red in h1. specialize (h1 x y). destruct h1 as [h1 | h1]. left; auto. destruct h1 as [h1 | h1]. subst. right. left. apply proj1_sig_injective; auto. right; auto. Qed. (*Wyckoff*) Lemma well_order_exist_rel : forall {T:Type} (R:relation T), well_order R -> forall (P:T->Prop), well_order (exist_rel R P). intros T R h1 P. destruct h1 as [h1 h2]. constructor. apply well_founded_exist_rel; auto. apply total_strict_order_exist_rel; auto. Qed. (*Wyckoff*) Lemma min_set_exist_rel_eq : forall {T:Type} (R:Relation T) (pfw:well_order R) (b:T) (pfi:Inhabited (Full_set T)) (pfi':Inhabited (Full_set ({t:T | R t b}))) (pfr: R (min_set _ R pfw _ pfi) b), exist (fun a : T => R a b) _ pfr = min_set _ (exist_rel R (fun a : T => R a b)) (well_order_exist_rel R pfw (fun a:T => R a b)) _ pfi'. intros T R h1 b h2 h3 h4. pose proof (wo_total_strict_order (well_order_exist_rel R h1 (fun a:T => R a b))) as h5. red in h5. specialize (h5 (exist (fun a : T => R a b) (min_set _ R h1 (Full_set T) h2) h4) (min_set _ (exist_rel R (fun a : T => R a b)) (well_order_exist_rel R h1 (fun a : T => R a b)) (Full_set {t : T | R t b}) h3)). destruct h5 as [h5 | h5]. pose proof (min_set_compat _ (exist_rel R (fun a : T => R a b)) (well_order_exist_rel R h1 (fun a : T => R a b)) (Full_set {t : T | R t b}) h3) as h6. simpl in h6. destruct h6 as [h6 h7]. specialize (h7 (exist (fun a : T => R a b) _ h4)). hfirst h7. constructor. specialize (h7 hf). contradiction. destruct h5 as [h5 | h5]; auto. pose proof (min_set_compat _ R h1 (Full_set T) h2) as h6. simpl in h6. destruct h6 as [h6 h7]. clear h6. unfold exist_rel in h5. simpl in h5. apply proj1_sig_injective; simpl. gtermr. specialize (h7 y). hfirst h7. constructor. specialize (h7 hf). contradiction. Qed. End WellOrder'. Section WOArith. (*Wyckoff*) Lemma wo_nat : well_order lt. constructor. apply lt_wf. red. intros m n. pose proof (lt_eq_lt_dec m n); tauto. Qed. (*Wyckoff*) Definition min_set_nat (A:Ensemble nat) (pfi:Inhabited A) := min_set _ lt wo_nat A pfi. (*Wyckoff*) Lemma min_set_nat_compat : forall (A:Ensemble nat) (pfi:Inhabited A) (n:nat), Inn A n -> Inn A (min_set _ lt wo_nat A pfi) /\ ~n < min_set _ lt wo_nat A pfi. intros A h1 n h2. pose proof (min_set_compat _ lt wo_nat A h1) as h3. simpl in h3. destruct h3 as [h3 h4]. specialize (h4 n h2). split; auto. Qed. (*Wyckoff*) Lemma min_set_nat_nin : forall (A:Ensemble nat) (pfi:Inhabited A), {min_set_nat A pfi = 0} + {~Inn A (pred (min_set_nat A pfi))}. intros A h1. destruct (zerop (min_set_nat A h1)) as [h2 | h3]. left; auto. right. intro h4. unfold min_set_nat in h4. pose proof (min_set_compat _ lt wo_nat A h1) as h5. simpl in h5. destruct h5 as [h5 h6]. specialize (h6 _ h4). unfold min_set_nat in h3. omega. Qed. (*Wyckoff*) Lemma min_set_nat_max_set_compat : forall (A:Ensemble nat) (pfi:Inhabited A), glb_nat A (min_set_nat A pfi). intros A h1. pose proof (min_set_compat _ lt wo_nat A h1) as h3. simpl in h3. destruct h3 as [h3 h4]. red. split. red. intros n h2. unfold min_set_nat. specialize (h4 _ h2). omega. intros b' h5. red in h5. specialize (h5 _ h3). assumption. Qed. End WOArith. Section WellOrderMinimum. Variable T:Type. Variable R:relation T. Hypothesis well_ord: well_order R. (*Schepler*) Definition WO_minimum: forall S:Ensemble T, Inhabited S -> { x:T | Inn S x /\ forall y:T, Inn S y -> y = x \/ R x y }. refine (fun S H => constructive_definite_description _ _). pose proof (WF_implies_MEP T R (wo_well_founded well_ord)). unfold minimal_element_property in H0. pose proof (H0 S H). destruct H1. destruct H1. exists x. red. split. split. assumption. intros. case (wo_total_strict_order well_ord x y). tauto. intro. case H4. auto. intro. contradict H5. auto. intros. destruct H3. case (wo_total_strict_order well_ord x x'). intro. pose proof (H4 x H1). case H6. trivial. intro. contradict H7. auto. intro. case H5. trivial. intro. contradict H6. auto. Defined. End WellOrderMinimum. Implicit Arguments WO_minimum [[T]]. Section WellOrderConstruction. Variable T:Type. (*Schepler*) Definition restriction_relation (R:relation T) (S:Ensemble T) : relation ({z:T | Inn S z}) := fun (x y:{z:T | Inn S z}) => R (proj1_sig x) (proj1_sig y). Record partial_WO : Type := { pwo_S: Ensemble T; pwo_R: relation T; pwo_R_lives_on_S: forall (x y:T), pwo_R x y -> Inn pwo_S x /\ Inn pwo_S y; pwo_wo: well_order (restriction_relation pwo_R pwo_S) }. (*Schepler*) (*The last condition below says that S WO1 is a downward closed subset of S WO2 *) Record partial_WO_ord (WO1 WO2:partial_WO) : Prop := { pwo_S_incl: Included (pwo_S WO1) (pwo_S WO2); pwo_restriction: forall x y:T, Inn (pwo_S WO1) x -> Inn (pwo_S WO1) y -> (pwo_R WO1 x y <-> pwo_R WO2 x y); pwo_downward_closed: forall x y:T, Inn (pwo_S WO1) y -> Inn (pwo_S WO2) x -> pwo_R WO2 x y -> Inn (pwo_S WO1) x }. (*Schepler*) Lemma partial_WO_preord : preorder partial_WO_ord. Proof. constructor. unfold reflexive. intro. destruct x. constructor; simpl. auto with sets. split. trivial. trivial. auto. unfold transitive. destruct x. destruct y. destruct z. intros. destruct H. destruct H0. simpl in pwo_S_incl0; simpl in pwo_restriction0; simpl in pwo_downward_closed0; simpl in pwo_S_incl1; simpl in pwo_restriction1; simpl in pwo_downward_closed1. constructor; simpl. auto with sets. intros. apply iff_trans with (pwo_R1 x y). apply pwo_restriction0; auto with sets. apply pwo_restriction1; auto with sets. intros. apply pwo_downward_closed0 with y; trivial. apply pwo_downward_closed1 with y; trivial. auto with sets. apply <- (pwo_restriction1 x y); trivial. auto with sets. apply pwo_downward_closed1 with y; trivial. auto with sets. Qed. End WellOrderConstruction. Section WO_implies_AC. (*Schepler*) Lemma WO_implies_AC: forall (A B:Type) (R: A -> B -> Prop) (WO:relation B), well_order WO -> (forall x:A, exists y:B, R x y) -> exists f:A->B, forall x:A, R x (f x). Proof. intros. assert (forall a:A, Inhabited [ b:B | R a b ]). intro. pose proof (H0 a). destruct H1. exists x. constructor; assumption. exists (fun a:A => proj1_sig (WO_minimum WO H [ b:B | R a b ] (H1 a))). intro. destruct @WO_minimum. simpl. destruct a. destruct H2. assumption. Qed. End WO_implies_AC. bool2-v0-3/GaloisConnection.v0000644000175000017500000000502313575574055016767 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2018, Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 Alpha bool2 -- certified -- Coq 8.4 pl4.*) (* This brief experimental file uses some definitions from Chapter 2 in Part II of Dietlinde Lau's "Function Algebras." It's not part of the core development of bool2.*) (*Relation on E_k := {0, . . . . k - 1} of arity [h]*) Require Import NaryFunctions. Require Import TypeUtilities. Require Import SetUtilities. Require Import FunctionAlgebrasBasics. Definition erelation h k : Type := Ensemble ({m | m < k}^h). Lemma finite_erelation : forall {h k} (R:erelation h k), Finite R. intros. eapply Finite_downward_closed. rewrite Finite_FiniteT_iff. apply finite_nprod. apply finite_nat_initial_segment. apply incl_full. Qed. (*This is an "elementary relation," (of arity [h]), which represents a full abstraction of the standard diagonal relation on a segment type.*) Inductive edelta {k} (h:nat) (R:Relation {m | m < k}) : erelation h k := | edelta_intro : forall (x:{m | m < k}^h), (forall i j (pfik:i < k) (pfjk:j < k), i < h -> j < h -> R (exist (fun m => m < k) _ pfjk) (exist (fun m => m < k) _ pfjk) -> x __ i = x __ j) -> Inn (edelta h R) x. (*This is the set of all [edelta]s parameterized by equivalence relations.*) Inductive D'_ (h k:nat) : Ensemble (erelation h k) := | D'__intro : forall {R}, Equivalence _ R -> Inn (D'_ h k) (edelta h R). (*The set of all [D'_]s across all arities, including 0*) Inductive D' (k:nat) : Ensemble {h:nat & (Ensemble (erelation h k))} := D'_intro : forall h, Inn (D' k) (existT _ h (D'_ h k)). bool2-v0-3/Topology.v0000644000175000017500000016746013575574055015363 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-2019, Daniel Wyckoff, except for the portions so labeleled which I got from Daniel Schepler's "Topology" package.*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Import SetUtilities. Require Import Image. Require Import Powerset. Require Import Classical_sets. Require Import Classical. Require Import FunctionalExtensionality. Require Import LogicUtilities. Require Import TypeUtilities. Require Import FunctionProperties. Require Import Basics. Require Import ProofIrrelevance. (*Schepler*) Record TopologicalSpace : Type := { point_set : Type; open : Ensemble point_set -> Prop; open_family_union : forall F : Family point_set, (forall S : Ensemble point_set, In F S -> open S) -> open (FamilyUnion F); open_intersection2: forall U V:Ensemble point_set, open U -> open V -> open (Intersection U V); open_full : open (Full_set point_set) }. Arguments open [_] _. (*Schepler*) Lemma open_empty: forall (X:TopologicalSpace), open (Empty_set (point_set X)). Proof. intros. rewrite <- empty_family_union. apply (open_family_union X). intros. destruct H. Qed. (*Schepler*) Lemma open_union2: forall (X:TopologicalSpace)(U V:Ensemble (point_set X)), open U -> open V -> open (Union U V). Proof. intros. assert (Union U V = FamilyUnion (Couple U V)). apply Extensionality_Ensembles; split; red; intros. destruct H1. exists U; auto with sets. exists V; auto with sets. destruct H1. destruct H1. left; trivial. right; trivial. rewrite H1; apply open_family_union. intros. destruct H2; trivial. Qed. (*Schepler*) Lemma open_indexed_union: forall {X:TopologicalSpace} {A:Type} (F:IndexedFamily A (point_set X)), (forall a:A, open (F a)) -> open (IndexedUnion F). Proof. intros. rewrite indexed_to_family_union. apply open_family_union. intros. destruct H0. rewrite H1; apply H. Qed. (*Schepler*) Lemma open_finite_indexed_intersection: forall {X:TopologicalSpace} {A:Type} (F:IndexedFamily A (point_set X)), FiniteT A -> (forall a:A, open (F a)) -> open (IndexedIntersection F). Proof. intros. induction H. rewrite empty_indexed_intersection. apply open_full. assert (IndexedIntersection F = Intersection (IndexedIntersection (fun x:T => F (Some x))) (F None)). apply Extensionality_Ensembles; split; red; intros. destruct H1. constructor. constructor. intros; apply H1. apply H1. destruct H1. destruct H1. constructor. destruct i. apply H1. apply H2. rewrite H1. apply open_intersection2. apply IHFiniteT. intros; apply H0. apply H0. destruct H1. assert (IndexedIntersection F = IndexedIntersection (fun x:X0 => F (f x))). apply Extensionality_Ensembles; split; red; intros. constructor. destruct H3. intro; apply H3. constructor. destruct H3. intro; rewrite <- H2 with i. apply H3. rewrite H3. apply IHFiniteT. intro; apply H0. Qed. Section Discrete. Variable T : Type. (*Wyckoff*) Definition open_dscrt (S:(Ensemble T)) := True. (*Wyckoff*) Lemma open_family_union_dscrt : forall (F: Family T), (forall S : Ensemble T, In F S -> open_dscrt S) -> open_dscrt (FamilyUnion F). intros. compute. tauto. Qed. (*Wyckoff*) Lemma open_intersection2_dscrt : forall (U V:Ensemble T), open_dscrt U -> open_dscrt V -> open_dscrt (Intersection U V). intros. compute. tauto. Qed. (*Wyckoff*) Lemma open_full_dscrt : open_dscrt (Full_set T). intros. compute. tauto. Qed. (*Wyckoff*) Definition discrete_top := Build_TopologicalSpace T open_dscrt open_family_union_dscrt open_intersection2_dscrt open_full_dscrt. End Discrete. Section Trivial. Variable T: Type. (*Wyckoff*) Hypothesis T_ne : Full_set(T) <> Empty_set T. Definition open_trivial (S:(Ensemble T)) := (S = (Full_set T) \/ S = (Empty_set T)). (*Wyckoff*) Lemma open_family_union_trivial : forall (F: Family T), (forall S : Ensemble T, In F S -> open_trivial S) -> open_trivial (FamilyUnion F). intros F h1. compute. unfold open_trivial in h1. assert (h2: let UF := (FamilyUnion F = Full_set T) in UF \/ ~UF). intro UF. tauto. destruct h2 as [h3 | h4]. (*left*) left. assumption. (*right*) assert (h5: forall S : Ensemble T, In F S -> S <> Full_set T). intros S h6. assert (h7: S = Full_set T \/ S = Empty_set T). apply h1. assumption. destruct h7 as [h8 | h9]. (*left*) assert (h9 : FamilyUnion F = Full_set T). apply Extensionality_Ensembles. unfold Same_set. split. (*left*) unfold Included. intros. apply Full_intro. (*right*) rewrite <- h8. unfold Included. intros x h9. apply family_union_intro with (S := S). assumption. assumption. contradiction. (*right*) intro h10. rewrite h10 in h9. contradiction. assert (h13 : forall (S : Ensemble T), In F S -> S = Empty_set T). intros S h14. assert (h15 : S = Full_set T \/ S = Empty_set T). apply h1. assumption. assert (h16 : S <> Full_set T). apply h5. assumption. tauto. right. apply Extensionality_Ensembles. unfold Same_set. split. (*left*) unfold Included. intros x h17. inversion h17 as [S y h18 h19 h20]. assert (h14 : S = Empty_set T). apply h13. assumption. rewrite h14 in h19. assumption. (*right*) unfold Included. intros x h21. assert (h22 : False). inversion h21. tauto. Qed. (*Wyckoff*) Lemma open_intersection2_trivial : forall (U V:Ensemble T), open_trivial U -> open_trivial V -> open_trivial (Intersection U V). unfold open_trivial. intros U V h1 h2. destruct h1 as [h3|h4]; destruct h2 as [h5|h6]. assert (h7: Intersection U V = U). apply Extensionality_Ensembles. unfold Same_set. unfold Included. split. intros x h7. inversion h7 as [y h8 h9 h10]. assumption. rewrite <- h5 in h3. rewrite h3. intros x h11. apply Intersection_intro. assumption. assumption. rewrite h7. left. apply h3. assert (h7: Intersection U V = V). apply Extensionality_Ensembles. unfold Same_set. unfold Included. split. intros x h7. inversion h7 as [y h8 h9 h10]. assumption. intros x h8. rewrite h6 in h8. inversion h8. rewrite h7. right. assumption. assert (h8: Intersection U V = U). apply Extensionality_Ensembles. unfold Same_set. unfold Included. split. intros x h7. inversion h7 as [y h8 h9 h10]. assumption. intros x h8. rewrite h4 in h8. inversion h8. rewrite h8. right. assumption. assert (h8: Intersection U V = U). apply Extensionality_Ensembles. unfold Same_set. unfold Included. split. intros x h7. inversion h7 as [y h8 h9 h10]. assumption. intros x h8. rewrite h4 in h8. inversion h8. rewrite h8. right. assumption. Qed. (*Wyckoff*) Lemma open_full_trivial : open_trivial (Full_set T). compute. left. reflexivity. Qed. (*Wyckoff*) Definition trivial_top := Build_TopologicalSpace T open_trivial open_family_union_trivial open_intersection2_trivial open_full_trivial. End Trivial. Section Cofinite. Variable T:Type. (*Wyckoff*) Definition open_cof (S:Ensemble T) := (S = Empty_set _) \/ Finite (Comp S). (*Wyckoff*) Lemma open_family_union_cof : forall F : Family T, (forall S : Ensemble T, In F S -> open_cof S) -> open_cof (FamilyUnion F). red. intros F h1. red in h1. pose (h2 := (forall S:Ensemble T, In F S -> S = Empty_set _)). assert (h3: h2 \/ (exists S:Ensemble T, In F S /\ Finite (Comp S))). assert (h3: h2 \/ ~h2). tauto. destruct h3 as [h4|h5]. tauto. unfold h2 in h5. pose proof (not_all_ex_not _ _ h5) as h6. elim h6. intros S h7. right. exists S. assert (h9: S <> Empty_set _). intro h10. tauto. assert (h10: ~~(In F S)). tauto. pose proof (NNPP _ h10) as h11. split. assumption. pose proof (h1 S h11) as h12. destruct h12. contradiction. assumption. destruct h3 as [h12 | h13]. left. apply Extensionality_Ensembles. unfold Same_set. unfold Included. split. intros S h14. inversion h14. unfold h2 in h12. pose proof (h12 S0 H). rewrite H2 in H0. contradiction. intros; contradiction. right. elim h13. intros S h15. elim h15. intros h16 h17. assert (h18: Included S (FamilyUnion F)). unfold Included. intros x h18. apply family_union_intro with (S). assumption. assumption. pose proof (complement_inclusion _ _ h18) as h19. apply Finite_downward_closed with (Comp S). assumption. assumption. Qed. (*Wyckoff*) Lemma open_intersection2_cof: forall U V:Ensemble T, open_cof U -> open_cof V -> open_cof (Intersection U V). intros U V h1 h2. red in h1. red in h2. destruct h1 as [h3 | h4]. destruct h2 as [h5 | h6]. (*case 1*) assert (h7: Intersection U V = Empty_set _ ). rewrite h3. rewrite h5. auto with sets. red. left. assumption. (*case 2*) assert (h8: Intersection U V = Empty_set _ ). rewrite h3. auto with sets. red. left. assumption. (*case 3*) red. destruct h2 as [h10 | h11]. left. rewrite h10. auto with sets. (*case 4*) right. pose proof (comp_int _ U V) as h12. rewrite h12. apply Union_preserves_Finite. assumption. assumption. Qed. Lemma open_full_cof : open_cof (Full_set T). red. right. pose proof (complement_full T) as h1. rewrite h1. constructor. Qed. (*Wyckoff*) Definition cof_top := Build_TopologicalSpace T open_cof open_family_union_cof open_intersection2_cof open_full_cof. End Cofinite. Section Top_Space_Equiv. (*Wyckoff*) Definition set_iso (T1 T2 : Type) (heq : T1 = T2) (S1 : Ensemble T1) : Ensemble T2. rewrite heq in S1. refine S1. Defined. (*Wyckoff*) Definition ts_eq (X1 X2 : TopologicalSpace) : Prop := exists heq:((point_set X1) = (point_set X2)), (forall (S1:Ensemble (point_set X1)) (S2:Ensemble (point_set X2)), S2 = set_iso _ _ heq S1 -> ((open S1) <-> (open S2))). End Top_Space_Equiv. Section Cofinite_Discrete_Finite. Variable T : Type. Definition X_dt := discrete_top T. Definition X_ct := cof_top T. (*Wyckoff*) Lemma dscrt_cof_same_fin : FiniteT T -> ts_eq X_dt X_ct. intro h1. red. simpl. exists (refl_equal T). intros S1 S2 h2. unfold open_dscrt. unfold open_cof. split. (*left*) intro h3. right. apply FiniteT_Finite. assumption. (*right*) intros. auto. Qed. End Cofinite_Discrete_Finite. Section Inherited. Variable X:TopologicalSpace. Variable Y:(Ensemble (point_set X)). (*Wyckoff*) Definition Yt := sig_set Y. (*Wyckoff*) Inductive open_inh (P:(Ensemble Yt)) : Prop := open_inh_intro : forall (Q:Ensemble (point_set X)), open Q -> im_proj1_sig P = Intersection Y Q -> open_inh P. (*This returns the set of all open sets whose intersection with [Y] yield [P]. Taking all of these instead of just one is crucial to avoid the Axiom of Choice / Indefinite Description in [open_fmaily_union_inh].*) Definition open_inh_family (P:Ensemble Yt) : Family (point_set X) := [Q:Ensemble (point_set X) | open Q /\ im_proj1_sig P = Intersection Y Q]. Lemma open_inh_family_open : forall (P:Ensemble Yt), open (FamilyUnion (open_inh_family P)). intros P. apply open_family_union. intros S h1. destruct h1 as [h1]. destruct h1; auto. Qed. Definition im_open_inh_family (F:Family Yt) : Family (point_set X) := Im F (fun P => FamilyUnion (open_inh_family P)). Lemma union_im_open_inh_family_open : forall F:Family Yt, open (FamilyUnion (im_open_inh_family F)). intro F. apply open_family_union. intros S h1. inversion h1; subst. apply open_family_union. intros S h2. destruct h2 as [h2]; destruct h2; auto. Qed. (*Wyckoff*) Definition Fam_sig (F:Family Yt) := sig_set F. (*Wyckoff*) Definition F_to_Fs (F:Family Yt) := Full_set (Fam_sig F). (*Wyckoff*) (*This is a bit stronger than the proof in [Givant/Halmos], in that, once again, no Axiom of Choice is used.*) Lemma open_family_union_inh : forall F:Family Yt, (forall S : (Fam_sig F), open_inh (proj1_sig S)) -> open_inh (FamilyUnion F). intros F h1. pose proof (open_inh_intro (FamilyUnion F) (FamilyUnion (im_open_inh_family F))) as h2. hfirst h2. apply open_family_union. intros S h3. destruct h3; subst. apply open_family_union. intros S h4. destruct h4 as [h4]. destruct h4; auto. specialize (h2 hf). apply h2; clear h2. apply Extensionality_Ensembles; red; split; red; intros x h2. inversion h2; subst; clear h2. destruct H as [S y h2 h3]. constructor. apply proj2_sig. specialize (h1 (exist _ _ h2)). inversion h1; clear h1; subst; simpl in H0. econstructor. econstructor. apply h2. reflexivity. econstructor. constructor. split. apply H. auto. pose proof (in_im_proj1_sig_iff Y S y) as h4. rewrite <- h4 in h3. rewrite H0 in h3. inversion h3; subst; auto. destruct h2 as [y h2 h3]. inversion h3; subst; clear h3. inversion H; clear H; subst. rename x into S. inversion H0; clear H0; subst. inversion H; clear H; subst. destruct H0 as [h3 h4]. specialize (h1 (exist _ _ H1)). inversion h1; clear h1; simpl in H0. apply Im_intro with (exist _ _ h2). econstructor. apply H1. pose proof (Intersection_intro _ Y S0 y h2 H2) as h5. rewrite <- h4 in h5. inversion h5; subst; clear h5. destruct x as [x hx]; simpl. pose proof (proof_irrelevance _ h2 hx); subst; auto. simpl; auto. Qed. (*Wyckoff*) Lemma open_family_union_inh' : forall F : Family Yt, (forall S : Ensemble Yt, In F S -> open_inh S) -> open_inh (FamilyUnion F). intros F h1. apply open_family_union_inh. intro S'. apply (h1 (proj1_sig S') (proj2_sig S')). Qed. (*Wyckoff*) Lemma open_intersection2_inh: forall U V:Ensemble Yt, open_inh U -> open_inh V -> open_inh (Intersection U V). intros U V h1 h2. destruct h1, h2. exists (Intersection Q Q0). apply (open_intersection2 X); auto. pose proof (Intersection_idempotent Y) as h10. rewrite <- h10 at 2. rewrite Intersection_associative at 1. rewrite <- (Intersection_associative Y Q Q0). rewrite (Intersection_commutative _ Y Q) at 1. rewrite Intersection_associative at 1. rewrite <- H2. rewrite <- Intersection_associative at 1. rewrite <- H0. rewrite intersection_im_proj1_sig_eq; auto. Qed. (*Wyckoff*) Lemma open_full_inh : open_inh (Full_set Yt). econstructor. apply open_full. pose proof (im_proj1_sig_full_set_sig Y) as h1. rewrite h1 at 1. symmetry. rewrite <- inclusion_iff_intersection_eq. red; intros; constructor. Qed. (*Wyckoff*) Definition inh_top := Build_TopologicalSpace Yt open_inh open_family_union_inh' open_intersection2_inh open_full_inh. End Inherited. Section Closed. Variable X : TopologicalSpace. (*Schepler*) Definition closed (F:Ensemble (point_set X)) := open (Comp F). (*Schepler*) Lemma closed_complement_open: forall (U:Ensemble (point_set X)), closed (Comp U) -> open U. Proof. intros. red in H. rewrite Complement_Complement in H. assumption. Qed. (*Schepler*) Lemma closed_empty : closed (Empty_set (point_set X)). red. pose proof (complement_empty (point_set X)) as h1. rewrite h1. apply open_full. Qed. (*Schepler*) Lemma closed_full : closed (Full_set (point_set X)). red. pose proof (complement_full (point_set X)) as h1. rewrite h1. apply open_empty. Qed. (*Schepler*) Lemma closed_union2: forall (F G:Ensemble (point_set X)), closed F -> closed G -> closed (Union F G). Proof. intros. red in H, H0. red. assert (Comp (Union F G) = Intersection (Comp F) (Comp G)). unfold Comp. apply Extensionality_Ensembles; split; red; intros. constructor. auto with sets. auto with sets. destruct H1. red; red; intro. destruct H3. apply (H1 H3). apply (H2 H3). rewrite H1. apply open_intersection2; assumption. Qed. (*Schepler*) Lemma closed_intersection2: forall (F G:Ensemble (point_set X)), closed F -> closed G -> closed (Intersection F G). Proof. intros. red in H, H0. red. assert (Comp (Intersection F G) = Union (Comp F) (Comp G)). apply Extensionality_Ensembles; split; red; intros. apply NNPP. red; intro. unfold Comp in H1. unfold In in H1. contradict H1. constructor. apply NNPP. red; intro. auto with sets. apply NNPP. red; intro. auto with sets. red; red; intro. destruct H2. destruct H1; auto with sets. rewrite H1; apply open_union2; trivial. Qed. (*Schepler*) Lemma closed_family_intersection: forall (F:Family (point_set X)), (forall S:Ensemble (point_set X), In F S -> closed S) -> closed (FamilyIntersection F). Proof. intros. unfold closed in H. red. assert (Comp (FamilyIntersection F) = FamilyUnion [ S:Ensemble (point_set X) | In F (Comp S) ]). apply Extensionality_Ensembles; split; red; intros. apply NNPP. red; intro. red in H0; red in H0. contradict H0. constructor. intros. apply NNPP. red; intro. contradict H1. exists (Comp S). constructor. rewrite Complement_Complement; assumption. assumption. destruct H0. red; red; intro. destruct H2. destruct H0. pose proof (H2 _ H0). contradiction H3. rewrite H0; apply open_family_union. intros. destruct H1. pose proof (H _ H1). rewrite Complement_Complement in H2; assumption. Qed. (*Schepler*) Lemma closed_indexed_intersection: forall {A:Type} (F:IndexedFamily A (point_set X)), (forall a:A, closed (F a)) -> closed (IndexedIntersection F). Proof. intros. rewrite indexed_to_family_intersection. apply closed_family_intersection. intros. destruct H0. rewrite H1; trivial. Qed. (*Schepler*) Lemma closed_finite_indexed_union: forall {A:Type} (F:IndexedFamily A (point_set X)), FiniteT A -> (forall a:A, closed (F a)) -> closed (IndexedUnion F). Proof. intros. red. assert (Comp (IndexedUnion F) = IndexedIntersection (fun a:A => Comp (F a))). apply Extensionality_Ensembles; split; red; intros. constructor. intros. red; red; intro. contradiction H1. exists i. assumption. destruct H1. red; red; intro. destruct H2. contradiction (H1 i). rewrite H1; apply open_finite_indexed_intersection; trivial. Qed. End Closed. Arguments closed [_] _. Section Closed_Discrete. Variable T:Type. (*Wyckoff*) Lemma all_closed_dscrt : forall S:Ensemble (point_set (X_dt T)), closed S. intro S. red. simpl. red. auto. Qed. End Closed_Discrete. Section Closed_Cofinite. Variable T:Type. (*Wyckoff*) Lemma closed_cof_fin : (forall S:Ensemble (point_set (X_ct T)), closed S <-> (Finite S) \/ (S = Full_set _)). intro S. split. (* -> *) intro h1. red in h1. simpl in h1. red in h1. destruct h1 as [h2 | h3]. (*h2*) right. simpl. pose proof (complement_empty T) as h4. rewrite <- h2 in h4. rewrite Complement_Complement in h4. assumption. (*h3*) left. rewrite Complement_Complement in h3. simpl. assumption. (* <- *) intro h5. simpl. red. destruct h5 as [h6 | h7]. right. simpl in h6. rewrite <- Complement_Complement in h6. assumption. left. simpl in h7. pose proof (complement_full T) as h8. rewrite <- h7 in h8. assumption. Qed. End Closed_Cofinite. Section Clopen. Variable X:TopologicalSpace. (*Wyckoff*) Definition clopen (S:Ensemble (point_set X)) := open S /\ closed S. (*Wyckoff*) Lemma clopen_full : clopen (Full_set (point_set X)). red. split. apply open_full. apply closed_full. Qed. (*Wyckoff*) Lemma clopen_empty : clopen (Full_set (point_set X)). red. split. apply open_full. apply closed_full. Qed. End Clopen. Arguments clopen [_] _. Section Clopen_Discrete. (*Wyckoff*) Lemma clopen_discrete : forall (T:Type) (S:(Ensemble (point_set (X_dt T)))), clopen S. intros T S. simpl in S. red. split. simpl. red. auto. simpl. red. auto. Qed. End Clopen_Discrete. (*This section is copied and pasted directly from the file InteriorsClosures.v from Daniel Schepler's Topology user contribution package *) Section interior_closure. Variable X:TopologicalSpace. Variable S:Ensemble (point_set X). (*Schepler*) Definition interior := FamilyUnion [ U:Ensemble (point_set X) | open U /\ Included U S ]. (*Schepler*) Definition closure := FamilyIntersection [ F:Ensemble (point_set X) | closed F /\ Included S F ]. (*Schepler*) Lemma interior_open : open interior. Proof. apply open_family_union. intros. destruct H as [[]]; trivial. Qed. (*Schepler*) Lemma interior_deflationary : Included interior S. Proof. red; intros. destruct H. destruct H as [[]]; auto with sets. Qed. (*Schepler*) Lemma interior_fixes_open: open S -> interior = S. Proof. intros. apply Extensionality_Ensembles; split. apply interior_deflationary. red; intros. exists S; trivial. constructor; auto with sets. Qed. (*Schepler*) Lemma interior_maximal: forall U:Ensemble (point_set X), open U -> Included U S -> Included U interior. Proof. intros. red; intros. exists U; trivial. constructor; split; trivial. Qed. (*Schepler*) Lemma closure_closed : closed closure. Proof. apply closed_family_intersection. intros. destruct H as [[]]; trivial. Qed. (*Schepler*) Lemma closure_inflationary : Included S closure. Proof. red; intros. constructor; intros. destruct H0 as [[]]. auto with sets. Qed. (*Schepler*) Lemma closure_fixes_closed : closed S -> closure = S. Proof. intro. apply Extensionality_Ensembles; split. red; intros. destruct H0. apply H0; split; auto with sets. apply closure_inflationary. Qed. (*Schepler*) Lemma closure_minimal: forall F:Ensemble (point_set X), closed F -> Included S F -> Included closure F. Proof. intros; red; intros. destruct H1. apply H1. constructor; split; trivial. Qed. End interior_closure. (*Wyckoff*) Notation "P 'o'" := (interior _ P) (at level 50). Notation "P -" := (closure _ P) (at level 50). Notation "P '" := (Comp P) (at level 50). Section interior_examples. Variable T:Type. (*Wyckoff*) Example discrete_int_eq : forall S:Ensemble (point_set (X_dt T)), S o = S. intro S. simpl in S. unfold interior. apply Extensionality_Ensembles. red. split. (*left*) red. intros x h2. inversion h2 as [U y h3 h4 h5]. inversion h3 as [h6]. destruct h6 as [h7 h8]. auto with sets. (*right*) simpl. red. intros x h1. apply family_union_intro with (S). constructor; split; red; trivial. assumption. Qed. (*Wyckoff*) (*In a space with the cofinite topology, the interior of a cofinite subset is itself, and the interior of every other subset is empty.*) Example cof_int_eq_empty : forall S : Ensemble (point_set (X_ct T)), (Finite (S ') -> S o = S) /\ (~ Finite (S ') -> S o = Empty_set _). intro S. simpl in S. split. (*left*) intro h1. simpl in h1. unfold interior. simpl. apply Extensionality_Ensembles. red. split. (* left *) unfold Included. intros x h2. inversion h2 as [U y h3]. inversion h3 as [h4]. destruct h4. auto. (* right *) unfold Included. intros x h5. apply family_union_intro with (S). constructor. split. (* left *) red. right; assumption. (* right *) auto. (* right *) assumption. (*right*) simpl. intro h6. apply Extensionality_Ensembles. red. split. (* left *) red. intros x h7. unfold interior in h7. inversion h7. inversion H. destruct H2. simpl in H2. red in H2. destruct H2. (*left*) rewrite H2 in H0. contradiction. (*right*) pose proof (complement_inclusion _ _ H3). pose proof (Finite_downward_closed _ _ H2). pose proof (H5 (S ')). pose proof (H6 H4). contradiction. (* right *) auto with sets. Qed. End interior_examples. Section closure_examples. Variable T:Type. (*Wyckoff*) Example discrete_clo_eq : forall S:Ensemble (point_set (X_dt T)), S - = S. intro S. simpl in S. unfold closure. apply Extensionality_Ensembles. red. split. (*left*) red. simpl. intros x h1. inversion h1 as [y h2]. apply h2. constructor. split. red. trivial. auto with sets. (*right*) red. simpl. intros x h2. constructor. intros S0 h3. inversion h3 as [h4]. destruct h4 as [[]h6]. auto with sets. Qed. (*Wyckoff*) Example cof_clo_eq_empty : forall S : Ensemble (point_set (X_ct T)), (Finite S -> S - = S) /\ (~ Finite S -> S - = Full_set _). simpl. intro S. split. (*left*) intro h1. unfold closure. simpl. apply Extensionality_Ensembles. red. split. (*left*) red. intros x h2. inversion h2 as [y h3]. apply h3. constructor. split. (*left*) red. right. pose proof (Complement_Complement _ S) as h4. rewrite h4. assumption. (*right*) auto with sets. (*right*) red. intros x h5. constructor. intros S0 h6. inversion h6 as [h7]. destruct h7 as [[]h8]; auto with sets. (*right*) intro h8. unfold closure. apply Extensionality_Ensembles. red. split. (*left*) red. intros x h9. constructor. (*right*) red. intros x h10. constructor. intros S0 h11. inversion h11 as [h12]. destruct h12 as [h13 h14]. red in h13. simpl in h13. red in h13. case h13. (*left*) intro h15. pose proof (complement_empty T) as h16. rewrite <- h15 in h16. pose proof (Complement_Complement _ S0) as h17. simpl in h17. rewrite h16 in h17. rewrite <- h17. auto with sets. (*right*) intro h18. pose proof (Complement_Complement _ S0) as h19. simpl in h19. rewrite h19 in h18. pose proof (Finite_downward_closed _ _ h18) as h20. pose proof (h20 S h14) as h21. contradiction. Qed. End closure_examples. Arguments interior [_] _ _. Arguments closure [_] _ _. Arguments interior_open [_] _. Arguments interior_deflationary [_] _ _ _. Arguments interior_fixes_open [_] _ _. Arguments interior_maximal [_] _ _ _ _ _ _. Arguments closure_closed [_] _. Arguments closure_inflationary [_] _ _ _. Arguments closure_fixes_closed [_] _ _. Arguments closure_minimal [_] _ _ _ _ _ _. (*This is also copied and pasted from the same file as mentioned above previous section *) Section interior_closure_relations. Variable X:TopologicalSpace. (*Schepler*) Lemma interior_increasing: forall S T:Ensemble (point_set X), Included S T -> Included (interior S) (interior T). Proof. intros. apply interior_maximal. apply interior_open. assert (Included (interior S) S). apply interior_deflationary. auto with sets. Qed. (*Schepler*) Lemma interior_intersection: forall S T:Ensemble (point_set X), interior (Intersection S T) = Intersection (interior S) (interior T). Proof. intros. apply Extensionality_Ensembles; split. assert (Included (interior (Intersection S T)) (interior S)). apply interior_increasing. auto with sets. assert (Included (interior (Intersection S T)) (interior T)). apply interior_increasing. auto with sets. auto with sets. apply interior_maximal. apply open_intersection2; apply interior_open. pose proof (interior_deflationary S). pose proof (interior_deflationary T). red; intros x H1; constructor; (apply H || apply H0); destruct H1; trivial. Qed. (*Schepler*) Lemma interior_union: forall S T:Ensemble (point_set X), Included (Union (interior S) (interior T)) (interior (Union S T)). Proof. intros. apply interior_maximal. apply open_union2; apply interior_open. pose proof (interior_deflationary S). pose proof (interior_deflationary T). auto with sets. Qed. (*Schepler*) Lemma interior_complement: forall S:Ensemble (point_set X), interior (Comp S) = Comp (closure S). Proof. intros. apply Extensionality_Ensembles; split. rewrite <- Complement_Complement with (A:=interior (Comp S)). apply complement_inclusion. apply closure_minimal. red. rewrite Complement_Complement. apply interior_open. pattern S at 1; rewrite <- Complement_Complement with (A:=S). apply complement_inclusion. apply interior_deflationary. apply interior_maximal. apply closure_closed. apply complement_inclusion. apply closure_inflationary. Qed. (*Schepler*) Lemma closure_increasing: forall S T:Ensemble (point_set X), Included S T -> Included (closure S) (closure T). Proof. intros. apply closure_minimal. apply closure_closed. pose proof (closure_inflationary T). auto with sets. Qed. (*Schepler*) Lemma closure_complement: forall S:Ensemble (point_set X), closure (Comp S) = Comp (interior S). Proof. intros. pose proof (interior_complement (Comp S)). rewrite Complement_Complement in H. rewrite H. rewrite Complement_Complement; reflexivity. Qed. (*Schepler*) Lemma closure_union: forall S T:Ensemble (point_set X), closure (Union S T) = Union (closure S) (closure T). Proof. intros. apply Extensionality_Ensembles; split. apply closure_minimal. apply closed_union2; apply closure_closed. pose proof (closure_inflationary S). pose proof (closure_inflationary T). auto with sets. assert (Included (closure S) (closure (Union S T))). apply closure_increasing; auto with sets. assert (Included (closure T) (closure (Union S T))). apply closure_increasing; auto with sets. auto with sets. Qed. (*Schepler*) Lemma closure_union_finite : forall (F:Family (point_set X)), Finite F -> closure (FamilyUnion F) = FamilyUnion (Im F (@closure X)). intros F h1. induction h1. (*Empty case*) assert (h2: FamilyUnion (Empty_set (Ensemble (point_set X))) - = Empty_set _). apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h3. pose proof (empty_family_union (point_set X)) as h4. rewrite h4 in h3. pose proof (closure_fixes_closed (Empty_set (point_set X))) as h5. pose proof (h5 (closed_empty X)) as h6. rewrite h6 in h3. contradiction. (* >= *) red. intros x h1. contradiction. rewrite h2. apply Extensionality_Ensembles. red. split. (*<=*) red. intros; contradiction. (*>=*) red. intros x h3. inversion h3 as [S y h4]. pose proof (image_empty _ _ (@closure X)) as h5. rewrite h5 in h4. contradiction. (*"Add" case*) rename x into S. rename A into F. pose proof (family_union_add F S) as h2. pose proof (f_equal (@closure X) h2) as h3. pose proof (closure_union S (FamilyUnion F)) as h4. rewrite IHh1 in h4. rewrite h3. rewrite h4. pose proof (family_union_add (Im F (@closure X)) (S -)) as h5. rewrite <- h5. pose proof (Im_add _ _ F S (@closure X)) as h6. apply f_equal; symmetry; assumption. Qed. (*Schepler*) Lemma closure_intersection: forall S T:Ensemble (point_set X), Included (closure (Intersection S T)) (Intersection (closure S) (closure T)). Proof. intros. assert (Included (closure (Intersection S T)) (closure S)). apply closure_increasing; auto with sets. assert (Included (closure (Intersection S T)) (closure T)). apply closure_increasing; auto with sets. auto with sets. Qed. (*Schepler*) Lemma meets_every_open_neighborhood_impl_closure: forall (S:Ensemble (point_set X)) (x:point_set X), (forall U:Ensemble (point_set X), open U -> In U x -> Inhabited (Intersection S U)) -> In (closure S) x. Proof. intros. apply NNPP; intro. pose proof (H (Comp (closure S))). destruct H1. apply closure_closed. exact H0. destruct H1. contradiction H2. apply closure_inflationary. assumption. Qed. (*Schepler*) Lemma closure_impl_meets_every_open_neighborhood: forall (S:Ensemble (point_set X)) (x:point_set X), In (closure S) x -> forall U:Ensemble (point_set X), open U -> In U x -> Inhabited (Intersection S U). Proof. intros. apply NNPP; intro. assert (Included (closure S) (Comp U)). apply closure_minimal. unfold closed. rewrite Complement_Complement; assumption. red; intros. intro. contradiction H2. exists x0; constructor; trivial. contradict H1. apply H3. assumption. Qed. (*Schepler*) Definition dense (S:Ensemble (point_set X)) : Prop := closure S = Full_set _. (*Schepler*) Lemma dense_meets_every_nonempty_open: forall S:Ensemble (point_set X), dense S -> forall U:Ensemble (point_set X), open U -> Inhabited U -> Inhabited (Intersection S U). Proof. intros. destruct H1. apply closure_impl_meets_every_open_neighborhood with x; trivial. rewrite H; constructor. Qed. (*Schepler*) Lemma meets_every_nonempty_open_impl_dense: forall S:Ensemble (point_set X), (forall U:Ensemble (point_set X), open U -> Inhabited U -> Inhabited (Intersection S U)) -> dense S. Proof. intros. apply Extensionality_Ensembles; split; red; intros. constructor. apply meets_every_open_neighborhood_impl_closure. intros. apply H; trivial. exists x; trivial. Qed. End interior_closure_relations. Arguments interior_increasing [_] _ _ _ _ _. Arguments interior_intersection [_] _ _. Arguments interior_union [_] _ _ _ _. Arguments interior_complement [_] _. Arguments closure_increasing [_] _ _ _ _ _. Arguments closure_complement [_] _. Arguments closure_union [_] _ _. Arguments closure_intersection [_] _ _ _ _. (*From Chapter 9 in Givant/Halmos*) Section more_interior_closure_relations1. Variable X:TopologicalSpace. (*Wyckoff*) Lemma closure_eq1 : forall (P:Ensemble (point_set X)), P - = P ' o '. intro P. apply Extensionality_Ensembles. red. split. (*left*) pose proof (interior_open (P ')) as h1. pose proof (interior_deflationary (P ')) as h2. pose proof (complement_inclusion ((P ') o) (P ')) as h3. pose proof (h3 h2) as h4. pose proof (Complement_Complement _ P) as h5. rewrite h5 in h4. assert (h6:(((P ') o) ') - = ((P ') o) '). unfold closure. apply Extensionality_Ensembles. unfold Same_set. split. (*left*) red. intros x h7. inversion h7 as [y h8]. pose proof (h8 P) as h9. assert (h10:closed (((P ') o )')). red. pose proof (Complement_Complement _ ((P ') o)) as h11. rewrite h11. assumption. assert (h11:Included (((P ') o) ') (((P ' ) o) ')). auto with sets. pose proof (conj h10 h11) as h12. assert (h13:In [F : Ensemble (point_set X) | closed F /\ Included (((P ') o) ') F] (((P ') o) ')). constructor; assumption. auto with sets. (*rigbt*) red. intros x h14. constructor. intros S h15. inversion h15 as [h16]. destruct h16 as [h17 h18]. auto with sets. pose proof (closure_closed (((P ') o) ')) as h19. rewrite h6 in h19. apply (closure_minimal _ _ h19 h4). (*right*) unfold closure. intros x h20. constructor. intros Q h21. inversion h21 as [h22]. destruct h22 as [h23 h24]. pose proof (complement_inclusion _ _ h24) as h25. red in h23. pose proof (interior_maximal _ _ h23 h25) as h26. pose proof (complement_inclusion _ _ h26) as h27. pose proof (Complement_Complement _ Q) as h28. rewrite h28 in h27. auto with sets. Qed. (*Wyckoff*) Lemma interior_eq: forall (P:Ensemble (point_set X)), P o = P ' - '. intro P. pose proof (closure_eq1 (P ')) as h1. pose proof (Complement_Complement _ P) as h2. rewrite h2 in h1. rewrite h1. simpl. rewrite Complement_Complement. reflexivity. Qed. Arguments dense [_] _. (*Wyckoff*) (*The reverse implication is above somewhere *) Lemma dense_impl_meets_every_nonempty_open: forall S:Ensemble (point_set X), dense S -> (forall U:Ensemble (point_set X), open U -> Inhabited U -> Inhabited (Intersection S U)). intros S h1 U h2 h3. red in h1. inversion h3 as [x h4]. pose proof (Full_intro _ x) as h5. rewrite <- h1 in h5. apply closure_impl_meets_every_open_neighborhood with x; assumption. Qed. (*Wyckoff*) Definition dense_in (P Q:(Ensemble (point_set X))) := Included Q (P -). Section Dense_Examples. Variable T:Type. (* In a discrete space, no proper subset of the space is dense.*) (*Wyckoff*) Lemma dense_discrete_none : let Xt := (point_set (X_dt T)) in forall S:(Ensemble Xt), S <> (Full_set Xt) -> ~dense S. intros Xt S h0. red. intro h1. red in h1. unfold closure in h1. unfold closed in h1. simpl in h1. assert (h2: open_dscrt _ (S ') /\ Included S S). split. (*left*) red; trivial. (*right*) auto with sets. assert (h4:Included (Full_set T) S). rewrite <- h1. red. intros x h4. inversion h4 as [y h5]. pose proof (h5 S) as h6. assert (h7:In [F : Ensemble T | open_dscrt T (F ') /\ Included S F] S). constructor. apply h2. apply (h6 h7). assert (h8 : S = (Full_set T)). apply Extensionality_Ensembles. red. split. (*left*) red. intros x h9. constructor. (*right*) assumption. contradiction. Qed. (*Wyckoff*) (*In an infinite space with cofinite topology, every infinite subset of the space is dense. *) Lemma infnt_cof_infnt_dense : let Xt := (point_set (X_ct T)) in forall S:(Ensemble (point_set (X_ct T))), ~ FiniteT T -> ~ Finite S -> dense S. intros Xt S h1 h2. red. simpl in h2. simpl. apply cof_clo_eq_empty. simpl. assumption. Qed. End Dense_Examples. (*Wyckoff*) (*A set P is defined to be nowhere dense if it is not dense in any non-empty open set.*) Definition now_dense (P: (Ensemble (point_set X))) : Prop := forall S:(Ensemble (point_set X)), open S -> Inhabited S -> ~ dense_in P S. (*Wyckoff*) Lemma closure_empty : (Empty_set (point_set X)) - = Empty_set (point_set X). apply Extensionality_Ensembles. red. split. (*left*) apply closure_minimal. apply closed_empty. auto with sets. (*right*) auto with sets. Qed. (*Wyckoff*) Lemma interior_empty : (Empty_set (point_set X)) o = Empty_set (point_set X). unfold interior. apply Extensionality_Ensembles. red. split. (*left*) red. intros x h1. inversion h1 as [U y h2 h3]. inversion h2 as [h4]. destruct h4 as [h5 h6]. auto with sets. (*right*) red. intros x h7. contradiction. Qed. (*Wyckoff*) Lemma closure_full : (Full_set (point_set X)) - = Full_set (point_set X). Proof. apply Extensionality_Ensembles. red. split. (*left*) unfold closure. red. intros. constructor. (*right*) apply closure_inflationary. Qed. (*Wyckoff*) Lemma interior_full : (Full_set (point_set X)) o = Full_set (point_set X). apply Extensionality_Ensembles. red. split. (*left*) apply interior_deflationary. (*right*) unfold interior. red. intros x h1. apply family_union_intro with (Full_set (point_set X)). constructor. split. (*left*) apply open_full. (*right*) auto with sets. constructor. Qed. (*Wyckoff*) (*Nowhere dense means that the interior of the closure of P is empty.*) Lemma now_dense_iff1 : forall P : (Ensemble (point_set X)), now_dense P <-> P - o = Empty_set (point_set X). intros P. split. (*left*) intro h1. apply Extensionality_Ensembles. red. split. (*left*) (*red. intros x h2. *) red in h1. unfold interior. unfold dense_in in h1. assert (h2:forall U:Ensemble (point_set X), open U -> Included U (P -) -> U = (Empty_set _)). intros U h3 h4. case (classic (U = (Empty_set _))); trivial. intro h5. pose proof (not_empty_Inhabited _ U h5) as h6. pose proof (h1 U h3 h6) as h7. contradiction. red. intros x h8. inversion h8 as [U y h9 h10]. inversion h9 as [h11]. destruct h11 as [h12 h13]. pose proof (h2 U h12 h13) as h14. rewrite h14 in h10. contradiction. (*right*) auto with sets. (*right*) intro h11. red. intros S h1 h2. pose proof (Inhabited_not_empty _ _ h2) as h3. unfold dense_in. intro h4. unfold interior in h11. assert (h12: In [U : Ensemble (point_set X) | open U /\ Included U (P -)] S). constructor. split; assumption. inversion h2 as [x h13]. pose proof (family_union_intro _ _ _ x h12 h13) as h14. rewrite h11 in h14. contradiction. Qed. (*Wyckoff*) Lemma now_dense_iff2 : forall P : (Ensemble (point_set X)), now_dense P <-> (forall U:Ensemble (point_set X), open U -> Included U (P -) -> U = (Empty_set _)). intros P. split. (*left*) intros h3 U h5 h6. case (classic (U = (Empty_set _))); trivial. intro h7. pose proof (not_empty_Inhabited _ U h7) as h8. red in h3. pose proof (h3 U h5 h8) as h9. unfold dense_in in h9. contradiction. (*right*) intros h1. red. intros S h2 h3. red. intro h4. red in h4. pose proof (h1 S h2 h4) as h5. rewrite h5 in h3. inversion h3 as [x h7]. contradiction. Qed. (*Wyckoff*) (* "To say that no non-empty open set is included in P − [equivalent to now_dense from above lemma] is equivalent to saying that every non- empty open set has a non-empty intersection with P - '" *) Lemma now_dense_iff3 : forall P : (Ensemble (point_set X)), now_dense P <-> forall (S : (Ensemble (point_set X))), open S -> Inhabited S -> Inhabited (Intersection (P - ') S). intro P. pose proof (now_dense_iff2 P) as h1. rewrite h1. clear h1. split. (*left*) intros h1 S h2 h3. case (classic (Included S (P -))) as [h4 | h5]. (*h4*) pose proof (h1 S h2 h4) as h6. rewrite h6 in h3. inversion h3 as [x h7]. contradiction. (*h5*) apply complement_meets_non_subset; trivial. (*right*) intros h1 U h2 h3. case (classic (Inhabited U)) as [h4 | h5]. (*h4*) pose proof (h1 _ h2 h4) as h6. pose proof (complement_meets_non_subset U (P -) h4) as h7. rewrite <- h7 in h6. contradiction. (*h5*) pose proof (not_empty_Inhabited (point_set X) U) as h6. case (classic (U = Empty_set (point_set X))); trivial. intro h7. pose proof (h6 h7) as h8; contradiction. Qed. (*Wyckoff*) (*Thus, P is nowhere dense if and only if P − 'is dense.*) Lemma now_dense_iff4 : forall P: Ensemble (point_set X), now_dense P <-> dense (P - '). intro P. rewrite now_dense_iff3. pose proof (meets_every_nonempty_open_impl_dense _ ((P -) ')) as h1. pose proof (dense_impl_meets_every_nonempty_open ((P -) ')) as h2. tauto. Qed. End more_interior_closure_relations1. Arguments now_dense [_] _. Section nowhere_dense_examples. Variable T:Type. (*Wyckoff*) (*In a discrete space only the empty set is nowhere dense.*) Example now_dense_dscrt : forall S:(Ensemble (point_set (X_dt T))), now_dense S -> S = (Empty_set _). intros S h1. pose proof (now_dense_iff2 (X_dt T) S) as h2. rewrite h2 in h1. simpl. simpl in h1. unfold open_dscrt in h1. apply (h1 S I (closure_inflationary S)). Qed. (*Wyckoff*) (*"In an infinite space with the cofinite topology, a set is nowhere dense just in case it is finite."*) Example now_dense_cof : ~FiniteT T -> forall S:(Ensemble (point_set (X_ct T))), now_dense S <-> Finite S. intros h1 S. split. (* -> *) intro h2. rewrite now_dense_iff1 in h2. pose proof (cof_int_eq_empty _ (S -)) as h3. destruct h3 as [h4 h5]. case (classic (Finite ((S -) '))) as [h6 | h7]. (*h6*) pose proof (h4 h6) as h8. case (classic (Finite S)); trivial. intro h9. pose proof (cof_clo_eq_empty _ S) as h10. destruct h10 as [h11 h12]. pose proof (h12 h9) as h13. assert (h14 : Full_set (point_set (X_ct T)) = Empty_set _). congruence. pose proof (Empty_is_finite (point_set (X_ct T))) as h15. rewrite <- h14 in h15. pose proof (Finite_FiniteT _ h15) as h16. contradiction. (*h7*) pose proof (h5 h7) as h8. pose proof (closure_closed S) as h9. unfold closed in h9. simpl in h9. red in h9. case h9 as [h10 | h11]. (*h10*) pose proof (complement_empty T) as h12. rewrite <- h10 in h12. pose proof (Complement_Complement _ (S -)) as h13. simpl in h12. simpl in h13. rewrite h13 in h12. pose proof (interior_full (X_ct T)) as h14. simpl in h14. rewrite <- h12 in h14 at 1. rewrite h8 in h14. simpl in h14. pose proof (Empty_is_finite T) as h15. rewrite h14 in h15. pose proof (Finite_FiniteT _ h15) as h16. contradiction. (*h11*) pose proof (Complement_Complement _ (S -)) as h12. simpl in h11. simpl in h12. rewrite h12 in h11. simpl. apply Finite_downward_closed with (S -). simpl. assumption. apply closure_inflationary. (* <- *) intro h2. rewrite now_dense_iff1. pose proof (cof_clo_eq_empty _ S) as h3. destruct h3 as [h4 h4']. pose proof (h4 h2) as h5. pose proof (cof_int_eq_empty _ S) as h6. destruct h6 as [h7 h8]. case (classic (Finite (S '))) as [h9 | h10]. (*h9*) pose proof (h7 h9) as h10. pose proof (inf_comp_fin h1) as h11. pose proof (h11 S h9) as h12. simpl in h2. contradiction. (*h10*) pose proof (h8 h10) as h11. rewrite <- h5 in h11. assumption. Qed. End nowhere_dense_examples. Section more_interior_closure_relations2. (*Wyckoff*) Lemma closure_eq2 : forall {X:TopologicalSpace} (S:Ensemble (point_set X)), closure S = [x:(point_set X) | forall U:Ensemble (point_set X), open U -> In U x -> Inhabited (Intersection S U)]. intros X S. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h1. constructor. intros U h2 h3. apply closure_impl_meets_every_open_neighborhood with x; assumption. (* >= *) red. intros x h2. inversion h2 as [h3]. apply meets_every_open_neighborhood_impl_closure; assumption. Qed. (*Wyckoff*) (* This is Exercise 13 c from Chapter 9. "(P \/ Q)-' = P-' /\ Q-' "*) Lemma closure_comp_distribute : forall {X:TopologicalSpace} (P Q:Ensemble (point_set X)), (Union P Q) - ' = Intersection (P - ') (Q - '). intros X P Q. pose proof (closure_union P Q) as h1. rewrite h1. apply comp_union. Qed. (*Wyckoff*) (* This is Exercise 13 d from Chapter 9. "P /\ Q - <= (P /\ Q) - whenever P is open."*) Lemma closure_over_intersection_open : forall {X:TopologicalSpace} (P Q:Ensemble (point_set X)), open P -> Included (Intersection P (Q -)) ((Intersection P Q) -). intros X P Q h1. red. intros x h2. destruct h2 as [x h2l h2r]. rewrite (closure_eq2 Q) in h2r. inversion h2r as [h3]. rewrite (closure_eq2 (Intersection P Q)). constructor. intros U h5 h6. pose proof (open_intersection2 _ _ _ h1 h5) as h7. pose proof (Intersection_intro _ _ _ _ h2l h6) as h8. pose proof (h3 _ h7 h8) as h9. rewrite (Intersection_commutative _ _ Q). rewrite Intersection_associative. assumption. Qed. (*Wyckoff*) Lemma union_now_dense : forall {X:TopologicalSpace} (P Q:Ensemble (point_set X)), now_dense P -> now_dense Q -> now_dense (Union P Q). intros X P Q h1 h2. pose proof (now_dense_iff4 _ P) as h3. pose proof (now_dense_iff4 _ Q) as h4. rewrite h3 in h1. rewrite h4 in h2. red in h1. red in h2. assert (h5: (Q - ') = Intersection (Full_set _) (Q - ')). apply Extensionality_Ensembles. red. split. (* <= *) red. intros. constructor. constructor. assumption. (* >= *) red. intros ? h6. inversion h6; assumption. rewrite <- h1 in h5. pose proof (closure_over_intersection_open (Q - ') (P - ')) as h6. pose proof (Intersection_commutative _ (P - ' -) (Q - ')) as h7. assert (h8:open (Q - ')). apply closed_complement_open. pose proof (Complement_Complement _ (Q -)) as h9. rewrite h9. apply closure_closed. apply h6 in h8. rewrite h7 in h5. rewrite <- h5 in h8. pose proof (closure_comp_distribute Q P) as h9. pose proof (Union_commutative P Q) as h10. rewrite h10. apply now_dense_iff4. red. pose proof (f_equal (closure (X:=X)) h9) as h11. rewrite h11. pose proof (closure_closed (Intersection (Q - ') (P - '))) as h12. pose proof (closure_minimal _ _ h12 h8) as h13. rewrite h2 in h13. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x ?. constructor. (* >= *) assumption. Qed. (*Wyckoff*) (*"A finite union of nowhere dense sets is again nowhere dense"*) Lemma finite_union_now_dense : forall {X:TopologicalSpace} (F:Family (point_set X)), Finite F -> (forall A, In F A -> now_dense A ) -> now_dense (FamilyUnion F). intros X F h1 h2. induction h1 as [|F A' h3 S h4]. (*Empty*) pose proof (now_dense_iff1 X (Empty_set _)) as h3. pose proof (empty_family_union (point_set X)) as h4. rewrite h4. rewrite h3. pose proof (closure_empty X) as h5. rewrite h5. apply interior_empty. (*Add*) (*unfold Add. unfold Add in h2.*) pose proof (family_union_add F S) as h5. rewrite h5. pose proof (Add_intro2 _ F S) as h6. pose proof (h2 _ h6) as h7. assert (h8:forall A:Ensemble (point_set X), In F A -> now_dense A). intros A h9. apply NNPP. intro h10. assert (h11: ~ In (Add F S) A). pose proof (h2 A) as h11. tauto. pose proof (Add_intro1 _ F S A) as h12. apply h12 in h9. contradiction. apply h3 in h8. apply union_now_dense; assumption. Qed. End more_interior_closure_relations2. (*Wyckoff*) (*"A set is said to be meager if it is the countable union of nowhere dense sets."*) Definition meager {X:TopologicalSpace} (S:Ensemble (point_set X)) := exists (It:Type) (F:(IndexedFamily It (point_set X))), CountableT It /\ (forall (i:It), now_dense (F i)) /\ (S = IndexedUnion F). Section meager_examples. Variable T:Type. (*Wyckoff*) (*"In a discrete space only the empty set is meager."*) Example meager_empty_dscrt : forall (S:Ensemble (point_set (X_dt T))), meager S -> S = (Empty_set _). intros S h1. (*simpl in S. simpl.*) red in h1. elim h1. intros It h2. elim h2. intros F h3. clear h1 h2. destruct h3 as [h4 h5]. destruct h5 as [h6 h7]. assert (h8: forall i:It, (F i) = (Empty_set _)). intro i. pose proof (h6 i) as h9. apply (now_dense_dscrt _ (F i) h9). rewrite h7. case (classic (Inhabited S)) as [h10 | h11]. (*h10*) inversion h10 as [x h11]. rewrite h7 in h11. inversion h11 as [i y h12]. pose proof (h8 i) as h9. rewrite h9 in h12. contradiction. (*h11*) rewrite <- h7. case (classic (S = Empty_set _)); trivial. intro h12. pose proof (not_empty_Inhabited _ _ h12) as h13. contradiction. Qed. (*Wyckoff*) (*"In an infinite space with the cofinite topology, a set is meager just in case it is countable."*) Example meager_countable_cof_inf : (~FiniteT T) -> forall S:(Ensemble (point_set (X_ct T))), (meager S <-> Countable S). intros h1 S. split. (*->*) intro h2. unfold meager in h2. destruct h2 as [It h3]. destruct h3 as [F h4]. destruct h4 as [h5 h6]. destruct h6 as [h7 h8]. rewrite h8. simpl. assert (h9:forall (i:It), Countable (F i)). intro i. pose proof (now_dense_cof _ h1 (F i)) as h10. apply h10 in h7. apply Finite_impl_Countable; assumption. apply countable_union; assumption. (* <- *) intro h2. red. red in h2. simpl in h2. simpl. pose proof CountableT_is_FiniteT_or_countably_infinite as h3. destruct (h3 _ h2) as [h0|h4]. (*h0 *) pose (SingF S) as F'. simpl in F'. pose (fun (s:{x:T | In S x}) => Singleton (proj1_sig s)) as F. exists {x:T | In S x}, F. split. apply FiniteT_impl_CountableT; assumption. split. intro s. pose proof (now_dense_cof _ h1 (F s)) as h4. rewrite h4. simpl. unfold F. apply Singleton_is_finite. pose proof (indexed_to_family_union _ _ F) as h4. rewrite h4. pose proof (union_singF_eq S) as h5. unfold SingF in h5. simpl in h5. unfold ImageFamily. unfold F. rewrite h5 at 1. assert (h6: Im S (Singleton (U:=T)) = Im (Full_set _) F). apply Extensionality_Ensembles. red. split. (* <= *) red. intros A h6. inversion h6 as [? h7 ? h8]. simpl in x. pose proof (Full_intro _ (exist _ x h7)) as h9. assert (h10:A = F (exist (In S) x h7)). unfold F. rewrite h8. simpl. reflexivity. rewrite h8 in h6. apply (Im_intro _ _ _ _ _ h9 _ h10). (* >= *) red. intros A h6. inversion h6 as [? ? ? h8]. unfold F in h8. apply (Im_intro _ _ _ _ _ (proj2_sig x) _ h8). f_equal. assumption. (* h4 *) destruct h4 as [F h5]. pose proof (bijective_impl_invertible _ h5) as h6. inversion h6 as [G h7 h8]. assert (h9 : bijective G). apply invertible_impl_bijective. apply (intro_invertible G F h8 h7). pose (fun n:nat => Singleton (proj1_sig (G n))) as G'. exists _, G'. split. apply (intro_nat_injection _ (@id nat)). red. intros x1 x2 h10. unfold id in h10. assumption. split. intro n. pose proof (now_dense_cof _ h1 (G' n)) as h10. rewrite h10. apply Singleton_is_finite. pose proof (union_singF_eq S) as h10. unfold X_ct. simpl. simpl. pose proof (indexed_to_family_union _ _ G') as h11. rewrite h11. rewrite h10. unfold SingF. unfold ImageFamily. unfold X_ct. simpl. assert (h12: Im S (Singleton (U:=T)) = Im (Full_set nat) G'). unfold G'. apply Extensionality_Ensembles. red. split. (* <= *) red. intros A h12. inversion h12 as [t h13 ? h14 h15]. simpl in t. red in h9. destruct h9 as [? h9]. red in h9. pose proof (h9 (exist _ t h13)) as h16. destruct h16 as [n h17]. pose proof (Full_intro _ n) as h18. symmetry in h17. pose proof (Im_intro _ _ _ _ _ h18 _ h17) as h19. pose proof (f_equal (@proj1_sig _ _ ) h17) as h20. unfold G'. eapply Im_intro. apply h18. rewrite <- h17. simpl. assumption. (* >= *) red. intros A h12. inversion h12 as [n ? B h14]. eapply Im_intro. apply (proj2_sig (G n)). assumption. f_equal; assumption. Qed. (*Wyckoff*) (*"In particular, if the space is countably infinite, then every subset is meager."*) Corollary meager_cof_countable : ~FiniteT T-> CountableT T -> forall S:(Ensemble (point_set (X_ct T))), meager S. intros h1 h2 S. pose proof (meager_countable_cof_inf h1 S) as h3. rewrite h3. apply CountableT_Countable. simpl. assumption. Qed. End meager_examples. Section Boundary. Variable X:TopologicalSpace. (*Wyckoff*) (*"A point x is a boundary point of a set P if every neighborhood of x contains points of P and points of P '."*) Definition boundary_point (x:point_set X) (P:Ensemble (point_set X)) : Prop := forall (U:Ensemble (point_set X)), open U -> In U x -> Inhabited (Intersection P U) /\ Inhabited (Intersection (P ') U). (* "In other words, x is a boundary point of P just in case it belongs to the closure of both P and P '." *) (*Wyckoff*) Lemma boundary_point_iff : forall (x:point_set X) (P:Ensemble (point_set X)), boundary_point x P <-> (In (P -) x) /\ (In (P ' -) x). intros x P. rewrite closure_eq2. rewrite closure_eq2. split. (* -> *) intro h1. unfold boundary_point in h1. split. (* P *) constructor. apply h1. (* P ' *) constructor. apply h1. (* < - *) intro h1. destruct h1 as [h1l h1r]. red. destruct h1l as [h1l]. destruct h1r as [h1r]. intros. split. apply h1l; assumption. apply h1r; assumption. Qed. (*Wyckoff*) (* "The boundary of P is the set of its boundary points" *) Definition boundary (P:Ensemble (point_set X)) : Ensemble (point_set X) := [x:(point_set X) | boundary_point x P]. (*Wyckoff*) (* " . . . that is to say, it is the set P - /\ (P ' − ) " *) Lemma boundary_eq : forall (P:Ensemble (point_set X)), boundary P = Intersection (P -) (P ' -). intro P. unfold boundary. apply Extensionality_Ensembles. red. intros. split. (*left*) red. intros x h1. destruct h1 as [h1]. apply boundary_point_iff in h1. constructor; apply h1. (*right*) red. intros x h1. destruct h1 as [x h1 h2]. pose proof (conj h1 h2) as h3. constructor. rewrite boundary_point_iff; assumption. Qed. (*Wyckoff*) (*"The boundary of a set is always closed, because it is the intersection of two closed sets."*) Lemma boundary_closed : forall (P:Ensemble (point_set X)), closed (boundary P). intros. rewrite boundary_eq. apply closed_intersection2; apply closure_closed. Qed. (*Wyckoff*) (*"If a set P is open, then its complement P is closed, and therefore P ' - = P ' ; in this case, the boundary of P is just the set-theoretic difference P - (- P)."*) Lemma open_boundary_difference : forall P:(Ensemble (point_set X)), open P -> boundary P = Setminus (P -) P. intros P h1. pose proof (Complement_Complement _ P) as h2. rewrite <- h2 in h1. change (open ((P ') ')) with (closed (P ')) in h1. rewrite boundary_eq. pose proof (closure_closed (P ')). assert (h3: (P ') - = P '). apply Extensionality_Ensembles. split. (* <= *) simpl in h1. apply closure_minimal. assumption. auto with sets. (* >= *) apply closure_inflationary. rewrite h3. rewrite setminus_int_complement; reflexivity. Qed. (*Wyckoff*) (*"More generally, the boundary of any nowhere dense set P in a topological space is just the closure P - "*) Lemma boundary_now_dense : forall (P:Ensemble (point_set X)), now_dense P -> boundary P = P -. intros P h1. rewrite now_dense_iff4 in h1. pose proof (closure_inflationary P) as h2. pose proof (complement_inclusion _ _ h2) as h3. pose proof (closure_increasing _ _ h3) as h4. red in h1. rewrite h1 in h4. assert (h5:(Full_set _ = P ' -)). apply Extensionality_Ensembles. red. split; auto with sets. red. intros; constructor. rewrite boundary_eq. rewrite <- h5. apply Extensionality_Ensembles. red. split; red. (* <= *) intros x h6. destruct h6; assumption. (* >= *) intros x h6. constructor; [assumption|constructor]. Qed. (*Wyckoff*) (*"The closure of a set P ought to be, and is, the union of P with its boundary. "*) Lemma closure_union_boundary : forall P:Ensemble (point_set X), let Q := (boundary P) in P - = Union P Q. intros P Q. apply Extensionality_Ensembles. red. split. (* >= *) Focus 2. pose proof (boundary_eq P) as h1. change (boundary P) with Q in h1. pose proof (closure_inflationary P) as h2. assert (h3: Included Q (P -)). red. intros x h3. rewrite h1 in h3. destruct h3; assumption. apply included2_impl_union_included. constructor; assumption. (* <= *) red. intros x h1. case (classic (In P x)). auto with sets. intro h2. change (~ In P x) with (In (P ') x) in h2. assert (h3:In Q x). unfold Q. rewrite boundary_eq. pose proof (closure_inflationary (P ')). constructor; auto with sets. auto with sets. Qed. End Boundary. Arguments boundary [_] _ _. Section Boundary_Examples. Variable T:Type. (*Wyckoff*) (*" In a discrete space every set of points has an empty boundary."*) Example boundary_discrete : forall (S:Ensemble (point_set (X_dt T))), boundary S = Empty_set _. intro S. (*simpl in S. simpl.*) rewrite open_boundary_difference. assert (h1: closed S). apply all_closed_dscrt. pose proof (closure_minimal S _ h1) as h2. assert (h3:Included S S); auto with sets. apply h2 in h3. apply setminus_sub_sup; assumption. simpl. red. trivial. Qed. (*Wyckoff*) (* "In an infinite space with the cofinite topology, the boundary of a finite set is itself, the boundary of a cofinite set is its complement, and the boundary of an infinite set with an infinite complement is the whole space."*) Example boundary_inf_cof : ~FiniteT T -> forall S:(Ensemble (point_set (X_ct T))), ((Finite S) -> boundary S = S) /\ ((Finite (S ')) -> boundary S = S ') /\ ((~ Finite S) -> (~Finite (S ')) -> boundary S = (Full_set _)). intros h1 S. repeat split. (* Conclusion I *) intro h2. pose proof closed_cof_fin as h3. pose proof (or_introl (S = Full_set _) h2) as h4. rewrite <- h3 in h4. rewrite boundary_eq. pose proof (closure_fixes_closed _ h4) as h5. rewrite h5. pose proof (Complement_Complement _ S) as h6. rewrite <- h6 in h2. pose proof (inf_comp_fin h1 _ h2) as h7. pose proof (infnt_cof_infnt_dense T _ h1 h7) as h8. red in h8. rewrite h8. apply intersection_full_set. (* Conclusion II *) intro h2. rewrite boundary_eq. assert (h3: open S). simpl. red. right. assumption. pose proof (Complement_Complement _ S) as h4. rewrite <- h4 in h3. assert (h5:closed (S ')). assumption. pose proof (closure_fixes_closed _ h5) as h6. rewrite h6. pose proof (inf_comp_fin h1 _ h2) as h7. pose proof (infnt_cof_infnt_dense _ _ h1 h7) as h8. red in h8. rewrite h8. rewrite Intersection_commutative. apply intersection_full_set. (* Conclusion III *) intros h2 h3. rewrite boundary_eq. pose proof (infnt_cof_infnt_dense _ _ h1 h2) as h4. pose proof (infnt_cof_infnt_dense _ _ h1 h3) as h5. red in h4. red in h5. rewrite h4. rewrite h5. apply intersection_full_set. Qed. End Boundary_Examples. Section Regular. Variable X:TopologicalSpace. (*Wyckoff*) (*"An open set is said to be regular if it coincides with the interior of its own closure."*) Definition regular (S:(Ensemble (point_set X))) : Prop := S = S - o. (*Wyckoff*) Lemma regular_open : forall (S:Ensemble (point_set X)), regular S -> open S. intros S h1. red in h1. rewrite h1. apply interior_open. Qed. (*Wyckoff*) Lemma regular_iff : forall (P:Ensemble (point_set X)), regular P <-> P = P - ' - '. intro P. unfold regular. rewrite interior_eq. tauto. Qed. (*Wyckoff*) Example example_open_not_regular : (exists (P:Ensemble (point_set X)), Inhabited P /\ now_dense P) -> exists (S:Ensemble (point_set X)), open S /\ ~ regular S. intro h1. destruct h1 as [P h1]. destruct h1 as [h1l h1r]. assert (h2:open (P - ')). change (open (P - ')) with (closed (P -)). apply closure_closed. assert (h3: (P - ') <> Full_set _). intro h4. pose proof (complement_full (point_set X)) as h5. rewrite <- h4 in h5. pose proof (Complement_Complement _ (P -)) as h6. rewrite h6 in h5. pose proof (closure_inflationary P) as h7. rewrite h5 in h7. assert (h8:P = Empty_set _). apply Extensionality_Ensembles. red. split; [assumption | auto with sets]. inversion h1l as [x h9]. rewrite h8 in h9. contradiction. exists (P - '). split; try assumption. assert (h4:(P - ' - ' - ') = (Full_set (point_set X)) -> ~ regular (P - ')). intro h4. rewrite <- h4 in h3. unfold regular. rewrite interior_eq. assumption. assert (h5:P - ' - ' - ' = Full_set (point_set X)). apply now_dense_iff4 in h1r. red in h1r. rewrite h1r. pose proof (complement_full (point_set X)) as h5. rewrite h5. pose proof (closure_empty X) as h6. rewrite h6. apply complement_empty. tauto. Qed. End Regular. bool2-v0-3/EqDec.vo0000644000175000017500000003036313575600573014671 0ustar dwyckoff76dwyckoff76 Є•¦¾Á[Р%EqDec 5BooleanAlgebrasIntro2@Àð“  +type_eq_decÐ@‘·!T”‘ A¶!xA¶!yB©›  ° &Specif $Init #Coq@@'sumbool @ ©›  ° %Logic   @@"eq@°C ©š ° @#not¶1”‘ HA”‘ LB° ¨+g+ S'+S'7"7 7"7%'  ‘‘%d  ’p  ‘‘,ˆ  ‘‘>”@€ÀÀ@°LAÀ@°A  2 $Init #Coq@N@AB°UB%@C  .type_eq_dec_eqÐ@’@¶!T”‘ zE¶"pf©š °{@vg¶#pf'©ši©›  °(@V@°©Wrv@ÀÀ@°@@1À@°MA3@AB° ARÀÀ@°#A7@A°$A  8FunctionalExtensionality %Logic #Coq@OÀ@°-A  PÀ@° ªEB1@ABCD  5impl_type_eq_dec_prodÐ@’A¶!T”‘ ¸\¶!U”‘ ¿]¶@©šC¤¶@©šG¨©šJ©›  ° )Datatypes $Init #Coq@@$prod @ DŸ@ÀÀ@°'A—À@°)A{À@°+A }Q@ABC°-A gÀÀ@°1A ƒu@A°,A£ÀÀ@°/AˆÀ@°1A @AB°2A !hÀ@°5A @ACDE  0match_eq_dec_symÐ@’B¶!T”‘ à¶!U”‘ á¶"pf©šŽï¶!x׶!y<¶"bl?¶"brE©›“°FÍÀÿB AA @D·!x©› ©›û°GUó©šò©›°[ù© ]û ·!e© °c·!n©©°j#ÍÀ"B# @D·"©›& ©›° t©š©›!°&z&©9 | ·!e©*°. ‚ ·!n©"©1°5'‰B@ÀÀ@°fA@A°gAÀÀ@°jA o@A°eA@BC  /type_dep_eq_decÐ@‘·c”‘ yü·!P¶@b”A¶!ac¶!bK¶@©Mi¶@©²l©k ©_°q·U©T¶!T”‘ ™ü¶ ¶@”A”‘ ¢ý°  +g() + S' +S' + S'+ S'7"7 7"7%'  ‘‘{Ü  ’qè  ‘‘‚  ‘‘” @€ÀÀÀ@°"AU@A°CA@B°DB$@C  6type_dep_eq_dec_compatÐ@’C¶!T”‘ Ê¶!P¶@³2¶@©šR³©š °Ð@h ¸@ÀÀ@°A•@A°AE@B  /dep_type_eq_decÐ@‘·!S”‘ ì·Û¶@Ó”‘ ò¶!sÖ©š °ú@ï©Ýⶔ‘ ¶!T¶@픑  ”‘ ° `+g()+ S' 7$'  ’!L@€ÀÀ@° BA À@°3AÀ@°AÙ@ABC°6BÀ@°2B@AD  *sig_eq_decÐ@’D¶!T”‘ 1 ¶"pf©š·¶!P¶@Ÿ©š¾©›  °@#sig@  ·!x©+0@ÀÀ@°@@ëÀ@° S@î@AB°$A ÀÀÀ@°(A 3A@A°*Aõ@B°+AyÀÀ@°.A 9WÀ@°1A÷@AB°2AzÀ@°4A@ACDE  &eq_decÐ@@¶Z”‘ pK©{W@À@° wKA0@A@o@ÀÀÀÀÀ@-@A+@BÀ@À@çÀ@e@ABCDçÀÀçåÀ@ÝÀ@¦@ABC¤£DE ÀÀŸœÀ›—À@•À@#@ABCD"ÀÀÀ# À@Ó@ABÒÀÀ@³@A²ÀÀÀ@t@ArÀ@?@AB;À@:@ACDE9À83ÀÀÀ@4À@#@AB-@CðÀÀ@ß@AzyBDEFGH @@@   0ProofIrrelevance %Logic #Coq@0Y\ø‚,&L«ç—ÜT^‚´¡   5ProofIrrelevanceFacts %Logic #Coq@0êî%¬JXæÃQeâë   *EqdepFacts %Logic #Coq@0FÀO!ZTßg$Jã®Íƒ  0 Chž^y»Õžf(:   *Logic_Type $Init #Coq@0¸°ûÅ4€K95X 7Rã   'Prelude $Init #Coq@0/ø^,² sâPzò±Úºö   'Tactics $Init #Coq@0æ”ÃäûE%‚QHìÿ   "Wf $Init #Coq@0OY+SË+â <Â]ôÏO   %Peano $Init #Coq@0"P¾auã1ñmÛ‰5,ª  Â0þ.'úŠ¦ÛK»¾MK  0ȽèÂDîS0Âû`ža˜•{  ˆ0Ä\_¨u‡*Å´;ÿ` –¾   )Notations $Init #Coq@0&H¥çd³&Ów,j'@@°ó  ð (CONSTANT°‘°@e@@‘@  #_18 )IMPLICITS   ŠŠà@A@@@@  ‘  @@@@  #_19 $HEAD ‘@  #_20 /ARGUMENTS-SCOPEÀ@‘ *type_scope@ @@  ¸ +°*@’A  #_22 '   °+@Ã'  ‘&@  #_23 $ ‘    #_24 $À@‘ # @ @@ @ ‘Ì ‘Ï@  #_25 L ²‘&KC    @ °è@ AA @ @@@@  ¶ g°f@’A  #_27 c   °g@Ác  ‘b@  #_28 ` ‘    #_29 `À@‘ _ a @ @@ @ @ ‘  ‘@  #_30 Œ ²‘*‹F    @ °ê@ AA °ç@ AA @ @ @ @@@@  ³ ­°¬@’A  #_32 ©   °­@¾©  ‘¨@  #_33 ¦ ‘    #_34 ¦À@‘ ¥ § @ @ @ @ @@ @ @ ‘U @ @ @ @@  #_35 Ö ²‘.ÕG    @ °ë@ AA °è@ AA @ @ @ @ @@@@  … ø°÷@‘@  #_37 ô   ((ó  ‘ò@  #_38 ð ‘ ï  #_39 îÀ@‘ í @@ @ A@  #_40  ²‘B    @ °Ž@ AA @@@@  h +°*@’A  #_42 '   °+@s'  ‘&@  #_43 $ ‘    #_44 $À@‘ # @ @@ @ A ‘Î@  #_45 K ²‘%JG    @ °—@ AA @ @ @ @ @ @@@@  † j°i@‘@  #_47 f   °j@‘f  ‘e@  #_48 c ‘ b  #_49 aÀ@‘ ` @@ @ A@  #_50 „ ²‘ƒB    @ °®@ AA @@@@  t ž°@’A  #_52 š   °ž@š  ‘™@  #_53 — ‘    #_54 —À@‘ – @ @@ @ ‘? A@  #_55 ¾ ²‘%½E    @ °£@ AA @ @ @ @@@@  p Û°Ú@A  #_57 ×   °Û@{×  ‘Ö@  #_58 Ô ‘    #_59 ÔÀ@‘ Ó@ @@  #_60 ô ²‘óC    @ °ò@ AA @ @@@@  #_62 (LIBTYPES Ð@¡@@¡ Ð@§ @@@§  ÐÐ@@¡@À@‘ nn@AÐ@A¡AÀÀ@‘ FF@A‘¨@B@AB¡@@¡ Ð@§ @@@§  Ð@@¡@ÀÀ@‘ ÉÉÀ@‘ ˜˜@AB‘h@C@A@@ Ð@B¡BÀÀ@ À@ @AB@C@A@@A@ Ð@A¡AÀÀ@À@@AB @C@A@Ð@£@@£ Ð@‘Ñ¡‘ÔÀÀ@‘ ëëÀ@‘ **@AB‘ ,,@C@A@ Ð@@¡@ÀÀ@ À@ @AB@C@A@@AC@@ ÐÐ@B¡BÀÀ@D@AB@B@A§ @@@§  ÐÐ@@¡@À@T@AÐ@A¡AÀÀ@@@A;@B@AB¡@@¡ Ð@§ @@@§  Ð@@¡@À@,@A@A@@ Ð@B¡BÀ@0@A@A@@A@ Ð@A¡AÀ@4@A@A@Ð@£@@£ Ð@‘¡‘ÀÀ@DÀ@a@AB@@C@A@ Ð@@¡@ÀÀ@JÀ@g@ABF@C@A@@AC@@ Ð@B¡BÀÀ@o@Aj@BÐ@§ @@@§  Ð@@¡@À@‘@AÐ@A¡AÀÀ@^À@{@AB\À@[@AC@AB@@ ÐÐ@B¡BÀÀÀ@ƒ@Ad@Bb@C@A§ @@@§  ÐÐ@@¡@À@¨@A@AA¡AÀ@s@A@B@@ Ð@B¡BÀ@w@AÐ@§ @@@§  Ð@£@@£ Ð@‘\¡À@¿@A@A@ Ð@@¡@À@Ã@A@A@@A@@ Ð@§ @@@§  Ð@A¡AÀ@Î@A@A@@ Ð@§ @@@§  Ð@A¡AÀ@Ù@A@A@@ Ð@B¡BÀ@Ý@A@A@@A@@A@@AB@@B@@AB@@B@@A@ Ð@A¡AÀÀÀ@Ë@AÆ@BÛ@CÐ@£@@£ ÐÐ@‘¡‘ŽÀÀ@ë@AºÀ@×@AB@A£@@£ ÐÐ@‘D¡À@Â@A@A£@@£ Ð@’—¡’šÀÀ@@AÌ@B@A@ ÐÐ@@¡@À@ @A@A£@@£ Ð@‘³¡À@Ú@A@A@ Ð@@¡@À@Þ@A@A@@B@@B@ ÐÐ@@¡@ÀÀ@æ@Aä@B@AÄ@@ @@Ä" Ð@¢@@¢ Ð@£@@£ Ð@£@@£ Ð@’3¡À@9@A@A@ Ð@£@@£ Ð@£@@£ Ð@£@@£ Ð@’:¡À@N@A@A@ Ð@@¡@À@R@A@A@@A@ Ð@@¡@À@V@A@A@@A@ Ð@@¡@À@Z@A@A@@A@@A@ Ð@£@@£ Ð@‘D¡À@e@A@A@ Ð@£@@£ Ð@£@@£ Ð@£@@£ Ð@’f¡À@z@A@A@ Ð@@¡@À@~@A@A@@A@ Ð@@¡@À@‚@A@A@@A@ Ð@@¡@À@†@A@A@@A@@A@@A@ Ð@@¡@À@Š@A@A@@A@ Ð@£@@£ Ð@£@@£ Ð@@¡@À@˜@A@A@ Ð@@¡@À@œ@A@A@@A@ Ð@@¡@À@ @A@A@@A@  Ð@¢@@¢ Ð@£@@£ Ð@£@@£ Ð@£@@£ Ð@’§¡À@»@A@A@ Ð@@¡@À@¿@A@A@@A@ Ð@@¡@À@Ã@A@A@@A@ Ð@@¡@À@Ç@A@A@@A@ Ð@@¡@À@Ë@A@A@@A@ Ð@¢@@¢ Ð@£@@£ Ð@‘º¡À@Û@A@A@ Ð@£@@£ Ð@£@@£ Ð@£@@£ Ð@’Ü¡À@ð@A@A@ Ð@@¡@À@ô@A@A@@A@ Ð@@¡@À@ø@A@A@@A@ Ð@@¡@À@ü@A@A@@A@@A@ Ð@@¡@À@@A@A@@A@@B@@B@ ÐÐÐ@@¡@ÀÀÀ@@AÎ@BÌ@C@A£@@£ Ð@£@@£ Ð@’f¡À@ß@AÐ@’÷¡À@@A@AB@ Ð@@¡@À@èÀ@@AB@A@@A@ Ð@@¡@À@í@AÐ@¢@@¢ Ð@@¡@À@@A@A@ Ð@£@@£ Ð@@¡@À@@A@A@ Ð@@¡@À@@A@A@@A@@AB@@BÄ@@ @@Ä Ð@¢@@¢ Ð@£@@£ Ð@£@@£ Ð@’L¡À@R@A@A@ Ð@£@@£ Ð@£@@£ Ð@£@@£ Ð@’S¡À@g@A@A@ Ð@@¡@À@k@A@A@@A@ Ð@@¡@À@o@A@A@@A@ Ð@@¡@À@s@A@A@@A@@A@ Ð@£@@£ Ð@‘]¡À@~@A@A@ Ð@£@@£ Ð@£@@£ Ð@£@@£ Ð@’¡À@“@A@A@ Ð@@¡@À@—@A@A@@A@ Ð@@¡@À@›@A@A@@A@ Ð@@¡@À@Ÿ@A@A@@A@@A@@A@ Ð@@¡@À@£@A@A@@A@ Ð@£@@£ Ð@£@@£ Ð@@¡@À@±@A@A@ Ð@@¡@À@µ@A@A@@A@ Ð@@¡@À@¹@A@A@@A@  Ð@¢@@¢ Ð@£@@£ Ð@£@@£ Ð@£@@£ Ð@’À¡À@Ô@A@A@ Ð@@¡@À@Ø@A@A@@A@ Ð@@¡@À@Ü@A@A@@A@ Ð@@¡@À@à@A@A@@A@ Ð@@¡@À@ä@A@A@@A@ Ð@¢@@¢ Ð@£@@£ Ð@‘Ó¡À@ô@A@A@ Ð@£@@£ Ð@£@@£ Ð@£@@£ Ð@’õ¡À@ @A@A@ Ð@@¡@À@ @A@A@@A@ Ð@@¡@À@@A@A@@A@ Ð@@¡@À@@A@A@@A@@A@ Ð@@¡@À@@A@A@@A@@C@@AB@@A@@@  HB  ÕM  mR  W  b\  mg  xr  ƒ}  Žˆ  ­“  ž˜  ©£  ´®@ a ƒ . v € Š ” ó Ÿ@„•¦¾0ä–Š@,7ô·Ï€®ÊFÉ„•¦¾áòK*з!T”‘  %EqDec 5BooleanAlgebrasIntro2@E·"h1©š ° @+type_eq_decA·"h2© B©š ° 8FunctionalExtensionality %Logic #Coq@@=functional_extensionality_dep ÐC·!x¶!yD©›  ° &Specif $Init @@'sumbool@ ©›  ° !   !@@"eq@°E/5©š ° @#not5;·&(©5Ð#·&$#©,A©<C·+)È!s©- CI-ÍÀ-B AA @D· ©2 ©'°F?M©©,°©; ©0°GAH©'^©+ EL` ·"h3È Í·"©K ©@°H8Q©7©E°©T ©I°I#A©@©œ \A° g{{ ·"h4©š ° *EqdepFacts x w@@5internal_eq_rew_r_depð8V·u·*©c°J=¶©g°K8©j°©y ©n°L†/©e©%°˜©'°  ·$©x°/oo©·!H©~°©€°YY©œ ƒA !]¯©· ©ˆ°”A©€©Ž°'^^©·©’° ©”°&TT©·©˜°©š°MRR·"H0© °©–©¤°NBB·"H1©ª°©¬°OFF©- Hˆ©š °°@(eq_transà©Ä ©¹°PLL©°©p°©> R©t° ’©š °Ã@'f_equalà¶@©Ø ©Í°Q44©Ä·!f ©ÿ©U  =©‹ "©Y ©à&(­  ©_ YO  ©b \Z©š ° 0ProofIrrelevance   @@1proof_irrelevance°rp·Ÿ°©ðªÌê· «·¾”¶¤©ì“©ú°‰©œ  B°’)·­©õ‹ÍÀ  °@%False@@€ @D·8›© °© ‚„©°„†© «â©°ˆŠ9©@–€-A·àïÈþåìÍú·Þ©°ש"°×Ô5II ·Îâ©ÌðÙû·8Ú·Òöî/©'°½-¶·ñ)Í'·Y!©,° €WC·Ûì©·µ©1°çH\©·¹©5°­©+©9°Ö ©·Á©=°µ©?°Øÿ©·Ç©C°»©E°×ü·ª©H°À©>©L°²é·§©O°©E©S°¯ìð,©¤à©d ©Y°­ëò©O©a°©c°/©à¶@ ©q ©f°­Ìø©\·˜©•©q ©í ©­à(Aœ ©ó í8  ©ö ðC£©”°1’¦·½”‘ »\·!U”‘ Â]·¾±·´²·¡©›  ° )Datatypes › °@@$prod@ ¨¯·«©  ’«ÍÀBB @D·!p© }›©³ ©¨°© zƒÕÏ© Ñ·"x1‰·"x2ŒÍ·©) w‡©Æ ©»°©/ t}©œ 2AÀxØæì©· Ú·"y1‡·"y2Š©·©Û ©аmàõ©Æ©·ýͶ·¼©å ©Ú°l³ê©Ð©ë ©à°©T Ht©%ÀJv´½©'ÀLxÝö©Ü   ·!eÍηÔ©ý ©ò°©f PZ©7ÀR\¶Æ©9ÀT^Ñï©î  ·"e0&©š °ÿ@(eq_rec_ràjû·Tk© © °©} _g©NÀaiÄ8©PÀckßè© ©àoì·fp©# ©°©[ÀnvEó ©©а© °.**©©¢ 2.äIôO·NK©'àƒ·z„©7 ©,°#©oÀ‚ŠYî!©$©6°©4°B©wÀ’œ@©,·©·Á¶@©>°5115©j©À ©¶ ‹˜©‡ÀšÍÀIB@ AD·^C·@©O° ©’À˜¥|¶@©T°©È ‰{©™À‹ŸOƒ·à©\°S©ŸÀ²º.Q©·æ©b° ©·È©f°­&©·Ì©j°!©š °k@)False_ind g©Jž©²à"½·#ÍÜ·Ù©í R¯§·!t·!u®©ÆÀ¸Ì<"/°±F«·»­Íi›­ · ²©Œàòƒˆ©’°©°ž©ÓÀîøR‰œ©ˆ·^©\ÍP·«·@©œ°Z©ßÀåò;nÉMÉ·&©¢°™©åÀø[}—©·,©¨°f d©·B©·©®°3@K©D §©ÝÞ©òàbé·ÍcÍ·@Så·>A·<Bâ©ÀòXiêëë·êç©Ç°ÓË·Ô©‹Í·Ú¿·@©Ë°‰©À!j„ø|ø·U©Ñ°È©À'/Š“Æ©·[©×°• “©·=u©·?/©·A©ß°&Ex·:©v Ù©£ ©$àom·ÿpÍN·K©_ 4rT·r7·p8©4Àx&…Œ©6Àz(’¬ ©4à¤?‚©;À-A“—¤%&©½ #·A”‘ ?à·„”‘ Dá·@3·"$· ·"bl!·"br ÈûâéÍ÷·ýÛ©°ÙÍû·0ÖÒF ·3Ö·!nÖ0Í·7©1 ©&°Ýÿ©ß©ù ! ·B#·=X ·XñÈÍ·©A ©6°Ó©,©:°×Í·Q©K ©@°Ù©6Ûm ·Z·'p ·Ç·©š °I@(eq_ind_rðç·d趩F©T°æ©V°è&· ÍÍV·@P©…©Û ï€tˆ·‡ÈE7=ÍA·G/©c°Z)Ž ·.1©ð'·z¶’©_©m°ÿ$˜©o°?/·™©f©t° 44Ís·@m ©¢©ø  1€J¥‘·DE©û t·¿”‘ ½·!P¶@³·½°·!a£·!bŸ·©¡··÷©‹ºb·×”‘ Õ ·"pfÑ·¶@Æ··©›  °±@#sig@ À·¿Á”·½©  ½·Ã¾šÍÀ BB @D·™© ¬·Ë­©Çã©Ç ©¼°© —·Ô˜©·ììæ©· è·Ú¼·"hxÍ·°©( œ·â]©Ü ©Ñ°©/ “·é”$©œ 3AÀ˜ïý©Î ñ·î¬·"hyo©·D©ð ©å°œõ ©ÛÍÉ·Ï©ø ©í°Šä©ã©þ ©ó°©Q Ž· ©”#©$À“Óñ©&À• ©ò ) ·c©ð›·œ¶>© ©°©f p· q©ž8©9Àu4:©;Àwê© ·N©Ì&©Ê°©°©x ®·2¯©·J©KÀ³L©MÀµ3© ©·­©)°©Ç"V;©·³©/°•©··©3°¶I¬·J‘©œb©·¿©;°”‘ x““©·Ç©C°ŠŠ·¨©F°¶\z¿·]{©u·¥©M°Òææ·"H2©R°©†æí©¥à©´ k·nl©¡†©‡Àpé‰ÀrÅñ©à¶@© Ê© U·}©|•·š ©—°€Ó©í ©­à©÷Ù*žŸ ©õ <Ë¢ ©ø GF¥ ©û ãV¨©™°W«Y«w·ä”©‡°’‡·¸“©·º©ˆ°©š °–@)proj1_sig° r©ºÀ"t`i©°$vq©J úm·ªû¶ù©Œ©š°öÅj¶Õ©Â”·©”s·Û©4sÍ¢·@œœ©Ë©' ýi€iÔÀy©èà’>©$ @’ Ù³bool2-v0-3/FiniteBags.v0000644000175000017500000043236613575574055015562 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-2019 Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import SetUtilities. Require Import FunctionalExtensionality. Require Import Description. Require Import DecidableDec. Require Import ProofIrrelevance. Require Import FunctionProperties. Require Import TypeUtilities. Require Import LogicUtilities. Require Import FiniteMaps. Require Import ArithUtilities. Require Import ListUtilities. Require Import EqDec. (*Version 0.3 bool2 Alpha -- certified -- Coq 8.4 pl4.*) Definition nonzero {T:Type} (S:Ensemble (T*nat)) := (forall x:T, ~ Inn S (x, 0)). Lemma nonzero_incl : forall {T:Type} (S S':Ensemble (T*nat)), Included S S' -> nonzero S' -> nonzero S. intros T S S' h1 h2. red. red in h2. intro x. specialize (h2 x). intro h3. red in h1. specialize (h1 _ h3). contradiction. Qed. Inductive Bag (T:Type) := bag_intro : type_eq_dec T -> forall S:Ensemble (T*nat), Finite S -> fun_well_defined S -> nonzero S -> Bag T. Lemma bag_intro_functional : forall {T:Type} (hdt hdt':type_eq_dec T) (S S':Ensemble (T*nat)) (hf:Finite S) (hf':Finite S') (hsf:fun_well_defined S) (hsf':fun_well_defined S') (hnz:nonzero S) (hnz':nonzero S'), S = S' -> bag_intro _ hdt S hf hsf hnz = bag_intro _ hdt' S' hf' hsf' hnz'. intros T hdt hdt' S S' h1 h2 h3 h4 h5 h6 h7. subst. pose proof (proof_irrelevance _ h1 h2). pose proof (proof_irrelevance _ h4 h3). pose proof (proof_irrelevance _ h5 h6). subst; f_equal. apply type_eq_dec_eq. Qed. Lemma bag_type_eq_dec : forall {T:Type} (B:Bag T), type_eq_dec T. intros ? B. destruct B; auto. Qed. Lemma bag_to_pair_set {T:Type} (B:Bag T) : Ensemble (T*nat). destruct B. refine S. Defined. Lemma bag_to_pair_set_inj : forall (T:Type), Injective (@bag_to_pair_set T). intro T. red. intros B B' h1. destruct B as [hdt S h2 h3 h4]; destruct B' as [hdt' S' h5 h6 h7]. simpl in h1. subst. pose proof (proof_irrelevance _ h2 h5). pose proof (proof_irrelevance _ h3 h6). pose proof (proof_irrelevance _ h4 h7). subst. f_equal. apply type_eq_dec_eq. Qed. Lemma finite_bag_to_pair_set : forall {T:Type} (B:Bag T), Finite (bag_to_pair_set B). intros T B. unfold bag_to_pair_set. destruct B. assumption. Qed. Lemma fun_well_defined_bag_to_pair_set : forall {T:Type} (B:Bag T), fun_well_defined (bag_to_pair_set B). intros T B. unfold bag_to_pair_set. destruct B. assumption. Qed. Lemma nonzero_bag_to_pair_set : forall {T:Type} (B:Bag T), nonzero (bag_to_pair_set B). intros T B. unfold bag_to_pair_set. destruct B. assumption. Qed. Definition bag_to_set {T:Type} (B:Bag T) := Im (bag_to_pair_set B) (@fst _ _). Lemma bag_to_set_dom_rel_compat : forall {T:Type} (B:Bag T), bag_to_set B = dom_rel (bag_to_pair_set B). unfold bag_to_set. intros. rewrite <- dom_rel_eq. reflexivity. Qed. Lemma finite_bag_to_set : forall {T:Type} (B:Bag T), Finite (bag_to_set B). intros T B. apply finite_image. apply finite_bag_to_pair_set. Qed. Lemma finite_dom_rel_bag_to_pair_set : forall {T:Type} (B:Bag T), Finite (dom_rel (bag_to_pair_set B)). intros T B. rewrite <- bag_to_set_dom_rel_compat. apply finite_bag_to_set. Qed. Lemma bag_to_fin_map_ex : forall {T:Type} (B:Bag T), let S:= bag_to_pair_set B in exists! F:Fin_map (dom_rel S) (ran_rel S) 0, S = fin_map_to_fps F. intros T B S. destruct B as [hdt S' hf hfp hnz]. pose proof (dom_rel_finite _ hf) as h3. pose proof (ran_rel_finite _ hf) as h4. exists (fin_map_intro _ _ _ hdt h3 h4 _ hfp). red. split. apply (fin_map_to_fps_s_compat hdt _ _ 0 h3 h4 _ hfp). intros B h5. destruct B as [hdt' h6 h7 S'' h8]. pose proof (proof_irrelevance _ h3 h6). pose proof (proof_irrelevance _ h4 h7). subst. pose proof (fin_map_to_fps_s_compat hdt' _ _ 0 h6 h7 _ h8) as h11. rewrite <- h11 in h5. clear h11. subst. subst. rewrite (proof_irrelevance _ hfp h8). unfold bag_to_pair_set; f_equal. apply type_eq_dec_eq. Qed. Definition bag_to_fin_map {T:Type} (B:Bag T) := proj1_sig (constructive_definite_description _ (bag_to_fin_map_ex B)). Lemma bag_to_fin_map_compat : forall {T:Type} (B:Bag T), (bag_to_pair_set B) = fin_map_to_fps (bag_to_fin_map B). intros T B. unfold bag_to_fin_map. destruct constructive_definite_description. simpl. assumption. Qed. Lemma bag_to_fin_map_inj : forall (T:Type) (B B':Bag T), fin_map_app (bag_to_fin_map B) = fin_map_app (bag_to_fin_map B') -> B = B'. intros T B B' h1. pose proof (bag_to_fin_map_compat B) as h2. pose proof (bag_to_fin_map_compat B') as h3. destruct B as [hdt S h4 h5 h6]; destruct B' as [hdt' S' h7 h8 h9]. unfold bag_to_pair_set in h2. unfold bag_to_pair_set in h3. assert (h10:S = S'). apply Extensionality_Ensembles. red. split. red. intros pr h11. destruct (zerop (snd pr)) as [h12 | h13]. red in h6. pose proof (h6 (fst pr)) as h13. rewrite <- h12 in h13. rewrite <-surjective_pairing in h13. contradiction. rewrite h2 in h11. assert (h12: snd pr = fin_map_app (bag_to_fin_map (bag_intro T hdt S h4 h5 h6)) (fst pr)). symmetry. apply fin_map_to_fps_fin_map_app_compat. assumption. rewrite surjective_pairing. rewrite h12. rewrite h1. rewrite h3 at 1. apply in_fin_map_to_fps. unfold bag_to_pair_set. apply NNPP. intro h14. pose proof (fin_map_app_def ((bag_to_fin_map (bag_intro T hdt' S' h7 h8 h9))) _ h14) as h15. rewrite <- h1 in h15. rewrite <- h12 in h15. omega. red. intros pr h11. destruct (zerop (snd pr)) as [h12 | h13]. red in h9. pose proof (h9 (fst pr)) as h13. rewrite <- h12 in h13. rewrite <-surjective_pairing in h13. contradiction. rewrite h3 in h11. assert (h12: snd pr = fin_map_app (bag_to_fin_map (bag_intro T hdt' S' h7 h8 h9)) (fst pr)). symmetry. apply fin_map_to_fps_fin_map_app_compat. assumption. rewrite surjective_pairing. rewrite h12. rewrite <- h1. rewrite h2 at 1. apply in_fin_map_to_fps. unfold bag_to_pair_set. apply NNPP. intro h14. pose proof (fin_map_app_def ((bag_to_fin_map (bag_intro T hdt S h4 h5 h6))) _ h14) as h15. rewrite h1 in h15. rewrite <- h12 in h15. omega. clear h2 h3. subst. pose proof (proof_irrelevance _ h4 h7). pose proof (proof_irrelevance _ h5 h8). pose proof (proof_irrelevance _ h6 h9). subst. f_equal. apply type_eq_dec_eq. Qed. Notation "B {->} x" := ((bag_to_fin_map B) |-> x) (at level 20). Lemma bag_to_fin_map_compat' : forall {T:Type} (B:Bag T) (x:T), B{->}x = fin_fps_to_f (bag_type_eq_dec B) (finite_dom_rel_bag_to_pair_set B) (bag_to_pair_set B) (fun_well_defined_bag_to_pair_set B) 0 x. intros T B x. pose proof (fin_map_app_compat (bag_to_fin_map B) x) as h1. rewrite h1. apply fin_fps_to_f_functional. rewrite bag_to_fin_map_compat; auto. Qed. Lemma bag_app_functional : forall {T:Type} (B B':Bag T), B = B' -> forall x:T, B{->}x = B'{->}x. intros T B B' h1. subst. auto. Qed. Definition Inb1 {T:Type} (B:Bag T) (x:T) : Prop := B {->} x > 0. Lemma inb1_dom_rel_compat : forall {T:Type} (B:Bag T) (x:T), Inb1 B x -> Inn (dom_rel (bag_to_pair_set B)) x. intros T B x h1. red in h1. destruct B as [S h2' h3' h4']. unfold bag_to_fin_map in h1. destruct constructive_definite_description as [F h5]. simpl in h1. apply NNPP. intro h6. pose proof (fin_map_app_def F x h6) as h7. destruct F. simpl in h1. simpl in h7. omega. Qed. Lemma not_inb1 : forall {T:Type} (B:Bag T) (x:T), ~Inb1 B x -> B {->} x = 0. intros T B x h1. destruct (zerop (B {->} x)) as [h2 | h3]. assumption. contradict h1. red. omega. Qed. Definition Inb2 {T:Type} (B:Bag T) (pr:(T*nat)) : Prop := B {->} (fst pr) = snd pr /\ snd pr > 0. Lemma inb2_impl_inb1 : forall {T:Type} (B:Bag T) (pr:(T*nat)), Inb2 B pr -> Inb1 B (fst pr). intros T B pr h1. red. red in h1. destruct h1 as [h1l h1r]. rewrite h1l. assumption. Qed. Lemma inb1_inb2_compat : forall {T:Type} (B:Bag T) (x:T), Inb1 B x -> Inb2 B (x, B{->}x). intros T B x h1. red. simpl. split; auto. Qed. Lemma inb1_dec : forall {T:Type} (B:Bag T) (x:T), {Inb1 B x} + {~Inb1 B x}. intros T B x. destruct (zerop (B{->}x)) as [h1 | h2]. right. red. intro h2. red in h2. omega. left. red. auto with arith. Qed. Lemma inb2_inj : forall {T:Type} (B:Bag T) (x:T) (m n:nat), Inb2 B (x, m) -> Inb2 B (x, n) -> m = n. intros T B x m n h1 h2. red in h1. red in h2. simpl in h1. simpl in h2. destruct h1; destruct h2. subst. reflexivity. Qed. Definition Inb2_le {T:Type} (B:Bag T) (pr:(T*nat)) : Prop := snd pr <= B {->} (fst pr) /\ 0 < snd pr. Lemma inb2_impl_inb2_le : forall {T:Type} (B:Bag T) (pr:(T*nat)), Inb2 B pr -> Inb2_le B pr. intros T B pr h1. red in h1. destruct h1 as [h1l h1r]. red. split; auto with arith. rewrite <- h1l. auto with arith. Qed. Lemma inb1_impl_inb2_le_1 : forall {T:Type} (B:Bag T) (x:T), Inb1 B x -> Inb2_le B (x, 1). intros T B x h1. red in h1. red. split; auto with sets. Qed. Lemma inb1_impl_inb2_le : forall {T:Type} (B:Bag T) (x:T), Inb1 B x -> Inb2_le B (x, B {->} x). intros; red; simpl; split; auto. Qed. Lemma inb2_le_impl_inb1 : forall {T:Type} (B:Bag T) (x:T) (n:nat), Inb2_le B (x, n) -> Inb1 B x. intros T B x n h1. red in h1. simpl in h1. red; destruct h1; omega. Qed. Definition Inclb {T:Type} (B B':Bag T) : Prop := forall x:T, Inb1 B x -> Inb2_le B' (x, B{->}x). Lemma inclb_refl : forall {T:Type} (B:Bag T), Inclb B B. intros; red; apply inb1_impl_inb2_le. Qed. Lemma inclb_trans : forall {T:Type} (B C D:Bag T), Inclb B C -> Inclb C D -> Inclb B D. intros T B C D h1 h2. red. red in h1. red in h2. intros x h3. specialize (h1 _ h3). pose proof (inb2_le_impl_inb1 _ _ _ h1) as h4. specialize (h2 _ h4). red in h1. red in h2. simpl in h1. simpl in h2. destruct h1 as [h1l h1r]. destruct h2 as [h2l h2r]. red. simpl. split. omega. assumption. Qed. Lemma bag_in_compat : forall {T:Type} (B:Bag T) (pr:T*nat), let S:= (bag_to_pair_set B) in Inn S pr -> snd pr = B {->} (fst pr). intros T B pr S h1. unfold bag_to_fin_map. destruct constructive_definite_description as [F h2]. simpl. symmetry. apply fin_map_to_fps_fin_map_app_compat. rewrite <- h2. assumption. Qed. Lemma in_dom_rel_inb1_compat : forall {T:Type} (B:Bag T) (x:T), Inn (dom_rel (bag_to_pair_set B)) x -> Inb1 B x. intros T B x h1. red. destruct h1 as [h1]. destruct h1 as [y h1]. pose proof (bag_in_compat _ _ h1) as h2. simpl in h2. rewrite <- h2. pose proof (nonzero_bag_to_pair_set B) as h3. red in h3. specialize (h3 x). destruct (zerop y) as [h4|]; auto with sets. rewrite h4 in h1. contradiction. Qed. Lemma inb1_dom_rel_compat_iff : forall {T:Type} (B:Bag T) (x:T), Inb1 B x <-> Inn (dom_rel (bag_to_pair_set B)) x. intros T B x. split. intro h1. apply inb1_dom_rel_compat; auto. apply in_dom_rel_inb1_compat. Qed. Lemma inb1_in_compat : forall {T:Type} (B:Bag T) (x:T), let S := bag_to_pair_set B in Inb1 B x -> Inn (bag_to_pair_set B) (x, B {->} x). intros T B x S h1. rewrite (bag_to_fin_map_compat B) at 1. apply in_fin_map_to_fps. apply inb1_dom_rel_compat. assumption. Qed. Lemma in_bag_to_set_iff : forall {T:Type} (B:Bag T) (x:T), Inn (bag_to_set B) x <-> Inb1 B x. intros T B x. split. intro h1. rewrite bag_to_set_dom_rel_compat in h1. apply in_dom_rel_inb1_compat. assumption. intro h1. apply Im_intro with (x, B{->}x); auto. apply inb1_in_compat; auto. Qed. Lemma inb2_in_compat : forall {T:Type} (B:Bag T) (pr:(T*nat)), Inb2 B pr -> Inn (bag_to_pair_set B) pr. intros T B pr h1. pose proof (inb2_impl_inb1 B pr h1) as h2. pose proof (inb1_in_compat B _ h2) as h3. rewrite bag_to_fin_map_compat. red in h1. destruct h1 as [h1l h1r]. rewrite surjective_pairing. rewrite <- h1l. apply in_fin_map_to_fps. constructor. exists (B{->}fst pr). assumption. Qed. Lemma in_inb2_compat : forall {T:Type} (B:Bag T) (pr:(T*nat)), Inn (bag_to_pair_set B) pr -> Inb2 B pr. intros T B pr h1. pose proof (fun_well_defined_bag_to_pair_set B) as h2. assert (h3:Inn (dom_rel (bag_to_pair_set B)) (fst pr)). constructor. exists (snd pr). rewrite <- surjective_pairing. assumption. pose proof (in_dom_rel_inb1_compat _ _ h3) as h4. pose proof (inb1_in_compat _ _ h4) as h5. rewrite surjective_pairing in h1. pose proof (fp_functional h2 _ _ _ h1 h5) as h6. red. split. rewrite h6. reflexivity. pose proof (nonzero_bag_to_pair_set B) as h7. red in h7. destruct (zerop (snd pr)) as [h8|]; auto with arith. specialize (h7 (fst pr)). rewrite <- h8 in h7. contradiction. Qed. Definition bag_set_unq_set {T:Type} (S:Ensemble (T*nat)) : Ensemble (T*nat) := [pr:(T*nat) | Inn S pr /\ (snd pr) > 0]. Lemma nonzero_bag_set_unq_set : forall {T:Type} (S:Ensemble (T*nat)), nonzero (bag_set_unq_set S). intros T S. red. intros x h1. inversion h1 as [h2]. clear h1. simpl in h2. destruct h2; omega. Qed. Lemma bag_set_unq_set_bag_to_pair_set : forall {T:Type} (B:Bag T), bag_set_unq_set (bag_to_pair_set B) = bag_to_pair_set B. intros T B. destruct B as [hdt S h1 h2 h3]. simpl. unfold bag_set_unq_set. red in h3. apply Extensionality_Ensembles. red. split. red. intros pr h4. destruct h4 as [h4]. destruct h4; assumption. red. intros pr h4. constructor. split; auto. destruct (zerop (snd pr)) as [h5|]; auto with arith. rewrite surjective_pairing in h4. rewrite h5 in h4. specialize (h3 (fst pr)). contradiction. Qed. Lemma bag_set_unq_set_incl : forall {T:Type} (S:Ensemble (T*nat)), Included (bag_set_unq_set S) S. intros T S pr h1. destruct h1 as [h1]. destruct h1; assumption. Qed. Lemma bag_set_unq_set_preserves_inclusion : forall {T:Type} (R S:Ensemble (T*nat)), Included R S -> Included (bag_set_unq_set R) (bag_set_unq_set S). intros T R S h1. red. intros pr h2. destruct h2 as [h2]. destruct h2 as [h2 h3]. constructor; auto with sets. Qed. Lemma bag_set_unq_set_fun_well_defined : forall {T:Type} {S:Ensemble (T*nat)}, fun_well_defined S -> fun_well_defined (bag_set_unq_set S). intros T S h1. pose proof (bag_set_unq_set_incl S) as h2. apply (fun_well_defined_incl _ _ h2). assumption. Qed. Lemma finite_bag_set_unq_set : forall {T:Type} (S:Ensemble (T*nat)), Finite S -> Finite (bag_set_unq_set S). intros T S h0. pose proof (bag_set_unq_set_incl S) as h1. apply (Finite_downward_closed _ _ h0 _ h1). Qed. Lemma im_fun1_nonzero : forall {T:Type} (A:Ensemble T), nonzero (Im A (fun x => (x, 1))). intros T A x. intro h1. inversion h1 as [x' h4 pr h5]. subst. inversion h5. Qed. Lemma im_fun1_fp : forall {T:Type} (A:Ensemble T), fun_well_defined (Im A (fun x => (x, 1))). intros T A. red. constructor. intros x h1. destruct h1 as [h2]. destruct h2 as [n h2]. inversion h2 as [x' h3 pr h4]. subst. inversion h4. subst. clear h4. exists 1. red. split. split. constructor. exists x'. assumption. assumption. intros n h4. destruct h4 as [h4l h4r]. inversion h4l as [h5]. destruct h5 as [x h5]. inversion h5 as [x'' h6 pr h7]. inversion h7. reflexivity. intros pr h1. inversion h1 as [x h2 pr' h3]. subst. simpl. split. constructor. exists 1. assumption. constructor. exists x. assumption. Qed. Definition set_to_bag {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (pf:Finite A) := bag_intro _ pfdt (Im A (fun x=>(x, 1))) (finite_image _ _ _ _ pf) (im_fun1_fp A) (im_fun1_nonzero A). Lemma empty_set_nonzero : forall (T:Type), nonzero (Empty_set (T*nat)). intros; red; auto with sets. Qed. Definition empty_bag (T:Type) (pfdt:type_eq_dec T) : Bag T := bag_intro _ pfdt _ (Empty_is_finite (T*nat)) (fun_well_defined_empty T nat) (empty_set_nonzero T). Lemma empty_bag_eq : forall {T:Type} (pfdt pfdt':type_eq_dec T), empty_bag _ pfdt = empty_bag _ pfdt'. intros. apply bag_intro_functional; auto. Qed. Lemma empty_pair_set_empty_bag_compat : forall {T:Type} (B:Bag T), (bag_to_pair_set B) = Empty_set _ -> B = empty_bag T (bag_type_eq_dec B). intros T B h1. destruct B as [hdt S h2 h3 h4]. simpl in h1. subst. unfold empty_bag. pose proof (proof_irrelevance _ h2 (Empty_is_finite (T*nat))). pose proof (proof_irrelevance _ h3 (fun_well_defined_empty T nat)). pose proof (proof_irrelevance _ h4 (empty_set_nonzero T)). subst; f_equal. apply type_eq_dec_eq. Qed. Lemma empty_bag_empty_pair_set_compat : forall (T:Type) (pf:type_eq_dec T), (bag_to_pair_set (empty_bag T pf)) = Empty_set _. intros T. unfold empty_bag. unfold bag_to_pair_set. reflexivity. Qed. Lemma empty_set_empty_bag_compat : forall {T:Type} (pf:type_eq_dec T) (B:Bag T), bag_to_set B = Empty_set _ -> B = empty_bag T pf. intros T hdt B h1. unfold bag_to_set in h1. pose proof (empty_image _ _ h1) as h2. pose proof (empty_pair_set_empty_bag_compat B h2) as h3. rewrite h3. f_equal. apply type_eq_dec_eq. Qed. Lemma empty_bag_empty_set_compat : forall (T:Type) (pfdt:type_eq_dec T), bag_to_set (empty_bag T pfdt) = Empty_set T. intros T hdt. unfold bag_to_set. rewrite empty_bag_empty_pair_set_compat. apply image_empty. Qed. Lemma empty_set_empty_bag_compat_iff : forall {T:Type} (B:Bag T), (bag_to_pair_set B) = Empty_set _ <-> B = empty_bag T (bag_type_eq_dec B). split. apply empty_pair_set_empty_bag_compat. intro h1. rewrite h1. apply empty_bag_empty_pair_set_compat. Qed. Lemma non_empty_bag_inhabited : forall {T:Type} (B:Bag T), B <> empty_bag T (bag_type_eq_dec B) -> exists x:T, Inb1 B x. intros T B h1. rewrite <- empty_set_empty_bag_compat_iff in h1. apply not_empty_Inhabited in h1. destruct h1 as [pr h1]. exists (fst pr). red. pose proof (nonzero_bag_to_pair_set B) as h2. red in h2. specialize (h2 (fst pr)). pose proof (bag_in_compat _ _ h1) as h3. rewrite <- h3. destruct (zerop (snd pr)) as [h4 | h5]. rewrite <- h4 in h2. rewrite <- surjective_pairing in h2. contradiction. red. assumption. Qed. Lemma all_not_inb1_empty_bag : forall {T:Type} (B:Bag T), (forall x:T, ~Inb1 B x) -> B = empty_bag T (bag_type_eq_dec B). intros T B h1. apply empty_pair_set_empty_bag_compat. apply Extensionality_Ensembles. red. split. red. intros pr h2. destruct B as [hdt S h3 h4 h5]. simpl in h2. red in h5. specialize (h1 (fst pr)). apply not_inb1 in h1. assert (h6:bag_to_pair_set (bag_intro T hdt S h3 h4 h5) = S). simpl. reflexivity. rewrite <- h6 in h2. pose proof (bag_in_compat _ _ h2) as h7. rewrite surjective_pairing in h2. rewrite h7 in h2. rewrite h1 in h2. pose proof (h5 (fst pr)). contradiction. auto with sets. Qed. Lemma empty_bag_O : forall {T:Type} (pf:type_eq_dec T) (x:T), (empty_bag T pf) {->} x = 0. intros T h0 x. unfold bag_to_fin_map. destruct (constructive_definite_description) as [F h1]. simpl. simpl in F. destruct F as [hdt h2 h3 S h4]. pose proof h4 as h4'. rewrite dom_rel_empty in h4'. pose proof (fp_empty1_s _ _ h4'). subst. simpl. apply fin_fps_to_f_def. intro h5. rewrite dom_rel_empty in h5. contradiction. Qed. Lemma inhabited_non_empty_bag : forall {T:Type} (B:Bag T), (exists x:T, Inb1 B x) -> B <> empty_bag T (bag_type_eq_dec B). intros T B h1 h2. destruct h1 as [x h1]. rewrite h2 in h1. red in h1. pose proof (empty_bag_O (bag_type_eq_dec B) x). omega. Qed. Lemma not_inb1_empty_bag : forall (T:Type) (pfdt:type_eq_dec T) (x:T), ~Inb1 (empty_bag T pfdt) x. intros T h0 x h1. red in h1. pose proof (empty_bag_O h0 x) as h2. omega. Qed. Lemma empty_bag_inclb : forall {T:Type} (B:Bag T), Inclb (empty_bag T (bag_type_eq_dec B)) B. intros T B. red. intros x ?. pose proof (not_inb1_empty_bag T (bag_type_eq_dec B) x). contradiction. Qed. Lemma empty_bag_eq_dec : forall {T:Type} (pfdt:type_eq_dec T) (B:Bag T), {B = empty_bag T pfdt} + {B <> empty_bag T pfdt}. intros T pfdt B. pose proof (finite_bag_to_pair_set B) as h1. destruct (finite_inh_or_empty_dec _ h1) as [h2 | h2]. right. intro. subst. destruct h2 as [pr h2]. simpl in h2. contradiction. left. erewrite empty_bag_eq. apply empty_pair_set_empty_bag_compat; auto. Qed. Section BagAdd. (*The peculiar definition accounts for "zero adds" and duplicate adds. For "cumulative" pair adds, see [bag_add_n] functionality *) Definition bag_to_pair_set_add {T:Type} (B:Bag T) (pr:T*nat) := bag_set_unq_set (Add ((Subtract (bag_to_pair_set B) (fst pr, B {->} fst pr))) pr). Lemma finite_bag_to_pair_set_add : forall {T:Type} (B:Bag T) (pr:T*nat), Finite (bag_to_pair_set_add B pr). intros T B pr. unfold bag_to_pair_set_add. pose proof (bag_set_unq_set_incl (Add (Subtract (bag_to_pair_set B) (fst pr, B{->}fst pr)) pr)) as h1. eapply Finite_downward_closed. pose proof (finite_bag_to_pair_set B) as h2. pose proof (incl_subtract (bag_to_pair_set B) (fst pr, B{->}fst pr)) as h3. pose proof (Finite_downward_closed _ _ h2 _ h3) as h4. pose proof (Add_preserves_Finite _ _ pr h4) as h5. apply h5. assumption. Qed. Lemma fun_well_defined_bag_to_pair_set_add : forall {T:Type} (B:Bag T) (pr:T*nat), fun_well_defined (bag_to_pair_set_add B pr). intros T B pr. unfold bag_to_pair_set_add. pose proof (bag_set_unq_set_incl (Add (Subtract (bag_to_pair_set B) (fst pr, B{->}fst pr)) pr)) as h1. eapply fun_well_defined_incl; simpl. apply h1. rewrite (surjective_pairing pr). apply fun_well_defined_add. simpl. pose proof (incl_subtract (bag_to_pair_set B) (fst pr, B{->}fst pr)) as h2. eapply fun_well_defined_incl; simpl. apply h2. apply fun_well_defined_bag_to_pair_set. intro h2. simpl in h2. destruct h2 as [h2]. destruct h2 as [n h2]. inversion h2 as [h3 h4]. pose proof (bag_in_compat _ _ h3) as h5. simpl in h5. contradict h4. subst. constructor. Qed. Lemma nonzero_bag_to_pair_set_add : forall {T:Type} (B:Bag T) (pr:T*nat), nonzero (bag_to_pair_set_add B pr). intros T B pr. red. intros x h1. inversion h1 as [h2]. clear h1. simpl in h2. destruct h2; omega. Qed. Definition bag_pair_add {T:Type} (B:Bag T) (pr:T*nat) := bag_intro _ (bag_type_eq_dec B) _ (finite_bag_to_pair_set_add B pr) (fun_well_defined_bag_to_pair_set_add B pr) (nonzero_bag_to_pair_set_add B pr). Lemma bag_pair_add_app_eq : forall {T:Type} (B:Bag T) (pr:T*nat), bag_pair_add B pr{->}fst pr = snd pr. intros T B pr. destruct (zerop (snd pr)) as [h1 | h2]. rewrite h1. pose proof (bag_to_fin_map_compat' (bag_pair_add B pr) (fst pr)) as h2. simpl in h2. simpl. rewrite h2. apply fin_fps_to_f_def. intro h3. destruct h3 as [h3]. destruct h3 as [n h3]. inversion h3 as [h4]. clear h3. simpl in h4; destruct h4 as [h4]. inversion h4 as [? h5 | ? h6]; subst. clear h4. contradict h5. intro h3. inversion h3 as [h4 h5]. contradict h5. pose proof (bag_in_compat _ _ h4) as h5. simpl in h5. subst. constructor. inversion h6 as [h7]. clear h6. rewrite (surjective_pairing pr) in h7. simpl in h7. inversion h7. subst. omega. pose proof h2 as h1. clear h2. pose proof (bag_to_fin_map_compat' (bag_pair_add B pr) (fst pr)) as h2. simpl in h2. simpl. rewrite h2. assert (h3:Inn (bag_to_pair_set_add B pr) pr). constructor. split; auto with arith. right. constructor. pose proof (fp_inv (bag_type_eq_dec B) _ _ _ 0 (finite_dom_rel_bag_to_pair_set (bag_pair_add B pr)) pr (fun_well_defined_bag_to_pair_set (bag_pair_add B pr)) h3) as h4. rewrite <- h4. f_equal. apply type_eq_dec_eq. Qed. Lemma bag_pair_add_intro1 : forall {T:Type} (B:Bag T) (pr pr':T*nat), (fst pr <> fst pr') -> Inb2 B pr -> Inb2 (bag_pair_add B pr') pr. intros T B pr pr' h0 h1. red. red in h1. destruct h1 as [h1l h1r]. split; auto. assert (h5:Inn (bag_to_pair_set_add B pr') pr). constructor. split; auto with arith. left. constructor. rewrite <- h1l in h1r. pose proof (inb1_in_compat _ _ h1r) as h5. rewrite h1l in h5. rewrite surjective_pairing. assumption. intro h5. rewrite surjective_pairing in h5. inversion h5. apply neq_sym in h0. contradiction. pose proof (fp_inv (bag_type_eq_dec B) _ _ _ 0 (finite_dom_rel_bag_to_pair_set (bag_pair_add B pr')) pr (fun_well_defined_bag_to_pair_set (bag_pair_add B pr')) h5) as h7. pose proof (bag_to_fin_map_compat' (bag_pair_add B pr') (fst pr)) as h6. rewrite h6. rewrite <- h7. f_equal. apply type_eq_dec_eq. Qed. Lemma bag_pair_add_intro2 : forall {T:Type} (B:Bag T) (pr:T*nat), (snd pr > 0) -> Inb2 (bag_pair_add B pr) pr. intros T B pr h1. apply in_inb2_compat. constructor. split; auto. right. constructor. Qed. Lemma bag_pair_add_app_neq : forall {T:Type} (B:Bag T) (pr:T*nat) (x:T), x <> fst pr -> (bag_pair_add B pr) {->} x = B {->} x. intros T B pr x h1. destruct (inb1_dec B x) as [h2 | h3]. pose proof (inb1_inb2_compat _ _ h2) as h3. assert (h4:x = fst (x, B{->}x)). reflexivity. rewrite h4 in h1. pose proof (bag_pair_add_intro1 _ _ _ h1 h3) as h5. pose proof (inb2_in_compat _ _ h5) as h6. pose proof (bag_in_compat _ _ h6) as h7. simpl in h7. rewrite h7. reflexivity. pose proof (not_inb1 _ _ h3) as h4. rewrite h4. assert (h5:~Inb1 (bag_pair_add B pr) x). intro h5. pose proof (inb1_dom_rel_compat _ _ h5) as h6. destruct h6 as [h6l]. destruct h6l as [n h6l]. inversion h6l as [h7]. clear h6l. simpl in h7. destruct h7 as [h7l h7r]. inversion h7l as [? h8 | ? h9]; subst. clear h7l. inversion h8 as [h9]. clear h8. pose proof (bag_in_compat _ _ h9) as h10. simpl in h10. omega. inversion h9. subst. simpl in h1. contradiction h1. reflexivity. apply not_inb1. assumption. Qed. Definition bag_fun_well_defined_add_set {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pfsf:fun_well_defined C) (x:T) := let n := fin_fps_to_f pfdt pff _ pfsf 0 x in (Subtract (Add C (x, S n)) (x, n)). Lemma in_bag_fun_well_defined_add_set : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pfsf:fun_well_defined C) (x:T) (n:nat), Inn (bag_fun_well_defined_add_set pfdt pff pfsf x) (x, n) -> n = S (fin_fps_to_f pfdt pff _ pfsf 0 x). intros T C hdt hf h1 x n h2. unfold bag_fun_well_defined_add_set in h2. rewrite sub_add_comm in h2. inversion h2 as [pr h2l | pr h2r]. subst. destruct h2l as [h3 h4]. rewrite (fin_fps_to_f_compat' hdt hf h1 0) in h3. destruct h3 as [h3]. simpl in h3. destruct h3; subst. contradict h4. constructor. subst. inversion h2r. reflexivity. intro h3. inversion h3; omega. Qed. Lemma in_bag_fun_well_defined_add_set_neq : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pfsf:fun_well_defined C) (x x':T) (n:nat), x <> x' -> Inn (bag_fun_well_defined_add_set pfdt pff pfsf x) (x', n) -> n = (fin_fps_to_f pfdt pff _ pfsf 0 x'). intros T C hdt hf h1 x x' n h2 h3. unfold bag_fun_well_defined_add_set in h3. rewrite sub_add_comm in h3. inversion h3 as [? h4 | ? h5]. subst. inversion h4 as [h5 h6]. rewrite (fin_fps_to_f_compat' hdt hf h1 0) in h5. destruct h5 as [h5]. simpl in h5. destruct h5 as [h5l h5r]. assumption. subst. inversion h5. subst. contradict h2. reflexivity. intro h4. inversion h4. omega. Qed. Lemma in_bag_fun_well_defined_add_set_neq' : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pfsf:fun_well_defined C) (x x':T), x <> x' -> Inn (dom_rel C) x' -> Inn (bag_fun_well_defined_add_set pfdt pff pfsf x) (x', fin_fps_to_f pfdt pff _ pfsf 0 x'). intros T C hdt hf h1 x x' h2 h3. constructor. left. rewrite (fin_fps_to_f_compat' hdt hf h1 0) at 1. constructor. simpl. split; auto. intro h4. inversion h4. contradiction. Qed. Lemma dom_rel_bag_fun_well_defined_add : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pf:fun_well_defined C) (x:T), dom_rel (bag_fun_well_defined_add_set pfdt pff pf x) = Add (dom_rel C) x. intros T C hdt hf h1 x. destruct (classic (Inn (dom_rel C) x)) as [h2 | h3]. unfold bag_fun_well_defined_add_set. rewrite sub_add_comm. apply Extensionality_Ensembles. red. split. red. intros x' h3. rewrite dom_rel_add in h3. inversion h3 as [? h4 | ? h5]. subst. pose proof (incl_subtract C (x, fin_fps_to_f hdt hf _ h1 0 x)) as h5. pose proof (dom_rel_incl _ _ h5) as h6. left. auto with sets. subst. simpl in h5. inversion h5. subst. right. constructor. rewrite dom_rel_add. simpl. rewrite in_add_eq. rewrite <- (dom_rel_add _ (x, fin_fps_to_f hdt hf _ h1 0 x)). apply dom_rel_incl. rewrite add_sub_compat_in. auto with sets. rewrite (fin_fps_to_f_compat' hdt hf h1 0) at 1. constructor. simpl. split; auto. assumption. intro h3. inversion h3. omega. unfold bag_fun_well_defined_add_set. rewrite subtract_nin. rewrite dom_rel_add. simpl. reflexivity. intro h4. inversion h4 as [? h5 | ? h6]; subst. pose proof (fp_in_dom _ _ _ h1 (x, fin_fps_to_f hdt hf _ h1 0 x) h5) as h6. simpl in h6. contradiction. inversion h6. omega. Qed. Lemma finite_bag_fun_well_defined_add_set : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pf:fun_well_defined C) (x:T), Finite C -> Finite (bag_fun_well_defined_add_set pfdt pff pf x). intros T C hdt hf h1 x h2. unfold bag_fun_well_defined_add_set. assert (h3:Included (Subtract (Add C (x, S (fin_fps_to_f hdt hf _ h1 0 x))) (x, fin_fps_to_f hdt hf _ h1 0 x)) (Add C (x, S (fin_fps_to_f hdt hf _ h1 0 x)))). apply incl_subtract. assert (h4:Finite (Add C (x, S (fin_fps_to_f hdt hf _ h1 0 x)))). apply Add_preserves_Finite. assumption. apply (Finite_downward_closed _ _ h4 _ h3). Qed. (*Get rid of classic IN some day*) Lemma fun_well_defined_bag_fun_well_defined_add_set : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pf:fun_well_defined C) (x:T), fun_well_defined (bag_fun_well_defined_add_set pfdt pff pf x). intros T C hdt hf h1 x. destruct (classic (Inn (dom_rel C) x)) as [h2 | h3]. constructor. intros x' h3. rewrite dom_rel_bag_fun_well_defined_add in h3. rewrite in_add_eq in h3. destruct (classic (x = x')) as [h4 | h5]. subst. exists (S (fin_fps_to_f hdt hf _ h1 0 x')). red. split. split. constructor. exists x'. unfold bag_fun_well_defined_add_set. rewrite sub_add_comm. right. constructor. intro h4. inversion h4. omega. unfold bag_fun_well_defined_add_set. rewrite sub_add_comm. right. constructor. intro h4. inversion h4. omega. intros n h4. destruct h4 as [h4l h4r]. symmetry. apply in_bag_fun_well_defined_add_set. assumption. exists (fin_fps_to_f hdt hf _ h1 0 x'). red. split. split. constructor. exists x'. apply in_bag_fun_well_defined_add_set_neq'; auto. apply in_bag_fun_well_defined_add_set_neq'; auto. intros n h6. destruct h6 as [h6l h6r]. symmetry. apply (in_bag_fun_well_defined_add_set_neq hdt hf h1 x x' n). assumption. assumption. assumption. intros pr h3. split. constructor. exists (snd pr). rewrite <- surjective_pairing. assumption. constructor. exists (fst pr). rewrite <- surjective_pairing. assumption. unfold bag_fun_well_defined_add_set. rewrite subtract_nin. apply fun_well_defined_add. assumption. assumption. intro h4. inversion h4 as [? h5 | ? h6]. subst. rewrite (fin_fps_to_f_compat' hdt hf h1 0) in h5 at 1. destruct h5 as [h5]. simpl in h5. destruct h5; contradiction. subst. inversion h6 as [h7]. omega. Qed. Lemma nonzero_bag_fun_well_defined_add_set : forall {T:Type} (pfdt:type_eq_dec T) (C:Ensemble (T*nat)) (pff:Finite (dom_rel C)), nonzero C -> forall (pffp:fun_well_defined C) (x:T), nonzero (bag_fun_well_defined_add_set pfdt pff pffp x). intros T hdt C hf h1 h2 x. red. red in h1. intros x' h3. inversion h3 as [h4 h5]. clear h3. inversion h4 as [? h6 | ? h7]. clear h4. subst. specialize (h1 x'). contradiction. subst. inversion h7; subst. Qed. Definition bag_add {T:Type} (B:Bag T) (x:T) := bag_intro _ (bag_type_eq_dec B) _ (finite_bag_fun_well_defined_add_set (bag_type_eq_dec B) (finite_dom_rel_bag_to_pair_set B) (fun_well_defined_bag_to_pair_set B) x (finite_bag_to_pair_set B)) (fun_well_defined_bag_fun_well_defined_add_set (bag_type_eq_dec B) (finite_dom_rel_bag_to_pair_set B) (fun_well_defined_bag_to_pair_set B) x) (nonzero_bag_fun_well_defined_add_set (bag_type_eq_dec B) _ (finite_dom_rel_bag_to_pair_set B) (nonzero_bag_to_pair_set B) (fun_well_defined_bag_to_pair_set B) x). Lemma bag_add_intro1 : forall {T:Type} (B:Bag T) (a:T), Inb1 B a -> forall x:T, Inb1 (bag_add B x) a. intros T B a h1 x. apply in_dom_rel_inb1_compat. pose proof (inb1_dom_rel_compat _ _ h1) as h2. simpl. rewrite dom_rel_bag_fun_well_defined_add. left. assumption. Qed. Lemma bag_add_intro2 : forall {T:Type} (B:Bag T) (x:T), Inb1 (bag_add B x) x. intros T B x. apply in_dom_rel_inb1_compat. simpl. rewrite dom_rel_bag_fun_well_defined_add. right. constructor. Qed. Lemma bag_add_app_eq : forall {T:Type} (B:Bag T) (x:T), (bag_add B x) {->} x = S (B {->} x). intros T B x. pose proof (bag_add_intro2 B x) as h1. pose proof (fun_well_defined_bag_to_pair_set B) as h2. pose proof (in_bag_fun_well_defined_add_set (bag_type_eq_dec B) (finite_dom_rel_bag_to_pair_set B) h2 x ((bag_add B x) {->} x)) as h3. pose proof (inb1_in_compat _ _ h1) as h4. simpl in h4. pose proof (proof_irrelevance _ h2 (fun_well_defined_bag_to_pair_set B)); subst. specialize (h3 h4). rewrite h3. f_equal. symmetry. apply bag_to_fin_map_compat'. Qed. Lemma bag_add_app_neq : forall {T:Type} (B:Bag T) (a x:T), a <> x -> (bag_add B a) {->} x = B {->} x. intros T B a x h1. destruct (zerop (B{->} x)) as [h2 | h3]. rewrite h2. destruct (zerop (bag_add B a{->}x)) as [h4 | h5]. rewrite h4. reflexivity. assert (h6:Inb1 (bag_add B a) x). red. auto with arith. pose proof (inb1_dom_rel_compat _ _ h6) as h7. simpl in h7. rewrite dom_rel_bag_fun_well_defined_add in h7. destruct h7 as [x h7l | x h7r]. apply in_dom_rel_inb1_compat in h7l. red in h7l. omega. destruct h7r. contradict h1. reflexivity. assert (h4:Inb1 B x). red. auto with arith. pose proof (bag_add_intro1 B x h4 a) as h7. pose proof (inb1_inb2_compat _ _ h7) as h8. pose proof (inb2_in_compat _ _ h8) as h9. simpl in h9. pose proof (in_bag_fun_well_defined_add_set_neq (bag_type_eq_dec B) (finite_dom_rel_bag_to_pair_set B) (fun_well_defined_bag_to_pair_set B) _ _ _ h1 h9) as h10. rewrite h10 at 1. rewrite bag_to_fin_map_compat'. reflexivity. Qed. Lemma bag_add_inb2_compat : forall {T:Type} (B:Bag T) (x:T) (n:nat), Inb2 B (x, n) -> Inb2 (bag_add B x) (x, S n). intros T B x n h1. pose proof (inb2_in_compat _ _ h1) as h2. pose proof (bag_in_compat _ _ h2) as h3. apply in_inb2_compat. simpl in h3. rewrite h3. simpl. constructor. right. pose proof (bag_to_fin_map_compat' B x) as h4. rewrite h4. constructor. intro h4. rewrite (bag_to_fin_map_compat' B x) in h4. inversion h4 as [h5]. omega. Qed. Lemma bag_add_inv : forall {T:Type} (pfdt:type_eq_dec T) (B:Bag T) (x:T) (pr:T*nat), Inb2 (bag_add B x) pr -> (Inb2 B pr \/ fst pr = x) /\ (fst pr = x <-> ~Inb2 B pr). intros T h0 B x pr h1. split. red in h1. destruct h1 as [h1l h1r]. destruct (h0 (fst pr) x) as [h2 | h3]. right. assumption. rewrite bag_add_app_neq in h1l. left. assert (h4:Inb1 B (fst pr)). red. rewrite h1l. assumption. pose proof (inb1_inb2_compat _ _ h4) as h5. rewrite h1l in h5. rewrite surjective_pairing. assumption. apply neq_sym. assumption. split. intros h2 h3. subst. pose proof (inb2_in_compat _ _ h3) as h4. clear h3. pose proof (inb2_in_compat _ _ h1) as h5. clear h1. simpl in h5. destruct h5 as [h5 h6]. contradict h6. pose proof (bag_in_compat _ _ h4) as h6. rewrite <- bag_to_fin_map_compat'. rewrite <- h6. rewrite surjective_pairing. constructor. intro h2. red in h1. destruct h1 as [h1l h1r]. apply NNPP. intro h3. apply neq_sym in h3. rewrite (bag_add_app_neq B _ _ h3) in h1l. assert (h4:Inb2 B pr). red. split; auto. contradiction. Qed. Lemma bag_to_set_add : forall {T:Type} (B:Bag T) (x:T), bag_to_set (bag_add B x) = Add (bag_to_set B) x. intros T B x. unfold bag_to_set. assert (h4:x = fst (x, B{->}x)). reflexivity. rewrite h4 at 2. rewrite <- Im_add. rewrite <- dom_rel_eq. rewrite <- dom_rel_eq. destruct (inb1_dec B x) as [h1 | h2]. rewrite in_add_eq. simpl. rewrite dom_rel_bag_fun_well_defined_add. pose proof (inb1_dom_rel_compat _ _ h1) as h5. rewrite in_add_eq; auto. apply inb2_in_compat. apply inb1_inb2_compat. assumption. pose proof (not_inb1 _ _ h2) as h3. rewrite h3. simpl. rewrite dom_rel_add. simpl. rewrite dom_rel_bag_fun_well_defined_add. f_equal. Qed. Fixpoint bag_add_n {T:Type} (B:Bag T) (x:T) (n:nat) := match n with | O => B | S n => bag_add (bag_add_n B x n) x end. Lemma le_bag_add_n_app : forall {T:Type} (B:Bag T) (x:T) (n:nat), n > 0 -> bag_add B x {->} x <= bag_add_n B x n {->} x. intros T B x n. revert T B x. induction n as [|n h1]; simpl. intros; omega. intros T B x h2. pose proof (bag_add_app_eq (bag_add_n B x n) x) as he. simpl. simpl in he. rewrite he. clear he. destruct (zerop n) as [h3 | h4]. subst. simpl. pose proof (bag_add_app_eq B x) as he. simpl in he. omega. specialize (h1 _ B x h4). simpl in h1. omega. Qed. Lemma bag_add_n_app_neq : forall {T:Type} (B:Bag T) (a x:T) (n:nat), a <> x -> (bag_add_n B a n) {->} x = B {->} x. intros T B a x n. revert T B a x. induction n as [|n h1]. simpl. auto. intros T B a x h2. specialize (h1 T B a x h2). simpl. pose proof (bag_add_app_neq (bag_add_n B a n) _ _ h2) as h3. simpl in h3. rewrite h3. assumption. Qed. (* Maybe add a version of this to bag_subtract*) Lemma bag_add_n_app_eq : forall {T:Type} (B:Bag T) (x:T) (n:nat), (bag_add_n B x n) {->} x = B {->} x + n. intros T B x n. revert B x. induction n as [|n h1]; simpl; auto with arith. intros B x. assert (h2:B{->}x + S n = S (B{->}x + n)). omega. rewrite h2. specialize (h1 B x). pose proof (bag_add_app_eq (bag_add_n B x n) x) as h3. simpl in h3. simpl. rewrite h3. rewrite h1. reflexivity. Qed. Lemma bag_add_n_bag_pair_add_compat : forall {T:Type} (B:Bag T) (x:T) (n:nat), ~Inb1 B x -> bag_add_n B x n = bag_pair_add B (x, n). intros T B a n h0. apply bag_to_fin_map_inj. apply functional_extensionality. intro x. destruct (eq_dec a x) as [h1 | h2]. subst. rewrite bag_pair_add_app_eq. simpl. rewrite bag_add_n_app_eq. pose proof (not_inb1 _ _ h0) as h1. rewrite h1. simpl. reflexivity. rewrite bag_pair_add_app_neq; auto. rewrite bag_add_n_app_neq; auto. Qed. Lemma bag_pair_add_bag_add_n_compat : forall {T:Type} (B:Bag T) (pr:T*nat), B{->}fst pr <= snd pr -> bag_pair_add B pr = bag_add_n B (fst pr) (snd pr - (B{->}(fst pr))). intros T B pr h0. apply bag_to_fin_map_inj. apply functional_extensionality. intro x. destruct (eq_dec (fst pr) x) as [h1 | h2]. subst. rewrite bag_pair_add_app_eq. rewrite bag_add_n_app_eq. simpl. omega. rewrite bag_pair_add_app_neq. rewrite bag_add_n_app_neq. reflexivity. assumption. auto. Qed. End BagAdd. Section BagSubtract. Lemma finite_bag_to_pair_set_subtract : forall {T:Type} (B:Bag T) (pr:T*nat), Finite (Subtract (bag_to_pair_set B) pr). intros T B pr. eapply Finite_downward_closed. apply finite_bag_to_pair_set. apply incl_subtract. Qed. Lemma fun_well_defined_bag_to_pair_set_subtract : forall {T:Type} (B:Bag T) (pr:T*nat), fun_well_defined (Subtract (bag_to_pair_set B) pr). intros T B pr. eapply fun_well_defined_incl. apply incl_subtract. apply fun_well_defined_bag_to_pair_set. Qed. Lemma nonzero_bag_to_pair_set_subtract : forall {T:Type} (B:Bag T) (pr:T*nat), nonzero (Subtract (bag_to_pair_set B) pr). intros T B R. eapply nonzero_incl. apply incl_subtract. apply nonzero_bag_to_pair_set. Qed. Definition bag_pair_subtract {T:Type} (B:Bag T) (pr:T*nat) := bag_intro _ (bag_type_eq_dec B) _ (finite_bag_to_pair_set_subtract B pr) (fun_well_defined_bag_to_pair_set_subtract B pr) (nonzero_bag_to_pair_set_subtract B pr). Lemma bag_to_pair_set_bag_pair_subtract : forall {T:Type} (B:Bag T) (pr:T*nat), bag_to_pair_set (bag_pair_subtract B pr) = Subtract (bag_to_pair_set B) pr. intros T B pr. unfold bag_to_pair_set, bag_pair_subtract, bag_to_pair_set. reflexivity. Qed. Lemma bag_pair_subtract_O : forall {T:Type} (B:Bag T) (x:T), bag_pair_subtract B (x, O) = B. intros T B x. destruct B as [hdt S h1 h2 h3]. unfold bag_pair_subtract. apply bag_intro_functional. simpl. assert (h4:~Inn S (x, 0)). intro h5. red in h3. specialize (h3 x). contradiction. apply subtract_nin. assumption. Qed. Lemma bag_pair_subtract_ninb2 : forall {T:Type} (B:Bag T) (pr:T*nat), ~Inb2 B pr -> bag_pair_subtract B pr = B. intros T B pr h1. destruct B as [hdt S h2 h3 h4]. unfold bag_pair_subtract. apply bag_intro_functional. unfold bag_to_pair_set. apply Extensionality_Ensembles. red. split. apply incl_subtract. red. intros pr' h5. constructor; auto. intro h6. destruct h6; subst. contradict h1. red. split. symmetry. apply bag_in_compat. simpl. assumption. red in h4. specialize (h4 (fst pr)). red in h3. pose proof h5 as h5'. rewrite (fin_fps_to_f_compat' hdt (dom_rel_finite _ h2) h3 0) in h5'. destruct h5' as [h5']. destruct h5' as [h5l h5r]. destruct (zerop (snd pr)) as [h6|]; auto with arith. rewrite <- h6 in h4. rewrite <- surjective_pairing in h4. contradiction. Qed. Lemma bag_to_pair_set_subtract_cons : forall {T:Type} (B:Bag T) (l:list (T*nat)) (pr:T*nat), NoDup (pr::l) -> bag_to_pair_set B = list_to_set (pr::l) -> bag_to_pair_set (bag_pair_subtract B pr) = list_to_set l. intros T B l pr h0 h1. pose proof (no_dup_cons_nin _ _ h0) as h0'. pose proof (fun_well_defined_bag_to_pair_set B) as h2'. rewrite h1 in h2'. simpl in h1. apply Extensionality_Ensembles. red. split. red. intros pr' h2. simpl in h2. destruct h2 as [h2 h3]. rewrite h1 in h2. destruct h2 as [pr' h2l | pr' h2r]. assumption. contradiction. red. intros pr' h2. simpl. constructor. pose proof (Add_intro1 _ _ pr _ h2) as h3. rewrite <- h1 in h3. assumption. rewrite <- list_to_set_in_iff in h2. intro h3. destruct h3. contradiction. Qed. Definition bag_fun_well_defined_subtract_set {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pfsf:fun_well_defined C) (x:T) := let n := fin_fps_to_f pfdt pff _ pfsf 0 x in bag_set_unq_set (Subtract (Add C (x, pred n)) (x, n)). Lemma finite_bag_fun_well_defined_subtract_set : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pf:fun_well_defined C) (x:T), Finite C -> Finite (bag_fun_well_defined_subtract_set pfdt pff pf x). intros T C hdt hf h1 x h2. unfold bag_fun_well_defined_subtract_set. assert (h3:Included (bag_set_unq_set (Subtract (Add C (x, pred (fin_fps_to_f hdt hf C h1 0 x))) (x, fin_fps_to_f hdt hf C h1 0 x))) (bag_set_unq_set (Add C (x, pred (fin_fps_to_f hdt hf C h1 0 x))))). assert (h3:Included (Subtract (Add C (x, pred (fin_fps_to_f hdt hf C h1 0 x))) (x, fin_fps_to_f hdt hf C h1 0 x)) (Add C (x, pred (fin_fps_to_f hdt hf C h1 0 x)))). apply incl_subtract. apply bag_set_unq_set_preserves_inclusion. assumption. assert (h4:Finite (Add C (x, pred (fin_fps_to_f hdt hf C h1 0 x)))). apply Add_preserves_Finite. assumption. pose proof (finite_bag_set_unq_set _ h4) as h5. apply (Finite_downward_closed _ _ h5 _ h3). Qed. Lemma dom_rel_subtract_fun_well_defined : forall {T U:Type} (S:Ensemble (T*U)) (pr:T*U), fun_well_defined S -> (~Inn S pr -> ~Inn (dom_rel S) (fst pr)) -> dom_rel (Subtract S pr) = Subtract (dom_rel S) (fst pr). intros T U S pr h0 him. destruct (classic_dec (Inn S pr)) as [hin | hnin]. apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [h1]. destruct h1 as [y h1]. inversion h1 as [h2 h3]. constructor. constructor. exists y. assumption. intro h5. destruct h5. rewrite surjective_pairing in hin. pose proof (fp_functional h0 _ _ _ hin h2). subst. contradict h3. rewrite <- surjective_pairing. constructor. red. intros x h1. constructor. destruct h1 as [h1 h2]. destruct h1 as [h1]. destruct h1 as [y h1]. exists y. constructor; auto. intro h3. inversion h3. subst. simpl in h2. contradict h2. constructor. rewrite subtract_nin. specialize (him hnin). rewrite subtract_nin; auto. assumption. Qed. Lemma bag_set_unq_set_sub_x_O : forall {T:Type} (C:Ensemble (T*nat)) (x:T), bag_set_unq_set (Subtract C (x, O)) = bag_set_unq_set C. intros T C x. unfold bag_set_unq_set. apply Extensionality_Ensembles. red. split. red. intros pr h1. constructor. destruct h1 as [h1]. destruct h1 as [h1l h1r]. destruct h1l; split; auto. red. intros pr h1. destruct h1 as [h1]. destruct h1 as [h1l h1r]. constructor. split. constructor. assumption. intros h2. destruct h2. simpl in h1r. omega. assumption. Qed. Lemma bag_set_unq_set_sub_x_n : forall {T:Type} (C:Ensemble (T*nat)) (x:T) (n:nat), bag_set_unq_set (Subtract C (x, n)) = Subtract (bag_set_unq_set C) (x, n). intros T C x n. apply Extensionality_Ensembles. red. split. red. intros pr h1. destruct h1 as [h1]. destruct h1 as [h1 h2]. constructor. constructor. destruct h1 as [h1l h1r]. split; auto. intro h3. destruct h3. destruct h1 as [h1l h1r]. contradict h1r. constructor. red. intros pr h1. destruct h1 as [h1]. destruct h1 as [h1]. destruct h1 as [h1l h1r]. constructor. split. constructor; auto. assumption. Qed. Lemma bag_set_unq_set_add_x_O : forall {T:Type} (C:Ensemble (T*nat)) (x:T), bag_set_unq_set (Add C (x, O)) = bag_set_unq_set C. intros T C x. unfold bag_set_unq_set. apply Extensionality_Ensembles. red. split. red. intros pr h1. destruct h1 as [h1]. destruct h1 as [h1l h1r]. destruct h1l as [pr h1a | pr h1b]. constructor. split; auto. destruct h1b. simpl in h1r. omega. red. intros pr h1. destruct h1 as [h1]. destruct h1 as [h1l h1r]. constructor. split. left. assumption. assumption. Qed. Lemma bag_set_unq_set_add_x_Sn : forall {T:Type} (C:Ensemble (T*nat)) (x:T) (n:nat), 0 < n -> bag_set_unq_set (Add C (x, n)) = Add (bag_set_unq_set C) (x, n). intros T C x n h1. apply Extensionality_Ensembles. red. split. red. intros pr h2. destruct h2 as [h2]. destruct h2 as [h2l h2r]. destruct h2l as [pr h2a | pr h2b]. left. constructor. split; auto. destruct h2b; subst. right. constructor. red. intros pr h2. constructor. split. destruct h2 as [pr h2l | pr h2r]. destruct h2l as [h2l]. left. destruct h2l; auto. destruct h2r; subst. right. constructor. destruct h2 as [pr h2l | pr h2r]. destruct h2l as [h2l]. destruct h2l; auto. destruct h2r. simpl. omega. Qed. (*do a in_fin_ens_dec destructiion here instead of classic.*) Lemma dom_rel_bag_fun_well_defined_subtract : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pf:fun_well_defined C) (x:T), let n := fin_fps_to_f pfdt pff _ pf O x in dom_rel (bag_fun_well_defined_subtract_set pfdt pff pf x) = if (zerop (pred n)) then Subtract (dom_rel (bag_set_unq_set C)) x else dom_rel (bag_set_unq_set C). intros T C hdt hff h1 x n. destruct (classic (Inn (dom_rel C) x)) as [h2 | h3]; destruct (zerop (pred n)) as [h4 | h5]. pose proof (pred0_dec _ h4) as h5. clear h4. destruct h5 as [h5l | h5r]; subst. unfold bag_fun_well_defined_subtract_set. rewrite h5l. simpl. destruct (classic (Inn C (x, O))) as [h6 | h7]. rewrite sub_add_same_in. rewrite bag_set_unq_set_sub_x_O. assert (h7:~Inn (dom_rel (bag_set_unq_set C)) x). intro h8. destruct h8 as [h8]. destruct h8 as [y h8]. destruct h8 as [h8]. destruct h8 as [h8l h8r]. simpl in h8r. pose proof (fp_functional h1 _ _ _ h6 h8l). subst. omega. rewrite subtract_nin. reflexivity. assumption. assumption. contradict h7. rewrite <- h5l. apply fin_fps_to_f_s_compat. assumption. unfold bag_fun_well_defined_subtract_set. rewrite h5r. simpl. rewrite sub_add_comm. rewrite bag_set_unq_set_add_x_O. rewrite bag_set_unq_set_sub_x_n. rewrite dom_rel_subtract_fun_well_defined. simpl. reflexivity. apply fun_well_defined_incl with C. red. intros pr h6. destruct h6 as [h6]. destruct h6; auto. assumption. simpl. rewrite <- h5r. intro h6. intro h7. contradict h6. pose proof (bag_set_unq_set_incl C) as h8. pose proof (dom_rel_incl _ _ h8) as h9. assert (h10:Inn (dom_rel C) x). auto with sets. pose proof (fin_fps_to_f_s_compat hdt hff h1 0 _ h10) as h11. constructor. split; auto. simpl. rewrite h5r. omega. intro h6. inversion h6. apply Extensionality_Ensembles. red. split. red. intros x' h6. destruct h6 as [h6]. destruct h6 as [m h6]. inversion h6 as [h6l]. destruct h6l as [h6l h6r]. rewrite sub_add_comm in h6l. inversion h6l as [pr h7 | pr h8]. subst. destruct h7 as [h7l h7r]. constructor. exists m. constructor. split; auto. subst. inversion h8 as [h9]. constructor. exists n. constructor. split. rewrite <- h9. apply fin_fps_to_f_s_compat. assumption. simpl. omega. intro h7. inversion h7. pose proof (O_lt_pred_n_lt_n _ h5). unfold n in H. omega. unfold bag_fun_well_defined_subtract_set. rewrite bag_set_unq_set_sub_x_n. rewrite bag_set_unq_set_add_x_Sn. red. intros x' h6. constructor. destruct (classic_dec (x = x')) as [h7 | h8]. rewrite <- h7. exists (pred n). constructor. right. constructor. intro h8. inversion h8 as [h9]. pose proof (O_lt_pred_n_lt_n _ h5) as h10. unfold n in h9. unfold n in h10. omega. destruct h6 as [h6]. destruct h6 as [r h6]. exists r. constructor. left. constructor. destruct h6 as [h6]. assumption. intro h9. inversion h9. contradiction. unfold n in h5. assumption. destruct (pred0_dec _ h4) as [h5 | h6]. subst. unfold bag_fun_well_defined_subtract_set. rewrite h5. simpl. rewrite bag_set_unq_set_sub_x_O. rewrite bag_set_unq_set_add_x_O. rewrite subtract_nin. reflexivity. intro h6. pose proof (bag_set_unq_set_incl C) as h7. pose proof (dom_rel_incl _ _ h7) as h8. auto with sets. subst. unfold bag_fun_well_defined_subtract_set. rewrite h6. simpl. rewrite sub_add_comm. rewrite bag_set_unq_set_add_x_O. rewrite bag_set_unq_set_sub_x_n. rewrite dom_rel_subtract_fun_well_defined. simpl. reflexivity. apply bag_set_unq_set_fun_well_defined; auto. simpl. intro. intro h7. pose proof (bag_set_unq_set_incl C) as h8. pose proof (dom_rel_incl _ _ h8) as h9. auto with sets. intro h7. inversion h7. unfold n in h5. rewrite fin_fps_to_f_def in h5. simpl in h5. omega. assumption. Qed. Lemma fun_well_defined_bag_fun_well_defined_subtract_set : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pf:fun_well_defined C) (x:T), fun_well_defined (bag_fun_well_defined_subtract_set pfdt pff pf x). intros T C hdt hf h1 x. constructor. intros x' h3. rewrite dom_rel_bag_fun_well_defined_subtract in h3. destruct (zerop (pred (fin_fps_to_f hdt hf C h1 0 x))) as [h4 | h5]. pose proof (pred0_dec _ h4) as h5. clear h4. destruct h5 as [h5l | h5r]. assert (h6:~Inn (dom_rel (bag_set_unq_set C)) x). intro h6. destruct h6 as [h6]. destruct h6 as [n h6]. destruct h6 as [h6]. simpl in h6. destruct h6 as [h6l h6r]. pose proof (fp_in_dom _ _ _ h1 _ h6l) as h7. simpl in h7. pose proof (fin_fps_to_f_s_compat hdt hf h1 0 _ h7) as h8. pose proof (fp_functional h1 _ _ _ h6l h8). subst. omega. rewrite subtract_nin in h3. assert (h7:x <> x'). intro h7. subst. contradiction. destruct h3 as [h3]. destruct h3 as [n h3]. exists n. red. split. split. constructor. exists x'. constructor. simpl. split. rewrite h5l. simpl. rewrite add_subtract_a. constructor. pose proof (bag_set_unq_set_incl C). auto with sets. intro h9. inversion h9; subst. contradict h7. reflexivity. destruct h3 as [h3]. destruct h3; auto. constructor. rewrite h5l. simpl. rewrite add_subtract_a. split. constructor. destruct h3 as [h3]. simpl in h3. destruct h3; auto. intro h8. inversion h8. contradiction. destruct h3 as [h3]. simpl in h3. destruct h3; auto. intros m h8. destruct h8 as [h8l h8r]. unfold bag_fun_well_defined_subtract_set in h8r. inversion h8r as [h9]. destruct h9 as [h9l h9r]. rewrite h5l in h9l. simpl in h9l. rewrite add_subtract_a in h9l. destruct h9l as [h9l]. destruct h3 as [h3]. destruct h3 as [h3l h3r]. apply (fp_functional h1 _ _ _ h3l h9l). assumption. unfold bag_fun_well_defined_subtract_set. rewrite h5r. simpl. destruct h3 as [h3 h4]. destruct h3 as [h3]. destruct h3 as [n h3]. exists n. red. split. split. constructor. exists x'. constructor. split. rewrite sub_add_comm. left. constructor. pose proof (bag_set_unq_set_incl C). auto with sets. intro h6. inversion h6. subst. contradict h4. constructor. intro h6. inversion h6. destruct h3 as [h3]. destruct h3; auto. rewrite bag_set_unq_set_sub_x_n. rewrite bag_set_unq_set_add_x_O. constructor. assumption. intro h6. inversion h6. subst. contradict h4. constructor. intros m h6. destruct h6 as [h6l h6r]. inversion h6r as [h7]. destruct h7 as [h7l h7r]. inversion h7l as [h8]. inversion h8 as [pr h8l | pr h8r]. subst. pose proof (bag_set_unq_set_incl C) as h9. assert (h10:Inn C (x', n)). auto with sets. apply (fp_functional h1 _ _ _ h10 h8l). subst. inversion h8r. subst. simpl in h7r. omega. destruct (classic_dec (x = x')) as [h6 | h7]. subst. destruct (classic (Inn (dom_rel C) x')) as [h6 | h7]. exists (pred (fin_fps_to_f hdt hf C h1 0 x')). red. split. split. constructor. exists x'. constructor. simpl. split; auto with arith. rewrite sub_add_comm. right. constructor. intro h7. inversion h7. pose proof (O_lt_pred_n_lt_n _ h5). omega. constructor. simpl. split; auto with arith. rewrite sub_add_comm. right. constructor. intro h7. inversion h7. pose proof (O_lt_pred_n_lt_n _ h5). omega. intros n h7. destruct h7 as [h7l h7r]. destruct h7r as [h7r]. simpl in h7r. rewrite sub_add_comm in h7r. destruct h7r as [h7a h7b]. inversion h7a as [? h8 | ? h9]. subst. destruct h8 as [h8l h8r]. contradict h8r. pose proof (fin_fps_to_f_s_compat hdt hf h1 0 _ h6) as h9. pose proof (fp_functional h1 _ _ _ h8l h9). subst. constructor. subst. inversion h9. reflexivity. intro h8. inversion h8. pose proof (O_lt_pred_n_lt_n _ h5). omega. pose proof (fin_fps_to_f_def hdt hf h1 0 _ h7) as h8. rewrite h8 in h5. simpl in h5. omega. exists (fin_fps_to_f hdt hf C h1 0 x'). red. split. split. constructor. exists x'. constructor. simpl. split. rewrite sub_add_comm. left. constructor. apply fin_fps_to_f_s_compat. destruct h3 as [h3]. destruct h3 as [n h3]. inversion h3 as [h4]. destruct h4 as [h4l h4r]. constructor. exists n. assumption. intro h8. inversion h8. subst. contradict h7. reflexivity. intro h8. inversion h8. pose proof (O_lt_pred_n_lt_n _ h5). omega. destruct h3 as [h3]. destruct h3 as [n h3]. inversion h3 as [h4]. destruct h4 as [h4l h4r]. simpl in h4r. pose proof (fp_in_dom _ _ _ h1 _ h4l) as h8. simpl in h8. pose proof (fin_fps_to_f_s_compat hdt hf h1 0 _ h8) as h9. pose proof (fp_functional h1 _ _ _ h4l h9). subst. assumption. constructor. simpl. split. constructor. left. pose proof (bag_set_unq_set_incl C) as h8. pose proof (dom_rel_incl _ _ h8) as h9. assert (h10:Inn (dom_rel C) x'). auto with sets. apply (fin_fps_to_f_s_compat hdt hf h1 0 _ h10). intro h8. inversion h8. subst. contradict h7. reflexivity. destruct h3 as [h3]. destruct h3 as [n h3]. inversion h3 as [h4]. simpl in h4. destruct h4 as [h4l h4r]. pose proof (fp_in_dom _ _ _ h1 _ h4l) as h8. simpl in h8. pose proof (fin_fps_to_f_s_compat hdt hf h1 0 _ h8) as h9. pose proof (fp_functional h1 _ _ _ h4l h9). subst. assumption. intros n h8. destruct h8 as [h8l h8r]. inversion h8r as [h9]. simpl in h9. destruct h9 as [h9l h9r]. inversion h9l as [h10]. clear h9l. inversion h10 as [? h11 | ? h12]. subst. pose proof (bag_set_unq_set_incl C) as h12. pose proof (dom_rel_incl _ _ h12) as h13. assert (h14:Inn (dom_rel C) x'). auto with sets. pose proof (fin_fps_to_f_s_compat hdt hf h1 0 _ h14) as h15. apply (fp_functional h1 _ _ _ h15 h11). subst. inversion h12. subst. contradict h7. reflexivity. intros pr h2. split. constructor. exists (snd pr). rewrite <- surjective_pairing. assumption. constructor. exists (fst pr). rewrite <- surjective_pairing. assumption. Qed. Lemma nonzero_bag_fun_well_defined_subtract_set : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite (dom_rel C)) (pf:fun_well_defined C) (x:T), nonzero (bag_fun_well_defined_subtract_set pfdt pff pf x). intros T C hdt hf h1 x. red. intros x' h2. inversion h2 as [h3]. clear h2. simpl in h3. destruct h3; omega. Qed. Definition bag_subtract {T:Type} (B:Bag T) (x:T) := let pfdt := bag_type_eq_dec B in let pff := finite_dom_rel_bag_to_pair_set B in bag_intro _ pfdt _ (finite_bag_fun_well_defined_subtract_set pfdt pff (fun_well_defined_bag_to_pair_set B) x (finite_bag_to_pair_set B)) (fun_well_defined_bag_fun_well_defined_subtract_set pfdt pff (fun_well_defined_bag_to_pair_set B) x) (nonzero_bag_fun_well_defined_subtract_set pfdt pff (fun_well_defined_bag_to_pair_set B) x). (*handle classic in_dec.*) Lemma fun_well_defined_bag_fun_well_defined_subtract_set_fin_fps_to_f : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite C) (pf:fun_well_defined C) (x:T), let pffd := dom_rel_finite _ pff in let pff' := finite_bag_fun_well_defined_subtract_set pfdt (dom_rel_finite _ pff) pf x pff in let pffd' := dom_rel_finite _ pff' in fin_fps_to_f pfdt pffd' _ (fun_well_defined_bag_fun_well_defined_subtract_set pfdt (dom_rel_finite _ pff) pf x) 0 x = pred (fin_fps_to_f pfdt pffd _ pf 0 x). intros T C hdt hf h1 x hfd hf' hfd'. gterml. redterm2 x0. fold c. generalize c. intro h0. destruct (classic (Inn (dom_rel (bag_fun_well_defined_subtract_set hdt hfd h1 x)) x)) as [h2 | h3]. destruct h2 as [h2]. destruct h2 as [n h2]. inversion h2 as [h3]. clear h2. destruct h3 as [h3l h3r]. inversion h3l as [h4 h5]. clear h3l. assert (h6:~Inn C (x, n)). intro h6. pose proof (fp_in_dom _ _ _ h1 _ h6) as h7. simpl in h7. pose proof (fin_fps_to_f_s_compat hdt hfd h1 0 _ h7) as h8. pose proof (fp_functional h1 _ _ _ h6 h8). subst. contradict h5. constructor. destruct (classic_dec (Inn (dom_rel C) x)) as [h7 | h8]. inversion h4 as [? h9 | ? h10]. contradiction. subst. inversion h10 as [h11]. clear h10. destruct h7 as [h7]. destruct h7 as [m h7]. pose proof (fp_in_dom _ _ _ h1 _ h7) as h12. simpl in h12. pose proof (fin_fps_to_f_s_compat hdt hfd h1 0 _ h12) as h13. pose proof (fp_functional h1 _ _ _ h7 h13). subst. assert (h14:Inn (bag_fun_well_defined_subtract_set hdt hfd h1 x) (x, pred (fin_fps_to_f hdt hfd C h1 0 x))). constructor. split; auto. constructor. right. constructor. assumption. pose proof (fp_in_dom _ _ _ h0 _ h14) as h15. simpl in h15. pose proof (fin_fps_to_f_s_compat hdt hfd' h0 0 _ h15) as h16. apply (fp_functional h0 _ _ _ h16 h14). pose proof (fin_fps_to_f_def hdt hfd h1 0 _ h8) as h9. rewrite h9. assert (h10:~Inn (dom_rel (bag_fun_well_defined_subtract_set hdt hfd h1 x)) x). contradict h8. destruct h8 as [h8]. destruct h8 as [m h8]. inversion h8 as [h10]. simpl in h10. destruct h10 as [h10l h10r]. inversion h10l as [h11 h12]. clear h10l. inversion h11 as [? h13 | ? h14]. subst. contradict h12. pose proof (fp_in_dom _ _ _ h1 _ h13) as h14. simpl in h14. pose proof (fin_fps_to_f_s_compat hdt hfd h1 0 _ h14) as h15. pose proof (fp_functional h1 _ _ _ h13 h15) as h16. subst. constructor. subst. inversion h14; subst. rewrite h9 in h10r. simpl in h10r. omega. pose proof (fin_fps_to_f_def hdt hfd' h0 0 _ h10) as h11. rewrite h11. simpl. reflexivity. pose proof (fin_fps_to_f_def hdt hfd' h0 0 _ h3) as h4. rewrite h4. destruct (zerop (pred (fin_fps_to_f hdt hfd C h1 0 x))) as [h5 | h6]. rewrite h5. reflexivity. contradict h3. constructor. exists (pred (fin_fps_to_f hdt hfd C h1 0 x)). constructor. split; auto. constructor. right. constructor. intro h7. inversion h7. pose proof (O_lt_pred_n_lt_n _ h6). omega. Qed. Lemma fun_well_defined_bag_fun_well_defined_subtract_set_fin_fps_to_f_neq : forall {T:Type} {C:Ensemble (T*nat)} (pfdt:type_eq_dec T) (pff:Finite C) (pf:fun_well_defined C) (a x:T), a <> x -> let pffd := dom_rel_finite _ pff in let pff' := finite_bag_fun_well_defined_subtract_set pfdt (dom_rel_finite _ pff) pf a pff in let pffd' := dom_rel_finite _ pff' in fin_fps_to_f pfdt pffd' _ (fun_well_defined_bag_fun_well_defined_subtract_set pfdt pffd pf a) 0 x = fin_fps_to_f pfdt pffd _ pf 0 x. intros T C hdt hf h1 a x h2 hfd hff' hfd'. gen2. intro h3. destruct (classic (Inn (dom_rel (bag_fun_well_defined_subtract_set hdt hfd h1 a)) x)) as [h4 | h5]. assert (h5:Inn (dom_rel C) x). destruct h4 as [h4]. destruct h4 as [n h4]. destruct h4 as [h4]. simpl in h4. destruct h4 as [h4 h5]. inversion h4 as [h6 h7]. clear h4. inversion h6 as [? h6l | ?r h6r]. subst. clear h6. constructor. exists n. assumption. inversion h6r. contradiction. pose proof (fin_fps_to_f_s_compat hdt hfd h1 0 _ h5) as h6. pose proof (fin_fps_to_f_s_compat hdt hfd' h3 0 _ h4) as h7. inversion h7 as [h8]. clear h7. simpl in h8. destruct h8 as [h8l h8r]. inversion h8l as [h9 h10]; clear h8l. inversion h9 as [? h11 | ? h12]; subst; clear h9. apply (fp_functional h1 _ _ _ h11 h6). inversion h12. contradiction. rewrite fin_fps_to_f_def; auto. destruct (zerop (fin_fps_to_f hdt hfd C h1 0 x)) as [h6 | h7]. rewrite h6. reflexivity. assert (h8:Inn (dom_rel C) x). apply NNPP. intro h8. pose proof (fin_fps_to_f_def hdt hfd h1 0 _ h8). omega. contradict h5. constructor. exists (fin_fps_to_f hdt hfd C h1 0 x). constructor. simpl. split; auto with arith. constructor. left. apply fin_fps_to_f_s_compat. assumption. intro h9. inversion h9. contradiction. Qed. Lemma bag_subtract_app_eq : forall {T:Type} (B:Bag T) (x:T), (bag_subtract B x) {->} x = pred (B {->} x). intros T B x. do 2 rewrite bag_to_fin_map_compat'. pose proof (fun_well_defined_bag_fun_well_defined_subtract_set_fin_fps_to_f (bag_type_eq_dec B) (finite_bag_to_pair_set B) (fun_well_defined_bag_to_pair_set B) x) as h1. simpl. simpl in h1. symmetry. erewrite fin_fps_to_f_functional. rewrite <- h1 at 1. apply fin_fps_to_f_functional. f_equal. apply proof_irrelevance. reflexivity. Qed. Lemma bag_subtract_ninb1 : forall {T:Type} (B:Bag T) (x:T), ~Inb1 B x -> bag_subtract B x = B. intros T B x h1. destruct B as [hdt C h2 h3 h4]. unfold bag_subtract. apply bag_intro_functional. pose proof h1 as h1'. rewrite inb1_dom_rel_compat_iff in h1. unfold bag_fun_well_defined_subtract_set. match goal with |- bag_set_unq_set (Subtract _ (_, ?x)) = _ => pose x as c end. assert (h5:c = 0). destruct (zerop c) as [h6 | h7]; auto. assert (h8:Inb1 (bag_intro T hdt C h2 h3 h4) x). red. rewrite bag_to_fin_map_compat'. auto with arith. contradiction. subst. rewrite h5. simpl. red in h4. pose proof (h4 x) as h4'. rewrite (sub_add_same_nin _ _ h4') at 1. apply Extensionality_Ensembles. red; split; try apply bag_set_unq_set_incl; auto. red. intros pr h6. constructor. split; auto. pose proof (h4 (fst pr)) as h4''. destruct (zerop (snd pr)) as [h7 | h8]. rewrite <- h7 in h4''. rewrite <- surjective_pairing in h4''. contradiction. auto with arith. Qed. Lemma bag_fun_well_defined_subtract_set_eq : forall {T:Type} (pfdt:type_eq_dec T) (B:Bag T) (a:T) (pffp:fun_well_defined (bag_to_pair_set B)), bag_fun_well_defined_subtract_set pfdt (dom_rel_finite (bag_to_pair_set B) (finite_bag_to_pair_set B)) pffp a = bag_to_pair_set (bag_subtract B a). intros T hdt B a h1. unfold bag_fun_well_defined_subtract_set. rewrite bag_set_unq_set_sub_x_n. destruct (inb1_dec B a) as [h2 | h2]. unfold bag_subtract. unfold bag_to_pair_set. unfold bag_fun_well_defined_subtract_set. unfold bag_set_unq_set. destruct B as [hdt' S h3 h4 h5]. apply Extensionality_Ensembles; red; split; red; intros c h6. destruct h6 as [h6 h7]. destruct h6 as [h6]. destruct h6 as [h6 h8]. constructor. split. constructor. pose proof (type_eq_dec_eq hdt (bag_type_eq_dec (bag_intro T hdt' S h3 h4 h5))) as h9. subst. erewrite fin_fps_to_f_functional. apply h6. simpl. reflexivity. rewrite in_sing_iff. rewrite in_sing_iff in h7. intro h9. rewrite <- h9 in h7. contradict h7. f_equal. apply fin_fps_to_f_functional; auto. assumption. destruct h6 as [h6]. destruct h6 as [h6 h7]. destruct h6 as [h6 h8]. rewrite in_sing_iff in h8. constructor. constructor. split; auto. erewrite fin_fps_to_f_functional. apply h6. simpl. reflexivity. rewrite in_sing_iff. erewrite fin_fps_to_f_functional. apply h8. simpl. reflexivity. rewrite bag_subtract_ninb1; auto. pose proof (not_inb1 _ _ h2) as h3. rewrite bag_to_fin_map_compat' in h3. erewrite fin_fps_to_f_functional; auto. rewrite h3 at 1. rewrite h3. simpl. rewrite bag_set_unq_set_add_x_O. pose proof (bag_pair_subtract_O B a) as h4. unfold bag_pair_subtract in h4. pose proof (f_equal (@bag_to_pair_set _) h4) as h5. simpl in h5. rewrite <- h5. f_equal. rewrite bag_set_unq_set_sub_x_n. rewrite bag_set_unq_set_bag_to_pair_set. assumption. Qed. Lemma bag_subtract_app_neq : forall {T:Type} (B:Bag T) (a x:T), a <> x -> (bag_subtract B a) {->} x = B {->} x. intros T B a x h1. do 2 rewrite bag_to_fin_map_compat'. pose proof (bag_type_eq_dec B) as hdt. pose proof (fun_well_defined_bag_to_pair_set B) as hfp. pose proof (fun_well_defined_bag_fun_well_defined_subtract_set_fin_fps_to_f_neq hdt (finite_bag_to_pair_set B) hfp _ _ h1) as h3. simpl in h3. symmetry. erewrite fin_fps_to_f_functional. rewrite <- h3 at 1. apply fin_fps_to_f_functional. apply bag_fun_well_defined_subtract_set_eq. reflexivity. Qed. Lemma ninb1_bag_pair_subtract_app : forall {T:Type} (B:Bag T) (x:T), ~Inb1 (bag_pair_subtract B (x, (B {->} x))) x. intros T B x h1. apply inb1_inb2_compat in h1. red in h1. simpl in h1. destruct h1 as [? h1]. assert (h2:Inb1 (bag_pair_subtract B (x, B{->}x)) x). red. assumption. pose proof (inb1_dom_rel_compat _ _ h2) as h3. simpl in h3. destruct h3 as [h3]. destruct h3 as [n h3]. assert (h4:Inn (bag_to_pair_set B) (x, n)). pose proof (incl_subtract (bag_to_pair_set B) (x, B{->}x)) as h4. auto with sets. pose proof (bag_in_compat _ _ h4) as h5. simpl in h5. subst. inversion h3 as [h5 h6]. contradict h6. constructor. Qed. Lemma bag_subtract_bag_pair_subtract_eq : forall {T:Type} (B:Bag T) (x:T), bag_subtract (bag_pair_subtract B (x, B{->}x)) x = bag_pair_subtract B (x, B{->}x). intros T B x. apply bag_subtract_ninb1. apply ninb1_bag_pair_subtract_app. Qed. Lemma bag_pair_subtract_app_eq_O : forall {T:Type} (B:Bag T) (x:T), bag_pair_subtract B (x, B{->}x) {->} x = 0. intros T B x. destruct (inb1_dec B x) as [h0 | h1]. destruct B as [hdt S h2 h3 h4]. unfold bag_pair_subtract. rewrite bag_to_fin_map_compat'. simpl. assert (h5:~Inn (dom_rel (Subtract S (x, bag_intro T hdt S h2 h3 h4 {->}x))) x). intro h6. rewrite dom_rel_subtract_fun_well_defined in h6; auto. simpl in h6. destruct h6 as [h6l h6r]. contradiction h6r. constructor. intro h7. pose proof (inb1_in_compat _ _ h0) as h8. simpl in h8. contradiction. apply fin_fps_to_f_def. assumption. assert (h2:~Inb1 (bag_pair_subtract B (x, B{->}x)) x). intro h2. pose proof (not_inb1 _ _ h1) as h3. rewrite h3 in h2. rewrite bag_pair_subtract_O in h2. contradiction. apply fin_map_app_def. intro h3. apply in_dom_rel_inb1_compat in h3. contradiction. Qed. Lemma bag_pair_subtract_app_neq : forall {T:Type} (B:Bag T) (a x:T) (n:nat), a <> x -> bag_pair_subtract B (a, n) {->} x = B {->} x. intros T B a x n h1. destruct (inb1_dec B x) as [h2 | h3]. pose proof (inb1_in_compat _ _ h2) as h3. pose proof (bag_in_compat _ _ h3) as h4. assert (h5:Inn (bag_to_pair_set (bag_pair_subtract B (a, n))) (x, B {->} x)). simpl. constructor. assumption. intro h5. inversion h5. contradiction. assert (h6:Inn (dom_rel (bag_to_pair_set (bag_pair_subtract B (a, n)))) x). constructor. exists (B{->}x). assumption. pose proof (bag_type_eq_dec B) as hdt. pose proof (finite_dom_rel_bag_to_pair_set (bag_pair_subtract B (a, n))) as hf. pose proof (fin_fps_to_f_s_compat hdt hf (fun_well_defined_bag_to_pair_set (bag_pair_subtract B (a, n))) 0 _ h6) as h8. pose proof (bag_to_fin_map_compat' (bag_pair_subtract B (a, n)) x) as h9. simpl in h8, h9. erewrite fin_fps_to_f_functional in h9. rewrite <- h9 in h8 at 1. pose proof (incl_subtract (bag_to_pair_set B) (a, n)) as h10. assert (h11:Inn (bag_to_pair_set B) (x, bag_pair_subtract B (a, n){->} x)). auto with sets. apply (fp_functional (fun_well_defined_bag_to_pair_set B) _ _ _ h11 h3). rewrite bag_to_pair_set_bag_pair_subtract; auto. pose proof (not_inb1 _ _ h3) as h4. rewrite h4. assert (h5:~Inb1 (bag_pair_subtract B (a, n)) x). intro h5. pose proof (inb1_in_compat _ _ h5) as h6. simpl in h6. inversion h6 as [h7 h8]. pose proof (bag_in_compat _ _ h7) as h9. simpl in h9. rewrite h4 in h9. rewrite h9 in h7. pose proof (nonzero_bag_to_pair_set B x) as h10. contradiction. apply not_inb1. assumption. Qed. Lemma bag_pair_add_undoes_bag_pair_subtract : forall {T:Type} (B:Bag T) (pr:T*nat), bag_pair_add (bag_pair_subtract B pr) (fst pr, B{->}fst pr) = B. intros T B pr. apply bag_to_fin_map_inj. apply functional_extensionality. intro x. destruct (eq_dec x (fst pr)) as [h1 | h2]. subst. pose proof (bag_pair_add_app_eq (bag_pair_subtract B pr) (fst pr, B{->}fst pr)) as h1. simpl in h1. simpl. rewrite h1. reflexivity. rewrite bag_pair_add_app_neq. apply neq_sym in h2. pose proof (bag_pair_subtract_app_neq B (fst pr) x (snd pr) h2) as h3. rewrite (surjective_pairing pr). assumption. simpl. assumption. Qed. Lemma bag_subtract_undoes_bag_add : forall {T:Type} (B:Bag T) (x:T), bag_subtract (bag_add B x) x = B. intros T B x. apply bag_to_pair_set_inj. simpl. apply Extensionality_Ensembles. red. split. red. intros pr h1. destruct h1 as [h1]. destruct h1 as [h1l h1r]. destruct h1l as [h1l h2]. destruct h1l as [pr h3 | pr h4]. destruct h3 as [h3 h5]. destruct h3 as [pr h3l | pr h3r]. assumption. destruct h3r; subst. contradict h2. rewrite <- bag_to_fin_map_compat'. rewrite <- bag_add_app_eq. rewrite bag_to_fin_map_compat'. simpl. constructor. destruct h4; subst. apply inb2_in_compat. red. simpl. simpl in h1r. split; auto. rewrite bag_to_fin_map_compat'. apply S_inj. let P := type of h1r in match P with pred ?c > 0 => assert (h3:c > 0) end. pose proof (O_lt_pred_n _ h1r). auto with arith. simpl in h3. rewrite <- (S_pred _ _ h3). rewrite <- bag_to_fin_map_compat'. rewrite <- bag_add_app_eq. rewrite bag_to_fin_map_compat'. reflexivity. red. intros pr h1. destruct (eq_dec (fst pr) x) as [h2 | h3]. subst. constructor. split. constructor. right. pose proof (bag_to_fin_map_compat' (bag_add B (fst pr)) (fst pr)) as h2. simpl in h2. rewrite <- h2. pose proof (bag_add_app_eq B (fst pr)) as h3. simpl. simpl in h3. rewrite h3. simpl. pose proof (bag_in_compat _ _ h1) as h4. rewrite <- h4. rewrite surjective_pairing. constructor. intro h2. rewrite surjective_pairing in h2. inversion h2 as [h3]. clear h2. pose proof (bag_to_fin_map_compat' (bag_add B (fst pr)) (fst pr)) as h4. simpl in h4. rewrite <- h4 in h3. pose proof (bag_in_compat _ _ h1) as h5. simpl in h3. pose proof (bag_add_app_eq B (fst pr)) as h6. simpl in h6. rewrite h6 in h3. rewrite h5 in h3. omega. pose proof (in_inb2_compat _ _ h1) as h2. red in h2. destruct h2; assumption. constructor. split. constructor. left. constructor. left. assumption. intro h4. rewrite surjective_pairing in h4. inversion h4 as [h5]. symmetry in h5. contradiction. intro h4. rewrite surjective_pairing in h4. inversion h4 as [h5]. symmetry in h5. contradiction. pose proof (in_inb2_compat _ _ h1) as h2. destruct h2 as [h2l h2r]. assumption. Qed. Lemma bag_add_undoes_bag_subtract_inb1 : forall {T:Type} (B:Bag T) (x:T), Inb1 B x -> bag_add (bag_subtract B x) x = B. intros T B a h1. apply bag_to_fin_map_inj. apply functional_extensionality. intro x. destruct ((bag_type_eq_dec B) a x) as [h2 | h3]. subst. rewrite (bag_add_app_eq (bag_subtract B x) x). rewrite (bag_subtract_app_eq B x). red in h1. rewrite <- (S_pred _ _ h1). reflexivity. rewrite (bag_add_app_neq (bag_subtract B a)). rewrite (bag_subtract_app_neq B a). reflexivity. assumption. assumption. Qed. Lemma bag_add_adds_bag_subtract_ninb1 : forall {T:Type} (B:Bag T) (x:T), ~Inb1 B x -> bag_add (bag_subtract B x) x = bag_add B x. intros T B a h1. apply bag_to_fin_map_inj. apply functional_extensionality. intro x. rewrite (bag_subtract_ninb1). reflexivity. assumption. Qed. Lemma bag_subtract_inb1_inj : forall {T:Type} (B B':Bag T) (x:T), Inb1 B x -> Inb1 B' x -> bag_subtract B x = bag_subtract B' x -> B = B'. intros T B B' x hin hnin h1. pose proof (f_equal (fun B=>bag_add B x) h1) as h3. simpl in h3. do 2 rewrite bag_add_undoes_bag_subtract_inb1 in h3; auto. Qed. Lemma bag_subtract_inb1_inj' : forall {T:Type} (B B':Bag T) (x:T), (Inb1 B x <-> Inb1 B' x) -> bag_subtract B x = bag_subtract B' x -> B = B'. intros T B B' x h0 h1. pose proof (f_equal (fun B=>bag_add B x) h1) as h3. simpl in h3. destruct (inb1_dec B x) as [hin | hnin]; destruct (inb1_dec B' x) as [hin' | hnin']. do 2 rewrite bag_add_undoes_bag_subtract_inb1 in h3; auto. rewrite h0 in hin. contradiction. rewrite h0 in hnin. contradiction. do 2 rewrite bag_add_adds_bag_subtract_ninb1 in h3; auto. do 2 rewrite bag_subtract_ninb1 in h1; auto. Qed. Lemma bag_add_inj : forall {T:Type} (B B':Bag T) (x:T), bag_add B x = bag_add B' x -> B = B'. intros T B B' x h1. pose proof (f_equal (fun B=>bag_subtract B x) h1) as h2. simpl in h2. do 2 rewrite bag_subtract_undoes_bag_add in h2. assumption. Qed. Lemma bag_subtract_inclb : forall {T:Type} (B:Bag T) (x:T), Inclb (bag_subtract B x) B. intros T B a. red. intros x h1. red. simpl. split. red in h1. destruct (eq_dec a x) as [h2 | h3]. subst. pose proof (bag_subtract_app_eq B x) as h2. simpl in h2. rewrite h2. simpl in h1. rewrite h2 in h1. omega. pose proof (bag_subtract_app_neq B _ _ h3) as h4. simpl in h4. rewrite h4. omega. red in h1. auto with arith. Qed. (*left off here.*) Lemma pred_bag_ex : forall {T:Type} (B:Bag T), B <> empty_bag T (bag_type_eq_dec B) -> exists (B':Bag T) (x:T), B = bag_add B' x /\ Inclb B' B /\ B' <> B. intros T B h1. pose proof (non_empty_bag_inhabited _ h1) as h2. destruct h2 as [x h2]. exists (bag_subtract B x). exists x. rewrite bag_add_undoes_bag_subtract_inb1; auto. split; auto. split. red. intros w h3. destruct (eq_dec w x) as [h4 | h5]. subst. pose proof (bag_subtract_app_eq B x) as h4. rewrite h4. red. simpl; split. omega. red in h3. rewrite h4 in h3. auto with arith. apply neq_sym in h5. pose proof (bag_subtract_app_neq B _ _ h5) as h6. simpl in h6. simpl. rewrite h6. red. simpl. split; auto. red in h3. simpl in h6. simpl in h3. rewrite h6 in h3. auto with arith. intro h3. pose proof (f_equal (fun B => B{->}x) h3) as h4. simpl in h4. pose proof (bag_subtract_app_eq B x) as h5. simpl in h5. simpl in h4. rewrite h5 in h4. red in h2. rewrite (S_pred _ _ h2) in h4. simpl in h4. omega. Qed. Fixpoint bag_subtract_n {T:Type} (B:Bag T) (x:T) (n:nat) := match n with | O => B | S n => bag_subtract (bag_subtract_n B x n) x end. Lemma bag_subtract_n_app_le : forall {T:Type} (B:Bag T) (x:T) (n:nat), n > 0 -> bag_subtract_n B x n {->} x<= bag_subtract B x {->} x. intros T B x n. revert T B x. induction n as [|n h1]; simpl. intros; omega. intros T B x h2. destruct (zerop n) as [h3 | h4]. subst. simpl. auto with arith. specialize (h1 _ B x h4). pose proof (bag_subtract_app_eq (bag_subtract_n B x n) x) as h5. simpl in h5. rewrite h5. apply le_S_n. destruct (zerop (pred (bag_subtract_n B x n{->} x))) as [h6 | h7]. rewrite h6. omega. assert (h8:0 < bag_subtract_n B x n{->}x). omega. rewrite <- (S_pred _ _ h8). auto with arith. Qed. Lemma bag_subtract_n_ninb1 : forall {T:Type} (B:Bag T) (x:T) (n:nat), ~Inb1 B x -> bag_subtract_n B x n = B. intros T B x n h1. revert B x h1. induction n as [|n h1]. simpl. auto. intros B x h2. simpl. rewrite h1; auto. apply bag_subtract_ninb1. assumption. Qed. Lemma bag_subtract_n_app_1 : forall {T:Type} (B:Bag T) (x:T) (n:nat), n > 0 -> B{->}x = 1 -> bag_subtract_n B x n = bag_subtract B x. intros T B x n h0 h1. revert B x h0 h1. induction n as [|n h1]. intros; omega. intros B x h0 h2. simpl. apply bag_to_fin_map_inj. apply functional_extensionality. intro y. destruct (zerop n) as [h3 | h4]. subst. simpl. reflexivity. specialize (h1 B x h4 h2). rewrite h1. destruct (classic (x = y)) as [h5 | h6]. subst. rewrite bag_subtract_app_eq. rewrite bag_subtract_app_eq. rewrite h2. simpl. reflexivity. rewrite bag_subtract_app_neq; auto. Qed. Lemma bag_subtract_n_app_neq : forall {T:Type} (B:Bag T) (a x:T) (n:nat), a <> x -> (bag_subtract_n B a n) {->} x = B {->} x. intros T B a x n. revert T B a x. induction n as [|n h1]. simpl. auto. intros T B a x h2. specialize (h1 T B a x h2). simpl. pose proof (bag_subtract_app_neq (bag_subtract_n B a n) _ _ h2) as h3. simpl in h3. rewrite h3. assumption. Qed. Lemma bag_subtract_n_eq_mono : forall {T:Type} (B:Bag T) (x:T) (n:nat), bag_subtract_n (bag_subtract B x) x n{->}x <= bag_subtract_n B x n{->}x. intros T B x n. revert T B x. induction n as [|n h1]. simpl. intros. pose proof (bag_subtract_app_eq B x) as h1. simpl in h1. rewrite h1. omega. intros T B x. simpl. pose proof (bag_subtract_app_eq (bag_subtract_n (bag_subtract B x) x n) x) as h2. simpl in h2. rewrite h2. pose proof (bag_subtract_app_eq (bag_subtract_n B x n) x) as h3. simpl in h3. rewrite h3. apply le_pred. apply h1. Qed. Lemma bag_subtract_n_eq_pred : forall {T:Type} (B:Bag T) (x:T) (n:nat), pred (bag_subtract_n B x n{->}x) = bag_subtract_n (bag_subtract B x) x n {->} x. intros T B x n. revert T B x. induction n as [|n h1]; simpl. intros; rewrite <- bag_subtract_app_eq. reflexivity. intros T B x. pose proof (bag_subtract_app_eq (bag_subtract_n B x n) x) as h2. simpl in h2. rewrite h2. pose proof (bag_subtract_app_eq (bag_subtract_n (bag_subtract B x) x n) x) as h3. simpl in h3. rewrite h3. f_equal. apply h1. Qed. Lemma bag_subtract_n_lt_constraint : forall {T:Type} (B:Bag T) (x:T) (n:nat), bag_subtract_n B x n{->}x > 0 -> n < B {->}x. intros T B x n. revert T B x. induction n as [|n h1]; simpl; auto with arith. intros T B x h2. pose proof (bag_subtract_app_eq (bag_subtract_n B x n) x) as h3. simpl in h3. rewrite h3 in h2. assert (h4:bag_subtract_n (bag_subtract B x) x n {->}x > 0). rewrite <- bag_subtract_n_eq_pred. assumption. specialize (h1 _ _ _ h4). pose proof (lt_n_S _ _ h1) as h5. rewrite bag_subtract_app_eq in h5. destruct (zerop (B{->}x)) as [h6 | h7]. rewrite h6 in h5. omega. rewrite <- (S_pred _ _ h7) in h5. assumption. Qed. Lemma bag_pair_subtract_ninb2_le_impl : forall {T:Type} (B:Bag T) (x:T) (n:nat), ~Inb2_le B (x, n) -> n = 0 \/ ~Inb1 B x \/ bag_subtract_n B x n = bag_pair_subtract B (x, B {->} x). intros T B x n. revert B x. induction n as [|n h1]. intros B x h2. left. reflexivity. intros B x h2. right. specialize (h1 (bag_subtract B x) x). assert (h3:~Inb2_le (bag_subtract B x) (x, n)). intro h4. red in h4. simpl in h4. destruct h4 as [h4l h4r]. unfold Inb2_le in h2. simpl in h2. simpl in h4l. pose proof (bag_subtract_app_eq B x) as h5. simpl in h5. rewrite h5 in h4l. omega. specialize (h1 h3). destruct h1 as [h1a | [h1b | h1c]]. subst. left. intro h4. contradict h2. red. simpl. split; auto with arith. destruct (inb1_dec B x) as [h4 | h5]. assert (h6:B{->}x = 1). unfold Inb1 in h1b. rewrite bag_subtract_app_eq in h1b. assert (h5:0 <= pred (B{->}x)). auto with arith. red in h4. omega. right. rewrite h6. pose proof (bag_subtract_bag_pair_subtract_eq B x) as h10. rewrite h6 in h10. rewrite <- h10. simpl. f_equal. apply bag_to_fin_map_inj. apply functional_extensionality. intro w. destruct (eq_dec w x) as [h8 | h9]. subst. destruct (zerop n) as [h8 | h9]. subst. simpl. contradict h2. red. simpl. split; auto with arith. rewrite bag_subtract_n_app_1; auto. rewrite bag_subtract_app_eq. rewrite h6. simpl. rewrite <- h6. symmetry. apply bag_pair_subtract_app_eq_O. rewrite bag_pair_subtract_app_neq; auto. rewrite bag_subtract_n_app_neq; auto. left; auto. right. simpl. simpl in h1c. simpl. destruct (zerop n) as [h4 | h5]. subst. simpl. assert (h4:~Inb1 B x). intro h4. contradict h2. red. simpl. split; auto. rewrite bag_subtract_ninb1. pose proof (not_inb1 _ _ h4) as h5. rewrite h5. rewrite bag_pair_subtract_O. reflexivity. assumption. apply bag_to_fin_map_inj. apply functional_extensionality. intro a. destruct (eq_dec a x) as [h6 | h7]. subst. rewrite bag_pair_subtract_app_eq_O. rewrite bag_subtract_app_eq. pose proof (bag_subtract_app_eq B x) as h6. simpl in h6. rewrite h6 in h1c. pose proof (bag_pair_subtract_app_eq_O (bag_subtract B x) x) as h7. rewrite bag_subtract_app_eq in h7. rewrite <- h1c in h7. rewrite <- bag_subtract_n_eq_pred in h7. destruct (pred0_dec _ h7) as [h8 | h9]. assumption. contradict h3. red. simpl. split; auto. assert (h10:bag_subtract_n B x n{->}x > 0). omega. pose proof (bag_subtract_n_lt_constraint _ _ _ h10). pose proof (bag_subtract_app_eq B x) as h11. simpl in h11. rewrite h11. omega. rewrite bag_subtract_app_neq. rewrite bag_pair_subtract_app_neq. rewrite bag_subtract_n_app_neq. reflexivity. apply neq_sym in h7. assumption. apply neq_sym in h7. assumption. apply neq_sym in h7. assumption. Qed. Lemma bag_subtract_n_app_ninb1 : forall {T:Type} (B:Bag T) (x:T), ~Inb1 (bag_subtract_n B x (B{->}x)) x. intros T B x h1. red in h1. pose proof (bag_subtract_n_lt_constraint _ _ _ h1) as h2. omega. Qed. Lemma bag_subtract_n_bag_pair_subtract_compat : forall {T:Type} (B:Bag T) (x:T), bag_subtract_n B x (B{->}x) = bag_pair_subtract B (x, B{->}x). intros T B a. apply bag_to_fin_map_inj. apply functional_extensionality. intro x. destruct (eq_dec a x) as [h1 | h2]. subst. pose proof (bag_subtract_n_app_ninb1 B x) as h1. pose proof (not_inb1 _ _ h1) as h2. rewrite h2. rewrite bag_pair_subtract_app_eq_O. reflexivity. rewrite bag_subtract_n_app_neq. rewrite bag_pair_subtract_app_neq. reflexivity. assumption. assumption. Qed. End BagSubtract. Section ImSetBag. Lemma im_set_bag_ex : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pf:Finite A) (f:T->U), exists! (B:Bag U), (Im A f = bag_to_set B) /\ (forall pr:U*nat, Inb2 B pr -> (cardinal _ [x:T | Inn A x /\ f x = fst pr] (snd pr))). intros T U h0 A h1 f. induction h1 as [|A h2 h3 x h4]. exists (empty_bag U h0). red. split. split. rewrite empty_bag_empty_set_compat. rewrite image_empty. reflexivity. intros pr h2. contradict h2. intro h3. pose proof (inb2_impl_inb1 _ _ h3) as h4. contradict h4. apply not_inb1_empty_bag. intros B h1. destruct h1 as [h1l h1r]. rewrite image_empty in h1l. symmetry. apply empty_set_empty_bag_compat; auto. destruct h3 as [B h3]. red in h3. destruct h3 as [h3l h3r]. destruct h3l as [h3a h3b]. exists (bag_add B (f x)). red. split. split. simpl. rewrite bag_to_set_add. rewrite Im_add. f_equal. assumption. intros pr h5. pose proof (bag_add_inv h0 _ _ _ h5) as h6. destruct h6 as [h6l h6r]. destruct h6l as [h7 | h8]. assert (hneq:fst pr <> f x). tauto. specialize (h3b _ h7). assert (h8: [x0 : T | Inn (Add A x) x0 /\ f x0 = fst pr] = [x0 : T | Inn A x0 /\ f x0 = fst pr]). apply Extensionality_Ensembles. red. split. red. intros t h8. destruct h8 as [h8]. destruct h8 as [h8l h8r]. destruct h8l as [t h9 | t h10]. constructor. split; auto. destruct h10. subst. contradict h7. auto. red. intros t h8. destruct h8 as [h8]. destruct h8 as [h8l h8r]. constructor. split; auto. left. assumption. rewrite h8. assumption. rewrite h8. assert (h9:~ Inb2 B pr). tauto. unfold Inb2 in h9. apply not_and_or in h9. destruct h9 as [h9l | h9r]. rewrite h8 in h9l. rewrite surjective_pairing in h5. rewrite h8 in h5. red in h5. simpl in h5. destruct h5 as [h5l h5r]. pose proof (bag_add_app_eq B (f x)) as h10. simpl in h10. rewrite h10 in h5l. pose proof (f_equal pred h5l) as h11. simpl in h11. destruct (zerop (pred (snd pr))) as [h12 |h13]. rewrite h12 in h11. assert (h13:snd pr = 1). omega. rewrite h13. assert (h14: [x0 : T | Inn (Add A x) x0 /\ f x0 = f x] = Singleton x). apply Extensionality_Ensembles. red. split. red. intros t h14. destruct h14 as [h14]. destruct h14 as [h14l h14r]. destruct h14l as [t h14a | t h14b]. assert (h15:Inn (Im A f) (f t)). apply Im_intro with t. assumption. reflexivity. rewrite h3a in h15. unfold bag_to_set in h15. destruct h15 as [pr' h15l]. subst. pose proof (bag_in_compat _ _ h15l) as h16. pose proof (nonzero_bag_to_pair_set B) as h17. red in h17. assert (h18:snd pr' <> 0). intro h18. specialize (h17 (fst pr')). rewrite surjective_pairing in h15l. rewrite h18 in h15l. contradiction. rewrite h16 in h18. rewrite <- H in h18. contradiction. assumption. red. intros x' h15. destruct h15. constructor. split; auto with sets. rewrite h14. apply card_sing. rewrite <- h11 in h13. assert (h14:Inb2 B ((f x), B{->}f x)). red. simpl. split; auto with arith. specialize (h3b _ h14). simpl in h3b. rewrite h11 in h3b. pose proof h9l as h9l'. clear h9l. assert (h9: [x0 : T | Inn (Add A x) x0 /\ f x0 = fst pr] = Add [x0 : T | Inn A x0 /\ f x0 = fst pr] x). apply Extensionality_Ensembles. red. split. red. intros t h9. destruct h9 as [h9]. destruct h9 as [h9l h9r]. destruct h9l as [t h9a |t h9b]. constructor. constructor. split; auto. destruct h9b. right. constructor. red. intros t h9. destruct h9 as [t h9l | t h9r]. destruct h9l as [h9l]. destruct h9l as [h9a h9b]. constructor. split. left. assumption. assumption. destruct h9r. constructor. split; auto. right. constructor. rewrite h8 in h9. rewrite h9. rewrite (S_pred _ _ h5r). constructor. assumption. intro h15. destruct h15 as [h15]. destruct h15; contradiction. assert (h10r:snd pr = 0). omega. rewrite surjective_pairing in h5. rewrite h10r in h5. red in h5. simpl in h5. destruct h5; omega. intros B' h5. destruct h5 as [h5l h5r]. rewrite Im_add in h5l. apply bag_to_fin_map_inj. apply functional_extensionality. intro c. destruct (eq_dec (f x) c) as [h6 | h7]. subst. pose proof (Add_intro2 _ (Im A f) (f x)) as h6. rewrite h5l in h6. rewrite in_bag_to_set_iff in h6. destruct (inb1_dec B (f x)) as [h7 | h8]. rewrite bag_add_app_eq. pose proof (inb1_inb2_compat _ _ h6) as h8. pose proof (inb1_inb2_compat _ _ h7) as h9. specialize (h3b _ h9). simpl in h3b. specialize (h5r _ h8). simpl in h5r. clear h9. assert (h9: [x0 : T | Inn (Add A x) x0 /\ f x0 = f x] = Add [x0 : T | Inn A x0 /\ f x0 = f x] x). apply Extensionality_Ensembles. red. split. red. intros t h9. destruct h9 as [h9]. destruct h9 as [h9l h9r]. destruct h9l as [t h9a |t h9b]. constructor. constructor. split; auto. destruct h9b. right. constructor. red. intros t h9. destruct h9 as [t h9l | t h9r]. destruct h9l as [h9l]. destruct h9l as [h9a h9b]. constructor. split. left. assumption. assumption. destruct h9r. constructor. split; auto. right. constructor. rewrite h9 in h5r. assert (h10:~Inn [x0 : T | Inn A x0 /\ f x0 = f x] x). intro h11. destruct h11 as [h11]. destruct h11; contradiction. pose proof (card_add_nin _ _ h3b _ h10) as h11. eapply cardinal_is_functional. apply h11. apply h5r. reflexivity. pose proof (not_inb1 _ _ h8) as h9. rewrite bag_add_app_eq. rewrite h9. assert (h14: [x0 : T | Inn (Add A x) x0 /\ f x0 = f x] = Singleton x). apply Extensionality_Ensembles. red. split. red. intros t h10. destruct h10 as [h10]. destruct h10 as [h10l h10r]. destruct h10l as [t h10a |t h10b]. assert (h11:Inn (Im A f) (f t)). apply Im_intro with t; auto. rewrite h10r in h11. rewrite h3a in h11. contradict h8. rewrite in_bag_to_set_iff in h11. assumption. assumption. red. intros t h10. destruct h10. constructor. split; try right; auto. constructor. pose proof (inb1_inb2_compat _ _ h6) as h15. specialize (h5r _ h15). simpl in h5r. rewrite h14 in h5r. pose proof (card_sing x) as h16. eapply cardinal_is_functional. apply h16. apply h5r. reflexivity. rewrite bag_add_app_neq; auto. destruct (inb1_dec B c) as [h8 | h9]; destruct (inb1_dec B' c) as [h10 | h11]. pose proof (inb1_inb2_compat _ _ h8) as h11. pose proof (inb1_inb2_compat _ _ h10) as h12. specialize (h3b _ h11). simpl in h3b. specialize (h5r _ h12). simpl in h5r. assert (h13:[x0 : T | Inn A x0 /\ f x0 = c] = [x0 : T | Inn (Add A x) x0 /\ f x0 = c]). apply Extensionality_Ensembles. red. split. red. intros t h13. destruct h13 as [h13]. destruct h13 as [h13l h13r]. constructor. split; try left; auto. red. intros t h13. destruct h13 as [h13]. destruct h13 as [h13l h13r]. destruct h13l as [t h13a | t h13b]. constructor. split; auto. destruct h13b. contradiction. rewrite <- h13 in h5r. eapply cardinal_is_functional. apply h3b. apply h5r. reflexivity. rewrite h3a in h5l. rewrite <- in_bag_to_set_iff in h8. assert (h12:Inn (bag_to_set B') c). rewrite <- h5l. left. assumption. rewrite in_bag_to_set_iff in h12. contradiction. rewrite h3a in h5l. apply inb1_dom_rel_compat in h10. rewrite <- bag_to_set_dom_rel_compat in h10. rewrite <- h5l in h10. destruct h10 as [c h10l | c h10r]. rewrite in_bag_to_set_iff in h10l. contradiction. destruct h10r. contradict h7. reflexivity. pose proof (not_inb1 _ _ h9) as h10. pose proof (not_inb1 _ _ h11) as h12. rewrite h10, h12. reflexivity. Qed. Definition im_set_bag {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pf:Finite A) (f:T->U) : Bag U := proj1_sig (constructive_definite_description _ (im_set_bag_ex pfdu _ pf f)). Lemma im_set_bag_compat : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pf:Finite A) (f:T->U), let B := (im_set_bag pfdu A pf f) in (Im A f = bag_to_set B) /\ (forall pr:U*nat, Inb2 B pr -> (cardinal _ [x:T | Inn A x /\ f x = fst pr] (snd pr))). unfold im_set_bag; intros; destruct constructive_definite_description; auto. Qed. Lemma im_fin_map_bag_ex : forall {T U:Type} {A:Ensemble T} {D:Ensemble U} {def:U} (pfdu:type_eq_dec U) (F:Fin_map A D def), exists! (B:Bag U), (im_fin_map F = bag_to_set B) /\ (forall pr:U*nat, Inb2 B pr -> (cardinal _ [x:T | Inn A x /\ F |-> x = fst pr] (snd pr))). intros T U A D def h0 F. pose proof (im_set_bag_ex h0 A (fin_map_fin_dom F) (fin_map_app F)) as h1. destruct h1 as [B h1]. red in h1. destruct h1 as [h1l h1r]. destruct h1l as [h1a h1b]. exists B. red. split. split. unfold im_fin_map. assumption. assumption. assumption. Qed. Definition im_fin_map_bag {T U:Type} {A:Ensemble T} {D:Ensemble U} {def:U} (pfdu:type_eq_dec U) (F:Fin_map A D def) := proj1_sig (constructive_definite_description _ (im_fin_map_bag_ex pfdu F)). Lemma im_fin_map_bag_compat : forall {T U:Type} {A:Ensemble T} {D:Ensemble U} {def:U} (pfdu:type_eq_dec U) (F:Fin_map A D def), let B := (im_fin_map_bag pfdu F) in (im_fin_map F = bag_to_set B) /\ (forall pr:U*nat, Inb2 B pr -> (cardinal _ [x:T | Inn A x /\ F |-> x = fst pr] (snd pr))). unfold im_fin_map_bag; intros; destruct constructive_definite_description; auto. Qed. Lemma im_set_bag_empty : forall {T U:Type} (pfdu:type_eq_dec U) (f:T->U), im_set_bag pfdu (Empty_set T) (Empty_is_finite T) f = empty_bag U pfdu. intros T U h0 f. pose proof (im_set_bag_compat h0 _ (Empty_is_finite T) f) as h1. destruct h1 as [h1l h1r]. rewrite image_empty in h1l. symmetry in h1l. eapply empty_set_empty_bag_compat in h1l. apply h1l. Qed. Lemma im_set_bag_empty' : forall {T U:Type} (pfdu:type_eq_dec U) (f:T->U) (pf:Finite (Empty_set T)), im_set_bag pfdu (Empty_set T) pf f = empty_bag U pfdu. intros T U h0 f h1. pose proof (im_set_bag_empty h0 f) as h2. pose proof (proof_irrelevance _ h1 (Empty_is_finite _)); subst; auto. Qed. Lemma im_set_bag_empty'' : forall {T U:Type} (pfdu:type_eq_dec U) (f:T->U) (A:Ensemble T) (pfeq:A = Empty_set _) (pf:Finite A), im_set_bag pfdu A pf f = empty_bag U pfdu. intros; subst. apply im_set_bag_empty'; auto. Qed. Lemma im_set_bag_functional1 : forall {T U:Type} (pfdu:type_eq_dec U) (A A':Ensemble T) (pf:Finite A) (pf':Finite A') (f:T->U), A = A' -> im_set_bag pfdu A pf f = im_set_bag pfdu A' pf' f. intros T U h0 A A' h1 h2 f h3. subst. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. Lemma im_set_bag_functional1' : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pf:Finite A) (f f':T->U), f = f' -> im_set_bag pfdu A pf f = im_set_bag pfdu A pf f'. intros; subst; reflexivity. Qed. Lemma im_set_bag_functional2 : forall {T U:Type} (pfdu:type_eq_dec U) (A A':Ensemble T) (pf:Finite A) (pf':Finite A') (f f':T->U), A = A' -> f = f' -> im_set_bag pfdu A pf f = im_set_bag pfdu A' pf' f'. intros T U h0 A A' h1 h2 f f' h3 h4. subst. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. Definition im_set_bag_in {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pfa:Finite A) (P:T->Prop) (f:{x:T|P x}->U) (pfx:(forall x:T, Inn A x -> P x)) : Bag U := im_set_bag pfdu (Full_set {x:T|Inn A x}) (iff1 (finite_full_sig_iff A) pfa) (fun x:{x:T|Inn A x}=>f (exist _ (proj1_sig x) (pfx _ (proj2_sig x)))). Lemma im_set_bag_in_empty : forall {T U:Type} (pfdu:type_eq_dec U) (pfe:Finite (Empty_set _)) (P:T->Prop) (f:{x:T|P x}->U) (pfx:(forall x:T, Inn (Empty_set _) x -> P x)), im_set_bag_in pfdu _ pfe _ f pfx = empty_bag U pfdu. intros T U h0 h1 P f h2. unfold im_set_bag_in. assert (h3:Full_set {x:T | Inn (Empty_set T) x} = Empty_set _). apply Extensionality_Ensembles; red; split; simpl; auto with sets. red. intro x. pose proof (proj2_sig x). contradiction. pose proof (subsetT_eq_compat _ _ _ _ (iff1 (finite_full_sig_iff (Empty_set T)) h1) (Empty_is_finite {x:T | Inn (Empty_set T) x}) h3) as h4. dependent rewrite -> h4. apply im_set_bag_empty. Qed. Lemma im_set_bag_add : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pf:Finite A) (x:T), ~Inn A x -> forall f:T->U, im_set_bag pfdu (Add A x) (Add_preserves_Finite _ _ x pf) f = bag_add (im_set_bag pfdu A pf f) (f x). intros T U h0 A h1 x h2 f. pose proof (im_set_bag_compat h0 _ (Add_preserves_Finite _ _ x h1) f) as h3. destruct h3 as [h3l h3r]. apply bag_to_fin_map_inj. apply functional_extensionality. intro y. destruct (eq_dec y (f x)) as [h4 | h5]. subst. rewrite bag_add_app_eq. pose proof (Add_intro2 _ A x) as h4. assert (h5:Inn (Im (Add A x) f) (f x)). apply Im_intro with x; auto. rewrite h3l in h5. rewrite in_bag_to_set_iff in h5. pose proof (inb1_inb2_compat _ _ h5) as h6. specialize (h3r _ h6). simpl in h3r. pose proof (im_set_bag_compat h0 _ h1 f) as h7. destruct h7 as [h7l h7r]. destruct (inb1_dec (im_set_bag h0 A h1 f) (f x)) as [h8 | h9]. pose proof (inb1_inb2_compat _ _ h8) as h10. specialize (h7r _ h10). simpl in h7r. assert (h11:~Inn [x0:T | Inn A x0 /\ f x0 = f x]x). intro h12. destruct h12 as [h12]. destruct h12; contradiction. pose proof (card_add_nin _ _ h7r _ h11) as h12. assert (h13:Add [x0 : T | Inn A x0 /\ f x0 = f x] x = [x0 : T | Inn (Add A x) x0 /\ f x0 = f x]). apply Extensionality_Ensembles. red. split. red. intros a h14. constructor. destruct h14 as [a h14l | a h14r]. destruct h14l as [h14l]. destruct h14l as [h14a h14b]. split; auto. left. assumption. destruct h14r; subst. split; auto with sets. red. intros a h14. destruct h14 as [h14]. destruct h14 as [h14l h14r]. destruct h14l as [a h14a | a h14b]. left. constructor. constructor; auto. right. assumption. rewrite h13 in h12. eapply cardinal_is_functional; auto. apply h3r. apply h12. pose proof (not_inb1 _ _ h9) as h10. rewrite h10. assert (h11: cardinal _ [x0 : T | Inn (Add A x) x0 /\ f x0 = f x] 1). assert (h12:[x0 : T | Inn (Add A x) x0 /\ f x0 = f x] = Singleton x). apply Extensionality_Ensembles. red. split. intros a h11. destruct h11 as [h11]. destruct h11 as [h11l h11r]. destruct h11l as [a h11a | a h11b]. assert (h12:Inn (Im A f) (f x)). rewrite <- h11r. apply Im_intro with a; auto. rewrite h7l in h12. rewrite in_bag_to_set_iff in h12. contradiction. assumption. red. intros a h11. destruct h11; subst. constructor. split; auto with sets. rewrite h12. apply card_sing. eapply cardinal_is_functional; auto. apply h3r. apply h11. apply neq_sym in h5. rewrite bag_add_app_neq; auto. destruct (inb1_dec (im_set_bag h0 A h1 f) y) as [h6 | h7]. pose proof (inb1_inb2_compat _ _ h6) as h7. pose proof (im_set_bag_compat h0 A h1 f) as h8. destruct h8 as [h8l h8r]. specialize (h8r _ h7). simpl in h8r. assert (h9:Inb1 (im_set_bag h0 (Add A x) (Add_preserves_Finite T A x h1) f) y). rewrite <- in_bag_to_set_iff. rewrite <- h3l. rewrite <- in_bag_to_set_iff in h6. rewrite <- h8l in h6. destruct h6 as [a h6 ? h9]. subst. apply Im_intro with a. left; auto. auto. pose proof (inb1_inb2_compat _ _ h9) as h10. specialize (h3r _ h10). simpl in h3r. assert (h11: [x0 : T | Inn (Add A x) x0 /\ f x0 = y] = [x0 : T | Inn A x0 /\ f x0 = y]). apply Extensionality_Ensembles. red. split. red. intros a h11. destruct h11 as [h11]. destruct h11 as [h11l h11r]. destruct h11l as [a h11a | a h11b]. subst. constructor. split; auto. destruct h11b; subst. contradict h5. reflexivity. red. intros a h11. destruct h11 as [h11]. destruct h11 as [h11l h11r]. subst. constructor. split; auto with sets. rewrite h11 in h3r. eapply cardinal_is_functional; auto. apply h3r. apply h8r. pose proof (not_inb1 _ _ h7) as h8. rewrite h8. assert (hnin:~Inb1 (im_set_bag h0 (Add A x) (Add_preserves_Finite T A x h1) f) y). intro h9. pose proof (inb1_inb2_compat _ _ h9) as h10. specialize (h3r _ h10). simpl in h3r. assert (h11: [x0 : T | Inn (Add A x) x0 /\ f x0 = y] = [x0 : T | Inn A x0 /\ f x0 = y]). apply Extensionality_Ensembles. red. split. red. intros a h11. destruct h11 as [h11]. destruct h11 as [h11l h11r]. destruct h11l as [a h11a | a h11b]. subst. constructor. split; auto. destruct h11b; subst. contradict h5. reflexivity. red. intros a h11. destruct h11 as [h11]. destruct h11 as [h11l h11r]. subst. constructor. split; auto with sets. rewrite h11 in h3r. red in h9. pose proof (card_gt_O_inh _ _ h9 h3r) as h12. destruct h12 as [a h12]. destruct h12 as [h12]. destruct h12 as [h12l h12r]. assert (h13:Inn (Im A f) (f a)). apply Im_intro with a; auto. pose proof (im_set_bag_compat h0 A h1 f) as h14. destruct h14 as [h14l h14r]. rewrite h14l in h13. rewrite in_bag_to_set_iff in h13. subst. contradiction. pose proof (not_inb1 _ _ hnin) as h9. rewrite h9. reflexivity. Qed. Lemma im_set_bag_add' : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (x:T) (pf:Finite (Add A x)) , ~Inn A x -> forall f:T->U, im_set_bag pfdu (Add A x) pf f = bag_add (im_set_bag pfdu A (add_downward_closed _ _ pf) f) (f x). intros T U h0 A x h1 h2 f. pose proof (im_set_bag_add h0 A (add_downward_closed _ _ h1) x h2 f) as h3. assert (h4:h1 = Add_preserves_Finite T A x (add_downward_closed A x h1)). apply proof_irrelevance. rewrite h4, h3. f_equal. f_equal. f_equal. apply proof_irrelevance. Qed. End ImSetBag. Section ImSetBagDep. Lemma im_in_ens_bag_ex : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pf:Finite A) (f:fun_in_ens A U), exists! (B:Bag U), (im_in_ens f = bag_to_set B) /\ (forall pr:U*nat, Inb2 B pr -> (cardinal _ [x:T | exists pf:Inn A x, f x pf = fst pr] (snd pr))). intros T U h0 A h1 f. induction h1 as [|A h2 h3 x h4]. exists (empty_bag U h0). red. split. split. rewrite empty_bag_empty_set_compat. rewrite im_in_ens_empty. reflexivity. intros pr h2. contradict h2. intro h3. pose proof (inb2_impl_inb1 _ _ h3) as h4. contradict h4. apply not_inb1_empty_bag. intros B h1. destruct h1 as [h1l h1r]. rewrite im_in_ens_empty in h1l. symmetry. apply empty_set_empty_bag_compat; auto. specialize (h3 (fun_from_add f)). destruct h3 as [B h3]. red in h3. destruct h3 as [h3l h3r]. destruct h3l as [h3a h3b]. exists (bag_add B (f x (Add_intro2 _ _ _))). red. split. split. simpl. rewrite bag_to_set_add. rewrite im_in_ens_add. f_equal. assumption. assumption. intros pr h5. pose proof (bag_add_inv h0 _ _ _ h5) as h6. destruct h6 as [h6l h6r]. destruct h6l as [h7 | h8]. assert (hneq:fst pr <> f x (Add_intro2 _ _ _)). tauto. specialize (h3b _ h7). assert (h8: [x0 : T | exists pf:Inn (Add A x) x0, f x0 pf = fst pr] = [x0 : T | exists pf:Inn A x0, fun_from_add f x0 pf = fst pr]). apply Extensionality_Ensembles. red. split. red. intros t h8. destruct h8 as [h8]. destruct h8 as [h8l h8r]. destruct h8l as [t h9 | t h10]. constructor. exists h9. rewrite <- h8r. unfold fun_from_add. f_equal. apply proof_irrelevance. destruct h10. subst. hl h6r. rewrite <- h8r. f_equal. apply proof_irrelevance. rewrite h6r in hl. contradiction. red. intros t h8. destruct h8 as [h8]. destruct h8 as [h8l h8r]. constructor. exists (Add_intro1 _ _ _ _ h8l). rewrite <- h8r. unfold fun_from_add; auto. rewrite h8. assumption. rewrite h8. assert (h9:~ Inb2 B pr). tauto. unfold Inb2 in h9. apply not_and_or in h9. destruct h9 as [h9l | h9r]. rewrite h8 in h9l. rewrite surjective_pairing in h5. rewrite h8 in h5. red in h5. simpl in h5. destruct h5 as [h5l h5r]. pose proof (bag_add_app_eq B (f x (Add_intro2 _ _ _))) as h10. simpl in h10. rewrite h10 in h5l. pose proof (f_equal pred h5l) as h11. simpl in h11. destruct (zerop (pred (snd pr))) as [h12 |h13]. rewrite h12 in h11. assert (h13:snd pr = 1). omega. rewrite h13. assert (h14: [x0 : T | exists pf:Inn (Add A x) x0, f x0 pf = f x (Add_intro2 _ _ _)] = Singleton x). apply Extensionality_Ensembles. red. split. red. intros t h14. destruct h14 as [h14]. destruct h14 as [h14l h14r]. destruct h14l as [t h14a | t h14b]. assert (h15:Inn (im_in_ens (fun_from_add f)) (f t (Add_intro1 _ _ _ _ h14a))). econstructor. unfold fun_from_add. reflexivity. rewrite h3a in h15. unfold bag_to_set in h15. inversion h15; subst. pose proof (bag_in_compat _ _ H) as h16. pose proof (nonzero_bag_to_pair_set B) as h17. red in h17. assert (h18:snd x0 <> 0). intro h18. specialize (h17 (fst x0)). rewrite surjective_pairing in H. rewrite h18 in H. contradiction. rewrite h16 in h18. rewrite <- H0 in h18. erewrite f_equal_dep in h18. rewrite h11 in h18 at 1. contradict h18; auto. erewrite f_equal_dep in h18. rewrite h14r in h18 at 1. contradiction. reflexivity. assumption. red. intros x' h15. destruct h15. constructor. exists (Add_intro2 _ _ _); auto. rewrite h14. apply card_sing. rewrite <- h11 in h13. assert (h14:Inb2 B ((f x (Add_intro2 _ _ _)), B{->}f x (Add_intro2 _ _ _))). red. simpl. split; auto with arith. specialize (h3b _ h14). simpl in h3b. rewrite h11 in h3b. pose proof h9l as h9l'. clear h9l. assert (h9: [x0 : T | exists pf:Inn (Add A x) x0, f x0 pf = fst pr] = Add [x0 : T | exists pf:Inn A x0, fun_from_add f x0 pf = fst pr] x). apply Extensionality_Ensembles. red. split. red. intros t h9. destruct h9 as [h9]. destruct h9 as [h9l h9r]. destruct h9l as [t h9a |t h9b]. constructor. constructor. exists h9a. rewrite <- h9r. unfold fun_from_add; f_equal; apply proof_irrelevance. destruct h9b. right. constructor. red. intros t h9. destruct h9 as [t h9l | t h9r]. destruct h9l as [h9l]. destruct h9l as [h9a h9b]. constructor. exists (Add_intro1 _ _ _ _ h9a). rewrite <- h9b. unfold fun_from_add; auto. destruct h9r. constructor. exists (Add_intro2 _ _ _). rewrite h8; auto. rewrite h8 in h9. rewrite h9. rewrite (S_pred _ _ h5r). constructor. assumption. intro h15. destruct h15 as [h15]. destruct h15; contradiction. assert (h10r:snd pr = 0). omega. rewrite surjective_pairing in h5. rewrite h10r in h5. red in h5. simpl in h5. destruct h5; omega. intros B' h5. destruct h5 as [h5l h5r]. rewrite im_in_ens_add in h5l. apply bag_to_fin_map_inj. apply functional_extensionality. intro c. destruct (eq_dec (f x (Add_intro2 _ _ _)) c) as [h6 | h7]. subst. pose proof (Add_intro2 _ (im_in_ens (fun_from_add f)) (f x (Add_intro2 _ _ _))) as h6. rewrite h5l in h6. rewrite in_bag_to_set_iff in h6. destruct (inb1_dec B (f x (Add_intro2 _ _ _))) as [h7 | h8]. rewrite bag_add_app_eq. pose proof (inb1_inb2_compat _ _ h6) as h8. pose proof (inb1_inb2_compat _ _ h7) as h9. specialize (h3b _ h9). simpl in h3b. specialize (h5r _ h8). simpl in h5r. clear h9. assert (h9: [x0 : T | exists pf:Inn (Add A x) x0, f x0 pf = f x (Add_intro2 _ _ _)] = Add [x0 : T | exists pf:Inn A x0, fun_from_add f x0 pf = f x (Add_intro2 _ _ _)] x). apply Extensionality_Ensembles. red. split. red. intros t h9. destruct h9 as [h9]. destruct h9 as [h9l h9r]. destruct h9l as [t h9a |t h9b]. constructor. constructor. exists h9a. unfold fun_from_add. rewrite <- h9r. f_equal. apply proof_irrelevance. destruct h9b. right. constructor. red. intros t h9. destruct h9 as [t h9l | t h9r]. destruct h9l as [h9l]. destruct h9l as [h9a h9b]. constructor. exists (Add_intro1 _ _ _ _ h9a). rewrite <- h9b. unfold fun_from_add; auto. destruct h9r. constructor. exists (Add_intro2 _ _ _); auto. rewrite h9 in h5r. assert (h10:~Inn [x0 : T | exists pf:Inn A x0, fun_from_add f x0 pf = f x (Add_intro2 _ _ _)] x). intro h11. destruct h11 as [h11]. destruct h11; contradiction. pose proof (card_add_nin _ _ h3b _ h10) as h11. eapply cardinal_is_functional. apply h11. apply h5r. reflexivity. pose proof (not_inb1 _ _ h8) as h9. rewrite bag_add_app_eq. rewrite h9. assert (h14: [x0 : T | exists pf:Inn (Add A x) x0, f x0 pf = f x (Add_intro2 _ _ _)] = Singleton x). apply Extensionality_Ensembles. red. split. red. intros t h10. destruct h10 as [h10]. destruct h10 as [h10l h10r]. destruct h10l as [t h10a |t h10b]. assert (h11:Inn (im_in_ens (fun_from_add f)) (f t (Add_intro1 _ _ _ _ h10a))). econstructor. unfold fun_from_add. reflexivity. erewrite f_equal_dep in h11. rewrite h10r in h11 at 1. rewrite h3a in h11. contradict h8. rewrite in_bag_to_set_iff in h11. assumption. reflexivity. assumption. red. intros t h10. destruct h10. constructor. exists (Add_intro2 _ _ _); auto. pose proof (inb1_inb2_compat _ _ h6) as h15. specialize (h5r _ h15). simpl in h5r. rewrite h14 in h5r. pose proof (card_sing x) as h16. eapply cardinal_is_functional. apply h16. apply h5r. reflexivity. rewrite bag_add_app_neq; auto. destruct (inb1_dec B c) as [h8 | h9]; destruct (inb1_dec B' c) as [h10 | h11]. pose proof (inb1_inb2_compat _ _ h8) as h11. pose proof (inb1_inb2_compat _ _ h10) as h12. specialize (h3b _ h11). simpl in h3b. specialize (h5r _ h12). simpl in h5r. assert (h13:[x0 : T | exists pf:Inn A x0, fun_from_add f x0 pf = c] = [x0 : T | exists pf:Inn (Add A x) x0, f x0 pf = c]). apply Extensionality_Ensembles. red. split. red. intros t h13. destruct h13 as [h13]. destruct h13 as [h13l h13r]. constructor. exists (Add_intro1 _ _ _ _ h13l). rewrite <- h13r. unfold fun_from_add; auto. red. intros t h13. destruct h13 as [h13]. destruct h13 as [h13l h13r]. destruct h13l as [t h13a | t h13b]. constructor. exists h13a. rewrite <- h13r. unfold fun_from_add. f_equal. apply proof_irrelevance. destruct h13b. contradict h7. rewrite <- h13r. f_equal; apply proof_irrelevance. rewrite <- h13 in h5r. eapply cardinal_is_functional. apply h3b. apply h5r. reflexivity. rewrite h3a in h5l. rewrite <- in_bag_to_set_iff in h8. assert (h12:Inn (bag_to_set B') c). rewrite <- h5l. left. assumption. rewrite in_bag_to_set_iff in h12. contradiction. rewrite h3a in h5l. apply inb1_dom_rel_compat in h10. rewrite <- bag_to_set_dom_rel_compat in h10. rewrite <- h5l in h10. destruct h10 as [c h10l | c h10r]. rewrite in_bag_to_set_iff in h10l. contradiction. destruct h10r. contradict h7. reflexivity. pose proof (not_inb1 _ _ h9) as h10. pose proof (not_inb1 _ _ h11) as h12. rewrite h10, h12. reflexivity. assumption. Qed. Definition im_in_ens_bag {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pf:Finite A) (f:fun_in_ens A U) : Bag U := proj1_sig (constructive_definite_description _ (im_in_ens_bag_ex pfdu _ pf f)). Lemma im_in_ens_bag_compat : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pf:Finite A) (f:fun_in_ens A U), let B := (im_in_ens_bag pfdu A pf f) in (im_in_ens f = bag_to_set B) /\ (forall pr:U*nat, Inb2 B pr -> (cardinal _ [x:T | exists pf:Inn A x, f x pf = fst pr] (snd pr))). unfold im_in_ens_bag; intros; destruct constructive_definite_description; auto. Qed. Lemma im_in_ens_bag_empty : forall {T U:Type} (pfdu:type_eq_dec U) (f:fun_in_ens (Empty_set T) U), im_in_ens_bag pfdu (Empty_set T) (Empty_is_finite T) f = empty_bag U pfdu. intros T U h0 f. pose proof (im_in_ens_bag_compat h0 _ (Empty_is_finite T) f) as h1. destruct h1 as [h1l h1r]. rewrite im_in_ens_empty in h1l. symmetry in h1l. eapply empty_set_empty_bag_compat in h1l. apply h1l. Qed. Lemma im_in_ens_bag_empty' : forall {T U:Type} (pfdu:type_eq_dec U) (f:fun_in_ens (Empty_set _) U) (pf:Finite (Empty_set T)), im_in_ens_bag pfdu (Empty_set T) pf f = empty_bag U pfdu. intros T U h0 f h1. pose proof (im_in_ens_bag_empty h0 f) as h2. assert (h1 = Empty_is_finite _). apply proof_irrelevance. subst. assumption. Qed. Lemma im_in_ens_bag_empty'' : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (f:fun_in_ens A U) (pfeq:A = Empty_set _) (pf:Finite A), im_in_ens_bag pfdu A pf f = empty_bag U pfdu. intros; subst. apply im_in_ens_bag_empty'; auto. Qed. Lemma im_in_ens_bag_functional1 : forall {T U:Type} (pfdu:type_eq_dec U) (A A':Ensemble T) (pf:Finite A) (pf':Finite A') (f:fun_in_ens A U) (pfe:A = A'), let f' := fun_in_ens_transfer f pfe in im_in_ens_bag pfdu A pf f = im_in_ens_bag pfdu A' pf' f'. intros T U h0 A A' h1 h2 f h3. subst; simpl. assert (h1 = h2). apply proof_irrelevance. subst. rewrite fun_in_ens_transfer_eq_refl; auto. Qed. Lemma im_in_ens_bag_functional1' : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pf:Finite A) (f f':fun_in_ens A U), f = f' -> im_in_ens_bag pfdu A pf f = im_in_ens_bag pfdu A pf f'. intros; subst; reflexivity. Qed. Lemma im_in_ens_bag_functional2 : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pf:Finite A) (pf':Finite A) (f:fun_in_ens A U), im_in_ens_bag pfdu A pf f = im_in_ens_bag pfdu A pf' f. intros; f_equal; apply proof_irrelevance. Qed. Lemma im_in_ens_bag_add : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (pf:Finite A) (x:T), ~Inn A x -> forall f:fun_in_ens (Add A x) U, let f' := fun_from_add f in im_in_ens_bag pfdu (Add A x) (Add_preserves_Finite _ _ x pf) f = bag_add (im_in_ens_bag pfdu A pf f') (f x (Add_intro2 _ _ _)). intros T U h0 A h1 x h2 f. pose proof (im_in_ens_bag_compat h0 _ (Add_preserves_Finite _ _ x h1) f) as h3. destruct h3 as [h3l h3r]. apply bag_to_fin_map_inj. apply functional_extensionality. intro y. destruct (eq_dec y (f x (Add_intro2 _ _ _))) as [h4 | h5]. subst. rewrite bag_add_app_eq. pose proof (Add_intro2 _ A x) as h4. assert (h5:Inn (im_in_ens f) (f x (Add_intro2 _ _ _))). econstructor. reflexivity. rewrite h3l in h5. rewrite in_bag_to_set_iff in h5. pose proof (inb1_inb2_compat _ _ h5) as h6. specialize (h3r _ h6). simpl in h3r. pose proof (im_in_ens_bag_compat h0 _ h1 (fun_from_add f)) as h7. destruct h7 as [h7l h7r]. destruct (inb1_dec (im_in_ens_bag h0 A h1 (fun_from_add f)) (f x (Add_intro2 _ _ _))) as [h8 | h9]. pose proof (inb1_inb2_compat _ _ h8) as h10. specialize (h7r _ h10). simpl in h7r. assert (h11:~Inn [x0:T | exists pf:Inn A x0, f x0 (Add_intro1 _ _ _ _ pf) = f x (Add_intro2 _ _ _)] x). intro h12. destruct h12 as [h12]. destruct h12; contradiction. pose proof (card_add_nin _ _ h7r _ h11) as h12. assert (h13:Add [x0 : T | exists pf:Inn A x0, (fun_from_add f) x0 pf = f x (Add_intro2 _ _ _) ] x = [x0 : T | exists pf:Inn (Add A x) x0, f x0 pf = f x (Add_intro2 _ _ _)]). apply Extensionality_Ensembles. red. split. red. intros a h14. constructor. destruct h14 as [a h14l | a h14r]. destruct h14l as [h14l]. destruct h14l as [h14a h14b]. ex_goal. left; auto. exists hex. erewrite f_equal_dep. unfold fun_from_add in h14b. rewrite h14b at 1. f_equal. reflexivity. destruct h14r; subst. exists (Add_intro2 _ _ _); auto. red. intros a h14. destruct h14 as [h14]. destruct h14 as [h14l h14r]. destruct h14l as [a h14a | a h14b]. left. constructor. exists h14a. unfold fun_from_add. erewrite f_equal_dep. rewrite h14r at 1; auto. reflexivity. right. assumption. rewrite h13 in h12. eapply cardinal_is_functional; auto. apply h3r. apply h12. pose proof (not_inb1 _ _ h9) as h10. rewrite h10. assert (h11: cardinal _ [x0 : T | exists pf:Inn (Add A x) x0, f x0 pf = f x (Add_intro2 _ _ _)] 1). assert (h12:[x0 : T | exists pf:Inn (Add A x) x0, f x0 pf = f x (Add_intro2 _ _ _)] = Singleton x). apply Extensionality_Ensembles. red. split. intros a h11. destruct h11 as [h11]. destruct h11 as [h11l h11r]. destruct h11l as [a h11a | a h11b]. assert (h12:Inn (im_in_ens (fun_from_add f)) (f x (Add_intro2 _ _ _))). rewrite <- h11r. econstructor. unfold fun_from_add. f_equal. apply proof_irrelevance. rewrite h7l in h12. rewrite in_bag_to_set_iff in h12. contradiction. assumption. red. intros a h11. destruct h11; subst. constructor. exists (Add_intro2 _ _ _). reflexivity. rewrite h12. apply card_sing. eapply cardinal_is_functional; auto. apply h3r. apply h11. apply neq_sym in h5. rewrite bag_add_app_neq; auto. destruct (inb1_dec (im_in_ens_bag h0 A h1 (fun_from_add f)) y) as [h6 | h7]. pose proof (inb1_inb2_compat _ _ h6) as h7. pose proof (im_in_ens_bag_compat h0 A h1 (fun_from_add f)) as h8. destruct h8 as [h8l h8r]. specialize (h8r _ h7). simpl in h8r. assert (h9:Inb1 (im_in_ens_bag h0 (Add A x) (Add_preserves_Finite T A x h1) f) y). rewrite <- in_bag_to_set_iff. rewrite <- h3l. rewrite <- in_bag_to_set_iff in h6. rewrite <- h8l in h6. destruct h6 as [a h6 ? h9]. subst. econstructor. unfold fun_from_add. reflexivity. pose proof (inb1_inb2_compat _ _ h9) as h10. specialize (h3r _ h10). simpl in h3r. assert (h11: [x0 : T | exists pf:Inn (Add A x) x0, f x0 pf = y] = [x0 : T | exists pf:Inn A x0, f x0 (Add_intro1 _ _ _ _ pf) = y]). apply Extensionality_Ensembles. red. split. red. intros a h11. destruct h11 as [h11]. destruct h11 as [h11l h11r]. destruct h11l as [a h11a | a h11b]. subst. constructor. exists h11a. f_equal; apply proof_irrelevance. destruct h11b; subst. contradict h5. f_equal; apply proof_irrelevance. red. intros a h11. destruct h11 as [h11]. destruct h11 as [h11l h11r]. subst. constructor. ex_goal. left; auto. exists hex. f_equal; apply proof_irrelevance. rewrite h11 in h3r. eapply cardinal_is_functional; auto. apply h3r. apply h8r. pose proof (not_inb1 _ _ h7) as h8. rewrite h8. assert (hnin:~Inb1 (im_in_ens_bag h0 (Add A x) (Add_preserves_Finite T A x h1) f) y). intro h9. pose proof (inb1_inb2_compat _ _ h9) as h10. specialize (h3r _ h10). simpl in h3r. assert (h11: [x0 : T | exists pf:Inn (Add A x) x0, f x0 pf = y] = [x0 : T | exists pf:Inn A x0, f x0 (Add_intro1 _ _ _ _ pf) = y]). apply Extensionality_Ensembles. red. split. red. intros a h11. destruct h11 as [h11]. destruct h11 as [h11l h11r]. destruct h11l as [a h11a | a h11b]. subst. constructor. exists h11a. f_equal. apply proof_irrelevance. destruct h11b; subst. contradict h5. f_equal. apply proof_irrelevance. red. intros a h11. destruct h11 as [h11]. destruct h11 as [h11l h11r]. subst. constructor. exists (Add_intro1 _ _ _ _ h11l). reflexivity. rewrite h11 in h3r. red in h9. pose proof (card_gt_O_inh _ _ h9 h3r) as h12. destruct h12 as [a h12]. destruct h12 as [h12]. destruct h12 as [h12l h12r]. assert (h13:Inn (im_in_ens (fun_from_add f)) (f a (Add_intro1 _ _ _ _ h12l))). econstructor. unfold fun_from_add. f_equal. pose proof (im_in_ens_bag_compat h0 A h1 (fun_from_add f)) as h14. destruct h14 as [h14l h14r]. rewrite h14l in h13. rewrite in_bag_to_set_iff in h13. subst. contradiction. pose proof (not_inb1 _ _ hnin) as h9. rewrite h9. reflexivity. Grab Existential Variables. assumption. Qed. Lemma im_in_ens_bag_add' : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (x:T) (pf:Finite (Add A x)) , ~Inn A x -> forall f:fun_in_ens (Add A x) U, im_in_ens_bag pfdu (Add A x) pf f = bag_add (im_in_ens_bag pfdu A (add_downward_closed _ _ pf) (fun_from_add f)) (f x (Add_intro2 _ _ _)). intros T U h0 A x h1 h2 f. pose proof (im_in_ens_bag_add h0 A (add_downward_closed _ _ h1) x h2 f) as h3. assert (h4:h1 = Add_preserves_Finite T A x (add_downward_closed A x h1)). apply proof_irrelevance. rewrite h4, h3. f_equal. f_equal. f_equal. apply proof_irrelevance. Qed. End ImSetBagDep. Lemma card_bag_ex : forall {T:Type} (pfdt:type_eq_dec T) (B:Bag T), exists! n:nat, forall (l:list (T*nat)), NoDup l -> list_to_set l = bag_to_pair_set B -> n = plusl (map (@snd _ _) l). intros T h0 B. pose proof (finite_bag_to_pair_set B) as h1. pose proof (finite_set_list_no_dup _ h1) as h2. destruct h2 as [l h2]. destruct h2 as [h2l h2r]. exists (plusl (map (snd (B:=nat)) l)). red. split. intros l' h3 h4. rewrite h2l in h4. apply plusl_nat_valued_fun_functional; auto. apply impl_type_eq_dec_prod; auto. apply nat_eq_dec. intros n h3. symmetry. symmetry in h2l. apply h3; auto. Qed. Definition card_bag {T:Type} (pfdt:type_eq_dec T) (B:Bag T) := proj1_sig (constructive_definite_description _ (card_bag_ex pfdt B)). Lemma card_bag_compat : forall {T:Type} (pfdt:type_eq_dec T) (B:Bag T) (l:list (T*nat)), NoDup l -> list_to_set l = bag_to_pair_set B -> (card_bag pfdt B) = plusl (map (@snd _ _) l). intros. unfold card_bag. destruct constructive_definite_description as [? h2]. simpl. apply h2; auto. Qed. (*The "power set" of a bag, except for the bag itself *) Definition power_bag {T:Type} (B:Bag T) := [B':Bag T | Inclb B' B]. (*maximum mulitiplicity of a bag *) Definition max_mult {T:Type} (B:Bag T) (l:list nat) (pfn:l <> nil) (pfe:(list_to_set l = Im (bag_to_set B) (fin_map_app (bag_to_fin_map B)))) : nat := maxl l pfn. Lemma power_bag_empty : forall (T:Type) (pfte:type_eq_dec T), power_bag (empty_bag T pfte) = Singleton (empty_bag T pfte). intros T hte. unfold power_bag. apply Extensionality_Ensembles; red; split. red. intros B h1. destruct h1 as [h1]. red in h1. destruct (empty_bag_eq_dec hte B) as [h2 | h3]. subst. constructor. erewrite empty_bag_eq in h3. eapply non_empty_bag_inhabited in h3. destruct h3 as [x h3]. specialize (h1 _ h3). red in h1. simpl in h1. destruct h1 as [h1l h1r]. pose proof (empty_bag_O hte x) as h4. simpl in h4. simpl in h1l. rewrite h4 in h1l. omega. red. intros B h1. destruct h1. constructor. apply inclb_refl. Qed. Lemma finite_power_bag : forall {T:Type} (pfdt:type_eq_dec T) (B:Bag T), Finite (power_bag B). intros T h0 B. pose (bag_to_set B) as A. pose proof (finite_bag_to_set B) as h1. pose proof (finite_image _ _ (bag_to_set B) (fin_map_app (bag_to_fin_map B)) h1) as h2. pose proof (finite_set_list _ h2) as h3. destruct h3 as [l h3]. symmetry in h3. destruct (nil_dec' l) as [|hn]; subst. simpl in h3. symmetry in h3. apply empty_image in h3. apply (empty_set_empty_bag_compat h0) in h3. subst. rewrite power_bag_empty. apply Singleton_is_finite. pose (max_mult _ _ hn h3) as mm. pose (seg_weak mm) as D. pose proof (finite_seg_weak mm) as h4. pose proof (cart_prod_fin _ _ h1 h4) as h5. pose proof (power_set_finitet _ h5) as h6. rewrite finite_full_sig_iff. rewrite Finite_FiniteT_iff. assert (h7:forall B':Bag T, Inn (power_bag B) B' -> Included (bag_to_pair_set B') (cart_prod (bag_to_set B) (seg_weak mm))). intros B' h7. destruct h7 as [h7]. red. intros pr h8. red in h7. pose proof (in_inb2_compat _ _ h8) as h9. pose proof (inb2_impl_inb1 _ _ h9) as h10. specialize (h7 _ h10). pose proof (bag_in_compat _ _ h8) as h11. rewrite <- h11 in h7. constructor. split. red in h7. simpl in h7. destruct h7 as [h7l h7r]. unfold bag_to_set. assert (h12:Inb1 B (fst pr)). red. omega. apply Im_intro with (fst pr, B{->}fst pr). apply inb1_in_compat. assumption. simpl. reflexivity. constructor. unfold mm. unfold max_mult. red in h7. simpl in h7. destruct h7 as [h7l h7r]. assert (h12:In (B{->}fst pr) l). rewrite list_to_set_in_iff. rewrite h3. apply Im_intro with (fst pr). unfold bag_to_set. apply Im_intro with (fst pr, B{->}fst pr). apply inb1_in_compat. red. omega. simpl. reflexivity. reflexivity. pose proof (maxl_compat _ _ h12) as h13. simpl in h13. assert (he:hn = in_not_nil l _ h12). apply proof_irrelevance. rewrite he. omega. pose (fun B':{x:Bag T | Inn (power_bag B) x} => (exist (fun S => Included S (cart_prod (bag_to_set B) (seg_weak mm))) (bag_to_pair_set (proj1_sig B')) (h7 _ (proj2_sig B')))) as f. apply (inj_finite _ _ f). assumption. red. intros B1 B2 h8. destruct B1; destruct B2. unfold f in h8. simpl in h8. pose proof (exist_injective _ _ _ _ _ h8) as h9. pose proof (bag_to_pair_set_inj _ _ _ h9). subst. apply proj1_sig_injective. simpl. reflexivity. intros. apply classic. Qed. Lemma empty_bag_in_power_bag : forall {T:Type} (B:Bag T), Inn (power_bag B) (empty_bag T (bag_type_eq_dec B)). intros T B. constructor. red. intros x h1. pose proof (not_inb1_empty_bag _ (bag_type_eq_dec B) x). contradiction. Qed. Lemma full_bag_in_power_bag : forall {T:Type} (B:Bag T), Inn (power_bag B) B. intros T B. constructor. apply inclb_refl. Qed. Lemma bag_inductive_step_coarsen : forall {T:Type} (P:(Bag T)->Prop), (forall (B:Bag T) (x:T), P B -> P (bag_add B x)) -> (forall (B:Bag T) (pr:T*nat), B{->}fst pr <= snd pr -> P B -> P (bag_pair_add B pr)). intros T P h1 B pr h2. rewrite (surjective_pairing pr). rewrite bag_pair_add_bag_add_n_compat; auto. simpl. revert h1 h2. induction (snd pr - B{->}fst pr) as [|n h1]. simpl. auto. intros h2 h3 h4. simpl. specialize (h1 h2). specialize (h1 h3 h4). specialize (h2 (bag_add_n B (fst pr) n) (fst pr)). specialize (h2 h1). assumption. Qed. Lemma bag_induction : forall {T:Type} (pfdt:type_eq_dec T) (P:(Bag T)->Prop), P (empty_bag T pfdt) -> (forall (B:Bag T) (x:T), P B -> P (bag_add B x)) -> forall B:(Bag T), P B. intros T hdt P h1 h2 B. pose proof (bag_inductive_step_coarsen _ h2) as h3. clear h2. pose proof (finite_bag_to_pair_set B) as h4. pose proof (finite_set_list_no_dup _ h4) as h5. destruct h5 as [l h5]. destruct h5 as [h5l h5r]. revert h5l h5r. clear h4. revert P h1 B h3. induction l as [|pr l h6]. intros. simpl in h5l. apply empty_pair_set_empty_bag_compat in h5l. erewrite (empty_bag_eq _ hdt) in h5l. subst. assumption. intros P h1 B h2 h3l h3r. pose proof (bag_to_pair_set_subtract_cons _ _ _ h3r h3l) as h4. pose proof (no_dup_cons _ _ h3r) as h9. specialize (h6 _ h1 _ h2 h4 h9). assert (h10:(bag_pair_subtract B pr){->}(fst pr) <= B{->}(fst pr)). destruct (eq_dec (snd pr) (B{->}fst pr)) as [h10 |h11]. rewrite (surjective_pairing pr). rewrite h10. rewrite bag_pair_subtract_app_eq_O. simpl. omega. assert (h12:~Inb2 B pr). intro h13. red in h13. destruct h13 as [h13l h13r]. symmetry in h13l. contradiction. pose proof (bag_pair_subtract_ninb2 _ _ h12) as h13. pose proof (bag_app_functional _ _ h13 (fst pr)). omega. specialize (h2 _ (fst pr, B{->}fst pr) h10 h6). pose proof (bag_pair_add_undoes_bag_pair_subtract B pr) as h11. rewrite h11 in h2. assumption. Qed. Lemma plus_bag_nat_ex : forall (B:Bag nat), exists! n:nat, forall l:list (nat*nat), NoDup l -> list_to_set l = bag_to_pair_set B -> (n = (plusl (map (fun pr=>(fst pr) * (snd pr)) l))). intro B. pose proof (finite_bag_to_pair_set B) as h1. pose proof (finite_set_list_no_dup _ h1) as h2. destruct h2 as [l h2]. destruct h2 as [h2l h2r]. exists (plusl (map (fun pr : nat * nat => fst pr * snd pr) l)). red. split. intros l' h3 h4. rewrite h2l in h4. erewrite plusl_nat_valued_fun_functional; auto. apply impl_type_eq_dec_prod; apply nat_eq_dec. intros n h3. symmetry. symmetry in h2l. apply h3; auto. Qed. Definition plus_bag_nat (B:Bag nat) := proj1_sig (constructive_definite_description _ (plus_bag_nat_ex B)). Lemma plus_bag_nat_compat : forall (B:Bag nat) (l:list (nat*nat)), NoDup l -> list_to_set l = bag_to_pair_set B -> (plus_bag_nat B) = (plusl (map (fun pr=>(fst pr) * (snd pr)) l)). intros. unfold plus_bag_nat. destruct constructive_definite_description as [n h4]. simpl. apply h4; auto. Qed. Lemma plus_bag_nat_empty : plus_bag_nat (empty_bag nat nat_eq_dec) = 0. pose proof (empty_bag_empty_pair_set_compat nat nat_eq_dec) as h1. assert (h2:list_to_set (@nil (nat*nat)) = Empty_set _). auto. assert (h3:NoDup (@nil (nat*nat))). constructor. rewrite <- h1 in h2. pose proof (plus_bag_nat_compat _ nil h3 h2) as h4. simpl in h4. assumption. Qed. Lemma plus_bag_nat_add : forall (B:Bag nat) (n:nat), plus_bag_nat (bag_add B n) = plus_bag_nat B + n. intros B n. pose proof (plus_bag_nat_compat (bag_add B n)) as h3. pose proof (plus_bag_nat_compat B) as h4. pose proof (finite_bag_to_pair_set (bag_add B n)) as h5. pose proof (finite_set_list_no_dup _ h5) as h6. destruct h6 as [l h6]. destruct h6 as [h6l h6r]. symmetry in h6l. specialize (h3 _ h6r h6l). rewrite h3. pose proof (finite_bag_to_pair_set B) as h7. pose proof (finite_set_list_no_dup _ h7) as h8. destruct h8 as [l' h8r]. destruct h8r as [h8a h8b]. symmetry in h8a. specialize (h4 _ h8b h8a). rewrite h4. destruct (inb1_dec B n) as [h9 | h10]. pose proof (inb1_inb2_compat _ _ h9) as h10. pose proof (inb2_in_compat _ _ h10) as h11. rewrite <- h8a in h11 at 1. rewrite <- list_to_set_in_iff in h11. pose proof (impl_type_eq_dec_prod eq_nat_dec eq_nat_dec) as hi. rewrite (plusl_map_remove hi l' _ _ h8b h11). simpl. pose proof (bag_add_intro1 B n h9 n) as h12. pose proof (inb1_inb2_compat _ _ h12) as h13. pose proof (inb2_in_compat _ _ h13) as h14. rewrite <- h6l in h14 at 1. rewrite <- list_to_set_in_iff in h14. rewrite (plusl_map_remove hi l _ _ h6r h14). simpl. simpl. pose proof (bag_add_app_eq B n) as h15. simpl in h15. rewrite h15. pose proof (no_dup_remove hi _ (n, (bag_add B n){->}n) h6r) as h16. simpl in h16. rewrite h15 in h16. pose proof (no_dup_remove hi _ (n, B{->}n) h8b) as h17. assert (h18:list_to_set (remove hi (n, S (B{->}n)) l) = list_to_set (remove hi (n, B{->}n) l')). apply Extensionality_Ensembles. red. split. red. intros x h18. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h18. apply in_remove_inv in h18. destruct h18 as [h18l h18r]. rewrite in_remove_iff. assert (hneq:fst x <> n). intro h19. subst. rewrite list_to_set_in_iff in h18l. rewrite h6l in h18l. apply in_inb2_compat in h18l. red in h18l. destruct h18l as [h18a h18b]. rewrite bag_add_app_eq in h18a. assert (x = (fst x, S (B{->}fst x))). apply injective_projections; auto. contradiction. split. rewrite list_to_set_in_iff in h18l. rewrite h6l in h18l. apply in_inb2_compat in h18l. apply bag_add_inv in h18l. destruct h18l as [h18a h18b]. destruct h18a as [h19 | h20]. apply inb2_in_compat in h19. rewrite <- h8a in h19. rewrite <- list_to_set_in_iff in h19. assumption. subst. contradict hneq. reflexivity. apply nat_eq_dec. intro h19. rewrite (surjective_pairing x) in h19. inversion h19. contradiction. red. intros x h18. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h18. apply in_remove_inv in h18. destruct h18 as [h18l h18r]. rewrite in_remove_iff. assert (hneq:fst x <> n). intro h19. subst. rewrite list_to_set_in_iff in h18l. rewrite h8a in h18l. apply in_inb2_compat in h18l. red in h18l. destruct h18l as [h18a h18b]. assert (x = (fst x, (B{->}fst x))). apply injective_projections; auto. contradiction. split. rewrite list_to_set_in_iff in h18l. rewrite h8a in h18l. apply in_inb2_compat in h18l. pose proof (inb2_impl_inb1 _ _ h18l) as hinb1. pose proof (bag_add_intro1 _ _ hinb1 n) as h19. pose proof (inb1_inb2_compat _ _ h19) as h20. rewrite bag_add_app_neq in h20. pose proof (bag_add_inv nat_eq_dec _ _ _ h20) as h21. simpl in h21. destruct h21 as [h21a h21b]. destruct h21a as [h21l | h21r]. rewrite surjective_pairing in h18l. pose proof (inb2_inj _ _ _ _ h18l h21l ) as h22. rewrite <- h22 in h20. rewrite list_to_set_in_iff. rewrite h6l. apply inb2_in_compat. rewrite surjective_pairing. assumption. contradiction. apply neq_sym. assumption. intro h19. rewrite (surjective_pairing x) in h19. inversion h19. contradiction. pose proof (plusl_nat_valued_fun_functional hi _ _ h16 h17 h18 (fun pr:nat*nat => fst pr * snd pr)) as h19. rewrite h19. ring. pose proof (not_inb1 _ _ h10) as h11. assert (h12:Inb2 (bag_add B n) (n, 1)). red. simpl. split. pose proof (bag_add_app_eq B n) as h13. simpl in h13. simpl. rewrite h13. rewrite h11. reflexivity. auto with arith. apply inb2_in_compat in h12. rewrite <- h6l in h12. rewrite <- list_to_set_in_iff in h12. pose proof (impl_type_eq_dec_prod eq_nat_dec eq_nat_dec) as hi. pose proof (plusl_map_remove hi _ (fun pr:nat*nat=>fst pr*snd pr) _ h6r h12) as h13. rewrite h13. simpl. pose proof (no_dup_remove hi _ (n, 1) h6r) as h14. assert (h15:list_to_set (remove hi (n, 1) l) = list_to_set l'). apply Extensionality_Ensembles. red. split. red. intros x h15. rewrite <- list_to_set_in_iff in h15. rewrite in_remove_iff in h15. destruct h15 as [h15l h15r]. rewrite h8a. rewrite list_to_set_in_iff in h15l. rewrite h6l in h15l. apply in_inb2_compat in h15l. assert (h16:fst x <> n). intro h17. subst. red in h15l. destruct h15l as [h15a h15b]. rewrite bag_add_app_eq in h15a. rewrite h11 in h15a. rewrite h15a in h15r. contradict h15r. rewrite (surjective_pairing x) at 1. reflexivity. pose proof (bag_add_inv nat_eq_dec _ _ _ h15l) as h17. destruct h17 as [h17l h17r]. destruct h17l as [h17a | h17b]. apply inb2_in_compat; auto. contradiction. red. intros x h15. rewrite <- list_to_set_in_iff. rewrite in_remove_iff. split. rewrite h8a in h15. apply in_inb2_compat in h15. pose proof (inb2_impl_inb1 _ _ h15) as h16. pose proof (bag_add_intro1 B (fst x) h16 n) as h17. apply inb1_in_compat in h17. rewrite <- h6l in h17 at 1. rewrite list_to_set_in_iff. assert (h18:n <> fst x). intro h18. subst. contradiction. rewrite bag_add_app_neq in h17; auto. red in h15. destruct h15 as [h15l h15r]. rewrite h15l in h17. rewrite surjective_pairing. assumption. intro h16. rewrite h8a in h15. apply in_inb2_compat in h15. apply inb2_impl_inb1 in h15. rewrite (surjective_pairing x) in h16. inversion h16. subst. contradiction. pose proof (plusl_nat_valued_fun_functional hi _ _ h14 h8b h15 (fun pr:nat*nat => fst pr * snd pr)) as h16. rewrite h16. ring. Qed. Lemma plus_bag_nat_im_set_bag_add : forall {T:Type} (A:Ensemble T) (pf:Finite A) (x:T) (f:T->nat), ~Inn A x -> plus_bag_nat (im_set_bag nat_eq_dec (Add A x) (Add_preserves_Finite _ _ x pf) f) = plus_bag_nat (im_set_bag nat_eq_dec A pf f) + (f x). intros T A h1 x f hnin. rewrite im_set_bag_add; auto. apply plus_bag_nat_add. Qed. Lemma plus_bag_nat_im_set_bag_eq_f : forall {T:Type} (A:Ensemble T) (pf:Finite A) (f:T->nat) (n:nat), (forall x:T, Inn A x -> f x = n) -> plus_bag_nat (im_set_bag nat_eq_dec A pf f) = (card_fun1 A) * n. intros T A h1. pose proof (finite_set_list_no_dup _ h1) as h2. destruct h2 as [l h2]. revert h1 h2. revert A. induction l as [|a l h1]; simpl. intros A h1 h2 f n h3. destruct h2 as [h2l h2r]. subst. rewrite card_fun1_empty. assert (h4:h1 = Empty_is_finite T). apply proof_irrelevance. subst. rewrite im_set_bag_empty. rewrite plus_bag_nat_empty. auto with arith. intros A h2 h3 f n h4. destruct h3 as [h3l h3r]. pose proof (no_dup_cons_nin _ _ h3r) as h5. pose proof (no_dup_cons _ _ h3r) as h6. subst. rewrite list_to_set_in_iff in h5. pose proof (list_to_set_finite l) as h7. pose proof (plus_bag_nat_im_set_bag_add _ h7 _ f h5) as h8. assert (h9:h2 = Add_preserves_Finite _ (list_to_set l) a h7). apply proof_irrelevance. subst. rewrite h8. pose proof (h4 _ (Add_intro2 _ (list_to_set l) a)) as h9. rewrite h9. assert (h10:forall x:T, Inn (list_to_set l) x -> f x = n). intros x h11. apply (h4 _ (Add_intro1 _ _ a _ h11)). specialize (h1 _ h7 (conj (eq_refl _) h6) f n h10). rewrite h1. rewrite card_add_nin'; auto. ring. Qed. Lemma plus_bag_nat_im_set_bag_eq_f' : forall {T:Type} (A A':Ensemble T) (pf:Finite A) (pf':Finite A') (f:T->nat), (exists g:T->T, injective g /\ A' = Im A g /\ forall x:T, Inn A x -> f x = f (g x)) -> plus_bag_nat (im_set_bag nat_eq_dec A pf f) = plus_bag_nat (im_set_bag nat_eq_dec A' pf' f). intros T A A' h1 h2 f h3. destruct h3 as [g [h0 [h3 h4]]]. subst. pose proof (finite_set_list_no_dup _ h1) as h3. destruct h3 as [l h3]. revert h2 h3. revert h4. revert h0. revert g. revert f h1. revert A. induction l as [|a l h1]; simpl. intros A f h1 g h0 h2 h3 h4. destruct h4 as [h4 h5]. subst. rewrite im_set_bag_empty'. rewrite im_set_bag_empty''; auto. rewrite image_empty; auto. intros A f h0 g hi h4 h5 h6. destruct h6 as [h6 h7]. subst. pose proof (no_dup_cons _ _ h7) as h8. pose proof (no_dup_cons_nin _ _ h7) as h9. rewrite im_set_bag_add'. rewrite plus_bag_nat_add. assert (h4' : forall x:T, Inn (list_to_set l) x -> f x = f (g x)). intros x h10. apply h4. left; auto. specialize (h1 _ f (list_to_set_finite _) g hi h4' (finite_image _ _ _ _ (list_to_set_finite _)) (conj (eq_refl _) h8)). assert (h10:list_to_set_finite l = add_downward_closed (list_to_set l) a h0). apply proof_irrelevance. rewrite h10 in h1. rewrite h1. symmetry. generalize h5. rewrite Im_add. intro h2. rewrite im_set_bag_add'. rewrite plus_bag_nat_add. rewrite <- h4. f_equal. f_equal. f_equal. apply proof_irrelevance. right. constructor. intro h3. inversion h3 as [x h11 q hq]. subst. apply hi in hq. subst. rewrite <- list_to_set_in_iff in h11. contradiction. rewrite <- list_to_set_in_iff; auto. Qed. Lemma plus_bag_nat_im_set_bag_eq_f_hetero : forall {T U:Type} (A:Ensemble T) (A':Ensemble U) (pf:Finite A) (pf':Finite A') (f:T->nat) (f':U->nat), (exists g:T->U, injective g /\ A' = Im A g /\ forall x:T, Inn A x -> f x = f' (g x)) -> plus_bag_nat (im_set_bag nat_eq_dec A pf f) = plus_bag_nat (im_set_bag nat_eq_dec A' pf' f'). intros T U A A' h1 h2 f f' h3. destruct h3 as [g [h0 [h3 h4]]]. subst. pose proof (finite_set_list_no_dup _ h1) as h3. destruct h3 as [l h3]. revert h2 h3. revert h4. revert h0. revert g. revert f h1. revert A. induction l as [|a l h1]; simpl. intros A f h1 g h0 h2 h3 h4. destruct h4 as [h4 h5]. subst. rewrite im_set_bag_empty'. rewrite im_set_bag_empty''; auto. rewrite image_empty; auto. intros A f h0 g hi h4 h5 h6. destruct h6 as [h6 h7]. subst. pose proof (no_dup_cons _ _ h7) as h8. pose proof (no_dup_cons_nin _ _ h7) as h9. rewrite im_set_bag_add'. rewrite plus_bag_nat_add. assert (h4' : forall x:T, Inn (list_to_set l) x -> f x = f' (g x)). intros x h10. apply h4. left; auto. specialize (h1 _ f (list_to_set_finite _) g hi h4' (finite_image _ _ _ _ (list_to_set_finite _)) (conj (eq_refl _) h8)). assert (h10:list_to_set_finite l = add_downward_closed (list_to_set l) a h0). apply proof_irrelevance. rewrite h10 in h1. rewrite h1. generalize h5. rewrite Im_add. intro h2. rewrite im_set_bag_add'. rewrite plus_bag_nat_add. rewrite <- h4. f_equal. f_equal. f_equal. apply proof_irrelevance. right. constructor. intro h3. inversion h3 as [x h11 q hq]. subst. apply hi in hq. subst. rewrite <- list_to_set_in_iff in h11. contradiction. rewrite <- list_to_set_in_iff; auto. Qed. Lemma plus_bag_nat_im_set_bag_eq_f_hetero_dep : forall {T U:Type} (A:Ensemble T) (A':Ensemble U) (pf:Finite A) (pf':Finite A') (f:T->nat) (f':U->nat), (exists g:fun_in_ens A U, inj_in_ens g /\ A' = im_in_ens g /\ forall (x:T) (pfa:Inn A x), f x = f' (g x pfa)) -> plus_bag_nat (im_set_bag nat_eq_dec A pf f) = plus_bag_nat (im_set_bag nat_eq_dec A' pf' f'). intros T U A A' h1 h2 f f' h3. destruct h3 as [g [h0 [h3 h4]]]. subst. pose proof (finite_set_list_no_dup _ h1) as hf. destruct hf as [l ha]. revert h2. revert h4. revert ha. revert f f' h1 h0. revert g. revert A. induction l as [|a l h1]; simpl. intros A g h1 f h0 h2 h3 h4 hi. destruct h3 as [h3 h5]. subst. rewrite im_set_bag_empty'. rewrite im_set_bag_empty''; auto. rewrite im_in_ens_empty; auto. intros A g f f' h3 hinj h5 h6. destruct h5 as [h5 h7]. subst. pose proof (no_dup_cons _ _ h7) as h8. pose proof (no_dup_cons_nin _ _ h7) as h9. rewrite im_set_bag_add'. rewrite plus_bag_nat_add. assert (h10:forall x (pf:Inn (list_to_set l) x), Inn (im_in_ens g) (fun_from_add g x pf)). intros x h11. econstructor. unfold fun_from_add. reflexivity. specialize (h1 (list_to_set l) (fun_from_add g) f f' (list_to_set_finite _) (inj_fun_from_add_compat _ hinj)). hfirst h1. split; auto. specialize (h1 hf). hfirst h1. intros x h11; simpl. unfold fun_from_add. apply h6. intro hfin. specialize (h1 hf0 (finite_im_in_ens _ (list_to_set_finite _))). erewrite im_set_bag_functional2. rewrite h1 at 1. pose proof (im_in_ens_add g) as h11. hfirst h11. contradict h9. rewrite list_to_set_in_iff; auto. specialize (h11 hf1). pose proof hfin as hfin'. rewrite h11 in hfin'. pose proof (im_set_bag_functional1 nat_eq_dec _ _ hfin hfin' f' h11) as h12. simpl in h12. rewrite h12. rewrite im_set_bag_add'. rewrite plus_bag_nat_add. repeat f_equal. apply proof_irrelevance. specialize (h6 _ (Add_intro2 _ _ _)); auto. intro hy. inversion hy; subst. unfold fun_from_add in H. apply hinj in H; subst. contradiction. reflexivity. reflexivity. rewrite <- list_to_set_in_iff; auto. Qed. Lemma plus_bag_nat_im_in_ens_bag_eq_f_hetero : forall {T U:Type} (A:Ensemble T) (A':Ensemble U) (pff:Finite A) (pff':Finite A') (f:fun_in_ens A nat) (f':fun_in_ens A' nat), (exists (g:fun_in_ens A U) (pfe:A' = im_in_ens g), inj_in_ens g /\ forall x (pfx:Inn A x), let pfgx := in_eq_set _ _ (eq_sym pfe) (g x pfx) (im_in_ens_intro g _ pfx _ eq_refl) in f x pfx = f' (g x pfx) pfgx) -> plus_bag_nat (im_in_ens_bag nat_eq_dec A pff f) = plus_bag_nat (im_in_ens_bag nat_eq_dec A' pff' f'). intros T U A A' h1 h2 f f' h3. destruct h3 as [g [h0 [h3 h4]]]. subst. pose proof (finite_set_list_no_dup _ h1) as hf. destruct hf as [l ha]. revert h2 h3. revert h4. revert ha. revert f f' h1. revert g. revert A. induction l as [|a l h1]; simpl. intros A g h1 f h0 h2 h3 h4 hi. destruct h2 as [h2 h5]. subst. rewrite im_in_ens_bag_empty'. rewrite im_in_ens_bag_empty''; auto. rewrite im_in_ens_empty; auto. intros A g f f' hi h4 h5 h6 hinj. destruct h4 as [h4 h7]. subst. pose proof (no_dup_cons _ _ h7) as h8. pose proof (no_dup_cons_nin _ _ h7) as h9. rewrite im_in_ens_bag_add'. rewrite plus_bag_nat_add. assert (h10:forall x (pf:Inn (list_to_set l) x), Inn (im_in_ens g) (fun_from_add g x pf)). intros x h11. econstructor. unfold fun_from_add. reflexivity. specialize (h1 (list_to_set l) (fun_from_add g) (fun_from_add f) (fun_from_add_im f' hinj) (list_to_set_finite _)). hfirst h1. split; auto. specialize (h1 hf). hfirst h1. intros x h11; simpl. erewrite fun_from_add_im_functional. rewrite fun_from_add_im_compat. specialize (h5 x (Add_intro1 _ _ _ _ h11)). unfold fun_from_add. rewrite h5. f_equal. apply proof_irrelevance. reflexivity. specialize (h1 hf0 (finite_im_in_ens _ (list_to_set_finite _)) (inj_fun_from_add_compat _ hinj)). erewrite im_in_ens_bag_functional2. rewrite h1 at 1. pose proof (im_in_ens_add g) as h11. hfirst h11. contradict h9. rewrite list_to_set_in_iff; auto. specialize (h11 hf1). pose proof h6 as h6'. rewrite h11 in h6'. pose proof (im_in_ens_bag_functional1 nat_eq_dec _ _ h6 h6' f' h11) as h12. simpl in h12. rewrite h12. rewrite im_in_ens_bag_add'. rewrite plus_bag_nat_add. repeat f_equal. apply proof_irrelevance. apply functional_extensionality_dep. intro y. apply functional_extensionality. intro hy. destruct hy as [x h13]; subst. specialize (hf0 _ h13). simpl in hf0. erewrite fun_from_add_im_functional. rewrite <- hf0 at 1. specialize (h5 x (Add_intro1 _ _ _ _ h13)). unfold fun_from_add. erewrite f_equal_dep. rewrite h5 at 1. rewrite fun_in_ens_transfer_compat. repeat f_equal. apply proof_irrelevance. reflexivity. reflexivity. specialize (h5 _ (Add_intro2 _ _ _)). rewrite h5. rewrite fun_in_ens_transfer_compat. repeat f_equal. apply proof_irrelevance. contradict hf1. inversion hf1; subst. unfold fun_from_add in H. apply hinj in H; subst; auto. rewrite <- list_to_set_in_iff; auto. Qed. bool2-v0-3/PartialAlgebrasBasics.v0000644000175000017500000034513713575574055017730 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2019, Daniel Wyckoff.*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- version 8.4 pl4 Coq.*) Require Import NaryFunctions. Require Import Description. Require Import ProofIrrelevance. Require Import FunctionalExtensionality. Require Import Equality. Require Import SetUtilities. Require Import TypeUtilities. Require Import TypeUtilities2. Require Import SetUtilities2. Require Import FunctionProperties. Require Import FunctionProperties2. Require Import FunctionProperties3. Require Import FunctionAlgebrasBasics. Require Import EqDec. Require Import ListUtilities. Require Import ArithUtilities. Require Import LogicUtilities. (*This implementation of many-sorted partial algebras is based on Reichel's "Structural Induction on Partial Algebras" (or equivalently, the second edition, "Initial Computability, Algebraic Specifications, and Partial Algebras."*) (*In anticipation of homomorphisms and subalgebras, we will name the operator symbols after a number in an initial segment of the natural numbers to enforce a universal interface. We will always use a finite number of operations on our algebras, which will be crucial when we induct over term algebras. It's a bit cumbersome, for the same reason that arbitrary S-sorted systems of variables (v:X->S) are so convenient. Term algebras can tolerate infinite systems of variables, but not infinite operations. *) Definition signature S (num_ops:nat) : Type := seg_fun num_ops (all_tuples S * S). Definition signature_seg S num_ops : Type := signature (S @) num_ops. Definition unary_sig {S num_ops} (sig:signature S num_ops) : Prop := forall op (pfop:op < num_ops), exists (pfe:projT1 (fst (sig op pfop)) = 1), projT2 (fst (sig op pfop)) = transfer_nprod_r pfe (snd (sig op pfop), tt). Lemma projT1_fst_sig0_impl : forall {S num_ops} (sig:signature S num_ops) op (pfop:op < num_ops) (pfe:projT1 (fst (sig op pfop)) = 0), (projT2 (fst (sig op pfop))) = transfer_dep_r pfe tt. intros. apply (transfer_dep_inj pfe). apply nprod0_eq. Qed. Definition op_to_fun {S} (A : S -> Type) (op:all_tuples S * S) : Type := tuple_prod_all A (fst op) -> A (snd op). Definition op_to_fun' {S num_ops} (A:S->Type) (sig:signature S num_ops) i (pfi:i < num_ops) : Type := op_to_fun A (sig i pfi). Definition op_to_fun_seg {S} (A :S ---> Type) (op:all_tuples (S @) * (S @)) : Type := op_to_fun A op. Definition op_to_fun_seg' {S num_ops} (A :S ---> Type) (sig:signature (S @) num_ops) i (pfi:i < num_ops) : Type := op_to_fun_seg A (sig i pfi). Definition op_to_fun_ens {S T} (A:S->Ensemble T) (op:all_tuples S * S) : Type := sig_set (tuple_prod_ens A (projT2 (fst op))) -> sig_set (A (snd op)). Definition op_to_fun_ens' {S T num_ops} (A:S->Ensemble T) (sig:signature S num_ops) i (pfi:i < num_ops) : Type := op_to_fun_ens A (sig i pfi). Definition op_to_fun_dep {S} {U:S->Type} (A:forall s:S, Ensemble (U s)) (op:all_tuples S * S) : Type := sig_set (tuple_prod_ens_dep A (projT2 (fst op))) -> sig_set (A (snd op)). Definition op_to_fun_dep' {S num_ops} {U:S->Type} (sig:signature S num_ops) (A:forall s:S, Ensemble (U s)) i (pfi:i < num_ops) : Type := op_to_fun_dep A (sig i pfi). Definition op_to_fun_ens_to_dep {S T} {A:S->Ensemble T} {op:all_tuples S * S} (f:op_to_fun_ens A op) : op_to_fun_dep A op (U := fun _:S => T) := fun (t:sig_set (tuple_prod_ens_dep A (projT2 (fst op)))) => let pfi := in_tuple_prod_ens_dep_const_compat A (proj1_sig t) (proj2_sig t) in f (exist (fun c => Inn (tuple_prod_ens A (projT2 (fst op))) c) _ pfi). Definition op_to_fun_ens_to_dep' {S T num_ops} (sig:signature S num_ops) (A:S->Ensemble T) i (pfi:i < num_ops) (f:op_to_fun_ens A (sig i pfi)) : op_to_fun_dep' sig A i pfi := op_to_fun_ens_to_dep f. Definition op_to_fun_ens_seg {S T} (A:S ---> Ensemble T) (op:all_tuples (S @) * S @) : Type := op_to_fun_ens A op. Definition op_to_fun_ens_seg' {S T num_ops} (sig:signature (S @) num_ops) (A:S ---> Ensemble T) i (pfi:i < num_ops) : Type := op_to_fun_ens_seg A (sig i pfi). Definition op_to_fun_seg_prod_type {n} (f:seg_fun n Type) (U:Type) : Type := seg_fun_prod_type f -> U. Definition op_to_pfun {S} (A : S -> Type) (op:all_tuples S * S) (dom:Ensemble (tuple_prod_all A (fst op))) : Type := fun_in_ens dom (A (snd op)). Definition op_to_pfun' {S num_ops} (sig:signature S num_ops) (A : S -> Type) i (pfi:i < num_ops) (dom:Ensemble (tuple_prod_all A (fst (sig i pfi)))) : Type := op_to_pfun A (sig i pfi) dom. Definition op_to_pfun_seg {S} (A : S ---> Type) (op:all_tuples (S @) * S @) (dom:Ensemble (tuple_prod_all A (fst op))) : Type := op_to_pfun A op dom. Definition op_to_pfun_seg' {S num_ops} (sig:signature (S @) num_ops) (A : S ---> Type) i (pfi:i < num_ops) (dom:Ensemble (tuple_prod_all A (fst (sig i pfi)))) : Type := op_to_pfun_seg A (sig i pfi) dom. Definition op_to_pfun_ens {S T} (A:S->Ensemble T) (op:all_tuples S * S) (dom:Ensemble (T ^ projT1 (fst op))) (pfd: Included dom (tuple_prod_ens A (projT2 (fst op)))) : Type := fun_in_ens dom (sig_set (A (snd op))). Definition op_to_pfun_ens' {S T num_ops} (A:S->Ensemble T) (sig:signature S num_ops) op (pfop:op < num_ops) (dom:Ensemble (T ^ (projT1 (fst (sig op pfop))))) (pfd: Included dom (tuple_prod_ens A (projT2 (fst (sig op pfop))))) : Type := op_to_pfun_ens A (sig op pfop) dom pfd. Definition assment {S X} (A:S->Type) (v:X->S) : Type := forall x:X, A (v x). Definition assment_seg {X S} (A:S ---> Type) (v:X ----> S) : Type := assment A v. Definition assment_ens {S U X} (A:S->Ensemble U) (v:X->S) : X -> Ensemble U := fun x => A (v x). Definition assment_ens_sig {S U X} (A:S->Ensemble U) (v:X->S) : Type := forall x:X, sig_set (A (v x)). Definition map_to_assment_ens_sig {S U X n} {A:S->Ensemble U} {v:X->S} (x:X^n) (a:assment_ens_sig A v) : U^n := map_nprod (fun j (pfj:j < n) => proj1_sig (a (nth_prod j pfj x))). Definition assment_ens_sig_w' {S U X n} {A:S->Ensemble U} {v:X->S} (a:assment_ens_sig A v) (w:X^n) : U^n := map_nprod (fun i (pfi:i < n) => (proj1_sig (a (nth_prod i pfi w)))). Lemma in_assment_ens_sig_w' : forall {S U X n} {A:S->Ensemble U} {v:X->S} (a:assment_ens_sig A v) (w:X^n), Inn (tuple_prod_ens_sys A v w) (assment_ens_sig_w' a w). intros S U X n A v a w. constructor. intros i hi. unfold tuple_prod_sys, assment_ens_sig_w'. rewrite nth_map_nprod. apply proj2_sig. Qed. Lemma in_assment_ens_sig_w_nth' : forall {S U X n} {A:S->Ensemble U} {v:X->S} (a:assment_ens_sig A v) (w:X^n) i (pfi:i < n), Inn (A (v (nth_prod i pfi w))) (nth_prod i pfi (assment_ens_sig_w' a w)). intros S U X n A v a w i hi. pose proof (in_assment_ens_sig_w' a w) as h1. inversion h1 as [h2]. specialize (h2 i hi). unfold tuple_prod_sys in h2; auto. Qed. Inductive in_assment_ens_sig {S U X} {A:S->Ensemble U} {v:X->S} (u:U) (a:assment_ens_sig A v) : Prop := assment_ens_sig_in_intro : forall x:X, u = proj1_sig (a x) -> in_assment_ens_sig u a. Lemma in_assment_ens_sig_compat : forall {S U X} {A:S->Ensemble U} {v:X->S} (a:assment_ens_sig A v) x, Inn (A (v x)) (proj1_sig (a x)). intros; apply proj2_sig. Qed. Definition w_in_assment_ens_sig {S U X n} {A:S->Ensemble U} {v:X->S} (u:U^n) (a:assment_ens_sig A v) : Prop := forall i (pfi:i < n), in_assment_ens_sig (nth_prod i pfi u) a. Lemma w_in_assment_ens_sig_compat : forall {S U X n} {A:S->Ensemble U} {v:X->S} (u:U^n) (a:assment_ens_sig A v) (x:X^n), Inn (tuple_prod_ens_sys A v x) (map_nprod (fun j (pfj:j < n) => proj1_sig (a (nth_prod j pfj x)))). intros S U X n A v u a x. constructor. intros j hj. unfold tuple_prod_sys. rewrite nth_map_nprod. apply proj2_sig. Qed. Definition in_assment_ens_sig_at_op {S U X num_ops} {sig:signature S num_ops} {A:S->Ensemble U} {v:X->S} op (pfop:op < num_ops) (u:U^projT1 (fst (sig op pfop))) (a:assment_ens_sig A v) : Prop := w_in_assment_ens_sig u a. Definition assment_dep {X S} {U:S->Type} (A:forall s:S, Ensemble (U s)) (v:X->S) : Type := forall x:X, sig_set (A (v x)). Definition assment_ens_sig_to {S U X} {A:S->Ensemble U} {v:X->S} (a:assment_ens_sig A v) : @assment_dep X S (fun _=> U) (fun s => (A s)) v := a. Definition assment_dep_const_to {S U X} {A:S->Ensemble U} {v:X->S} (a:@assment_dep X S (fun _=> U) (fun s => (A s)) v) : assment_ens_sig A v := a. Definition assment_w {S n} (A:S -> Type) (w:S^n) : Type := tuple_prod A w. Definition assment_w_ens {S U n} (A:S -> Ensemble U) (w:S^n) : Type := sig_set (tuple_prod_ens A w). Definition assment_w_dep {S n} {U:S -> Type} (A:forall s:S, Ensemble (U s)) (w:S^n) : Type := sig_set (tuple_prod_ens_dep A w). Fixpoint in_pre_w {S X n} : forall (w:S^n) (v:X->S), Prop := match n with | 0 => fun _ _ => True | S n' => fun w v => let (fw, sw) := w in in_pre fw v /\ (in_pre_w sw v) end. Definition pre_sys_aux {S X} (v:X->S) (s:S) : Ensemble X := [x:X | v x = s]. Lemma assment_w_tt : forall {S} (A:S -> Type), assment_w A tt (n:=0) = unit. auto. Qed. Lemma assment_w_tt_ : forall {S} (A:S -> Type) (w:S^0), assment_w A tt (n:=0) = unit. auto. Qed. Lemma assment_w_tt' : forall {S U} (A:S -> Ensemble U) (w:S^0), assment_w_ens A w = sig_set (Singleton tt). unfold assment_w_ens; intros; rewrite tuple_prod_ens0; auto. Qed. Lemma assment_w_dep_tt : forall {S} {U:S -> Type} (A:forall s:S, Ensemble (U s)) (w:S^0), assment_w_dep A w = sig_set (Singleton tt). unfold assment_w_dep; intros; rewrite tuple_prod_ens_dep0; auto. Qed. Lemma assment_w_S' : forall {T U n} (A:T -> Ensemble U) (w:T^(S n)), assment_w_ens A w = sig_set [u:U^(S n) | Inn (A (fst w)) (fst u) /\ Inn (tuple_prod_ens A (snd w)) (snd u)]. unfold assment_w_ens; intros; rewrite tuple_prod_ens_S; destruct w; auto. Qed. Lemma assment_w_dep_S : forall {T n} {U:T -> Type} (A:forall t:T, Ensemble (U t)) (w:T^(S n)), assment_w_dep A w = sig_set [t:tuple_prod U w | Inn (A (fst w)) (fst t) /\ Inn (tuple_prod_ens_dep A (snd w)) (snd t)]. unfold assment_w_dep; intros; rewrite tuple_prod_ens_dep_S; destruct w; auto. Qed. Lemma in_assment_w_dep_S : forall {T n} {U:T -> Type} {A:forall t:T, Ensemble (U t)} {w:T^(S n)} (p:assment_w_dep A w), Inn [t:tuple_prod U w | Inn (A (fst w)) (fst t) /\ Inn (tuple_prod_ens_dep A (snd w)) (snd t)] (proj1_sig p). intros ? ? ? ? w p ; constructor. pose proof (tuple_prod_ens_dep_S A w) as h1. pose proof (proj2_sig p) as h2; simpl in h2. rewrite h1 in h2 at 1. destruct h2; auto. Qed. Lemma in_assment_w_dep_S' : forall {T n} {U:T -> Type} {A:forall t:T, Ensemble (U t)} {w:T^(S n)} (p:sig_set [t:tuple_prod U w | Inn (A (fst w)) (fst t) /\ Inn (tuple_prod_ens_dep A (snd w)) (snd t)]), Inn (tuple_prod_ens_dep A w) (proj1_sig p). intros ? ? ? ? w p. destruct p as [p hp]. destruct hp as [hp]; simpl. destruct hp as [h1 h2]. constructor. intros i hi; destruct i as [|i]; simpl; auto. destruct h2; auto. Qed. Lemma in_assment_w_dep_snd : forall {T n} {U:T -> Type} {A:forall t:T, Ensemble (U t)} {w:T^(S n)} (a:assment_w_dep A w), Inn (tuple_prod_ens_dep A (snd w)) (snd (proj1_sig a)). intros ? ? ? ? ? a. pose proof (in_assment_w_dep_S a) as h1. destruct h1 as [h1]; destruct h1; auto. Qed. Definition assment_w_dep_to_S {T n} {U:T -> Type} {A:forall t:T, Ensemble (U t)} {w:T^(S n)} (a:assment_w_dep A w) : assment_w_dep A (snd w) := exist (fun c => Inn (tuple_prod_ens_dep A (snd w)) c) (snd (proj1_sig a)) (in_assment_w_dep_snd a). Lemma transfer_r_assment_w_dep_S : forall {T n} {U:T->Type} {A:forall t, Ensemble (U t)} {w:T^(S n)} (p:sig_set [t:tuple_prod U w | Inn (A (fst w)) (fst t) /\ Inn (tuple_prod_ens_dep A (snd w)) (snd t)]), transfer_r (assment_w_dep_S A w) p = exist (fun x : tuple_prod U w => (Inn (tuple_prod_ens_dep A w) x)) (proj1_sig p) (in_assment_w_dep_S' _). intros T n U A w p. rewrite transfer_r_eq_iff. apply f_equal_sig_set. rewrite <- (tuple_prod_ens_dep_S); auto; simpl. unfold proj1_sig; auto. Qed. Lemma transfer_assment_w_dep_S : forall {T n} {U:T->Type} {A:forall t, Ensemble (U t)} {w:T^(S n)} (p:sig_set (tuple_prod_ens_dep A w)), transfer (assment_w_dep_S A w) p = exist (fun x : tuple_prod U w => Inn [t:tuple_prod U w | Inn (A (fst w)) (fst t) /\ Inn (tuple_prod_ens_dep A (snd w)) (snd t)] x) (proj1_sig p) (in_assment_w_dep_S _). intros T n U A w p. rewrite transfer_eq_iff. apply f_equal_sig_set. rewrite <- (tuple_prod_ens_dep_S); auto; simpl. unfold proj1_sig; auto. Qed. Lemma fst_transfer_assment_w_dep_S: forall {T n} {U:T->Type} {A:forall t, Ensemble (U t)} (w:T^(S n)) (p:sig_set (tuple_prod_ens_dep A w)), fst (proj1_sig (transfer (assment_w_dep_S A w) p)) = fst (proj1_sig p). intros; rewrite transfer_assment_w_dep_S; simpl; auto. Qed. Lemma fst_transfer_r_assment_w_dep_S : forall {T n} {U:T->Type} {A:forall t, Ensemble (U t)} (w:T^(S n)) (p:sig_set [t:tuple_prod U w | Inn (A (fst w)) (fst t) /\ Inn (tuple_prod_ens_dep A (snd w)) (snd t)]), fst (proj1_sig (transfer_r (assment_w_dep_S A w) p)) = fst (proj1_sig p). intros; rewrite transfer_r_assment_w_dep_S; simpl; auto. Qed. Definition seg_fun_prod_ens_to_assment {S n} {U:S->Type} (A:forall s, Ensemble (U s)) (x:S^n) (f:seg_fun_depT_nth U x) (pfi:forall i (pfi:i < n), Inn (A (nth_prod i pfi x)) (f i pfi)) : assment_w_dep A x := exist (fun c => Inn (tuple_prod_ens_dep A x) c) (seg_fun_depT_nth_to_tuple_prod x f) (in_seg_fun_depT_nth_to_tuple_prod _ _ _ pfi). Definition assment_w_seg {S n} (A:S ---> Type) (w:(S @)^n) : Type := assment_w A w. Lemma assment_w_seg_tt : forall {S} (A:S ---> Type), assment_w_seg A tt (n:=0) = unit. auto. Qed. Lemma assment_w_seg_tt_ : forall {S} (A:S ---> Type) (w:(S @)^0), assment_w_seg A w (n:=0) = unit. auto. Qed. Fixpoint map_assment_w {S n} : forall {A B:S -> Type} {w:S^n} (f:Sfun A B) (p:assment_w A w), assment_w B w := match n with | 0 => fun A B w f p => transfer_r (assment_w_tt_ B w) tt | S n' => fun A B w f p => let (ft, st) := w in let (fp, sp) := p in let f' := map_assment_w f in (f _ fp, f' sp) end. Fixpoint map_assment_w_dep {S n} {U V:S -> Type} : forall {A:forall s:S, Ensemble (U s)} {B:forall s:S, Ensemble (V s)} {w:S^n} (f:Sfun_dep A B) (p:assment_w_dep A w), assment_w_dep B w := match n with | 0 => fun A B w f p => transfer_r (assment_w_dep_tt B w) (exist (fun c => Inn (Singleton tt) c) tt (In_singleton _ tt)) | S n' => fun A B w f p => let pfe := assment_w_dep_S A w in let pfe' := assment_w_dep_S B w in let p' := transfer pfe p in let p'' := proj1_sig p' in let hp'' := proj2_sig p' in let pf1 := sat_in1 _ _ _ hp'' in let pf2 := sat_in2 _ _ _ hp'' in let ff := f _ _ pf1 in let a := map_assment_w_dep f (exist (fun c => Inn (tuple_prod_ens_dep A (snd w)) c) _ pf2) (w := snd w) in let pff := intro_characteristic_sat (fun t : tuple_prod V w => Inn (B (fst w)) (fst t) /\ Inn (tuple_prod_ens_dep B (snd w)) (snd t)) (proj1_sig ff, proj1_sig a) (conj (proj2_sig ff) (proj2_sig a)) in let x := exist _ _ pff in let x' := transfer_r pfe' x in x' end. Lemma proj1_sig_map_assment_w_dep_eq_iff : forall {T n} {U V:T->Type} {A:forall t:T, Ensemble (U t)} {B:forall t:T, Ensemble (V t)} {w:T^n} (a:assment_w_dep A w) (f:forall (t:T) (u:U t), Inn (A t) u -> {v : V t | Inn (B t) v}) (t:tuple_prod V w), proj1_sig (map_assment_w_dep f a) = t <-> (forall j (pfj: j < n), let pfaj := iff1 (in_tuple_prod_ens_dep_nth_iff A (proj1_sig a)) (proj2_sig a) j pfj in proj1_sig (f _ _ pfaj) = nth_tuple_prod t j pfj). intros T n U V A B w a f t; simpl; split. intros ? j hj; subst. destruct a as [t ht]. destruct ht as [ht]. generalize dependent U; generalize dependent V; generalize dependent j; revert w. induction n as [|n ih]; simpl. intros; omega. simpl in ih. intro w. destruct w as [fw sw]. intro j; destruct j as [|j]; simpl. intros hj V B U A t. destruct t as [ft st]. intros ht f; simpl. symmetry. gterml. redterm0 x. redterm0 c. redterm0 c0. fold c1. rewrite (fst_transfer_r_assment_w_dep_S (fw, sw) c1) at 1. unfold c, c1; simpl. repeat f_equal. apply f_equal_dep. gterml. redterm0 x0. redterm0 c2. redterm0 c3. fold c4. rewrite (fst_transfer_assment_w_dep_S (fw, sw) c4) at 1. unfold c4; simpl; auto. intros hj V B U A t ht f. specialize (ih sw j (lt_S_n _ _ hj) V B U A (snd t)). hfirst ih. intros i hi. specialize (ht (S i) (lt_n_S _ _ hi)). simpl in ht. pose proof (proof_irrelevance _ hi (lt_S_n _ _ (lt_n_S _ _ hi))) as h1. rewrite <- h1 in ht; auto. specialize (ih hf f). terml ih. redterm0 x. redterm0 c. fold c0 in ih. gterml. redterm0 x0. redterm0 c1. fold c2. pose proof (proof_irrelevance _ c0 c2) as h2. rewrite <- h2; auto. unfold c0. pose proof (transfer_r_sig_set_eq _ _ (tuple_prod_ens_dep_S B (fw, sw)) (assment_w_dep_S B (fw, sw))) as h3. gtermr. redterm0 y. redterm2 y. redterm0 c4. redterm0 c5. redterm0 c6. specialize (h3 c7). pose proof (proj1_sig_transfer_r_sig_set c7 (tuple_prod_ens_dep_S B (fw, sw)) (assment_w_dep_S B (fw, sw))) as h4. pose proof (f_equal (@snd _ _ ) h4) as h5. terml h5. fold x1 c7. pose proof (nth_tuple_prod_functional1 x1 (snd (proj1_sig c7)) j j (lt_S_n _ _ hj) h5 (eq_refl _)) as h6. unfold x1 in h6. rewrite h6 at 1. simpl. gtermr. redterm1 y0. pose proof (proof_irrelevance _ (eq_lt_trans _ _ _ eq_refl (lt_S_n _ _ hj)) (lt_S_n _ _ hj)) as h7. rewrite h7, transfer_dep_eq_refl. fold c0. rewrite ih at 1. repeat f_equal; apply proj1_sig_injective; simpl. unfold c7 in h4; simpl in h4. pose proof (f_equal (@snd _ _) h4) as h8; simpl in h8. termr h8. gtermr. fold y1 in h3. redterm0 y2. fold c9. f_equal. unfold c9. redterm0 c9. redterm0 c10. fold c11. rewrite (proj1_sig_transfer_sig_set c11 (tuple_prod_ens_dep_S A (fw, sw)) (assment_w_dep_S A (fw, sw))) at 1. unfold c11; simpl; auto. intro h1. generalize dependent U; generalize dependent V; revert w. induction n as [|n ih]; simpl. intros; apply unit_eq. intro w; destruct w as [fw sw]; simpl. intros V B t. destruct t as [ft st]. intros U A a f h1. specialize (ih sw V B st U A (assment_w_dep_to_S a) f). hfirst ih. intros j hj. specialize (h1 (S j) (lt_n_S _ _ hj)); simpl in h1. pose proof (proof_irrelevance _ hj (lt_S_n _ _ (lt_n_S _ _ hj))) as h2. rewrite h2, <- h1; auto. repeat f_equal; apply proof_irrelevance. specialize (ih hf); subst. gterml. redterm0 x. redterm0 x. pose proof (proj1_sig_transfer_sig_set c0 (tuple_prod_ens_dep_S B (fw, sw)) (assment_w_dep_S B (fw, sw))) as h8. unfold c0 in h8. rewrite transfer_undoes_transfer_r in h8; simpl in h8. rewrite <- h8. f_equal. specialize (h1 0 (lt_0_Sn _)); simpl in h1. rewrite <- h1. f_equal. apply f_equal_dep. repeat f_equal; auto. apply (proj1_sig_transfer_sig_set a (tuple_prod_ens_dep_S A (fw, sw)) (assment_w_dep_S A (fw, sw))). repeat f_equal. unfold assment_w_dep_to_S. apply proj1_sig_injective; simpl. f_equal. apply (proj1_sig_transfer_sig_set a (tuple_prod_ens_dep_S A (fw, sw)) (assment_w_dep_S A (fw, sw))). Qed. Lemma nth_tuple_prod_map_assment_w_dep : forall {T n} {U:T->Type} {V:T->Type} {A:forall t:T, Ensemble (U t)} {B:forall t:T, Ensemble (V t)} {w:T^n} (a:assment_w_dep A w) (f:forall (t:T) (u:U t), Inn (A t) u -> {v:V t | Inn (B t) v}) j (pfj:j < n), nth_tuple_prod (proj1_sig (map_assment_w_dep f a)) j pfj = proj1_sig (f (nth_prod j pfj w) (nth_tuple_prod (proj1_sig a) j pfj) ((iff1 (in_tuple_prod_ens_dep_nth_iff A (proj1_sig a)) (proj2_sig a)) j pfj)). intros T n U V A B w a f j hj. pose proof (proj1_sig_map_assment_w_dep_eq_iff a f (proj1_sig (map_assment_w_dep f a))) as h1; simpl in h1. hl h1; auto. rewrite h1 in hl. specialize (hl j hj); auto. Qed. (*Total sigma-algebra*) (*Note the many sorts [A_sigma_t] are not necessarily finite. This is to account for things like the term algebra. The usual convention is a finite a number of operations, and a countable set of type variables, but I'm allowing a primitive, arbitrary type of type variables, which eventually in the development will be tamed.*) Record TAlgebra {S num_ops} (sig:signature S num_ops) : Type := { A_sigma_t : S -> Type; ops_sigma_t : forall i (pf:i < num_ops), op_to_fun' A_sigma_t sig i pf }. Record TAlgebra_ens {S num_ops} U (sig:signature S num_ops) : Type := { type_eq_decS_t' : type_eq_dec S; A_sigma_t' : S -> Ensemble U; ops_sigma_t' : forall i (pf:i < num_ops), op_to_fun_ens' A_sigma_t' sig i pf }. (*This is the central variant [_dep], which is the most universal and most primitive, and if it seems that the friction between type theory and set theory is too unnecessary here, consider inductively defined sets like [Term_algebra_ens] further down.*) Record TAlgebra_dep {S num_ops} (sig:signature S num_ops) : Type := { U_dep_t : S -> Type; type_eq_decS_t : type_eq_dec S; A_dep_t : forall (s:S), Ensemble (U_dep_t s); ops_dep_t : forall i (pf:i < num_ops), op_to_fun_dep' sig A_dep_t i pf }. Definition TAlgebra_to_dep {S U num_ops} {sig:signature S num_ops} (A:TAlgebra_ens U sig) : TAlgebra_dep sig := Build_TAlgebra_dep S num_ops sig (fun _:S => U) (type_eq_decS_t' _ _ A) (A_sigma_t' U sig A) (fun i (pfi:i < num_ops) => (op_to_fun_ens_to_dep' sig (A_sigma_t' U sig A) i pfi (ops_sigma_t' U sig A i pfi))). Lemma unary_pointwise : forall {S num_ops} (sig:signature S num_ops) (A:TAlgebra_dep sig) op (pfop:op < num_ops) (pfe:projT1 (fst (sig op pfop)) = 1), projT2 (fst (sig op pfop)) = transfer_nprod_r pfe (snd (sig op pfop), tt) -> pointwise_ind_eq (U_dep_t _ A) (U_dep_t _ A) (projT2 (fst (sig op pfop))) (transfer_nprod_r pfe (snd (sig op pfop), tt)). intros S num_ops sig A op hop he he'. red; intros i hi. pose proof (lt_eq_trans _ _ _ hi he) as h4. apply lt1 in h4; subst. repeat f_equal; auto. Qed. (*This predicate returns true if a given S-set generates a specific point in a Total Algebra.*) Inductive point_gen {S num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} (E:forall s:S, Ensemble (U_dep_t sig A s)) (pfincl:forall s:S, Included (E s) (A_dep_t sig A s)) : forall (s:S) (u:U_dep_t sig A s), Inn (A_dep_t sig A s) u -> Prop := | point_gen_point : forall (s:S) (u:U_dep_t sig A s) (pfin:Inn (E s) u), point_gen E pfincl s u (pfincl s _ pfin) | point_gen_nil_op : forall op (pfop:op < num_ops), let n := projT1 (fst (sig op pfop)) in let w := projT2 (fst (sig op pfop)) in let s := snd (sig op pfop) in forall (pfeq:projT1 (fst (sig op pfop)) = 0), let pfintt := in_transfer_tt_tuple_prod_ens_dep (U_dep_t sig A) (A_dep_t sig A) w pfeq in let u := ops_dep_t sig A op pfop (exist (fun c => Inn (tuple_prod_ens_dep (A_dep_t sig A) w) c) _ pfintt) in point_gen E pfincl s (proj1_sig u) (proj2_sig u) | point_gen_app : forall op (pfop:op < num_ops), let n := projT1 (fst (sig op pfop)) in let w := projT2 (fst (sig op pfop)) in let s := snd (sig op pfop) in forall (p:tuple_prod (U_dep_t sig A) w) (pfintpl:Inn (tuple_prod_ens_dep (A_dep_t sig A) w) p), (forall i (pfi:i < (projT1 (fst (sig op pfop)))), point_gen E pfincl (nth_prod i pfi (projT2 (fst (sig op pfop)))) (nth_tuple_prod p i pfi) (in_nth_tuple_prod_ens_dep pfintpl i pfi)) -> let u := ops_dep_t sig A op pfop (exist (fun c => Inn (tuple_prod_ens_dep (A_dep_t sig A) w) c) p pfintpl) in point_gen E pfincl s (proj1_sig u) (proj2_sig u). Inductive alg_gen {S num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} (E:forall s:S, Ensemble (U_dep_t sig A s)) (pfincl:forall s:S, Included (E s) (A_dep_t sig A s)) : Prop := | alg_gen_intro : (forall (s:S) (u:U_dep_t sig A s) (pfin:Inn (A_dep_t sig A s) u), point_gen E pfincl s u pfin) -> alg_gen E pfincl. (*This returns true for something like [in_nest _ . . . _ e (op3 (op2 (op1 e))) 3], where [Inn (E _) e]. So these are always built starting with the "generators", i.e. [nest_level] is an absolute nest level. It returns a count from inward to outward of the number of left (right) outward parenthesis bounding an element in [E]. Of course there can be many posibilities for constructor parameters, due to the algebra not necessarily being free, and no explicit term-rewriting system being specified. [E] is presumably a set of generators, although we don't require it, all the more so in light of the wide range of notions of free partial algebras. This will primarily be used for the Term Algebra, which is the free algebra on the class of total algebras of signature [sig], so we'll be fine. See below if you're looking for a relative parenthesis count starting from an arbitrary element (not just a "generator").*) Inductive in_nest_gen_n {S num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {E:forall s:S, Ensemble (U_dep_t sig A s)} (pfincle:forall s:S, Included (E s) (A_dep_t sig A s)) : forall (s t:S) (u:U_dep_t _ A s) (v:U_dep_t _ A t) (pfu:Inn (A_dep_t _ A s) u) (pfv:Inn (A_dep_t _ A t) v) (nest_level:nat), Prop := | in_nest0 : forall (s:S) (e:U_dep_t _ A s) (pfe:Inn (E s) e), let pfe' := pfincle s _ pfe in in_nest_gen_n pfincle s s e e pfe' pfe' 0 | in_nestS : forall (s t:S) (a:U_dep_t _ A s) (b:U_dep_t _ A t) (pfa:Inn (A_dep_t _ A s) a) (pfb:Inn (A_dep_t _ A t) b) n (pfinnest: in_nest_gen_n pfincle s t a b pfa pfb n) op (pfop:op < num_ops) (p:tuple_prod (U_dep_t _ A) (projT2 (fst (sig op pfop)))) (pfp:Inn (tuple_prod_ens_dep (A_dep_t _ A) (projT2 (fst (sig op pfop)))) p), in_tuple_prod (type_eq_decS_t _ A) b p (A := U_dep_t _ A) -> in_nest_gen_n pfincle s (snd (sig op pfop)) a (proj1_sig (ops_dep_t _ A op pfop (exist (fun c => Inn (tuple_prod_ens_dep _ _) c) p pfp))) pfa (proj2_sig _) (Datatypes.S n). Inductive in_nest_gen {S num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {E:forall s:S, Ensemble (U_dep_t sig A s)} (pfincle:forall s:S, Included (E s) (A_dep_t sig A s)) : forall (s t:S) (u:U_dep_t _ A s) (v:U_dep_t _ A t) (pfu:Inn (A_dep_t _ A s) u) (pfv:Inn (A_dep_t _ A t) v), Prop := | in_nest_gen_intro : forall (s t:S) (u:U_dep_t _ A s) (v:U_dep_t _ A t) (pfu:Inn (A_dep_t _ A s) u) (pfv:Inn (A_dep_t _ A t) v) (nest_level:nat), in_nest_gen_n pfincle s t u v pfu pfv (Datatypes.S nest_level) -> in_nest_gen pfincle s t u v pfu pfv. (*If this were used in anything but a free algebra (like the term algebra), it would be almost useless in current form. This is the relative form of [in_nest_n].*) Definition in_nest_n_rel {S num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} (s t:S) (u:U_dep_t _ A s) (v:U_dep_t _ A t) (pfu:Inn (A_dep_t _ A s) u) (pfv:Inn (A_dep_t _ A t) v) (nest_level:nat) : Prop := in_nest_gen_n (fun c:S => inclusion_reflexive (A_dep_t _ A c)) s t u v pfu pfv nest_level. Definition in_nest_rel {S num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} (s t:S) (u:U_dep_t _ A s) (v:U_dep_t _ A t) (pfu:Inn (A_dep_t _ A s) u) (pfv:Inn (A_dep_t _ A t) v) (nest_level:nat) : Prop := in_nest_gen (fun c:S => inclusion_reflexive (A_dep_t _ A c)) s t u v pfu pfv. Lemma in_nest_gen_S_iff : forall {S num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {E:forall s:S, Ensemble (U_dep_t sig A s)} {s t:S} (pfds:type_eq_dec S) (pfincle:forall s:S, Included (E s) (A_dep_t sig A s)) (u:U_dep_t _ A s) (v:U_dep_t _ A t) (pfu:Inn (A_dep_t _ A s) u) (pfv:Inn (A_dep_t _ A t) v) n, in_nest_gen_n pfincle s t u v pfu pfv (Datatypes.S n) <-> (exists op (pfop:op < num_ops) (p:tuple_prod (U_dep_t _ A) (projT2 (fst (sig op pfop)))) (pfp:Inn (tuple_prod_ens_dep _ _) p) (t':S) (w:U_dep_t sig A t') (pfw:Inn (A_dep_t sig A t') w) (pfe:t = snd (sig op pfop)), (in_tuple_prod pfds w p /\ v = transfer_dep_r pfe (proj1_sig (ops_dep_t _ A op pfop (exist _ p pfp)))) /\ ((n = 0 /\ Inn (E s) u /\ exists (pfe':s = t'), u = transfer_dep_r pfe' w) \/ (n > 0 /\ in_nest_gen_n pfincle s t' u w pfu pfw n))). intros S num_ops sig A E s t hds he u v hu hv n; split; intro h1. inversion h1; subst; clear h1. clear H6 H7 H5. apply existT_injective2 in H4; subst. apply existT_injective2 in H; subst. exists op, pfop, p, pfp, t0, b, pfb, (eq_refl _). rewrite transfer_dep_r_eq_refl. split; auto. split; auto. erewrite type_eq_dec_eq at 1. apply H3. destruct (zerop n) as [|hn]; subst. inversion pfinnest; subst. apply existT_injective2 in H9; subst. apply existT_injective2 in H7; subst. rewrite in_tuple_prod_iff in H3. destruct H3 as [i [hi [h9 h10]]]; subst. rewrite transfer_dep_eq_refl in h10; subst. left; auto. inversion pfinnest; subst. apply existT_injective2 in H4; subst; split; auto. apply existT_injective2 in H7; subst. split; auto. exists (eq_refl _). rewrite transfer_dep_r_eq_refl; auto. right. pose proof (proof_irrelevance _ hu pfa); subst; auto with arith. destruct h1 as [op [hop [p [hp [t' [w' [hw' [he' [h1 h2]]]]]]]]]; destruct h1 as [h1 h3]; subst. revert hv. rewrite transfer_dep_r_eq_refl. intro hv. destruct h2 as [h2|h2]. destruct h2 as [? [h2 [h3 h4]]]; subst. rewrite transfer_dep_r_eq_refl in h2. revert hu. rewrite transfer_dep_r_eq_refl. intro hu. pose proof (in_nest0 he t' w' h2) as h6. pose proof (in_nestS he _ _ _ _ _ _ _ h6 op hop p hp) as h7. hfirst h7. erewrite type_eq_dec_eq at 1. apply h1. specialize (h7 hf). erewrite f_equal_dep. apply h7. apply proof_irrelevance. destruct h2 as [h2 h3]. pose proof (in_nestS he _ _ _ _ _ _ _ h3 op hop p hp) as h7. hfirst h7. erewrite type_eq_dec_eq at 1. apply h1. specialize (h7 hf). erewrite f_equal_dep. apply h7. auto. Qed. Definition sub_elts {S num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {E:forall s:S, Ensemble (U_dep_t sig A s)} {pfincle:forall s:S, Included (E s) (A_dep_t sig A s)} (t:S) (v:U_dep_t _ A t) (pfv:Inn (A_dep_t _ A t) v) : forall (s:S), Ensemble (U_dep_t _ A s) := fun s => [u:U_dep_t _ A s | exists (pfu:Inn (A_dep_t _ A s) u), in_nest_gen pfincle s t u v pfu pfv]. Definition sub_elts_n {S num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {E:forall s:S, Ensemble (U_dep_t sig A s)} {pfincle:forall s:S, Included (E s) (A_dep_t sig A s)} (t:S) (v:U_dep_t _ A t) (pfv:Inn (A_dep_t _ A t) v) n : forall (s:S), Ensemble (U_dep_t _ A s) := fun s => [u:U_dep_t _ A s | exists (pfu:Inn (A_dep_t _ A s) u), in_nest_gen_n pfincle s t u v pfu pfv n]. (*Total homomorphism between two total algebras.*) Inductive Thomo {S num_ops} {sig:signature S num_ops} {A B:TAlgebra sig} (f:Sfun (A_sigma_t sig A) (A_sigma_t sig B)) : Prop := Thomo_intro : (forall i (pfi:i < num_ops), let w := fst (sig i pfi) in let s := snd (sig i pfi) in let opA := ops_sigma_t sig A i pfi in let opB := ops_sigma_t sig B i pfi in forall (a:assment_w (A_sigma_t _ A) (projT2 w)), let f' := map_assment_w f in f _ (opA a) = opB (f' a)) -> Thomo f. (*Dependent-Ensemble type form -- the most general, and metaphysically pure.*) Inductive Thomo_dep {S num_ops} {sig:signature S num_ops} {A B:TAlgebra_dep sig} (f:Sfun_dep (A_dep_t sig A) (A_dep_t sig B)) : Prop := Thomo_dep_intro : (forall i (pfi:i < num_ops), let w := fst (sig i pfi) in let s := snd (sig i pfi) in let opA := ops_dep_t sig A i pfi in let opB := ops_dep_t sig B i pfi in forall (a:assment_w_dep (A_dep_t _ A) (projT2 w)), let opa := proj1_sig (opA a) in let pfopa := proj2_sig (opA a) in let f' := map_assment_w_dep f in f _ opa pfopa = opB (f' a)) -> Thomo_dep f. Lemma homos_eq_on_generators : forall {S num_ops} {sig:signature S num_ops} {A B:TAlgebra_dep sig} (E:forall s:S, Ensemble (U_dep_t sig A s)) (pfincl:forall s:S, Included (E s) (A_dep_t sig A s)) (f g:Sfun_dep (A_dep_t sig A) (A_dep_t sig B)), Thomo_dep f -> Thomo_dep g -> alg_gen E pfincl -> (forall (s:S) (u:U_dep_t sig A s) (pfin:Inn (E s) u), proj1_sig (f s u (pfincl _ _ pfin)) = proj1_sig (g s u (pfincl _ _ pfin))) -> (forall op (pfop:op < num_ops), let n := projT1 (fst (sig op pfop)) in let w := projT2 (fst (sig op pfop)) in let s := snd (sig op pfop) in forall (pfeq:projT1 (fst (sig op pfop)) = 0), let pfintt := in_transfer_tt_tuple_prod_ens_dep (U_dep_t sig A) (A_dep_t sig A) w pfeq in let u := ops_dep_t sig A op pfop (exist (fun c => Inn (tuple_prod_ens_dep (A_dep_t sig A) w) c) _ pfintt) in f s (proj1_sig u) (proj2_sig u) = g s (proj1_sig u) (proj2_sig u)) -> f = g. intros S num_ops sig A B E hincl f g h1 h2 h3 h4 hne. destruct h1 as [h1], h2 as [h2]. apply functional_extensionality_dep. intro s. apply functional_extensionality_dep. intro u. apply functional_extensionality_dep. intro hu. destruct h3 as [h3]. specialize (h3 s u hu). induction h3. apply proj1_sig_injective; auto. apply h4. specialize (h1 op pfop). specialize (h2 op pfop). simpl in h1, h2. specialize (hne op pfop pfeq); simpl in hne. pose proof h4 as h4'. specialize (h4 s (proj1_sig u)). specialize (h1 (exist (fun c => Inn (tuple_prod_ens_dep (A_dep_t sig A) w) c) _ pfintt)). terml hne; terml h1. assert (he:x = x0). unfold x, x0. repeat f_equal. fold x0 in h1, hne. fold x in hne; fold x0 in h1. termr hne; gtermr. assert (he':y = y0). auto. fold y0. rewrite <- he'. fold y in hne. rewrite <- hne. unfold x, s,u. repeat f_equal. specialize (h1 op pfop); specialize (h2 op pfop); simpl in h1, h2. specialize (h1 (exist (fun c => Inn (tuple_prod_ens_dep (A_dep_t sig A) w) c) _ pfintpl)). specialize (h2 (exist (fun c => Inn (tuple_prod_ens_dep (A_dep_t sig A) w) c) _ pfintpl)). terml h2. redterm1 x. unfold u. unfold s, w. unfold s, w in p, pfintpl, h1, h2, H, H0. clear c x u s n w. rewrite h1. rewrite h2 at 1. f_equal. apply proj1_sig_injective; simpl. pose proof (proj1_sig_map_assment_w_dep_eq_iff (exist (fun c : tuple_prod (U_dep_t sig A) (projT2 (fst (sig op pfop))) => Inn (tuple_prod_ens_dep (A_dep_t sig A) (projT2 (fst (sig op pfop)))) c) p pfintpl) f (proj1_sig (map_assment_w_dep g (exist (fun c : tuple_prod (U_dep_t sig A) (projT2 (fst (sig op pfop))) => Inn (tuple_prod_ens_dep (A_dep_t sig A) (projT2 (fst (sig op pfop)))) c) p pfintpl)))) as h3; simpl in h3. rewrite h3. intros j hj. specialize (H0 j hj). terml H0. redterm0 x. gterml. redterm0 x0. redterm0 c0. pose proof (proof_irrelevance _ c1 c) as h6. fold c1. fold c in H0. rewrite h6, H0. rewrite nth_tuple_prod_map_assment_w_dep; simpl. repeat f_equal. apply proof_irrelevance. Qed. (*Partial sigma-algebra*) Record PAlgebra {S num_ops} (sigma:signature S num_ops) : Type := { A_sigma : S -> Type; doms : forall i (pfi:i < num_ops), (Ensemble (tuple_prod_all A_sigma (fst (sigma i pfi)))); ops_sigma : forall i (pfi:i < num_ops), op_to_pfun A_sigma (sigma i pfi) (doms i pfi) }. Record PAlgebra_ens {S num_ops} U (sigma:signature S num_ops) : Type := { A_sigma_ens : S -> Ensemble U; doms_ens : forall op (pfop:op < num_ops), Ensemble (U^projT1 (fst (sigma op pfop))); incl_doms_ens : forall op (pfop:op < num_ops), Included (doms_ens op pfop) (tuple_prod_ens A_sigma_ens (projT2 (fst (sigma op pfop)))); ops_sigma_ens : forall op (pfop:op < num_ops), op_to_pfun_ens A_sigma_ens (sigma op pfop) (doms_ens op pfop) (incl_doms_ens op pfop) }. Lemma pointwise_ind_eq_seg_fun_prod_type_op : forall {S num_ops} {sig:signature S num_ops} (A:TAlgebra_dep sig) op (pfop:op < num_ops) (f:seg_fun_prod_type (fun (j : nat) (pfj : j < projT1 (fst (sig op pfop))) => U_dep_t sig A (nth_prod j pfj (projT2 (fst (sig op pfop)))))), pointwise_ind_eq (seg_fun_type_to_ind (fun (j : nat) (pfj : j < projT1 (fst (sig op pfop))) => U_dep_t sig A (nth_prod j pfj (projT2 (fst (sig op pfop)))))) (U_dep_t sig A) (full_segT_ind (projT1 (fst (sig op pfop)))) (projT2 (fst (sig op pfop))). intros S num_ops sig A op hop f; red; intros i hi. unfold seg_fun_type_to_ind. f_equal. unfold full_segT_ind. rewrite nth_map_nprod; simpl; auto. Qed. Definition term X num_ops : Type := list (X + (num_ops @)). Definition pre_sys {S X} num_ops (v:X->S) (s:S) : Ensemble (term X num_ops) := Im (pre_sys_aux v s) (fun x => inl x::nil). Lemma in_pre_sys : forall {S X} num_ops (v:X->S) (s:S) x, Inn (pre_sys num_ops v (v x)) (inl x::nil). intros; apply Im_intro with x; constructor; auto. Qed. (*Prefix Polish Notation.*) Inductive Term_algebra_ens {S X num_ops} (sig:signature S num_ops) (v:X->S) : S -> Ensemble (term X num_ops) := | term_algebra_ensX : forall (x:X), Inn (Term_algebra_ens sig v (v x)) ((inl x)::nil) | term_algebra_nil_op : forall i (pfi:i < num_ops), projT1 (fst (sig i pfi)) = 0 -> Inn (Term_algebra_ens sig v (snd (sig i pfi))) (inr (exist (fun c => c < num_ops) i pfi) :: nil) | term_algebra_ens_app : forall i (pfi: i < num_ops) (t:(term X num_ops)^projT1 (fst (sig i pfi))), (forall j (pfj:j < projT1 (fst (sig i pfi))), Inn (Term_algebra_ens sig v (nth_prod j pfj (projT2 (fst (sig i pfi))))) (nth_prod j pfj t)) -> Inn (Term_algebra_ens sig v (snd (sig i pfi))) (inr (exist (fun c => c < num_ops) i pfi) :: app_list (nprod_to_list _ _ t)). (*I know inductive sets don't need something like this, but oftentimes the actual parameters are not in form required by constructors.*) Lemma in_term_algebra_ens_iff : forall {S X num_ops} (sig:signature S num_ops) (v:X->S) (s:S) (t:term X num_ops), Inn (Term_algebra_ens sig v s) t <-> (exists x, s = v x /\ t = (inl x)::nil) \/ (exists i (pfi:i < num_ops), s = snd (sig i pfi) /\ projT1 (fst ((sig i pfi))) = 0 /\ t = inr _ (exist (fun c => c < num_ops) i pfi) :: nil) \/ (exists i (pfi:i < num_ops) (t':(term X num_ops)^projT1 (fst (sig i pfi))), s = snd (sig i pfi) /\ (forall j (pfj:j < projT1 (fst (sig i pfi))), Inn (Term_algebra_ens sig v (nth_prod j pfj (projT2 (fst (sig i pfi))))) (nth_prod j pfj t')) /\ t = inr (exist (fun c => c < num_ops) i pfi) :: (app_list (nprod_to_list _ _ t'))). intros S X num_ops sig v s t; split; intro h2. inversion h2; subst. left. exists x; auto. right; left. exists i, pfi. split; auto. right; right. exists i, pfi, t0. split; auto. destruct h2 as [h2|[h2|h2]]. destruct h2 as [x [h2 ?]]; subst. constructor. destruct h2 as [i [h2 [h3 [h4 ?]]]]; subst. constructor. assumption. destruct h2 as [i [h2 [t' [? [h3 ?]]]]]; subst. constructor; auto. Qed. Lemma in_term_algebra_not_nil : forall {S X num_ops} (sig:signature S num_ops) (v:X->S) (t:term X num_ops) (s:S), Inn (Term_algebra_ens sig v s) t -> t <> nil. intros ? ? ? ? ? ? ? h1; induction h1; apply cons_neq_nil. Qed. Lemma in_nth_tuple_prod_term_algebra : forall {S X num_ops} {sig:signature S num_ops} {v:X->S} {op} {pfop:op < num_ops} (a':assment_w_dep (Term_algebra_ens sig v) (projT2 (fst (sig op pfop)))) j (hj:j < projT1 (fst (sig op pfop))) x (pfe:nth_prod j hj (projT2 (fst (sig op pfop))) = v x), Inn (Term_algebra_ens sig v (v x)) (nth_tuple_prod (proj1_sig a') j hj). intros ? ? ? ? ? ? ? a'. destruct a' as [p hp]. destruct hp as [hp]. intros j hj; pose proof (hp j hj) as h1; simpl. intro x. inversion h1; clear h1. intro. rewrite <- pfe. constructor. intro. rewrite <- pfe. constructor; auto. intro h1. rewrite <- h1. constructor. intros k hk. apply H1. Qed. Definition Term_algebra_at_op {S X num_ops} (sig:signature S num_ops) (v:X->S) op (pfop:op < num_ops) : Ensemble ((term X num_ops)^(projT1 (fst (sig op pfop)))) := tuple_prod_ens (Term_algebra_ens sig v) (projT2 (fst (sig op pfop))). Definition Term_algebra_ens_op {S X num_ops} (sig:signature S num_ops) (v:X->S) op (pfop:op < num_ops) : Ensemble (term X num_ops) := Term_algebra_ens sig v (snd (sig op pfop)). Definition Term_algebra_ens_dep {S X num_ops n} (sig:signature S num_ops) (v:X->S) (w:S^n) : Ensemble (tuple_prod (fun _:S => term X num_ops) w) := tuple_prod_ens_dep (Term_algebra_ens sig v) w. Definition Term_algebra_ens_dep_at_op {S X num_ops} (sig:signature S num_ops) (v:X->S) op (pfop:op < num_ops) : Ensemble (tuple_prod (fun _:S => term X num_ops) (projT2 (fst (sig op pfop)))) := tuple_prod_ens_dep (Term_algebra_ens sig v) (projT2 (fst (sig op pfop))). Definition Term_algebra_w {S X num_ops n} (sig:signature S num_ops) (v:X -> S) (w:S^n) : Ensemble ((term X num_ops)^n) := tuple_prod_ens (Term_algebra_ens sig v) w. Definition Term_algebra_sys {S X T num_ops} (sig:signature S num_ops) (v:X -> S) (w:T->S) : T -> Ensemble (term X num_ops) := tuple_prod_sys (Term_algebra_ens sig v) w. Definition Term_algebra_op_aux {S X num_ops} (sig:signature S num_ops) (v:X -> S) op (pfop:op < num_ops) : fun_in_ens (Term_algebra_at_op sig v op pfop) (term X num_ops) := fun w pf => (inr (exist (fun c => c < num_ops) op pfop) :: app_list (nprod_to_list _ _ w)). Lemma in_term_algebra_op_aux : forall {S X num_ops} (sig:signature S num_ops) (v:X -> S) op (pfop:op < num_ops) (w:term X num_ops ^ projT1 (fst (sig op pfop))) (pf:Inn (Term_algebra_at_op sig v op pfop) w), Inn (Term_algebra_ens sig v (snd (sig op pfop))) (Term_algebra_op_aux sig v op pfop w pf). intros S X num_ops sig v i hi w hw. unfold Term_algebra_op_aux. constructor. intros j hj. destruct hw as [h1]; subst; auto. Qed. Definition Term_algebra_op {S X num_ops} (sig:signature S num_ops) (v:X -> S) op (pfop:op < num_ops) : sig_set (Term_algebra_at_op sig v op pfop) -> sig_set (Term_algebra_ens sig v (snd (sig op pfop))) := fun x' => let (x, pfx) := x' in let y := Term_algebra_op_aux sig v op pfop x pfx in let pfy := in_term_algebra_op_aux sig v op pfop x pfx in exist _ y pfy. Definition Term_algebra {S X num_ops} (pfds:type_eq_dec S) (sig:signature S num_ops) (v:X -> S) : TAlgebra_ens (term X num_ops) sig := Build_TAlgebra_ens S num_ops (term X num_ops) sig pfds (Term_algebra_ens sig v) (Term_algebra_op sig v). Definition op_term_to_term {S X num_ops} {sig:signature S num_ops} {v:X->S} {op} {pfop:op < num_ops} (x:sig_set (Term_algebra_at_op sig v op pfop)) : term X num_ops := app_list (nprod_to_list _ _ (proj1_sig x)). Definition op_term_to_app_term {S X num_ops} {sig:signature S num_ops} {v:X->S} {op} {pfop:op < num_ops} (x:sig_set (Term_algebra_at_op sig v op pfop)) : term X num_ops := (inr (exist (fun c => c < num_ops) _ pfop))::op_term_to_term x. Lemma op_term_to_app_term_neq : forall {S X num_ops} {sig:signature S num_ops} {v:X->S} {op} {pfop:op < num_ops} (y:sig_set (Term_algebra_at_op sig v op pfop)) (x:X), op_term_to_app_term y <> inl x::nil. unfold op_term_to_app_term; intros; intro; discriminate. Qed. Lemma in_op_term_to_app_term : forall {S X num_ops} {sig:signature S num_ops} {v:X->S} {op} {pfop:op < num_ops} (x:sig_set (Term_algebra_at_op sig v op pfop)), Inn (Term_algebra_ens sig v (snd (sig op pfop))) (op_term_to_app_term x). unfold op_term_to_app_term; intros S X num_ops sig v hop h1 x; destruct x as [x hx]; constructor; intros j hj; inversion hx; auto. Qed. Lemma op_term_to_term_sig_tuple_prod_ens_dep_const_to_term_algebra : forall {S X num_ops} {sig:signature S num_ops} (v:X->S) op (pfop: op < num_ops) (t:term X num_ops ^ projT1 (fst (sig op pfop))) (pft:forall j (pfj:j < projT1 (fst (sig op pfop))), Inn (Term_algebra_ens sig v (nth_prod j pfj (projT2 (fst (sig op pfop))))) (nth_prod j pfj t)), op_term_to_term (sig_tuple_prod_ens_dep_const_to (Term_algebra_ens sig v) (projT2 (fst (sig op pfop))) (seg_fun_prod_ens_to_assment (Term_algebra_ens sig v) (projT2 (fst (sig op pfop))) (fun j (pfj:j < projT1 (fst (sig op pfop))) => nth_prod j pfj t) pft)) = app_list (nprod_to_list (term X num_ops) (projT1 (fst (sig op pfop))) t). unfold op_term_to_term, seg_fun_prod_ens_to_assment; intros; f_equal. pose proof (sig_tuple_prod_ens_dep_const_to_seg_fun_depT_nth_to_tuple_prod (Term_algebra_ens sig v) (projT2 (fst (sig op pfop))) t pft) as h1. f_equal. simpl in h1; simpl; auto. Qed. Definition Term_algebra_op_aux_dep {S X num_ops} (sig:signature S num_ops) (v:X -> S) i (pfi:i < num_ops) : fun_in_ens (Term_algebra_ens_dep_at_op sig v i pfi) (term X num_ops) := fun p pfp => let pfp' := in_tuple_prod_ens_dep_const_compat _ _ pfp in op_term_to_app_term (exist _ _ pfp'). Lemma in_term_algebra_op_aux_dep : forall {S X num_ops} (sig:signature S num_ops) (v:X -> S) i (pfi:i < num_ops) (p:tuple_prod (fun _:S => term X num_ops) (projT2 (fst (sig i pfi)))) (pfp:Inn (Term_algebra_ens_dep_at_op sig v i pfi) p), let pfp' := in_tuple_prod_ens_dep_const_compat _ _ pfp in Inn (Term_algebra_ens sig v (snd (sig i pfi))) (op_term_to_app_term (exist _ _ pfp')). unfold op_term_to_app_term; intros S X num_ops sig v i hi p hp; simpl. constructor; simpl. intros; inversion hp; auto. specialize (H j pfj). rewrite nth_prod_transfer_tuple_prod_const; auto. Qed. Definition Term_algebra_dep' {S X num_ops} (sig:signature S num_ops) (v:X -> S) i (pfi:i < num_ops) : op_to_fun_dep' sig (Term_algebra_ens sig v) i pfi := fun (x':sig_set (Term_algebra_ens_dep_at_op sig v i pfi)) => let (x, pfx) := x' in let y := Term_algebra_op_aux_dep sig v i pfi x pfx in let pfx' := in_term_algebra_op_aux_dep sig v i pfi x pfx in exist (fun c => Inn (Term_algebra_ens sig v (snd (sig i pfi))) c) y pfx'. Definition Term_algebra_dep {S X num_ops} (pfds:type_eq_dec S) (sig:signature S num_ops) (v:X->S) : TAlgebra_dep sig := Build_TAlgebra_dep S num_ops sig (fun s => term X num_ops) pfds (Term_algebra_ens sig v) (Term_algebra_dep' sig v). Lemma incl_term_algebra : forall {S X num_ops} (pfds:type_eq_dec S) (sig:signature S num_ops) (v:X->S) s, (Included (Term_algebra_ens sig v s) (A_dep_t _ (Term_algebra_dep pfds sig v) s)). intros; red; intros; simpl; auto. Qed. Lemma in_assment_w_dep_term_algebra_ens : forall {S X num_ops} {sig:signature S num_ops} {v:X->S} op (pfop:op < num_ops) (a:assment_w_dep (Term_algebra_ens sig v) (projT2 (fst (sig op pfop)))), Inn (Term_algebra_ens_dep_at_op sig v op pfop) (proj1_sig a). unfold Term_algebra_ens_dep_at_op; intros; constructor; intros i hi. pose proof (proj2_sig a) as h1; simpl in h1. inversion h1; auto. Qed. Lemma proj1_sig_term_algebra_dep_eq' : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} op (pfop:op < num_ops) (a:assment_w_dep (Term_algebra_ens sig v) (projT2 (fst (sig op pfop)))), proj1_sig (Term_algebra_dep' sig v op pfop a) = inr (exist (fun c : nat => c < num_ops) op pfop) :: op_term_to_term (sig_tuple_prod_ens_dep_const_to (Term_algebra_ens sig v) (projT2 (fst (sig op pfop))) a). intros S X num_ops sig A v op hop a. unfold Term_algebra_dep', Term_algebra_op_aux_dep, op_term_to_app_term. destruct a; auto. Qed. Lemma incl_pre_sys : forall {S X} {num_ops} (pfds:type_eq_dec S) (sig:signature S num_ops) (v:X->S) s, Included (pre_sys num_ops v s) (A_dep_t _ (Term_algebra_dep pfds sig v) s). intros S X num_ops hds sig v s; red; intros t ht. destruct ht as [x ht]; subst. destruct ht as [ht]; subst. constructor. Qed. Lemma pre_sys_generates_term_algebra' : forall {S X num_ops} (pfds:type_eq_dec S) (sig:signature S num_ops) (v:X->S), alg_gen (pre_sys num_ops v) (incl_pre_sys pfds sig v) (A:=Term_algebra_dep pfds sig v). intros S X num_ops hds sig v. constructor. intros s u h1. pose proof h1 as h2. pose proof (point_gen_point (pre_sys num_ops v) (A := Term_algebra_dep hds sig v)) as h3. induction h2. pose proof (proof_irrelevance _ h1 (term_algebra_ensX sig v x)); subst. hfirst h3. intros s u h1. destruct h1; destruct H; subst; simpl; constructor. specialize (h3 hf (v x) (inl x ::nil)). hfirst h3. unfold pre_sys. apply Im_intro with x. constructor; auto. reflexivity. specialize (h3 hf0). pose proof (proof_irrelevance _ (hf (v x) (inl x::nil) hf0) (term_algebra_ensX sig v x)) as h4. pose proof (proof_irrelevance _ hf (incl_pre_sys hds sig v)) as h5; subst. rewrite <- h4; auto. simpl in h1. hfirst h3. intros s u h4. destruct h4; destruct H; destruct H0; subst; simpl; constructor. pose proof (point_gen_nil_op (pre_sys num_ops v) (incl_pre_sys hds sig v) (A := Term_algebra_dep hds sig v) i pfi H) as h4; simpl in h4. unfold Term_algebra_op_aux_dep, op_term_to_app_term, op_term_to_term in h4; simpl in h4. term1 h4. redterm0 c. assert (h5:c0 = nil). unfold c0. rewrite app_list_eq_nil_iff. left. pose proof @nprod_to_list_eq_nil_iff. gterml. redterm0 x. pose proof (nprod_to_list_eq_nil_iff c1) as h6. fold c1. rewrite h6 at 1; auto. pose proof (f_equal (fun c => inr (exist (fun q => q < num_ops) i pfi) :: c) h5) as h6; simpl in h6. term0 h4. fold c1; fold c1 in h4. pose proof c1 as h7; simpl in h7. unfold Term_algebra_op_aux_dep, op_term_to_app_term, op_term_to_term in h7; simpl in h7, c1. fold c0 in h7. rewrite h6 in h7. pose proof (subsetT_eq_compat _ _ _ _ c1 h7 h6) as h8. unfold Term_algebra_op_aux_dep, op_term_to_app_term, op_term_to_term in h8; simpl in h8. dependent rewrite -> h8 in h4. pose proof (proof_irrelevance _ h7 h1); subst; auto. pose proof (point_gen_app (pre_sys num_ops v) (A := Term_algebra_dep hds sig v)) as h8. hfirst h8. intros s q hq. destruct hq as [q]; subst. destruct H1; subst. simpl. constructor. specialize (h8 hf i pfi); simpl in h8. specialize (h8 (transfer_to_tuple_prod_const _ t)). hfirst h8. constructor. intros j hj. rewrite nth_tuple_prod_transfer_to_tuple_prod_const; auto. specialize (h8 hf0). hfirst h8. intros j hj. gen0. rewrite nth_tuple_prod_transfer_to_tuple_prod_const. intro h9. specialize (H0 j hj h9); auto. pose proof (proof_irrelevance _ hf (incl_pre_sys hds sig v)); subst; auto. specialize (h8 hf1). pose proof (proof_irrelevance _ hf (incl_pre_sys hds sig v)); subst; auto. term0 h8; simpl in c. unfold Term_algebra_op_aux_dep, op_term_to_app_term, op_term_to_term in h8; simpl in h8. pose proof (transfer_tuple_prod_const_to_compat (projT2 (fst (sig i pfi))) t) as h9. pose proof (f_equal (fun c => (inr (exist (fun c0 : nat => c0 < num_ops) i pfi) :: app_list (nprod_to_list (term X num_ops) (projT1 (fst (sig i pfi))) c))) h9) as h10; simpl in h10. term0 h8. pose proof c0 as h11. simpl in h11. unfold Term_algebra_op_aux_dep, op_term_to_app_term, op_term_to_term in h11; simpl in h11. rewrite h10 in h11. pose proof (subsetT_eq_compat _ _ _ _ c0 h11 h10) as h12. unfold Term_algebra_op_aux_dep, op_term_to_app_term, op_term_to_term in h12; simpl in h12. fold c0 in h8. dependent rewrite -> h12 in h8. pose proof (proof_irrelevance _ h1 h11). unfold c0 in h12. clear c c0 h9. subst; auto. Qed. Lemma in_nest_gen_n_var_iff : forall {S X num_ops} {sig:signature S num_ops} {v:X->S} (pfds:type_eq_dec S) x y (pfx:Inn (Term_algebra_ens sig v (v x)) (inl x::nil)) (pfy:Inn (Term_algebra_ens sig v (v y)) (inl y::nil)) n, in_nest_gen_n (incl_pre_sys pfds sig v) (v x) (v y) (inl x::nil) (inl y::nil) pfx pfy n (A := Term_algebra_dep pfds sig v) <-> n = 0 /\ x = y. intros S X num_ops sig v hds x y hx hy n; split; intro h1. inversion h1; subst. inversion H2; tauto. destruct h1; subst. pose proof (in_nest0 (incl_pre_sys hds sig v) (v y) (inl y::nil) (in_pre_sys _ _ (v y) _) (A := Term_algebra_dep hds sig v)) as h1; simpl in h1. term1 h1. pose proof (proof_irrelevance _ c hx) as he. rewrite <- he; clear he. pose proof (proof_irrelevance _ c hy) as he. rewrite <- he; clear he. unfold c; auto. Qed. Lemma in_nest_gen_discriminate : forall {S X num_ops} {sig:signature S num_ops} {v:X->S} (pfds:type_eq_dec S) op (pfop:op < num_ops) x n (pfr: Inn (Term_algebra_ens sig v (snd (sig op pfop))) (inr (exist (fun c : nat => c < num_ops) op pfop) :: nil)), ~ in_nest_gen_n (incl_pre_sys pfds sig v) (v x) (snd (sig op pfop)) (inl x :: nil) (inr (exist (fun c : nat => c < num_ops) op pfop) :: nil) (term_algebra_ensX sig v x) pfr n (A := Term_algebra_dep pfds sig v). intros; intro h1. dependent induction h1. rename IHh1 into ih. destruct pfb. rewrite in_nest_gen_n_var_iff in h1. destruct h1; subst. pose proof (proof_irrelevance _ pfop pfop0); subst. clear x1. rename x into hpfe. rename H into h1. rename x2 into heqn. rewrite in_tuple_prod_iff in h1. destruct h1 as [i [h1 [h2 h3]]]. unfold op_term_to_term in heqn. rewrite app_list_eq_nil_iff in heqn; simpl in heqn. destruct heqn as [h4|h4]. rewrite (nprod_to_list_eq_nil_iff (transfer_tuple_prod_const p)) in h4 at 1. omega. red in h4. specialize (h4 (inl x3::nil)). pose proof (f_equal (transfer_dep_r h2) h3) as h5. rewrite transfer_dep_r_undoes_transfer_dep in h5. rewrite h5 in h4 at 1. hfirst h4. unfold term. erewrite (in_transfer_dep_r_nprod_to_list_const_iff h2 (nth_tuple_prod p i h1)). rewrite <- (nth_prod_transfer_tuple_prod_const p). apply in_nth_prod. specialize (h4 hf); discriminate. pose proof (proof_irrelevance _ pfop pfop0); subst. specialize (ih i pfi x0). hfirst ih. constructor; auto. specialize (ih hf). hfirst ih. constructor. specialize (ih hf0). hfirst ih. constructor. specialize (ih hf1 (eq_refl _) (eq_refl _)). hfirst ih. constructor. specialize (ih hf2). hfirst ih. constructor. specialize (ih hf3). hfirst ih. constructor. specialize (ih hf4). hfirst ih. apply eq_JMeq. apply proof_irrelevance. specialize (ih hf5); auto. rename x2 into h2. rename H into h3. unfold op_term_to_term in h2. simpl in h2. rewrite app_list_eq_nil_iff in h2. rewrite in_tuple_prod_iff in h3. destruct h3 as [j [hj [h4 h5]]]. destruct h2 as [h2|h2]. rewrite (nprod_to_list_eq_nil_iff (transfer_tuple_prod_const p)) in h2. pose proof (proof_irrelevance _ pfop pfop0); subst; omega. red in h2. specialize (h2 (nth_tuple_prod p j hj)). hfirst h2. rewrite <- in_nprod_iff', <- (nth_prod_transfer_tuple_prod_const p). apply in_nth_prod. specialize (h2 hf). pose proof pfp as hp. destruct hp as [hp]. specialize (hp j hj). simpl in hp. rewrite h2 in hp at 1. apply in_term_algebra_not_nil in hp. contradict hp; auto. Qed. Inductive Recursion_homo_aux_var {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) : forall (s:S) (t:term X num_ops) (u:U_dep_t sig A s) (pfu:Inn (A_dep_t sig A s) u), Prop := recursion_homo_aux_var_intro : forall x, Recursion_homo_aux_var a (v x) (inl x :: nil) (proj1_sig (a x)) (proj2_sig _). Lemma Recursion_homo_aux_var_functional : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) (s s':S) (t t':term X num_ops) (u:U_dep_t sig A s) (u':U_dep_t sig A s') (pfu: Inn (A_dep_t sig A s) u) (pfu':Inn (A_dep_t sig A s') u') (pfe:s = s'), t = t' -> u = transfer_dep_r pfe u' -> Recursion_homo_aux_var a s t u pfu -> Recursion_homo_aux_var a s' t' u' pfu'. intros; subst. pose proof (proof_irrelevance _ pfu pfu'); subst; auto. Qed. (*This "operator variant" exists to enable better abstraction of formal parameters, i.e. fail-safe inversion, i.e. one that uses primitive inductive parameters in the actual parameters -- so it seems, and is verbose, especially in the second constructor. If you would a prefer a simpler (concise) version that fails on some inversions, then consult the repository tree for just that!*) Inductive Recursion_homo_aux_op {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) op (pfop:op < num_ops) : forall (t:term X num_ops) (u:U_dep_t sig A (snd (sig op pfop))) (pfu:Inn (A_dep_t sig A (snd (sig op pfop))) u), Prop := | recursion_homo_aux_op_nil_op_intro : let n := projT1 (fst (sig op pfop)) in let w := projT2 (fst (sig op pfop)) in let s := snd (sig op pfop) in forall (pfeq:projT1 (fst (sig op pfop)) = 0), let pfintt := in_transfer_tt_tuple_prod_ens_dep (U_dep_t sig A) (A_dep_t sig A) w pfeq in let u := ops_dep_t sig A op pfop (exist (fun c => Inn (tuple_prod_ens_dep (A_dep_t sig A) w) c) _ pfintt) in Recursion_homo_aux_op a op pfop (inr (exist (fun c => c < num_ops) op pfop) :: nil) (proj1_sig u) (proj2_sig u) | rec_homo_aux_op_app : let n := projT1 (fst (sig op pfop)) in let w := projT2 (fst (sig op pfop)) in let s := snd (sig op pfop) in forall (ttm:assment_w_dep (Term_algebra_ens sig v) w) (ta:assment_w_dep (A_dep_t sig A) w), (forall j (pfj:j < n), let wj := nth_prod j pfj w in let jttm := nth_tuple_prod (proj1_sig ttm) j pfj in let jta := nth_tuple_prod (proj1_sig ta) j pfj in let pftaj := iff1 (in_tuple_prod_ens_dep_nth_iff (A_dep_t sig A) _) (proj2_sig ta) j pfj in (Recursion_homo_aux_var a wj (nth_tuple_prod (proj1_sig ttm) j pfj) (nth_tuple_prod (proj1_sig ta) j pfj) pftaj \/ exists (pfjttm:jttm <> nil) op' (pfop':op' < num_ops) (pfe:wj = snd (sig op' pfop')), let jta' := transfer_dep pfe jta in let pftaj' := in_transfer_dep pfe _ pftaj in hd' jttm pfjttm = inr (exist (fun c => c < num_ops) op' pfop') /\ Recursion_homo_aux_op a op' pfop' jttm jta' pftaj')) -> let opta := ops_dep_t sig A op pfop (exist (fun c => Inn (tuple_prod_ens_dep (A_dep_t sig A) w) c) (proj1_sig ta) (proj2_sig ta)) in let ttm' := sig_tuple_prod_ens_dep_const_to _ _ ttm in let tm := inr (exist (fun c => c < num_ops) op pfop) :: (op_term_to_term ttm') in Recursion_homo_aux_op a op pfop tm (proj1_sig opta) (proj2_sig opta). Lemma recursion_homo_aux_op_functional : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) op op' (pfop:op < num_ops) (pfop':op' < num_ops) (t t':term X num_ops) (u:U_dep_t sig A (snd (sig op pfop))) (u':U_dep_t sig A (snd (sig op' pfop'))) (pfu:Inn (A_dep_t sig A (snd (sig op pfop))) u) (pfu':Inn (A_dep_t sig A (snd (sig op' pfop'))) u') (pfope:op = op'), t = t' -> u = transfer_seg_fun_depT2_r pfope pfop u' (f := fun i pfi => snd (sig i pfi)) -> Recursion_homo_aux_op a op pfop t u pfu -> Recursion_homo_aux_op a op' pfop' t' u' pfu'. intros; subst. pose proof (proof_irrelevance _ pfop pfop'); subst. revert H1. Gen0. rewrite transfer_seg_fun_depT2_r_eq_refl; auto. intros h1 h2. pose proof (proof_irrelevance _ h1 pfu'); subst; auto. Qed. Lemma recursion_homo_aux_op_hd : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) op (pfop:op < num_ops) (t:term X num_ops) (u:U_dep_t sig A (snd (sig op pfop))) (pfu:Inn (A_dep_t sig A (snd (sig op pfop))) u), Recursion_homo_aux_op a op pfop t u pfu -> exists pf:t <> nil, hd' t pf = inr (exist (fun c => c < num_ops) op pfop). intros ? ? ? ? ? ? ? ? ? ? ? ? h1. destruct h1; try unfold tm; exists (cons_neq_nil _ _); simpl; auto. Qed. Inductive Recursion_homo_aux {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) : forall (s:S) (t:term X num_ops) (u:U_dep_t sig A s) (pfu:Inn (A_dep_t sig A s) u), Prop := recursion_homo_aux_var : forall x, Recursion_homo_aux a (v x) (inl x :: nil) (proj1_sig (a x)) (proj2_sig _) | recursion_homo_aux_op : forall op (pfop:op < num_ops) (t:term X num_ops) (u:U_dep_t sig A (snd (sig op pfop))) (pfu:Inn (A_dep_t sig A (snd (sig op pfop))) u), Recursion_homo_aux_op a op pfop t u pfu -> Recursion_homo_aux a (snd (sig op pfop)) t u pfu. Lemma Recursion_homo_aux_functional : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) (s s':S) (t t':term X num_ops) (u:U_dep_t sig A s) (u':U_dep_t sig A s') (pfu: Inn (A_dep_t sig A s) u) (pfu':Inn (A_dep_t sig A s') u') (pfe:s = s'), t = t' -> u = transfer_dep_r pfe u' -> Recursion_homo_aux a s t u pfu -> Recursion_homo_aux a s' t' u' pfu'. intros; subst; pose proof (proof_irrelevance _ pfu pfu'); subst; auto. Qed. Lemma recursion_homo_aux_var_well_defined : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) (s:S) (t:term X num_ops) (u u':U_dep_t sig A s) (pfu:Inn (A_dep_t sig A s) u) (pfu':Inn (A_dep_t sig A s) u'), Recursion_homo_aux_var a s t u pfu -> Recursion_homo_aux_var a s t u' pfu' -> u = u'. intros S X num_ops sig A v a s t u u' hu hu' h1 h2. inversion h1; subst. inversion h2; subst. apply existT_injective2 in H7; subst. apply existT_injective2 in H5; subst; auto. Qed. (*This monstrosity is a leftover from something that might or might not have worked out or work out, so I'll leave it for now, for aesthetics, but will make a case for its existence or (non-) in due time.*) Lemma split_built_same_ind : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} {a:assment_dep (A_dep_t sig A) v} {R:forall {op0} {pfop0:op0 < num_ops} {op1} {pfop1:op1 < num_ops} {t t':term X num_ops} {u:U_dep_t _ A (snd (sig op0 pfop0))} {u':U_dep_t _ A (snd (sig op1 pfop1))} {pfu:Inn (A_dep_t _ A (snd (sig op0 pfop0))) u} {pfu':Inn (A_dep_t _ A (snd (sig op1 pfop1))) u'}, Recursion_homo_aux_op a op0 pfop0 t u pfu -> Recursion_homo_aux_op a op1 pfop1 t' u' pfu' -> Prop} {op0} {pfop0:op0 < num_ops} {op1} {pfop1:op1 < num_ops}, let n0 := projT1 (fst (sig op0 pfop0)) in let n1 := projT1 (fst (sig op1 pfop1)) in let w0 := projT2 (fst (sig op0 pfop0)) in let w1 := projT2 (fst (sig op1 pfop1)) in let s0 := snd (sig op0 pfop0) in let s1 := snd (sig op1 pfop1) in forall (ttm: assment_w_dep (Term_algebra_ens sig v) w0) (ttm':assment_w_dep (Term_algebra_ens sig v) w1) (ta:assment_w_dep (A_dep_t sig A) w0) (ta':assment_w_dep (A_dep_t sig A) w1) (pfen:n0 = n1), (forall j (pfj:j < n0), let pfj' := lt_eq_trans _ _ _ pfj pfen in let wj0 := nth_prod j pfj w0 in let wj1 := nth_prod j pfj' w1 in let jttm := nth_tuple_prod (proj1_sig ttm) j pfj in let jttm' := nth_tuple_prod (proj1_sig ttm') j pfj' in let jta := nth_tuple_prod (proj1_sig ta) j pfj in let jta' := nth_tuple_prod (proj1_sig ta') j pfj' in let pftaj := iff1 (in_tuple_prod_ens_dep_nth_iff (A_dep_t sig A) _) (proj2_sig ta) j pfj in let pftaj' := iff1 (in_tuple_prod_ens_dep_nth_iff (A_dep_t sig A) _) (proj2_sig ta') j pfj' in Recursion_homo_aux_var a wj0 jttm jta pftaj /\ Recursion_homo_aux_var a wj1 jttm' jta' pftaj' \/ (exists (pfjttm: jttm <> nil) (pfjttm':jttm' <> nil) op2 (pfop2:op2 < num_ops) op3 (pfop3:op3 < num_ops) (pfe0:wj0 = snd (sig op2 pfop2)) (pfe1:wj1 = snd (sig op3 pfop3)), let jta0 := transfer_dep pfe0 jta in let jta0' := transfer_dep pfe1 jta' in let pftaj0 := in_transfer_dep pfe0 _ pftaj in let pftaj0' := in_transfer_dep pfe1 _ pftaj' in exists (pfr2:Recursion_homo_aux_op a op2 pfop2 jttm jta0 pftaj0) (pfr3:Recursion_homo_aux_op a op3 pfop3 jttm' jta0' pftaj0'), hd' jttm pfjttm = inr (exist (fun c => c < num_ops) op2 pfop2) /\ hd' jttm' pfjttm' = inr (exist (fun c => c < num_ops) op3 pfop3) /\ R pfr2 pfr3)) -> (forall j (pfj:j < n0), let wj := nth_prod j pfj w0 in let jttm := nth_tuple_prod (proj1_sig ttm) j pfj in let jta := nth_tuple_prod (proj1_sig ta) j pfj in let pftaj := iff1 (in_tuple_prod_ens_dep_nth_iff (A_dep_t sig A) _) (proj2_sig ta) j pfj in Recursion_homo_aux_var a wj jttm jta pftaj \/ (exists (pfjttm: jttm <> nil) op2 (pfop2:op2 < num_ops) (pfe:wj = snd (sig op2 pfop2)), let jta0 := transfer_dep pfe jta in let pftaj0 := in_transfer_dep pfe _ pftaj in hd' jttm pfjttm = inr (exist (fun c => c < num_ops) op2 pfop2) /\ Recursion_homo_aux_op a op2 pfop2 jttm jta0 pftaj0)) /\ (forall j (pfj:j < n1), let wj := nth_prod j pfj w1 in let jttm' := nth_tuple_prod (proj1_sig ttm') j pfj in let jta' := nth_tuple_prod (proj1_sig ta') j pfj in let pftaj' := iff1 (in_tuple_prod_ens_dep_nth_iff (A_dep_t sig A) _) (proj2_sig ta') j pfj in Recursion_homo_aux_var a wj jttm' jta' pftaj' \/ (exists (pfjttm': jttm' <> nil) op3 (pfop3:op3 < num_ops) (pfe:wj = snd (sig op3 pfop3)), let jta0' := transfer_dep pfe jta' in let pftaj0' := in_transfer_dep pfe _ pftaj' in hd' jttm' pfjttm' = inr (exist (fun c => c < num_ops) op3 pfop3) /\ Recursion_homo_aux_op a op3 pfop3 jttm' jta0' pftaj0')). intros S X num_ops sig A v a R op0 hop0 op1 hop1; simpl; intros ttm ttm' ta ta' hne h1; split; intros j hj. specialize (h1 j hj); destruct h1 as [h1 | h1]; destruct h1 as [h1 h2]. left; auto. destruct h2 as [h2 [op2 [hop2 [op3 [hop3 [h3 [h4 [h5 [h6 [h7 [h8 h9]]]]]]]]]]]. right. exists h1, op2, hop2, h3; split; auto. specialize (h1 j (lt_eq_trans _ _ _ hj (eq_sym hne))). destruct h1 as [h1|h1]; destruct h1 as [h1 h2]. left. term0 h2. term1 h2. term2 h2. term3 h2. gterm0. pose proof (Recursion_homo_aux_var_functional a c2 (nth_prod j hj (projT2 (fst (sig op1 hop1)))) c1 (nth_tuple_prod (proj1_sig ttm') j hj) c0 (nth_tuple_prod (proj1_sig ta') j hj) c c3) as h10. unfold c, c0, c1, c2, c3 in h10. hfirst h10. apply nth_prod_functional3; auto. specialize (h10 hf). hfirst h10. f_equal; apply proof_irrelevance. specialize (h10 hf0). hfirst h10. symmetry. rewrite <- transfer_dep_r_eq_iff. gtermr. redterm0 y. redterm0 c4. fold c5. pose proof (proof_irrelevance _ c5 hj) as h12. rewrite h12; auto. specialize (h10 hf1 h2); auto. destruct h2 as [h2 [op2 [hop2 [op3 [hop3 [h3 [h4 [h5 [h6 [h7 [h8 h9]]]]]]]]]]]. right. ex_goal. nterml h2. redterm0 x. pose proof (proof_irrelevance _ c hj) as h12. rewrite <- h12; auto. exists hex. exists op3, hop3. ex_goal. rewrite <- h4. f_equal; apply proof_irrelevance. exists hex0; split; auto. rewrite <- h8. apply f_equal_dep. repeat f_equal; apply proof_irrelevance. pose proof h6 as h6'. gen0. intro h13. revert h9; revert h6. Gen0. intros h14 h15 h16. term0 h13; term0 h14. redterm0 c0. redterm0 c1. assert (h17:c = c0). unfold c, c0. rewrite transfer_dep_hetero. pose proof (nth_tuple_prod_functional3 (proj1_sig ta') j hj c2) as h18. rewrite h18. fold c2. f_equal; apply proof_irrelevance. symmetry in h17. pose proof (subsetT_eq_compat _ _ _ _ h14 h13 h17) as h18. dependent rewrite <- h18. pose proof h15 as h15'. fold c0 c c2 in h15'. fold c0. apply (recursion_homo_aux_op_functional a op3 op3 hop3 hop3 (nth_tuple_prod (proj1_sig ttm') j c2) (nth_tuple_prod (proj1_sig ttm') j hj) c0 c0 h14 h14 (eq_refl _)). f_equal; apply proof_irrelevance. rewrite transfer_seg_fun_depT2_r_eq_refl; auto. assumption. Qed. (*Turns out "subterm injectivity" is non-trivial, though still very manageable, and rather than postpone a Winter release, I thought I would just include the last step in [recursion_homo_ex] in a bit of a clumsy way -- by leaving it as an Axiom, more or less an Admitted, and the proof is in the works, as you can see by all the nesting functions and lemmas, and an easy induction on depth should cinch this, and something a bit stronger that will be reworked in [recursion_homo_ex]. Release date over purity of release is what it's about, so long as the releases come yearly at Winter Solstice!*) Axiom terms_app_list_eq_length : forall {S X num_ops n} {sig:signature S num_ops} (v:X->S) (p:S^n) (t u:term X num_ops^n), (forall i (pfi:i < n), Inn (Term_algebra_ens sig v (nth_prod i pfi p)) (nth_prod i pfi t)) -> (forall i (pfi:i < n), Inn (Term_algebra_ens sig v (nth_prod i pfi p)) (nth_prod i pfi u)) -> (app_list (nprod_to_list _ _ t) = app_list (nprod_to_list _ _ u)) -> forall i (pfi:i < n), length (nth_prod i pfi t) = length (nth_prod i pfi u). Definition recursion_homo_ex : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) (s:S) (t:term X num_ops) (pft:Inn (Term_algebra_ens sig v s) t), exists! u:U_dep_t sig A s, exists pfu:Inn (A_dep_t sig A s) u, Recursion_homo_aux a s t u pfu. intros S X num_ops sig A v a s t ht. revert a; revert A. induction ht. intros A a. exists (proj1_sig (a x)). red; split. exists (proj2_sig _). constructor. intros u hu. destruct hu as [hu h1]. inversion h1; subst. clear h1. apply existT_injective2 in H2; auto. inversion H2. intros A a. exists (proj1_sig (ops_dep_t sig A i pfi (exist (fun c => Inn (tuple_prod_ens_dep (A_dep_t sig A) (projT2 (fst (sig i pfi))) ) c) _ (in_transfer_tt_tuple_prod_ens_dep (U_dep_t sig A) (A_dep_t sig A) (projT2 (fst (sig i pfi))) H)))). red. split. exists (proj2_sig _). constructor. constructor. intros u hu. destruct hu as [h1 h2]. inversion h2; clear h2; subst. inversion H3; clear H3 H9. unfold u1, pfintt0, w, n, s, w0 in H2. clear u1 pfintt0 w n s w0. subst. pose proof (proof_irrelevance _ pfop pfi); subst. pose proof (proof_irrelevance _ H pfeq0); subst. apply existT_injective2 in H6; auto. clear H10. unfold opta, w, n, s, w0, ttm'0, w0 in H7, ttm, ta0, ttm', ttm, ttm0, tm0, H8. clear opta w n s w0 ttm'0 tm0 ttm'. subst. pose proof (proof_irrelevance _ pfi pfop); subst. apply existT_injective2 in H6. rewrite <- H6. repeat f_equal. apply proj1_sig_injective; simpl. apply tuple_prod_n0_eq; auto. intros A a. pose (fun j (pfj:j < projT1 (fst (sig i pfi))) => (proj1_sig (constructive_definite_description _ (H0 j pfj A a)))) as f. pose proof (rec_homo_aux_op_app a i pfi) as h3. simpl in h3. pose (seg_fun_prod_ens_to_assment (Term_algebra_ens sig v) (projT2 (fst (sig i pfi))) _ H) as tm_w. specialize (h3 tm_w). pose (seg_fun_prod_ens_to_assment (A_dep_t sig A) (projT2 (fst (sig i pfi))) f) as ta. hfirst ta. intros j hj. unfold f. destruct constructive_definite_description; simpl. destruct e; auto. pose (ta hf) as ta'. specialize (h3 ta'). hfirst h3. unfold ta', ta, tm_w. intros j hj. unfold f. simpl. match goal with |- Recursion_homo_aux_var _ _ _ _ ?pf' \/ _ => generalize pf' end. do 2 rewrite nth_tuple_prod_seg_fun_depT_nth_to_tuple_prod. intro h9. match goal with |- Recursion_homo_aux_var _ _ _ ?c' _ \/ _ => pose c' as c3 end. redterm0 c3. pose proof (proj2_sig c) as h10. simpl in h10. destruct h10 as [h10 h11]. inversion h11; clear h11. left. gen0. unfold c in H7. dependent rewrite <- H7. rewrite <- H3. intro. pose proof (proof_irrelevance _ h0 (proj2_sig (a x))) as h12. rewrite h12. constructor. right. subst. unfold c in H7. pose proof (recursion_homo_aux_op_hd _ _ _ _ _ _ H4) as h8. destruct h8 as [h8 h13]. exists h8, op, pfop, (eq_sym H6). split; auto. pose proof H7 as h14. symmetry in h14. rewrite (transfer_dep_eq_iff _ _ (eq_sym H6)) in h14. terml h14. pose proof (recursion_homo_aux_op_functional a op op pfop pfop (nth_prod j hj t) (nth_prod j hj t) u x pfu) as h15. gterm0. fold c0. specialize (h15 c0 (eq_refl _) (eq_refl _)). hfirst h15. unfold x. rewrite h14, transfer_seg_fun_depT2_r_eq_refl; auto. specialize (h15 hf0). unfold x in h15. apply h15; auto. specialize (h3 hf0). term1 h3. exists c. red; split. exists (proj2_sig _). constructor. unfold c. unfold tm_w in h3. rewrite op_term_to_term_sig_tuple_prod_ens_dep_const_to_term_algebra in h3; auto. intros u h1. destruct h1 as [h1 h4]. unfold c. inversion h4; clear h4. subst. inversion H4; clear H4. clear n s. unfold w0, pfintt0 in u1. unfold u1, w0, pfintt0, w0 in H3. unfold w0 in u1. unfold c. clear H11 w0 pfintt0 u1 w c. subst. pose proof (proof_irrelevance _ pfi pfop). clear H7. unfold f, ta', ta, f. unfold f, ta', ta, f, tm_w in hf, h3, hf0. clear f ta' ta tm_w. subst. apply existT_injective2 in H5; subst. repeat f_equal. apply proj1_sig_injective; simpl. rewrite seg_fun_depT_nth_to_tuple_prod_eq_iff. intros k hk. destruct constructive_definite_description; simpl. destruct e as [h10 h11]. inversion h11; omega; clear h11. unfold n in H10. clear n. unfold w. unfold w in ttm. clear ttm' w. unfold opta in H8, H11. clear opta. unfold w0 in ta1, ttm0, H10. unfold ttm'0, w0 in tm0, H9. clear ttm'0. unfold w0 in tm0, H8, H11. clear w0. subst. unfold ta', ta in h3, hf0. unfold ta', ta. clear c. clear ta' ta. unfold op_term_to_term in H9. pose proof (proof_irrelevance _ pfi pfop). unfold f. unfold f, tm_w in hf, h3, hf0. clear f tm_w. subst. apply existT_injective2 in H7. subst. repeat f_equal. apply proj1_sig_injective; simpl. rewrite seg_fun_depT_nth_to_tuple_prod_eq_iff. intros k hk. destruct constructive_definite_description; simpl. destruct e as [e he]. pose proof (H0 k hk A a) as h12. destruct h12 as [u h12]. red in h12. destruct h12 as [[h12 h13] h14]. symmetry. pose proof (h14 x) as h15. hfirst h15. exists e; auto. specialize (h15 hf1). subst. symmetry. apply h14. exists (in_nth_tuple_prod_ens_dep (proj2_sig ta1) k hk). apply app_list_nprod_to_list_inj in H9; subst. pose proof (nth_prod_sig_tuple_prod_ens_dep_const_to (Term_algebra_ens sig v) (projT2 (fst (sig i pfop))) ttm0 k hk) as h15. terml h15. termr h15. gterm0. symmetry in h15. pose proof (Recursion_homo_aux_functional a (nth_prod k hk (projT2 (fst (sig i pfop)))) (nth_prod k hk (projT2 (fst (sig i pfop)))) y x0 (nth_tuple_prod (proj1_sig ta1) k hk) (nth_tuple_prod (proj1_sig ta1) k hk) c c (eq_refl _)) as h16. unfold x0, y in h16. specialize (h16 h15 (eq_refl _)). unfold c in h16. apply h16; clear h16. pose proof (H10 k hk) as h17. destruct h17 as [h17|h17]. inversion h17; clear h17. rewrite (transfer_dep_eq_iff _ _ H12) in H13. pose proof (Recursion_homo_aux_functional a (v x1) (nth_prod k hk (projT2 (fst (sig i pfop)))) (inl x1::nil) (inl x1::nil) (proj1_sig (a x1)) (nth_tuple_prod (proj1_sig ta1) k hk) (proj2_sig (a x1)) (in_nth_tuple_prod_ens_dep (proj2_sig ta1) k hk) H4 (eq_refl _)) as h18. hfirst h18. symmetry. rewrite <- transfer_dep_r_eq_iff; auto. specialize (h18 hf2). apply h18. constructor. destruct h17 as [hn [op [hop [h19 [h20 h21]]]]]. term0 h21. eapply (Recursion_homo_aux_functional a (snd (sig op hop)) (nth_prod k hk (projT2 (fst (sig i pfop)))) (nth_tuple_prod (proj1_sig ttm0) k hk) (nth_tuple_prod (proj1_sig ttm0) k hk) (transfer_dep h19 (nth_tuple_prod (proj1_sig ta1) k hk))); auto. erewrite transfer_dep_transfer_dep_r_compat. reflexivity. constructor. apply h21. eapply terms_app_list_eq_length; auto. intros j hj. rewrite (nth_prod_sig_tuple_prod_ens_dep_const_to _ _ ttm0). pose proof (proj2_sig ttm0) as h15. simpl in h15. destruct h15 as [h15]; auto. Grab Existential Variables. auto. Qed. Lemma recursion_homo_aux_well_defined : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) (s:S) (t:term X num_ops) (pft:Inn (Term_algebra_ens sig v s) t) u u' (pfu:Inn (A_dep_t sig A s) u) (pfu':Inn (A_dep_t sig A s) u'), Recursion_homo_aux a s t u pfu -> Recursion_homo_aux a s t u' pfu' -> u = u'. intros S X num_ops sig A v a s t ht u u' hu hu' h1 h2. pose proof (recursion_homo_ex a s t ht) as h3. destruct h3 as [c [h4 h5]]. destruct h4 as [h4 h6]. pose proof (h5 u) as h7. hfirst h7. exists hu; auto. specialize (h7 hf); subst. apply h5. exists hu'; auto. Qed. Lemma recursion_homo_aux_op_well_defined : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) op (pfop:op < num_ops) (t:term X num_ops) (pft:Inn (Term_algebra_ens sig v (snd (sig op pfop))) t) u u' (pfu:Inn (A_dep_t sig A (snd (sig op pfop))) u) (pfu':Inn (A_dep_t sig A (snd (sig op pfop))) u'), Recursion_homo_aux_op a op pfop t u pfu -> Recursion_homo_aux_op a op pfop t u' pfu' -> u = u'. intros S X num_ops sig A v a op pfop t ht u u' hu hu' h1 h2. pose proof (recursion_homo_aux_op a op pfop t u hu h1) as h3. pose proof (recursion_homo_aux_op a op pfop t u' hu' h2) as h4. pose proof (recursion_homo_aux_well_defined _ _ _ ht _ _ _ _ h3 h4); auto. Qed. Definition recursion_homo {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) (s:S) (t:term X num_ops) (pft:Inn (Term_algebra_ens sig v s) t) : U_dep_t sig A s := proj1_sig (constructive_definite_description _ (recursion_homo_ex a s t pft)). Lemma recursion_homo_compat : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) (s:S) (t:term X num_ops) (pft:Inn (Term_algebra_ens sig v s) t), exists pfu:Inn (A_dep_t sig A s) (recursion_homo a s t pft), Recursion_homo_aux a s t (recursion_homo a s t pft) pfu. unfold recursion_homo; intros; destruct constructive_definite_description; auto. Qed. Lemma in_recursion_homo : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) (s:S) (t:term X num_ops) (pft:Inn (Term_algebra_ens sig v s) t), Inn (A_dep_t _ A s) (recursion_homo a s t pft). intros S X num_ops sig A v a s t ht. pose proof (recursion_homo_compat a s t ht) as h1. destruct h1; auto. Qed. Lemma recursion_homo_eq_iff : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) (s:S) (t:term X num_ops) (pft:Inn (Term_algebra_ens sig v s) t) (u:U_dep_t _ A s), recursion_homo a s t pft = u <-> exists (pfu:Inn (A_dep_t _ A s) u), Recursion_homo_aux a s t u pfu. intros S X num_ops sig A v a s t ht u. pose proof (recursion_homo_compat a s t ht) as h1. destruct h1 as [h1 h2]. split; intro h3; subst. exists h1; auto. destruct h3 as [h3 h4]. eapply recursion_homo_aux_well_defined. apply ht. apply h2. apply h4. Qed. Lemma recursion_homo_vx : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) x (pfin:Inn (Term_algebra_ens sig v (v x)) (inl x :: nil)), recursion_homo a (v x) (inl x :: nil) pfin = proj1_sig (a x). intros S X num_ops sig A v a x h1. pose proof (recursion_homo_compat a (v x) (inl x ::nil) h1) as h2. destruct h2 as [h2 h3]. inversion h3; clear h3; subst. clear H3. apply existT_injective2 in H2; auto. inversion H2. Qed. Lemma recursion_homo_nil_op : forall {S X num_ops} {sig:signature S num_ops} {v:X->S} {op} {pfop:op < num_ops} (A:TAlgebra_dep sig) (a:assment_dep (A_dep_t sig A) v) s (pfe:projT1 (fst (sig op pfop)) = 0) (pfint:Inn (Term_algebra_ens sig v s) (inr (exist (fun c3 : nat => c3 < num_ops) op pfop) :: nil)) (pfe':snd (sig op pfop) = s), recursion_homo a s (inr (exist (fun c3 : nat => c3 < num_ops) op pfop) :: nil) pfint = transfer_dep pfe' (proj1_sig (ops_dep_t sig A op pfop (exist (fun c => Inn (tuple_prod_ens_dep (A_dep_t sig A) (projT2 (fst (sig op pfop)))) c) _ (in_transfer_tt_tuple_prod_ens_dep (U_dep_t sig A) (A_dep_t sig A) (projT2 (fst (sig op pfop))) pfe)))). intros S X num_ops sig v op hop A a x h1 h2 h3; subst. rewrite transfer_dep_eq_refl. pose proof (recursion_homo_compat a (snd (sig op hop)) (inr (exist (fun c => c < num_ops) op hop) :: nil) h2) as h4. destruct h4 as [h4 h5]. inversion h5; clear h5; subst. inversion H2; clear H2. clear s n. unfold u0. unfold u0, pfintt0, w0, w in H1, H8. clear u0 pfintt0 w0 w. subst. pose proof (proof_irrelevance _ pfop hop); subst. pose proof (proof_irrelevance _ h1 pfeq0); subst. rewrite transfer_dep_eq_iff in H5. rewrite <- H5. rewrite transfer_dep_eq_refl. reflexivity. clear tm0. unfold ttm'0 in H7. clear ttm'0 s ttm'. unfold opta, w0, w, n in H6, H9, ttm, ta0, ttm0, H8, H7. clear opta w0 w n. subst. rewrite transfer_dep_eq_iff in H5. rewrite <- H5. pose proof (proof_irrelevance _ hop pfop); subst. pose proof (@tuple_prod_n0_eq _ _ _ _ (proj1_sig ta0) (transfer_tt_tuple_prod (U_dep_t sig A) h1 (projT2 (fst (sig op pfop)))) h1) as h5. generalize (proj2_sig ta0). generalize h1. rewrite h5. intros. rewrite transfer_dep_eq_refl. repeat f_equal. apply proj1_sig_injective; simpl; auto. f_equal. apply proof_irrelevance. Grab Existential Variables. repeat f_equal. apply proof_irrelevance. reflexivity. Qed. Definition recursion_homo_Sfun_dep {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) : Sfun_dep (Term_algebra_ens sig v) (A_dep_t sig A) := fun s (t:term X num_ops) (pft:Inn (Term_algebra_ens sig v s) t) => let (pfu, _) := constructive_definite_proof (recursion_homo_compat a s t pft) in exist (fun c => Inn (A_dep_t sig A s) c) (recursion_homo a s t pft) pfu. Lemma recursion_extends : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) (x:X), recursion_homo_Sfun_dep a (v x) (inl x::nil) (term_algebra_ensX sig v x) = a x. intros S X num_ops sig A v a. unfold recursion_homo_Sfun_dep. intro x. unfold constructive_definite_proof; simpl. pose proof (recursion_homo_compat a (v x) (inl x::nil) (term_algebra_ensX sig v x)) as h1. destruct h1 as [h1 h2]. pose proof (recursion_homo_aux_var a x) as h3. apply proj1_sig_injective; simpl. apply (recursion_homo_aux_well_defined _ _ _ (term_algebra_ensX _ _ _) _ _ _ _ h2 h3). Qed. Lemma homo_recursion : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v), Thomo_dep (recursion_homo_Sfun_dep a) (A := Term_algebra_dep (type_eq_decS_t _ A) sig v) (B := A). intros S X num_ops sig A v a. constructor. intros i hi; simpl. intros a'. destruct a' as [p hp]. pose proof (recursion_homo_compat a _ _ (proj2_sig (Term_algebra_dep' sig v i hi (exist _ _ hp)))) as h1. destruct h1 as [h1 h2]. unfold recursion_homo_Sfun_dep. unfold constructive_definite_proof. destruct constructive_definite_description. apply proj1_sig_injective; simpl. pose proof (proof_irrelevance _ x h1); subst. rename r into h10. inversion h10; clear h10; subst. inversion H2; clear H2. clear H9. unfold u0 in H1; clear H1 u0 s n pfintt0 w0 w. subst. unfold Term_algebra_op_aux_dep, op_term_to_app_term. termr H7. redterm0 y. redterm0 c. gterml. redterm1 x. redterm0 c1. redterm0 c2. redterm0 c3. pose proof (proof_irrelevance _ c4 c0) as h8. gen0. intro h10; simpl in h10. pose proof h10 as h10'. unfold op_term_to_app_term in h10'. rewrite <- H7 in h10'. pose proof (f_equal (fun c => inr (exist (fun c => c < num_ops) i hi)::c) H7) as h12. pose proof (subsetT_eq_compat _ _ _ _ h10' h10 h12) as h13. simpl in h13; simpl. unfold op_term_to_app_term in h13. dependent rewrite <- h13. pose proof (proof_irrelevance _ pfop hi); subst. erewrite recursion_homo_nil_op. rewrite transfer_dep_eq_refl. repeat f_equal. apply proj1_sig_injective; simpl. apply tuple_prod_n0_eq; auto. clear tm0. unfold opta, ttm'0, w0, w, n in H7, ta0, ttm0, H8, H6, H9, ttm. clear opta ttm'0 s ttm' w0 w n. subst. pose proof (proof_irrelevance _ pfop hi); subst. apply existT_injective2 in H5. rewrite <- H5. repeat f_equal. apply proj1_sig_injective; simpl. symmetry. rewrite proj1_sig_map_assment_w_dep_eq_iff. intros k hk; simpl. rewrite recursion_homo_eq_iff. exists (in_nth_tuple_prod_ens_dep (proj2_sig ta0) k hk). rename H7 into h12. unfold op_term_to_term in h12. apply app_list_nprod_to_list_inj in h12; simpl in h12. unfold sig_tuple_prod_ens_dep_const_to in h12; simpl in h12. destruct ttm0 as [q hq]; simpl in h12. apply inj_transfer_tuple_prod_const in h12; subst. specialize (H8 k hk); simpl in H8. destruct H8 as [h10 |h10]. inversion h10; clear h10. pose proof (recursion_homo_aux_var a x) as h11. gen0. dependent rewrite <- H11. intro h12. pose proof (proof_irrelevance _ h12 (proj2_sig (a _))); subst. constructor. destruct h10 as [h10 [op [hop [h13 h14]]]]. destruct h14 as [h14 h15]. pose proof (recursion_homo_aux_op a op hop (nth_tuple_prod p k hk) (transfer_dep h13 (nth_tuple_prod (proj1_sig ta0) k hk)) _ h15) as h16. pose proof (Recursion_homo_aux_functional a (snd (sig op hop)) (nth_prod k hk (projT2 (fst (sig i hi)))) (nth_tuple_prod p k hk) (nth_tuple_prod p k hk) (transfer_dep h13 (nth_tuple_prod (proj1_sig ta0) k hk)) (nth_tuple_prod (proj1_sig ta0) k hk)) as h17. eapply h17; auto. rewrite (transfer_dep_transfer_dep_r_compat h13 (eq_sym h13)). reflexivity. apply h16. intros; simpl. eapply terms_app_list_eq_length. intros j hj. pose proof (proj2_sig ttm0) as h15. simpl in h15. destruct h15 as [h15]. specialize (h15 j hj). rewrite <- (nth_prod_sig_tuple_prod_ens_dep_const_to _ _ ttm0) in h15. apply h15. intros j hj. pose proof hp as h16. destruct h16 as [h16]. rewrite nth_prod_transfer_tuple_prod_const; auto. auto. Grab Existential Variables. reflexivity. auto. Qed. Lemma transfer_dep_nth_tuple_prod_proj1_map_assment_w_dep : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} {op} {pfop:op < num_ops} (a:assment_dep (A_dep_t sig A) v) (a':assment_w_dep (Term_algebra_ens sig v) (projT2 (fst (sig op pfop)))) j (hj:j < projT1 (fst (sig op pfop))) x (pfe:nth_prod j hj (projT2 (fst (sig op pfop))) = v x), transfer_dep pfe (nth_tuple_prod (proj1_sig (map_assment_w_dep (recursion_homo_Sfun_dep a) a')) j hj) = recursion_homo a (v x) (nth_tuple_prod (proj1_sig a') j hj) (in_nth_tuple_prod_term_algebra a' j hj x pfe). unfold recursion_homo_Sfun_dep; intros. rewrite nth_tuple_prod_map_assment_w_dep. unfold constructive_definite_proof; simpl. rewrite <- transfer_dep_eq_iff. gterml. gtermr. redterm0 x0. redterm0 c. redterm0 y. redterm0 c1. fold c2 c0. generalize c0, c2; clear c0 c2. clear c1 c y x0. rewrite pfe. intros; repeat f_equal. apply proof_irrelevance. Qed. Lemma transfer_dep_nth_tuple_prod_proj1_map_assment_w_dep' : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} {op} {pfop:op < num_ops} (a:assment_dep (A_dep_t sig A) v) (a':assment_w_dep (Term_algebra_ens sig v) (projT2 (fst (sig op pfop)))) x j (pfj:j < projT1 (fst (sig op pfop))) (pfe:nth_prod j pfj (projT2 (fst (sig op pfop))) = v x), transfer_dep pfe (nth_tuple_prod (proj1_sig (map_assment_w_dep (fun (s : S) (t : term X num_ops) (pft : Inn (Term_algebra_ens sig v s) t) => exist (fun c1 : U_dep_t sig A s => Inn (A_dep_t sig A s) c1) (recursion_homo a s t pft) (proj1_sig (constructive_definite_description (Recursion_homo_aux a s t (recursion_homo a s t pft)) (propositional_ex_unq (Inn (A_dep_t sig A s) (recursion_homo a s t pft)) (Recursion_homo_aux a s t (recursion_homo a s t pft)) (recursion_homo_compat a s t pft))))) a')) j pfj) = recursion_homo a (v x) (nth_tuple_prod (proj1_sig a') j pfj) (in_nth_tuple_prod_term_algebra a' j pfj x pfe). intros; rewrite <- transfer_dep_nth_tuple_prod_proj1_map_assment_w_dep; auto. Qed. Lemma recursion_homo_aux_pred_op : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v) i (pfi:i < num_ops) (a':assment_w_dep (Term_algebra_ens sig v) (projT2 (fst (sig i pfi)))) (pfr:Inn (A_dep_t sig A (snd (sig i pfi))) (recursion_homo a (snd (sig i pfi)) (proj1_sig (Term_algebra_dep' sig v i pfi a')) (proj2_sig (Term_algebra_dep' sig v i pfi a')))), Recursion_homo_aux a (snd (sig i pfi)) (proj1_sig (Term_algebra_dep' sig v i pfi a')) (recursion_homo a (snd (sig i pfi)) (proj1_sig (Term_algebra_dep' sig v i pfi a')) (proj2_sig (Term_algebra_dep' sig v i pfi a'))) pfr <-> forall j (pfj:j < projT1 (fst (sig i pfi))), (exists x:X, v x = nth_prod j pfj (projT2 (fst (sig i pfi)))) \/ (exists op (pfop:op < num_ops), snd (sig op pfop) = nth_prod j pfj (projT2 (fst (sig i pfi)))). intros S X num_ops sig A v a i hi a' hr. split; intro h1. inversion h1; clear h1. intros j hj. left. exists x. unfold Term_algebra_dep' in H. destruct a' as [y z]. simpl in H. unfold Term_algebra_op_aux_dep, op_term_to_app_term in H. discriminate. intros j hj. unfold Term_algebra_dep' in H5; simpl in H5, H3. unfold Term_algebra_op_aux_dep, op_term_to_app_term in H5, H3, H2, hr. destruct a' as [p hp]. subst. simpl in H5, H3, H2, hr. clear H4 H3. inversion H2; clear H2. clear s n. unfold u0, pfintt0, w0, w in H1, H7. clear u0 pfintt0 w0 w. subst. pose proof (proof_irrelevance _ hi pfop); subst. clear H0. right. pose proof hj as hj'. rewrite pfeq in hj'. contradict (lt_n_0 _ hj'). clear H7. unfold opta, w0, ttm'0, tm0, n in H3, H4, ta0, ttm0, H6, H4. clear opta w0 ttm'0 tm0 s ttm' n. unfold w in ttm. clear w. subst. pose proof (proof_irrelevance _ hi pfop); subst. specialize (H6 j hj). destruct H6 as [h5 | h5]. inversion h5; clear h5. left. exists x; auto. destruct h5 as [h5 [op [hop [h8 h9]]]]. right. exists op, hop; auto. pose proof (recursion_homo_compat a (snd (sig i hi)) (proj1_sig (Term_algebra_dep' sig v i hi a')) (proj2_sig _)) as h2. destruct h2. pose proof (proof_irrelevance _ x hr);subst; auto. Qed. Theorem RecursionTheorem : forall {S X num_ops} {sig:signature S num_ops} {A:TAlgebra_dep sig} {v:X->S} (a:assment_dep (A_dep_t sig A) v), exists! (f:Sfun_dep (Term_algebra_ens sig v) (A_dep_t sig A)), Thomo_dep f (A := Term_algebra_dep (type_eq_decS_t _ A) sig v) (B := A) /\ forall x:X, let pfx := term_algebra_ensX sig v x in let fx := proj1_sig (f _ _ pfx) in let ax := proj1_sig (a x) in fx = ax. intros S X num_ops sig A v a. exists (recursion_homo_Sfun_dep a). red; split. split. apply homo_recursion. intro; f_equal; apply recursion_extends. intros f h1. destruct h1 as [h1 h2]. eapply (homos_eq_on_generators (A := Term_algebra_dep (type_eq_decS_t _ A) sig v) (B := A)); auto. apply homo_recursion. eapply pre_sys_generates_term_algebra'. intros s u hin. destruct hin as [x]; subst. destruct i as [i]; subst. erewrite f_equal_dep. rewrite recursion_extends. rewrite <- h2; repeat f_equal; apply proof_irrelevance. reflexivity. intros op hop n w s he; simpl. destruct h1 as [h1]. specialize (h1 op hop); simpl in h1. pose proof (in_transfer_tt_tuple_prod_ens_dep (U_dep_t sig A) (A_dep_t sig A) w he) as h3. pose proof (in_transfer_tt_tuple_prod_ens_dep (fun _ => term X num_ops) (Term_algebra_ens sig v) w he) as h4. specialize (h1 (exist (fun c => Inn (tuple_prod_ens_dep (Term_algebra_ens sig v) w) c) _ h4)). unfold s, n, w; unfold w in h4, h1, h3; clear s n w. unfold Term_algebra_dep' in h1; simpl in h1. unfold recursion_homo_Sfun_dep. unfold constructive_definite_proof. destruct constructive_definite_description; simpl. apply proj1_sig_injective; simpl. pose proof (recursion_homo_nil_op A a (snd (sig op hop)) he) as h12. unfold Term_algebra_op_aux_dep, op_term_to_app_term, op_term_to_term; simpl. gterml. redterm1 x0. redterm0 c. assert (h8:c0 = nil). unfold c0. rewrite app_list_eq_nil_iff. left. gterml. redterm0 x1. pose proof (nprod_to_list_eq_nil_iff c1) as h9. fold c1. rewrite h9 at 1; auto. fold c0. pose proof (f_equal (cons (inr (exist (fun q => q < num_ops) op hop))) h8) as h9. redterm0 x0. fold c1. pose proof c1 as hc1; simpl in hc1. unfold Term_algebra_op_aux_dep, op_term_to_app_term, op_term_to_term in hc1, c1; simpl in hc1, c1. fold c0 in hc1. rewrite h9 in hc1. pose proof (subsetT_eq_compat _ _ _ _ c1 hc1 h9) as h10. fold c0 c1 in h10. dependent rewrite -> h10. terml h1. redterm0 x1. specialize (h12 hc1 (eq_refl _)). rewrite h12 at 1. rewrite transfer_dep_eq_refl. unfold c2, Term_algebra_op_aux_dep, op_term_to_app_term, op_term_to_term in h1, c2; simpl in h1, c2. fold c0 c2 in h1. pose proof (proof_irrelevance _ c1 c2) as h20. rewrite <- h20 in h1. dependent rewrite -> h10 in h1. rewrite h1. repeat f_equal. apply proj1_sig_injective; simpl. apply tuple_prod_n0_eq; auto. Qed. (*Existence equation.*) Inductive ex_eq {S X num_ops} (v:X->S) (t t':term X num_ops) : Prop := ex_eq_intro : forall (sig:signature S num_ops) s, Inn (Term_algebra_ens sig v s) t -> Inn (Term_algebra_ens sig v s) t' -> ex_eq v t t'. Definition ex_eq_ens {S X} num_ops (v:X->S) : Type := Ensemble {pr:(term X num_ops) * (term X num_ops) | ex_eq v (fst pr) (snd pr)}. Definition dom_cond {S X} num_ops (v:X->S) : Type := ex_eq_ens num_ops v. (*Equational partial signature.*) Definition ep_signature {X S} (v:X->S) num_ops : Type := seg_fun num_ops ((all_tuples S * S) * (dom_cond num_ops v))%type. Definition ep_signature_to_sig {X S} {v:X->S} {num_ops} (theta:ep_signature v num_ops) : signature S num_ops := fun op (pfop:op < num_ops) => fst (theta op pfop). Inductive Interpretation_aux {S X U num_ops} {sig:signature S num_ops} {A:PAlgebra_ens U sig} {v:X->S} (a:assment_ens_sig (A_sigma_ens U sig A) v) : forall s:S, term X num_ops -> forall u:U, Inn (A_sigma_ens U sig A s) u -> Prop := | interpretationX : forall x:X, Interpretation_aux a (v x) (inl x::nil) (proj1_sig (a x)) (proj2_sig (a x)) | interpretation_app : forall op (pfop: op < num_ops), let w := projT2 (fst (sig op pfop)) in let n := projT1 (fst (sig op pfop)) in let s := snd (sig op pfop) in forall (t:(term X num_ops)^n) (pft:Inn (Term_algebra_at_op sig v op pfop) t) (ufun:seg_fun n U), (forall j (pfj:j < n), exists pfu:Inn (A_sigma_ens U sig A (nth_prod j pfj w)) (ufun j pfj), Interpretation_aux a (nth_prod j pfj w) (nth_prod j pfj t) (ufun j pfj) pfu) -> let aw := map_nprod ufun in forall (pfaw:Inn (doms_ens U sig A op pfop) aw), let u := ops_sigma_ens U sig A op pfop aw pfaw in Interpretation_aux a s (op_term_to_app_term (exist (fun c => Inn (Term_algebra_at_op sig v op pfop) c) t pft)) (proj1_sig u) (proj2_sig u). Inductive Interpretation {S X U num_ops} {sig:signature S num_ops} {A:PAlgebra_ens U sig} {v:X->S} (a:assment_ens_sig (A_sigma_ens U sig A) v) (s:S) (t:term X num_ops) (u:U) (pfu:Inn (A_sigma_ens U sig A s) u) : Prop := | interpretation_aux_intro : Interpretation_aux a s t u pfu -> (forall u' (pfu':Inn (A_sigma_ens U sig A s) u'), Interpretation_aux a s t u' pfu' -> u = u') -> Interpretation a s t u pfu. Lemma Interpretation_well_defined : forall {S X U num_ops} {sig:signature S num_ops} (v:X->S) (A:PAlgebra_ens U sig) (a:assment_ens_sig (A_sigma_ens U sig A) v) (s:S) (t u:term X num_ops) (y z: U) (pfy:Inn (A_sigma_ens U sig A s) y) (pfz:Inn (A_sigma_ens U sig A s) z), t = u -> Interpretation a s t y pfy -> Interpretation a s u z pfz -> y = z. intros S X U num_ops sig v A a s t u y z hy hz h1 h2 h3; subst. destruct h3, h2. specialize (H0 _ hy H1); auto. Qed. Inductive Interpretation_dom {S X U num_ops} {sig:signature S num_ops} {A:PAlgebra_ens U sig} {v:X->S} (a:assment_ens_sig (A_sigma_ens U sig A) v) : forall (s:S), Ensemble (term X num_ops) := | interpretation_domX : forall x, Inn (Interpretation_dom a (v x)) (inl x::nil) | interpretation_dom_app : forall op (pfop: op < num_ops), let w := projT2 (fst (sig op pfop)) in let n := projT1 (fst (sig op pfop)) in let s := snd (sig op pfop) in forall (t:(term X num_ops)^n) (pft:Inn (Term_algebra_at_op sig v op pfop) t) (ufun:seg_fun n U), let tm := op_term_to_app_term (exist (fun c => Inn (Term_algebra_at_op sig v op pfop) c) t pft) in (forall j (pfj:j < n), Inn (Interpretation_dom a (nth_prod j pfj w)) (nth_prod j pfj t)) -> forall (pfin:Inn (doms_ens U sig A op pfop) (map_nprod ufun)), let u := ops_sigma_ens U sig A op pfop _ pfin in Interpretation a s tm (proj1_sig u) (proj2_sig u) -> Inn (Interpretation_dom a s) tm. Lemma interpretation_ex : forall {S X U num_ops} {sig:signature S num_ops} {A:PAlgebra_ens U sig} {v:X->S} (a:assment_ens_sig (A_sigma_ens U sig A) v) (s:S) (t:term X num_ops), Inn (Interpretation_dom a s) t -> exists! u:U, exists (pfu:Inn (A_sigma_ens U sig A s) u), Interpretation a s t u pfu. intros S X U num_ops sig A v a s t h1. inversion h1; subst. exists (proj1_sig (a x)). red; split. exists (proj2_sig _). constructor. constructor. intros u' hu' h2. inversion h2; subst; auto. intros u h2. destruct h2 as [h2 h3]. inversion h3. inversion H; subst; auto. subst. exists (proj1_sig (ops_sigma_ens U sig A op pfop _ pfin)). red; split. exists (proj2_sig _). assumption. intros u' hu'. destruct hu' as [h3 h4]. inversion h4. inversion H0. specialize (H2 _ (proj2_sig u) H3); subst; auto. Qed. Definition interpretation {S X U num_ops} {sig:signature S num_ops} {A:PAlgebra_ens U sig} {v:X->S} (a:assment_ens_sig (A_sigma_ens U sig A) v) (s:S) : fun_in_ens (Interpretation_dom a s) U := fun (t:term X num_ops) (pfin:Inn (Interpretation_dom a s) t) => proj1_sig (constructive_definite_description _ (interpretation_ex a s t pfin)). Lemma interpretation_compat : forall {S X U num_ops} {sig:signature S num_ops} {A:PAlgebra_ens U sig} {v:X->S} (a:assment_ens_sig (A_sigma_ens U sig A) v) (s:S) (t:term X num_ops) (pfin:Inn (Interpretation_dom a s) t), exists (pffi:Inn (A_sigma_ens U sig A s) (interpretation a s t pfin)), Interpretation a s t (interpretation a s t pfin) pffi. Proof. unfold interpretation; intros; apply proj2_sig. Qed. Lemma in_interpretation : forall {S X U num_ops} {sig:signature S num_ops} {A:PAlgebra_ens U sig} {v:X->S} (a:assment_ens_sig (A_sigma_ens U sig A) v) (s:S) (t:term X num_ops) (pfin:Inn (Interpretation_dom a s) t), Inn (A_sigma_ens U sig A s) (interpretation a s t pfin). intros; pose proof (interpretation_compat a s t pfin) as h1; simpl in h1; destruct h1; auto. Qed. bool2-v0-3/Homomorphisms.v0000644000175000017500000033766613575574055016422 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2015, Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import BoolAlgBasics. Require Import FunctionProperties. Require Import Subalgebras. Require Import SetUtilities. Require Import FunctionalExtensionality. Require Import LogicUtilities. Require Import Description. Require Import FieldsOfSets. Require Import TypeUtilities. Require Import Basics. Require Import FiniteOperations. Require Import ListUtilities. Require Import DecidableDec. Record homomorphism {A B:Bool_Alg} (f:bt A->bt B) : Prop := { homo_times : forall x y, f (x*y) = (f x) * (f y); homo_plus : forall x y, f (x+y) = (f x) + (f y); homo_comp : forall x, f (-x) = - (f x) }. Record homomorphism_general {A:Bool_Alg} {T:Type} (f:(bt A)->T) (homo_ut : T -> T -> T) (homo_up : T -> T -> T) (homo_uc : T -> T) : Prop := { homo_gen_times : forall x y, f (x*y) = homo_ut (f x) (f y); homo_gen_plus : forall x y, f (x+y) = homo_up (f x) (f y); homo_gen_comp : forall x, f (-x) = homo_uc (f x) }. Lemma homo_sym_diff : forall {A B:Bool_Alg} (f:bt A->bt B), homomorphism f -> forall x y, f (x /_\ y) = (f x) /_\ (f y). intros A B f h1 x y. unfold sym_diff. destruct h1 as [h1a h1b h1c]. rewrite h1b. do 2 rewrite h1a. do 2 rewrite h1c. reflexivity. Qed. Lemma homo_id : forall A:Bool_Alg, homomorphism id (A:=A). intros; constructor; auto. Qed. Lemma universal_homomorphic_ops_imply_bool : forall {A:Bool_Alg} {T:Type} (f:(bt A)->T) (ut:T->T->T) (up:T->T->T) (uc:T->T) (pf:homomorphism_general f ut up uc), surjective f -> exists! B:Bool_Alg, Bc B = Build_Bconst T (Full_set T) up ut (f 1) (f 0) uc. intros A T f ut up uc h1 h2. red in h2. destruct h1 as [h1a h1b h1c]. pose (Build_Bconst T (Full_set T) up ut (f 1) (f 0) uc) as Bc'. assert (my_und_set : BS Bc' = Full_set (Btype Bc')). simpl. reflexivity. assert (my_assoc_sum : forall n m p : Btype Bc', n + (m + p) = n + m + p). intros p q r. simpl. pose proof (h2 p) as h3. pose proof (h2 q) as h4. pose proof (h2 r) as h5. destruct h3 as [u h3]. destruct h4 as [v h4]. destruct h5 as [w h5]. subst. rewrite <- h1b at 1. rewrite <- h1b at 1. rewrite assoc_sum. rewrite h1b. rewrite h1b. reflexivity. assert (my_assoc_prod : forall n m p : Btype Bc', n * (m * p) = n * m * p). intros p q r. simpl. pose proof (h2 p) as h3. pose proof (h2 q) as h4. pose proof (h2 r) as h5. destruct h3 as [u h3]. destruct h4 as [v h4]. destruct h5 as [w h5]. subst. rewrite <- h1a at 1. rewrite <- h1a at 1. rewrite assoc_prod. rewrite h1a. rewrite h1a. reflexivity. assert (my_comm_sum : forall n m : Btype Bc', n + m = m + n). intros p q. simpl. pose proof (h2 p) as h3. pose proof (h2 q) as h4. destruct h3 as [u h3]. destruct h4 as [v h4]. subst. rewrite <- h1b. rewrite comm_sum. rewrite h1b. reflexivity. assert (my_comm_prod : forall n m : Btype Bc', n * m = m * n). intros p q. simpl. pose proof (h2 p) as h3. pose proof (h2 q) as h4. destruct h3 as [u h3]. destruct h4 as [v h4]. subst. rewrite <- h1a. rewrite comm_prod. rewrite h1a. reflexivity. assert (my_abs_sum : forall n m : Btype Bc', n + n * m = n). intros p q. simpl. pose proof (h2 p) as h3. pose proof (h2 q) as h4. destruct h3 as [u h3]. destruct h4 as [v h4]. subst. rewrite <- h1a. rewrite <- h1b. rewrite abs_sum. reflexivity. assert (my_abs_prod : forall n m : Btype Bc', n * (n + m) = n). intros p q. simpl. pose proof (h2 p) as h3. pose proof (h2 q) as h4. destruct h3 as [u h3]. destruct h4 as [v h4]. subst. rewrite <- h1b. rewrite <- h1a. rewrite abs_prod. reflexivity. assert (my_dist_sum : forall n m p : Btype Bc', p * (n + m) = p * n + p * m). intros p q r. pose proof (h2 p) as h3. pose proof (h2 q) as h4. pose proof (h2 r) as h5. destruct h3 as [u h3]. destruct h4 as [v h4]. destruct h5 as [w h5]. subst. simpl. rewrite <- h1b at 1. rewrite <- h1a at 1. rewrite dist_sum. rewrite h1b at 1. do 2 rewrite h1a at 1. reflexivity. assert (my_dist_prod : forall n m p : Btype Bc', p + (n * m) = (p + n) * (p + m)). intros p q r. pose proof (h2 p) as h3. pose proof (h2 q) as h4. pose proof (h2 r) as h5. destruct h3 as [u h3]. destruct h4 as [v h4]. destruct h5 as [w h5]. subst. simpl. rewrite <- h1a at 1. rewrite <- h1b at 1. rewrite dist_prod. rewrite h1a at 1. do 2 rewrite h1b at 1. reflexivity. assert (my_comp_sum : forall n : Btype Bc', n + - n = 1). intro n. pose proof (h2 n) as h3. destruct h3 as [u h3]. subst. simpl. rewrite <- h1c. rewrite <- h1b. rewrite comp_sum. reflexivity. assert (my_comp_prod : forall n : Btype Bc', n * - n = 0). intro n. pose proof (h2 n) as h3. destruct h3 as [u h3]. subst. simpl. rewrite <- h1c. rewrite <- h1a. rewrite comp_prod. reflexivity. exists (Build_Bool_Alg _ my_und_set my_assoc_sum my_assoc_prod my_comm_sum my_comm_prod my_abs_sum my_abs_prod my_dist_sum my_dist_prod my_comp_sum my_comp_prod). red. split. simpl. reflexivity. intros B h1. apply bc_inj. simpl. rewrite h1. reflexivity. Qed. Lemma homo_zero : forall {A B:Bool_Alg} (f:bt A->bt B), homomorphism f -> f 0 = 0. intros A B f h1. destruct h1 as [h1 h2 h3]. pose proof (comp_prod A 0) as h4. pose proof (f_equal f (eq_refl (0*-0))) as h5. rewrite h4 in h5 at 1. rewrite h1 in h5. rewrite h3 in h5. rewrite comp_prod in h5. assumption. Qed. Lemma homo_one : forall {A B:Bool_Alg} (f:bt A->bt B), homomorphism f -> f 1 = 1. intros A B f h1. destruct h1 as [h1 h2 h3]. pose proof (comp_sum A 1) as h4. pose proof (f_equal f (eq_refl ((1+(-(1)))))) as h5. rewrite h4 in h5 at 1. rewrite h2 in h5. rewrite h3 in h5. rewrite comp_sum in h5. assumption. Qed. Lemma homo_two_ops_plus : forall {A B:Bool_Alg} (f:bt A->bt B), (forall x y, f (x + y) = (f x) + (f y)) -> (forall x, f (-x) = - (f x)) -> homomorphism f. intros A B f h1 h2. constructor; auto. intros x y. rewrite (doub_neg (f x)). rewrite (doub_neg (f y)). rewrite <- de_mor_sum. rewrite <- h2. rewrite <- h2. rewrite <- h1. rewrite <- h2. rewrite de_mor_sum. do 2 rewrite <- doub_neg. reflexivity. Qed. Lemma homo_two_ops_times : forall {A B:Bool_Alg} (f:bt A->bt B), (forall x y, f (x * y) = (f x) * (f y)) -> (forall x, f (-x) = - (f x)) -> homomorphism f. intros A B f h1 h2. constructor; auto. intros x y. rewrite (doub_neg (f x)). rewrite (doub_neg (f y)). rewrite <- de_mor_prod. rewrite <- h2. rewrite <- h2. rewrite <- h1. rewrite <- h2. rewrite de_mor_prod. do 2 rewrite <- doub_neg. reflexivity. Qed. Lemma homo_times_list : forall {A B:Bool_Alg} (f:bt A->bt B), homomorphism f -> forall (l:list (bt A)), NoDup l -> f (times_list l) = times_list (map f l). intros A B f h1 l. induction l as [|a l h2]. intros; simpl; auto. apply homo_one; auto. intro h3; simpl. rewrite homo_times; auto. f_equal. apply h2. apply (no_dup_cons l a). assumption. Qed. Lemma homo_plus_list : forall {A B:Bool_Alg} (f:bt A->bt B), homomorphism f -> forall (l:list (bt A)), NoDup l -> f (plus_list l) = plus_list (map f l). intros A B f h1 l. induction l as [|a l h2]. intros; simpl; auto. apply homo_zero; auto. intro h3; simpl. rewrite homo_plus; auto. f_equal. apply h2. apply (no_dup_cons l a). assumption. Qed. Lemma homo_times_set : forall {A B:Bool_Alg} (f:bt A->bt B), homomorphism f -> forall (E:Ensemble (bt A)) (pf:Finite E), f (times_set E pf) = times_set (Im E f) (finite_image _ _ _ f pf). intros A B f h1 E h2. pose proof (finite_set_list_no_dup _ h2) as h3. destruct h3 as [l [h3 h4]]. subst. pose proof (times_set_compat' (list_to_set l) _ h2 (eq_refl _)) as h3. rewrite h3 at 1. pose proof (map_im_compat f l) as h6. symmetry in h6. pose proof (times_set_compat' (Im (list_to_set l) f) (map f l) (finite_image (bt A) (bt B) (list_to_set l) f h2) h6) as h5. rewrite h5 at 1. apply homo_times_list; auto. Qed. Lemma homo_plus_set : forall {A B:Bool_Alg} (f:bt A->bt B), homomorphism f -> forall (E:Ensemble (bt A)) (pf:Finite E), f (plus_set E pf) = plus_set (Im E f) (finite_image _ _ _ f pf). intros A B f h1 E h2. pose proof (finite_set_list_no_dup _ h2) as h3. destruct h3 as [l [h3 h4]]. subst. pose proof (plus_set_compat' (list_to_set l) _ h2 (eq_refl _)) as h3. rewrite h3 at 1. pose proof (map_im_compat f l) as h6. symmetry in h6. pose proof (plus_set_compat' (Im (list_to_set l) f) (map f l) (finite_image (bt A) (bt B) (list_to_set l) f h2) h6) as h5. rewrite h5 at 1. apply homo_plus_list; auto. Qed. Lemma homo_compose : forall {A B C:Bool_Alg} (f:bt A->bt B) (g:bt B->bt C), homomorphism f -> homomorphism g -> homomorphism (compose g f). intros A B C f g h1 h2. apply homo_two_ops_times. intros x y. unfold compose. rewrite (homo_times f); auto. rewrite (homo_times g); auto. intro x. unfold compose. rewrite (homo_comp f); auto. rewrite (homo_comp g); auto. Qed. Lemma homo_inv : forall {A B:Bool_Alg} (f:bt A->bt B) (pf:invertible f), homomorphism f -> homomorphism (proj1_sig (function_inverse _ pf)). intros A B f h1 h0. pose proof h1 as h2. apply invertible_impl_bijective in h2. destruct h2 as [h2l h2r]. apply homo_two_ops_plus. intros p q. pose proof (h2r p) as h3. pose proof (h2r q) as h4. destruct h3 as [p' h3]. destruct h4 as [q' h4]. subst. rewrite <- homo_plus. do 3 rewrite (match (proj2_sig (function_inverse f h1)) with | conj pf _ => pf end). reflexivity. assumption. intro x. specialize (h2r x). destruct h2r as [x' hr]. subst. rewrite <- homo_comp. do 2 rewrite (match (proj2_sig (function_inverse f h1)) with | conj pf _ => pf end). reflexivity. assumption. Qed. Lemma homo_mono : forall {A B:Bool_Alg} (f:bt A->bt B), homomorphism f -> forall x y, le x y -> le (f x) (f y). intros A B f h1. destruct h1 as [h1a h1b h1c]. intros x y h2. red. rewrite <- h1b. red in h2. rewrite h2. reflexivity. Qed. Lemma homo_im_closed : forall {A B:Bool_Alg} (f:bt A->bt B), homomorphism f -> alg_closed (Im (BS (Bc A)) f). intros A B f h1. pose proof h1 as h1'. destruct h1 as [h1a h1b h1c]. constructor. red. intros x y. unfold Bplus_sub. destruct x as [x h2]. destruct y as [y h3]. simpl. destruct h2 as [x h2]. destruct h3 as [y h3]. subst. rewrite <- h1b. apply Im_intro with (x+y); auto. rewrite und_set. constructor. intros x y. unfold Btimes_sub. destruct x as [x h2]. destruct y as [y h3]. simpl. destruct h2 as [x h2]. destruct h3 as [y h3]. subst. rewrite <- h1a. apply Im_intro with (x*y); auto. rewrite und_set. constructor. red. apply Im_intro with 1. rewrite und_set. constructor. rewrite homo_one; auto. red. apply Im_intro with 0. rewrite und_set. constructor. rewrite homo_zero; auto. intro x. unfold Bcomp_sub. destruct x as [x h2]. simpl. destruct h2 as [x h2]. subst. rewrite <- h1c. apply Im_intro with (-x); auto. rewrite und_set. constructor. Qed. (*This returns the same homomorphism, only with the range replaced with its image.*) Definition homo_onto {A B:Bool_Alg} (f:bt A->bt B) (pf:homomorphism f) : bt A->bt (Subalg _ (homo_im_closed _ pf)). intros x. unfold bt, Subalg, SubBtype. simpl. rewrite und_set. refine (exist _ _ (proj2_sig (sig_im_fun f x))). Defined. Lemma homo_onto_compat : forall {A B:Bool_Alg} (f:bt A->bt B) (pf:homomorphism f), surjective (homo_onto _ pf) /\ forall (x:bt A), f x = proj1_sig (homo_onto _ pf x). intros A B f h1. assert (h2: (forall x : bt A, f x = proj1_sig (homo_onto f h1 x))). unfold homo_onto. unfold eq_rect_r. unfold eq_rect. simpl. destruct (eq_sym (und_set A)). simpl. reflexivity. split; auto. pose proof (surj_sig_im_fun f) as h0. red in h0. red. intro y. simpl in y. destruct y as [y h3]. pose proof h3 as h3'. rewrite und_set in h3'. specialize (h0 (exist _ _ h3')). destruct h0 as [x h0]. exists x. apply proj1_sig_injective. simpl. rewrite <- h2. destruct h3. subst. pose proof (f_equal (@proj1_sig _ _) h0) as h4. simpl in h4. assumption. Qed. Lemma homo_homo_onto : forall {A B:Bool_Alg} (f:bt A->bt B) (pf:homomorphism f), homomorphism (homo_onto _ pf). intros A B f h1. pose proof (homo_onto_compat _ h1) as h2. destruct h2 as [h2l h2r]. apply homo_two_ops_times. intros x y. apply proj1_sig_injective. simpl. rewrite <- h2r. rewrite homo_times; auto. unfold Btimes_sub. do 2 rewrite <- h2r. reflexivity. intro x. apply proj1_sig_injective. simpl. rewrite <- h2r. rewrite homo_comp; auto. unfold Bcomp_sub. f_equal. apply h2r. Qed. Lemma inj_homo_onto_iff : forall {A B:Bool_Alg} (f:bt A->bt B) (pf:homomorphism f), Injective f <-> Injective (homo_onto _ pf). intros A B f h1. pose proof (homo_onto_compat _ h1) as h2. destruct h2 as [h2l h2r]. split. intro h2. red in h2. red. intros a b h3. apply h2. do 2 rewrite h2r. f_equal. assumption. intro h3. red in h3. red. intros a b h4. do 2 rewrite h2r in h4. apply proj1_sig_injective in h4. apply h3; auto. Qed. Lemma homo_transfer_iff : forall (A B C:Bool_Alg), A = B -> forall (f:bt A -> bt C) (pf:bt A = bt B), homomorphism f <-> homomorphism (transfer_fun pf f). intros A B C h1 f h2. subst. pose proof (proof_irrelevance _ h2 (eq_refl _)); subst; subst. rewrite transfer_fun_eq_refl. tauto. Qed. Definition rel_homo {A:Bool_Alg} (p:(bt A)) := (fun x=>x*p). Definition relativization_set {A:Bool_Alg} (p:bt A) := [x:(bt A) | le x p]. Definition relativization_set_zero : forall {A:Bool_Alg} (p:bt A), Inn (relativization_set p) 0. intros A p. constructor. apply zero_min. Qed. Definition relativization_set_one : forall {A:Bool_Alg} (p:bt A), Inn (relativization_set p) p. intros A p. constructor. apply refl_le. Qed. Lemma rel_homo_im_in_rel_set : forall {A:Bool_Alg} (p:(bt A)), Included (Im (BS (Bc A)) (rel_homo p)) (relativization_set p). intros A p. red. intros x h1. destruct h1 as [x h1]. subst. constructor. unfold rel_homo. rewrite comm_prod. apply times_le. Qed. Lemma rel_homo_sig_ex : forall {A:Bool_Alg} (p:bt A), exists! f:(bt A)->sig_set (relativization_set p), forall x, proj1_sig (f x) = rel_homo p x. intros A p. assert (h1:forall x:(bt A), Inn (Im (BS (Bc A)) (rel_homo p)) (x*p)). intros x. apply Im_intro with x. apply in_bs. reflexivity. exists (fun x => exist _ _ (rel_homo_im_in_rel_set _ _ ((h1 x)))). red. simpl. split. auto. intros f h2. apply functional_extensionality. intro x. apply proj1_sig_injective. simpl. rewrite h2. reflexivity. Qed. Definition rel_homo_sig {A:Bool_Alg} (p:bt A) : (bt A) -> sig_set (relativization_set p) := proj1_sig (constructive_definite_description _ (rel_homo_sig_ex p)). Lemma rel_homo_sig_compat : forall {A:Bool_Alg} (p:bt A), let f:=rel_homo_sig p in forall x, proj1_sig (f x) = rel_homo p x. intros A p f. unfold f. unfold rel_homo_sig. destruct constructive_definite_description as [f' h2]. simpl. assumption. Qed. Definition rel_times_closed : forall {B:Bool_Alg} (p:bt B) (x y:bt B), Inn (relativization_set p) x -> Inn (relativization_set p) y -> Inn (relativization_set p) (x * y). intros B p x y h1 h2. destruct h1 as [h1]. destruct h2 as [h2]. constructor. rewrite <- (idem_prod p). apply mono_prod; auto. Qed. Definition rel_plus_closed : forall {B:Bool_Alg} (p:bt B) (x y:bt B), Inn (relativization_set p) x -> Inn (relativization_set p) y -> Inn (relativization_set p) (x + y). intros B p x y h1 h2. destruct h1 as [h1]. destruct h2 as [h2]. constructor. rewrite <- (idem_sum p). apply mono_sum; auto. Qed. Definition rel_comp_closed : forall {B:Bool_Alg} (p:bt B) (x:bt B), Inn (relativization_set p) x -> Inn (relativization_set p) (-x * p). intros B p x h1. constructor. rewrite comm_prod. apply times_le. Qed. Lemma rel_times_ex : forall {B:Bool_Alg} (p:bt B), exists! (rt:sig_set (relativization_set p) -> sig_set (relativization_set p) -> sig_set (relativization_set p)), forall x y, proj1_sig (rt x y) = (proj1_sig x) * (proj1_sig y). intros B p. exists (fun (x y:sig_set (relativization_set p)) => exist _ _ (rel_times_closed p (proj1_sig x) (proj1_sig y) (proj2_sig x) (proj2_sig y))). red. simpl. split. auto. intros f h1. apply functional_extensionality. intro x. apply functional_extensionality. intro y. apply proj1_sig_injective. simpl. rewrite h1. reflexivity. Qed. Definition rel_times {B:Bool_Alg} (p:bt B) := proj1_sig (constructive_definite_description _ (rel_times_ex p)). Lemma rel_times_compat : forall {B:Bool_Alg} (p:bt B), let rt := rel_times p in forall x y, proj1_sig (rt x y) = (proj1_sig x) * (proj1_sig y). intros B p rt x y. unfold rt. unfold rel_times. destruct constructive_definite_description as [h1 h2]. simpl. apply h2. Qed. Lemma rel_times_compat' : forall {B:Bool_Alg} (p:bt B), rel_times p = fun x y => exist _ _ (rel_times_closed p (proj1_sig x) (proj1_sig y) (proj2_sig x) (proj2_sig y)). intros B p. apply functional_extensionality. intro x. apply functional_extensionality. intro y. apply proj1_sig_injective. simpl. apply rel_times_compat. Qed. Lemma rel_plus_ex : forall {B:Bool_Alg} (p:bt B), exists! (rp:sig_set (relativization_set p) -> sig_set (relativization_set p) -> sig_set (relativization_set p)), forall x y, proj1_sig (rp x y) = (proj1_sig x) + (proj1_sig y). intros B p. exists (fun (x y:sig_set (relativization_set p)) => exist _ _ (rel_plus_closed p (proj1_sig x) (proj1_sig y) (proj2_sig x) (proj2_sig y))). red. simpl. split. auto. intros f h1. apply functional_extensionality. intro x. apply functional_extensionality. intro y. apply proj1_sig_injective. simpl. rewrite h1. reflexivity. Qed. Definition rel_plus {B:Bool_Alg} (p:bt B) := proj1_sig (constructive_definite_description _ (rel_plus_ex p)). Lemma rel_plus_compat : forall {B:Bool_Alg} (p:bt B), let rp := rel_plus p in forall x y, proj1_sig (rp x y) = (proj1_sig x) + (proj1_sig y). intros B p rp x y. unfold rp. unfold rel_plus. destruct constructive_definite_description as [h1 h2]. simpl. apply h2. Qed. Lemma rel_plus_compat' : forall {B:Bool_Alg} (p:bt B), rel_plus p = fun x y => exist _ _ (rel_plus_closed p (proj1_sig x) (proj1_sig y) (proj2_sig x) (proj2_sig y)). intros B p. apply functional_extensionality. intro x. apply functional_extensionality. intro y. apply proj1_sig_injective. simpl. apply rel_plus_compat. Qed. Lemma rel_comp_ex : forall {B:Bool_Alg} (p:bt B), exists! (rc:sig_set (relativization_set p) -> sig_set (relativization_set p)), forall x, proj1_sig (rc x) = - (proj1_sig x) * p. intros B p. exists (fun x:sig_set (relativization_set p) => exist _ _ (rel_comp_closed p (proj1_sig x) (proj2_sig x))). red. simpl. split. auto. intros f h1. apply functional_extensionality. intro x. apply proj1_sig_injective. rewrite h1. simpl. reflexivity. Qed. Definition rel_comp {B:Bool_Alg} (p:bt B) := proj1_sig (constructive_definite_description _ (rel_comp_ex p)). Lemma rel_comp_compat : forall {B:Bool_Alg} (p:bt B), let rc := rel_comp p in forall x, proj1_sig (rc x) = - (proj1_sig x) * p. intros B p rp x. unfold rp. unfold rel_comp. destruct constructive_definite_description as [h1 h2]. simpl. apply h2. Qed. Lemma rel_comp_compat' : forall {B:Bool_Alg} (p:bt B), rel_comp p = fun x => exist _ _ (rel_comp_closed p (proj1_sig x) (proj2_sig x)). intros B p. apply functional_extensionality. intro x. apply proj1_sig_injective. simpl. apply rel_comp_compat. Qed. Definition rel_bc {B:Bool_Alg} (p:bt B) := (Build_Bconst (sig_set (relativization_set p)) (full_sig (relativization_set p)) (rel_plus p) (rel_times p) (exist _ _ (relativization_set_one p)) (exist _ _ (relativization_set_zero p)) (rel_comp p)). Lemma homo_gen_rel_homo_sig : forall {A:Bool_Alg} (p:bt A), homomorphism_general (rel_homo_sig p) (rel_times p) (rel_plus p) (rel_comp p). intros A p. constructor. intros x y. pose proof (rel_homo_sig_compat p (x*y)) as h2. apply proj1_sig_injective. rewrite h2 at 1. pose proof (rel_times_compat p (rel_homo_sig p x) (rel_homo_sig p y)) as h1. rewrite h1 at 1. pose proof (rel_homo_sig_compat p x) as h3. pose proof (rel_homo_sig_compat p y) as h4. rewrite h3. rewrite h4. unfold rel_homo. rewrite (comm_prod _ y p). rewrite assoc_prod. rewrite <- (assoc_prod _ x p p). rewrite idem_prod. rewrite <- assoc_prod. rewrite (comm_prod _ y p). rewrite assoc_prod. reflexivity. intros x y. pose proof (rel_homo_sig_compat p (x+y)) as h2. apply proj1_sig_injective. rewrite h2 at 1. pose proof (rel_plus_compat p (rel_homo_sig p x) (rel_homo_sig p y)) as h1. rewrite h1 at 1. pose proof (rel_homo_sig_compat p x) as h3. pose proof (rel_homo_sig_compat p y) as h4. rewrite h3. rewrite h4. unfold rel_homo. apply dist_sum_r. intros x. pose proof (rel_homo_sig_compat p (-x)) as h1. apply proj1_sig_injective. rewrite h1 at 1. pose proof (rel_comp_compat p (rel_homo_sig p x)) as h2. rewrite h2 at 1. pose proof (rel_homo_sig_compat p x) as h3. rewrite h3. unfold rel_homo. rewrite de_mor_prod. rewrite dist_sum_r. rewrite (comm_prod _ (-p) p). rewrite comp_prod. rewrite zero_sum. reflexivity. Qed. Lemma surj_rel_homo_sig : forall {A:Bool_Alg} (p:bt A), surjective (rel_homo_sig p). intros A p. red. intros b. destruct b as [b h1]. destruct h1 as [h1]. red in h1. pose proof h1 as h1'. pose proof (eq_ord b p) as h2. rewrite h2 in h1'. exists b. apply proj1_sig_injective. simpl. pose proof (rel_homo_sig_compat p b) as h3. rewrite h3 at 1. rewrite <- h1' at 2. reflexivity. Qed. Definition rel_ba {B:Bool_Alg} (p:bt B) := proj1_sig (constructive_definite_description _ (universal_homomorphic_ops_imply_bool _ _ _ _ (homo_gen_rel_homo_sig p) (surj_rel_homo_sig p))). Lemma rel_ba_compat : forall {B:Bool_Alg} (p:bt B), let B' := rel_ba p in Bc B' = rel_bc p. intros B p B'. unfold B'. unfold rel_ba. destruct constructive_definite_description as [A h1]. simpl. rewrite h1. unfold rel_bc. assert (h2:Full_set (sig_set (relativization_set p)) = full_sig (relativization_set p)). reflexivity. rewrite h2. f_equal; try apply proj1_sig_injective; simpl. pose proof (rel_homo_sig_compat p 1) as h3. unfold rel_homo in h3. rewrite comm_prod in h3. rewrite one_prod in h3. assumption. pose proof (rel_homo_sig_compat p 0) as h3. unfold rel_homo in h3. rewrite comm_prod in h3. rewrite zero_prod in h3. assumption. Qed. Lemma btype_bc_rel_ba : forall {B:Bool_Alg} (p:bt B), bt (rel_ba p) = sig_set (relativization_set p). intros B p. unfold bt. simpl. rewrite rel_ba_compat. simpl. reflexivity. Qed. Lemma rel_ba_times_compat : forall {B:Bool_Alg} (p:bt B), (Btimes (Bc (rel_ba p))) = (transfer_dep_r (btype_bc_rel_ba p) (U:=fun T=>T->T->T) (Btimes (rel_bc p))). intros B p. simpl. rewrite rel_times_compat'. rewrite <- (transfer_dep_transfer_dep_r_compat (eq_sym (btype_bc_rel_ba p)) (btype_bc_rel_ba p)). symmetry. assert (h1:@transfer_dep Type (fun x:Type => x-> x-> x) _ _ (eq_sym (btype_bc_rel_ba p)) (fun x0 y0 : sig_set (relativization_set p) => exist (Inn (relativization_set p)) (proj1_sig x0 * proj1_sig y0) (rel_times_closed p (proj1_sig x0) (proj1_sig y0) (proj2_sig x0) (proj2_sig y0))) = (Btimes (Bc (rel_ba p)))). rewrite <- (@transfer_dep_eq_iff Type (fun x:Type => x->x->x) _ _ (eq_sym (btype_bc_rel_ba p)) (fun x0 y0 : sig_set (relativization_set p) => exist (Inn (relativization_set p)) (proj1_sig x0 * proj1_sig y0) (rel_times_closed p (proj1_sig x0) (proj1_sig y0) (proj2_sig x0) (proj2_sig y0))) (Btimes (Bc (rel_ba p)))). assert (h1:(fun x0 y0 : sig_set (relativization_set p) => exist (Inn (relativization_set p)) (proj1_sig x0 * proj1_sig y0) (rel_times_closed p (proj1_sig x0) (proj1_sig y0) (proj2_sig x0) (proj2_sig y0))) = (rel_times p)). apply functional_extensionality. intro x'. apply functional_extensionality. intro y'. apply proj1_sig_injective. simpl. rewrite <- (rel_times_compat p x' y'). reflexivity. rewrite h1. unfold bt. rewrite (rel_ba_compat p). simpl. reflexivity. assumption. Qed. Lemma rel_ba_plus_compat : forall {B:Bool_Alg} (p:bt B), (Bplus (Bc (rel_ba p))) = (transfer_dep_r (btype_bc_rel_ba p) (U:=fun T=>T->T->T) (Bplus (rel_bc p))). intros B p. simpl. rewrite rel_plus_compat'. rewrite <- (transfer_dep_transfer_dep_r_compat (eq_sym (btype_bc_rel_ba p)) (btype_bc_rel_ba p)). symmetry. assert (h1:@transfer_dep Type (fun x:Type => x-> x-> x) _ _ (eq_sym (btype_bc_rel_ba p)) (fun x0 y0 : sig_set (relativization_set p) => exist (Inn (relativization_set p)) (proj1_sig x0 + proj1_sig y0) (rel_plus_closed p (proj1_sig x0) (proj1_sig y0) (proj2_sig x0) (proj2_sig y0))) = (Bplus (Bc (rel_ba p)))). rewrite <- (@transfer_dep_eq_iff Type (fun x:Type => x->x->x) _ _ (eq_sym (btype_bc_rel_ba p)) (fun x0 y0 : sig_set (relativization_set p) => exist (Inn (relativization_set p)) (proj1_sig x0 + proj1_sig y0) (rel_plus_closed p (proj1_sig x0) (proj1_sig y0) (proj2_sig x0) (proj2_sig y0))) (Bplus (Bc (rel_ba p)))). assert (h1:(fun x0 y0 : sig_set (relativization_set p) => exist (Inn (relativization_set p)) (proj1_sig x0 + proj1_sig y0) (rel_plus_closed p (proj1_sig x0) (proj1_sig y0) (proj2_sig x0) (proj2_sig y0))) = (rel_plus p)). apply functional_extensionality. intro x'. apply functional_extensionality. intro y'. apply proj1_sig_injective. simpl. rewrite <- (rel_plus_compat p x' y'). reflexivity. rewrite h1. unfold bt. rewrite (rel_ba_compat p). simpl. reflexivity. assumption. Qed. Lemma rel_ba_comp_compat : forall {B:Bool_Alg} (p:bt B), (Bcomp (Bc (rel_ba p))) = (transfer_dep_r (btype_bc_rel_ba p) (U:=fun T=>T->T) (Bcomp (rel_bc p))). intros B p. simpl. rewrite rel_comp_compat'. rewrite <- (transfer_dep_transfer_dep_r_compat (eq_sym (btype_bc_rel_ba p)) (btype_bc_rel_ba p)). symmetry. assert (h1:@transfer_dep Type (fun x:Type => x-> x) _ _ (eq_sym (btype_bc_rel_ba p)) (fun x0: sig_set (relativization_set p) => exist ( Inn (relativization_set p)) (- proj1_sig x0 * p) (rel_comp_closed p (proj1_sig x0) (proj2_sig x0) )) = (Bcomp (Bc (rel_ba p)))). rewrite <- (@transfer_dep_eq_iff Type (fun x:Type => x->x) _ _ (eq_sym (btype_bc_rel_ba p)) (fun x0 : sig_set (relativization_set p) => exist (Inn (relativization_set p)) (-proj1_sig x0 * p) (rel_comp_closed p (proj1_sig x0) (proj2_sig x0))) (Bcomp (Bc (rel_ba p)))). assert (h1:(fun x0 : sig_set (relativization_set p) => exist (Inn (relativization_set p)) (- proj1_sig x0 * p) (rel_comp_closed p (proj1_sig x0) (proj2_sig x0))) = (rel_comp p)). apply functional_extensionality. intro x'. apply proj1_sig_injective. simpl. rewrite <- (rel_comp_compat p x'). reflexivity. rewrite h1. unfold bt. rewrite (rel_ba_compat p). simpl. reflexivity. assumption. Qed. Definition point_induced_homo (F:Field_of_Sets) (x0:Xt F) : bt (fos_ba F) -> (bt two_bool_ba). unfold bt. rewrite two_bool_ba_compat. simpl. intro P. destruct P as [P]. refine (char_fun P x0). Defined. Lemma homo_point_induced_homo : forall (F:Field_of_Sets) (x0:Xt F), homomorphism (point_induced_homo F x0). intros F x0. apply homo_two_ops_times. unfold fos_ba. simpl. intros P Q. simpl in P. simpl in Q. simpl. destruct P as [P h1]. destruct Q as [Q h2]. unfold Int_fos. simpl. unfold point_induced_homo. simpl. unfold eq_rect_r. unfold eq_rect. unfold bt. destruct (eq_sym two_bool_ba_compat). simpl. apply char_fun_int. intro P. simpl in P. destruct P as [P h1]. unfold point_induced_homo. simpl. unfold eq_rect_r, eq_rect. unfold bt. destruct (eq_sym two_bool_ba_compat). unfold Comp_fos. simpl. apply char_fun_comp. Qed. Definition function_induced_homo (A B:Field_of_Sets) (f:(Xt A) -> (Xt B)) (pf:forall P:(bt (fos_ba B)), Inn (F A) (inv_im (proj1_sig P) f)) : (bt (fos_ba B)) -> (bt (fos_ba A)). intro P. specialize (pf P). refine (exist _ _ pf). Defined. Lemma homo_function_induced_homo : forall (A B:Field_of_Sets) (f:(Xt A) -> (Xt B)) (pf:forall P:(bt (fos_ba B)), Inn (F A) (inv_im (proj1_sig P) f)), homomorphism (function_induced_homo A B f pf). intros A B f h1. unfold function_induced_homo. apply homo_two_ops_times. intros P Q. simpl. unfold Int_fos. simpl. destruct P as [P h2]. destruct Q as [Q h3]. simpl. apply proj1_sig_injective. simpl. apply inv_im_int. intro P. simpl. unfold Comp_fos. simpl. destruct P as [P h2]. simpl. apply proj1_sig_injective. simpl. apply inv_im_comp. Qed. Inductive monomorphism {A B:Bool_Alg} (f:(bt A)->(bt B)) : Prop := | monomorphism_intro : homomorphism f -> Injective f -> monomorphism f. Lemma mono_homo : forall {A B:Bool_Alg} (f:(bt A)->(bt B)), monomorphism f -> homomorphism f. intros A B f h1. destruct h1; auto. Qed. Inductive epimorphism {A B:Bool_Alg} (f:(bt A)->(bt B)) : Prop := | epimorphism_intro : homomorphism f -> surjective f -> epimorphism f. Lemma epi_homo : forall {A B:Bool_Alg} (f:(bt A)->(bt B)), epimorphism f -> homomorphism f. intros A B f h1. destruct h1; auto. Qed. Lemma mono_zero_rev : forall {A B:Bool_Alg} (f:bt A->bt B), monomorphism f -> forall (x:bt A), f x = 0 -> x = 0. intros A B f h1 x h2. destruct h1 as [h1 h3]. rewrite <- (homo_zero f) in h2; auto. Qed. Lemma mono_one_rev : forall {A B:Bool_Alg} (f:bt A->bt B), monomorphism f -> forall (x:bt A), f x = 1 -> x = 1. intros A B f h1 x h2. destruct h1 as [h1 h3]. rewrite <- (homo_one f) in h2; auto. Qed. Inductive isomorphism {A B:Bool_Alg} (f:(bt A)->(bt B)) : Prop := | isomorphism_intro : homomorphism f -> bijective f -> isomorphism f. Lemma iso_homo : forall {A B:Bool_Alg} (f:(bt A)->(bt B)), isomorphism f -> homomorphism f. intros A B f h1. destruct h1; auto. Qed. Lemma iso_epi_mono_compat_iff : forall {A B:Bool_Alg} (f:(bt A)->(bt B)), isomorphism f <-> monomorphism f /\ epimorphism f. intros A B f. split. intro h1. destruct h1 as [h1 h2]. red in h2. destruct h2 as [h2l h2r]. split. constructor; auto. constructor; auto. intro h1. destruct h1 as [h1 h2]. destruct h1 as [h1 h3]. destruct h2 as [h1' h4]. constructor; auto. red. split; auto. Qed. Definition iso_inv {A B:Bool_Alg} (f:bt A->bt B) (pf:isomorphism f) : bt B -> bt A. destruct pf as [h1 h2]. apply bijective_impl_invertible in h2. refine (proj1_sig (function_inverse _ h2)). Defined. Lemma iso_inv_compat : forall {A B:Bool_Alg} (f:bt A->bt B) (pf:isomorphism f), isomorphism (iso_inv f pf) /\ (forall x, (iso_inv f pf) (f x) = x) /\ (forall x, f (iso_inv f pf x) = x). intros A B f h1. destruct h1 as [h1 h2]. unfold iso_inv. split. split. unfold iso_inv. apply homo_inv; auto. apply invertible_impl_bijective. apply invertible_impl_inv_invertible. unfold function_inverse. destruct constructive_definite_description. simpl. assumption. Qed. Lemma mono_transfer_iff : forall (A B C:Bool_Alg), A = B -> forall (f:bt A -> bt C) (pf:bt A = bt B), monomorphism f <-> monomorphism (transfer_fun pf f). intros A B C h1 f h2. subst. pose proof (proof_irrelevance _ h2 (eq_refl _)); subst. rewrite transfer_fun_eq_refl. tauto. Qed. Definition automorphism {A:Bool_Alg} (f:bt A->bt A) : Prop := isomorphism f. Lemma automorphism_id : forall A:Bool_Alg, @automorphism A id. intro A. red. constructor; [apply homo_id | apply bij_id]. Qed. Definition isomorphic (A B:Bool_Alg) : Prop := exists phi:bt A->bt B, isomorphism phi. Notation "A =~ B" := (isomorphic A B) (at level 30). Lemma refl_iso : forall A:Bool_Alg, A =~ A. red. intro A. exists id. apply automorphism_id. Qed. Lemma symm_iso : forall A B:Bool_Alg, A =~ B -> B =~ A. intros A B h1. destruct h1 as [f h1]. destruct h1 as [h1 h2]. apply bijective_impl_invertible in h2. exists (proj1_sig (function_inverse _ h2)). constructor. apply homo_inv. assumption. apply invertible_impl_bijective. apply invertible_impl_inv_invertible; auto. Qed. Lemma trans_iso : forall A B C:Bool_Alg, A =~ B -> B =~ C -> A =~ C. red. intros A B C h1 h2. destruct h1 as [f h1]. destruct h2 as [g h2]. destruct h1 as [h1l h1r]. destruct h2 as [h2l h2r]. exists (compose g f). constructor. apply homo_compose; auto. apply bij_compose; auto. Qed. Definition sub_isomorphic (A B:Bool_Alg) := exists (C:Ensemble (bt B)) (pf:alg_closed C), A =~ Subalg _ pf. Notation "A <=~ B" := (sub_isomorphic A B) (at level 30). Lemma rel_homo_psa : forall {Xt:Type} (Y P:bt (psa Xt)), rel_homo Y P = Intersection P Y. intros Xt Y P. unfold rel_homo; simpl; auto. Qed. Lemma rel_set_psa : forall {Xt:Type} (Y:bt (psa Xt)), relativization_set Y = power_set Y. intros Xt Y. unfold relativization_set. simpl. unfold le. simpl. assert (h1:forall P:bt (psa Xt), Union P Y = Y <-> Included P Y). intros; rewrite inclusion_iff_union. tauto. simpl in Y. simpl in h1. rewrite (sat_iff _ _ h1). unfold power_set. reflexivity. Qed. Lemma btype_bc_rel_ba_psa : forall {Xt:Type} (Y:bt (psa Xt)), bt (rel_ba Y) = sig_set (power_set Y). intros Xt Y. unfold bt. pose proof (rel_ba_compat Y) as h1. rewrite h1. simpl. f_equal. apply rel_set_psa. Qed. Lemma sig_set_psa : forall {Xt:Type} (Y:bt (psa Xt)), sig_set (relativization_set Y) = sig_set (power_set Y). intros Xt Y. f_equal. apply rel_set_psa. Qed. Definition proj1_sig_rel_ba_psa {Xt:Type} (Y:bt (psa Xt)) (N:bt (rel_ba Y)) : Ensemble Xt := (proj1_sig (transfer (btype_bc_rel_ba_psa Y) N)). Lemma proj1_sig_rel_ba_psa_compat : forall {Xt:Type} (Y:bt (psa Xt)) (N:bt (rel_ba Y)), Included (proj1_sig_rel_ba_psa Y N) Y. intros Xt Y N. apply (proj2_sig (transfer (btype_bc_rel_ba_psa Y) N)). Qed. Lemma proj1_sig_rel_ba_psa_compat' : forall {Xt:Type} (Y:bt (psa Xt)) (N:bt (rel_ba Y)), Inn (power_set Y) (proj1_sig_rel_ba_psa Y N). intros Xt Y N. constructor. apply proj1_sig_rel_ba_psa_compat. Qed. Lemma proj1_sig_rel_ba_psa_inj : forall {Xt:Type} (Y:bt (psa Xt)) (A B:bt (rel_ba Y)), proj1_sig_rel_ba_psa Y A = proj1_sig_rel_ba_psa Y B -> A = B. intros Xt Y A B h1. unfold proj1_sig_rel_ba_psa in h1. apply proj1_sig_injective in h1. apply transfer_inj in h1. assumption. Qed. Lemma rel_ba_psa_times_compat : forall {Xt:Type} (Y:bt (psa Xt)), Btimes (Bc (rel_ba Y)) = transfer_dep_r (btype_bc_rel_ba_psa Y) (U:=fun T=>T->T->T) (fun A B : sig_set (power_set Y) => exist _ _ (in_power_intersection (proj1_sig A) (proj1_sig B) Y (proj2_sig A))). intros Xt Y. pose proof (rel_ba_times_compat Y) as h1. simpl in h1. rewrite rel_times_compat' in h1. rewrite h1. simpl. rewrite (@transfer_dep_r_hetero _ (fun y=>y->y->y) _ _ _ (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y)). apply functional_extensionality. intro A. apply functional_extensionality. intro B. apply proj1_sig_injective. rewrite transfer_dep_r_fun2_eq. pose proof (rel_set_psa Y) as h2. do 2 rewrite (transfer_sig_set_eq _ _ h2 (eq2 (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y))). simpl. destruct A as [A h3]. destruct B as [B h4]. simpl. pose proof (eq2 (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y)) as h5. rewrite (transfer_r_sig_set_eq _ _ h2 (eq2 (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y))). simpl. reflexivity. Qed. Lemma rel_ba_psa_times_compat' : forall {Xt : Type} (Y:bt (psa Xt)) (A B:bt (rel_ba Y)), A * B = transfer_r (btype_bc_rel_ba_psa Y) (exist _ _ (in_power_intersection (proj1_sig_rel_ba_psa Y A) (proj1_sig_rel_ba_psa Y B) Y (proj1_sig_rel_ba_psa_compat' _ A))). intros Xt Y A B. rewrite rel_ba_psa_times_compat. rewrite transfer_dep_r_fun2_eq. f_equal. apply proj1_sig_injective. simpl. unfold proj1_sig_rel_ba_psa. reflexivity. Qed. Lemma rel_ba_psa_plus_compat : forall {Xt:Type} (Y:bt (psa Xt)), Bplus (Bc (rel_ba Y)) = transfer_dep_r (btype_bc_rel_ba_psa Y) (U:=fun T=>T->T->T) (fun A B : sig_set (power_set Y) => exist _ _ (in_power_union (proj1_sig A) (proj1_sig B) Y (proj2_sig A) (proj2_sig B))). intros Xt Y. pose proof (rel_ba_plus_compat Y) as h1. simpl in h1. rewrite rel_plus_compat' in h1. rewrite h1. simpl. rewrite (@transfer_dep_r_hetero _ (fun y=>y->y->y) _ _ _ (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y)). apply functional_extensionality. intro A. apply functional_extensionality. intro B. apply proj1_sig_injective. rewrite transfer_dep_r_fun2_eq. pose proof (rel_set_psa Y) as h2. do 2 rewrite (transfer_sig_set_eq _ _ h2 (eq2 (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y))). simpl. destruct A as [A h3]. destruct B as [B h4]. simpl. pose proof (eq2 (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y)) as h5. rewrite (transfer_r_sig_set_eq _ _ h2 (eq2 (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y))). simpl. reflexivity. Qed. Lemma rel_ba_psa_plus_compat' : forall {Xt : Type} (Y:bt (psa Xt)) (A B:bt (rel_ba Y)), A + B = transfer_r (btype_bc_rel_ba_psa Y) (exist _ _ (in_power_union (proj1_sig_rel_ba_psa Y A) (proj1_sig_rel_ba_psa Y B) Y (proj1_sig_rel_ba_psa_compat' _ A) (proj1_sig_rel_ba_psa_compat' _ B) )). intros Xt Y A B. rewrite rel_ba_psa_plus_compat. rewrite transfer_dep_r_fun2_eq. f_equal. apply proj1_sig_injective. simpl. unfold proj1_sig_rel_ba_psa. reflexivity. Qed. Lemma rel_ba_psa_comp_compat : forall {Xt:Type} (Y:bt (psa Xt)), Bcomp (Bc (rel_ba Y)) = transfer_dep_r (btype_bc_rel_ba_psa Y) (U:=fun T=>T->T) (fun A: sig_set (power_set Y) => exist _ _ (in_power_comp (proj1_sig A) Y (proj2_sig A))). intros Xt Y. pose proof (rel_ba_comp_compat Y) as h1. simpl in h1. rewrite rel_comp_compat' in h1. rewrite h1. simpl. rewrite (@transfer_dep_r_hetero _ (fun y=>y->y) _ _ _ (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y)). apply functional_extensionality. intro A. apply proj1_sig_injective. rewrite transfer_dep_r_fun1_eq. pose proof (rel_set_psa Y) as h2. rewrite (transfer_sig_set_eq _ _ h2 (eq2 (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y))). simpl. destruct A as [A h3]. pose proof (eq2 (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y)) as h5. rewrite (transfer_r_sig_set_eq _ _ h2 (eq2 (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y))). simpl. rewrite setminus_int_complement. rewrite Intersection_commutative. reflexivity. Qed. Lemma rel_ba_psa_comp_compat' : forall {Xt : Type} (Y:bt (psa Xt)) (A:bt (rel_ba Y)), - A = transfer_r (btype_bc_rel_ba_psa Y) (exist _ _ (in_power_comp (proj1_sig_rel_ba_psa Y A) Y (proj1_sig_rel_ba_psa_compat' _ A) )). intros Xt Y A. rewrite rel_ba_psa_comp_compat. rewrite transfer_dep_r_fun1_eq. f_equal. apply proj1_sig_injective. simpl. unfold proj1_sig_rel_ba_psa. reflexivity. Qed. Lemma proj1_sig_rel_ba_psa_times : forall {Xt : Type} (Y:bt (psa Xt)) (A B:bt (rel_ba Y)), proj1_sig_rel_ba_psa _ (A * B) = Intersection (proj1_sig_rel_ba_psa _ A) (proj1_sig_rel_ba_psa _ B). intros Xt Y A B. pose proof (rel_ba_psa_times_compat' Y A B) as h0. pose proof (f_equal (transfer (btype_bc_rel_ba_psa Y)) h0) as h2. rewrite transfer_undoes_transfer_r in h2. pose proof (f_equal (@proj1_sig _ _) h2) as h3. simpl in h3. unfold proj1_sig_rel_ba_psa. assumption. Qed. Lemma proj1_sig_rel_ba_psa_plus : forall {Xt : Type} (Y:bt (psa Xt)) (A B:bt (rel_ba Y)), proj1_sig_rel_ba_psa _ (A + B) = Union (proj1_sig_rel_ba_psa _ A) (proj1_sig_rel_ba_psa _ B). intros Xt Y A B. pose proof (rel_ba_psa_plus_compat' Y A B) as h0. pose proof (f_equal (transfer (btype_bc_rel_ba_psa Y)) h0) as h2. rewrite transfer_undoes_transfer_r in h2. pose proof (f_equal (@proj1_sig _ _) h2) as h3. simpl in h3. unfold proj1_sig_rel_ba_psa. assumption. Qed. Lemma proj1_sig_rel_ba_psa_comp : forall {Xt : Type} (Y:bt (psa Xt)) (A:bt (rel_ba Y)), proj1_sig_rel_ba_psa _ (- A) = Setminus Y (proj1_sig_rel_ba_psa _ A). intros Xt Y A. pose proof (rel_ba_psa_comp_compat' Y A) as h0. pose proof (f_equal (transfer (btype_bc_rel_ba_psa Y)) h0) as h2. rewrite transfer_undoes_transfer_r in h2. pose proof (f_equal (@proj1_sig _ _) h2) as h3. simpl in h3. unfold proj1_sig_rel_ba_psa. assumption. Qed. Lemma rel_ba_psa_sig_set_iso : forall {Xt:Type} (Y:(bt (psa Xt))), rel_ba Y =~ psa (sig_set Y). intros Xt Y. red. pose proof (rel_ba_compat Y) as h1. simpl in h1. assert (h2:forall P, exists! V:Btype (Bc (psa (sig_set Y))), Im V (@proj1_sig _ _) = rel_homo P Y). intro P. assert (h2:Included (rel_homo P Y) Y). unfold rel_homo. simpl. auto with sets. exists (Im (full_sig (rel_homo P Y)) (fun v=>(exist _ _ (h2 _ (proj2_sig v))))). red. split. rewrite im_im. simpl. apply Extensionality_Ensembles. red. split. red. intros a h3. destruct h3 as [a h3]. subst. unfold rel_homo. simpl. apply proj2_sig. red. intros a h3. apply Im_intro with (exist _ _ h3). constructor. simpl. reflexivity. intros V h3. apply Extensionality_Ensembles. red. split. red. intros y h4. destruct y as [y h5]. simpl. inversion h4 as [v h4b]. subst. apply exist_injective in H. subst. pose proof (proj2_sig v) as h7. simpl in h7. rewrite <- h3 in h7 at 1. inversion h7 as [v' h8]. subst. simpl in V. destruct v. simpl. simpl in H. subst. destruct v'. simpl. simpl in h5. pose proof (proof_irrelevance _ i0 h5); subst. assumption. red. intros x h4. destruct x as [x h5]. assert (h6:Inn (rel_homo P Y) x). rewrite <- h3 at 1. apply Im_intro with (exist (fun x0:Xt => Inn Y x0) x h5). assumption. simpl. reflexivity. apply Im_intro with (exist _ _ h6). constructor. apply proj1_sig_injective. simpl. reflexivity. pose (fun P:Btype (Bc (rel_ba Y)) => (proj1_sig (constructive_definite_description _ (h2 (proj1_sig_rel_ba_psa Y P))))) as f. unfold psa. simpl. exists f. assert (h3:bijective f). red. split. red. unfold f. intros A B. destruct constructive_definite_description as [A' h4]. destruct constructive_definite_description as [B' h5]. simpl. intro h6. subst. rewrite h5 in h4. unfold rel_homo in h4. simpl in h4. pose proof (proj1_sig_rel_ba_psa_compat _ A) as h7. pose proof (proj1_sig_rel_ba_psa_compat _ B) as h8. rewrite inclusion_iff_intersection_eq in h7. rewrite inclusion_iff_intersection_eq in h8. rewrite Intersection_commutative in h7. rewrite Intersection_commutative in h8. rewrite h7 in h4. rewrite h8 in h4. apply proj1_sig_rel_ba_psa_inj in h4. subst. reflexivity. red. simpl. intro B. simpl in f. assert (h3:Inn (relativization_set Y) (Im B (@proj1_sig _ _))). constructor. unfold le. simpl. rewrite <- inclusion_iff_union. red. intros x h3. destruct h3 as [x h3]. subst. apply proj2_sig. exists (transfer_r (btype_bc_rel_ba Y) (exist _ _ h3)). unfold f. destruct constructive_definite_description as [V h4]. simpl. unfold rel_homo in h4. simpl in h4. pose proof (proj1_sig_rel_ba_psa_compat _ (transfer_r (btype_bc_rel_ba Y) (exist (Inn (relativization_set Y)) (Im B (proj1_sig (P:=fun x : Xt => Inn Y x))) h3))) as h5. rewrite inclusion_iff_intersection_eq in h5. rewrite Intersection_commutative in h5. rewrite h5 in h4 at 1. unfold proj1_sig_rel_ba_psa in h4. rewrite transfer_transfer_r_hetero in h4. pose proof (rel_set_psa Y) as h6. pose proof (transfer_sig_set_eq _ _ h6 (eq2 (btype_bc_rel_ba Y) (btype_bc_rel_ba_psa Y)) (exist (Inn (relativization_set Y)) (Im B (proj1_sig (P:=fun x : Xt => Inn Y x))) h3)) as h7. rewrite h7 in h4 at 1. simpl in h4. apply im_inj_inj in h4; auto. red. intros; apply proj1_sig_injective; auto. constructor. Focus 2. assumption. apply homo_two_ops_times. intros A B. unfold f. assert (h0:Injective (proj1_sig (P:=fun x : Xt => Inn Y x))). red; intros; apply proj1_sig_injective; auto. pose proof (proj1_sig_rel_ba_psa_times Y A B) as h7. destruct constructive_definite_description as [A' h4]. destruct constructive_definite_description as [B' h5]. destruct constructive_definite_description as [V h6]. simpl. unfold rel_homo in h4. simpl in h4. unfold rel_homo in h5. simpl in h5. unfold rel_homo in h6. simpl in h6. pose proof (proj1_sig_rel_ba_psa_compat Y (A*B)) as h8. pose proof (proj1_sig_rel_ba_psa_compat Y A) as h9. pose proof (proj1_sig_rel_ba_psa_compat Y B) as h10. rewrite inclusion_iff_intersection_eq in h8. rewrite inclusion_iff_intersection_eq in h9. rewrite inclusion_iff_intersection_eq in h10. rewrite Intersection_commutative in h8 at 1. rewrite Intersection_commutative in h9 at 1. rewrite Intersection_commutative in h10 at 1. rewrite h8 in h4. rewrite h9 in h5. rewrite h10 in h6. clear h8 h9 h10. rewrite h7 in h4 at 1. rewrite <- h6 in h4. rewrite <- h5 in h4. rewrite <- im_intersection_inj in h4; auto. apply im_inj_inj in h4; auto. intro A. pose proof (proj1_sig_rel_ba_psa_comp Y A) as h7. unfold f. destruct constructive_definite_description as [A' h4]. destruct constructive_definite_description as [V h6]. simpl. unfold rel_homo in h4. simpl in h4. unfold rel_homo in h6. simpl in h6. pose proof (proj1_sig_rel_ba_psa_compat Y (-A)) as h8. pose proof (proj1_sig_rel_ba_psa_compat Y A) as h9. rewrite inclusion_iff_intersection_eq in h8. rewrite inclusion_iff_intersection_eq in h9. rewrite Intersection_commutative in h8 at 1. rewrite Intersection_commutative in h9 at 1. rewrite h8 in h4. rewrite h9 in h6. clear h8 h9. rewrite h7 in h4 at 1. rewrite <- h6 in h4. simpl in V. simpl in Y. assert (h9: Setminus Y (Im V (proj1_sig (P:=fun x : Xt => Inn Y x))) = Im (Comp V) (@proj1_sig _ _)). apply Extensionality_Ensembles. red. split. intros x h10. destruct h10 as [h10 h11]. apply Im_intro with (exist _ _ h10). intro h12. contradict h11. apply Im_intro with (exist _ _ h10). assumption. simpl. reflexivity. simpl. reflexivity. red. intros x h10. destruct h10 as [x h10]. subst. constructor. apply proj2_sig. intro h11. inversion h11 as [x' h12 y h13]. subst. apply proj1_sig_injective in h13. subst. contradiction. rewrite h9 in h4. apply im_inj_inj in h4; auto. red. intros; apply proj1_sig_injective; auto. Qed. Lemma same_card_type_iso_psa : forall (Xt Yt:Type), (exists f:Xt->Yt, bijective f) -> psa Yt =~ psa Xt. intros Wt Yt h1. destruct h1 as [f h1]. pose proof (fos_psa_compat Wt) as h3. pose proof (fos_psa_compat Yt) as h4. destruct h3 as [h3l h3r]. destruct h4 as [h4l h4r]. simpl. pose (fun P:Ensemble Yt => inv_im P f) as f''. pose proof h3l as h3l'. symmetry in h3l'. pose proof h4l as h4l'. symmetry in h4l'. pose (fun x:(Xt (fos_psa Wt)) => (transfer h4l' (f (transfer h3l x)))) as f'. pose proof (transfer_fun_pred' h3l h4l' f (fun pr:{pT:Type*Type & fst pT -> snd pT} => bijective (projT2 pr))) as ht. simpl in ht. pose proof h1 as h1'. rewrite ht in h1'. assert (hbij:bijective f'). unfold f'. unfold transfer at 1 in h1'. unfold eq_rect_r in h1'. rewrite <- eq_rect_eq in h1'. assumption. assert (h5:(forall P : bt (fos_ba (fos_psa Yt)), Inn (F (fos_psa Wt)) (inv_im (proj1_sig P) f'))). rewrite h3r. intros. constructor. red; intros; constructor. pose (function_induced_homo (fos_psa Wt) (fos_psa Yt) f' h5) as phi. simpl in phi. unfold Ft in phi. pose (fun B:Ensemble Yt => proj1_sig_fos_psa (phi (exist_fos_psa B))) as g. exists g. constructor. pose proof (homo_function_induced_homo (fos_psa Wt) (fos_psa Yt) f' h5) as h6. destruct h6 as [h6a h6b h6c]. apply homo_two_ops_times. intros A B. simpl in A, B. simpl in h6a. unfold Ft in h6a. specialize (h6a (exist_fos_psa A) (exist_fos_psa B)). unfold Int_fos in h6a. simpl in h6a. unfold function_induced_homo in h6a. simpl in h6a. pose proof (f_equal (@proj1_sig _ _) h6a) as h6a'. simpl in h6a'. unfold g. simpl. destruct (fos_psa_compat Wt) as [h8l h8r]. rewrite <- h8l. do 3 rewrite <- eq_rect_eq. rewrite proj1_sig_exist_fos_psa_int. rewrite h6a'. reflexivity. intro A. simpl in A. simpl in h6c. specialize (h6c (exist_fos_psa A)). unfold Comp_fos in h6c. simpl in h6c. unfold function_induced_homo in h6c. simpl in h6c. pose proof (f_equal (@proj1_sig _ _) h6c) as h6c'. simpl in h6c'. unfold g. simpl. destruct (fos_psa_compat Wt) as [h8l h8r]. rewrite <- h8l. do 2 rewrite <- eq_rect_eq. rewrite proj1_sig_exist_fos_psa_comp. assumption. pose proof hbij as hbij'. red in hbij'. destruct hbij' as [hb1 hb2]. red in h1. destruct h1 as [h1l h1r]. red in h1l. red in h1r. red. split. red. intros A B h1. unfold g in h1. unfold phi in h1. unfold exist_fos_psa in h1. destruct (fos_psa_compat Yt) as [h2 h3]. unfold proj1_sig_fos_psa in h1. destruct (fos_psa_compat Wt) as [h6 h7]. simpl in h1. destruct h6. do 2 rewrite <- eq_rect_eq in h1. unfold eq_rect_r in h1. pose proof (inv_im_surj_inj f' hb2 _ _ h1) as hp. apply proj1_sig_injective in hp. rewrite h3 in hp. simpl in hp. pose proof (f_equal (@proj1_sig _ _) hp) as hp'. rewrite h2 in hp'. simpl in hp'. assumption. red. intro A. unfold g. red in hb2. simpl in A. unfold phi. simpl. exists (Im A f). destruct (fos_psa_compat Wt) as [h7 h8]. apply Extensionality_Ensembles. red. split. red. intros x h9. rewrite in_eq_rect in h9. destruct h9 as [h9]. pose proof (exist_fos_psa_compat (Im A f)) as h10. rewrite h10 in h9. unfold f' in h9. destruct (fos_psa_compat Yt) as [h11 h12]. pose proof h11 as h11'. symmetry in h11'. rewrite <- (transfer_dep_transfer_dep_r_compat h11' h11) in h9. pose proof (proof_irrelevance _ h4l' h11') as h13. rewrite h13 in h9. rewrite <- transfer_in in h9. inversion h9 as [w h15 y h17]. apply h1l in h17. rewrite <- h17 in h15. pose proof (proof_irrelevance _ h7 h3l) as h18. rewrite h18 in h15. rewrite <- transfer_r_eq in h15. rewrite transfer_undoes_transfer_r in h15. assumption. red. intros x h9. rewrite in_eq_rect. constructor. pose proof (exist_fos_psa_compat (Im A f)) as h10. rewrite h10. destruct (fos_psa_compat Yt) as [h11 h12]. pose proof h11 as h11'. symmetry in h11'. rewrite <- (transfer_dep_transfer_dep_r_compat h11' h11). simpl. unfold f'. pose proof (proof_irrelevance _ h11' h4l') as h13. rewrite h13. rewrite <- transfer_in. apply Im_intro with (transfer h3l (eq_rect_r id x h7)). rewrite <- transfer_r_eq. pose proof (proof_irrelevance _ h3l h7) as h14. rewrite h14, transfer_undoes_transfer_r. assumption. reflexivity. Qed. Lemma same_card_type_iso_psa' : forall (Xt Yt:Type), (exists f:Xt->Yt, bijective f) -> psa Xt =~ psa Yt. intros Xt Yt h1. assert (h2:exists g:Yt -> Xt, bijective g). destruct h1 as [f h1]. apply bijective_impl_invertible in h1. exists (proj1_sig (function_inverse f h1)). apply invertible_impl_bijective. apply invertible_impl_inv_invertible. eapply same_card_type_iso_psa; auto. Qed. Inductive complete_homo {A B:Bool_Alg} (f:bt A->bt B) : Prop := complete_homo_intro : homomorphism f -> (forall (S:Ensemble (bt A)) p, sup S p -> sup (Im S f) (f p)) -> complete_homo f. Inductive complete_homo_inv {A B:Bool_Alg} (f:bt A->bt B) : Prop := complete_homo_inv_intro : homomorphism f -> (forall (S:Ensemble (bt B)) p, sup S (f p) -> sup (inv_im S f) p) -> complete_homo_inv f. Lemma iso_complete : forall {A B:Bool_Alg} (f:bt A->bt B), isomorphism f -> complete_homo f. intros A B f h1. destruct h1 as [h2 h3]. constructor. assumption. intros S p h4. red in h4. destruct h4 as [h4l h4r]. assert (h5:ub (Im S f) (f p)). red. intros s h5. destruct h5 as [s h5]. subst. red in h4l. specialize (h4l _ h5). apply homo_mono; auto. red. split. auto. intros b' h6. apply bijective_impl_invertible in h3. pose (function_inverse _ h3) as f'. red in h6. apply NNPP. intro h7. assert (h8: ~ le p (proj1_sig f' b')). intro h8. apply (homo_mono _ h2) in h8. pose proof (proj2_sig (function_inverse f h3)) as h9. simpl in h9. destruct h9 as [h9l h9r]. rewrite h9r in h8. contradiction. specialize (h4r (proj1_sig f' b')). assert (h9:~ub S (proj1_sig f' b')). tauto. unfold ub in h9. apply not_all_ex_not in h9. destruct h9 as [c h9]. assert (h10:Inn S c /\ ~le c (proj1_sig f' b')). tauto. clear h9. destruct h10 as [h10l h10r]. specialize (h4l _ h10l). assert (h11:Inn (Im S f) (f c)). apply Im_intro with c; auto. specialize (h6 _ h11). pose proof (homo_inv _ h3 h2) as h12. apply (homo_mono _ h12) in h6. rewrite (match (proj2_sig (function_inverse f h3)) with | conj P _ => P end) in h6. contradiction. Qed. Lemma iso_complete_inv : forall {A B:Bool_Alg} (f:bt A->bt B), isomorphism f -> complete_homo_inv f. intros A B f h1. pose proof (iso_inv_compat _ h1) as h2. destruct h2 as [h2l h2r]. apply iso_complete in h2l. destruct h2l as [h3 h4]. constructor. destruct h1; auto. intros S p h5. specialize (h4 _ _ h5). destruct h2r as [h2a h2b]. rewrite h2a in h4. assert (h6:Im S (iso_inv f h1) = inv_im S f). apply Extensionality_Ensembles. red. split. red. intros x h6. destruct h6 as [y h6]. subst. constructor. rewrite h2b. assumption. red. intros x h6. destruct h6 as [h6]. apply Im_intro with (f x); auto. rewrite h6 in h4. assumption. Qed. Lemma mono_iso_bij_sig_im_fun : forall {A B:Bool_Alg} (f:bt A->bt B) (pf:monomorphism f), isomorphism (homo_onto _ (mono_homo _ pf)). intros A B f h1. generalize (mono_homo f h1). intro h2. destruct h1 as [h1l h1r]. pose proof (proof_irrelevance _ h1l h2); subst. constructor. apply homo_homo_onto. red. split. rewrite <- (inj_homo_onto_iff f h2); auto. pose proof (homo_onto_compat _ h2) as h3. destruct h3; assumption. Qed. Lemma mono_iso_bij_sig_im_fun_iff : forall {A B:Bool_Alg} (f:bt A->bt B) (pf:homomorphism f), Injective f <-> isomorphism (homo_onto _ pf). intros A B f h1. pose proof (homo_onto_compat _ h1) as h0. destruct h0 as [h0l h0r]. split. intro h2. constructor; auto. apply homo_homo_onto. red. split; auto. rewrite (inj_homo_onto_iff _ h1) in h2. assumption. intro h2. destruct h2 as [h2l h2r]. red in h2r. destruct h2r as [h2a h2b]. rewrite <- (inj_homo_onto_iff f h1) in h2a. assumption. Qed. Lemma mono_complete_regular_image : forall {A B:Bool_Alg} (f:bt B->bt A), monomorphism f -> (complete_homo f <-> regular_subalg _ (Im (BS (Bc B)) f)). intros A B f h1. split. Focus 2. intro h2. pose proof (mono_iso_bij_sig_im_fun _ h1) as h3. constructor; auto. destruct h1 as [h1l h1r]. assumption. intros Pi p h4. apply iso_complete in h3. destruct h3 as [h3a h3b]. specialize (h3b _ _ h4). destruct h2 as [h2]. destruct h1 as [h1a h1b]. specialize (h2 (homo_im_closed _ h1a)). assert (h5:Included (Im Pi f) (Im (BS (Bc B)) f)). red. intros x h5. destruct h5 as [x h5]. subst. apply Im_intro with x. rewrite und_set. constructor. reflexivity. specialize (h2 _ h5). pose proof (homo_onto_compat f (mono_homo f (monomorphism_intro f h1a h1b))) as h7. destruct h7 as [h7l h7r]. assert (h6:subset_sig (Im Pi f) (Im (BS (Bc B)) f) h5 = (Im Pi (homo_onto f (mono_homo f (monomorphism_intro f h1a h1b))))). apply Extensionality_Ensembles. red. split. red. intros x h6. destruct h6 as [x h6]. subst. destruct x as [x h7]. simpl. destruct h7 as [x h8]. subst. apply Im_intro with x. assumption. apply proj1_sig_injective. simpl. apply h7r. red. intros x h8. destruct h8 as [x h8]. subst. unfold subset_sig. assert (h9:Inn (Im Pi f) (f x)). apply Im_intro with x; auto. assert (h10:Inn (full_sig (Im Pi f)) (exist _ _ h9)). constructor. apply Im_intro with (exist (Inn (Im Pi f)) (f x) h9). assumption. apply proj1_sig_injective. simpl. symmetry. apply h7r. rewrite <- h6 in h3b at 1. assert (h8: Inn (Im (BS (Bc B)) f) (f p)). apply Im_intro with p. rewrite und_set. constructor. reflexivity. specialize (h2 _ h8). assert (h9:homo_onto f (mono_homo f (monomorphism_intro f h1a h1b)) p = (exist (Inn (Im (BS (Bc B)) f)) (f p) h8)). apply proj1_sig_injective. simpl. symmetry. apply h7r. rewrite h9 in h3b. apply h2. simpl in h3b. unfold sup. unfold sup in h3b. pose proof (proof_irrelevance _ h1a (mono_homo f (monomorphism_intro f h1a h1b))) as h10. rewrite h10. assumption. intro h2. constructor. intros h3 Qi h4 q h5 h6. destruct h5 as [p h5]. subst. destruct h2 as [h2l h2r]. pose proof (mono_iso_bij_sig_im_fun _ h1) as h7. pose (homo_onto f (mono_homo f h1)) as f'. pose proof (proof_irrelevance _ h3 (homo_im_closed f (mono_homo f h1))); subst. apply iso_complete_inv in h7. destruct h7 as [h7ll h7r]. specialize (h7r (subset_sig Qi (Im (BS (Bc B)) f) h4)). assert (h8: homo_onto f (mono_homo f h1) p = exist (Inn (Im (BS (Bc B)) f)) (f p) (Im_intro (bt B) (bt A) (BS (Bc B)) f p h5 (f p) eq_refl)). apply proj1_sig_injective. simpl. pose proof (homo_onto_compat f (mono_homo f h1)) as h7. destruct h7 as [h8 h9]. symmetry. apply h9. rewrite <- h8 in h6 at 1. specialize (h7r _ h6). apply h2r in h7r. pose proof (homo_onto_compat f (mono_homo f h1)) as h10. destruct h10 as [h10l h10r]. assert (h9: Im (inv_im (subset_sig Qi (Im (BS (Bc B)) f) h4) (homo_onto f (mono_homo f h1))) f = Qi). apply Extensionality_Ensembles. red. split. red. intros x h9. destruct h9 as [y h9]. subst. destruct h9 as [h9]. rewrite subset_sig_compat in h9. rewrite <- h10r in h9. assumption. red. intros x h9. red in h10l. pose proof h9 as h9'. apply h4 in h9. specialize (h10l (exist _ _ h9)). destruct h10l as [b h10l]. apply Im_intro with b. constructor. rewrite subset_sig_compat. pose proof (f_equal (@proj1_sig _ _) h10l) as h12. simpl in h12. subst. assumption. destruct h9. subst. pose proof (f_equal (@proj1_sig _ _) h10l) as h12. simpl in h12. rewrite <- h10r in h12. symmetry. assumption. simpl in h9. simpl in h7r. rewrite h9 in h7r at 1. assumption. Qed. Lemma degen_all_homo : forall {A B:Bool_Alg} (f:bt A->bt B), degen B -> homomorphism f. intros A B f h1. pose proof (degen_eq _ h1) as h2. apply homo_two_ops_plus. intros x y. apply (h2 (f (x + y)) (f x + f y)). intro x. apply (h2 (f (-x)) (- f x)). Qed. (*Analogues of all the above theorems for Bool_Alg_p instead of just Bool_Alg*) Section ParametricAnalogues. Variable T:Type. Record homomorphism_p {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp) : Prop := { homo_times_p : forall x y, f (x%*y) = (f x) %* (f y); homo_plus_p : forall x y, f (x%+y) = (f x) %+ (f y); homo_comp_p : forall x, f (%-x) = %- (f x) }. Record homomorphism_p1 {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B) : Prop := { homo_times_p1 : forall x y, f (x%*y) = (f x) * (f y); homo_plus_p1 : forall x y, f (x%+y) = (f x) + (f y); homo_comp_p1 : forall x, f (%-x) = - (f x) }. Record homomorphism_p2 {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) : Prop := { homo_times_p2 : forall x y, f (x*y) = (f x) %* (f y); homo_plus_p2 : forall x y, f (x+y) = (f x) %+ (f y); homo_comp_p2 : forall x, f (-x) = %- (f x) }. Lemma homomorphism_p_iff : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), homomorphism_p f <-> homomorphism (ba_conv_fun f). intros Ap Bp f. split. intro h1. destruct h1. constructor; auto. intro h1. destruct h1. constructor; auto. Qed. Lemma homomorphism_p1_iff : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 f <-> homomorphism (ba_conv_fun1 f). intros Ap Bp f. split. intro h1. destruct h1. constructor; auto. intro h1. destruct h1. constructor; auto. Qed. Lemma homomorphism_p2_iff : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), homomorphism_p2 f <-> homomorphism (ba_conv_fun2 f). intros A Bp f. split. intro h1. destruct h1. constructor; auto. intro h1. destruct h1. constructor; auto. Qed. Lemma homo_transfer_iff_p : forall (A B C:Bool_Alg_p T) (pf:A = B) (f:btp A -> btp C), homomorphism_p f <-> homomorphism_p (transfer (f_equal (fun D => btp D->btp C) pf) f). intros A B C h1 f. subst. simpl. rewrite transfer_eq_refl. tauto. Qed. Lemma homo_transfer_r_iff_p : forall (A B C:Bool_Alg_p T) (pf:A = B) (f:btp B -> btp C), homomorphism_p f <-> homomorphism_p (transfer_r (f_equal (fun D => btp D->btp C) pf) f). intros A B C h1 f. subst. simpl. rewrite transfer_r_eq_refl. tauto. Qed. Lemma homo_transfer_iff_p1 : forall (A B:Bool_Alg_p T) (C:Bool_Alg) (pf:A = B) (f:btp A -> bt C), homomorphism_p1 f <-> homomorphism_p1 (transfer (f_equal (fun D => btp D->bt C) pf) f). intros A B C h1 f. subst. simpl. rewrite transfer_eq_refl. tauto. Qed. Lemma homo_transfer_r_iff_p1 : forall (A B:Bool_Alg_p T) (C:Bool_Alg) (pf:A = B) (f:btp B -> bt C), homomorphism_p1 f <-> homomorphism_p1 (transfer_r (f_equal (fun D => btp D->bt C) pf) f). intros A B C h1 f. subst. simpl. rewrite transfer_r_eq_refl. tauto. Qed. Lemma homo_transfer_fun_iff_p : forall (A B C:Bool_Alg_p T), A = B -> forall (f:btp A -> btp C) (pf:btp A = btp B), homomorphism_p f <-> homomorphism_p (transfer_fun pf f). intros A B C h1 f h2. subst. pose proof (proof_irrelevance _ h2 (eq_refl _)); subst. rewrite transfer_fun_eq_refl. tauto. Qed. Lemma homo_transfer_fun_iff_p1 : forall (A B:Bool_Alg_p T) (C:Bool_Alg), A = B -> forall (f:btp A -> bt C) (pf:btp A = btp B), homomorphism_p1 f <-> homomorphism_p1 (transfer_fun pf f). intros A B C h1 f h2. subst. pose proof (proof_irrelevance _ h2 (eq_refl _)); subst. rewrite transfer_fun_eq_refl. tauto. Qed. Record homomorphism_general_p {Ap:Bool_Alg_p T} {A:Ensemble T} (f:btp Ap->sig_set A) (homo_ut_p : sig_set A -> sig_set A -> sig_set A) (homo_up_p : sig_set A -> sig_set A -> sig_set A) (homo_uc_p : sig_set A-> sig_set A) : Prop := { homo_gen_times_p : forall x y, f (x%*y) = homo_ut_p (f x) (f y); homo_gen_plus_p : forall x y, f (x%+y) = homo_up_p (f x) (f y); homo_gen_comp_p : forall x, f (%-x) = homo_uc_p (f x) }. Lemma homomorphism_general_p_iff : forall {Ap:Bool_Alg_p T} {A:Ensemble T} (f:btp Ap->sig_set A) (homo_ut_p : sig_set A -> sig_set A -> sig_set A) (homo_up_p : sig_set A -> sig_set A -> sig_set A) (homo_uc_p : sig_set A-> sig_set A), homomorphism_general_p f homo_ut_p homo_up_p homo_uc_p <-> homomorphism_general (ba_conv_fun1 f) homo_ut_p homo_up_p homo_uc_p. intros Ap T' f t p c. split. intro h1. destruct h1; constructor; auto. intro h1. destruct h1; constructor; auto. Qed. Lemma homo_sym_diff_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), homomorphism_p f -> forall x y, f (x /%\ y) = (f x) /%\ (f y). intros Ap Bp f h1 x y. rewrite homomorphism_p_iff in h1. apply (@homo_sym_diff (ba_conv Ap) (ba_conv Bp)); auto. Qed. Lemma homo_sym_diff_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 f -> forall x y, f (x /%\ y) = (f x) /_\ (f y). intros Ap B f h1 x y. rewrite homomorphism_p1_iff in h1. apply (@homo_sym_diff (ba_conv Ap) B); auto. Qed. Lemma homo_sym_diff_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), homomorphism_p2 f -> forall x y, f (x /_\ y) = (f x) /%\ (f y). intros A Bp f h1 x y. rewrite homomorphism_p2_iff in h1. apply (@homo_sym_diff A (ba_conv Bp)); auto. Qed. Lemma homo_id_p : forall Ap:Bool_Alg_p T, homomorphism_p id (Ap:=Ap). intros; constructor; auto. Qed. Definition homo2_p1 (Ap:Bool_Alg_p T) (B:Bool_Alg) : btp (two_bap Ap) -> bt B. intro x. destruct x as [x h1]. simpl in h1. apply couple_inv_dec in h1. destruct h1. refine 0. refine 1. Defined. Definition iso2_p1 (Ap:Bool_Alg_p T) (B:Bool_Alg) := homo2_p1 Ap (two B). Lemma universal_homomorphic_ops_imply_bool_p : forall {Ap:Bool_Alg_p T} {A:Ensemble T} (f:btp Ap->sig_set A) (ut_p : sig_set A -> sig_set A -> sig_set A) (up_p : sig_set A -> sig_set A -> sig_set A) (uc_p : sig_set A-> sig_set A) (pf:homomorphism_general_p f ut_p up_p uc_p), surjective f -> exists! Bp:Bool_Alg_p T, Bc_p T Bp = Build_Bconst_p T A (full_sig A) up_p ut_p (f %1) (f %0) uc_p. intros Ap A f t p c h1 h2. rewrite homomorphism_general_p_iff in h1. pose proof (@universal_homomorphic_ops_imply_bool (ba_conv Ap) (sig_set A) f t p c h1 h2) as h3. destruct h3 as [B h3]. red in h3. destruct h3 as [h3l h3r]. assert (h4:bt B = sig_set A). unfold bt. rewrite h3l. simpl. reflexivity. exists (ba_sig_set_conv A _ h4). red. split. rewrite ba_sig_set_conv_bc_p_compat. destruct B. simpl. simpl in h3l. simpl in h4. subst. unfold bc_sig_set_conv. simpl. simpl in h4. pose proof (proof_irrelevance _ h4 (eq_refl _)); subst. do 4 rewrite transfer_dep_eq_refl. do 2 rewrite transfer_eq_refl. f_equal. intros Bp h5. specialize (h3r (ba_conv Bp)). simpl in h3r. rewrite h5 in h3r. unfold bconst_conv in h3r. simpl in h3r. specialize (h3r (eq_refl _)). subst. apply bc_inj_p. rewrite ba_sig_set_conv_bc_p_compat. simpl. unfold bconst_conv. unfold bc_sig_set_conv. simpl. destruct Bp. destruct Bc_p. simpl. simpl in h4. simpl in h5. simpl in h3l. rewrite h5. inversion h3l as [ha]. unfold bt in h4. simpl in h4. rewrite (transfer_dep_eq_iff _ _ h4) in H. rewrite (transfer_dep_eq_iff _ _ h4) in H0. rewrite (transfer_dep_eq_iff _ _ h4) in H1. rewrite (transfer_dep_eq_iff _ _ h4) in H2. rewrite (transfer_dep_eq_iff _ _ h4) in H3. rewrite (transfer_dep_eq_iff _ _ h4) in H4. rewrite H, H0, H1, H2, H3, H4 at 1. reflexivity. Qed. Lemma homo_zero_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), homomorphism_p f -> f %0 = %0. intros Ap Bp f h1. rewrite homomorphism_p_iff in h1. apply (@homo_zero (ba_conv Ap) (ba_conv Bp)); auto. Qed. Lemma homo_zero_p' : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), homomorphism_p f -> forall (pf:Inn (ba_p_ens Ap) (proj1_sig (bzero_p Ap))), f (exist _ (proj1_sig %0) pf) = %0. intros Ap Bp f h1 h2. pose proof (proof_irrelevance _ h2 (proj2_sig (bzero_p Ap))); subst. rewrite <- unfold_sig. apply homo_zero_p; auto. Qed. Lemma homo_zero_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 f -> f %0 = 0. intros Ap B f h1. rewrite homomorphism_p1_iff in h1. apply (@homo_zero (ba_conv Ap) B); auto. Qed. Lemma homo_zero_p1' : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 f -> forall (pf:Inn (ba_p_ens Ap) (proj1_sig (bzero_p Ap))), f (exist _ (proj1_sig %0) pf) = 0. intros Ap Bp f h1 h2. pose proof (proof_irrelevance _ h2 (proj2_sig (bzero_p Ap))); subst. rewrite <- unfold_sig. apply homo_zero_p1; auto. Qed. Lemma homo_zero_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), homomorphism_p2 f -> f 0 = %0. intros A Bp f h1. rewrite homomorphism_p2_iff in h1. apply (@homo_zero A (ba_conv Bp)); auto. Qed. Lemma homo_one_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), homomorphism_p f -> f %1 = %1. intros Ap Bp f h1. rewrite homomorphism_p_iff in h1. apply (@homo_one (ba_conv Ap) (ba_conv Bp)); auto. Qed. Lemma homo_one_p' : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), homomorphism_p f -> forall (pf:Inn (ba_p_ens Ap) (proj1_sig (bone_p Ap))), f (exist _ (proj1_sig %1) pf) = %1. intros Ap Bp f h1 h2. pose proof (proof_irrelevance _ h2 (proj2_sig (bone_p Ap))); subst. rewrite <- unfold_sig. apply homo_one_p; auto. Qed. Lemma homo_one_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 f -> f %1 = 1. intros Ap B f h1. rewrite homomorphism_p1_iff in h1. apply (@homo_one (ba_conv Ap) B); auto. Qed. Lemma homo_one_p1' : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 f -> forall (pf:Inn (ba_p_ens Ap) (proj1_sig (bone_p Ap))), f (exist _ (proj1_sig %1) pf) = 1. intros Ap Bp f h1 h2. pose proof (proof_irrelevance _ h2 (proj2_sig (bone_p Ap))); subst. rewrite <- unfold_sig. apply homo_one_p1; auto. Qed. Lemma homo_one_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), homomorphism_p2 f -> f 1 = %1. intros A Bp f h1. rewrite homomorphism_p2_iff in h1. apply (@homo_one A (ba_conv Bp)); auto. Qed. Lemma homo_two_ops_plus_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), (forall x y, f (x %+ y) = (f x) %+ (f y)) -> (forall x, f (%-x) = %- (f x)) -> homomorphism_p f. intros Ap Bp f h1 h2. rewrite homomorphism_p_iff. apply (homo_two_ops_plus (ba_conv_fun f) h1 h2). Qed. Lemma homo_two_ops_plus_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), (forall x y, f (x %+ y) = (f x) + (f y)) -> (forall x, f (%-x) = - (f x)) -> homomorphism_p1 f. intros Ap B f h1 h2. rewrite homomorphism_p1_iff. apply (homo_two_ops_plus (ba_conv_fun1 f) h1 h2). Qed. Lemma homo_two_ops_plus_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), (forall x y, f (x + y) = (f x) %+ (f y)) -> (forall x, f (-x) = %- (f x)) -> homomorphism_p2 f. intros Ap B f h1 h2. rewrite homomorphism_p2_iff. apply (homo_two_ops_plus (ba_conv_fun2 f) h1 h2). Qed. Lemma homo_two_ops_times_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), (forall x y, f (x %* y) = (f x) %* (f y)) -> (forall x, f (%-x) = %- (f x)) -> homomorphism_p f. intros Ap Bp f h1 h2. rewrite homomorphism_p_iff. apply (homo_two_ops_times (ba_conv_fun f) h1 h2). Qed. Lemma homo_two_ops_times_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), (forall x y, f (x %* y) = (f x) * (f y)) -> (forall x, f (%-x) = - (f x)) -> homomorphism_p1 f. intros Ap B f h1 h2. rewrite homomorphism_p1_iff. apply (homo_two_ops_times (ba_conv_fun1 f) h1 h2). Qed. Lemma homo_two_ops_times_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), (forall x y, f (x * y) = (f x) %* (f y)) -> (forall x, f (-x) = %- (f x)) -> homomorphism_p2 f. intros Ap B f h1 h2. rewrite homomorphism_p2_iff. apply (homo_two_ops_times (ba_conv_fun2 f) h1 h2). Qed. Lemma ba_conv_fun_compose : forall {Ap Bp Cp:Bool_Alg_p T} (f:btp Ap->btp Bp) (g:btp Bp->btp Cp), ba_conv_fun (compose g f) = compose (ba_conv_fun g) (ba_conv_fun f). intros Ap Bp Cp f g. unfold ba_conv_fun. unfold compose. reflexivity. Qed. Lemma homo_compose_p : forall {Ap Bp Cp:Bool_Alg_p T} (f:btp Ap->btp Bp) (g:btp Bp->btp Cp), homomorphism_p f -> homomorphism_p g -> homomorphism_p (compose g f). intros Ap Bp Cp f g h1 h2. rewrite homomorphism_p_iff in h1. rewrite homomorphism_p_iff in h2. pose proof (homo_compose _ _ h1 h2) as h3. rewrite homomorphism_p_iff. rewrite ba_conv_fun_compose. assumption. Qed. Lemma ba_conv_fun_inv_iff : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), invertible f <-> invertible (ba_conv_fun f). intros Ap Bp f. split. intro h1. destruct h1 as [g h2 h3]. apply (intro_invertible (ba_conv_fun f) (ba_conv_fun g)); auto. intro h1. destruct h1 as [g h2 h3]. apply (intro_invertible f g); auto. Qed. Lemma ba_conv_fun1_inv_iff : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), invertible f <-> invertible (ba_conv_fun1 f). intros Ap B f. split. intro h1. destruct h1 as [g h2 h3]. apply (intro_invertible (ba_conv_fun1 f) (ba_conv_fun2 g)); auto. intro h1. destruct h1 as [g h2 h3]. apply (intro_invertible f g); auto. Qed. Lemma ba_conv_fun2_inv_iff : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), invertible f <-> invertible (ba_conv_fun2 f). intros A Bp f. split. intro h1. destruct h1 as [g h2 h3]. apply (intro_invertible (ba_conv_fun2 f) (ba_conv_fun1 g)); auto. intro h1. destruct h1 as [g h2 h3]. apply (intro_invertible f g); auto. Qed. Lemma ba_conv_fun_function_inverse : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp) (pf:invertible f), ba_conv_fun (proj1_sig (function_inverse f pf)) = proj1_sig (function_inverse (ba_conv_fun f) (iff1 (ba_conv_fun_inv_iff f) pf)). intros Ap Bp f h1. unfold ba_conv_fun, function_inverse. apply functional_extensionality. intro x. rewrite (proof_irrelevance _ (unique_inverse f h1) (unique_inverse (fun x0 : Btype (Bc (ba_conv Ap)) => f x0) (iff1 (ba_conv_fun_inv_iff f) h1))); auto. Qed. Lemma ba_conv_fun1_function_inverse : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B) (pf:invertible f), ba_conv_fun2 (proj1_sig (function_inverse f pf)) = proj1_sig (function_inverse (ba_conv_fun1 f) (iff1 (ba_conv_fun1_inv_iff f) pf)). intros Ap B f h1. unfold ba_conv_fun1, ba_conv_fun2, function_inverse. apply functional_extensionality. intro x. rewrite (proof_irrelevance _ (unique_inverse f h1) (unique_inverse (fun x0 : Btype (Bc (ba_conv Ap)) => f x0) (iff1 (ba_conv_fun1_inv_iff f) h1))); auto. Qed. Lemma ba_conv_fun2_function_inverse : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) (pf:invertible f), ba_conv_fun1 (proj1_sig (function_inverse f pf)) = proj1_sig (function_inverse (ba_conv_fun2 f) (iff1 (ba_conv_fun2_inv_iff f) pf)). intros A Bp f h1. unfold ba_conv_fun1, ba_conv_fun2, function_inverse. apply functional_extensionality. intro x. rewrite (proof_irrelevance _ (unique_inverse f h1) (unique_inverse (fun x0 : bt A => f x0) (iff1 (ba_conv_fun2_inv_iff f) h1))); auto. Qed. Lemma homo_inv_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp) (pf:invertible f), homomorphism_p f -> homomorphism_p (proj1_sig (function_inverse _ pf)). intros Ap Bp f h1 h2. rewrite homomorphism_p_iff in h2. pose proof (homo_inv (ba_conv_fun f) (iff1 (ba_conv_fun_inv_iff f) h1) h2) as h3. rewrite homomorphism_p_iff. rewrite ba_conv_fun_function_inverse. assumption. Qed. Lemma homo_inv_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B) (pf:invertible f), homomorphism_p1 f -> homomorphism_p2 (proj1_sig (function_inverse _ pf)). intros Ap B f h1 h2. rewrite homomorphism_p1_iff in h2. pose proof (homo_inv (ba_conv_fun1 f) (iff1 (ba_conv_fun1_inv_iff f) h1) h2) as h3. rewrite homomorphism_p2_iff. rewrite ba_conv_fun1_function_inverse. assumption. Qed. Lemma homo_inv_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) (pf:invertible f), homomorphism_p2 f -> homomorphism_p1 (proj1_sig (function_inverse _ pf)). intros A Bp f h1 h2. rewrite homomorphism_p2_iff in h2. pose proof (homo_inv (ba_conv_fun2 f) (iff1 (ba_conv_fun2_inv_iff f) h1) h2) as h3. rewrite homomorphism_p1_iff. rewrite ba_conv_fun2_function_inverse. assumption. Qed. Lemma homo_mono_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), homomorphism_p f -> forall x y, le_p x y -> le_p (f x) (f y). intros Ap Bp f h1 x y h2. rewrite homomorphism_p_iff in h1. rewrite le_p_iff in h2. rewrite le_p_iff. apply (homo_mono _ h1); auto. Qed. Lemma homo_mono_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 f -> forall x y, le_p x y -> le (f x) (f y). intros Ap B f h1 x y h2. rewrite homomorphism_p1_iff in h1. rewrite le_p_iff in h2. apply (homo_mono _ h1); auto. Qed. Lemma homo_mono_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), homomorphism_p2 f -> forall x y, le x y -> le_p (f x) (f y). intros A Bp f h1 x y h2. rewrite homomorphism_p2_iff in h1. rewrite le_p_iff. apply (homo_mono _ h1); auto. Qed. Lemma im_ba_conv_fun_eq : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp) (pf:Included (Im (Im (BS_p T (Bc_p T Ap)) f) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x))) (A_p T (Bc_p T Bp))), Im (BS (Bc (ba_conv Ap))) (ba_conv_fun f) = (ba_conv_und_subalg T Bp (Im (Im (BS_p T (Bc_p T Ap)) f) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x))) pf). intros Ap Bp f h1 . apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. unfold ba_conv_und_subalg. unfold ba_conv_set. unfold transfer_dep. unfold eq_rect_r. simpl. pose proof (Im_intro _ _ _ f _ h2 _ (eq_refl _)) as h3. pose proof (Im_intro _ _ _ (@proj1_sig _ _) _ h3 _ (eq_refl _)) as h4. pose (exist _ _ h4) as x'. apply Im_intro with x'. constructor. unfold x'. simpl. unfold ba_conv_fun. apply proj1_sig_injective. simpl. reflexivity. red. intros x h2. destruct h2 as [x h2]. subst. destruct x as [x h3]. simpl. inversion h3 as [y h4]. subst. destruct h4 as [y h4]. subst. apply Im_intro with y. assumption. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma im_ba_conv_fun1_eq : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), Im (BS (Bc (ba_conv Ap))) (ba_conv_fun1 f) = (Im (BS_p T (Bc_p T Ap)) f). intros Ap B f. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. destruct x as [x h0]. apply Im_intro with (exist _ _ h0). apply in_bs_p. unfold ba_conv_fun1. f_equal. red. intros x h1. destruct h1 as [x h1]. subst. apply Im_intro with x. apply in_bs. unfold ba_conv_fun1. reflexivity. Qed. Lemma im_ba_conv_fun2_eq : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A-> btp Bp) (pf:Included (Im (Im (ba_ens A) f) (@proj1_sig _ _)) (A_p T (Bc_p T Bp))), Im (BS (Bc A)) (ba_conv_fun2 f) = ba_conv_und_subalg T Bp (Im (Im (ba_ens A) f) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x))) pf. intros A Bp f h1. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. unfold ba_conv_und_subalg. unfold ba_conv_set. unfold transfer_dep. unfold eq_rect_r. simpl. rewrite und_set in h2. pose proof (Im_intro _ _ _ f _ h2 _ (eq_refl _)) as h3. pose proof (Im_intro _ _ _ (@proj1_sig _ _) _ h3 _ (eq_refl _)) as h4. pose (exist _ _ h4) as x'. apply Im_intro with x'. constructor. unfold x'. simpl. unfold ba_conv_fun. apply proj1_sig_injective. simpl. reflexivity. red. intros x h2. destruct h2 as [x h2]. subst. destruct x as [x h3]. simpl. inversion h3 as [y h4]. subst. destruct h4 as [y h4]. subst. apply Im_intro with y. unfold ba_ens, bt in h4. rewrite <- (und_set A) in h4. assumption. apply proj1_sig_injective. simpl. reflexivity. Qed. Lemma homo_im_closed_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), homomorphism_p f -> exists pf:Included (Im (Im (BS_p T (Bc_p T Ap)) f) (@proj1_sig _ _)) (A_p T (Bc_p T Bp)), alg_closed_p (Im (Im (BS_p T (Bc_p T Ap)) f) (@proj1_sig _ _)) pf. intros Ap Bp f h1. rewrite homomorphism_p_iff in h1. ex_goal. rewrite im_im. red. intros x h2. destruct h2 as [x h2]. subst. apply proj2_sig. exists hex. rewrite alg_closed_p_iff. pose proof (homo_im_closed _ h1) as h3. rewrite (im_ba_conv_fun_eq _ hex) in h3 at 1. assumption. Qed. Lemma homo_im_closed_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 f -> alg_closed (Im (BS_p T (Bc_p T Ap)) f). intros Ap B f h1. rewrite homomorphism_p1_iff in h1. pose proof (homo_im_closed _ h1) as h3. pose proof (im_ba_conv_fun1_eq f) as h4. unfold bt in h3, h4. rewrite h4 in h3. assumption. Qed. Lemma homo_im_closed_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), homomorphism_p2 f -> exists pf:Included (Im (Im (ba_ens A) f) (@proj1_sig _ _)) (A_p T (Bc_p T Bp)), alg_closed_p (Im (Im (ba_ens A) f) (@proj1_sig _ _)) pf. intros A Bp f h1. rewrite homomorphism_p2_iff in h1. assert (h2:Included (Im (Im (ba_ens A) f) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x))) (A_p T (Bc_p T Bp))). rewrite im_im. red. intros x h2. destruct h2 as [x h2]. subst. apply proj2_sig. exists h2. rewrite alg_closed_p_iff. pose proof (homo_im_closed _ h1) as h3. pose proof (im_ba_conv_fun2_eq f h2) as h4. unfold bt in h3, h4. rewrite h4 in h3. assumption. Qed. Lemma homo_subalg_im_p_ex : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), homomorphism_p f -> exists! (Cp:Bool_Alg_p T), exists (pfincl:Included (Im (Im (BS_p T (Bc_p T Ap)) f) (@proj1_sig _ _)) (A_p T (Bc_p T Bp))) (pfcl:alg_closed_p (Im (Im (BS_p T (Bc_p T Ap)) f) (@proj1_sig _ _)) pfincl), Cp = Subalg_p _ _ pfincl pfcl. intros Ap Bp f h1. pose proof (homo_im_closed_p _ h1) as h2. destruct h2 as [h2 h3]. exists (Subalg_p _ _ h2 h3). red. split. exists h2. exists h3. reflexivity. intros Cp h4. destruct h4 as [h4 [h5 h6]]. subst. pose proof (proof_irrelevance _ h4 h2); subst. pose proof (proof_irrelevance _ h3 h5); subst; auto. Qed. Definition homo_subalg_im_p {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp) (pf:homomorphism_p f) : Bool_Alg_p T := proj1_sig (constructive_definite_description _ (homo_subalg_im_p_ex f pf)). Lemma homo_subalg_im_p_compat : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp) (pf:homomorphism_p f), let Cp := homo_subalg_im_p _ pf in exists (pfincl:Included (Im (Im (BS_p T (Bc_p T Ap)) f) (@proj1_sig _ _)) (A_p T (Bc_p T Bp))) (pfcl:alg_closed_p (Im (Im (BS_p T (Bc_p T Ap)) f) (@proj1_sig _ _)) pfincl), Cp = Subalg_p _ _ pfincl pfcl. intros Ap f h1 h2. simpl. unfold homo_subalg_im_p. destruct constructive_definite_description as [Cp h3]. simpl. assumption. Qed. Lemma homo_subalg_im_p1_ex : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 f -> exists! (C:Bool_Alg), exists (pfcl:alg_closed (Im (BS_p T (Bc_p T Ap)) f)), C = Subalg _ pfcl. intros Ap B f h1. pose proof (homo_im_closed_p1 _ h1) as h2. exists (Subalg _ h2). red. split. exists h2. reflexivity. intros Cp h4. destruct h4 as [h4 h5]. subst. pose proof (proof_irrelevance _ h4 h2); subst; auto. Qed. Definition homo_subalg_im_p1 {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B) (pf:homomorphism_p1 f) : Bool_Alg := proj1_sig (constructive_definite_description _ (homo_subalg_im_p1_ex f pf)). Lemma homo_subalg_im_p1_compat : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B) (pf:homomorphism_p1 f), let C:= homo_subalg_im_p1 f pf in exists (pfcl:alg_closed (Im (BS_p T (Bc_p T Ap)) f)), C = Subalg _ pfcl. intros Ap B f h1 C. unfold C. unfold homo_subalg_im_p1. destruct constructive_definite_description as [C' h2]. simpl. assumption. Qed. Lemma homo_subalg_im_p2_ex : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), homomorphism_p2 f -> exists! (Cp:Bool_Alg_p T), exists (pfincl:Included (Im (Im (ba_ens A) f) (@proj1_sig _ _)) (A_p T (Bc_p T Bp))) (pfcl:alg_closed_p (Im (Im (ba_ens A) f) (@proj1_sig _ _)) pfincl), Cp = Subalg_p _ _ pfincl pfcl. intros A Bp f h1. pose proof (homo_im_closed_p2 _ h1) as h2. destruct h2 as [h2 h3]. exists (Subalg_p _ _ h2 h3). red. split. exists h2. exists h3. reflexivity. intros x h4. destruct h4 as [h4 [h5 h6]]. subst. pose proof (proof_irrelevance _ h2 h4); subst. pose proof (proof_irrelevance _ h3 h5); subst; auto. Qed. Definition homo_subalg_im_p2 {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) (pf:homomorphism_p2 f) := proj1_sig (constructive_definite_description _ (homo_subalg_im_p2_ex f pf)). Lemma homo_subalg_im_p2_compat : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) (pf:homomorphism_p2 f), let Cp := homo_subalg_im_p2 f pf in exists (pfincl:Included (Im (Im (ba_ens A) f) (@proj1_sig _ _)) (A_p T (Bc_p T Bp))) (pfcl:alg_closed_p (Im (Im (ba_ens A) f) (@proj1_sig _ _)) pfincl), Cp = Subalg_p _ _ pfincl pfcl. intros A Bp f h1 Cp. unfold Cp, homo_subalg_im_p2. destruct constructive_definite_description as [C' h2]. simpl. assumption. Qed. Lemma transfer_homo_subalg_im_p_ex : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp) (pf:homomorphism_p f) (x:btp Ap), exists! y:btp (homo_subalg_im_p f pf), proj1_sig y = proj1_sig (f x). intros Ap Bp f h1 x. pose proof (homo_subalg_im_p_compat f h1) as h2. simpl in h2. destruct h2 as [h2 [h3 h4]]. assert (h5:Inn (ba_p_ens (homo_subalg_im_p f h1)) (proj1_sig (f x))). unfold ba_p_ens. rewrite h4. unfold Subalg_p. simpl. apply Im_intro with (f x); auto. apply Im_intro with x; auto. rewrite und_set_p. constructor. exists (exist _ _ h5). red. split. simpl. reflexivity. intros y h6. apply proj1_sig_injective. simpl. rewrite h6 at 1. reflexivity. Qed. Definition transfer_homo_subalg_im_p {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp) (pf:homomorphism_p f) (x:btp Ap) : btp (homo_subalg_im_p f pf) := proj1_sig (constructive_definite_description _ (transfer_homo_subalg_im_p_ex _ pf x)). Lemma transfer_homo_subalg_im_p_compat : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp) (pf:homomorphism_p f) (x:btp Ap), let y := transfer_homo_subalg_im_p f pf x in proj1_sig y = proj1_sig (f x). intros Ap Bp f h1 x y. unfold transfer_homo_subalg_im_p in y. destruct constructive_definite_description as [z h2]. simpl in y. unfold y. assumption. Qed. Lemma transfer_homo_subalg_im_p1_ex : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B) (pf:homomorphism_p1 f) (x:btp Ap), let C:=homo_subalg_im_p1 f pf in exists! y:bt C, forall pfcl: alg_closed (Im (BS_p T (Bc_p T Ap)) f), exists (pfeq:C = Subalg _ pfcl), proj1_sig (transfer_ba_elt pfeq y) = f x. intros Ap B f h1 x C. pose proof (homo_subalg_im_p1_compat f h1) as h2. simpl in h2. destruct h2 as [h2 h3]. assert (h4:Inn (Im (BS_p T (Bc_p T Ap)) f) (f x)). apply Im_intro with x. apply in_bs_p. reflexivity. exists (transfer_ba_elt_r h3 (exist _ _ h4)). red. split. intro h5. pose proof (proof_irrelevance _ h2 h5); subst. exists h3. rewrite transfer_ba_elt_undoes_transfer_ba_elt_r. simpl. reflexivity. intros x' h5. specialize (h5 h2). destruct h5 as [h5 h6]. subst. apply (transfer_ba_elt_inj h3). rewrite transfer_ba_elt_undoes_transfer_ba_elt_r. apply proj1_sig_injective. pose proof (proof_irrelevance _ h3 h5); subst. rewrite h6 at 1. reflexivity. Qed. Definition transfer_homo_subalg_im_p1 {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B) (pf:homomorphism_p1 f) (x:btp Ap) := proj1_sig (constructive_definite_description _ (transfer_homo_subalg_im_p1_ex f pf x)). Lemma transfer_homo_subalg_im_p1_compat : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B) (pf:homomorphism_p1 f) (x:btp Ap), let C := homo_subalg_im_p1 f pf in let y := transfer_homo_subalg_im_p1 f pf x in forall pfcl: alg_closed (Im (BS_p T (Bc_p T Ap)) f), exists (pfeq:C = Subalg _ pfcl), proj1_sig (transfer_ba_elt pfeq y) = f x. intros Ap B f h1 x C y. unfold y, C, transfer_homo_subalg_im_p1. destruct constructive_definite_description as [y' h2]. simpl. assumption. Qed. Lemma transfer_homo_subalg_im_p2_ex : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) (pf:homomorphism_p2 f) (x:bt A), exists! y:btp (homo_subalg_im_p2 f pf), proj1_sig y = proj1_sig (f x). intros Ap Bp f h1 x. pose proof (homo_subalg_im_p2_compat f h1) as h2. simpl in h2. destruct h2 as [h2 [h3 h4]]. assert (h5:Inn (ba_p_ens (homo_subalg_im_p2 f h1)) (proj1_sig (f x))). unfold ba_p_ens. rewrite h4. unfold Subalg_p. simpl. apply Im_intro with (f x); auto. apply Im_intro with x; auto. constructor. exists (exist _ _ h5). red. split. simpl. reflexivity. intros y h6. apply proj1_sig_injective. simpl. rewrite h6 at 1. reflexivity. Qed. Definition transfer_homo_subalg_im_p2 {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) (pf:homomorphism_p2 f) (x:bt A) := proj1_sig (constructive_definite_description _ (transfer_homo_subalg_im_p2_ex f pf x)). Lemma transfer_homo_subalg_im_p2_compat : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) (pf:homomorphism_p2 f) (x:bt A), let y:= transfer_homo_subalg_im_p2 f pf x in proj1_sig y = proj1_sig (f x). intros A Bp f h1 x y. unfold y, transfer_homo_subalg_im_p2. destruct constructive_definite_description as [y' h2]. simpl. assumption. Qed. Definition homo_onto_p {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp) (pf:homomorphism_p f) : btp Ap->btp (homo_subalg_im_p _ pf) := (fun x => (transfer_homo_subalg_im_p f pf x)). Definition homo_onto_p1 {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B) (pf:homomorphism_p1 f) : btp Ap->bt (homo_subalg_im_p1 _ pf) := (fun x => (transfer_homo_subalg_im_p1 f pf x)). Definition homo_onto_p2 {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) (pf:homomorphism_p2 f) : bt A->btp (homo_subalg_im_p2 _ pf) := (fun x => (transfer_homo_subalg_im_p2 f pf x)). Lemma homomorphism_p1_restriction_sig : forall {Bp:Bool_Alg_p T} (A:Bool_Alg) (f:(btp Bp)->(bt A)), homomorphism_p1 f -> forall (Cp:Bool_Alg_p T) (pf:Included (ba_p_ens Cp) (ba_p_ens Bp)), subalg_of_p Cp Bp -> homomorphism_p1 (restriction_sig f _ pf). intros Bp A f h1 Cp h2 h3. red in h3. destruct h3 as [h3a [h3b h3c]]. destruct h1 as [h1a h1b h1c]. apply homo_two_ops_plus_p1. intros x y. destruct x as [x h4], y as [y h5]. unfold restriction_sig. simpl. assert (h8:Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h3a h3b)) x). rewrite <- h3c. assumption. assert (h9: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h3a h3b)) y). rewrite <- h3c. assumption. pose proof (ba_p_subst_plus _ _ h3c _ _ h4 h5 h8 h9) as h6. rewrite <- h1b. f_equal. apply proj1_sig_injective. simpl. rewrite h6 at 1. unfold Subalg_p. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. intros x. destruct x as [x h4]. unfold restriction_sig. simpl. assert (h8:Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h3a h3b)) x). rewrite <- h3c. assumption. pose proof (ba_p_subst_comp _ _ h3c _ h4 h8) as h6. rewrite <- h1c. f_equal. apply proj1_sig_injective. simpl. rewrite h6 at 1. unfold Subalg_p. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. Qed. Lemma ba_p_subst_homomorphism_p1 : forall (Ap Bp:Bool_Alg_p T) (C:Bool_Alg) (pf:Ap = Bp) (pf':btp Ap = btp Bp) (f:btp Ap->bt C), homomorphism_p1 f <-> homomorphism_p1 (transfer_fun pf' f). intros Ap Bp C h0 h1 f. subst. pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. rewrite transfer_fun_eq_refl. tauto. Qed. Inductive monomorphism_p {Ap Bp:Bool_Alg_p T} (f:(btp Ap)->(btp Bp)) : Prop := | monomorphism_intro_p : homomorphism_p f -> Injective f -> monomorphism_p f. Inductive monomorphism_p1 {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:(btp Ap)->(bt B)) : Prop := | monomorphism_intro_p1 : homomorphism_p1 f -> Injective f -> monomorphism_p1 f. Inductive monomorphism_p2 {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) : Prop := | monomorphism_intro_p2 : homomorphism_p2 f -> Injective f -> monomorphism_p2 f. Lemma mono_homo_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), monomorphism_p f -> homomorphism_p f. intros A B f h1; destruct h1; auto. Qed. Lemma mono_homo_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:(btp Ap)->(bt B)), monomorphism_p1 f -> homomorphism_p1 f. intros A B f h1. destruct h1; auto. Qed. Lemma mono_homo_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), monomorphism_p2 f -> homomorphism_p2 f. intros A B f h1. destruct h1; auto. Qed. Lemma monomorphism_p_iff : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap -> btp Bp), monomorphism_p f <-> monomorphism (ba_conv_fun f). intros Ap Bp f. split. intro h1. destruct h1 as [h1 h2]. rewrite homomorphism_p_iff in h1. constructor; auto. intro h1. destruct h1 as [h1 h2]. constructor; auto. rewrite homomorphism_p_iff; auto. Qed. Lemma monomorphism_p1_iff : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:(btp Ap)->(bt B)), monomorphism_p1 f <-> monomorphism (ba_conv_fun1 f). intros Ap B f. split. intro h1. destruct h1 as [h1 h2]. rewrite homomorphism_p1_iff in h1. constructor; auto. intro h1. destruct h1 as [h1 h2]. constructor; auto. rewrite homomorphism_p1_iff; auto. Qed. Lemma monomorphism_p2_iff : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), monomorphism_p2 f <-> monomorphism (ba_conv_fun2 f). intros A Bp f. split. intro h1. destruct h1 as [h1 h2]. rewrite homomorphism_p2_iff in h1. constructor; auto. intro h1. destruct h1 as [h1 h2]. constructor; auto. rewrite homomorphism_p2_iff; auto. Qed. Lemma mono_transfer_iff_p : forall (Ap Bp Cp:Bool_Alg_p T), Ap = Bp -> forall (f:btp Ap -> btp Cp) (pf:btp Ap = btp Bp), monomorphism_p f <-> monomorphism_p (transfer_fun pf f). intros A B C h1 f h2. subst. pose proof (proof_irrelevance _ h2 (eq_refl _)); subst. rewrite transfer_fun_eq_refl. tauto. Qed. Lemma mono_transfer_iff_p1 : forall (Ap Bp:Bool_Alg_p T) (C:Bool_Alg), Ap = Bp -> forall (f:btp Ap -> bt C) (pf:btp Ap = btp Bp), monomorphism_p1 f <-> monomorphism_p1 (transfer_fun pf f). intros A B C h1 f h2. subst. pose proof (proof_irrelevance _ h2 (eq_refl _)); subst. rewrite transfer_fun_eq_refl. tauto. Qed. Lemma mono_transfer_iff_p2 : forall (A B:Bool_Alg) (Cp:Bool_Alg_p T), A = B -> forall (f:bt A -> btp Cp) (pf:bt A = bt B), monomorphism_p2 f <-> monomorphism_p2 (transfer_fun pf f). intros A B Cp h1 f h2. subst. pose proof (proof_irrelevance _ h2 (eq_refl _)); subst. rewrite transfer_fun_eq_refl. tauto. Qed. Lemma mono_transfer_sig_iff_p1 : forall (Ap Bp:Bool_Alg_p T) (C:Bool_Alg), Ap = Bp -> forall (f:btp Ap -> bt C) (pf:ba_p_ens Ap = ba_p_ens Bp), monomorphism_p1 f <-> monomorphism_p1 (transfer_fun_sig pf f). intros Ap Bp C h1 f h2. subst. pose proof (proof_irrelevance _ h2 (eq_refl _)); subst. rewrite transfer_fun_sig_eq_refl. tauto. Qed. Lemma mono_zero_rev_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), monomorphism_p f -> forall (x:btp Ap), f x = %0 -> x = %0. intros Ap Bp f h1 x h2. rewrite monomorphism_p_iff in h1. pose proof (mono_zero_rev _ h1 _ h2). subst. reflexivity. Qed. Lemma mono_zero_rev_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), monomorphism_p1 f -> forall (x:btp Ap), f x = 0 -> x = %0. intros Ap Bp f h1 x h2. rewrite monomorphism_p1_iff in h1. pose proof (mono_zero_rev _ h1 _ h2). subst. reflexivity. Qed. Lemma mono_zero_rev_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), monomorphism_p2 f -> forall (x:bt A), f x = %0 -> x = 0. intros Ap Bp f h1 x h2. rewrite monomorphism_p2_iff in h1. pose proof (mono_zero_rev _ h1 _ h2). subst. reflexivity. Qed. Lemma mono_one_rev_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), monomorphism_p f -> forall (x:btp Ap), f x = %1 -> x = %1. intros Ap Bp f h1 x h2. rewrite monomorphism_p_iff in h1. pose proof (mono_one_rev _ h1 _ h2). subst. reflexivity. Qed. Lemma mono_one_rev_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), monomorphism_p1 f -> forall (x:btp Ap), f x = 1 -> x = %1. intros Ap Bp f h1 x h2. rewrite monomorphism_p1_iff in h1. pose proof (mono_one_rev _ h1 _ h2). subst. reflexivity. Qed. Lemma mono_one_rev_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), monomorphism_p2 f -> forall (x:bt A), f x = %1 -> x = 1. intros Ap Bp f h1 x h2. rewrite monomorphism_p2_iff in h1. pose proof (mono_one_rev _ h1 _ h2). subst. reflexivity. Qed. Inductive isomorphism_p1 {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:(btp Ap)->(bt B)) : Prop := | isomorphism_intro_p1 : homomorphism_p1 f -> bijective f -> isomorphism_p1 f. Inductive isomorphism_p2 {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) := | isomorphism_intro_p2 : homomorphism_p2 f -> bijective f -> isomorphism_p2 f. Lemma iso_mono_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:(btp Ap)->(bt B)), isomorphism_p1 f -> monomorphism_p1 f. intros A B f h1. destruct h1 as [h1 h2]. destruct h2 as [h2 h3]. constructor; auto. Qed. Lemma iso_mono_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), isomorphism_p2 f -> monomorphism_p2 f. intros A B f h1. destruct h1 as [h1 h2]. destruct h2 as [h2 h3]. constructor; auto. Qed. Lemma iso_homo_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:(btp Ap)->(bt B)), isomorphism_p1 f -> homomorphism_p1 f. intros A B f h1. destruct h1; auto. Qed. Lemma iso_homo_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), isomorphism_p2 f -> homomorphism_p2 f. intros A Bp f h1. destruct h1; auto. Qed. Lemma isomorphism_p1_iff : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:(btp Ap)->(bt B)), isomorphism_p1 f <-> isomorphism (ba_conv_fun1 f). intros Ap B f. split. intro h1. destruct h1 as [h1 h2]. rewrite homomorphism_p1_iff in h1. constructor; auto. intro h1. destruct h1 as [h1 h2]. constructor; auto. rewrite homomorphism_p1_iff; auto. Qed. Lemma isomorphism_p2_iff : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp), isomorphism_p2 f <-> isomorphism (ba_conv_fun2 f). intros A Bp f. split. intro h1. destruct h1 as [h1 h2]. rewrite homomorphism_p2_iff in h1. constructor; auto. intro h1. destruct h1 as [h1 h2]. constructor; auto. rewrite homomorphism_p2_iff; auto. Qed. Lemma homo_times_set_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 f -> forall (E:Ensemble (btp Ap)) (pf:Finite E), f (times_set_p E pf) = times_set (Im E f) (finite_image _ _ _ f pf). intros Ap B f h1 E h2. rewrite times_set_p_eq. rewrite homomorphism_p1_iff in h1. unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. rewrite homo_times_set; auto. Qed. Lemma homo_plus_set_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 f -> forall (E:Ensemble (btp Ap)) (pf:Finite E), f (plus_set_p E pf) = plus_set (Im E f) (finite_image _ _ _ f pf). intros Ap B f h1 E h2. rewrite plus_set_p_eq. rewrite homomorphism_p1_iff in h1. unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. rewrite homo_plus_set; auto. Qed. Lemma degen_all_homo_p : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), degen_p Bp -> homomorphism_p f. intros Ap Bp f h1. rewrite degen_p_iff in h1. pose proof (degen_all_homo (ba_conv_fun f) h1) as h2. rewrite homomorphism_p_iff. assumption. Qed. Lemma degen_all_homo_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), degen B -> homomorphism_p1 f. intros Ap B f h1. pose proof (degen_all_homo (ba_conv_fun1 f) h1) as h2. rewrite homomorphism_p1_iff. assumption. Qed. Lemma homo_homo2_p1 : forall (Ap:Bool_Alg_p T) (B:Bool_Alg), ~degen_p Ap -> homomorphism_p1 (homo2_p1 Ap B). intros Ap B hnd. unfold degen_p in hnd. unfold homo2_p1. apply homo_two_ops_plus_p1. intros x y. destruct x as [x h1], y as [y h2]. simpl. destruct (couple_inv_dec (proj1_sig %0) (proj1_sig %1)) as [h3 | h3]; destruct (couple_inv_dec (proj1_sig %0) (proj1_sig %1)) as [h4 | h4]; destruct (couple_inv_dec (proj1_sig %0) (proj1_sig %1)) as [h5 | h5]; subst. rewrite zero_sum. reflexivity. simpl in h3. pose proof (proof_irrelevance _ (proj2_sig _) (incl_two_p Ap (proj1_sig %1) h2)) as h4. pose proof (proof_irrelevance _ (proj2_sig _) (incl_two_p Ap (proj1_sig %0) h1)) as h4'. rewrite <- h4, <- h4' in h3. do 2 rewrite <- unfold_sig in h3. rewrite one_sum_p in h3. apply proj1_sig_injective in h3. contradict hnd. rewrite h3. reflexivity. pose proof (proof_irrelevance _ (proj2_sig _) (incl_two_p Ap (proj1_sig %1) h1)) as h4. pose proof (proof_irrelevance _ (proj2_sig _) (incl_two_p Ap (proj1_sig %0) h2)) as h4'. rewrite <- h4, <- h4' in h3. do 2 rewrite <- unfold_sig in h3. rewrite zero_sum_p in h3. apply proj1_sig_injective in h3. contradict hnd. rewrite h3. reflexivity. pose proof (proof_irrelevance _ (proj2_sig _) (incl_two_p Ap (proj1_sig %1) h1)) as h4. pose proof (proof_irrelevance _ (proj2_sig _) (incl_two_p Ap (proj1_sig %1) h2)) as h4'. rewrite <- h4, <- h4' in h3. rewrite <- unfold_sig in h3. rewrite one_sum_p in h3. apply proj1_sig_injective in h3. contradict hnd. rewrite h3. reflexivity. pose proof (proof_irrelevance _ (proj2_sig _) (incl_two_p Ap (proj1_sig %0) h1)) as h4. pose proof (proof_irrelevance _ (proj2_sig _) (incl_two_p Ap (proj1_sig %0) h2)) as h4'. rewrite <- h4, <- h4' in h3. rewrite <- unfold_sig in h3. rewrite zero_sum_p in h3. apply proj1_sig_injective in h3. contradict hnd. rewrite h3. reflexivity. rewrite one_sum. reflexivity. rewrite zero_sum. reflexivity. rewrite one_sum. reflexivity. intro x. destruct x as [x h1]. simpl. destruct (couple_inv_dec (proj1_sig %0) (proj1_sig %1)) as [h3 | h3]; destruct (couple_inv_dec (proj1_sig %0) (proj1_sig %1)) as [h4 | h4]; subst. pose proof (proof_irrelevance _ (proj2_sig _) (incl_two_p Ap (proj1_sig %0) h1)) as h4. rewrite <- h4 in h3. rewrite <- unfold_sig in h3. rewrite zero_comp_p in h3. apply proj1_sig_injective in h3. contradict hnd. rewrite h3. reflexivity. rewrite one_comp. reflexivity. rewrite zero_comp. reflexivity. pose proof (proof_irrelevance _ (proj2_sig _) (incl_two_p Ap (proj1_sig %1) h1)) as h4. rewrite <- h4 in h3. rewrite <- unfold_sig in h3. rewrite one_comp_p in h3. apply proj1_sig_injective in h3. contradict hnd. rewrite h3. reflexivity. Qed. Lemma inj_homo2_p1 : forall (Ap:Bool_Alg_p T) (B:Bool_Alg), ~degen_p Ap -> ~degen B -> Injective (homo2_p1 Ap B). intros Ap B hnda hndb; red; intros x y h1. destruct (two_bap_dec x) as [h2 | h2], (two_bap_dec y) as [h3 | h3]; subst; auto. contradict hndb. rewrite homo_zero_p1, homo_one_p1 in h1. red; auto. apply homo_homo2_p1; auto. apply homo_homo2_p1; auto. contradict hndb. red. rewrite homo_zero_p1, homo_one_p1 in h1; auto. apply homo_homo2_p1; auto. apply homo_homo2_p1; auto. Qed. Lemma mono_homo2_p1 : forall (Ap:Bool_Alg_p T) (B:Bool_Alg), ~degen_p Ap -> ~degen B -> monomorphism_p1 (homo2_p1 Ap B). intros; constructor. apply homo_homo2_p1; auto. apply inj_homo2_p1; auto. Qed. Lemma unq_homo2_p1 : forall (Ap:Bool_Alg_p T) (B:Bool_Alg) (g:btp (two_bap Ap)->bt B), ~degen_p Ap -> homomorphism_p1 g -> g = homo2_p1 Ap B. intros; apply functional_extensionality. intro x. destruct (two_bap_dec x); subst. rewrite homo_zero_p1. rewrite homo_zero_p1; auto. apply homo_homo2_p1; auto. assumption. rewrite homo_one_p1, homo_one_p1; auto. apply homo_homo2_p1; auto. Qed. Lemma iso_iso2_p1 : forall (Ap:Bool_Alg_p T) (B:Bool_Alg), ~degen_p Ap -> ~degen B -> isomorphism_p1 (iso2_p1 Ap B). intros Ap B h1 h2. unfold iso2_p1. constructor. apply homo_homo2_p1; auto. red. split. red. intros x y h3. unfold homo2_p1 in h3. destruct x as [x h4], y as [y h5]. apply proj1_sig_injective. simpl. destruct (couple_inv_dec (proj1_sig %0) (proj1_sig %1) x h4) as [h6 | h6], (couple_inv_dec (proj1_sig %0) (proj1_sig %1) y h5) as [h7 | h7]. congruence. contradict h2. red. simpl in h3. pose proof (f_equal (@proj1_sig _ _) h3) as h8. simpl in h8. assumption. contradict h2. simpl in h3. red. pose proof (f_equal (@proj1_sig _ _) h3) as h8. simpl in h8. rewrite h8. reflexivity. congruence. unfold homo2_p1. red. intro y. destruct y as [y h3]; simpl. destruct h3; simpl. exists %0. simpl. destruct ( couple_inv_dec (proj1_sig %0) (proj1_sig %1) (proj1_sig %0) (Z_c_p T Ap (two_p Ap) (incl_two_p Ap) (alg_closed_two_p Ap))) as [h3 | h3]. apply proj1_sig_injective; auto. contradict h1. red. apply proj1_sig_injective. assumption. exists %1. simpl. destruct (couple_inv_dec (proj1_sig %0) (proj1_sig %1) (proj1_sig %1) (O_c_p T Ap (two_p Ap) (incl_two_p Ap) (alg_closed_two_p Ap))) as [h3 | h3]. contradict h1. red. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. Qed. Definition iso_inv_p1 {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B) (pf:isomorphism_p1 f) : bt B -> btp Ap. destruct pf as [h1 h2]. apply bijective_impl_invertible in h2. refine (proj1_sig (function_inverse _ h2)). Defined. Lemma iso_inv_compat_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B) (pf:isomorphism_p1 f), let f' := iso_inv_p1 f pf in isomorphism_p2 f' /\ (forall x, f' (f x) = x) /\ (forall x, f (f' x) = x). intros A B f h1. destruct h1 as [h1 h2]. unfold iso_inv_p1. split. split. unfold iso_inv. apply homo_inv_p1; auto. apply invertible_impl_bijective. apply invertible_impl_inv_invertible. unfold function_inverse. destruct constructive_definite_description. simpl. assumption. Qed. Definition iso_inv_p2 {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) (pf:isomorphism_p2 f) : btp Bp -> bt A. destruct pf as [h1 h2]. apply bijective_impl_invertible in h2. refine (proj1_sig (function_inverse _ h2)). Defined. Lemma iso_inv_compat_p2 : forall {A:Bool_Alg} {Bp:Bool_Alg_p T} (f:bt A->btp Bp) (pf:isomorphism_p2 f), let f' := iso_inv_p2 f pf in isomorphism_p1 f' /\ (forall x, f' (f x) = x) /\ (forall x, f (f' x) = x). intros A B f h1. destruct h1 as [h1 h2]. unfold iso_inv_p2. split. split. unfold iso_inv_p2. apply homo_inv_p2; auto. apply invertible_impl_bijective. apply invertible_impl_inv_invertible. unfold function_inverse. destruct constructive_definite_description. simpl. assumption. Qed. Definition trivial_ba_p_mono (Bp:Bool_Alg_p T) (A:Bool_Alg) (pf:ba_p_ens Bp = im_proj1_sig (Couple (Bzero_p T (Bc_p T Bp)) (Bone_p T (Bc_p T Bp)))) : btp Bp->bt A. intro x. destruct (trivial_ba_p_dec Bp pf x). refine 0. refine 1. Defined. Lemma mono_trivial_ba_p_mono : forall (Bp:Bool_Alg_p T) (A:Bool_Alg) (pf:ba_p_ens Bp = im_proj1_sig (Couple %0 %1)), ~degen_p Bp -> ~degen A -> monomorphism_p1 (trivial_ba_p_mono Bp A pf). intros Bp A h1 h0 h0'. constructor. apply homo_two_ops_plus_p1. intros x y. unfold trivial_ba_p_mono. destruct (trivial_ba_p_dec _ h1 (x %+ y)) as [h2 | h2]; destruct (trivial_ba_p_dec Bp h1 x) as [h3 | h3]; destruct (trivial_ba_p_dec Bp h1 y) as [h4 | h4]; subst; try rewrite zero_sum; try rewrite one_sum; auto. rewrite one_sum_p in h2. contradict h0. red. rewrite h2; auto. rewrite zero_sum_p in h2. contradict h0. red. rewrite h2; auto. rewrite one_sum_p in h2. contradict h0. red. rewrite h2; auto. rewrite zero_sum_p in h2. contradict h0. red. assumption. intro x. unfold trivial_ba_p_mono. destruct (trivial_ba_p_dec Bp h1 (%-x)) as [h2 | h2]; destruct (trivial_ba_p_dec Bp h1 x) as [h3 | h3]; try rewrite one_comp; try rewrite zero_comp; subst; auto. rewrite zero_comp_p in h2. contradict h0. red. rewrite h2; auto. rewrite one_comp_p in h2. contradict h0. red. rewrite h2; auto. red. intros x y h2. unfold trivial_ba_p_mono in h2. destruct (trivial_ba_p_dec Bp h1 x) as [h3 | h3]; destruct (trivial_ba_p_dec Bp h1 y) as [h4 | h4]; subst; (reflexivity || (contradict h0'; red)); auto. Qed. Lemma unq_trivial_ba_p_mono : forall (Bp : Bool_Alg_p T) (A : Bool_Alg), ~ degen_p Bp -> ~ degen A -> forall (pf:ba_p_ens Bp = im_proj1_sig (Couple %0 %1)) (f:btp Bp -> bt A), monomorphism_p1 f -> trivial_ba_p_mono Bp A pf = f. intros Bp A h1 h2 h3 f h4. apply functional_extensionality. intro x. unfold trivial_ba_p_mono. destruct h4 as [h4 h5]. destruct (trivial_ba_p_dec _ h3 x) as [h6 | h6]; subst; try rewrite homo_zero_p1; try rewrite homo_one_p1; auto. Qed. End ParametricAnalogues. Arguments homomorphism_p [T] [Ap] [Bp] _. Arguments homomorphism_p_iff [T] [Ap] [Bp] _. Arguments homomorphism_general_p [T] [Ap] [A] _ _ _ _. Arguments homomorphism_general_p_iff [T] [Ap] [A] _ _ _ _. Arguments homo_sym_diff_p [T] [Ap] [Bp] _ _ _ _. Arguments homo_id_p [T] _. Arguments universal_homomorphic_ops_imply_bool_p [T] [Ap] [A] _ _ _ _ _ _. Arguments homo_zero_p [T] [Ap] [Bp] _ _. Arguments homo_one_p [T] [Ap] [Bp] _ _. Arguments homo_two_ops_plus_p [T] [Ap] [Bp] _ _ _. Arguments homo_two_ops_times_p [T] [Ap] [Bp] _ _ _. Arguments ba_conv_fun_compose [T] [Ap] [Bp] _ _ _. Arguments homo_compose_p [T] [Ap] [Bp] _ _ _ _ _. Arguments ba_conv_fun_inv_iff [T] [Ap] [Bp] _. Arguments ba_conv_fun_function_inverse [T] [Ap] [Bp] _ _. Arguments homo_inv_p [T] [Ap] [Bp] _ _ _. Arguments homo_mono_p [T] [Ap] [Bp] _ _ _ _ _. Arguments im_ba_conv_fun_eq [T] [Ap] [Bp] _ _. Arguments homo_im_closed_p [T] [Ap] [Bp] _ _. Arguments homo_subalg_im_p_ex [T] [Ap] [Bp] _ _. Arguments homo_subalg_im_p [T] [Ap] [Bp] _ _. Arguments homo_subalg_im_p_compat [T] [Ap] [Bp] _ _. Arguments transfer_homo_subalg_im_p_ex [T] [Ap] [Bp] _ _ _. Arguments transfer_homo_subalg_im_p [T] [Ap] [Bp] _ _ _. Arguments transfer_homo_subalg_im_p_compat [T] [Ap] [Bp] _ _ _. Arguments homo_onto_p [T] [Ap] [Bp] _ _ _. Arguments isomorphism_p1 [T] [Ap] [B] _. Arguments isomorphism_p2 [T] [A] [Bp] _. Arguments iso_mono_p1 [T] [Ap] [B] _ _. Arguments iso_mono_p2 [T] [A] [Bp] _ _. Arguments iso_homo_p1 [T] [Ap] [B] _ _. Arguments iso_homo_p2 [T] [A] [Bp] _ _. Arguments isomorphism_p1_iff [T] [Ap] [B] _. Arguments isomorphism_p2_iff [T] [A] [Bp] _. Arguments homo_homo2_p1 [T] [Ap] [B] _. Arguments iso_iso2_p1 [T] [Ap] [B] _ _. Arguments iso_inv_p1 [T] [Ap] [B] _ _ _. Arguments iso_inv_compat_p1 [T] [Ap] [B] _ _. Arguments iso_inv_p2 [T] [A] [Bp] _ _ _. Arguments iso_inv_compat_p2 [T] [A] [Bp] _ _. Arguments homo2_p1 [T] _ _ _. Arguments iso2_p1 [T] _ _ _. bool2-v0-3/ExtensionsOfHomomorphisms.v0000644000175000017500000130010513575574055020742 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2015-2018 Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import SetUtilities. Require Import SetUtilities1_5. Require Import BoolAlgBasics. Require Import Homomorphisms. Require Import Subalgebras. Require Import TypeUtilities. Require Import LogicUtilities. Require Import Description. Require Import FunctionalExtensionality. Require Import FiniteMaps. Require Import FiniteMaps2. Require Import FiniteOperations. Require Import ProofIrrelevance. Require Import DecidableDec. Require Import FunctionProperties. Require Import InfiniteOperations. Require Import WellOrders. Require Import ArithUtilities. Require Import SetUtilities2. Require Import ListUtilities. Require Import EqDec. Require Import LogicUtilities. Section Extensions. Variable T:Type. Inductive extends_homo {Ap Bp Cp:Bool_Alg_p T} (f:btp Bp->btp Cp) (g:btp Ap->btp Cp) : Prop := extends_homo_intro : subalg_of_p Ap Bp -> homomorphism_p f -> homomorphism_p g -> extends_sig f g -> extends_homo f g. Lemma refl_extends_homo : forall {Ap Bp:Bool_Alg_p T} (f:btp Ap->btp Bp), homomorphism_p f -> extends_homo f f. intros Ap Bp f h1. constructor; auto. apply refl_subalg_of_p. apply refl_extends_sig. Qed. Inductive extends_homo_p1 {Ap Bp:Bool_Alg_p T} {C:Bool_Alg} (f:btp Bp->bt C) (g:btp Ap->bt C) : Prop := extends_homo_p1_intro : subalg_of_p Ap Bp -> homomorphism_p1 _ f -> homomorphism_p1 _ g -> extends_sig f g -> extends_homo_p1 f g. Lemma refl_extends_homo_p1 : forall {Ap:Bool_Alg_p T} {B:Bool_Alg} (f:btp Ap->bt B), homomorphism_p1 _ f -> extends_homo_p1 f f. intros Ap Bp f h1. constructor; auto. apply refl_subalg_of_p. apply refl_extends_sig. Qed. Lemma extends_homo_p1_iff : forall {Ap Bp Cp:Bool_Alg_p T} (f:btp Bp->btp Cp) (g:btp Ap->btp Cp), extends_homo f g <-> extends_homo_p1 (ba_conv_fun2 f) (ba_conv_fun2 g). intros Ap Bp Cp f g. split. intro h1. destruct h1 as [h1 h2 h3 h4]. rewrite homomorphism_p_iff in h2, h3. constructor; auto. rewrite homomorphism_p1_iff. assumption. rewrite homomorphism_p1_iff. assumption. intro h1. destruct h1 as [h1 h2 h3 h4]. rewrite homomorphism_p1_iff in h2, h3. constructor; auto. rewrite homomorphism_p_iff. assumption. rewrite homomorphism_p_iff. assumption. Qed. Definition extends_homo2 {Ap Ap' Bp Cp:Bool_Alg_p T} (f:btp Bp->btp Cp) (g:btp Ap->btp Cp) (g':btp Ap'->btp Cp): Prop := extends_homo f g /\ extends_homo f g'. Definition extends_homo2_p1 {Ap Ap' Bp:Bool_Alg_p T} {C:Bool_Alg} (f:btp Bp->bt C) (g:btp Ap->bt C) (g':btp Ap'->bt C): Prop := extends_homo_p1 f g /\ extends_homo_p1 f g'. Definition extends_homo2_p1_iff : forall {Ap Ap' Bp Cp:Bool_Alg_p T} (f:btp Bp->btp Cp) (g:btp Ap->btp Cp) (g':btp Ap'->btp Cp), extends_homo2 f g g' <-> extends_homo2_p1 (ba_conv_fun2 f) (ba_conv_fun2 g) (ba_conv_fun2 g'). intros Ap Ap' Bp Cp f g g'. split. intro h1. destruct h1 as [h1 h2]. rewrite extends_homo_p1_iff in h1. rewrite extends_homo_p1_iff in h2. constructor; auto. intro h1. destruct h1 as [h1 h2]. rewrite <- extends_homo_p1_iff in h1. rewrite <- extends_homo_p1_iff in h2. constructor; auto. Qed. Lemma extends_homo_iff : forall {Ap Bp Cp:Bool_Alg_p T} (f:btp Bp->btp Cp) (g:btp Ap->btp Cp), extends_homo f g <-> homomorphism_p f /\ homomorphism_p g /\ exists pf:subalg_of_p Ap Bp, restriction_sig f _ (subalg_of_p_incl _ _ pf) = g. intros Ap Bp Cp f g. split. intro h1. destruct h1 as [h1 h2 h3 h4]. do 2 (split; auto). exists h1. rewrite extends_sig_iff in h4. destruct h4 as [h5 h6]. rewrite <- h6. f_equal. apply proof_irrelevance. intro h1. destruct h1 as [h1 [h2 h3]]. destruct h3 as [h3 h4]. constructor; auto. red in h3. destruct h3 as [h3a [h3b h3c]]. rewrite extends_sig_iff. exists h3a. rewrite <- h4. f_equal. apply proof_irrelevance. Qed. Lemma extends_homo_iff_p1 : forall {Ap Bp:Bool_Alg_p T} {C:Bool_Alg} (f:btp Bp->bt C) (g:btp Ap->bt C), extends_homo_p1 f g <-> homomorphism_p1 _ f /\ homomorphism_p1 _ g /\ exists pf:subalg_of_p Ap Bp, restriction_sig f _ (subalg_of_p_incl _ _ pf) = g. intros Ap Bp C f g. split. intro h1. destruct h1 as [h1 h2 h3 h4]. do 2 (split; auto). exists h1. rewrite extends_sig_iff in h4. destruct h4 as [h4 h5]. rewrite <- h5. f_equal. apply proof_irrelevance. intro h1. destruct h1 as [h1 [h2 h3]]. destruct h3 as [h3 h4]. constructor; auto. rewrite extends_sig_iff. destruct h3 as [h3a [h3b h3c]]. exists h3a. rewrite <- h4. f_equal. apply proof_irrelevance. Qed. Definition fam_fun_ba_one_range (Bp:Bool_Alg_p T) := Ensemble {Ap:(Bool_Alg_p T) & (btp Ap)->(btp Bp)}. Definition fam_fun_ba_one_range_p1 (B:Bool_Alg) := Ensemble {Ap:(Bool_Alg_p T) & (btp Ap)->(bt B)}. Definition ba_conv_fam_fun_ba_one_range {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp) : fam_fun_ba_one_range_p1 (ba_conv Bp) := Im F (fun p=>existT (fun S=>btp S->bt (ba_conv Bp)) (projT1 p) (ba_conv_fun2 (projT2 p))). Lemma in_ba_conv_fam_fun_ba_one_range_iff : forall {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp) (pr:{A:Bool_Alg_p T & (btp A)->(btp Bp)}), Inn F pr <-> Inn (ba_conv_fam_fun_ba_one_range F) (existT (fun S=>btp S->bt (ba_conv Bp)) (projT1 pr) (ba_conv_fun2 (projT2 pr))). intros Bp F pr. split. intro h1. apply Im_intro with pr; auto. intro h2. inversion h2 as [pr' h3 pr'' h4]. subst. clear h2. destruct pr as [A f], pr' as [A' f']. simpl in h4, h3. inversion h4. subst. clear H1. apply inj_pair2 in h4. rewrite <- h4 in h3 at 1. assumption. Qed. Definition fam_homo {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp) : Prop := forall f, Inn F f -> homomorphism_p (projT2 f). Definition fam_homo_p1 {B:Bool_Alg} (F:fam_fun_ba_one_range_p1 B) : Prop := forall f, Inn F f -> homomorphism_p1 _ (projT2 f). Lemma fam_homo_p1_iff : forall {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp), fam_homo F <-> fam_homo_p1 (ba_conv_fam_fun_ba_one_range F). intros Bp F. split. intro h1. red in h1. red. intros pr h2. destruct h2 as [pr h2]. subst. simpl. specialize (h1 pr h2). rewrite homomorphism_p1_iff. rewrite homomorphism_p_iff in h1. assumption. intro h1. red in h1. red. intros pr h2. rewrite in_ba_conv_fam_fun_ba_one_range_iff in h2. specialize (h1 _ h2). simpl in h1. rewrite homomorphism_p_iff. rewrite homomorphism_p1_iff in h1. assumption. Qed. Definition common_extension_fam_homo {Ap Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp) (f:btp Ap->btp Bp) : Prop := fam_homo F /\ forall g, Inn F g -> extends_homo f (projT2 g). Definition common_extension_fam_homo_p1 {Ap:Bool_Alg_p T} {B:Bool_Alg} (F:fam_fun_ba_one_range_p1 B) (f:btp Ap->bt B) : Prop := fam_homo_p1 F /\ forall g, Inn F g -> extends_homo_p1 f (projT2 g). Lemma common_extension_fam_homo_p1_iff : forall {Ap Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp) (f:btp Ap->btp Bp), common_extension_fam_homo F f <-> common_extension_fam_homo_p1 (ba_conv_fam_fun_ba_one_range F) (ba_conv_fun2 f). intros Ap Bp F f. split. intro h1. red in h1. red. destruct h1 as [h1 h2]. rewrite fam_homo_p1_iff in h1. split; auto. intros g h3. destruct g as [A' k]. inversion h3 as [pr h4 ? h5]. subst. inversion h5. subst. clear H1. rewrite h5 in h3. simpl. destruct pr as [A' k']. simpl in h3, h4, h5. apply inj_pair2 in h5. subst. specialize (h2 _ h4). rewrite <- extends_homo_p1_iff. simpl in h2. assumption. intro h1. red in h1. red. destruct h1 as [h1 h2]. rewrite <- fam_homo_p1_iff in h1. split; auto. intros g h3. destruct g as [A' k]. rewrite in_ba_conv_fam_fun_ba_one_range_iff in h3. specialize (h2 _ h3). simpl in h2, h3. rewrite extends_homo_p1_iff. assumption. Qed. Lemma common_extension_fam_homo_p1_homo_p1 : forall (Bp:Bool_Alg_p T) (A:Bool_Alg) (f:btp Bp->bt A) (fam:fam_fun_ba_one_range_p1 A), Inhabited fam -> common_extension_fam_homo_p1 fam f -> homomorphism_p1 _ f. intros Bp A f fam h1 h2. destruct h1 as [pr h1b]. red in h2. destruct h2 as [h2 h3]. specialize (h3 pr h1b). destruct h3. assumption. Qed. Definition directed_fam_homo {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp) : Prop := fam_homo F /\ forall f g, Inn F f -> Inn F g -> exists k, Inn F k /\ extends_homo2 (projT2 k) (projT2 f) (projT2 g). Definition directed_fam_homo_p1 {B:Bool_Alg} (F:fam_fun_ba_one_range_p1 B) : Prop := fam_homo_p1 F /\ forall f g, Inn F f -> Inn F g -> exists k, Inn F k /\ extends_homo2_p1 (projT2 k) (projT2 f) (projT2 g). Lemma directed_fam_homo_p1_iff : forall {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp), directed_fam_homo F <-> directed_fam_homo_p1 (ba_conv_fam_fun_ba_one_range F). intros Bp F. split. intro h1. red. red in h1. destruct h1 as [h1 h2]. rewrite fam_homo_p1_iff in h1. split; auto. intros f g h3 h4. destruct f as [A' f], g as [A'' f']. rewrite <- (in_ba_conv_fam_fun_ba_one_range_iff F (existT _ A' (ba_conv_fun1 f))) in h3. rewrite <- (in_ba_conv_fam_fun_ba_one_range_iff F (existT _ A'' (ba_conv_fun1 f'))) in h4. specialize (h2 _ _ h3 h4). simpl in h2. destruct h2 as [k h2]. exists k. simpl. destruct h2 as [h2a h2b]. split; auto. rewrite in_ba_conv_fam_fun_ba_one_range_iff in h2a. destruct k as [C k]. simpl in h2a. assumption. rewrite extends_homo2_p1_iff in h2b. assumption. intro h2. red in h2. red. destruct h2 as [h2 h3]. split. rewrite fam_homo_p1_iff. assumption. intros f g h4 h5. rewrite in_ba_conv_fam_fun_ba_one_range_iff in h4, h5. specialize (h3 _ _ h4 h5). simpl in h3. destruct h3 as [k h3]. exists k. destruct h3 as [h3a h3b]. split. rewrite in_ba_conv_fam_fun_ba_one_range_iff. destruct k as [A k]. simpl. assumption. rewrite extends_homo2_p1_iff. assumption. Qed. Definition inj_fam_homo {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp) : Prop := fam_homo F /\ forall f, Inn F f -> Injective (projT2 f). Definition inj_fam_homo_p1 {B:Bool_Alg} (F:fam_fun_ba_one_range_p1 B) : Prop := fam_homo_p1 F /\ forall f, Inn F f -> Injective (projT2 f). Lemma inj_fam_homo_p1_iff : forall {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp), inj_fam_homo F <-> inj_fam_homo_p1 (ba_conv_fam_fun_ba_one_range F). intros Bp F. split. intro h1. red. red in h1. destruct h1 as [h1 h2]. split. rewrite fam_homo_p1_iff in h1. assumption. intros f h3. destruct f as [A' f]. rewrite <- (in_ba_conv_fam_fun_ba_one_range_iff F (existT _ _ (ba_conv_fun2 f))) in h3. specialize (h2 _ h3). simpl in h2. simpl. assumption. intro h1. red in h1. red. destruct h1 as [h1 h2]. split. rewrite fam_homo_p1_iff. assumption. intros f h3. rewrite in_ba_conv_fam_fun_ba_one_range_iff in h3. specialize (h2 _ h3). simpl in h2. assumption. Qed. Definition fam_fun_ba_domains {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp) := Im F (@projT1 _ _). Definition fam_fun_ba_domains_p1 {B:Bool_Alg} (F:fam_fun_ba_one_range_p1 B) := Im F (@projT1 _ _). Lemma fam_fun_ba_domains_p1_eq : forall {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp), fam_fun_ba_domains F = fam_fun_ba_domains_p1 (ba_conv_fam_fun_ba_one_range F). intros Bp F. unfold fam_fun_ba_domains, fam_fun_ba_domains_p1, ba_conv_fam_fun_ba_one_range. rewrite im_im. simpl. reflexivity. Qed. Lemma directed_fam_homo_directed_domains : forall {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp), directed_fam_homo F -> directed (fam_fun_ba_domains F). intros Bp F h1. red in h1. destruct h1 as [h1 h2]. red. intros Ap Ap' h3 h4. destruct h3 as [fa h3]. destruct h4 as [fb h4]. subst. specialize (h2 _ _ h3 h4). destruct h2 as [k [h2a [h2b h2c]]]. destruct k as [Cp fc]. exists Cp. split. unfold fam_fun_ba_domains. apply Im_intro with (existT (fun Ap : Bool_Alg_p T => btp Ap -> btp Bp) Cp fc). assumption. simpl. reflexivity. destruct h2b as [h5]. simpl in h5. destruct h2c as [h6]. simpl in h6. split; auto. Qed. Lemma directed_fam_homo_directed_domains_p1 : forall {B:Bool_Alg} (F:fam_fun_ba_one_range_p1 B), directed_fam_homo_p1 F -> directed (fam_fun_ba_domains_p1 F). intros B F h1. red in h1. destruct h1 as [h1 h2]. red. intros Ap Ap' h3 h4. destruct h3 as [fa h3]. destruct h4 as [fb h4]. subst. specialize (h2 _ _ h3 h4). destruct h2 as [k [h2a [h2b h2c]]]. destruct k as [Cp fc]. exists Cp. split. unfold fam_fun_ba_domains_p1. apply Im_intro with (existT (fun Ap : Bool_Alg_p T => btp Ap -> bt B) Cp fc). assumption. simpl. reflexivity. destruct h2b as [h5]. simpl in h5. destruct h2c as [h6]. simpl in h6. split; auto. Qed. Lemma directed_common_extension_p1 : forall {B:Bool_Alg} (F:fam_fun_ba_one_range_p1 B), directed_fam_homo_p1 F -> Inhabited F -> exists (pfd: directed (fam_fun_ba_domains_p1 F)) (pfi:Inhabited (fam_fun_ba_domains_p1 F)), let Ap:=directed_ba_p pfd pfi in exists (f:btp Ap->bt B), common_extension_fam_homo_p1 F f /\ (inj_fam_homo_p1 F -> Injective f). intros B F h1 h2. pose proof (directed_fam_homo_directed_domains_p1 F h1) as h3. exists h3. assert (h4:Inhabited (fam_fun_ba_domains_p1 F)). destruct h2 as [f h4]. apply Inhabited_intro with (projT1 f). unfold fam_fun_ba_domains_p1. apply Im_intro with f; auto. exists h4. intro Ap. red in h1. destruct h1 as [h1a h1b]. red in h1a. assert (h5:forall x:btp Ap, exists! y:bt B, forall g, Inn F g -> forall (pf:Inn (ba_p_ens (projT1 g)) (proj1_sig x)), (projT2 g) (exist _ _ pf) = y). intro x. destruct x as [x h5]. simpl in h5. destruct h5 as [A x h5 h6]. destruct h5 as [A h5]. subst. destruct h5 as [f h5]. subst. simpl. exists ((projT2 f) (exist _ _ h6)). red. split. intros g h7 h8. specialize (h1b _ _ h5 h7). destruct h1b as [k [h1c h1d]]. red in h1d. destruct h1d as [h1d h1e]. destruct h1d as [h9 h10 h11 h12], h1e as [h13 h14 h15 h16]. red in h12. destruct h12 as [h12a h12b], h16 as [h16a h16b]. specialize (h12b (exist (Inn (ba_p_ens (projT1 f))) x h6)). specialize (h16b (exist (Inn (ba_p_ens (projT1 g))) x h8)). rewrite h12b, h16b. f_equal. apply proj1_sig_injective. simpl. reflexivity. intros x' h7. specialize (h7 _ h5 h6). assumption. exists (fun x => (proj1_sig (constructive_definite_description _ (h5 x)))). split. red. split. intros g h6. apply h1a; auto. intros g h6. constructor. apply directed_ba_p_subalg. unfold fam_fun_ba_domains_p1. apply Im_intro with g; auto. apply homo_two_ops_plus_p1. intros x y. destruct constructive_definite_description as [z h7]. simpl. destruct constructive_definite_description as [a h8]. simpl. destruct constructive_definite_description as [b h9]. simpl. destruct x as [x h10], y as [y h11]. simpl in h7, h8, h9. unfold Ap in h10, h11. simpl in h10, h11. destruct h10 as [Xp x h10 h12], h11 as [Yp y h13 h14]. destruct h10 as [Xp h10], h13 as [Yp h13]. subst. destruct h10 as [f h10], h13 as [f' h13]. subst. specialize (h8 _ h10 h12). specialize (h9 _ h13 h14). subst. simpl in h7. specialize (h1b _ _ h10 h13). destruct h1b as [f'' [h1b h1c]]. destruct h1c as [h15 h16]. destruct h15 as [h15a h15b h15c h15d], h16 as [h16a h16b h16c h16d]. red in h15d, h16d. destruct h15d as [h15d h15e], h16d as [h16d h16e]. specialize (h15e (exist _ _ h12)). specialize (h16e (exist _ _ h14)). rewrite h15e, h16e. specialize (h7 _ h1b). simpl in h7. pose proof (directed_plus_compat _ h3 (exist (fun x0 : T => Inn (FamilyUnion (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)))) x0) x (family_union_intro T (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T))) (ba_p_ens (projT1 f)) x (Im_intro (Bool_Alg_p T) (Ensemble T) (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)) (projT1 f) (Im_intro {Ap0 : Bool_Alg_p T & btp Ap0 -> bt B} (Bool_Alg_p T) F (projT1 (P:=fun Ap0 : Bool_Alg_p T => btp Ap0 -> bt B)) f h10 (projT1 f) eq_refl) (ba_p_ens (projT1 f)) eq_refl) h12)) (exist (fun x0 : T => Inn (FamilyUnion (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)))) x0) y (family_union_intro T (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T))) (ba_p_ens (projT1 f')) y (Im_intro (Bool_Alg_p T) (Ensemble T) (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)) (projT1 f') (Im_intro {Ap0 : Bool_Alg_p T & btp Ap0 -> bt B} (Bool_Alg_p T) F (projT1 (P:=fun Ap0 : Bool_Alg_p T => btp Ap0 -> bt B)) f' h13 (projT1 f') eq_refl) (ba_p_ens (projT1 f')) eq_refl) h14))) as h15. simpl in h15. specialize (h15 _ _ h12 h14). assert (h16: Inn (fam_fun_ba_domains_p1 F) (projT1 f)). unfold fam_fun_ba_domains_p1. apply Im_intro with f; auto. assert (h17:Inn (fam_fun_ba_domains_p1 F) (projT1 f')). apply Im_intro with f'; auto. specialize (h15 h16 h17). destruct h15 as [Cp [h18 [h19 [h20 [h21 [h22 [h23 [h24 [h25 h27]]]]]]]]]. assert (h29: (proj1_sig (exist (Inn (ba_p_ens (projT1 f''))) x (h15d x h12) %+ exist (Inn (ba_p_ens (projT1 f''))) y (h16d y h14))) = (proj1_sig (exist (Inn (ba_p_ens Cp)) x (h18 x h12) %+ exist (Inn (ba_p_ens Cp)) y (h19 y h14)))). assert (h30:Inn (fam_fun_ba_domains_p1 F) (projT1 f'')). apply Im_intro with f''; auto. pose proof (h3 _ _ h30 h25) as h31. destruct h31 as [Dp [h31 [h32 h33]]]. red in h32, h33. destruct h32 as [h32a [h32b h32c]], h33 as [h33a [h33b h33c]]. pose proof (ba_p_subst_plus _ _ h32c _ _ (h15d x h12) (h16d y h14)) as h34. assert (h35:Inn (ba_p_ens (Subalg_p Dp (ba_p_ens (projT1 f'')) h32a h32b)) x). rewrite <- h32c. apply (h15d x h12). assert (h36 : Inn (ba_p_ens (Subalg_p Dp (ba_p_ens (projT1 f'')) h32a h32b)) y). rewrite <- h32c. apply (h16d y h14). specialize (h34 h35 h36). pose proof (ba_p_subst_plus _ _ h33c _ _ (h18 x h12) (h19 y h14)) as h37. assert (h38: Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Cp) h33a h33b)) x). rewrite <- h33c. apply (h18 x h12). assert (h39 : Inn (ba_p_ens (Subalg_p Dp (ba_p_ens Cp) h33a h33b)) y). rewrite <- h33c. apply (h19 y h14). specialize (h37 h38 h39). rewrite h34, h37. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. assert (h0: Inn (ba_p_ens (projT1 f'')) (proj1_sig (directed_plus T h3 (exist (fun x0 : T => Inn (FamilyUnion (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)))) x0) x (family_union_intro T (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T))) (ba_p_ens (projT1 f)) x (Im_intro (Bool_Alg_p T) (Ensemble T) (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)) (projT1 f) (Im_intro {Ap0 : Bool_Alg_p T & btp Ap0 -> bt B} (Bool_Alg_p T) F (projT1 (P:=fun Ap0 : Bool_Alg_p T => btp Ap0 -> bt B)) f h10 (projT1 f) eq_refl) (ba_p_ens (projT1 f)) eq_refl) h12)) (exist (fun x0 : T => Inn (FamilyUnion (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)))) x0) y (family_union_intro T (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T))) (ba_p_ens (projT1 f')) y (Im_intro (Bool_Alg_p T) (Ensemble T) (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)) (projT1 f') (Im_intro {Ap0 : Bool_Alg_p T & btp Ap0 -> bt B} (Bool_Alg_p T) F (projT1 (P:=fun Ap0 : Bool_Alg_p T => btp Ap0 -> bt B)) f' h13 (projT1 f') eq_refl) (ba_p_ens (projT1 f')) eq_refl) h14))))). rewrite h27 at 1. pose (exist _ _ (h15d _ h12) %+ exist _ _ (h16d _ h14)) as z'. pose proof (in_ba_p_ens_plus _ (projT1 f'') _ _ (h15d x h12) (h16d y h14)) as h28. rewrite h29 in h28. assumption. specialize (h7 h0). rewrite <- h7. rewrite <- homo_plus_p1. f_equal. apply proj1_sig_injective. simpl. rewrite h27. rewrite h29 at 1. reflexivity. assumption. intro x. destruct constructive_definite_description as [a h7]. destruct constructive_definite_description as [b h8]. simpl. destruct x as [x h9]. simpl in h7, h8. unfold Ap in h9. destruct h9 as [X x h9 h10]. destruct h9 as [Xp h9]. subst. destruct h9 as [f h9]. subst. specialize (h8 _ h9 h10). rewrite <- h8 at 1. pose proof (directed_comp_compat _ h3 (exist (fun x0 : T => Inn (FamilyUnion (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)))) x0) x (family_union_intro T (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T))) (ba_p_ens (projT1 f)) x (Im_intro (Bool_Alg_p T) (Ensemble T) (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)) (projT1 f) (Im_intro {Ap0 : Bool_Alg_p T & btp Ap0 -> bt B} (Bool_Alg_p T) F (projT1 (P:=fun Ap0 : Bool_Alg_p T => btp Ap0 -> bt B)) f h9 (projT1 f) eq_refl) (ba_p_ens (projT1 f)) eq_refl) h10))) as h11. simpl in h11. specialize (h11 _ h10). assert (h12:Inn (fam_fun_ba_domains_p1 F) (projT1 f)). unfold fam_fun_ba_domains_p1. apply Im_intro with f; auto. specialize (h11 h12). destruct h11 as [Cp [h18 [h19 [h20 h21]]]]. specialize (h7 _ h9). assert (h22:Inn (ba_p_ens (projT1 f)) (proj1_sig (directed_comp T h3 (exist (fun x0 : T => Inn (FamilyUnion (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)))) x0) x (family_union_intro T (Im (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T))) (ba_p_ens (projT1 f)) x (Im_intro (Bool_Alg_p T) (Ensemble T) (fam_fun_ba_domains_p1 F) (ba_p_ens (T:=T)) (projT1 f) (Im_intro {Ap0 : Bool_Alg_p T & btp Ap0 -> bt B} (Bool_Alg_p T) F (projT1 (P:=fun Ap0 : Bool_Alg_p T => btp Ap0 -> bt B)) f h9 (projT1 f) eq_refl) (ba_p_ens (projT1 f)) eq_refl) h10))))). rewrite h21. apply in_ba_p_ens_comp. specialize (h7 h22). rewrite <- h7 at 1. rewrite <- homo_comp_p1. f_equal. apply proj1_sig_injective. simpl. rewrite h21. reflexivity. apply h1a; auto. apply h1a; auto. red. assert (h7: Included (A_p T (Bc_p T (projT1 g))) (A_p T (Bc_p T Ap))). red. intros x h7. simpl. apply family_union_intro with (ba_p_ens (projT1 g)). apply Im_intro with (projT1 g). unfold fam_fun_ba_domains. apply Im_intro with g; auto. reflexivity. assumption. exists h7. intro x. destruct constructive_definite_description as [a h8]. simpl. specialize (h8 _ h6). simpl in h8. specialize (h8 (proj2_sig _)). rewrite <- h8. f_equal. apply proj1_sig_injective. simpl. reflexivity. intro h6. red. intros x y. destruct constructive_definite_description as [a h7]. destruct constructive_definite_description as [b h8]. simpl. intro h9. subst. destruct x as [x h9], y as [y h10]. simpl in h9, h10. destruct h9 as [X x h9 h11], h10 as [Y y h10 h12]. destruct h9 as [X h9], h10 as [Y h10]. subst. destruct h9 as [f h9], h10 as [g h10]. subst. specialize (h7 _ h9). specialize (h8 _ h10). simpl in h7, h8. specialize (h7 h11). specialize (h8 h12). apply proj1_sig_injective. simpl. subst. red in h6. destruct h6 as [h6a h6b]. specialize (h1b _ _ h9 h10). destruct h1b as [k [h1b h1c]]. red in h1c. destruct h1c as [h1c h1d]. destruct h1c as [h1ca h1cb h1cc h1cd]. destruct h1d as [h1da h1db h1dc h1dd]. red in h1cd, h1dd. destruct h1cd as [h13 h14], h1dd as [h15 h16]. specialize (h14 (exist (Inn (ba_p_ens (projT1 f))) x h11)). specialize (h16 (exist (Inn (ba_p_ens (projT1 g))) y h12)). rewrite h14, h16 in h7. simpl in h7. specialize (h6b _ h1b). red in h6b. apply h6b in h7. pose proof (f_equal (@proj1_sig _ _) h7) as h17. simpl in h17. assumption. Qed. Lemma directed_common_extension : forall {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp), directed_fam_homo F -> Inhabited F -> exists (pfd: directed (fam_fun_ba_domains (ba_conv_fam_fun_ba_one_range F))) (pfi:Inhabited (fam_fun_ba_domains (ba_conv_fam_fun_ba_one_range F))), let Ap := directed_ba_p pfd pfi in exists (f:btp Ap->btp Bp), common_extension_fam_homo F f /\ (inj_fam_homo F -> Injective f). intros Bp F h1 h2. rewrite directed_fam_homo_p1_iff in h1. assert (h3:Inhabited (ba_conv_fam_fun_ba_one_range F)). destruct h2 as [x h2]. eapply Inhabited_intro. rewrite in_ba_conv_fam_fun_ba_one_range_iff in h2. apply h2. pose proof (directed_common_extension_p1 _ h1 h3) as h4. simpl in h4. destruct h4 as [h4 [h5 [f [h6 h7]]]]. exists h4, h5. simpl. exists f. rewrite common_extension_fam_homo_p1_iff. split. assumption. intro h8. rewrite inj_fam_homo_p1_iff in h8. apply h7; auto. Qed. Definition chain_fam_homo {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp) : Prop := fam_homo F /\ forall f g, Inn F f -> Inn F g -> extends_homo (projT2 f) (projT2 g) \/ extends_homo (projT2 g) (projT2 f). Definition chain_fam_homo_p1 {B:Bool_Alg} (F:fam_fun_ba_one_range_p1 B) : Prop := fam_homo_p1 F /\ forall f g, Inn F f -> Inn F g -> extends_homo_p1 (projT2 f) (projT2 g) \/ extends_homo_p1 (projT2 g) (projT2 f). Lemma chain_fam_homo_p1_iff : forall {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp), chain_fam_homo F <-> chain_fam_homo_p1 (ba_conv_fam_fun_ba_one_range F). intros Bp F. split. intro h1. red in h1. red. destruct h1 as [h1 h2]. rewrite fam_homo_p1_iff in h1. split; auto. intros f g h3 h4. destruct f as [A' f], g as [B' g]. simpl. rewrite <- (in_ba_conv_fam_fun_ba_one_range_iff F (existT _ _ (ba_conv_fun2 f))) in h3. rewrite <- (in_ba_conv_fam_fun_ba_one_range_iff F (existT _ _ (ba_conv_fun2 g))) in h4. specialize (h2 _ _ h3 h4). simpl in h2. do 2 rewrite extends_homo_p1_iff in h2. assumption. intro h1. red in h1. red. destruct h1 as [h1 h2]. split. rewrite fam_homo_p1_iff. assumption. intros f g h3 h4. rewrite in_ba_conv_fam_fun_ba_one_range_iff in h3, h4. specialize (h2 _ _ h3 h4). simpl in h2. do 2 rewrite extends_homo_p1_iff. assumption. Qed. Lemma chain_fam_homo_impl_directed_fam_homo_p1 : forall {B:Bool_Alg} (F:fam_fun_ba_one_range_p1 B), chain_fam_homo_p1 F -> directed_fam_homo_p1 F. intros B F h1. red. red in h1. destruct h1 as [h1a h1b]. split; auto. intros f g h2 h3. specialize (h1b _ _ h2 h3). destruct h1b as [h4 | h5]. exists f. split; auto. red. split; auto. apply refl_extends_homo_p1. destruct h4; auto. exists g. split; auto. red. split; auto. apply refl_extends_homo_p1. destruct h5; auto. Qed. Lemma chain_fam_homo_impl_directed_fam_homo : forall {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp), chain_fam_homo F -> directed_fam_homo F. intros Bp F h1. rewrite chain_fam_homo_p1_iff in h1. rewrite directed_fam_homo_p1_iff. apply chain_fam_homo_impl_directed_fam_homo_p1; auto. Qed. Lemma chain_common_extension : forall {Bp:Bool_Alg_p T} (F:fam_fun_ba_one_range Bp), chain_fam_homo F -> Inhabited F -> exists (pfd: directed (fam_fun_ba_domains (ba_conv_fam_fun_ba_one_range F))) (pfi:Inhabited (fam_fun_ba_domains (ba_conv_fam_fun_ba_one_range F))), let Ap := directed_ba_p pfd pfi in exists (f:btp Ap->btp Bp), common_extension_fam_homo F f /\ (inj_fam_homo F -> Injective f). intros Bp F h1 h2. pose proof (chain_fam_homo_impl_directed_fam_homo F h1) as h3. apply directed_common_extension; auto. Qed. Lemma chain_common_extension_p1 : forall {B:Bool_Alg} (F:fam_fun_ba_one_range_p1 B), chain_fam_homo_p1 F -> Inhabited F -> exists (pfd: directed (fam_fun_ba_domains_p1 F)) (pfi:Inhabited (fam_fun_ba_domains_p1 F)), let Ap:=directed_ba_p pfd pfi in exists (f:btp Ap->bt B), common_extension_fam_homo_p1 F f /\ (inj_fam_homo_p1 F -> Injective f). intros Bp F h1 h2. pose proof (chain_fam_homo_impl_directed_fam_homo_p1 F h1) as h3. apply directed_common_extension_p1; auto. Qed. Lemma gen_ens_determines_homo : forall {A B:Bool_Alg} (f g:bt B -> bt A), homomorphism f -> homomorphism g -> forall E:Ensemble (bt B), ba_ens B = Gen_Ens E -> agree_on f g E -> f = g. intros A B f g h1 h2 E h3 h4. destruct (classic (E = (Empty_set _))) as [h5 | h6]. subst. rewrite gen_ens_empty in h3. unfold ba_ens in h3. apply functional_extensionality. intros x. pose proof (Full_intro _ x) as h5. unfold bt in h5. unfold bt in h3. rewrite h3 in h5. destruct h5. rewrite (homo_zero f); auto. rewrite (homo_zero g); auto. rewrite (homo_one f); auto. rewrite (homo_one g); auto. apply not_empty_Inhabited in h6. pose [x:bt B | f x = g x] as C. assert (h5:Included E C). red. intros x h7. red in h4. specialize (h4 _ h7). constructor. assumption. assert (h7:alg_closed C). apply two_ops_times_closed. destruct h6 as [x h6]. apply Inhabited_intro with x. apply h5; auto. red. intros x. destruct x as [x h7]. destruct h7 as [h7]. unfold Bcomp_sub. simpl. constructor. rewrite (homo_comp f); auto. rewrite (homo_comp g); auto. f_equal. assumption. red. intros x y. destruct x as [x h7], y as [y h8]. destruct h7 as [h7], h8 as [h8]. unfold Btimes_sub. simpl. constructor. rewrite (homo_times f); auto. rewrite (homo_times g); auto. rewrite h7, h8. reflexivity. apply gen_ens_closed_eq in h7. apply gen_ens_preserves_inclusion in h5. rewrite h7 in h5. rewrite <- h3 in h5. assert (h8:C = ba_ens B). apply Extensionality_Ensembles. red. split. red. intros. constructor. assumption. unfold C in h8. apply functional_extensionality. intro x. pose proof (Full_intro _ x) as h9. unfold ba_ens in h8. unfold bt in h9. unfold bt in h8. rewrite <- h8 in h9. destruct h9. assumption. Qed. Lemma gen_ens_determines_homo_p1 : forall {Bp:Bool_Alg_p T} {A:Bool_Alg} (f g:btp Bp -> bt A), homomorphism_p1 _ f -> homomorphism_p1 _ g -> forall (E:Ensemble T) (pf:Included E (ba_p_ens Bp)), ba_p_ens Bp = Gen_Ens_p E pf -> agree_on f g (im_proj2_sig E pf) -> f = g. intros Bp A f g h1 h2 E h3 h4 h5. rewrite homomorphism_p1_iff in h1. rewrite homomorphism_p1_iff in h2. rewrite ba_p_ens_eq, gen_ens_p_eq in h4. apply im_proj1_sig_injective in h4. pose proof (gen_ens_determines_homo (ba_conv_fun1 f) (ba_conv_fun1 g) h1 h2 _ h4 h5) as h6. assumption. Qed. End Extensions. Section HomoExtensionCriterion. Lemma finite_homo_extension_criterion_if : forall {A B:Bool_Alg} {E:Ensemble (bt B)}, Finite E -> ba_ens B = Gen_Ens E -> forall (g:sig_set E->bt A), (forall F:Ensemble (bt B), Finite F -> Included F E -> forall a:Fin_map F signe mns, el_prod _ a = 0 -> el_prod_compose _ (sig_fun_app g 0) a = 0) -> exists! h:(bt B)->(bt A), homomorphism h /\ extends_sig1 h g. intros A B E h3 h1 g h2. pose proof (normal_form_gen_ba _ _ h3 h1) as h4. destruct h4 as [h4l h4r]. pose proof (normal_form_gen_im _ _ h3 g) as h5. simpl in h5. destruct h5 as [h5l h5r]. pose proof (gen_e_eq_c_unq _ _ h3) as h6. simpl in h6. rewrite <- h1 in h6. assert (h7:forall b:(bt B), exists! S : Ensemble (Fin_map E signe mns), exists (pf:Included S (non_zero_el_prod_maps B E)), plus_subset_non_zero_el_prod_maps B E h3 _ pf = b). intro b. assert (h7:Inn (ba_ens B) b). constructor. rewrite h6 in h7. destruct h7 as [h7]. destruct h7 as [S h7]. exists (non_zero_el_prod_maps_of_set B h3 S). red in h7. destruct h7 as [h7 h8]. red. destruct h7 as [h7 h9]. split. exists (incl_non_zero_el_prod_maps_of_set _ h3 S). rewrite <- h9. pose proof (plus_subset_el_prod_maps_eq_same_non_zero B h3 S) as h10. unfold plus_subset_non_zero_el_prod_maps. assert (h11: finite_image (Fin_map E signe mns) (Btype (Bc B)) S (el_prod B) (finite_fin_map_ens S h3 signe_finite) = finite_image (Fin_map E signe mns) (Btype (Bc B)) S (el_prod B) (Finite_downward_closed (Fin_map E signe mns) (non_zero_el_prod_maps B E) (non_zero_el_prod_maps_fin B E h3) S h7)). apply proof_irrelevance. rewrite h11 in h10 at 1. rewrite h10 at 1. apply plus_set_functional. reflexivity. intros S' h10. destruct h10 as [h10 h11]. rewrite <- h9 in h11. apply plus_subset_non_zero_el_prod_maps_inj in h11. subst. apply Extensionality_Ensembles. red. split. red. intros f h11. destruct h11. assumption. red. intros f h11. specialize (h10 _ h11). destruct h10 as [h10]. constructor; auto. pose (fun b:bt B => (proj1_sig (constructive_definite_description _ (h7 b)))) as fa. assert (h8:forall b:(bt B), Finite (fa b)). intro b. unfold fa. destruct constructive_definite_description as [S h9]. simpl. destruct h9 as [h9 h10]. unfold plus_subset_non_zero_el_prod_maps in h10. apply (Finite_downward_closed (Fin_map E signe mns) (non_zero_el_prod_maps B E) (non_zero_el_prod_maps_fin B E h3) S h9). pose (fun b:(bt B) => plus_subset_non_zero_el_prod_compose_maps _ _ h3 (sig_fun_app g 0) (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) (fa b)) (incl_non_zero_el_prod_compose_maps_of_set _ _ _ _ )) as f'. assert (h9:forall b:(bt B), Inn (Gen_Ens (Im (full_sig E) g)) (f' b)). intro b. unfold f'. assert (h10:Im E (sig_fun_app g 0) = Im (full_sig E) g). apply Extensionality_Ensembles. red. split. red. intros x h10. destruct h10 as [x h10]. subst. unfold sig_fun_app. destruct (DecidableDec.classic_dec (Inn E x)) as [h11 | h12]. apply Im_intro with (exist _ _ h11). constructor. reflexivity. contradiction. red. intros x h10. destruct h10 as [x h10]. subst. clear h10. destruct x as [x h10]. apply Im_intro with x; auto. unfold sig_fun_app. destruct (DecidableDec.classic_dec (Inn E x)) as [h11 | h12]. f_equal. apply proj1_sig_injective. reflexivity. contradiction. rewrite <- h10. apply in_gen_ens_im_g_plus_subset_el_prod_compose_maps. assert (hhom:homomorphism f'). pose (fun (b:bt B) => exist _ _ (h9 b)) as f. assert (h10:homomorphism (B:=Gen (Im (full_sig E) g)) f). apply homo_two_ops_plus. intros r s. unfold f. apply proj1_sig_injective. simpl. unfold Bplus_sub. simpl. unfold f'. unfold fa. destruct constructive_definite_description as [XY h10]. destruct constructive_definite_description as [X h11]. destruct constructive_definite_description as [Y h12]. simpl. destruct h10 as [h10a h10b], h11 as [h11a h11b], h12 as [h12a h12b]. subst. rewrite plus_subset_non_zero_el_prod_maps_union in h10b. apply plus_subset_non_zero_el_prod_maps_inj in h10b. subst. rewrite <- plus_subset_non_zero_el_prod_compose_maps_union. unfold plus_subset_non_zero_el_prod_compose_maps. pose proof (plus_subset_el_prod_compose_maps_eq_same_non_zero _ h3 (sig_fun_app g 0) (Union X Y)) as h11. assert (h12:(finite_image (Fin_map E signe mns) (Btype (Bc A)) (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) (Union X Y)) (el_prod_compose A (sig_fun_app g 0)) (finite_non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) (Union X Y))) = (finite_image (Fin_map E signe mns) (Btype (Bc A)) (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) (Union X Y)) (el_prod_compose A (sig_fun_app g 0)) (finite_fin_map_ens (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) (Union X Y)) h3 signe_finite))). apply proof_irrelevance. rewrite <- h12 at 1. rewrite <- h11 at 1. pose proof (im_union X Y (el_prod_compose A (sig_fun_app g 0))) as h13. assert (h14:Finite (Union (Im X (el_prod_compose A (sig_fun_app g 0))) (Im Y (el_prod_compose A (sig_fun_app g 0))))). rewrite <- h13. apply finite_image. apply finite_fin_map_ens. assumption. apply signe_finite. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Fin_map E signe mns) (Btype (Bc A)) (Union X Y) (el_prod_compose A (sig_fun_app g 0)) (finite_fin_map_ens (Union X Y) h3 signe_finite)) h14 h13) as h15. dependent rewrite -> h15. pose proof (finite_image _ _ _ (el_prod_compose A (sig_fun_app g 0)) (finite_fin_map_ens X h3 signe_finite)) as h16. pose proof (finite_image _ _ _ (el_prod_compose A (sig_fun_app g 0)) (finite_fin_map_ens Y h3 signe_finite)) as h17. rewrite (plus_set_union' _ _ h16 h17 h14) at 1. pose proof (im_union (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) X) (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) Y) (el_prod_compose A (sig_fun_app g 0))) as h18. assert (h19:Finite (Union (Im (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) X) (el_prod_compose A (sig_fun_app g 0))) (Im (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) Y) (el_prod_compose A (sig_fun_app g 0))))). rewrite <- h18. apply finite_image. apply finite_fin_map_ens; auto. apply signe_finite. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Fin_map E signe mns) (Btype (Bc A)) (Union (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) X) (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) Y)) (el_prod_compose A (sig_fun_app g 0)) (finite_fin_map_ens (Union (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) X) (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) Y)) h3 signe_finite)) h19 h18) as h20. dependent rewrite -> h20. assert (h21:Finite (Im (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) X) (el_prod_compose A (sig_fun_app g 0)))). apply finite_image. apply finite_non_zero_el_prod_compose_maps_of_set. assert (h22:Finite (Im (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) Y) (el_prod_compose A (sig_fun_app g 0)))). apply finite_image. apply finite_non_zero_el_prod_compose_maps_of_set. rewrite (plus_set_union' _ _ h21 h22). f_equal. pose proof (plus_subset_el_prod_compose_maps_eq_same_non_zero _ h3 (sig_fun_app g 0) X) as h23. assert (h24:plus_set (Im (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) X) (el_prod_compose A (sig_fun_app g 0))) (finite_image (Fin_map E signe mns) (Btype (Bc A)) (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) X) (el_prod_compose A (sig_fun_app g 0)) (finite_non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) X)) = plus_set (Im (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) X) (el_prod_compose A (sig_fun_app g 0))) h21). apply plus_set_functional; auto. rewrite h24 in h23. rewrite <- h23. apply plus_set_functional. reflexivity. pose proof (plus_subset_el_prod_compose_maps_eq_same_non_zero _ h3 (sig_fun_app g 0) Y) as h23. assert (h24:plus_set (Im (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) Y) (el_prod_compose A (sig_fun_app g 0))) (finite_image (Fin_map E signe mns) (Btype (Bc A)) (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) Y) (el_prod_compose A (sig_fun_app g 0)) (finite_non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) Y)) = plus_set (Im (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) Y) (el_prod_compose A (sig_fun_app g 0))) h22). apply plus_set_functional; auto. rewrite h24 in h23. rewrite <- h23. apply plus_set_functional. reflexivity. unfold f. intro r. apply proj1_sig_injective. simpl. unfold Bcomp_sub. simpl. unfold f', fa. destruct constructive_definite_description as [X' h10]. destruct constructive_definite_description as [X h11]. simpl. destruct h10 as [h10a h10b], h11 as [h11a h11b]. subst. pose proof (plus_subset_non_zero_el_prod_maps_comp B E h3 _ h11a) as h12. rewrite h12 in h10b. clear h12. apply plus_subset_non_zero_el_prod_maps_inj in h10b. subst. rewrite <- plus_subset_non_zero_el_prod_compose_maps_comp. apply plus_set_functional. f_equal. apply Extensionality_Ensembles. red. split. red. intros k h12. destruct h12 as [k h12 h13]. destruct h12 as [h12 h14]. destruct h12 as [h12]. destruct h13 as [h13]. constructor. constructor. assumption. intro h15. destruct h15 as [k h15 h16]. contradiction. red. intros k h12. destruct h12 as [h12 h13]. destruct h12 as [h12]. pose proof (contrapos _ _ (h2 _ h3 (inclusion_reflexive _) k)) as h17. pose proof (h17 h12) as h17'. constructor. constructor. constructor. assumption. contradict h13. constructor. assumption. constructor. assumption. constructor. assumption. unfold f in h10. destruct h10 as [h10a h10b h10c]. constructor. intros x y. pose proof (h10a x y) as h11. pose proof (f_equal (@proj1_sig _ _) h11) as h12. simpl in h12. rewrite h12. unfold Btimes_sub. simpl. reflexivity. intros x y. pose proof (h10b x y) as h11. pose proof (f_equal (@proj1_sig _ _) h11) as h12. simpl in h12. rewrite h12. unfold Bplus_sub. simpl. reflexivity. intro x. pose proof (h10c x) as h11. pose proof (f_equal (@proj1_sig _ _) h11) as h12. simpl in h12. rewrite h12. unfold Bcomp_sub. simpl. reflexivity. assert (hex:extends_sig1 f' g). red. intro x. destruct x as [x h9']. simpl. unfold f'. pose proof (elt_eq_plus_subset_non_zero_el_prod_maps_at_one _ _ h3 _ h9') as h10'. simpl in h10'. destruct h10' as [h10' h11']. pose proof (elt_eq_plus_subset_non_zero_el_prod_compose_maps_at_one _ _ h3 (sig_fun_app g 0) _ h9') as h12'. destruct h12' as [h12 h13]. unfold fa. destruct constructive_definite_description as [S h15]. simpl. destruct h15 as [h15 h16]. rewrite h11' in h16. apply plus_subset_non_zero_el_prod_maps_inj in h16. rewrite h16. clear h15. clear h16. assert (heq:sig_fun_app g 0 x = g (exist _ _ h9')). unfold sig_fun_app. destruct DecidableDec.classic_dec as [h14 | h15]. f_equal. apply proj1_sig_injective. reflexivity. contradiction. rewrite <- heq at 1. rewrite h13. assert (h14: [a : Fin_map E signe mns | Inn (non_zero_el_prod_compose_maps A (sig_fun_app g 0) E) a /\ a |-> x = pls] = (non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) [a : Fin_map E signe mns | Inn (non_zero_el_prod_maps B E) a /\ a |-> x = pls])). apply Extensionality_Ensembles. red. split. red. intros k h14. destruct h14 as [h14]. destruct h14 as [h14 h15]. destruct h14 as [h14]. constructor. constructor. split; auto. pose proof (h2 _ h3 (inclusion_reflexive _) k) as h16. apply contrapos in h16. constructor. assumption. assumption. constructor. assumption. red. intros k h14. destruct h14 as [k h14 h15]. destruct h14 as [h14]. destruct h14 as [h14a h14b]. destruct h15 as [h15]. constructor. split. constructor. assumption. assumption. pose proof (subsetT_eq_compat _ (fun S => Included S (non_zero_el_prod_compose_maps A (sig_fun_app g 0) E)) _ _ h12 (incl_non_zero_el_prod_compose_maps_of_set A h3 (sig_fun_app g 0) [a : Fin_map E signe mns | Inn (non_zero_el_prod_maps B E) a /\ a |-> x = pls]) h14) as h15. dependent rewrite -> h15. reflexivity. exists f'. red. split. split; auto. intros k h10. destruct h10 as [h10 h11]. assert (hag:agree_on f' k E). red in hex. red in h11. red. intros x h12. specialize (hex (exist _ _ h12)). specialize (h11 (exist _ _ h12)). simpl in hex, h11. congruence. pose proof (gen_ens_determines_homo f' k hhom h10 _ h1 hag) as h12. assumption. Qed. Lemma finite_homo_extension_criterion_if_p : forall {T:Type} {Ap Bp:Bool_Alg_p T} {E:Ensemble (btp Bp)}, Finite E -> ba_p_ens Bp = Gen_Ens_p _ (incl_im_proj1_sig_ba_p_ens E) -> forall (g:sig_set E->btp Ap), (forall F:Ensemble (btp Bp), Finite F -> Included F E -> forall a:Fin_map F signe mns, el_prod_p _ a = %0 -> el_prod_compose_p _ _ (sig_fun_app g %0) a = %0) -> exists! h:(btp Bp)->(btp Ap), homomorphism_p h /\ extends_sig1 h g. intros T Ap Bp E h1 h2 g h3. rewrite gen_ens_p_eq in h2. rewrite ba_p_ens_eq in h2 at 1. apply im_proj1_sig_injective in h2. rewrite <- im_proj2_sig_undoes_im_proj1_sig' in h2 at 1. assert (h5: (forall F : Ensemble (bt (ba_conv Bp)), Finite F -> Included F E -> forall a : Fin_map F signe mns, el_prod (ba_conv Bp) a = 0 -> el_prod_compose (ba_conv Ap) (sig_fun_app (ba_conv_sig_fun1 g) %0) a = %0)). intros F h5 h6 a h7. assert (h7': el_prod (ba_conv Bp) (ba_conv_fin_map_dom a) = 0). assumption. rewrite <- el_prod_p_eq in h7'. specialize (h3 F h5 h6 a h7'). rewrite el_prod_compose_p_eq in h3. assumption. pose proof (finite_homo_extension_criterion_if (A:=ba_conv Ap) (B:=ba_conv Bp) h1 h2 (ba_conv_sig_fun1 g) h5) as h4. destruct h4 as [k h4]. red in h4. destruct h4 as [h4a h4b]. exists k. red. split. rewrite homomorphism_p_iff. assumption. intro f. specialize (h4b f). rewrite homomorphism_p_iff. assumption. Qed. Lemma finite_homo_extension_criterion_if_p1 : forall {T':Type} {Bp:Bool_Alg_p T'} {A:Bool_Alg} {E:Ensemble (btp Bp)}, Finite E -> ba_p_ens Bp = Gen_Ens_p _ (incl_im_proj1_sig_ba_p_ens E) -> forall (g:sig_set E->bt A), (forall F:Ensemble (btp Bp), Finite F -> Included F E -> forall a:Fin_map F signe mns, el_prod_p _ a = %0 -> el_prod_compose_p1 _ (sig_fun_app g 0) a = 0) -> exists! h:(btp Bp)->(bt A), homomorphism_p1 _ h /\ extends_sig1 h g. intros T' Bp A E h1 h2 g h3. rewrite gen_ens_p_eq in h2. rewrite <- im_proj2_sig_undoes_im_proj1_sig' in h2. rewrite ba_p_ens_eq in h2. apply im_proj1_sig_injective in h2. assert (h5:forall F : Ensemble (bt (ba_conv Bp)), Finite F -> Included F E -> forall a : Fin_map F signe mns, el_prod (ba_conv Bp) a = 0 -> el_prod_compose A (sig_fun_app g 0) a (B:=ba_conv Bp) = 0). intros F h6 h7 a h8. rewrite <- (el_prod_p_eq _ a) in h8 at 1. specialize (h3 _ h6 h7 a h8). rewrite el_prod_compose_p1_eq in h3. assumption. pose proof (finite_homo_extension_criterion_if (B:=ba_conv Bp) h1 h2 g h5) as h4. destruct h4 as [k h4]. red in h4. destruct h4 as [h4 h6]. exists k. red. split. rewrite homomorphism_p1_iff. assumption. intros x' h7. rewrite homomorphism_p1_iff in h7. specialize (h6 _ h7). subst. reflexivity. Qed. (*You might be curious as to why I use [Bp:Bool_alg_p T], and not [B:Bool_alg]. At some point, I did just that, and was required at some point to consider families of subalgebras of it, and that required me to paramaterize and use a conversion function like [ba_to_ba_p], from [Bool_alg]s to [Bool_alg_p]s, which caused universe inconsistencies during some of my previous efforts to use well-founded induction. The meat is the same, and this is a neater proof than the non-parametric version, and more stable against universe inconsistencies.*) Theorem homo_extension_criterion_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)}, ba_p_ens Bp = Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E) -> forall (g:sig_set E->bt A), (exists h:(btp Bp)->(bt A), homomorphism_p1 _ h /\ extends_sig1 h g) <-> (forall F : Ensemble (btp Bp), Finite F -> Included F E -> forall a : Fin_map F signe mns, el_prod_p _ a = %0 -> el_prod_compose_p1 T (sig_fun_app g 0) a = 0). intros T Bp A E h1 g. split. Focus 2. (* <- *) intro h2. assert (h0:forall F:(finc E), Included (im_proj1_sig (proj1_sig F)) (ba_p_ens Bp)). intro F. destruct F as [F h3]. destruct h3 as [h3a h3b]. simpl. red. intros x ha. destruct ha as [x ha]. subst. apply proj2_sig. pose (fun F:(finc E) => Gen_p _ (h0 F)) as B_F. assert (h0':forall F:(finc E), Included (im_proj1_sig (proj1_sig F)) (ba_p_ens (B_F F))). unfold B_F. unfold B_F. intro F. rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. pose (fun F:(finc E) => im_proj2_sig _ (h0' F)) as finc_bp. assert (h8:forall F:(finc E), (im_proj1_sig (proj1_sig F)) = (im_proj1_sig (finc_bp F))). intro F. unfold finc_bp. rewrite im_proj1_sig_undoes_im_proj2_sig at 1. reflexivity. assert (hi:forall F:finc E, Included (ba_p_ens (B_F F)) (ba_p_ens Bp)). intro F. rewrite h1. unfold B_F. rewrite ba_p_ens_gen_p_eq. apply gen_ens_preserves_inclusion_p. destruct F as [F h9]. destruct h9 as [h9 h10]. simpl. rewrite incl_im_proj1_sig_iff. assumption. assert (h9:forall (F:(finc E)) (x:sig_set (finc_bp F)), Inn E (exist _ _ (hi F _ (proj2_sig (proj1_sig x))))). intros F x. destruct x as [x h10]. simpl. destruct x as [x h11]. simpl. destruct F as [F h12]. destruct h12 as [h12 h13]. inversion h10 as [y h14 ? h15]. subst. simpl in y. clear h14. apply exist_injective in h15. subst. destruct y as [y ha]. simpl. destruct ha as [y ha]. subst. apply h13. inversion h10 as [y' hc ? hd]. subst. apply exist_injective in hd. simpl in hd. simpl in y'. destruct y' as [y' hy]. destruct hy as [y' hy]. subst. assert (he:y' = exist (Inn (ba_p_ens Bp)) (proj1_sig y) (hi (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h12 h13)) (proj1_sig y) h11)). apply proj1_sig_injective. simpl. simpl in hd. rewrite hd. reflexivity. subst. assumption. pose (fun F:(finc E) => (fun x:sig_set (finc_bp F)=> g (exist _ _ (h9 _ x)))) as g_F. assert (h4:forall F:(finc E), Finite (finc_bp F)). intro F. unfold finc_bp. apply finite_image. rewrite <- finite_full_sig_iff. destruct F as [F [ha hb]]. simpl. apply finite_image; auto. assert (h3:forall F:(finc E), ba_p_ens (B_F F) = Gen_Ens_p (im_proj1_sig (finc_bp F)) (incl_im_proj1_sig_ba_p_ens (finc_bp F))). intro F. unfold B_F. rewrite ba_p_ens_gen_p_eq at 1. pose proof (gen_p_subalg_of_p_compat _ (h0 F)) as h3. assert (h5:Included (im_proj1_sig (finc_bp F)) (ba_p_ens Bp)). red. intros x ha. destruct ha as [x ha]. subst. destruct x as [x hb]. inversion ha as [y hc q hd]. subst. simpl. apply exist_injective in hd. subst. apply (hi F). assumption. pose proof (gen_ens_p_subalg_of_p h3 _ (incl_im_proj1_sig_ba_p_ens (finc_bp F)) h5) as h6. unfold bt, ba_p_ens, B_F. unfold bt, ba_p_ens, B_F in h6. rewrite h6 at 1. apply gen_ens_p_functional; auto. apply h8. assert (h10:forall (F:(finc E)) (D : Ensemble (btp (B_F F))), Finite D -> Included D (finc_bp F) -> forall a : Fin_map D signe mns, el_prod_p _ a = %0 -> el_prod_compose_p1 T (sig_fun_app (g_F F) 0) a = 0). intros F D h11 h12 a h13. pose (Im D (fun x => (exist _ _ (hi _ _ (proj2_sig x))))) as D'. assert (h14:Finite D'). unfold D'. apply finite_image. assumption. assert (h15:Included D' E). red in h12. red. intros x h15. destruct h15 as [x h15]. subst. destruct x as [x h16]. simpl. specialize (h12 _ h15). inversion h12 as [c h17 ? h18]. subst. apply exist_injective in h18. subst. destruct F as [F h18]. simpl in h12, h15, h16, c, h17. destruct h18 as [h18a h18b]. red in h18b. destruct c as [c ha]. simpl. destruct ha as [c ha]. pose proof (h18b _ ha) as hb. subst. assert (hc:c = exist (Inn (ba_p_ens Bp)) (proj1_sig c) (hi (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h18a h18b)) (proj1_sig c) h16)). apply proj1_sig_injective. simpl. reflexivity. rewrite <- hc at 1. assumption. assert (h16:forall x:sig_set D', Inn (Gen_Ens_p _ (h0 F)) (proj1_sig (proj1_sig x))). intro x. destruct x as [x h16]. simpl. destruct h16 as [x h16]. subst. destruct x as [x h17]. simpl. assumption. pose (fun x:sig_set D' => exist _ _ (h16 x)) as fp. pose (fun x:sig_set D' => a |-> (fp x)) as fa. pose (sig_fun_to_fin_map ba_eq_dec_p fa h14 mns) as ga. assert (h17:Included (Im (full_sig D') fa) signe). red. intros. destruct x. left. right. pose (fin_map_new_ran ga signe_finite h17) as fa'. assert (hg:subalg_of_p (B_F F) Bp). unfold B_F. apply gen_p_subalg_of_p_compat. assert (h18:el_prod_p _ fa' = %0). pose proof hg as hg'. apply proj1_sig_injective. pose proof (f_equal (@proj1_sig _ _) h13) as ha. red in hg. destruct hg as [hga [hgb hgc]]. pose proof (ba_p_subst_zero _ _ hgc) as hb. rewrite <- hb at 1. rewrite <- h13. unfold el_prod_p. apply times_set_sub_compat_p; auto. unfold im_proj1_sig, D'. do 3 rewrite im_im. apply im_ext_in. intros x hc. unfold eps_p, fa'. assert (he: (if fin_map_new_ran ga signe_finite h17 |-> exist (Inn (ba_p_ens Bp)) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x)) then exist (Inn (ba_p_ens Bp)) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x)) else %-exist (Inn (ba_p_ens Bp)) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x))) = (if a |-> x then exist (Inn (ba_p_ens Bp)) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x)) else %-exist (Inn (ba_p_ens Bp)) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x)))). pose proof (fin_map_new_ran_compat ga signe_finite h17 (exist (Inn (ba_p_ens Bp)) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x)))) as hd. rewrite <- hd. unfold ga. assert (hf:Inn D' (exist (Inn (ba_p_ens Bp)) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x)))). unfold D'. apply Im_intro with x; auto. pose proof (sig_fun_to_fin_map_compat ba_eq_dec_p fa h14 mns (exist (Inn (ba_p_ens Bp)) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x))) hf) as he. rewrite he. unfold fa, fp. simpl. assert (hg:a |-> x = a |-> exist _ _ (h16 (exist (Inn D') (exist (Inn (ba_p_ens Bp)) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x))) hf))). f_equal. apply proj1_sig_injective; auto. simpl in hg. rewrite <- hg at 1. reflexivity. unfold D'. apply Im_intro with x; auto; apply proj1_sig_injective; auto. pose proof (f_equal (@proj1_sig _ _) he) as hf. unfold ba_p_ens. unfold ba_p_ens. unfold btp in hf. unfold btp. unfold Btype_p. unfold Btype_p in hf. unfold ba_p_ens. unfold ba_p_ens in hf. unfold sig_set. unfold sig_set in hf. unfold D'. unfold D' in hf. rewrite hf at 1. assert (hg: (proj1_sig (if a |-> x then exist (Inn (A_p T (Bc_p T Bp))) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x)) else %-exist (Inn (A_p T (Bc_p T Bp))) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x)))) = (if a |-> x then (proj1_sig x) else (proj1_sig (%-exist (Inn (A_p T (Bc_p T Bp))) (proj1_sig x) (hi F (proj1_sig x) (proj2_sig x)))))). destruct (a |-> x). simpl. reflexivity. reflexivity. rewrite hg at 1. assert (hh:proj1_sig (if a |-> x then x else %-x) = (if a |-> x then proj1_sig x else proj1_sig (%-x))). destruct (a |-> x); auto. rewrite hh at 1. destruct (a |-> x); auto. assert (hk:Inn (ba_p_ens (Subalg_p Bp (ba_p_ens (B_F F)) hga hgb)) (proj1_sig x)). rewrite <- hgc. apply proj2_sig. pose proof (ba_p_subst_comp _ _ hgc _ (proj2_sig x) hk) as hj. simpl in hj. assert (hl: incl_gen_ens_p (im_proj1_sig (proj1_sig F)) (h0 F) (proj1_sig x) (proj2_sig x) = hi F _ (proj2_sig x)). apply proof_irrelevance. rewrite <- hl. unfold B_F in hj. unfold B_F. rewrite hj at 1. f_equal. simpl. f_equal. destruct x as [x hx]. simpl. f_equal. apply proj1_sig_injective; auto. pose proof (h2 D' h14 h15 fa' h18) as h20. rewrite <- h20. unfold el_prod_compose_p1. apply times_set_functional. unfold D', im_proj1_sig. rewrite im_im. apply im_ext_in. intros x hin. f_equal. unfold sig_fun_app. destruct (classic_dec (Inn (finc_bp F) x)) as [h21 | h22]. unfold g_F. destruct classic_dec as [h23 | h24]. f_equal. apply proj1_sig_injective; auto. contradict h24. destruct x as [x h22]. simpl. pose proof (h9 F (exist _ _ h21)) as h23. simpl in h23. assumption. contradict h22. apply h12. assumption. unfold fa'. unfold ba_conv_fin_map_dom. rewrite <- (fin_map_new_ran_compat ga signe_finite h17). unfold ga. assert (h21:Inn D' (exist _ _ (hi F (proj1_sig x) (proj2_sig x)))). apply Im_intro with x. assumption. reflexivity. rewrite (sig_fun_to_fin_map_compat ba_eq_dec_p fa h14 mns _ h21). unfold fa. unfold fp. simpl. f_equal. apply proj1_sig_injective. simpl. reflexivity. unfold D'. apply Im_intro with x; auto; apply proj1_sig_injective; auto. pose (fun F:(finc E) => proj1_sig (constructive_definite_description _ (finite_homo_extension_criterion_if_p1 (A:=A) (h4 F) (h3 F) (g_F F) (h10 F)))) as psi. pose (Im (Full_set (finc E)) (fun F => existT (fun C:(Bool_Alg_p T) => (btp C->bt A)) (B_F F) (psi F))) as fam. assert (h11:directed_fam_homo_p1 _ fam). red. split. red. intros f h11. destruct h11 as [F h11]. subst. simpl. unfold psi. destruct constructive_definite_description as [f h12]. simpl. destruct h12 as [h12 h13]. assumption. intros F G h11 h12. destruct h11 as [F h11], h12 as [G h12]. subst. destruct F as [F h13], G as [G h14]. destruct h13 as [h13a h13b], h14 as [h14a h14b]. simpl. pose (Union F G) as H'. assert (h15:Finite H'). apply Union_preserves_Finite; auto. assert (h16:Included H' E). unfold H'. auto with sets. pose (exist (fun S=>Finite S /\ Included S E) _ (conj h13a h13b)) as F'. pose (exist (fun S=>Finite S /\ Included S E) _ (conj h15 h16)) as H''. assert (h17:Included (ba_p_ens (B_F F')) (ba_p_ens (B_F H''))). unfold B_F. do 2 rewrite ba_p_ens_gen_p_eq. apply gen_ens_preserves_inclusion_p. unfold F', H''. simpl. unfold H'. rewrite incl_im_proj1_sig_iff. auto with sets. exists (existT (fun C:(Bool_Alg_p T) => (btp C->bt A)) (B_F H'') (psi H'')). split. unfold fam. apply Im_intro with H''. constructor. reflexivity. simpl. red. split. split. red. exists h17. unfold B_F. simpl. assert (h18: alg_closed_p (ba_p_ens (Gen_p (im_proj1_sig F) (h0 (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h13a h13b))))) h17). pose proof (ba_p_ens_gen_p_eq (im_proj1_sig F) (h0 (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h13a h13b)))) as h18. assert (h19:Included (Gen_Ens_p (im_proj1_sig F) (h0 (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h13a h13b)))) (ba_p_ens (B_F H''))). rewrite <- h18. assumption. pose proof (subsetT_eq_compat _ (fun S => Included S (ba_p_ens (B_F H''))) _ _ h17 h19 h18) as h20. unfold B_F, F' in h20. simpl in h20. dependent rewrite -> h20. pose proof (gen_ens_includes_p _ (h0 F')) as h22. unfold F' in h22. simpl in h22. assert (h23:Included (im_proj1_sig F) (ba_p_ens (B_F H''))). auto with sets. pose proof (closed_gen_ens_p _ h23) as h24. assert (h25:subalg_of_p (B_F H'') Bp). unfold B_F. apply gen_p_subalg_of_p_compat. pose proof (gen_ens_p_subalg_of_p h25 _ h23 (h0 F')) as h26. pose proof (subsetT_eq_compat _ (fun S => Included S (ba_p_ens (B_F H''))) _ _ (incl_gen_ens_p (im_proj1_sig F) h23) h19 h26) as h27. dependent rewrite -> h27 in h24. assumption. exists h18. assert (h19:Included (im_proj1_sig F) (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig H'')) (h0 H'')))). rewrite ba_p_ens_gen_p_eq. assert (h20:Included F H'). unfold H'. auto with sets. pose proof (gen_ens_includes_p _ (h0 H'')) as h21. unfold H'' in h21. simpl in h21. auto with sets. pose proof (incl_im_proj1_sig_iff F H') as ha. pose proof (iff2 ha h20) as hb. unfold H''. simpl. red. intros x hc. apply hb in hc. apply h21 in hc. assumption. assert (h20: Subalg_p (Gen_p (im_proj1_sig H') (h0 H'')) (ba_p_ens (Gen_p (im_proj1_sig F) (h0 (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h13a h13b))))) h17 h18 = Gen_p (im_proj1_sig F) h19). apply subalg_functional_p. rewrite ba_p_ens_gen_p_eq at 1. symmetry. apply gen_ens_p_subalg_of_p. apply gen_p_subalg_of_p_compat. rewrite h20 at 1. symmetry. apply gen_p_subalg_of_p. apply gen_p_subalg_of_p_compat. unfold psi. destruct constructive_definite_description as [f h18]. simpl. destruct h18 as [h18 h19]. assumption. unfold psi. destruct constructive_definite_description as [f h18]. simpl. destruct h18; auto. unfold psi. destruct constructive_definite_description as [f h18]. destruct constructive_definite_description as [k h19]. simpl. destruct h18 as [h18a h18b]. destruct h19 as [h19a h19b]. red in h18b. red in h19b. pose (restriction_sig f _ h17) as f'. assert (hinc:Included F H'). unfold H'. auto with sets. pose proof (incl_im_proj1_sig_iff F H') as ha. pose proof (iff2 ha hinc) as hinc'. assert (h20:f' = k). pose proof (subalg_of_p_gen_p_gen_p _ _ _ hinc' (h0 F') (h0 H'')) as h22. assert (h23:Included (ba_p_ens (Gen_p _ (h0 F'))) (ba_p_ens (Gen_p _ (h0 H'')))). do 2 rewrite ba_p_ens_gen_p_eq. assumption. pose proof (homomorphism_p1_restriction_sig _ _ f h18a _ h23 h22) as h24. assert (h25:Included (im_proj1_sig F) (ba_p_ens (Gen_p _ (h0 F')))). rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. assert (h26: ba_p_ens (Gen_p (im_proj1_sig F) (h0 F')) = Gen_Ens_p _ h25). rewrite ba_p_ens_gen_p_eq. assert (h26:subalg_of_p (Gen_p _ (h0 F')) Bp). apply gen_p_subalg_of_p_compat. rewrite (gen_ens_p_subalg_of_p h26 (im_proj1_sig F) h25 (h0 F')). reflexivity. assert (h27: agree_on (restriction_sig f (ba_p_ens (Gen_p _ (h0 F'))) h23) k (im_proj2_sig (im_proj1_sig F) h25)). red. intros x h27. assert (h29:Inn (finc_bp (exist _ _ (conj h13a h13b))) x). destruct h27 as [x h27]. subst. unfold finc_bp. simpl. apply Im_intro with x. assumption. apply proj1_sig_injective. simpl. reflexivity. specialize (h19b (exist _ _ h29)). simpl in h19b. rewrite <- h19b. destruct h27 as [x h27]. subst. destruct x as [x h30]. simpl. assert (h31:Inn (ba_p_ens (Gen_p _ (h0 H''))) x). apply h23. apply h25. assumption. assert (h32:Inn (finc_bp H'') (exist _ _ h31)). unfold finc_bp. apply Im_intro with (exist _ _ (hinc' _ h30)). constructor. apply proj1_sig_injective. simpl. reflexivity. unfold restriction_sig. simpl. specialize (h18b (exist _ _ h32)). simpl in h18b. assert (h33: f (exist (Inn (Gen_Ens_p _ (h0 H''))) x (h23 x (h25 x h30))) = f (exist (Inn (ba_p_ens (Gen_p _ (h0 H'')))) x h31)). f_equal. apply proj1_sig_injective. simpl. reflexivity. unfold H'' in h33. unfold H''. rewrite h33 at 1. simpl. unfold H'' in h18b. rewrite <- h18b. unfold g_F. f_equal. apply proj1_sig_injective. simpl. f_equal. apply proof_irrelevance. unfold f'. pose proof (gen_ens_determines_homo_p1 _ _ _ h24 h19a (im_proj1_sig F) h25 h26 h27) as h28. rewrite <- h28. f_equal. apply proof_irrelevance. red. assert (h21:Included (Gen_Ens_p (im_proj1_sig F) (h0 (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h13a h13b)))) (Gen_Ens_p _ (h0 H''))). apply gen_ens_preserves_inclusion_p. unfold H'. auto with sets. exists h21. intro x. rewrite <- h20. unfold f', restriction_sig. f_equal. apply proj1_sig_injective. simpl. reflexivity. pose (exist (fun S=>Finite S /\ Included S E) _ (conj h14a h14b)) as G'. assert (h17':Included (ba_p_ens (B_F G')) (ba_p_ens (B_F H''))). unfold B_F. do 2 rewrite ba_p_ens_gen_p_eq. apply gen_ens_preserves_inclusion_p. unfold G', H''. simpl. unfold H'. rewrite incl_im_proj1_sig_iff. auto with sets. unfold psi. destruct constructive_definite_description as [f h18]. destruct constructive_definite_description as [k h19]. simpl. destruct h18 as [h18a h18b]. destruct h19 as [h19a h19b]. red in h18b. red in h19b. pose (restriction_sig f _ h17') as f'. assert (hinc:Included (im_proj1_sig G) (im_proj1_sig H')). rewrite incl_im_proj1_sig_iff. unfold H'. auto with sets. assert (h20:f' = k). pose proof (subalg_of_p_gen_p_gen_p _ _ _ hinc (h0 G') (h0 H'')) as h22. assert (h23:Included (ba_p_ens (Gen_p _ (h0 G'))) (ba_p_ens (Gen_p _ (h0 H'')))). do 2 rewrite ba_p_ens_gen_p_eq. assumption. pose proof (homomorphism_p1_restriction_sig _ _ f h18a _ h23 h22) as h24. assert (h25:Included (im_proj1_sig G) (ba_p_ens (Gen_p _ (h0 G')))). rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. assert (h26: ba_p_ens (Gen_p _ (h0 G')) = Gen_Ens_p _ h25). rewrite ba_p_ens_gen_p_eq. assert (h26:subalg_of_p (Gen_p _ (h0 G')) Bp). apply gen_p_subalg_of_p_compat. rewrite (gen_ens_p_subalg_of_p h26 _ h25 (h0 G')). reflexivity. assert (h27: agree_on (restriction_sig f (ba_p_ens (Gen_p _ (h0 G'))) h23) k (im_proj2_sig _ h25)). red. intros x h27. assert (h29:Inn (finc_bp (exist _ _ (conj h14a h14b))) x). destruct h27 as [x h27]. subst. unfold finc_bp. simpl. apply Im_intro with x. assumption. apply proj1_sig_injective. simpl. reflexivity. specialize (h19b (exist _ _ h29)). simpl in h19b. rewrite <- h19b. destruct h27 as [x h27]. subst. destruct x as [x h30]. simpl. assert (h31:Inn (ba_p_ens (Gen_p _ (h0 H''))) x). apply h23. apply h25. assumption. assert (h32:Inn (finc_bp H'') (exist _ _ h31)). unfold finc_bp. apply Im_intro with (exist _ _ (hinc _ h30)). constructor. apply proj1_sig_injective. simpl. reflexivity. unfold restriction_sig. simpl. specialize (h18b (exist _ _ h32)). simpl in h18b. assert (h33: f (exist (Inn (Gen_Ens_p _ (h0 H''))) x (h23 x (h25 x h30))) = f (exist (Inn (ba_p_ens (Gen_p _ (h0 H'')))) x h31)). f_equal. apply proj1_sig_injective. simpl. reflexivity. unfold H'' in h33. unfold H''. rewrite h33 at 1. rewrite <- h18b at 1. unfold g_F. f_equal. apply proj1_sig_injective. simpl. apply proj1_sig_injective; auto. unfold f'. pose proof (gen_ens_determines_homo_p1 _ _ _ h24 h19a _ h25 h26 h27) as h28. rewrite <- h28. f_equal. apply proof_irrelevance. rewrite <- h20. constructor. unfold B_F. simpl. apply subalg_of_p_gen_p_gen_p. unfold H'. auto with sets. assumption. apply homomorphism_p1_restriction_sig. assumption. apply subalg_of_p_gen_p_gen_p. unfold G', H''. simpl. unfold H'. auto with sets. unfold f'. red. assert (h21: Included (A_p T (Bc_p T (B_F (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) G (conj h14a h14b))))) (A_p T (Bc_p T (B_F H'')))). simpl. apply gen_ens_preserves_inclusion_p. unfold H'. auto with sets. exists h21. intro x. unfold restriction_sig. f_equal. apply proj1_sig_injective. simpl. reflexivity. assert (h12:Inhabited (Full_set (finc E))). assert (h13:Finite (Empty_set (btp Bp))). auto with sets. assert (h14:Included (Empty_set (btp Bp)) E). auto with sets. unfold finc. apply Inhabited_intro with (exist (fun F=>Finite F /\ Included F E) _ (conj h13 h14)). constructor. assert (h13:Inhabited fam). unfold fam. destruct h12 as [K h12]. apply Inhabited_intro with ((fun F : (finc E) => existT (fun C : Bool_Alg_p T => btp C -> bt A) (B_F F) (psi F)) K). apply Im_intro with K. assumption. reflexivity. pose proof (directed_common_extension_p1 _ _ h11 h13) as h14. destruct h14 as [h14 [h15 [f [h17 h18]]]]. assert (hb':forall Cp:Bool_Alg_p T, Inn (fam_fun_ba_domains_p1 T fam) Cp -> subalg_of_p Cp Bp). intros Cp hc. destruct hc as [Cp hc]. subst. destruct hc as [Cp hc]. subst. simpl. unfold B_F. apply gen_p_subalg_of_p_compat. assert (h19:directed_ba_p h14 h15 = Bp). assert (h22:ba_p_ens (directed_ba_p h14 h15) = ba_p_ens Bp). apply Extensionality_Ensembles. red. split. red. intros x ha. unfold ba_p_ens, directed_ba_p in ha. simpl in ha. destruct ha as [S x hb hc]. destruct hb as [Ap hb]. subst. destruct hb as [Ap hb]. subst. destruct hb as [Ap hb]. subst. simpl in hc. apply (hi Ap); auto. rewrite h1 at 1. apply gen_ens_minimal_p. constructor. split. red. intros x h22. destruct h22 as [x h22]. subst. destruct x as [x hx]. simpl. assert (h23:Included (Singleton x) (ba_p_ens Bp)). red. intros x' hx'. destruct hx'. assumption. assert (h23':Included (Singleton (exist _ _ hx)) E). red. intros x' h23'. destruct h23'; auto. assert (h24:Finite (Singleton (exist _ _ hx))). apply Singleton_is_finite. unfold ba_p_ens, directed_ba_p. simpl. apply family_union_intro with (Gen_Ens_p (Singleton x) h23). apply Im_intro with (Gen_p (Singleton x) h23). unfold fam, fam_fun_ba_domains_p1. rewrite im_im. simpl. unfold finc. apply Im_intro with (existT (fun S => Finite S /\ Included S E) _ (conj h24 h23')). constructor. unfold B_F. simpl. assert (h25:im_proj1_sig (Singleton (exist _ _ hx)) = Singleton x). apply Extensionality_Ensembles. red. split. red. intros x' h25. destruct h25 as [x' h25]. subst. destruct h25. simpl. constructor. red. intros x' h25. destruct h25. apply Im_intro with (exist _ _ hx). constructor. simpl. reflexivity. symmetry. gen0. simpl. rewrite h25. intro h26. f_equal. apply proof_irrelevance. rewrite ba_p_ens_gen_p_eq. reflexivity. apply gen_ens_includes_p. constructor. assert (ha: Included (ba_p_ens (directed_ba_p h14 h15)) (ba_p_ens Bp)). red. intros x ha. unfold ba_p_ens, directed_ba_p in ha. simpl in ha. destruct ha as [S x hb hc]. destruct hb as [Ap hb]. subst. destruct hb as [Ap hb]. subst. destruct hb as [Ap hb]. subst. simpl in hc. apply (hi Ap); auto. exists ha. assert (hc:ha = incl_ba_p_ens_directed_ba_p h14 h15 _ hb'). apply proof_irrelevance. subst. apply alg_closed_p_ba_p_ens_directed_ba_p. apply subalg_of_p_eq; auto. apply subalg_of_p_directed_ba_p; auto. assert (h20:btp (directed_ba_p h14 h15) = btp Bp). f_equal; auto. pose (transfer_fun h20 f) as f'. pose proof (common_extension_fam_homo_p1_homo_p1 _ _ _ f fam h13 h17) as h21. exists f'. split. unfold f'. rewrite (ba_p_subst_homomorphism_p1 _ _ _ _ h19 h20) in h21. rewrite homomorphism_p1_iff in h21. unfold ba_conv_fun1 in h21. unfold Btype, ba_conv in h21. simpl in h21. simpl in h21. rewrite homomorphism_p1_iff. assumption. red. intro x. destruct h17 as [h17a h17b]. destruct x as [x h22]. simpl. destruct x as [x h22']. unfold f'. rewrite transfer_fun_eq. assert (h23:Included (Singleton (exist _ _ h22')) E). red; intros x' h24. destruct h24; auto. assert (h23':Included (Singleton x) (im_proj1_sig E)). red; intros x' h24. destruct h24; auto. apply Im_intro with (exist _ _ h22'). assumption. reflexivity. assert (h24:Included (Singleton x) (ba_p_ens Bp)). red; intros x' h25. destruct h25; auto. assert (h25:Inn (ba_p_ens (Gen_p (Singleton x) h24)) x). rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. constructor. assert (h26:Included (im_proj1_sig E) (ba_p_ens Bp)). red; intros x' hd. destruct hd as [x' hd q he]. rewrite he. clear he q. apply proj2_sig. pose proof (subalg_of_p_gen_p_gen_p Bp (Singleton x) _ h23' h24 h26) as h27. pose proof h27 as h27'. destruct h27' as [h27a [h27b h27c]]. assert (h28:Finite (Singleton x)). apply Singleton_is_finite. assert (h30:Finite (im_proj2_sig (Singleton x) h24)). apply finite_image; auto. rewrite <- finite_full_sig_iff. assumption. assert (h31:Included (Singleton x) (ba_p_ens (Gen_p (Singleton x) h24))). red. intros x' h32. destruct h32. rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. constructor. assert (h28':Finite (Singleton (exist _ _ h22'))). apply Singleton_is_finite. pose (exist (fun C=>Finite C /\ Included C E) _ (conj h28' h23)) as sp. pose (existT (fun Cp:Bool_Alg_p T=>btp Cp->bt A) _ (psi sp)) as gp. specialize (h17b gp). assert (h32:Inn fam gp). unfold fam. apply Im_intro with (exist (fun C=>Finite C /\ Included C E) _ (conj h28' h23)). constructor. reflexivity. specialize (h17b h32). pose proof h17b as h17. destruct h17 as [h33 h34 h35 h36]. red in h36. unfold gp in h36. simpl in h36. unfold psi in h36. destruct h36 as [h36 h37]. assert (he: Singleton x = im_proj1_sig (Singleton (exist (Inn (A_p T (Bc_p T Bp))) x h22'))). apply Extensionality_Ensembles. red; split; red; intros a ha. destruct ha. apply Im_intro with (exist _ _ h22'). constructor. reflexivity. destruct ha as [a ha q hb]. rewrite hb. clear hb q. destruct a as [a ha']. simpl. inversion ha. constructor. pose proof (gen_ens_p_functional (Singleton x) (im_proj1_sig (Singleton (exist _ _ h22'))) h24 (h0 sp) he) as hf. assert (hin:forall x:sig_set (Gen_Ens_p _ h24), Inn (Gen_Ens_p (im_proj1_sig (Singleton (exist _ _ h22'))) (h0 sp)) (proj1_sig x)). intro x'. destruct x' as [x' hg]. simpl. rewrite <- hf. assumption. assert (hin':forall x:sig_set (Gen_Ens_p _ h24), Inn (FamilyUnion (Im (fam_fun_ba_domains_p1 T fam) (ba_p_ens (T:=T)))) (proj1_sig x)). intro x'. destruct x' as [x' hx]. simpl. apply family_union_intro with (ba_p_ens ((Gen_p (Singleton x) h24))). apply Im_intro with (Gen_p (Singleton x) h24). unfold fam_fun_ba_domains_p1, fam. rewrite im_im. simpl. apply Im_intro with sp. constructor. unfold B_F, sp. simpl. apply gen_p_functional. assumption. reflexivity. rewrite ba_p_ens_gen_p_eq. assumption. assert (h37':forall x0:sig_set (Gen_Ens_p _ h24), proj1_sig (constructive_definite_description (fun h : btp (B_F sp) -> bt A => homomorphism_p1 T h /\ extends_sig1 h (g_F sp)) (finite_homo_extension_criterion_if_p1 (h4 sp) (h3 sp) (g_F sp) (h10 sp))) (exist _ _ (hin x0)) = f (exist _ _ (hin' x0))). intro x'. rewrite h37. f_equal. apply proj1_sig_injective. simpl. reflexivity. assert (hh:Inn (Gen_Ens_p (Singleton x) h24) x). apply gen_ens_includes_p. constructor. assert (h39:f (exist _ _ (hin' (exist _ _ hh))) = f (transfer_r h20 (exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h22'))). f_equal. apply proj1_sig_injective. simpl. unfold btp, directed_ba_p. simpl. unfold Btype_p. simpl. pose proof (f_equal (@ba_p_ens _) h19) as h19'. pose proof (transfer_r_sig_set_eq _ _ h19' h20 (exist _ _ h22')) as h38. rewrite h38 at 1. simpl. reflexivity. rewrite <- h39 at 1. simpl. pose proof (h37 (exist _ _ (hin (exist _ _ hh)))) as h40. simpl in h40. assert (h41:h36 _ (hin (exist _ _ hh)) = hin' (exist _ _ hh)). apply proof_irrelevance. simpl in h41. rewrite h41 in h40. rewrite <- h40. destruct constructive_definite_description as [k h42]. simpl. pose proof h42 as h42'. destruct h42' as [h42a h42b]. red in h42b. unfold finc_bp in h42b. assert (h43:Inn (ba_p_ens (B_F sp)) x). unfold ba_p_ens, B_F. simpl. apply gen_ens_includes_p. apply Im_intro with (exist _ _ h22'). constructor. simpl. reflexivity. assert (h44:Inn (im_proj2_sig (im_proj1_sig (proj1_sig sp)) (h0' sp)) (exist _ _ h43)). unfold im_proj1_sig, im_proj2_sig. rewrite im_full_sig_im_eq. simpl. apply Im_intro with (exist _ _ (In_singleton _ (exist _ _ h22'))). constructor. simpl. apply proj1_sig_injective; auto. simpl in h40. specialize (h42b (exist _ _ h44)). unfold g_F in h42b. simpl in h42b. assert (h45:g (exist _ _ (h9 sp (exist _ _ h44))) = g (exist _ _ h22)). f_equal. apply proj1_sig_injective. simpl. apply proj1_sig_injective. simpl. reflexivity. simpl in h45. rewrite h42b in h45 at 1. rewrite <- h45 at 1. f_equal. apply proj1_sig_injective; auto. (* -> *) intro h2. destruct h2 as [f [h2 h3]]. red in h3. intros F h4 h5 a h6. assert (h7:f (el_prod_p _ a) = 0). rewrite h6. apply homo_zero_p1. assumption. unfold el_prod_p in h7. unfold el_prod_compose_p1. rewrite <- h7 at 1. rewrite homo_times_set_p1; auto. apply times_set_functional. rewrite im_im. apply im_ext_in. intros x h8. specialize (h3 (exist _ _(h5 _ h8))). simpl in h3. unfold eps. assert (h9: f (if a |-> x then x else %- x) = (if a |-> x then (f x) else - (f x))). destruct (a|->x); auto. rewrite homo_comp_p1; auto. rewrite h9 at 1. rewrite <- h3 at 1. destruct (a |-> x). unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h10 | h11]. f_equal. apply proj1_sig_injective. simpl. reflexivity. contradict h11. apply h5; auto. rewrite <- h3. unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h10 | h11]. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. contradict h11. apply h5; auto. Qed. (*This is the same as above but utilizes unique existence.*) Corollary homo_extension_criterion_unq_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)}, ba_p_ens Bp = Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E) -> forall (g:sig_set E->bt A), (exists! h:(btp Bp)->(bt A), homomorphism_p1 _ h /\ extends_sig1 h g) <-> (forall F : Ensemble (btp Bp), Finite F -> Included F E -> forall a : Fin_map F signe mns, el_prod_p _ a = %0 -> el_prod_compose_p1 T (sig_fun_app g 0) a = 0). intros T Bp A E h1 g. pose proof (homo_extension_criterion_p1 h1 g) as h2. split. intros h3 F h4 h5 a h6. assert (h7: exists h : btp Bp -> bt A, homomorphism_p1 T h /\ extends_sig1 h g). destruct h3 as [h h3]. red in h3. destruct h3 as [h3a h3b]. exists h; auto. rewrite h2 in h7. apply h7; auto. intro h3. rewrite <- h2 in h3. destruct h3 as [h h4]. exists h. red. split; auto. intros h' h5. destruct h4 as [h4a h4b], h5 as [h5a h5b]. eapply gen_ens_determines_homo_p1; auto. apply h1. rewrite <- im_proj2_sig_undoes_im_proj1_sig'. red in h4b, h5b. red. intros x h6. specialize (h4b (exist _ _ h6)). specialize (h5b (exist _ _ h6)). simpl in h4b, h5b. congruence. Qed. Lemma homo_extension_compat : forall {A B:Bool_Alg} {E:Ensemble (bt B)} (g:sig_set E->bt A) (f:bt B->bt A), homomorphism f -> extends_sig1 f g -> forall (F:Ensemble (bt B)) (pffin:Finite F), Included F E -> forall (X:Ensemble (Fin_map F signe mns)) (pf:Included X (non_zero_el_prod_maps _ F)) (pf':Included X (non_zero_el_prod_compose_maps _ (sig_fun_app g 0) F)), f (plus_subset_non_zero_el_prod_maps _ _ pffin _ pf) = plus_subset_non_zero_el_prod_compose_maps _ _ pffin _ _ pf'. intros A B E g f h1 h2 F h3 h0 X h4 h5. unfold plus_subset_non_zero_el_prod_maps, plus_subset_non_zero_el_prod_compose_maps. rewrite homo_plus_set. apply plus_set_functional. rewrite im_im. apply im_ext_in. intros a h6. unfold el_prod, el_prod_compose. rewrite homo_times_set. apply times_set_functional. rewrite im_im. apply im_ext_in. intros x h7. assert (h8:f (eps x (a |-> x)) = eps (f x) (a |-> x)). destruct (a |-> x). simpl. reflexivity. simpl. rewrite homo_comp; auto. rewrite h8 at 1. red in h2. specialize (h2 (exist _ _ (h0 _ h7))). simpl in h2. rewrite <- h2. f_equal. unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h9 | h10]. f_equal. apply proj1_sig_injective; auto. contradict h10. apply h0; auto. assumption. assumption. Qed. Lemma finite_mono_extension_criterion_if : forall {A B:Bool_Alg} {E:Ensemble (bt B)}, Finite E -> ba_ens B = Gen_Ens E -> forall (g:sig_set E->bt A), (forall F:Ensemble (bt B), Finite F -> Included F E -> forall a:Fin_map F signe mns, (el_prod _ a = 0 <-> el_prod_compose _ (sig_fun_app g 0) a = 0)) -> exists! h:(bt B)->(bt A), monomorphism h /\ extends_sig1 h g. intros A B E h1 h2 g h3. pose proof (gen_e_eq_im_c). pose proof (gen_e_eq_c _ E h1) as h4. pose proof (gen_e_eq_im_c A (sig_fun_app g 0) E h1) as h5. assert (h6:forall X:Ensemble (Fin_map E signe mns), Included X (non_zero_el_prod_maps _ E) -> Included X (non_zero_el_prod_compose_maps _ (sig_fun_app g 0) E)). intros X h7. red in h7. red. intros a h8. specialize (h7 _ h8). destruct h7 as [h7]. constructor. specialize (h3 E h1 (inclusion_reflexive _ ) a). rewrite <- h3. assumption. assert (h3':forall F : Ensemble (bt B), Finite F -> Included F E -> forall a : Fin_map F signe mns, el_prod B a = 0 -> el_prod_compose A (sig_fun_app g 0) a = 0). intros F h3' h8 a h9. rewrite h3 in h9; auto. pose proof (finite_homo_extension_criterion_if h1 h2 g h3') as h7. destruct h7 as [f h7]. red in h7. destruct h7 as [h7 h8]. destruct h7 as [h7a h7b]. exists f. red. split. split. constructor. assumption. red. intros x y h9. rewrite <- h2 in h4. assert (h10:Inn (ba_ens B) x). constructor. assert (h11:Inn (ba_ens B) y). constructor. rewrite h4 in h10, h11. destruct h10 as [h10], h11 as [h11]. destruct h10 as [X h10], h11 as [Y h11]. destruct h10 as [h10a h10b], h11 as [h11a h11b]. subst. pose proof (homo_extension_compat g f h7a h7b E h1 (inclusion_reflexive _) X h10a (h6 _ h10a)) as h13. pose proof (homo_extension_compat g f h7a h7b E h1 (inclusion_reflexive _) Y h11a (h6 _ h11a)) as h14. rewrite h13, h14 in h9. apply plus_subset_non_zero_el_prod_compose_maps_inj in h9. subst. assert (h10a = h11a). apply proof_irrelevance. subst. reflexivity. assumption. intros f' h9. assert (h10:homomorphism f' /\ extends_sig1 f' g). destruct h9 as [h9a h9b]. destruct h9a. split; auto. apply h8; auto. Qed. Lemma finite_mono_extension_criterion_if_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)}, Finite E -> ba_p_ens Bp = Gen_Ens_p _ (incl_im_proj1_sig_ba_p_ens E) -> forall (g:sig_set E->bt A), (forall F:Ensemble (btp Bp), Finite F -> Included F E -> forall a:Fin_map F signe mns, el_prod_p _ a = %0 <-> el_prod_compose_p1 _ (sig_fun_app g 0) a = 0) -> exists! h:(btp Bp)->(bt A), monomorphism_p1 _ h /\ extends_sig1 h g. intros T Bp A E h1 h2 g h3. rewrite ba_p_ens_eq, gen_ens_p_eq in h2. apply im_proj1_sig_injective in h2. rewrite <- im_proj2_sig_undoes_im_proj1_sig' in h2 at 1. assert (h4: (forall F : Ensemble (bt (ba_conv Bp)), Finite F -> Included F (ba_conv_set E) -> forall a : Fin_map F signe mns, el_prod (ba_conv Bp) a = 0 <-> el_prod_compose A (sig_fun_app g 0) a (B:=ba_conv Bp) = 0)). intros F h4 h5 h6. specialize (h3 F h4 h5 h6). rewrite el_prod_p_eq, el_prod_compose_p1_eq in h3. assumption. pose proof (finite_mono_extension_criterion_if (E:=ba_conv_set E) h1 h2 g h4) as h5. destruct h5 as [f h5]. red in h5. destruct h5 as [h5a h5b]. destruct h5a as [h5l h5r]. unfold bt, ba_conv in f. simpl in f. exists f. red. split. split. rewrite monomorphism_p1_iff. assumption. assumption. intros f' h6. destruct h6 as [h6 h7]. rewrite monomorphism_p1_iff in h6. apply h5b. split; auto. Qed. Corollary mono_extension_criterion_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)}, ba_p_ens Bp = Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E) -> forall (g:sig_set E->bt A), (exists h:(btp Bp)->(bt A), monomorphism_p1 _ h /\ extends_sig1 h g) <-> (forall F : Ensemble (btp Bp), Finite F -> Included F E -> forall a : Fin_map F signe mns, el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0). intros T Bp A E h1 g. split. Focus 2. intro h2. assert (h3:forall F:finc E, Included (proj1_sig F) E). intro F. destruct F as [F h3]. simpl. destruct h3; auto. assert (h3':forall F:finc E, Included (im_proj1_sig (proj1_sig F)) (ba_p_ens Bp)). intro F. intros x h4. destruct h4 as [x h4]. subst. apply proj2_sig. assert (h4:forall F:finc E, ba_p_ens (Gen_p _ (h3' F)) = Gen_Ens_p _ (h3' F)). intro F. rewrite ba_p_ens_gen_p_eq. reflexivity. assert (h5:forall F:finc E, Included (im_proj1_sig (proj1_sig F)) (ba_p_ens (Gen_p _ (h3' F)))). intro F. destruct F as [F h6]. simpl. rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. assert (h6:forall F:finc E, Finite (im_proj2_sig (im_proj1_sig (proj1_sig F)) (h5 F))). intro F. apply finite_image. rewrite <- finite_full_sig_iff. destruct F as [F h6]. destruct h6; auto. simpl. apply finite_image; auto. assert (h4' : forall F:finc E, ba_p_ens (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F)) = Gen_Ens_p (im_proj1_sig (im_proj2_sig (im_proj1_sig (proj1_sig F)) (h5 F))) (incl_im_proj1_sig_ba_p_ens (im_proj2_sig (im_proj1_sig (proj1_sig F)) (h5 F)))). intro F. assert (h7 : im_proj1_sig (im_proj2_sig (im_proj1_sig (proj1_sig F)) (h5 F)) = im_proj1_sig (proj1_sig F)). unfold im_proj1_sig, im_proj2_sig. rewrite im_im. simpl. rewrite <- im_full_sig_proj1_sig. reflexivity. pose proof (subsetT_eq_compat _ (fun S=>Included S (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F)))) _ _ (incl_im_proj1_sig_ba_p_ens (im_proj2_sig (im_proj1_sig (proj1_sig F)) (h5 F))) (h5 F) h7) as h8. unfold bt, ba_p_ens in h8. unfold bt, ba_p_ens. dependent rewrite -> h8. simpl. symmetry. apply gen_ens_p_subalg_of_p. apply gen_p_subalg_of_p_compat. assert (h8:forall (F:finc E) (x:sig_set (im_proj2_sig (im_proj1_sig (proj1_sig F)) (h5 F))), Inn (ba_p_ens Bp) (proj1_sig (proj1_sig x))). intros F x. destruct x as [x h9]. simpl. destruct x as [x h10]. simpl. inversion h9 as [x' h11 q h12]. subst. apply exist_injective in h12. subst. destruct x' as [x' h12]. simpl. destruct h12 as [x' h12]. subst. apply proj2_sig. assert (h9:forall (F:finc E) (x:sig_set (im_proj2_sig (im_proj1_sig (proj1_sig F)) (h5 F))), Inn E (exist _ _ (h8 F x))). intros F x. destruct x as [x h10]. simpl. destruct x as [x h11]. simpl. inversion h10 as [q h12 ? h14]. subst. apply exist_injective in h14. subst. destruct q as [q h13]. simpl. destruct h13 as [q h13]. subst. clear h12. simpl in h11, h10. apply (h3 F) in h13. assert (hq:q = exist _ _ (h8 F (exist (fun x : {x : T | Inn (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F))) x} => Inn (im_proj2_sig (im_proj1_sig (proj1_sig F)) (h5 F)) x) (exist (fun x : T => Inn (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F))) x) (proj1_sig q) h11) h10))). apply proj1_sig_injective. simpl. reflexivity. rewrite <- hq at 1. assumption. assert (h18:forall (F:finc E), Included (im_proj1_sig (proj1_sig F)) (ba_p_ens Bp)). intros. simpl. red. intros x h19. destruct h19 as [x h19]. subst. apply proj2_sig. pose (fun F:finc E => (fun x:sig_set (im_proj2_sig (im_proj1_sig (proj1_sig F)) (h5 F)) => g (exist _ _ (h9 F x)))) as g'. assert (h10: forall F : finc E, (forall F0 : Ensemble (btp (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F))), Finite F0 -> Included F0 (im_proj2_sig (im_proj1_sig (proj1_sig F)) (h5 F)) -> forall a : Fin_map F0 signe mns, el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app (g' F) 0) a = 0)). intro F. intros D h10 h11 a. assert (hd:forall x: (btp (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F))), Inn (ba_p_ens Bp) (proj1_sig x)). intro x. destruct x as [x hx]. simpl. pose proof (ba_p_ens_gen_p_eq (im_proj1_sig (proj1_sig F)) (h3' F)) as hg. unfold ba_p_ens in hg. rewrite hg in hx. apply incl_gen_ens_p in hx. assumption. pose (Im D (fun x => (exist _ _ (hd x)))) as D'. specialize (h2 D'). assert (h12:Included D' E). red. intros x h13. destruct h13 as [x h13]. subst. apply h11 in h13. destruct h13 as [x h13]. subst. simpl. clear h13. destruct x as [x h13]. simpl. destruct h13 as [x h13]. subst. pose proof (h3 F _ h13) as h14. assert (h15:x = exist _ _ (hd (exist (Inn (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F)))) (proj1_sig x) (h5 F (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (proj1_sig F) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h13 (proj1_sig x) eq_refl))))). apply proj1_sig_injective; auto. rewrite <- h15 at 1; auto. assert (h13:Finite D'). apply finite_image. assumption. assert (h14:forall x:sig_set D', Inn (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F))) (proj1_sig (proj1_sig x))). intro x. destruct x as [x h14]. simpl. rewrite ba_p_ens_gen_p_eq. destruct h14 as [x h14]. subst. apply gen_ens_includes_p. pose proof (h11 _ h14) as h15. destruct h15 as [x h15]. subst. simpl. apply proj2_sig. assert (h15:Included (Im (full_sig D') (fun x : sig_set D' => a |-> exist (Inn (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F)))) (proj1_sig (proj1_sig x)) (h14 x))) signe). red. intros x h15. destruct x. left. right. pose (sig_fun_to_fin_map_ran ba_eq_dec_p (fun x=>a |-> (exist _ _ (h14 x))) h13 mns signe signe_finite h15) as a'. specialize (h2 h13 h12 a'). assert (h21: (forall x : btp (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F)), Inn D x -> a |-> x = a' |-> (exist _ _ (hd x)))). intros x h22. unfold a'. rewrite sig_fun_to_fin_map_ran_compat at 1. simpl. f_equal. apply proj1_sig_injective. simpl. reflexivity. split. intro h16. assert (hdeq:im_proj1_sig D = im_proj1_sig D'). unfold im_proj1_sig, D'. rewrite im_im. apply im_ext_in. intros; auto. pose proof (f_equal (@proj1_sig _ _) h16) as h17. simpl in h17. pose proof (gen_p_subalg_of_p_compat _ (h3' F)) as h19. assert (h11':Included D' E). red. intros x hi. destruct hi as [x hi]. subst. apply h12; auto. apply Im_intro with x; auto. assert (h23: forall (x : btp (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F))), Inn D x -> sig_fun_app (g' F) 0 x = sig_fun_app g 0 (exist _ _ (hd x))). intros x h24. unfold g', sig_fun_app. destruct classic_dec as [h25 | h26]. destruct classic_dec as [h27 | h28]. f_equal. apply proj1_sig_injective; auto. simpl. apply proj1_sig_injective; auto. contradict h28. destruct h25 as [x h25]. subst. simpl. clear h25. destruct x as [x h25]. simpl. destruct h25 as [x h25]. subst. pose proof (proj2_sig F) as h25'. simpl in h25'. destruct h25' as [h25a h25b]. pose proof (h25b _ h25) as h26. assert (h27:x = exist _ _ (hd (exist (Inn (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F)))) (proj1_sig x) (h5 F (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (proj1_sig F) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h25 (proj1_sig x) eq_refl))))). apply proj1_sig_injective; auto. rewrite <- h27 at 1. assumption. contradict h26. apply h11; auto. pose proof (el_prod_compose_compat_fin_map_eq_base A h19 g (sig_fun_app (g' F) 0) D D' hdeq h11' hd h23 a a' h21) as h24. rewrite <- h24 at 1. rewrite <- h2. pose proof (el_prod_compat_fin_map_eq_base h19 D D' hdeq a a' hd h21) as h20. apply proj1_sig_injective. rewrite <- h20 at 1. assumption. intro h16. pose proof (gen_p_subalg_of_p_compat _ (h3' F)) as h19. assert (h11':Included D' E). red. intros x hi. destruct hi as [x hi]. subst. apply h12; auto. apply Im_intro with x; auto. assert (h23: forall (x : btp (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F))), Inn D x -> sig_fun_app (g' F) 0 x = sig_fun_app g 0 (exist _ _ (hd x))). intros x h24. unfold g', sig_fun_app. destruct classic_dec as [h25 | h26]. destruct classic_dec as [h27 | h28]. f_equal. apply proj1_sig_injective; auto. simpl. apply proj1_sig_injective; auto. contradict h28. destruct h25 as [x h25]. subst. simpl. clear h25. destruct x as [x h25]. simpl. destruct h25 as [x h25]. subst. pose proof (proj2_sig F) as h25'. simpl in h25'. destruct h25' as [h25a h25b]. pose proof (h25b _ h25) as h26. assert (h27:x = exist _ _ (hd (exist (Inn (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F)))) (proj1_sig x) (h5 F (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (proj1_sig F) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h25 (proj1_sig x) eq_refl))))). apply proj1_sig_injective; auto. rewrite <- h27 at 1. assumption. contradict h26. apply h11; auto. assert (hdeq:im_proj1_sig D = im_proj1_sig D'). unfold im_proj1_sig, D'. rewrite im_im. apply im_ext_in. intros; auto. pose proof (el_prod_compose_compat_fin_map_eq_base A h19 g (sig_fun_app (g' F) 0) D D' hdeq h11' hd h23 a a' h21) as h24. rewrite <- h24 in h16. pose proof (el_prod_compat_fin_map_eq_base h19 D D' hdeq a a' hd h21) as h20. rewrite <- h2 in h16. apply proj1_sig_injective. rewrite h20 at 1. rewrite h16. reflexivity. pose (fun F:finc E => finite_mono_extension_criterion_if_p1 (A:=A) (E:=im_proj2_sig (im_proj1_sig (proj1_sig F)) (h5 F)) (h6 F) (h4' F) (g' F) (h10 F)) as monos. pose (fun F => (proj1_sig (constructive_definite_description _ (monos F)))) as psi. pose (Im (Full_set (finc E)) (fun F => existT (fun C:(Bool_Alg_p T) => (btp C->bt A)) (Gen_p (im_proj1_sig (proj1_sig F)) (h3' F)) (psi F))) as fam. assert (h11:directed_fam_homo_p1 _ fam). red. split. red. intros f h11. destruct h11 as [F h11]. subst. simpl. unfold psi. destruct constructive_definite_description as [f h12]. simpl. destruct h12 as [h12 h13]. destruct h12 as [h12 h12']. assumption. intros F G h11 h12. destruct h11 as [F h11], h12 as [G h12]. subst. destruct F as [F h13], G as [G h14]. destruct h13 as [h13a h13b], h14 as [h14a h14b]. simpl. pose (Union F G) as H'. assert (h15:Finite H'). apply Union_preserves_Finite; auto. assert (h16:Included H' E). unfold H'. auto with sets. pose (exist (fun S=>Finite S /\ Included S E) _ (conj h13a h13b)) as F'. pose (exist (fun S=>Finite S /\ Included S E) _ (conj h15 h16)) as H''. assert (h17:Included (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig F')) (h3' F'))) (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig H'')) (h3' H'')))). do 2 rewrite ba_p_ens_gen_p_eq. apply gen_ens_preserves_inclusion_p. unfold F', H''. simpl. unfold H'. rewrite incl_im_proj1_sig_iff. auto with sets. exists (existT (fun C:(Bool_Alg_p T) => (btp C->bt A)) (Gen_p (im_proj1_sig (proj1_sig H'')) (h3' H'')) (psi H'')). split. unfold fam. apply Im_intro with H''. constructor. reflexivity. simpl. red. split. split. red. exists h17. simpl. assert (h18': alg_closed_p (ba_p_ens (Gen_p (im_proj1_sig F) (h3' (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h13a h13b))))) h17). pose proof (ba_p_ens_gen_p_eq (im_proj1_sig F) (h3' (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h13a h13b)))) as h18'. assert (h19:Included (Gen_Ens_p (im_proj1_sig F) (h3' (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h13a h13b)))) (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig H'')) (h3' H'')))). rewrite <- h18'. assumption. pose proof (subsetT_eq_compat _ (fun S => Included S (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig H'')) (h3' H'')))) _ _ h17 h19 h18') as h20. unfold F' in h20. simpl in h20. dependent rewrite -> h20. pose proof (gen_ens_includes_p _ (h3' F')) as h22. unfold F' in h22. simpl in h22. assert (h23:Included (im_proj1_sig F) (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig H'')) (h3' H'')))). auto with sets. pose proof (closed_gen_ens_p _ h23) as h24. assert (h25:subalg_of_p (Gen_p (im_proj1_sig (proj1_sig H'')) (h3' H'')) Bp). apply gen_p_subalg_of_p_compat. pose proof (gen_ens_p_subalg_of_p h25 _ h23 (h3' F')) as h26. pose proof (subsetT_eq_compat _ (fun S => Included S (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig H'')) (h3' H'')))) _ _ (incl_gen_ens_p (im_proj1_sig F) h23) h19 h26) as h27. dependent rewrite -> h27 in h24. assumption. exists h18'. assert (h19:Included (im_proj1_sig F) (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig H'')) (h3' H'')))). rewrite ba_p_ens_gen_p_eq. assert (h20:Included F H'). unfold H'. auto with sets. pose proof (gen_ens_includes_p _ (h3' H'')) as h21. unfold H'' in h21. simpl in h21. red. intros a h22. destruct h22 as [a h22]. subst. apply h20 in h22. assert (hh:Inn (im_proj1_sig H') (proj1_sig a)). apply Im_intro with a; auto. apply h21 in hh. unfold H''. simpl. assumption. assert (h20: Subalg_p (Gen_p (im_proj1_sig H') (h3' H'')) (ba_p_ens (Gen_p (im_proj1_sig F) (h3' (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h13a h13b))))) h17 h18' = Gen_p (im_proj1_sig F) h19). apply subalg_functional_p. rewrite ba_p_ens_gen_p_eq at 1. symmetry. apply gen_ens_p_subalg_of_p. apply gen_p_subalg_of_p_compat. rewrite h20 at 1. symmetry. apply gen_p_subalg_of_p. apply gen_p_subalg_of_p_compat. unfold psi. destruct constructive_definite_description as [f h18']. simpl. destruct h18' as [h18' h19]. destruct h18' as [h18' h18'']. assumption. unfold psi. destruct constructive_definite_description as [f h18']. simpl. destruct h18' as [h18' h18b]. destruct h18' as [h18' h18'']. assumption. unfold psi. destruct constructive_definite_description as [f h18']. destruct constructive_definite_description as [k h19]. simpl. destruct h18' as [h18a h18b]. destruct h19 as [h19a h19b]. red in h18b. red in h19b. pose (restriction_sig f _ h17) as f'. assert (hinc:Included (im_proj1_sig F) (im_proj1_sig H')). rewrite incl_im_proj1_sig_iff. unfold H'. auto with sets. assert (h20:f' = k). pose proof (subalg_of_p_gen_p_gen_p _ _ _ hinc (h3' F') (h3' H'')) as h22. assert (h23:Included (ba_p_ens (Gen_p (im_proj1_sig F) (h3' F'))) (ba_p_ens (Gen_p (im_proj1_sig H') (h3' H'')))). do 2 rewrite ba_p_ens_gen_p_eq. assumption. destruct h18a as [h18a h18a']. pose proof (homomorphism_p1_restriction_sig _ _ f h18a _ h23 h22) as h24. assert (h25:Included (im_proj1_sig F) (ba_p_ens (Gen_p (im_proj1_sig F) (h3' F')))). rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. assert (h26: ba_p_ens (Gen_p (im_proj1_sig F) (h3' F')) = Gen_Ens_p (im_proj1_sig F) h25). rewrite ba_p_ens_gen_p_eq. assert (h26:subalg_of_p (Gen_p (im_proj1_sig F) (h3' F')) Bp). apply gen_p_subalg_of_p_compat. rewrite (gen_ens_p_subalg_of_p h26 (im_proj1_sig F) h25 (h3' F')). reflexivity. assert (h27: agree_on (restriction_sig f (ba_p_ens (Gen_p (im_proj1_sig F) (h3' F'))) h23) k (im_proj2_sig (im_proj1_sig F) h25)). red. intros x h27. pose (fun F:(finc E) => im_proj2_sig _ (h5 F)) as finc_bp. assert (h29:Inn (finc_bp (exist _ _ (conj h13a h13b))) x). destruct h27 as [x h27]. subst. unfold finc_bp. simpl. apply Im_intro with x. assumption. apply proj1_sig_injective. simpl. reflexivity. specialize (h19b (exist _ _ h29)). simpl in h19b. rewrite <- h19b. destruct h27 as [x h27]. subst. destruct x as [x h30]. simpl. assert (h31:Inn (ba_p_ens (Gen_p (im_proj1_sig H') (h3' H''))) x). apply h23. apply h25. assumption. assert (h32:Inn (finc_bp H'') (exist _ _ h31)). unfold finc_bp. apply Im_intro with (exist _ _ (hinc _ h30)). constructor. apply proj1_sig_injective. simpl. reflexivity. unfold restriction_sig. simpl. specialize (h18b (exist _ _ h32)). simpl in h18b. assert (h33: f (exist (Inn (Gen_Ens_p (im_proj1_sig H') (h3' H''))) x (h23 x (h25 x h30))) = f (exist (Inn (ba_p_ens (Gen_p (im_proj1_sig H') (h3' H'')))) x h31)). f_equal. apply proj1_sig_injective. simpl. reflexivity. rewrite h33. rewrite <- h18b. unfold g'. f_equal. apply proj1_sig_injective; simpl. apply proj1_sig_injective; auto. destruct h19a as [h19a h19a']. pose proof (gen_ens_determines_homo_p1 _ _ _ h24 h19a (im_proj1_sig F) h25 h26 h27) as h28. rewrite <- h28. unfold f'. f_equal. apply proof_irrelevance. subst. subst. red. assert (h21:Included (Gen_Ens_p (im_proj1_sig F) (h3' (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) F (conj h13a h13b)))) (Gen_Ens_p (im_proj1_sig H') (h3' H''))). apply gen_ens_preserves_inclusion_p. unfold H'. auto with sets. exists h21. intro x. unfold restriction_sig. f_equal. apply proj1_sig_injective. simpl. reflexivity. pose (exist (fun S=>Finite S /\ Included S E) _ (conj h14a h14b)) as G'. pose (fun F:(finc E) => Gen_p (im_proj1_sig (proj1_sig F)) (h3' F)) as B_F. assert (h17':Included (ba_p_ens (B_F G')) (ba_p_ens (B_F H''))). unfold B_F. do 2 rewrite ba_p_ens_gen_p_eq. apply gen_ens_preserves_inclusion_p. unfold G', H''. simpl. unfold H'. rewrite incl_im_proj1_sig_iff. auto with sets. unfold psi. destruct constructive_definite_description as [f h18']. destruct constructive_definite_description as [k h19]. simpl. destruct h18' as [h18a h18b]. destruct h19 as [h19a h19b]. red in h18b. red in h19b. pose (restriction_sig f _ h17') as f'. assert (hinc:Included (im_proj1_sig G) (im_proj1_sig H')). unfold H'. rewrite incl_im_proj1_sig_iff. auto with sets. assert (h20:f' = k). pose proof (subalg_of_p_gen_p_gen_p _ _ _ hinc (h3' G') (h3' H'')) as h22. assert (h23:Included (ba_p_ens (Gen_p (im_proj1_sig G) (h3' G'))) (ba_p_ens (Gen_p (im_proj1_sig H') (h3' H'')))). do 2 rewrite ba_p_ens_gen_p_eq. assumption. destruct h18a as [h18a h18a']. pose proof (homomorphism_p1_restriction_sig _ _ f h18a _ h23 h22) as h24. assert (h25:Included (im_proj1_sig G) (ba_p_ens (Gen_p (im_proj1_sig G) (h3' G')))). rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. assert (h26: ba_p_ens (Gen_p (im_proj1_sig G) (h3' G')) = Gen_Ens_p (im_proj1_sig G) h25). rewrite ba_p_ens_gen_p_eq. assert (h26:subalg_of_p (Gen_p (im_proj1_sig G) (h3' G')) Bp). apply gen_p_subalg_of_p_compat. rewrite (gen_ens_p_subalg_of_p h26 (im_proj1_sig G) h25 (h3' G')). reflexivity. assert (h27: agree_on (restriction_sig f (ba_p_ens (Gen_p (im_proj1_sig G) (h3' G'))) h23) k (im_proj2_sig (im_proj1_sig G) h25)). red. intros x h27. pose (fun F:(finc E) => im_proj2_sig _ (h5 F)) as finc_bp. assert (h29:Inn (finc_bp (exist _ _ (conj h14a h14b))) x). destruct h27 as [x h27]. subst. unfold finc_bp. simpl. apply Im_intro with x. assumption. apply proj1_sig_injective. simpl. reflexivity. specialize (h19b (exist _ _ h29)). simpl in h19b. rewrite <- h19b. destruct h27 as [x h27]. subst. destruct x as [x h30]. simpl. assert (h31:Inn (ba_p_ens (Gen_p (im_proj1_sig H') (h3' H''))) x). apply h23. apply h25. assumption. assert (h32:Inn (finc_bp H'') (exist _ _ h31)). unfold finc_bp. apply Im_intro with (exist _ _ (hinc _ h30)). constructor. apply proj1_sig_injective. simpl. reflexivity. unfold restriction_sig. simpl. specialize (h18b (exist _ _ h32)). simpl in h18b. assert (h33: f (exist (Inn (Gen_Ens_p (im_proj1_sig H') (h3' H''))) x (h23 x (h25 x h30))) = f (exist (Inn (ba_p_ens (Gen_p (im_proj1_sig H') (h3' H'')))) x h31)). f_equal. apply proj1_sig_injective. simpl. reflexivity. rewrite h33. rewrite <- h18b. unfold g'. f_equal. apply proj1_sig_injective. simpl. apply proj1_sig_injective; auto. unfold f'. destruct h19a as [h19a h19a']. pose proof (gen_ens_determines_homo_p1 _ _ _ h24 h19a (im_proj1_sig G) h25 h26 h27) as h28. rewrite <- h28. f_equal. apply proof_irrelevance. subst. subst. constructor. unfold B_F. simpl. apply subalg_of_p_gen_p_gen_p. unfold H'. auto with sets. destruct h18a as [h18a h18a']. assumption. apply homomorphism_p1_restriction_sig. destruct h18a as [h18a h18a']. assumption. apply subalg_of_p_gen_p_gen_p. unfold G', H''. simpl. unfold H'. auto with sets. red. assert (h21: Included (A_p T (Bc_p T (Gen_p (im_proj1_sig G) (h3' (exist (fun F0 : Ensemble (btp Bp) => Finite F0 /\ Included F0 E) G (conj h14a h14b)))))) (A_p T (Bc_p T (Gen_p (im_proj1_sig H') (h3' H''))))). simpl. apply gen_ens_preserves_inclusion_p. unfold H'. auto with sets. exists h21. intro x. unfold restriction_sig. f_equal. apply proj1_sig_injective. simpl. reflexivity. assert (h12:Inhabited (Full_set (finc E))). assert (h13:Finite (Empty_set (btp Bp))). auto with sets. assert (h14:Included (Empty_set (btp Bp)) E). auto with sets. unfold finc. apply Inhabited_intro with (exist (fun F=>Finite F /\ Included F E) _ (conj h13 h14)). constructor. pose (fun F:(finc E) => Gen_p (im_proj1_sig (proj1_sig F)) (h3' F)) as B_F. assert (h13:Inhabited fam). unfold fam. destruct h12 as [K h12]. apply Inhabited_intro with ((fun F : (finc E) => existT (fun C : Bool_Alg_p T => btp C -> bt A) (B_F F) (psi F)) K). apply Im_intro with K. assumption. reflexivity. pose proof (directed_common_extension_p1 _ _ h11 h13) as h14. destruct h14 as [h14 [h15 [f [h17 h18']]]]. assert (h19:directed_ba_p h14 h15 = Bp). assert (h21: forall Dp : Bool_Alg_p T, Inn (fam_fun_ba_domains_p1 T fam) Dp -> subalg_of_p Dp Bp) . intros Dp he. destruct he as [Dp he]. subst. destruct he as [S he]. subst. simpl. apply gen_p_subalg_of_p_compat. pose proof (subalg_of_p_directed_ba_p (Bp:=Bp) _ h14 h15 h21) as h20. assert (h22:ba_p_ens (directed_ba_p h14 h15) = ba_p_ens Bp). apply Extensionality_Ensembles. red; split. apply (incl_ba_p_ens_directed_ba_p h14 h15 Bp h21). rewrite h1 at 1. pose proof (alg_closed_p_ba_p_ens_directed_ba_p h14 h15 _ h21) as hb. pose proof (gen_ens_minimal_p' (im_proj1_sig E) (ba_p_ens (directed_ba_p h14 h15)) (incl_im_proj1_sig_ba_p_ens E) (incl_ba_p_ens_directed_ba_p h14 h15 Bp h21) hb) as ha. apply ha. unfold ba_p_ens, directed_ba_p. simpl. red. intros x h22. assert (hin:Inn (ba_p_ens Bp) x). rewrite h1 at 1. apply gen_ens_includes_p. assumption. assert (h23:Included (Singleton x) (ba_p_ens Bp)). red; intros x' h23; destruct h23; auto. assert (h23':Included (Singleton x) (im_proj1_sig E)). red. intros x' h23'. destruct h23'; auto. assert (h23'':Included (Singleton (exist _ _ hin)) E). red. intros x' h23''. destruct h23''; auto. rewrite <- in_im_proj1_sig_iff; auto. assert (h24:Finite (Singleton (exist _ _ hin))). apply Singleton_is_finite. apply family_union_intro with (Gen_Ens_p (Singleton x) h23). apply Im_intro with (Gen_p (Singleton x) h23). unfold fam, fam_fun_ba_domains_p1. rewrite im_im. simpl. unfold finc. apply Im_intro with (existT (fun S => Finite S /\ Included S E) _ (conj h24 h23'')). constructor. unfold B_F. simpl. apply gen_p_functional. apply Extensionality_Ensembles; red; split; red; intros a hc. destruct hc. apply Im_intro with (exist _ _ hin). constructor. reflexivity. subst. destruct hc as [a' hc ? hd]. subst. destruct hc. simpl. constructor; auto. rewrite ba_p_ens_gen_p_eq. reflexivity. apply gen_ens_includes_p. constructor. apply bc_inj_p. destruct h20 as [h20a [h20b h20c]]. rewrite h20c at 1. unfold Subalg_p. simpl. assert (h24: A_p T (Bc_p' T Bp (ba_p_ens (directed_ba_p h14 h15)) h20a h20b) = A_p T (Bc_p T Bp)). apply Extensionality_Ensembles. red. split. red. intros x h24. unfold bc_sig_set_conv. simpl. eapply incl_ba_p_ens_directed_ba_p. apply h21. apply h24. red. intros x h24. unfold Bc_p'. simpl. rewrite h22. assumption. apply (bconst_ext_p (Bc_p' T Bp (ba_p_ens (directed_ba_p h14 h15)) h20a h20b) _ h24). apply Extensionality_Ensembles. red. split. red. intros x h25. simpl in h25. clear h25. destruct x as [x h25]. unfold ba_p_ens, directed_ba_p in h24. simpl in h24. unfold ba_p_ens, directed_ba_p. simpl. rewrite <- (transfer_r_undoes_transfer (sig_set_eq (FamilyUnion (Im (fam_fun_ba_domains_p1 T fam) (ba_p_ens (T:=T)))) _ h24)). rewrite <- transfer_in_r at 1. rewrite (transfer_sig_set_eq _ _ h24 (sig_set_eq (FamilyUnion (Im (fam_fun_ba_domains_p1 T fam) (ba_p_ens (T:=T)))) _ h24)). simpl. apply in_bs_p. red; intros; constructor. rewrite <- transfer_fun2_r_transfer_dep_r_compat'. rewrite transfer_fun2_r_eq'. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite (transfer_r_sig_set_eq _ _ h24) at 1. simpl. apply proj1_sig_injective. simpl. do 2 rewrite (transfer_sig_set_eq _ _ h24) at 1. unfold Bplus_sub_p; simpl. destruct x as [x h26], y as [y h27]. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. rewrite <- transfer_fun2_r_transfer_dep_r_compat'. rewrite transfer_fun2_r_eq'. apply functional_extensionality. intro x. apply functional_extensionality. intro y. rewrite (transfer_r_sig_set_eq _ _ h24) at 1. simpl. apply proj1_sig_injective. simpl. do 2 rewrite (transfer_sig_set_eq _ _ h24) at 1. unfold Btimes_sub_p; simpl. destruct x as [x h26], y as [y h27]. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. apply proj1_sig_injective; auto. rewrite transfer_dep_r_id_transfer_r_compat. rewrite (transfer_r_sig_set_eq _ _ h24). apply proj1_sig_injective. simpl. reflexivity. rewrite transfer_dep_r_id_transfer_r_compat. rewrite (transfer_r_sig_set_eq _ _ h24). apply proj1_sig_injective. simpl. reflexivity. rewrite <- transfer_fun_r_transfer_dep_r_compat'. rewrite transfer_fun_r_eq'. apply functional_extensionality. intro x. rewrite (transfer_r_sig_set_eq _ _ h24) at 1. simpl. apply proj1_sig_injective. simpl. rewrite (transfer_sig_set_eq _ _ h24) at 1. unfold Bcomp_sub, Bcomp_sub_p. simpl. destruct x as [x h26]. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. assert (h20:btp (directed_ba_p h14 h15) = btp Bp). f_equal; auto. pose (transfer_fun h20 f) as f'. pose proof (common_extension_fam_homo_p1_homo_p1 _ _ _ f fam h13 h17) as h21. exists f'. split. unfold f'. rewrite (ba_p_subst_homomorphism_p1 _ _ _ _ h19 h20) in h21. rewrite homomorphism_p1_iff in h21. unfold ba_conv_fun1 in h21. unfold Btype, ba_conv in h21. simpl in h21. simpl in h21. constructor. rewrite homomorphism_p1_iff. assumption. assert (hinj:inj_fam_homo_p1 T fam ). constructor. red. intros rho hr. destruct hr as [rho hr]. rewrite H. simpl. unfold psi. destruct constructive_definite_description as [rho' ha]. simpl. destruct ha as [ha ha']. destruct ha as [ha1 ha2]. assumption. intros rho hr. destruct hr as [rho hr]. rewrite H. simpl. unfold psi. destruct constructive_definite_description as [rho' ha]. simpl. destruct ha as [ha ha']. destruct ha as [ha1 ha2]. assumption. pose proof (h18' hinj) as hinj'. red in hinj'. red. intros b1 b2 hb. specialize (hinj' (transfer_r h20 b1) (transfer_r h20 b2)). rewrite <- (transfer_undoes_transfer_r h20 b1) in hb at 1. rewrite <- (transfer_undoes_transfer_r h20 b2) in hb at 1. do 2 rewrite transfer_fun_compat in hb. specialize (hinj' hb). pose proof (f_equal (@ba_p_ens T) h19) as h19'. do 2 rewrite (transfer_r_sig_set_eq _ _ h19' h20) in hinj' at 1. apply exist_injective in hinj'. simpl in hinj'. apply proj1_sig_injective; auto. red. intro x. destruct h17 as [h17a h17b]. destruct x as [x h22]. simpl. unfold f'. rewrite transfer_fun_eq. assert (h23:Included (Singleton x) E). red; intros x' h24. destruct h24; auto. assert (h23':Included (Singleton (proj1_sig x)) (im_proj1_sig E)). red. intros a h24. destruct h24. rewrite in_im_proj1_sig_iff. assumption. assert (h24:Included (Singleton (proj1_sig x)) (ba_p_ens Bp)). intros x' h26. destruct h26. apply proj2_sig. assert (h25:Inn (ba_p_ens (Gen_p _ h24)) (proj1_sig x)). rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. constructor. assert (h26:Included (im_proj1_sig E) (ba_p_ens Bp)). red; intros x' ha. destruct ha as [x' hb q hc]. rewrite hc. clear hc. apply proj2_sig. pose proof (subalg_of_p_gen_p_gen_p Bp (Singleton (proj1_sig x)) (im_proj1_sig E) h23' h24 h26) as h27. pose proof h27 as h27'. destruct h27' as [h27a [h27b h27c]]. assert (h28:Finite (Singleton (proj1_sig x))). apply Singleton_is_finite. assert (h28':Finite (Singleton x)). apply Singleton_is_finite. assert (h30:Finite (im_proj2_sig (Singleton (proj1_sig x)) h24)). apply finite_image; auto. rewrite <- finite_full_sig_iff. assumption. assert (h31:Included (Singleton (proj1_sig x)) (ba_p_ens (Gen_p _ h24))). red. intros x' h32. destruct h32. rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. constructor. pose (exist (fun C=>Finite C /\ Included C E) _ (conj h28' h23)) as sp. pose (existT (fun Bp=>(btp Bp)->bt A) (Gen_p _ (h3' sp)) (psi sp)) as gp. specialize (h17b gp). assert (h32:Inn fam gp). unfold fam. apply Im_intro with (exist (fun C=>Finite C /\ Included C E) _ (conj h28' h23)). constructor. reflexivity. specialize (h17b h32). pose proof h17b as h17. destruct h17 as [h33 h34 h35 h36]. red in h36. unfold gp in h36. simpl in h36. unfold psi in h36. destruct h36 as [h36 h37]. assert (h38:Inn (Gen_Ens_p (Singleton (proj1_sig x)) h24) (proj1_sig x)). apply gen_ens_includes_p. constructor. pose proof h24 as h24'. rewrite <- im_proj1_sig_sing in h24'. assert (heq:Gen_Ens_p _ h24 = Gen_Ens_p _ h24'). apply gen_ens_p_functional. rewrite im_proj1_sig_sing. reflexivity. rewrite heq in h38. specialize (h37 (exist _ _ h38)). simpl in h37. destruct constructive_definite_description as [k h39]. simpl in h37. pose proof h39 as h39'. destruct h39' as [h39a h39b]. red in h39b. assert (h40: Inn (ba_p_ens (B_F sp)) (proj1_sig x)). unfold B_F. rewrite ba_p_ens_gen_p_eq. unfold sp. simpl. apply gen_ens_includes_p. rewrite in_im_proj1_sig_iff. constructor. pose (fun F:(finc E) => im_proj2_sig _ (h5 F)) as finc_bp. assert (h41:Inn (finc_bp sp) (exist _ _ h40)). unfold finc_bp, sp. simpl. unfold im_proj2_sig. assert (hin:Inn (im_proj1_sig (Singleton x)) (proj1_sig x)). apply Im_intro with x; auto. constructor. apply Im_intro with (exist _ _ hin). constructor. apply proj1_sig_injective. simpl. reflexivity. specialize (h39b (exist _ _ h41)). unfold g', sp in h39b. simpl in h39b. assert (h42: g (exist (Inn E) (exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h8 (exist (fun C : Ensemble (btp Bp) => Finite C /\ Included C E) (Singleton x) (conj h28' h23)) (exist (Inn (finc_bp (exist (fun C : Ensemble (btp Bp) => Finite C /\ Included C E) (Singleton x) (conj h28' h23)))) (exist (Inn (ba_p_ens (B_F (exist (fun C : Ensemble (btp Bp) => Finite C /\ Included C E) (Singleton x) (conj h28' h23))))) (proj1_sig x) h40) h41))) (h9 (exist (fun C : Ensemble (btp Bp) => Finite C /\ Included C E) (Singleton x) (conj h28' h23)) (exist (Inn (finc_bp (exist (fun C : Ensemble (btp Bp) => Finite C /\ Included C E) (Singleton x) (conj h28' h23)))) (exist (Inn (ba_p_ens (B_F (exist (fun C : Ensemble (btp Bp) => Finite C /\ Included C E) (Singleton x) (conj h28' h23))))) (proj1_sig x) h40) h41))) = g (exist (fun x0 : btp Bp => Inn E x0) x h22)). f_equal. apply proj1_sig_injective. simpl. apply proj1_sig_injective; auto. rewrite <- h42 at 1. unfold bt, finc_bp, B_F in h39b. unfold bt, finc_bp, B_F. rewrite h39b at 1. assert (h43: k (exist (Inn (ba_p_ens (Gen_p (im_proj1_sig (proj1_sig (exist (fun C : Ensemble (btp Bp) => Finite C /\ Included C E) (Singleton x) (conj h28' h23)))) (h3' (exist (fun C : Ensemble (btp Bp) => Finite C /\ Included C E) (Singleton x) (conj h28' h23)))))) (proj1_sig x) h40) = k (exist (Inn (Gen_Ens_p (im_proj1_sig (Singleton x)) h24')) (proj1_sig x) h38)). f_equal. apply proj1_sig_injective. simpl. reflexivity. unfold bt, finc_bp, B_F in h43. unfold bt, finc_bp, B_F. rewrite h43 at 1. rewrite h37 at 1. f_equal. pose h20 as h20'. pose proof (f_equal (@ba_p_ens _) h19) as h44. unfold btp, directed_ba_p, Btype_p, bc_sig_set_conv, directed_bcp in h44. simpl in h44. unfold ba_sig_set_conv in h44. simpl in h44. unfold ba_p_ens, bc_sig_set_conv in h44. simpl in h44. unfold btp, directed_ba_p, Btype_p in h20'. unfold Bc_p, ba_sig_set_conv in h20'. unfold bc_sig_set_conv in h20'. simpl in h20'. rewrite (transfer_r_sig_set_eq _ _ h44 h20). apply proj1_sig_injective. simpl. reflexivity. (* -> *) intro h2. destruct h2 as [f [h2 h3]]. destruct h2 as [h2 h2']. red in h3. intros F h4 h5 a. split. intro h6. assert (h7:f (el_prod_p _ a) = 0). rewrite h6. apply homo_zero_p1. assumption. unfold el_prod_p in h7. unfold el_prod_compose_p1. rewrite <- h7 at 1. rewrite homo_times_set_p1; auto. apply times_set_functional. rewrite im_im. apply im_ext_in. intros x h8. specialize (h3 (exist _ _(h5 _ h8))). simpl in h3. unfold eps_p, eps. assert (h9: f (if a |-> x then x else %-x) = (if a |-> x then (f x) else - (f x))). destruct (a|->x); auto. rewrite homo_comp_p1; auto. rewrite h9 at 1. rewrite <- h3 at 1. destruct (a |-> x). unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h10 | h11]. f_equal. apply proj1_sig_injective. simpl. reflexivity. contradict h11. apply h5; auto. rewrite <- h3. unfold sig_fun_app. destruct (classic_dec (Inn E x)) as [h10 | h11]. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. contradict h11. apply h5; auto. intro h6. apply h2'. rewrite homo_zero_p1; auto. unfold el_prod_p. rewrite homo_times_set_p1; auto. pose proof (im_im F (fun i : btp Bp => eps_p i (a |-> i)) f) as h7. assert (h8:Finite (Im F (fun x : btp Bp => f ((fun i : btp Bp => eps_p i (a |-> i)) x)))). rewrite <- h7. rewrite im_im. apply finite_image; auto. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (btp Bp) (bt A) (Im F (fun i : btp Bp => eps_p i (a |-> i))) f (finite_image (btp Bp) (btp Bp) F (fun i : btp Bp => eps_p i (a |-> i)) (fin_map_fin_dom a))) h8 h7) as h9. unfold bt in h9. unfold bt. dependent rewrite -> h9. rewrite <- h6. unfold el_prod_compose. apply times_set_functional. apply im_ext_in. intros x h10. unfold eps_p, eps. assert (h11: f (if a |-> x then x else %- x) = (if a |-> x then (f x) else - (f x))). destruct (a|->x); auto. rewrite homo_comp_p1; auto. rewrite h11 at 1. destruct (a |-> x). unfold sig_fun_app. destruct classic_dec as [h12 | h13]. specialize (h3 (exist _ _ h12)). rewrite <- h3 at 1. f_equal. contradict h13. apply h5; auto. f_equal. unfold sig_fun_app. destruct classic_dec as [h12 | h13]. specialize (h3 (exist _ _ h12)). rewrite <- h3 at 1. f_equal. contradict h13. apply h5; auto. Grab Existential Variables. apply h15. apply h14. apply Im_intro with x; auto. Qed. Corollary mono_extension_criterion_unq_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)}, ba_p_ens Bp = Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E) -> forall (g:sig_set E->bt A), (exists! h:(btp Bp)->(bt A), monomorphism_p1 _ h /\ extends_sig1 h g) <-> (forall F : Ensemble (btp Bp), Finite F -> Included F E -> forall a : Fin_map F signe mns, el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0). intros T Bp A E h1 g. pose proof (mono_extension_criterion_p1 h1 g) as h2. split. intro h3. apply unq_ex_impl_ex in h3. rewrite h2 in h3. apply h3. intro h3. rewrite <- h2 in h3. destruct h3 as [h h3]. destruct h3 as [h3 h4]. exists h. red. split. split; auto. intros h' h5. destruct h5 as [h5 h6]. destruct h3 as [h3 h3b], h5 as [h5 h5b]. eapply gen_ens_determines_homo_p1; auto. apply h1. red in h4, h6. red. intros x h7. rewrite <- im_proj2_sig_undoes_im_proj1_sig' in h7. specialize (h4 (exist _ _ h7)). specialize (h6 (exist _ _ h7)). simpl in h4, h6. congruence. Qed. (*This is Exercise 6 in Chapter 13 of Givant/Halmos.*) Lemma countable_homo_ext_crit_iff : forall {A B:Bool_Alg} {E:Ensemble (bt B)} (i_E:nat->sig_set E), bijective i_E -> forall g:sig_set E->bt A, ((forall F:Ensemble (bt B), Finite F -> Included F E -> forall a:Fin_map F signe mns, el_prod _ a = 0 -> el_prod_compose _ (sig_fun_app g 0) a = 0) <-> (forall (n:nat) (a:Fin_map (Im (seg n) (fun i => (proj1_sig (i_E i)))) signe mns), el_prod B a = 0 -> el_prod_compose _ (sig_fun_app g 0) a = 0)). intros A B E i_E h1 g. split. intros h2 n a h3. apply h2. apply finite_image. apply finite_seg. red. intros x h4. destruct h4 as [x h4]. subst. apply proj2_sig. assumption. intros h2 F h3 h4 a h5. apply bijective_impl_invertible in h1. pose (proj1_sig (function_inverse _ h1)) as i_E'. pose (Im (full_sig F) (fun a => i_E' (exist _ _ (h4 _ (proj2_sig a))))) as inds. pose proof (finite_inh_or_empty _ h3) as h6. destruct h6 as [h6a | h6b]. assert (h7:Finite inds). unfold inds. apply finite_image. rewrite <- finite_full_sig_iff. assumption. assert (h8:Inhabited inds). unfold inds. destruct h6a as [r h6a]. apply Inhabited_intro with ((fun a0 : sig_set F => i_E' (exist (Inn E) (proj1_sig a0) (h4 (proj1_sig a0) (proj2_sig a0)))) (exist _ _ h6a)). apply Im_intro with (exist _ _ h6a). constructor. reflexivity. pose (max_set_nat inds h7 h8) as n. assert (h9:forall (b:Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns), extends b a -> el_prod B b = 0). intros b h9. pose proof (el_prod_extends_le _ _ _ h9) as h10. rewrite h5 in h10. apply le_x_0. assumption. assert (h10: forall (b:Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns), extends b a -> el_prod_compose A (sig_fun_app g 0) b = 0). intros b h10. apply h2; apply h9; auto. assert (h11:Included F (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i)))). red. intros x h11. apply Im_intro with (i_E' (exist _ _ (h4 _ h11))). constructor. unfold n. pose proof (max_set_nat_compat inds h7 h8) as h12. simpl in h12. destruct h12 as [h12 h13]. destruct (nat_eq_dec (i_E' (exist _ _ (h4 x h11))) (max_set_nat inds h7 h8)) as [he | he]. omega. specialize (h13 (i_E' (exist (Inn E) x (h4 x h11)))). hfirst h13. constructor. unfold inds. apply Im_intro with (exist _ _ h11). constructor. f_equal. rewrite in_sing_iff. apply neq_sym; auto. specialize (h13 hf). clear hf. omega. pose proof (proj2_sig (function_inverse i_E h1)) as h15. simpl in h15. destruct h15 as [h15 h16]. unfold i_E'. rewrite h16. simpl. reflexivity. assert (h12:Finite (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i)))). apply finite_image. apply finite_seg. pose proof (el_prod_compose_eq_plus_subset_non_zero_el_prod_compose_maps_extends _ _ _ h12 h11 (sig_fun_app g 0) a) as h13. simpl in h13. destruct h13 as [h13 h14]. rewrite h14. unfold plus_subset_non_zero_el_prod_compose_maps. assert (h15:Im [a0 : Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns | Inn (non_zero_el_prod_compose_maps A (sig_fun_app g 0) (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i)))) a0 /\ extends a0 a] (el_prod_compose A (sig_fun_app g 0)) = Empty_set _). apply Extensionality_Ensembles; split; red; auto with sets. intros x h16. destruct h16 as [c h16]. subst. destruct h16 as [h16]. destruct h16 as [h16 h17]. destruct h16 as [h16]. pose proof (h10 c h17) as h18. contradiction. intros; contradiction. pose proof (Empty_is_finite (Btype (Bc A))) as h16. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns) (Btype (Bc A)) [a0 : Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns | Inn (non_zero_el_prod_compose_maps A (sig_fun_app g 0) (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i)))) a0 /\ extends a0 a] (el_prod_compose A (sig_fun_app g 0)) (finite_fin_map_ens [a0 : Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns | Inn (non_zero_el_prod_compose_maps A (sig_fun_app g 0) (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i)))) a0 /\ extends a0 a] h12 signe_finite)) h16 h15) as h17. dependent rewrite -> h17. apply plus_set_empty'. subst. unfold el_prod in h5. assert (h6:Im (Empty_set (bt B)) (fun i : Btype (Bc B) => eps i (a |-> i)) = Empty_set _). apply image_empty. pose proof (Empty_is_finite (bt B)) as h7. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Btype (Bc B)) (bt B) (Empty_set (bt B)) (fun i : Btype (Bc B) => eps i (a |-> i)) (fin_map_fin_dom a)) h7 h6) as h8. unfold bt in h8, h5. simpl in h8. simpl in h5. dependent rewrite -> h8 in h5. pose proof (times_set_empty' h7) as h9. rewrite h5 in h9 at 1. assert (h10:Included (Im (seg 0) (fun i : nat => proj1_sig (i_E i))) (Empty_set _)). red. intros x h10. destruct h10 as [x h10]. destruct h10. omega. assert (h10': Im (seg 0) (fun i : nat => proj1_sig (i_E i)) = Empty_set _). auto with sets. pose (restriction a h10) as a'. assert (h11:el_prod B a' = 0). unfold el_prod. assert (h12: Im (Im (seg 0) (fun i : nat => proj1_sig (i_E i))) (fun i : Btype (Bc B) => eps i (a' |-> i)) = Empty_set _). rewrite h10' at 1. apply image_empty. pose proof (Empty_is_finite (bt B)) as h13. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Btype (Bc B)) (bt B) (Im (seg 0) (fun i : nat => proj1_sig (i_E i))) (fun i : Btype (Bc B) => eps i (a' |-> i)) (fin_map_fin_dom a')) h13 h12) as h14. unfold bt in h14. unfold bt. dependent rewrite -> h14. rewrite <- h5. f_equal. apply proof_irrelevance. specialize (h2 _ a' h11). rewrite <- h2. unfold el_prod_compose. apply times_set_functional. rewrite image_empty. rewrite h10' at 1. rewrite image_empty. reflexivity. Qed. (*This alternate form of the above is directly desgined for use with atomless Boolean algebras. Note the [1] index, which is appropariate for back-and-forth arguments.*) Lemma countable_homo_ext_crit_iff' : forall {A B:Bool_Alg} {E:Ensemble (bt B)} (i_E:{n:nat | 1 <= n}->sig_set E), surjective i_E -> ~degen B -> forall g:sig_set E->bt A, ((forall F:Ensemble (bt B), Finite F -> Included F E -> forall a:Fin_map F signe mns, el_prod _ a = 0 -> el_prod_compose _ (sig_fun_app g 0) a = 0) <-> (forall (n:{m:nat | 1 <= m}) (a:Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i => (proj1_sig (i_E (exist _ _ (seg_one_t_pos _ i)))))) signe mns), el_prod B a = 0 -> el_prod_compose _ (sig_fun_app g 0) a = 0)). intros A B E i_E h1 h0 g. split. intros h2 n a h3. apply h2. apply finite_image. rewrite Finite_FiniteT_iff. apply finite_seg_one_t. red. intros x h4. destruct h4 as [x h4]. subst. apply proj2_sig. assumption. intros h2 F h3 h4 a h5. pose proof wo_nat as hw. pose proof (well_order_exist_rel _ hw (fun m => 1 <= m)) as hw'. pose (Im (full_sig F) (fun a => inv_im (Singleton (exist (fun e => Inn E e) _ (h4 _ (proj2_sig a)))) i_E)) as inds_ens. assert (hinh:forall S, Inn inds_ens S -> Inhabited S). intros S ha. destruct ha as [x ha]. subst. clear ha. destruct x as [x ha]. simpl. red in h1. specialize (h1 (exist _ _ (h4 x ha))). destruct h1 as [m h1]. unfold inv_im. apply Inhabited_intro with m. constructor. rewrite h1. constructor. pose (Im (full_sig inds_ens) (fun S => min_set _ _ hw' _ (hinh _ (proj2_sig S)))) as inds. pose proof (finite_inh_or_empty _ h3) as h6. destruct h6 as [h6a | h6b]. assert (h7:Finite inds). unfold inds, full_sig. apply finite_image. fold (full_sig inds_ens). rewrite <- finite_full_sig_iff. unfold inds_ens. apply finite_image. rewrite <- finite_full_sig_iff. assumption. assert (h8:Inhabited inds). unfold inds. destruct h6a as [r h6a]. assert (hin:Inn inds_ens (inv_im (Singleton (exist (fun e : bt B => Inn E e) r (h4 r h6a))) i_E)). unfold inds_ens. apply Im_intro with (exist _ _ h6a). constructor. f_equal. assert (hin':Inn (Im (full_sig inds_ens) (fun S : sig_set inds_ens => min_set {t : nat | 1 <= t} (exist_rel Peano.lt (fun m : nat => 1 <= m)) hw' (proj1_sig S) (hinh (proj1_sig S) (proj2_sig S)))) ((fun S : sig_set inds_ens => min_set {t : nat | 1 <= t} (exist_rel Peano.lt (fun m : nat => 1 <= m)) hw' (proj1_sig S) (hinh (proj1_sig S) (proj2_sig S))) (exist _ _ hin))). apply Im_intro with (exist _ _ hin). constructor. simpl. reflexivity. eapply Inhabited_intro. apply hin'. assert (hf:Finite (im_proj1_sig inds)). apply finite_image; auto. assert (hinh':Inhabited (im_proj1_sig inds)). destruct h8 as [c h8]. apply Inhabited_intro with (proj1_sig c). apply Im_intro with c. assumption. reflexivity. pose (max_set_nat (im_proj1_sig inds) hf hinh') as n. assert (hle:1 <= n). unfold n. pose proof (max_set_nat_compat (im_proj1_sig inds) hf hinh') as ha. simpl in ha. destruct ha as [ha hb]. destruct ha as [n' ha]. subst. apply proj2_sig. assert (h9:forall b:Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns, extends b a -> el_prod B b = 0). intros b h9. pose proof (el_prod_extends_le _ _ _ h9) as h10. rewrite h5 in h10. apply le_x_0. assumption. assert (h10: forall b:Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns, extends b a -> el_prod_compose A (sig_fun_app g 0) b = 0). intros b h10. apply (h2 (exist _ _ hle) b). apply h9; auto. assert (h11:Included F (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i)))))). red. intros x h11. pose proof (max_set_nat_compat _ hf hinh') as h12. simpl in h12. destruct h12 as [h12 h13]. pose ( (fun a0 : sig_set F => inv_im (Singleton (exist (fun e : bt B => Inn E e) (proj1_sig a0) (h4 (proj1_sig a0) (proj2_sig a0)))) i_E) (exist _ _ h11)) as Ix. simpl in Ix. assert (hc:Inn inds_ens Ix). unfold inds_ens. apply Im_intro with (exist _ _ h11). constructor. unfold Ix. reflexivity. pose (proj1_sig (( (fun S : sig_set inds_ens => min_set {t : nat | (fun m : nat => 1 <= m) t} (exist_rel Peano.lt (fun m : nat => 1 <= m)) hw' (proj1_sig S) (hinh (proj1_sig S) (proj2_sig S))) (exist _ _ hc)))) as i. simpl in i. assert (hd:1 <= i). apply proj2_sig. assert (he:i <= n). destruct (nat_eq_dec i n) as [he | he]. rewrite he; auto with arith. apply lt_le_weak. apply h13. constructor. eapply Im_intro. eapply Im_intro with (exist _ _ hc). constructor. reflexivity. unfold i. reflexivity. rewrite in_sing_iff. fold n; apply neq_sym; auto. assert (hg:Inn (Full_set (seg_one_t n)) (exist _ _ (conj hd he))). constructor. eapply Im_intro. apply hg. simpl. pose proof (min_set_compat {t : nat | 1 <= t} (exist_rel Peano.lt (fun m : nat => 1 <= m)) hw' Ix (hinh Ix hc)) as hh. simpl in hh. destruct hh as [hha hhb]. destruct hha as [hha]. inversion hha as [hhc]. pose proof (f_equal (@proj1_sig _ _) hhc) as hi. simpl in hi. rewrite hi. f_equal. f_equal. unfold i. apply proj1_sig_injective. simpl. f_equal. assert (h12: Finite (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i)))))). apply finite_image. rewrite Finite_FiniteT_iff. apply finite_seg_one_t. pose proof (el_prod_compose_eq_plus_subset_non_zero_el_prod_compose_maps_extends _ _ _ h12 h11 (sig_fun_app g 0) a) as h13. simpl in h13. destruct h13 as [h13 h14]. rewrite h14. unfold plus_subset_non_zero_el_prod_compose_maps. assert (h15: Im [a0 : Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns | Inn (non_zero_el_prod_compose_maps A (sig_fun_app g 0) (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i)))))) a0 /\ extends a0 a] (el_prod_compose A (sig_fun_app g 0)) = Empty_set _). apply Extensionality_Ensembles; split; red; auto with sets. intros x h16. destruct h16 as [c h16]. subst. destruct h16 as [h16]. destruct h16 as [h16 h17]. destruct h16 as [h16]. pose proof (h10 c h17) as h18. contradiction. intros; contradiction. pose proof (Empty_is_finite (Btype (Bc A))) as h16. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns) (Btype (Bc A)) [a0 : Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns | Inn (non_zero_el_prod_compose_maps A (sig_fun_app g 0) (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i)))))) a0 /\ extends a0 a] (el_prod_compose A (sig_fun_app g 0)) (finite_fin_map_ens [a0 : Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns | Inn (non_zero_el_prod_compose_maps A (sig_fun_app g 0) (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i)))))) a0 /\ extends a0 a] h12 signe_finite)) h16 h15) as h17. dependent rewrite -> h17. apply plus_set_empty'. unfold el_prod_compose. unfold el_prod in h5. assert (he: Im F (fun i : Btype (Bc B) => eps i (a |-> i)) = Empty_set _). rewrite h6b at 1. apply image_empty. assert (he':Im F (fun i : Btype (Bc B) => eps i (a |-> i)) = F). rewrite h6b. assumption. pose proof (Empty_is_finite (bt B)) as h7. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Btype (Bc B)) (bt B) F (fun i : Btype (Bc B) => eps i (a |-> i)) (fin_map_fin_dom a)) h3 he') as h8. unfold bt in h8, h5. simpl in h8. simpl in h5. dependent rewrite -> h8 in h5. assert (h9:Bone (Bc B) = Bzero (Bc B)). rewrite <- h5. pose proof (subsetT_eq_compat _ _ _ _ h3 h7 h6b) as hb. dependent rewrite -> hb. rewrite times_set_empty'. reflexivity. contradict h0. red. rewrite h9. reflexivity. Qed. Lemma countable_homo_ext_crit_iff_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} (i_E:{n:nat | 1 <= n}->sig_set E), surjective i_E -> ~degen_p Bp -> forall (g:sig_set E->bt A), ((forall F : Ensemble (btp Bp), Finite F -> Included F E -> forall a : Fin_map F signe mns, el_prod_p _ a = %0 -> el_prod_compose_p1 T (sig_fun_app g 0) a = 0) <-> (forall (n:{m:nat | 1 <= m}) (a:Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i => (proj1_sig (i_E (exist _ _ (seg_one_t_pos _ i)))))) signe mns), el_prod_p _ a = %0 -> el_prod_compose_p1 T (sig_fun_app g 0) a = 0)). intros T Bp A E i_E h1 h2 g. assert (h3:forall x:btp Bp, Inn E x -> Inn (ba_conv_set E) (ba_conv_elt x)). intros; auto. pose (fun m:{n:nat | 1 <= n} => (exist _ _ (h3 _ (proj2_sig (i_E m))))) as i_E'. assert (h4:surjective i_E'). unfold i_E'. red. red in h1. intro y. specialize (h1 y). destruct h1 as [m h1]. exists m. rewrite <- h1. apply proj1_sig_injective. simpl. reflexivity. rewrite degen_p_iff in h2. pose proof (countable_homo_ext_crit_iff' i_E' h4 h2 g) as h5. split. intro h6. assert (h7:forall F : Ensemble (bt (ba_conv Bp)), Finite F -> Included F (ba_conv_set E) -> forall a : Fin_map F signe mns, el_prod (ba_conv Bp) a = 0 -> el_prod_compose A (B:=ba_conv Bp) (sig_fun_app g 0) a = 0). intros F h8 h9 a h10. assert (h10':el_prod_p _ a = %0). rewrite el_prod_p_eq. assumption. specialize (h6 F h8 h9 a h10'). rewrite el_prod_compose_p1_eq in h6. assumption. rewrite h5 in h7. intros n a h8. rewrite el_prod_p_eq in h8. specialize (h7 n a h8). rewrite el_prod_compose_p1_eq. assumption. intro h6. assert (h7: forall (n : {m : nat | 1 <= m}) (a : Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i : seg_one_t (proj1_sig n) => proj1_sig (i_E' (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos (proj1_sig n) i))))) signe mns), el_prod (ba_conv Bp) a = 0 -> el_prod_compose A (B:=ba_conv Bp) (sig_fun_app g 0) a = 0). intros n a h8. assert (h9: el_prod_p _ a = %0). rewrite el_prod_p_eq; auto. specialize (h6 n a h9). rewrite el_prod_compose_p1_eq in h6. assumption. rewrite <- h5 in h7. intros F h8 h9 a h10. assert (h10' : el_prod (ba_conv Bp) a = 0). rewrite el_prod_p_eq in h10. assumption. specialize (h7 F h8 h9 a h10'). rewrite el_prod_compose_p1_eq. assumption. Qed. Lemma countable_mono_ext_crit_iff : forall {A B:Bool_Alg} {E:Ensemble (bt B)} (i_E:nat->sig_set E), bijective i_E -> forall g:sig_set E->bt A, ((forall F:Ensemble (bt B), Finite F -> Included F E -> forall a:Fin_map F signe mns, el_prod _ a = 0 <-> el_prod_compose _ (sig_fun_app g 0) a = 0) <-> (forall (n:nat) (a:Fin_map (Im (seg n) (fun i => (proj1_sig (i_E i)))) signe mns), el_prod B a = 0 <-> el_prod_compose _ (sig_fun_app g 0) a = 0)). intros A B E i_E h1 g. split. intros h2 n a. apply h2. apply finite_image. apply finite_seg. red. intros x h4. destruct h4 as [x h4]. subst. apply proj2_sig. intros h2 F h3 h4 a. pose proof (countable_homo_ext_crit_iff i_E h1 g) as hc. split. apply hc. intros n b h5. rewrite <- h2; auto. assumption. assumption. intro h5. apply bijective_impl_invertible in h1. pose (proj1_sig (function_inverse _ h1)) as i_E'. pose (Im (full_sig F) (fun a => i_E' (exist _ _ (h4 _ (proj2_sig a))))) as inds. pose proof (finite_inh_or_empty _ h3) as h6. destruct h6 as [h6a | h6b]. assert (h7:Finite inds). unfold inds. apply finite_image. rewrite <- finite_full_sig_iff. assumption. assert (h8:Inhabited inds). unfold inds. destruct h6a as [r h6a]. apply Inhabited_intro with ((fun a0 : sig_set F => i_E' (exist (Inn E) (proj1_sig a0) (h4 (proj1_sig a0) (proj2_sig a0)))) (exist _ _ h6a)). apply Im_intro with (exist _ _ h6a). constructor. reflexivity. pose (max_set_nat inds h7 h8) as n. assert (h9:forall (b:Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns), extends b a -> el_prod_compose A (sig_fun_app g 0) b = 0). intros b h9. pose proof (el_prod_compose_extends_le _ _ _ (sig_fun_app g 0) h9) as h10. rewrite h5 in h10. apply le_x_0. assumption. assert (h10: forall (b:Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns), extends b a -> el_prod B b = 0). intros b h10. rewrite h2; auto. assert (h11:Included F (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i)))). red. intros x h11. apply Im_intro with (i_E' (exist _ _ (h4 _ h11))). constructor. unfold n. pose proof (max_set_nat_compat inds h7 h8) as h12. simpl in h12. destruct h12 as [h12 h13]. specialize (h13 (i_E' (exist (Inn E) x (h4 x h11)))). assert (h14:Inn inds (i_E' (exist (Inn E) x (h4 x h11)))). unfold inds. apply Im_intro with (exist _ _ h11). constructor. f_equal. destruct (nat_eq_dec (i_E' (exist _ _ (h4 x h11))) (max_set_nat inds h7 h8)) as [he | he]. omega. hfirst h13. constructor; auto. rewrite in_sing_iff. apply neq_sym; auto. specialize (h13 hf). clear hf. omega. pose proof (proj2_sig (function_inverse i_E h1)) as h15. simpl in h15. destruct h15 as [h15 h16]. unfold i_E'. rewrite h16. simpl. reflexivity. assert (h12:Finite (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i)))). apply finite_image. apply finite_seg. pose proof (el_prod_eq_plus_subset_non_zero_el_prod_maps_extends _ _ _ h12 h11 a) as h13. simpl in h13. destruct h13 as [h13 h14]. rewrite h14. unfold plus_subset_non_zero_el_prod_maps. assert (h15:Im [a0 : Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns | Inn (non_zero_el_prod_maps B (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i)))) a0 /\ extends a0 a] (el_prod B ) = Empty_set _). apply Extensionality_Ensembles; split; red; auto with sets. intros x h16. destruct h16 as [c h16]. subst. destruct h16 as [h16]. destruct h16 as [h16 h17]. destruct h16 as [h16]. pose proof (h10 c h17) as h18. contradiction. intros; contradiction. pose proof (Empty_is_finite (Btype (Bc B))) as h16. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns) (Btype (Bc B)) [a0 : Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns | Inn (non_zero_el_prod_maps B (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i)))) a0 /\ extends a0 a] (el_prod B) (Finite_downward_closed (Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns) (non_zero_el_prod_maps B (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i)))) (non_zero_el_prod_maps_fin B (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) h12) [a0 : Fin_map (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i))) signe mns | Inn (non_zero_el_prod_maps B (Im (seg (S n)) (fun i : nat => proj1_sig (i_E i)))) a0 /\ extends a0 a] h13)) h16 h15) as h17. unfold bt in h17. unfold bt. dependent rewrite -> h17. apply plus_set_empty'. subst. unfold el_prod_compose in h5. assert (h6: Im (Empty_set (bt B)) (fun i : bt B => eps (sig_fun_app g 0 i) (a |-> i)) = Empty_set _). apply image_empty. pose proof (Empty_is_finite (bt A)) as h7. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (bt B) (bt A) (Empty_set (bt B)) (fun i : bt B => eps (sig_fun_app g 0 i) (a |-> i)) (fin_map_fin_dom a)) h7 h6) as h8. unfold bt in h8, h5. simpl in h8. simpl in h5. dependent rewrite -> h8 in h5. pose proof (times_set_empty' h7) as h9. rewrite h5 in h9 at 1. assert (h10:Included (Im (seg 0) (fun i : nat => proj1_sig (i_E i))) (Empty_set _)). red. intros x h10. destruct h10 as [x h10]. destruct h10. omega. assert (h10': Im (seg 0) (fun i : nat => proj1_sig (i_E i)) = Empty_set _). auto with sets. pose (restriction a h10) as a'. assert (h11:el_prod_compose A (sig_fun_app g 0) a' = 0). unfold el_prod_compose. assert (h12: Im (Im (seg 0) (fun i : nat => proj1_sig (i_E i))) (fun i : bt B => eps (sig_fun_app g 0 i) (a' |-> i)) = Empty_set _). rewrite h10' at 1. apply image_empty. pose proof (Empty_is_finite (bt A)) as h13. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (bt B) (bt A) (Im (seg 0) (fun i : nat => proj1_sig (i_E i))) (fun i : bt B => eps (sig_fun_app g 0 i) (a' |-> i)) (fin_map_fin_dom a')) h13 h12) as h14. unfold bt in h14. unfold bt. dependent rewrite -> h14. rewrite <- h5. f_equal. apply proof_irrelevance. specialize (h2 _ a'). rewrite <- h2 in h11. rewrite <- h11. unfold el_prod. apply times_set_functional. rewrite image_empty. rewrite h10' at 1. rewrite image_empty. reflexivity. Qed. (*This alternate form of the above is directly desgined for use with atomless Boolean algebras.*) Lemma countable_mono_ext_crit_iff' : forall {A B:Bool_Alg} {E:Ensemble (bt B)} (i_E:{n:nat | 1 <= n}->sig_set E), surjective i_E -> ~degen A -> ~degen B -> forall g:sig_set E->bt A, ((forall F:Ensemble (bt B), Finite F -> Included F E -> forall a:Fin_map F signe mns, el_prod _ a = 0 <-> el_prod_compose _ (sig_fun_app g 0) a = 0) <-> (forall (n:{m:nat | 1 <= m}) (a:Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i => (proj1_sig (i_E (exist _ _ (seg_one_t_pos _ i)))))) signe mns), el_prod B a = 0 <-> el_prod_compose _ (sig_fun_app g 0) a = 0)). intros A B E i_E h1 h0' h0 g. split. intros h2 n a. apply h2. apply finite_image. apply Finite_FiniteT_iff. apply finite_seg_one_t. red. intros x h4. destruct h4 as [x h4]. subst. apply proj2_sig. intros h2 F h3 h4 a. pose proof wo_nat as hw. pose proof (well_order_exist_rel _ hw (fun m => 1 <= m)) as hw'. pose (Im (full_sig F) (fun a => inv_im (Singleton (exist (fun e => Inn E e) _ (h4 _ (proj2_sig a)))) i_E)) as inds_ens. assert (hinh:forall S, Inn inds_ens S -> Inhabited S). intros S ha. destruct ha as [x ha]. subst. clear ha. destruct x as [x ha]. simpl. red in h1. specialize (h1 (exist _ _ (h4 x ha))). destruct h1 as [m h1]. unfold inv_im. apply Inhabited_intro with m. constructor. rewrite h1. constructor. pose (Im (full_sig inds_ens) (fun S => min_set _ _ hw' _ (hinh _ (proj2_sig S)))) as inds. pose proof (countable_homo_ext_crit_iff' i_E h1 h0 g) as hcnt. split. apply hcnt. intros n b h5. rewrite <- h2; auto. assumption. assumption. intro h5. pose proof (finite_inh_or_empty _ h3) as h6. destruct h6 as [h6a | h6b]. assert (h7:Finite inds). unfold inds, full_sig. apply finite_image. fold (full_sig inds_ens). rewrite <- finite_full_sig_iff. unfold inds_ens. apply finite_image. rewrite <- finite_full_sig_iff. assumption. assert (h8:Inhabited inds). unfold inds. destruct h6a as [r h6a]. assert (hin:Inn inds_ens (inv_im (Singleton (exist (fun e : bt B => Inn E e) r (h4 r h6a))) i_E)). unfold inds_ens. apply Im_intro with (exist _ _ h6a). constructor. f_equal. assert (hin':Inn (Im (full_sig inds_ens) (fun S : sig_set inds_ens => min_set {t : nat | 1 <= t} (exist_rel Peano.lt (fun m : nat => 1 <= m)) hw' (proj1_sig S) (hinh (proj1_sig S) (proj2_sig S)))) ((fun S : sig_set inds_ens => min_set {t : nat | 1 <= t} (exist_rel Peano.lt (fun m : nat => 1 <= m)) hw' (proj1_sig S) (hinh (proj1_sig S) (proj2_sig S))) (exist _ _ hin))). apply Im_intro with (exist _ _ hin). constructor. simpl. reflexivity. eapply Inhabited_intro. apply hin'. assert (hf:Finite (im_proj1_sig inds)). apply finite_image; auto. assert (hinh':Inhabited (im_proj1_sig inds)). destruct h8 as [c h8]. apply Inhabited_intro with (proj1_sig c). apply Im_intro with c. assumption. reflexivity. pose (max_set_nat (im_proj1_sig inds) hf hinh') as n. assert (hle:1 <= n). unfold n. pose proof (max_set_nat_compat (im_proj1_sig inds) hf hinh') as ha. simpl in ha. destruct ha as [ha hb]. destruct ha as [n' ha]. subst. apply proj2_sig. assert (h9:forall b:Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns, extends b a -> el_prod_compose A (sig_fun_app g 0) b = 0). intros b h9. pose proof (el_prod_compose_extends_le _ _ _ (sig_fun_app g 0) h9) as h10. rewrite h5 in h10. apply le_x_0. assumption. assert (h10: forall b:Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns, extends b a -> el_prod B b = 0). intros b h10. rewrite (h2 (exist _ _ hle)); auto. assert (h11:Included F (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i)))))). red. intros x h11. pose proof (max_set_nat_compat _ hf hinh') as h12. simpl in h12. destruct h12 as [h12 h13]. pose ( (fun a0 : sig_set F => inv_im (Singleton (exist (fun e : bt B => Inn E e) (proj1_sig a0) (h4 (proj1_sig a0) (proj2_sig a0)))) i_E) (exist _ _ h11)) as Ix. simpl in Ix. assert (hc:Inn inds_ens Ix). unfold inds_ens. apply Im_intro with (exist _ _ h11). constructor. unfold Ix. reflexivity. pose (proj1_sig (( (fun S : sig_set inds_ens => min_set {t : nat | (fun m : nat => 1 <= m) t} (exist_rel Peano.lt (fun m : nat => 1 <= m)) hw' (proj1_sig S) (hinh (proj1_sig S) (proj2_sig S))) (exist _ _ hc)))) as i. simpl in i. assert (hd:1 <= i). apply proj2_sig. assert (he:i <= n). destruct (nat_eq_dec i n) as [he | he]. omega. apply lt_le_weak. apply h13. constructor. eapply Im_intro. eapply Im_intro with (exist _ _ hc). constructor. reflexivity. unfold i. reflexivity. rewrite in_sing_iff. fold n; apply neq_sym; auto. assert (hg:Inn (Full_set (seg_one_t n)) (exist _ _ (conj hd he))). constructor. eapply Im_intro. apply hg. simpl. pose proof (min_set_compat {t : nat | 1 <= t} (exist_rel Peano.lt (fun m : nat => 1 <= m)) hw' Ix (hinh Ix hc)) as hh. simpl in hh. destruct hh as [hha hhb]. destruct hha as [hha]. inversion hha as [hhc]. pose proof (f_equal (@proj1_sig _ _) hhc) as hi. simpl in hi. rewrite hi. f_equal. f_equal. unfold i. apply proj1_sig_injective. simpl. f_equal. assert (h12: Finite (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i)))))). apply finite_image. rewrite Finite_FiniteT_iff. apply finite_seg_one_t. pose proof (el_prod_eq_plus_subset_non_zero_el_prod_maps_extends _ _ _ h12 h11 a) as h13. simpl in h13. destruct h13 as [h13 h14]. rewrite h14. unfold plus_subset_non_zero_el_prod_maps. assert (h15: Im [a0 : Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns | Inn (non_zero_el_prod_maps B (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i)))))) a0 /\ extends a0 a] (el_prod B) = Empty_set _). apply Extensionality_Ensembles; split; red; auto with sets. intros x h16. destruct h16 as [c h16]. subst. destruct h16 as [h16]. destruct h16 as [h16 h17]. destruct h16 as [h16]. pose proof (h10 c h17) as h18. contradiction. intros; contradiction. pose proof (Empty_is_finite (Btype (Bc B))) as h16. simpl. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns) (Btype (Bc B)) [a0 : Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns | Inn (non_zero_el_prod_maps B (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i)))))) a0 /\ extends a0 a] (el_prod B) (Finite_downward_closed (Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns) (non_zero_el_prod_maps B (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i)))))) (non_zero_el_prod_maps_fin B (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) h12) [a0 : Fin_map (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i))))) signe mns | Inn (non_zero_el_prod_maps B (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos n i)))))) a0 /\ extends a0 a] h13)) h16 h15) as h17. unfold bt in h17. unfold bt. dependent rewrite -> h17. apply plus_set_empty'. unfold el_prod_compose in h5. assert (he:Im F (fun i : bt B => eps (sig_fun_app g 0 i) (a |-> i)) = Empty_set _). rewrite h6b at 1. apply image_empty. pose proof (Empty_is_finite (bt A)) as h7. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (bt B) (bt A) F (fun i : bt B => eps (sig_fun_app g 0 i) (a |-> i)) (fin_map_fin_dom a)) h7 he) as h8. unfold bt in h8, h5. simpl in h8. simpl in h5. dependent rewrite -> h8 in h5. assert (h9:Bone (Bc A) = Bzero (Bc A)). rewrite <- h5 at 1. rewrite times_set_empty'. reflexivity. contradict h0'. red. rewrite h9. reflexivity. Qed. Lemma countable_mono_ext_one_compat_iff : forall {A B:Bool_Alg} {E:Ensemble (bt B)} (g:sig_set E->bt A), ~degen A -> ~degen B -> ((forall (i_E:nat->sig_set E) (n:nat) (a:Fin_map (Im (seg n) (fun i => (proj1_sig (i_E i)))) signe mns), el_prod B a = 0 <-> el_prod_compose _ (sig_fun_app g 0) a = 0) <-> (forall (i_E:{n:nat | 1 <= n}->sig_set E) (n:{m:nat | 1 <= m}) (a:Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i => (proj1_sig (i_E (exist _ _ (seg_one_t_pos _ i)))))) signe mns), el_prod B a = 0 <-> el_prod_compose _ (sig_fun_app g 0) a = 0)). intros A B E g hnda hndb. split. intros h1 i_E n a. pose (fun n:nat => i_E (exist _ (S n) (S_le_1 n))) as i_E'. specialize (h1 i_E' (proj1_sig n)). let P := type of h1 in match P with forall a:Fin_map ?S _ _, _ => pose S as C end. let P := type of a in match P with Fin_map ?S _ _ => pose S as D end. assert (h2 : C = D). apply Extensionality_Ensembles; red; split; red; intros m h2. destruct h2 as [m h2]. subst. destruct h2 as [h2]. assert (h3:1 <= S m <= proj1_sig n). omega. apply Im_intro with (exist (fun a => 1 <= a <= proj1_sig n) _ h3). constructor. unfold i_E'. f_equal. f_equal. apply proj1_sig_injective; auto. destruct h2 as [m h2]. subst. clear h2. destruct m as [m h3]. assert (h4:pred m < proj1_sig n). omega. apply Im_intro with (pred m). constructor; auto. unfold i_E'. f_equal. f_equal. apply proj1_sig_injective; simpl. assert (h5:0 < m). omega. rewrite <- (S_pred _ _ h5) at 1. reflexivity. symmetry in h2. pose (fin_map_dom_subst h2 a) as a'. specialize (h1 a'). unfold a' in h1. unfold el_prod in h1. unfold el_prod. let P := type of h1 in match P with times_set ?S _ = 0 <-> _ => pose S as G end. match goal with |- times_set ?S _ = 0 <-> _ => pose S as K end. fold K. fold G in h1. let P := type of h1 in match P with _ <-> ?x' = _ => pose x' as x end. match goal with |- _ <-> ?x' = _ => pose x' as y end. assert (h3:x = y). unfold x, y. symmetry. eapply el_prod_compose_fin_map_dom_subst_compat. fold y. fold x in h1. rewrite h3 in h1. rewrite <- h1. assert (h4:K = G). unfold K, G. do 2 rewrite im_im. match goal with |- Im _ ?f'= _ => pose f' as f end. fold f. assert (h11n : 1 <= 1 <= proj1_sig n). auto with arith. split. auto with arith. apply proj2_sig. unfold seg_one_t. apply Extensionality_Ensembles; red; split; red; intros z h14. destruct h14 as [z h14]. subst. clear h14. apply Im_intro with (pred (proj1_sig z)). constructor. destruct z as [z h14]. simpl. omega. unfold f. f_equal. f_equal. unfold i_E'. f_equal. apply proj1_sig_injective; simpl. destruct z as [z h12]. simpl. assert (h13:0 < z). omega. rewrite <- (S_pred _ _ h13). reflexivity. rewrite <- (fin_map_dom_subst_compat h2 a). f_equal. f_equal. unfold i_E'. f_equal. apply proj1_sig_injective; simpl. destruct z as [z h14]. assert (h13:0 < z). omega. simpl. rewrite <- (S_pred _ _ h13). reflexivity. unfold i_E'. apply Im_intro with z. constructor. f_equal. f_equal. apply proj1_sig_injective; simpl. destruct z as [z h14]. simpl. assert (h13:0 < z). omega. simpl. rewrite <- (S_pred _ _ h13). reflexivity. destruct h14 as [z h14]. subst. destruct h14 as [h14]. assert (h15:1 <= S z <= proj1_sig n). omega. apply Im_intro with (exist (fun a => 1 <= a <= proj1_sig n) _ h15). constructor. unfold f. f_equal. unfold i_E'. f_equal. f_equal. apply proj1_sig_injective; simpl. reflexivity. rewrite (fin_map_dom_subst_compat h2 a) at 1. f_equal. f_equal. unfold i_E'. f_equal. apply proj1_sig_injective; simpl. reflexivity. simpl. apply Im_intro with (exist (fun a => 1 <= a <= proj1_sig n) _ h15). constructor. f_equal. split. intro h5. rewrite <- h5. apply times_set_functional. rewrite h4. reflexivity. intro h5. rewrite <- h5. apply times_set_functional; auto. intros h1 i_E n a. destruct (zerop n) as [h2 | h2]. subst. assert (h2: (Im (seg 0) (fun i : nat => proj1_sig (i_E i))) = Empty_set _). rewrite seg_0_eq. rewrite image_empty. reflexivity. pose (fin_map_dom_subst h2 a) as a'. pose proof (el_prod_fin_map_dom_subst_compat _ a h2) as h4. pose proof (el_prod_compose_fin_map_dom_subst_compat _ a (sig_fun_app g 0) h2) as h3. rewrite h3, h4. rewrite el_prod_compose_empty, el_prod_empty. split. intro h5. contradict hndb. red. rewrite h5. reflexivity. intro h5. contradict hnda. red. rewrite h5. reflexivity. assert (h2':1 <= n). auto with arith. pose (fun i:{n:nat | 1 <= n} => i_E (pred (proj1_sig i))) as i_E'. pose (exist _ _ h2') as n'. specialize (h1 i_E' n'). unfold i_E', n' in h1. simpl in h1. let P := type of h1 in match P with forall a:Fin_map ?S _ _ , _ => pose S as C end. let P := type of a in match P with Fin_map ?S _ _ => pose S as D end. fold C in h1. fold D in a. assert (h3:C = D). apply Extensionality_Ensembles; red; split; red; intros m h4; destruct h4 as [m h4]; subst. unfold D. unfold i_E'. simpl. destruct m as [m h3]. destruct h3 as [h3 h5]. simpl. apply Im_intro with (pred m). constructor. omega. reflexivity. unfold C. destruct h4 as [h4]. assert (h5 : 1 <= S m <= n). omega. apply Im_intro with (exist (fun a => 1 <= a <= n) _ h5). constructor. simpl. reflexivity. unfold el_prod. fold D in h1. symmetry in h3. pose (fin_map_dom_subst h3 a) as a'. specialize (h1 a'). unfold a' in h1. unfold el_prod in h1. unfold el_prod. let P := type of h1 in match P with times_set ?S _ = 0 <-> _ => pose S as G end. match goal with |- times_set ?S _ = 0 <-> _ => pose S as K end. fold K. fold G in h1. let P := type of h1 in match P with _ <-> ?x' = _ => pose x' as x end. match goal with |- _ <-> ?x' = _ => pose x' as y end. assert (h5:x = y). unfold x, y. symmetry. eapply el_prod_compose_fin_map_dom_subst_compat. fold y. fold x in h1. rewrite h5 in h1. rewrite <- h1. assert (h4:K = G). unfold K, G, C. do 2 rewrite im_im. match goal with |- Im _ ?f'= _ => pose f' as f end. fold f. assert (h11n : 1 <= 1 <= proj1_sig (exist _ _ h2)). auto with arith. unfold seg_one_t. apply Extensionality_Ensembles; red; split; red; intros z h14. destruct h14 as [z h14]. subst. destruct h14 as [h14]. assert (h15:1 <= S z <= n). omega. apply Im_intro with (exist (fun a => 1 <= a <= n) _ h15). constructor. unfold f. f_equal. f_equal. rewrite (fin_map_dom_subst_compat h3 a). f_equal. f_equal. unfold i_E'. f_equal. apply Im_intro with z. constructor. assumption. reflexivity. destruct h14 as [z h14]. subst. clear h14. destruct z as [z h14]. apply Im_intro with (pred z). constructor. omega. unfold f. f_equal. rewrite (fin_map_dom_subst_compat h3 a). f_equal. apply Im_intro with (pred z). constructor. omega. reflexivity. split. intro h6. rewrite <- h6. apply times_set_functional; auto. intro h6. rewrite <- h6. apply times_set_functional; auto. Qed. (*Slight variation of the above.*) Lemma countable_mono_ext_one_compat_iff' : forall {A B:Bool_Alg} {E:Ensemble (bt B)} (g:sig_set E->bt A), ~degen A -> ~degen B -> forall (i_E:nat->sig_set E), let i_E' := fun n':{n:nat | 1 <= n} => i_E (pred (proj1_sig n')) in (forall (n:nat) (a:Fin_map (Im (seg n) (fun i => (proj1_sig (i_E i)))) signe mns), el_prod B a = 0 <-> el_prod_compose _ (sig_fun_app g 0) a = 0) <-> (forall (n:{m:nat | 1 <= m}) (a:Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i => (proj1_sig (i_E' (exist _ _ (seg_one_t_pos _ i)))))) signe mns), el_prod B a = 0 <-> el_prod_compose _ (sig_fun_app g 0) a = 0). intros A B E g hnda hndb i_E i_E'. split. intros h1 n a. destruct n as [n hn]. specialize (h1 n). simpl in a. assert (hn': 0 < n). omega. let P := type of h1 in match P with forall a:Fin_map ?S _ _, _ => pose S as C end. let P := type of a in match P with Fin_map ?S _ _ => pose S as D end. assert (h2: C = D). apply Extensionality_Ensembles; red; split; red; intros m h2. destruct h2 as [m h2]. subst. destruct h2 as [h2]. unfold D. assert (h3:1 <= S m <= n). omega. apply Im_intro with (exist (fun a => 1 <= a <= n) _ h3). constructor. unfold i_E'. reflexivity. destruct h2 as [m h2]. subst. clear h2. simpl in m. unfold i_E'. simpl. unfold C. destruct m as [m h3]. simpl. simpl in h3, D. apply Im_intro with (pred m). constructor. omega. reflexivity. symmetry in h2. pose (fin_map_dom_subst h2 a) as a'. specialize (h1 a'). unfold a' in h1. unfold el_prod in h1. unfold el_prod. let P := type of h1 in match P with times_set ?S _ = 0 <-> _ => pose S as G end. match goal with |- times_set ?S _ = 0 <-> _ => pose S as K end. fold K. fold G in h1. let P := type of h1 in match P with _ <-> ?x' = _ => pose x' as x end. match goal with |- _ <-> ?x' = _ => pose x' as y end. assert (h3:x = y). unfold x, y. symmetry. eapply el_prod_compose_fin_map_dom_subst_compat. fold y. fold x in h1. rewrite h3 in h1. rewrite <- h1. assert (h4:K = G). unfold K, G. do 2 rewrite im_im. match goal with |- Im _ ?f'= _ => pose f' as f end. fold f. apply Extensionality_Ensembles; red; split; red; intros z h14. destruct h14 as [z h14]. subst. clear h14. apply Im_intro with (pred (proj1_sig z)). constructor. destruct z as [z h14]. simpl. simpl in h14. omega. unfold f. f_equal. f_equal. unfold i_E'. f_equal. rewrite (fin_map_dom_subst_compat h2 a). f_equal. destruct z as [z h12]. simpl. assert (h13:0 < z). omega. unfold i_E'. simpl. simpl in h12. apply Im_intro with (exist (fun a => 1 <= a <= n) _ h12). constructor. simpl. reflexivity. destruct h14 as [z h14]. subst. destruct h14 as [h14]. assert (h15:1 <= S z <= n). omega. apply Im_intro with (exist (fun a => 1 <= a <= n) _ h15). constructor. unfold f. f_equal. unfold i_E'. f_equal. rewrite (fin_map_dom_subst_compat h2 a). f_equal. apply Im_intro with (exist (fun a => 1 <= a <= n) _ h15). constructor. f_equal. split. intro h6. rewrite <- h6. apply times_set_functional. rewrite h4; auto. intro h5. rewrite <- h5. apply times_set_functional. rewrite h4; auto. intros h1 n a. destruct (zerop n) as [h2 | h2]. subst. assert (h2 : Im (seg 0) (fun i : nat => proj1_sig (i_E i)) = Empty_set _). rewrite seg_0_eq. rewrite image_empty. reflexivity. pose (fin_map_dom_subst_compat h2 a) as a'. unfold el_prod, el_prod_compose. split. intro h3. let P := type of h3 in match P with times_set ?S _ = _ => assert (h4: S = Empty_set _) end. rewrite im_im. rewrite seg_0_eq at 1. rewrite image_empty. reflexivity. let P := type of h3 in match P with times_set _ ?pf = _ => pose pf as hyp end. fold hyp in h3. pose proof (times_set_functional _ _ h4 hyp (Empty_is_finite _)) as h5. rewrite h5 in h3. rewrite times_set_empty in h3. contradict hndb. red. rewrite h3. reflexivity. intro h3. let P := type of h3 in match P with times_set ?S _ = _ => assert (h4: S = Empty_set _) end. rewrite im_im. rewrite seg_0_eq at 1. rewrite image_empty. reflexivity. let P := type of h3 in match P with times_set _ ?pf = _ => pose pf as hyp end. fold hyp in h3. pose proof (times_set_functional _ _ h4 hyp (Empty_is_finite _)) as h5. rewrite h5 in h3. rewrite times_set_empty in h3. contradict hnda. red. rewrite h3. reflexivity. assert (h3: 1 <= n). omega. specialize (h1 (exist _ _ h3)). simpl in h1. let P := type of h1 in match P with forall a:Fin_map ?S _ _, _ => pose S as C end. let P := type of a in match P with Fin_map ?S _ _ => pose S as D end. assert (hcd: C = D). apply Extensionality_Ensembles; red; split; red; intros m hm. destruct hm as [m hm]. subst. clear hm. destruct m as [m h4]. unfold i_E', D. simpl. apply Im_intro with (pred m). constructor. omega. reflexivity. destruct hm as [m hm]; subst. destruct hm as [hm]. unfold C. unfold i_E'. simpl. assert (h4:1 <= S m <= n). omega. apply Im_intro with (exist (fun a => 1 <= a <= n) _ h4). constructor. f_equal. symmetry in hcd. pose (fin_map_dom_subst hcd a) as a'. specialize (h1 a'). unfold a' in h1. unfold el_prod in h1. unfold el_prod. let P := type of h1 in match P with times_set ?S _ = 0 <-> _ => pose S as G end. match goal with |- times_set ?S _ = 0 <-> _ => pose S as K end. fold K. fold G in h1. let P := type of h1 in match P with _ <-> ?x' = _ => pose x' as x end. match goal with |- _ <-> ?x' = _ => pose x' as y end. assert (hxy:x = y). unfold x, y. symmetry. eapply el_prod_compose_fin_map_dom_subst_compat. fold y. fold x in h1. rewrite hxy in h1. rewrite <- h1. assert (h4:K = G). unfold K, G. do 2 rewrite im_im. match goal with |- Im _ ?f'= _ => pose f' as f end. fold f. apply Extensionality_Ensembles; red; split; red; intros z h14. destruct h14 as [z h14]. subst. destruct h14 as [h14]. assert (hzn : 1 <= S z <= n). omega. apply Im_intro with (exist (fun a => 1 <= a <= n) _ hzn). constructor. unfold f. f_equal. f_equal. unfold i_E'. f_equal. rewrite (fin_map_dom_subst_compat hcd a). f_equal. apply Im_intro with z. constructor. assumption. reflexivity. destruct h14 as [z h14]. subst. clear h14. destruct z as [z h15]. unfold i_E'. simpl. apply Im_intro with (pred z). constructor. omega. unfold f. f_equal. rewrite (fin_map_dom_subst_compat hcd a). f_equal. apply Im_intro with (pred z). constructor. omega. reflexivity. split. intro h6. rewrite <- h6. apply times_set_functional. rewrite h4; auto. intro h5. rewrite <- h5. apply times_set_functional. rewrite h4; auto. Qed. Lemma countable_mono_ext_crit_iff'' : forall {A B:Bool_Alg} {E:Ensemble (bt B)} (i_E:nat->sig_set E), surjective i_E -> ~degen A -> ~degen B -> forall g:sig_set E->bt A, ((forall F:Ensemble (bt B), Finite F -> Included F E -> forall a:Fin_map F signe mns, el_prod _ a = 0 <-> el_prod_compose _ (sig_fun_app g 0) a = 0) <-> (forall (n:nat) (a:Fin_map (Im (seg n) (fun i => (proj1_sig (i_E i)))) signe mns), el_prod B a = 0 <-> el_prod_compose _ (sig_fun_app g 0) a = 0)). intros A B E i_E h1 h2 h3 g. pose proof (countable_mono_ext_one_compat_iff' g h2 h3 i_E) as h8. simpl in h8. pose (fun i:{n:nat | 1 <= n} => i_E (pred (proj1_sig i))) as i_E'. assert (h6:surjective i_E'). red. intro y. red in h1. specialize (h1 y). destruct h1 as [z h1]. subst. unfold i_E'. exists (exist _ _ (S_le_1 z)). simpl. reflexivity. pose proof (countable_mono_ext_crit_iff' i_E' h6 h2 h3 g) as h5. rewrite h5. rewrite h8. unfold i_E'. simpl. tauto. Qed. Definition l_finite_mono_ext_crit_finite {A:Bool_Alg} {B:Bool_Alg} (la:list (bt A)) (lb:list (bt B)) (pfeq:length la = length lb) : Prop := forall (m:nat) (pfm:m <= length la) (F:Fin_map (seg m) signe mns), el_prod_l_firstn pfm F = 0 <-> el_prod_l_firstn (le_eq_trans _ _ _ pfm pfeq) F = 0. Definition l_finite_mono_ext_crit_finite_p {T:Type} {Ap Bp:Bool_Alg_p T} (la:list (btp Ap)) (lb:list (btp Bp)) (pfeq:length la = length lb) : Prop := forall (m:nat) (pfm:m <= length la) (F:Fin_map (seg m) signe mns), el_prod_l_firstn_p pfm F = %0 <-> el_prod_l_firstn_p (le_eq_trans _ _ _ pfm pfeq) F = %0. Definition l_finite_mono_ext_crit_finite_p1 {T:Type} {Ap:Bool_Alg_p T} {B:Bool_Alg} (la:list (btp Ap)) (lb:list (bt B)) (pfeq:length la = length lb) : Prop := forall (m:nat) (pfm:m <= length la) (F:Fin_map (seg m) signe mns), el_prod_l_firstn_p pfm F = %0 <-> el_prod_l_firstn (le_eq_trans _ _ _ pfm pfeq) F = 0. Lemma l_finite_mono_ext_crit_finite_p1_nil : forall (T:Type) (Ap:Bool_Alg_p T) (B:Bool_Alg), ~degen_p Ap -> ~degen B -> @l_finite_mono_ext_crit_finite_p1 _ Ap B nil nil (eq_refl _). intros T Ap B hd hd'; red; intros m h1 F. simpl in h1. pose proof (le0 _ h1); subst; simpl. rewrite el_prod_l_firstn_p_eq. rewrite ba_conv_zero. assert (h2:h1 = le_eq_trans _ _ _ h1 eq_refl). apply proof_irrelevance. rewrite <- h2. unfold el_prod_l_firstn, el_prod_l. simpl. rewrite eq_sym_iff. change (degen_p Ap <-> (@eq (bt B) (Bone (Bc B)) (Bzero (Bc B)))). rewrite eq_sym_iff. change (degen_p Ap <-> degen B). split; intro; contradiction. Qed. Definition l_countable_mono_ext_crit {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} {D:Ensemble (bt A)} (i_E:nat->sig_set E) (i_D:nat->sig_set D) : Prop := forall (n:nat), let fb := (fun i (pfi:i < n) => (proj1_sig (i_E i))) in let fa := (fun i (pfi:i < n) => (proj1_sig (i_D i))) in let lb := map_seg fb in let la := map_seg fa in forall (Fb:Fin_map (seg (length lb)) signe mns), let pfe := eq_trans (length_map_seg (fun i (pfi:i < n) => (proj1_sig (i_E i)))) (eq_sym (length_map_seg (fun i (pfi:i < n) => (proj1_sig (i_D i))))) in let Fa := fin_map_seg_transfer Fb pfe in (el_prod_l_p Fb = %0 <-> el_prod_l Fa = 0). Lemma countable_mono_ext_crit_iff_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} (i_E:nat->sig_set E), surjective i_E -> ~degen A -> ~degen_p Bp -> forall (g:sig_set E->bt A), ((forall F : Ensemble (btp Bp), Finite F -> Included F E -> forall a : Fin_map F signe mns, el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0) <-> (forall (n:nat) (a:Fin_map (Im (seg n) (fun i => (proj1_sig (i_E i)))) signe mns), el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0)). intros T Bp A E i_E h1 h2 h2' g. rewrite degen_p_iff in h2'. assert (h3:forall x:btp Bp, Inn E x -> Inn (ba_conv_set E) (ba_conv_elt x)). intros; auto. pose (fun m:nat => (exist _ _ (h3 _ (proj2_sig (i_E m))))) as i_E'. assert (h4:surjective i_E'). unfold i_E'. red. red in h1. intro y. specialize (h1 y). destruct h1 as [m h1]. exists m. rewrite <- h1. apply proj1_sig_injective. simpl. reflexivity. pose proof (countable_mono_ext_crit_iff'' i_E' h4 h2 h2' g) as h5. split. intros h6 n a. let P := type of h5 in match P with ?h <-> _ => assert (h7:h) end. intros F h7 h8 a'. specialize (h6 F h7 h8 a'). rewrite <- h6. rewrite el_prod_p_eq. tauto. rewrite h5 in h7. rewrite el_prod_compose_p1_eq, el_prod_p_eq. apply h7. intros h6. let P := type of h5 in match P with _ <-> ?h => assert (h7:h) end. intros n a. unfold i_E' in a. simpl in a. specialize (h6 n a). rewrite el_prod_p_eq, el_prod_compose_p1_eq in h6. assumption. rewrite <- h5 in h7. intros F h8 h9 a. rewrite el_prod_p_eq, el_prod_compose_p1_eq. apply h7; auto. Qed. Lemma countable_mono_ext_crit_iff_p1' : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} (i_E:{n:nat | 1 <= n}->sig_set E), surjective i_E -> ~degen A -> ~degen_p Bp -> forall (g:sig_set E->bt A), ((forall F : Ensemble (btp Bp), Finite F -> Included F E -> forall a : Fin_map F signe mns, el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0) <-> (forall (n:{m:nat | 1 <= m}) (a:Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i => (proj1_sig (i_E (exist _ _ (seg_one_t_pos _ i)))))) signe mns), el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0)). intros T Bp A E i_E h1 h2 h2' g. assert (h3:forall x:btp Bp, Inn E x -> Inn (ba_conv_set E) (ba_conv_elt x)). intros; auto. pose (fun m:{n:nat | 1 <= n} => (exist _ _ (h3 _ (proj2_sig (i_E m))))) as i_E'. assert (h4:surjective i_E'). unfold i_E'. red. red in h1. intro y. specialize (h1 y). destruct h1 as [m h1]. exists m. rewrite <- h1. apply proj1_sig_injective. simpl. reflexivity. rewrite degen_p_iff in h2'. pose proof (countable_mono_ext_crit_iff' i_E' h4 h2 h2' g) as h5. split. intro h6. assert (h7:forall F : Ensemble (bt (ba_conv Bp)), Finite F -> Included F (ba_conv_set E) -> forall a : Fin_map F signe mns, el_prod (ba_conv Bp) a = 0 <-> el_prod_compose A (B:=ba_conv Bp) (sig_fun_app g 0) a = 0). intros F h8 h9 a. specialize (h6 F h8 h9 a). rewrite el_prod_p_eq, el_prod_compose_p1_eq in h6. assumption. rewrite h5 in h7. intros n a. specialize (h7 n a). rewrite el_prod_p_eq, el_prod_compose_p1_eq. assumption. intro h6. assert (h7: forall (n : {m : nat | 1 <= m}) (a : Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i : seg_one_t (proj1_sig n) => proj1_sig (i_E' (exist (fun n0 : nat => 1 <= n0) (proj1_sig i) (seg_one_t_pos (proj1_sig n) i))))) signe mns), el_prod (ba_conv Bp) a = 0 <-> el_prod_compose A (B:=ba_conv Bp) (sig_fun_app g 0) a = 0). intros n a. specialize (h6 n a). rewrite el_prod_p_eq, el_prod_compose_p1_eq in h6. assumption. rewrite <- h5 in h7. intros F h8 h9 a. specialize (h7 F h8 h9 a). rewrite el_prod_p_eq, el_prod_compose_p1_eq. assumption. Qed. Lemma countable_mono_ext_crit_iff_prop_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} (i_E:nat->sig_set E), surjective i_E -> ~degen A -> ~degen_p Bp -> ((forall (F : Ensemble (btp Bp)) (g:sig_set F->bt A), Finite F -> Included F E -> forall a : Fin_map F signe mns, el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0) <-> (forall (n:nat), let F := Im (seg n) (fun i => (proj1_sig (i_E i))) in forall (g:sig_set F -> bt A) (a:Fin_map F signe mns), el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0)). intros T Bp A E i_E h1 h2 h2'. split. intros h3 n F g a. pose proof (finite_image _ _ (seg n) (fun i:nat => proj1_sig (i_E i)) (finite_seg _)) as h4. specialize (h3 F g h4). hfirst h3. red; intros x h6. unfold F in h6. inversion h6; subst. apply proj2_sig. specialize (h3 hf a); auto. intros h3 F g h4 h5 a'. pose (fun x:sig_set E => match (in_fin_ens_dec ba_eq_dec_p F h4 (proj1_sig x)) with | left P => g (exist _ _ P) | right _ => 0 end) as g'. pose proof (countable_mono_ext_crit_iff_p1 i_E h1 h2 h2' g') as h6. hr h6. intros n a. specialize (h3 n); simpl in h3. assert (h7:Included (Im (seg n) (fun i : nat => proj1_sig (i_E i))) E). red; intros x h8. destruct h8 as [? h8]; subst. apply proj2_sig. pose (fun x : sig_set (Im (seg n) (fun i : nat => proj1_sig (i_E i))) => (g' (exist _ _ (h7 _ (proj2_sig x))))) as g''. specialize (h3 g'' a). rewrite h3. unfold g', g'', g'. simpl. apply el_prod_compose_p1_sig_fun_app_im_seg; auto. rewrite <- h6 in hr. specialize (hr F h4 h5 a'). rewrite hr. unfold g'. apply el_prod_compose_p1_sig_fun_app_comp; auto. Qed. Corollary l_countable_mono_extension_criterion_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} {D:Ensemble (bt A)} (i_E:nat->sig_set E) (i_D:nat->sig_set D) (pfse:surjective i_E) (pfsd:surjective i_D), ~degen_p Bp -> ~degen A -> ba_p_ens Bp = Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E) -> ba_ens A = Gen_Ens D -> fun_linds_eq_nat (sig_set_eq_dec ba_eq_dec_p E) (sig_set_eq_dec ba_eq_dec D) i_E i_D -> let g := surj_nat_to_fun_sig i_E i_D pfse in (exists h:(btp Bp)->(bt A), monomorphism_p1 _ h /\ extends_sig1 h g) <-> l_countable_mono_ext_crit i_E i_D. intros T Bp A E D i_E i_D hse hsd hdb hda hbe had hfe g. rewrite (mono_extension_criterion_p1 hbe g). split; intro h2. red. intros n Fb. specialize (h2 (Im (seg n) (proj1_sig_fun i_E))). hfirst h2. apply finite_image. apply finite_seg. specialize (h2 hf). hfirst h2. red; intros x h3. destruct h3 as [x h3]; subst. rewrite proj1_sig_fun_compat. apply proj2_sig. specialize (h2 hf0). pose proof (length_map_seg (fun_to_seg (proj1_sig_fun i_E) n)) as h3. unfold fun_to_seg, proj1_sig_fun in h3. pose (fin_map_seg_transfer Fb h3) as Fb'. pose proof (only_has_sames_p_dec' _ (map_seg (fun_to_seg (proj1_sig_fun i_E) n)) (seg n) Fb (finite_seg _)) as h8. destruct h8 as [h8 | h8]. pose (only_has_sames_map_seg_to_fin_map_p _ _ _ h8) as F. specialize (h2 F). pose proof (only_has_sames_map_el_prod_p _ _ _ h8) as h4. simpl in h4. unfold F in h2. rewrite <- h4 in h2. rewrite h2. unfold g. pose proof (el_prod_compose_sig_fun_app_only_has_sames_p1 _ i_E i_D hse hsd n Fb h8 hfe) as h5. rewrite h5. tauto. erewrite el_prod_l_has_comps_map_seg, el_prod_l_has_comps_map_seg_transfer; try tauto. refine i_D. rewrite countable_mono_ext_crit_iff_p1; auto. intros n F. red in h2. specialize (h2 n). pose (fin_map_im_seg_to_seg_length_p _ (proj1_sig_fun i_E) F) as F'. specialize (h2 F'). unfold g. erewrite el_prod_l_fin_map_im_seg_to_seg_length_p, el_prod_compose_sig_fun_app_p1. unfold F' in h2; auto. assumption. assumption. assumption. Qed. Corollary countable_mono_extension_criterion_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} (i_E:nat->sig_set E), ~degen A -> ~degen_p Bp -> surjective i_E -> ba_p_ens Bp = Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E) -> forall (g:sig_set E->bt A), (exists! h:(btp Bp)->(bt A), monomorphism_p1 _ h /\ extends_sig1 h g) <-> (forall (n:nat) (a:Fin_map (Im (seg n) (fun i => (proj1_sig (i_E i)))) signe mns), el_prod_p _ a = %0 <-> el_prod_compose_p1 _ (sig_fun_app g 0) a = 0). intros T Bp A E i_E hda hdb h1 h2 g. rewrite <- countable_mono_ext_crit_iff_p1; auto. apply mono_extension_criterion_unq_p1; auto. Qed. Corollary countable_mono_extension_criterion_p1' : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} (i_E:{n:nat | 1 <= n}->sig_set E), ~degen A -> ~degen_p Bp -> surjective i_E -> ba_p_ens Bp = Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E) -> forall (g:sig_set E->bt A), (exists! h:(btp Bp)->(bt A), monomorphism_p1 _ h /\ extends_sig1 h g) <-> (forall (n:{m:nat | 1 <= m}) (a:Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i => (proj1_sig (i_E (exist _ _ (seg_one_t_pos _ i)))))) signe mns), el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0). intros T Bp A E i_E hda hdb h1 h2 g. rewrite <- countable_mono_ext_crit_iff_p1'; auto. apply mono_extension_criterion_unq_p1; auto. Qed. Corollary finite_mono_extension_criterion_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} (i_E:{n:nat | 1 <= n <= card_fun1 E}->sig_set E), Finite E -> ~degen A -> ~degen_p Bp -> surjective i_E -> ba_p_ens Bp = Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E) -> forall (g:sig_set E->bt A), (exists! h:(btp Bp)->(bt A), monomorphism_p1 _ h /\ extends_sig1 h g) <-> (forall (n:{m:nat | 1 <= m <= card_fun1 E}) (a:Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i => (proj1_sig (i_E (exist (fun x=> 1 <= x <= card_fun1 E) _ (match (proj2_sig i) with | conj P Q => conj P (le_trans _ _ _ Q (match (proj2_sig n) with | conj P Q => Q end)) end)))))) signe mns), el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0). intros T Bp A E i_E h0 h1 h2 h3 h4 g. pose proof (finite_inh_or_empty _ h0) as h5. destruct h5 as [h5 | h5]. pose proof (card_fun1_ge_1 _ h5 h0) as h6. pose (fun n:{m:nat | 1 <= m} => match (le_lt_dec (proj1_sig n) (card_fun1 E)) with | left P => i_E (exist _ _ (conj (proj2_sig n) P)) | _ => i_E (exist _ _ (conj (le_refl 1) h6)) end) as i_E'. assert (h7:surjective i_E'). red. intros y. red in h3. specialize (h3 y). destruct h3 as [m h3]. subst. destruct m as [m h3a]. destruct h3a as [h3a h3b]. exists (exist _ _ h3a). unfold i_E'. simpl. destruct le_lt_dec as [h7 | h7]. f_equal. f_equal. apply proof_irrelevance. omega. pose proof (countable_mono_extension_criterion_p1' i_E' h1 h2 h7 h4 g) as h8. unfold i_E' in h8. simpl in h8. rewrite h8. split. intros h9 m a. destruct m as [m h11]. destruct h11 as [h11 h12]. simpl in a. specialize (h9 (exist _ _ h11)). assert (h13:Included (Im (Full_set (seg_one_t (proj1_sig (exist (fun m0 : nat => 1 <= m0) m h11)))) (fun i : seg_one_t (proj1_sig (exist (fun m0 : nat => 1 <= m0) m h11)) => proj1_sig match le_lt_dec (proj1_sig i) (card_fun1 E) with | left P => i_E (exist (fun n : nat => 1 <= n <= card_fun1 E) (proj1_sig i) (conj (seg_one_t_pos (proj1_sig (exist (fun m0 : nat => 1 <= m0) m h11)) i) P)) | right _ => i_E (exist (fun n : nat => 1 <= n <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end)) (Im (Full_set (seg_one_t m)) (fun i : seg_one_t m => proj1_sig (i_E (exist (fun x : nat => 1 <= x <= card_fun1 E) (proj1_sig i) match proj2_sig i with | conj P Q => conj P (le_trans (proj1_sig i) m (card_fun1 E) Q h12) end))))). red. intros x h13. destruct h13 as [x h13]. subst. clear h13. simpl in x. apply Im_intro with x. constructor. f_equal. destruct le_lt_dec as [h14 | h14]. f_equal. apply proj1_sig_injective. simpl. reflexivity. destruct x as [x h15]. simpl in h14. omega. pose (restriction a h13) as a'. specialize (h9 a'). assert (h15:el_prod_p _ a = el_prod_p _ a'). unfold el_prod_p. simpl. apply times_set_functional_p. do 2 rewrite im_im. apply im_ext_in. intros x h16. unfold a'. rewrite restriction_compat. f_equal. f_equal. destruct le_lt_dec as [h18 | h18]. f_equal. apply proj1_sig_injective. simpl. reflexivity. destruct x as [x h17]. simpl in h18. omega. f_equal. f_equal. destruct le_lt_dec as [h18 | h18]. f_equal. apply proj1_sig_injective. simpl. reflexivity. destruct x as [x h17]. simpl in h18. omega. simpl. apply Im_intro with x. constructor. reflexivity. assert (h16: el_prod_compose_p1 T (sig_fun_app g 0) a' = el_prod_compose_p1 T (sig_fun_app g 0) a). unfold el_prod_compose_p1. simpl. apply times_set_functional. do 2 rewrite im_im. apply im_ext_in. intros x h16. unfold a'. rewrite restriction_compat. f_equal. f_equal. destruct le_lt_dec as [h18 | h18]. f_equal. apply proj1_sig_injective. f_equal. f_equal. apply proj1_sig_injective; auto. destruct x as [x h17]. simpl in h18. omega. f_equal. destruct le_lt_dec as [h18 | h18]. f_equal. apply proj1_sig_injective. f_equal. f_equal. apply proj1_sig_injective; auto. destruct x as [x h17]. simpl in h18. omega. simpl. apply Im_intro with x; auto. rewrite h15 at 1. rewrite <- h16 at 1. assumption. intros h9 m a. assert (h12: 1 <= card_fun1 E <= card_fun1 E). split; auto. pose (match (le_lt_dec (proj1_sig m) (card_fun1 E)) with | left P => exist (fun x => 1 <= x <= card_fun1 E) _ (conj (proj2_sig m) P) | _ => exist _ _ h12 end) as n. specialize (h9 n). assert (h13:Included (Im (Full_set (seg_one_t (proj1_sig n))) (fun i : seg_one_t (proj1_sig n) => proj1_sig (i_E (exist (fun x : nat => 1 <= x <= card_fun1 E) (proj1_sig i) match proj2_sig i with | conj P Q => conj P (le_trans (proj1_sig i) (proj1_sig n) (card_fun1 E) Q match proj2_sig n with | conj _ Q0 => Q0 end) end)))) (Im (Full_set (seg_one_t (proj1_sig m))) (fun i : seg_one_t (proj1_sig m) => proj1_sig match le_lt_dec (proj1_sig i) (card_fun1 E) with | left P => i_E (exist (fun n : nat => 1 <= n <= card_fun1 E) (proj1_sig i) (conj (seg_one_t_pos (proj1_sig m) i) P)) | right _ => i_E (exist (fun n : nat => 1 <= n <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end))). red. intros x h14. destruct h14 as [x h14]. subst. clear h14. simpl. unfold n in x. destruct le_lt_dec as [h15 | h15]. simpl in x. apply Im_intro with x. constructor. f_equal. destruct le_lt_dec as [h13 | h13]. f_equal. apply proj1_sig_injective. simpl. reflexivity. destruct x as [x h14]. simpl in h13. omega. simpl in x. destruct x as [x h16]. assert (h17:1 <= x <= (proj1_sig m)). omega. apply Im_intro with (exist (fun c => 1 <= c <= (proj1_sig m)) _ h17). constructor. simpl. destruct h16 as [h16a h16b]. destruct h12 as [h12a h12b]. destruct le_lt_dec as [h18 | h18]. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. f_equal. f_equal. apply proj1_sig_injective. simpl. omega. pose (restriction a h13) as a'. specialize (h9 a'). assert (h15:el_prod_p _ a = el_prod_p _ a'). unfold el_prod_p. simpl. apply times_set_functional_p. do 2 rewrite im_im. unfold n. destruct le_lt_dec as [hl | hl]. simpl. apply im_ext_in. intros x h16. unfold a'. rewrite restriction_compat. f_equal. f_equal. destruct le_lt_dec as [h18 | h18]. f_equal. apply proj1_sig_injective. simpl. reflexivity. destruct x as [x h17]. simpl in h18. omega. f_equal. f_equal. destruct le_lt_dec as [h18 | h18]. f_equal. apply proj1_sig_injective. simpl. reflexivity. destruct x as [x h17]. simpl in h18. omega. simpl. apply Im_intro with x. constructor. reflexivity. destruct m as [m hm]. simpl. simpl in hl. rewrite (im_seg_one_t_decompose (card_fun1 E) m _ hl). simpl. assert (h19: Im (Full_set {r : nat | card_fun1 E < r <= m}) (restriction_sig_prop (fun a0 : nat => card_fun1 E < a0 <= m) (fun a0 : nat => 1 <= a0 <= m) (lt_le_prop_seg_one_t (card_fun1 E) m hl) (fun x : seg_one_t m => eps_p (proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end) (a |-> proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end))) = Singleton (eps_p (proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)))) (a |-> proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)))))). apply Extensionality_Ensembles; red; split; intros x h20. destruct h20 as [x h20]. subst. clear h20. destruct x as [x h20]. simpl. destruct le_lt_dec as [h21 | h21]. omega. constructor. destruct h20. assert (h21:card_fun1 E < m <= m). omega. apply Im_intro with (exist (fun r => card_fun1 E < r <= m) _ h21). constructor. simpl. f_equal. f_equal. destruct le_lt_dec as [h22 | h22]. omega. reflexivity. f_equal. f_equal. destruct le_lt_dec as [h22 | h22]. omega. reflexivity. rewrite h19 at 1. assert (h20:Included (Singleton (eps_p (proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)))) (a |-> proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)))))) (Im (Full_set (seg_one_t (card_fun1 E))) (restriction_sig_prop (fun a0 : nat => 1 <= a0 <= card_fun1 E) (fun a0 : nat => 1 <= a0 <= m) (lt_prop_seg_one_t (card_fun1 E) m hl) (fun x : seg_one_t m => eps_p (proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end) (a |-> proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end))))). red; intros x h23. destruct h23. apply Im_intro with (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)). constructor. unfold restriction_sig_prop. simpl. f_equal. f_equal. destruct (card_fun1 E) as [|h24]. reflexivity. f_equal. apply proj1_sig_injective; auto. f_equal. f_equal. destruct (card_fun1 E) as [|h24]. reflexivity. f_equal. apply proj1_sig_injective; auto. pose proof (inclusion_iff_union (Singleton (eps_p (proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)))) (a |-> proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)))))) (Im (Full_set (seg_one_t (card_fun1 E))) (restriction_sig_prop (fun a0 : nat => 1 <= a0 <= card_fun1 E) (fun a0 : nat => 1 <= a0 <= m) (lt_prop_seg_one_t (card_fun1 E) m hl) (fun x : seg_one_t m => eps_p (proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end) (a |-> proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end))))) as h25. pose proof (iff1 h25 h20) as h26. clear h25 h20. rewrite Union_commutative. rewrite h26 at 1. apply im_ext_in. intros x h27. clear h27. destruct x as [x h27]. unfold a'. rewrite restriction_compat. simpl. f_equal. f_equal. destruct le_lt_dec as [h28 | h28]. f_equal. apply proj1_sig_injective; auto. omega. destruct le_lt_dec as [h24 | h24]. f_equal. f_equal. f_equal. apply proj1_sig_injective; auto. omega. simpl. apply Im_intro with (exist (fun r => 1 <= r <= (card_fun1 E)) _ h27). constructor. f_equal. assert (h16:el_prod_compose_p1 T (sig_fun_app g 0) a = el_prod_compose_p1 T (sig_fun_app g 0) a'). unfold el_prod_compose_p1. simpl. apply times_set_functional. do 2 rewrite im_im. unfold n. destruct le_lt_dec as [hl | hl]. simpl. apply im_ext_in. intros x h16. unfold a'. rewrite restriction_compat. f_equal. f_equal. destruct le_lt_dec as [h18 | h18]. f_equal. apply proj1_sig_injective. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. destruct x as [x h17]. simpl in h18. omega. f_equal. f_equal. destruct le_lt_dec as [h18 | h18]. f_equal. apply proj1_sig_injective. simpl. reflexivity. destruct x as [x h17]. simpl in h18. omega. simpl. apply Im_intro with x. constructor. reflexivity. destruct m as [m hm]. simpl. simpl in hl. rewrite (im_seg_one_t_decompose (card_fun1 E) m _ hl). simpl. assert (h19:Im (Full_set {r : nat | card_fun1 E < r <= m}) (restriction_sig_prop (fun a0 : nat => card_fun1 E < a0 <= m) (fun a0 : nat => 1 <= a0 <= m) (lt_le_prop_seg_one_t (card_fun1 E) m hl) (fun x : seg_one_t m => eps (sig_fun_app g 0 (proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end)) (a |-> proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end))) = Singleton (eps (sig_fun_app g 0 (proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6))))) (a |-> proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)))))). apply Extensionality_Ensembles; red; split; intros x h20. destruct h20 as [x h20]. subst. clear h20. destruct x as [x h20]. simpl. destruct le_lt_dec as [h21 | h21]. omega. constructor. destruct h20. assert (h21:card_fun1 E < m <= m). omega. apply Im_intro with (exist (fun r => card_fun1 E < r <= m) _ h21). constructor. simpl. f_equal. f_equal. destruct le_lt_dec as [h22 | h22]. omega. reflexivity. f_equal. f_equal. destruct le_lt_dec as [h22 | h22]. omega. reflexivity. rewrite h19 at 1. assert (h20:Included (Singleton (eps (sig_fun_app g 0 (proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6))))) (a |-> proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)))))) (Im (Full_set (seg_one_t (card_fun1 E))) (restriction_sig_prop (fun a0 : nat => 1 <= a0 <= card_fun1 E) (fun a0 : nat => 1 <= a0 <= m) (lt_prop_seg_one_t (card_fun1 E) m hl) (fun x : seg_one_t m => eps (sig_fun_app g 0 (proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end)) (a |-> proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1%nat) h6)) end))))). red; intros x h23. destruct h23. apply Im_intro with (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)). constructor. simpl. f_equal. f_equal. destruct (card_fun1 E) as [|c]. reflexivity. f_equal. apply proj1_sig_injective; auto. f_equal. f_equal. apply proj1_sig_injective; auto. f_equal. apply proj1_sig_injective; auto. f_equal. f_equal. destruct (card_fun1 E) as [|c]. reflexivity. f_equal. apply proj1_sig_injective; auto. pose proof (inclusion_iff_union (Singleton (eps (sig_fun_app g 0 (proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6))))) (a |-> proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)))))) (Im (Full_set (seg_one_t (card_fun1 E))) (restriction_sig_prop (fun a0 : nat => 1 <= a0 <= card_fun1 E) (fun a0 : nat => 1 <= a0 <= m) (lt_prop_seg_one_t (card_fun1 E) m hl) (fun x : seg_one_t m => eps (sig_fun_app g 0 (proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1) h6)) end)) (a |-> proj1_sig match le_lt_dec (proj1_sig x) (card_fun1 E) with | left P => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig x) (conj (seg_one_t_pos m x) P)) | right _ => i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) 1%nat (conj (le_refl 1%nat) h6)) end))))) as h25. pose proof (iff1 h25 h20) as h26. clear h25 h20. rewrite Union_commutative. rewrite h26 at 1. apply im_ext_in. intros x h27. clear h27. destruct x as [x h27]. unfold a'. rewrite restriction_compat. simpl. f_equal. f_equal. f_equal. destruct le_lt_dec as [h28 | h28]. f_equal. apply proj1_sig_injective; auto. omega. destruct le_lt_dec as [h24 | h24]. f_equal. f_equal. f_equal. apply proj1_sig_injective; auto. omega. simpl. apply Im_intro with (exist (fun r => 1 <= r <= (card_fun1 E)) _ h27). constructor. f_equal. rewrite h15, h16. assumption. subst. rewrite gen_ens_p_eq in h4. rewrite <- im_proj2_sig_undoes_im_proj1_sig' in h4 at 1. rewrite gen_ens_empty in h4. split. intros h5 n ha. destruct n as [n h6]. pose proof h6 as h6'. rewrite card_fun1_empty in h6'. destruct h6' as [h7 h8]. omega. intro h5. assert (h6:forall x:btp Bp, {x = %0} + {x = %1}). intro x. destruct x as [x h6]. unfold ba_p_ens in h4. pose proof h6 as h7. rewrite h4 in h7. assert (h8:x = proj1_sig (exist _ _ h6)). auto. rewrite h8 in h7. rewrite in_im_proj1_sig_iff in h7 at 1. apply couple_inv_dec in h7. destruct h7 as [h7 | h7]. left. apply proj1_sig_injective. simpl. pose proof (f_equal (@proj1_sig _ _) h7) as h9. simpl in h9. assumption. right. apply proj1_sig_injective. simpl. pose proof (f_equal (@proj1_sig _ _) h7) as h9. simpl in h9. assumption. simpl in h4. pose (fun x:btp Bp => if (h6 x) then (Bzero (Bc A)) else (Bone (Bc A))) as k. exists (trivial_ba_p_mono T Bp A h4). red. split. split. apply mono_trivial_ba_p_mono; auto. red. intro x. destruct x as [x h7]. destruct h7. intros f h7. destruct h7 as [h7 h8]. apply unq_trivial_ba_p_mono; auto. Qed. Corollary finite_mono_extension_criterion_p1' : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} (i_E:{n:nat | 1 <= n <= card_fun1 E}->sig_set E), Finite E -> ~degen A -> ~degen_p Bp -> surjective i_E -> let Cp := Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E) in forall (g:sig_set E->bt A), (exists! h:(btp Cp)->(bt A), monomorphism_p1 _ h /\ extends_sig_sig h g) <-> (forall (n:{m:nat | 1 <= m <= card_fun1 E}) (a:Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i => (proj1_sig (i_E (exist (fun x=> 1 <= x <= card_fun1 E) _ (match (proj2_sig i) with | conj P Q => conj P (le_trans _ _ _ Q (match (proj2_sig n) with | conj P Q => Q end)) end)))))) signe mns), el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0). intros T Bp A E i_E h1 h2 h3 h4 Cp g. pose proof (gen_p_subalg_of_p_compat _ (incl_im_proj1_sig_ba_p_ens E)) as h6. pose proof (degen_subalg_of_compat_p _ _ h6) as h5. rewrite <- h5 in h3. pose proof (gen_ens_includes_p _ (incl_im_proj1_sig_ba_p_ens E)) as h7. rewrite <- ba_p_ens_gen_p_eq in h7. pose (im_proj2_sig _ h7) as E'. assert (h1':Finite E'). unfold E'. unfold im_proj2_sig, im_proj1_sig. apply finite_image. rewrite <- finite_full_sig_iff. apply finite_image; auto. assert (h8:forall x:sig_set E, Inn (Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)) (proj1_sig (proj1_sig x))). intro x. destruct x as [x ha]. simpl. apply h7. rewrite in_im_proj1_sig_iff; auto. rewrite <- ba_p_ens_gen_p_eq in h8. assert (h9:forall x:sig_set E, Inn E' (exist _ _ (h8 x))). intro x. destruct x as [x h9]. unfold E', im_proj2_sig. assert (h10:Inn (im_proj1_sig E) (proj1_sig x)). apply Im_intro with x; auto. apply Im_intro with (exist _ _ h10). constructor. apply proj1_sig_injective; auto. pose (fun m:{n : nat | 1 <= n <= card_fun1 E} => (exist _ _ (h9 (i_E m)))) as i_E'. assert (h10:card_fun1 E = card_fun1 E'). unfold E'. erewrite <- card_fun1_im_proj2_sig; auto. rewrite <- card_fun1_im_proj1_sig; auto. apply finite_image; auto. assert (h15:forall n:nat, 1 <= n <= card_fun1 E' -> 1 <= n <= card_fun1 E). intros; omega. assert (h15':forall n:nat, 1 <= n <= card_fun1 E -> 1 <= n <= card_fun1 E'). intros; omega. pose (fun m:{n:nat | 1 <= n <= card_fun1 E'} => i_E' (exist _ _ (h15 _ (proj2_sig m)))) as i_E''. assert (h11:surjective i_E''). red. intro y. destruct y as [y ha]. destruct ha as [y ha]. subst. destruct y as [y hb]. simpl. destruct hb as [y hb]. subst. pose proof (h4 (exist _ _ hb)) as hc. destruct hc as [m hc]. exists (exist (fun s => 1 <= s<= card_fun1 E') _ (h15' _ (proj2_sig m))). unfold i_E'', i_E'. apply proj1_sig_injective; apply proj1_sig_injective. simpl. destruct y as [y hd]. simpl. assert (he:exist _ _ (h15 _ (h15' _ (proj2_sig m))) = m). apply proj1_sig_injective. simpl. reflexivity. rewrite he. rewrite hc. simpl. reflexivity. assert (h12:ba_p_ens (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)) = Gen_Ens_p (im_proj1_sig E') (incl_im_proj1_sig_ba_p_ens E')). unfold E'. rewrite ba_p_ens_gen_p_eq at 1. erewrite <- gen_ens_p_subalg_of_p. apply gen_ens_p_functional. rewrite <- im_proj1_sig_undoes_im_proj2_sig. reflexivity. apply gen_p_subalg_of_p_compat. assert (h8':forall x:sig_set E', Inn (im_proj1_sig E) (proj1_sig (proj1_sig x))). intro x. destruct x as [x ha]. simpl. destruct x as [x hb]. simpl. inversion ha as [x' hc hd he]. subst. clear hc. apply exist_injective in he. subst. destruct x' as [x'' hc]. simpl. assumption. assert (h8'':forall x:sig_set E', Inn E (exist _ _ (incl_im_proj1_sig_ba_p_ens E _ (h8' x)))). intro x. destruct x as [x ha]. simpl. destruct x as [x hb]. simpl. inversion ha as [x' hc hd he]. subst. apply exist_injective in he. subst. destruct x' as [x'' hf]. simpl. destruct hf as [x'' hf]. subst. assert (hg:x'' = (exist (Inn (ba_p_ens Bp)) (proj1_sig x'') (incl_im_proj1_sig_ba_p_ens E (proj1_sig x'') (h8' (exist (fun x : {x : T | Inn (ba_p_ens (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E))) x} => Inn E' x) (exist (fun x : T => Inn (ba_p_ens (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E))) x) (proj1_sig x'') hb) ha))))). apply proj1_sig_injective; auto. rewrite <- hg; auto. pose (fun x:sig_set E' => (g (exist _ _ (h8'' x)))) as g'. pose proof (finite_mono_extension_criterion_p1 (A:=A) i_E'' h1' h2 h3 h11 h12 g') as h16. unfold Cp. split. intro h17. destruct h17 as [h h17]. assert (h18: (exists ! h : btp (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)) -> bt A, monomorphism_p1 T h /\ extends_sig1 h g')). exists h. red in h17. red. destruct h17 as [h17 h18]. destruct h17 as [h17a h17b]. split. split; auto. red in h17b. red. destruct h17b as [h19 h20]. intro x. destruct x as [x h21]. simpl. destruct h21 as [x h21]. subst. specialize (h20 x). unfold g'. unfold ba_p_ens. assert (h22:h19 _ (proj2_sig x) = h7 _ (proj2_sig x)). apply proof_irrelevance. rewrite h22 in h20. rewrite <- h20. f_equal. apply proj1_sig_injective; simpl. apply proj1_sig_injective. simpl. reflexivity. intros h' h19. apply h18. destruct h19 as [h19 h20]. split; auto. red. red in h20. exists h7. intro x. destruct x as [x h21]. pose proof (h7 _ h21) as h22. assert (h23:Inn E' (exist _ _ h22)). unfold E', im_proj1_sig, im_proj2_sig. rewrite im_full_sig_im_eq. simpl. destruct h21 as [x h21]. subst. apply Im_intro with (exist _ _ h21). constructor. apply proj1_sig_injective; auto. specialize (h20 (exist _ _ h23)). simpl in h20. simpl. assert (h24:h' (exist _ _ h22) = h' (exist _ _ (h7 x h21))). f_equal. apply proj1_sig_injective; auto. rewrite <- h24 at 1. rewrite <- h20. unfold g'. f_equal. apply proj1_sig_injective; simpl. apply proj1_sig_injective; auto. rewrite h16 in h18. intros n a. destruct n as [n h19]. simpl in a. assert (h20:1 <= n <= card_fun1 E'). rewrite <- h10. assumption. specialize (h18 (exist _ _ h20)). assert (h21: Finite (ba_p_ens (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)))). rewrite ba_p_ens_gen_p_eq. apply finite_fin_gen_p. apply finite_image; auto. pose (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun x : nat => 1 <= x <= card_fun1 E) (proj1_sig i) match proj2_sig i with | conj P Q => conj P (le_trans (proj1_sig i) n (card_fun1 E) Q match h19 with | conj _ Q0 => Q0 end) end)))) as S. pose (Im (Full_set (seg_one_t (proj1_sig (exist (fun m : nat => 1 <= m <= card_fun1 E') n h20)))) (fun i : seg_one_t (proj1_sig (exist (fun m : nat => 1 <= m <= card_fun1 E') n h20)) => proj1_sig (i_E'' (exist (fun x : nat => 1 <= x <= card_fun1 E') (proj1_sig i) match proj2_sig i with | conj P Q => conj P (le_trans (proj1_sig i) (proj1_sig (exist (fun m : nat => 1 <= m <= card_fun1 E') n h20)) (card_fun1 E') Q match proj2_sig (exist (fun m : nat => 1 <= m <= card_fun1 E') n h20) with | conj _ Q0 => Q0 end) end)))) as S'. assert (h22:forall x, Inn S' x -> Inn (ba_p_ens Bp) (proj1_sig x)). intros x h23. destruct x as [x h24]. simpl. pose proof h24 as h24'. rewrite ba_p_ens_gen_p_eq in h24'. apply incl_gen_ens_p in h24'. assumption. assert (h23:Finite S'). unfold S. simpl. apply finite_image. apply finite_full_set_seg_one_t. assert (h25:Included (Im (full_sig S') (fun x : sig_set S' => a |-> exist (Inn (ba_p_ens Bp)) (proj1_sig (proj1_sig x)) (h22 (proj1_sig x) (proj2_sig x)))) signe). red; intros x h26. destruct h26 as [x h26]. subst. destruct (a |-> exist (Inn (ba_p_ens Bp)) (proj1_sig (proj1_sig x)) (h22 (proj1_sig x) (proj2_sig x))). left. right. simpl in h18. pose (sig_fun_to_fin_map_ran (@ba_eq_dec_p _ (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E))) (fun x:sig_set S' => a |-> (exist _ _ (h22 _ (proj2_sig x)))) h23 mns signe signe_finite h25) as a'. specialize (h18 a'). assert (hie: im_proj1_sig S' = im_proj1_sig S). unfold S', S, im_proj1_sig. do 2 rewrite im_im. simpl. apply im_ext_in. intros x ha. f_equal. f_equal. f_equal. apply proj1_sig_injective; auto. assert (h26: forall x : btp (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)), Inn (ba_p_ens Bp) (proj1_sig x)). intro x. destruct x as [x ha]. simpl. eapply incl_gen_ens_p. rewrite <- ba_p_ens_gen_p_eq. apply ha. assert (h27: forall x : btp (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)), Inn S' x -> a' |-> x = a |-> exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h26 x)). intros x ha. unfold a'. pose proof (sig_fun_to_fin_map_ran_compat (@ba_eq_dec_p _ (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E))) (fun x0 : sig_set S' => a |-> exist (Inn (ba_p_ens Bp)) (proj1_sig (proj1_sig x0)) (h22 (proj1_sig x0) (proj2_sig x0))) h23 mns signe signe_finite h25 x ha) as hb. simpl in hb. rewrite hb at 1. f_equal. apply proj1_sig_injective. simpl. reflexivity. pose proof (el_prod_compat_fin_map_eq_base h6 S' S hie a' a h26 h27) as h28. assert (h29:forall x:sig_set E', Inn (ba_p_ens (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E))) (proj1_sig (proj1_sig x))). intro x. destruct x as [x ha]. simpl. apply proj2_sig. assert (h30:Included S E). red. intros x h31. destruct h31 as [x h31]. subst. apply proj2_sig. assert (h31: forall x : btp (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)), Inn S' x -> sig_fun_app g' 0 x = sig_fun_app g 0 (exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h26 x))). intros x ha. unfold sig_fun_app. destruct classic_dec as [hb | hb]. destruct classic_dec as [hc | hc]. unfold g'. f_equal. apply proj1_sig_injective; simpl. simpl. apply proj1_sig_injective; auto. contradict hc. destruct hb as [x hb]. subst. simpl. apply h30. unfold S. inversion ha as [x' hc hd he]. subst. simpl in hc. apply Im_intro with x'. constructor. apply proj1_sig_injective. simpl. pose proof (f_equal (@proj1_sig _ _) he) as hf. simpl in hf. rewrite hf. f_equal. f_equal. f_equal. apply proj1_sig_injective; auto. destruct classic_dec as [hc | hc]. pose proof (h9 (exist _ _ hc)) as h9'. assert (he:x = exist _ _ (h8 (exist _ _ hc))). apply proj1_sig_injective; auto. rewrite <- he in h9'. contradiction. reflexivity. pose proof (el_prod_compose_compat_fin_map_eq_base A h6 g (sig_fun_app g' 0) S' S hie h30 h26 h31 a' a h27) as h32. unfold S in h32. rewrite h32 at 1. assert (h33:el_prod_p _ a = %0 <-> el_prod_p _ a' = %0). split. intro ha. pose proof (f_equal (@proj1_sig _ _) ha) as ha'. unfold S in h28, ha'. rewrite <- h28 in ha' at 1. apply proj1_sig_injective; auto. intro ha. pose proof (f_equal (@proj1_sig _ _) ha) as ha'. unfold S in h28, ha'. rewrite h28 in ha'. apply proj1_sig_injective; auto. rewrite h33. assumption. intro h17. assert (h18: (forall (n : {m : nat | 1 <= m <= card_fun1 E'}) (a : Fin_map (Im (Full_set (seg_one_t (proj1_sig n))) (fun i : seg_one_t (proj1_sig n) => proj1_sig (i_E'' (exist (fun x : nat => 1 <= x <= card_fun1 E') (proj1_sig i) match proj2_sig i with | conj P Q => conj P (le_trans (proj1_sig i) (proj1_sig n) (card_fun1 E') Q match proj2_sig n with | conj _ Q0 => Q0 end) end)))) signe mns), el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g' 0) a = 0)). intros n a. destruct n as [n ha]. assert (ha':1 <= n <= card_fun1 E). rewrite h10. assumption. specialize (h17 (exist _ _ ha')). simpl in a. pose (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => exist (Inn (ba_p_ens (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)))) (proj1_sig (proj1_sig (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig i) (h15 (proj1_sig i) match proj2_sig i with | conj P Q => conj P (le_trans (proj1_sig i) n (card_fun1 E') Q match ha with | conj _ Q0 => Q0 end) end))))) (h8 (i_E (exist (fun n0 : nat => 1 <= n0 <= card_fun1 E) (proj1_sig i) (h15 (proj1_sig i) match proj2_sig i with | conj P Q => conj P (le_trans (proj1_sig i) n (card_fun1 E') Q match ha with | conj _ Q0 => Q0 end) end)))))) as S. simpl in h17. pose (Im (Full_set (seg_one_t n)) (fun i : seg_one_t n => proj1_sig (i_E (exist (fun x : nat => 1 <= x <= card_fun1 E) (proj1_sig i) match proj2_sig i with | conj P Q => conj P (le_trans (proj1_sig i) n (card_fun1 E) Q match ha' with | conj _ Q0 => Q0 end) end)))) as S'. assert (hd:forall x, Inn S' x -> Inn (ba_p_ens (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E))) (proj1_sig x)). intros x hb. rewrite ba_p_ens_gen_p_eq. destruct hb as [x hb]. subst. destruct x as [x hc]. simpl. apply gen_ens_includes_p. apply Im_intro with (proj1_sig (i_E (exist (fun x0 : nat => 1 <= x0 <= card_fun1 E) x match hc with | conj P Q => conj P (le_trans x n (card_fun1 E) Q match ha' with | conj _ Q0 => Q0 end) end))). apply proj2_sig. reflexivity. assert (he:Finite S'). unfold S'. apply finite_image. apply finite_full_set_seg_one_t. assert (hf: Included (Im (full_sig S') (fun x : sig_set S' => a |-> exist (Inn (ba_p_ens (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)))) (proj1_sig (proj1_sig x)) (hd (proj1_sig x) (proj2_sig x)))) signe). red; intros x ha1. destruct x; constructor. pose (sig_fun_to_fin_map_ran (@ba_eq_dec_p _ Bp) (fun x:sig_set S' => a |-> (exist _ _ (hd _ (proj2_sig x)))) he mns signe signe_finite hf) as a'. specialize (h17 a'). assert (hie: im_proj1_sig S' = im_proj1_sig S). unfold S', S, im_proj1_sig. do 2 rewrite im_im. simpl. apply im_ext_in. intros x ha1. f_equal. f_equal. f_equal. apply proj1_sig_injective; auto. symmetry in hie. assert (h26: forall x : btp (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)), Inn (ba_p_ens Bp) (proj1_sig x)). intro x. destruct x as [x ha1]. simpl. eapply incl_gen_ens_p. rewrite <- ba_p_ens_gen_p_eq. apply ha1. assert (h27: forall x : btp (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)), Inn S x -> a |-> x = a' |-> exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h26 x)). intros x ha1. unfold a'. assert (ha3:Inn S' (exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h26 x))). unfold S'. destruct ha1 as [x ha1]. subst. simpl. apply Im_intro with x. constructor. apply proj1_sig_injective; simpl. f_equal. f_equal. f_equal. apply proj1_sig_injective; simpl;auto. pose proof (sig_fun_to_fin_map_ran_compat (@ba_eq_dec_p _ Bp)(fun x0 : sig_set S' => a |-> exist (Inn (ba_p_ens (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)))) (proj1_sig (proj1_sig x0)) (hd (proj1_sig x0) (proj2_sig x0))) he mns signe signe_finite hf (exist _ _ (h26 x)) ha3) as ha2. rewrite ha2. f_equal. apply proj1_sig_injective; auto. pose proof (el_prod_compat_fin_map_eq_base h6 S S' hie a a' h26 h27) as hepc. assert (h30:Included S' E). red; intros x ha1. destruct ha1 as [x ha1]. subst. apply proj2_sig. assert (h31: forall x : btp (Gen_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E)), Inn S x -> sig_fun_app g' 0 x = sig_fun_app g 0 (exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h26 x))). intros x ha1. unfold sig_fun_app. destruct classic_dec as [hb | hb]. destruct classic_dec as [hc | hc]. unfold g'. f_equal. apply proj1_sig_injective; simpl. simpl. apply proj1_sig_injective; auto. contradict hc. destruct hb as [x hb]. subst. simpl. apply h30. unfold S'. inversion ha1 as [x' hc hd1 he1]. subst. simpl in hc. apply Im_intro with x'. constructor. apply proj1_sig_injective. simpl. pose proof (f_equal (@proj1_sig _ _) he1) as hf1. simpl in hf1. rewrite hf1. f_equal. f_equal. f_equal. apply proj1_sig_injective; auto. destruct classic_dec as [hc | hc]. pose proof (h9 (exist _ _ hc)) as h9'. assert (he1:x = exist _ _ (h8 (exist _ _ hc))). apply proj1_sig_injective; auto. rewrite <- he1 in h9'. contradiction. reflexivity. pose proof (el_prod_compose_compat_fin_map_eq_base A h6 g (sig_fun_app g' 0) S S' hie h30 h26 h31 a a' h27) as h32. unfold S, S' in h32. unfold S, S'. simpl in h32. simpl. rewrite <- h32 at 1. split. intro ha2. pose proof (f_equal (@proj1_sig _ _) ha2) as h33. simpl in h33. rewrite <- h17. apply proj1_sig_injective; simpl. unfold S, S' in hepc. unfold S, S'. rewrite <- hepc at 1. assumption. intro h33. apply proj1_sig_injective; simpl. unfold S, S' in hepc. rewrite hepc at 1. rewrite <- h17 in h33. f_equal. assumption. rewrite <- h16 in h18. destruct h18 as [h h18]. red in h18. destruct h18 as [h18 h19]. destruct h18 as [h18 h20]. exists h. red. split. split; auto. red in h20. red. exists h7. intro x. pose E'. assert (h21:forall x:sig_set (im_proj1_sig E), Inn (ba_p_ens (Gen_p _ (incl_im_proj1_sig_ba_p_ens E))) (proj1_sig x)). intro a. rewrite ba_p_ens_gen_p_eq. eapply gen_ens_includes_p. apply proj2_sig. assert (h22:forall x:sig_set (im_proj1_sig E), Inn E' (exist _ _ (h21 x))). intro c. unfold E'. apply Im_intro with c. constructor. apply proj1_sig_injective; auto. specialize (h20 (exist _ _ (h22 x))). simpl in h20. assert (h24:h (exist _ _ (h21 x)) = h (exist _ _ (h7 _ (proj2_sig x)))). f_equal. apply proj1_sig_injective; auto. rewrite <- h24 at 1. rewrite <- h20. unfold g'. f_equal. apply proj1_sig_injective; simpl. apply proj1_sig_injective; auto. intros h' h21. destruct h21 as [h21 h22]. apply h19. split; auto. red in h22. red. destruct h22 as [h22 h23]. intro x. assert (h24:Inn (im_proj1_sig E) (proj1_sig (proj1_sig x))). destruct x as [x h25]. simpl. destruct h25 as [x h25]. subst. simpl. destruct x as [x h26]. simpl. assumption. specialize (h23 (exist _ _ h24)). simpl in h23. assert (h25:exist _ _ (h22 _ h24) = proj1_sig x). apply proj1_sig_injective; auto. rewrite <- h25. rewrite <- h23 at 1. unfold g'. f_equal. apply proj1_sig_injective; simpl. apply proj1_sig_injective; auto. Grab Existential Variables. apply h7. Qed. Corollary finite_mono_extension_criterion_seg_p1 : forall {T:Type} {Bp:Bool_Alg_p T} {A:Bool_Alg} {E:Ensemble (btp Bp)} (i_E:(card_fun1 E) @->sig_set E), Finite E -> ~degen A -> ~degen_p Bp -> surjective i_E -> ba_p_ens Bp = Gen_Ens_p (im_proj1_sig E) (incl_im_proj1_sig_ba_p_ens E) -> forall (g:sig_set E->bt A), (exists! h:(btp Bp)->(bt A), monomorphism_p1 _ h /\ extends_sig1 h g) <-> (forall (n:(card_fun1 E) @) (a:Fin_map (Im (Full_set ((S (proj1_sig n)) @)) (fun i => (proj1_sig (i_E (exist (fun x=> x < card_fun1 E) _ (lt_le_trans _ _ _ (proj2_sig i) (proj2_sig n))))))) signe mns), el_prod_p _ a = %0 <-> el_prod_compose_p1 T (sig_fun_app g 0) a = 0). intros T Bp A E i_E h1 h2 h3 h4 h5 g. pose (segT_fun_to_seg_one i_E) as i_E'. pose proof (finite_mono_extension_criterion_p1 i_E' h1 h2 h3) as h6. hfirst h6. unfold i_E'. rewrite (surj_segT_fun_to_seg_one_iff i_E); auto. specialize (h6 hf h5 g). rewrite h6. split; intros h8 n a. specialize (h8 (st_to_sot n)). let P := type of a in match P with Fin_map ?S' _ _ => pose S' as S end. let P := type of h8 in match P with forall a:Fin_map ?S' _ _, _ => assert (h11:S' = S) end. unfold S. rewrite im_ens_sot. rewrite ens_sot_to_st_full_set. pose proof (segT_functional (proj1_sig (st_to_sot n)) (Datatypes.S (proj1_sig n))) as h10. hfirst h10. rewrite st_to_sot_compat; auto. specialize (h10 hf0). match goal with |- Im _ ?f' = Im _ _ => pose f' as k end. match goal with |- Im _ _ = Im _ ?g' => pose g' as h end. pose proof (im_ext_full_hetero_set1 h10 k h) as h11. unfold k, h in h11. simpl in h11; simpl. apply h11. unfold compose. intro x. destruct x as [x hx]; simpl. repeat f_equal. apply proj1_sig_injective; simpl. rewrite <- proj1_segT_transfer; simpl; auto. unfold S in h11. symmetry in h11. specialize (h8 (fin_map_dom_subst h11 a)). rewrite <- el_prod_fin_map_dom_subst_compat_p in h8. rewrite <- el_prod_compose_fin_map_dom_subst_compat_p1 in h8; auto. specialize (h8 (sot_to_st n)). let P := type of a in match P with Fin_map ?S' _ _ => pose S' as S end. let P := type of h8 in match P with forall a:Fin_map ?S' _ _, _ => assert (h11:S' = S) end. unfold S. rewrite im_ens_st. rewrite ens_st_to_sot_full_set. pose proof (f_equal Datatypes.S (sot_to_st_compat n)) as he. pose proof (S_pred (proj1_sig n) 0) as hs. hfirst hs. destruct n as [? [hn ?]]; simpl; auto with arith. specialize (hs hf0). rewrite <- hs in he. symmetry in he. pose proof (seg_one_t_functional _ _ hs) as h10. match goal with |- Im _ ?f' = Im _ _ => pose f' as k end. match goal with |- Im _ _ = Im _ ?g' => pose g' as h end. pose proof (eq_sym h10) as h10'. pose proof (im_ext_full_hetero_set1 h10' k h) as h11. unfold k, h in h11. simpl in h11; simpl. apply h11. unfold compose. intro x. f_equal. pose (proj2_sig (s_transfer h10' x)) as hp. pose proof hp as hp'; simpl in hp'. pose proof (proj1_seg_one_t_transfer (eq_sym hs) x h10') as h12. symmetry in h12. rewrite h12 in hp'. pose proof (subsetT_eq_compat _ _ _ _ hp hp' h12) as h13. unfold hp in h13. dependent rewrite h13. simpl. destruct x as [x hx]. simpl. destruct x as [|x]. omega. repeat f_equal. apply proof_irrelevance. pose (fin_map_dom_subst (eq_sym h11) a) as a'. specialize (h8 a'). unfold a' in h8. rewrite <- el_prod_fin_map_dom_subst_compat_p in h8 at 1. rewrite <- el_prod_compose_fin_map_dom_subst_compat_p1 in h8 at 1; auto. Qed. End HomoExtensionCriterion. (*There is insufficient time to translate the proof of Sikorski's Homomorphism Extension Theorem proper, unfortunately, but here are some nice preliminaries! Eventually one needs transfinite induction, and that's when the preliminaries have to be justified!*) Section HomoExtensionTheorem. Inductive one_step_extends {T:Type} (Bp Cp:Bool_Alg_p T) (r:T) : Prop := | one_step_extends_intro : subalg_of_p Cp Bp -> forall (pf:Included (Add (ba_p_ens Cp) r) (ba_p_ens Bp)), Bp = Gen_p _ pf -> one_step_extends Bp Cp r. Lemma one_step_extends_incl : forall {T:Type} (Bp Cp:Bool_Alg_p T) (r:T), one_step_extends Bp Cp r -> Included (Add (ba_p_ens Cp) r) (ba_p_ens Bp). intros T Bp Cp r h1. destruct h1; auto. Qed. Lemma one_step_extends_in : forall {T:Type} (Bp Cp:Bool_Alg_p T) (r:T), one_step_extends Bp Cp r -> Inn (ba_p_ens Bp) r. intros T Bp Cp r h1. destruct h1; auto with sets. Qed. Lemma trivial_step_extends_eq : forall {T:Type} (Bp Cp:Bool_Alg_p T) (r:T), Inn (ba_p_ens Cp) r -> one_step_extends Bp Cp r -> Bp = Cp. intros T Bp Cp r h1 h2. destruct h2 as [h2 h3 h4]. pose proof (in_add_eq _ _ h1) as h5. assert (h6:Included (ba_p_ens Cp) (ba_p_ens Bp)). auto with sets. pose proof (subsetT_eq_compat _ (fun S=>Included S (ba_p_ens Bp)) _ _ h3 h6 h5) as h7. dependent rewrite -> h7 in h4. assert (h8:h6 = subalg_of_p_incl _ _ h2). apply proof_irrelevance. rewrite h8 in h4. rewrite gen_p_subalg_of_p_eq in h4 at 1. assumption. Qed. Definition one_step_extension {T:Type} (Bp Cp:Bool_Alg_p T) (pfinc:Included (ba_p_ens Cp) (ba_p_ens Bp)) (pfs:subalg_of_p Cp Bp) (r:btp Bp) : Bool_Alg_p T := Gen_p (Add (ba_p_ens Cp) (proj1_sig r)) (ba_p_in_add_ens _ _ _ pfinc r). Lemma one_step_extension_compat : forall {T:Type} (Bp Cp:Bool_Alg_p T) (pfinc:Included (ba_p_ens Cp) (ba_p_ens Bp)) (pfs:subalg_of_p Cp Bp) (r:btp Bp), one_step_extends (one_step_extension Bp Cp pfinc pfs r) Cp (proj1_sig r). intros T Bp Cp h1 h2 r. unfold one_step_extension. eapply one_step_extends_intro. pose proof (gen_p_subalg_of_p_compat _ h1). pose proof (subalg_of_p_gen_p2 h1 h2) as h3. assert (h4:Included (ba_p_ens Cp) (Add (ba_p_ens Cp) (proj1_sig r))). auto with sets. pose proof (subalg_of_p_gen_p_gen_p _ _ _ h4 h1 (ba_p_in_add_ens T Bp (ba_p_ens Cp) h1 r)) as h5. pose proof (trans_subalg_of_p T) as h6. red in h6. specialize (h6 _ _ _ h3 h5). assumption. pose proof (gen_p_subalg_of_p_compat _ (ba_p_in_add_ens T Bp (ba_p_ens Cp) h1 r)) as h3. assert (h4:Included (Add (ba_p_ens Cp) (proj1_sig r)) (ba_p_ens (Gen_p (Add (ba_p_ens Cp) (proj1_sig r)) (ba_p_in_add_ens T Bp (ba_p_ens Cp) h1 r)))). rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. pose proof (gen_p_subalg_of_p h3 (Add (ba_p_ens Cp) (proj1_sig r)) h4 (ba_p_in_add_ens T Bp (ba_p_ens Cp) h1 r)) as h5. rewrite <- h5 at 1. apply gen_p_functional. reflexivity. Grab Existential Variables. rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. Qed. Lemma one_step_extends_homo_ext_crit_iff : forall {T:Type} (Bp Cp:Bool_Alg_p T) (r:T) (pf:one_step_extends Bp Cp r), ~Inn (ba_p_ens Cp) r -> forall (A:Bool_Alg) (g:btp Cp -> bt A) (g' : sig_set (Add (ba_p_ens Cp) r) -> bt A), homomorphism_p1 T g -> extends_sig g' g -> let Cr := im_proj2_sig (Add (ba_p_ens Cp) r) (one_step_extends_incl _ _ _ pf) in forall (g'':sig_set Cr->bt A), (forall x:sig_set Cr, exists pfin:Inn (Add (ba_p_ens Cp) r) (proj1_sig (proj1_sig x)), g'' x = g' (exist _ _ pfin)) -> ((forall F:(Ensemble (btp Bp)), Finite F -> Included (im_proj1_sig F) (Add (ba_p_ens Cp) r) -> forall a:Fin_map F signe mns, el_prod_p _ a = %0 -> el_prod_compose_p1 _ (sig_fun_app g'' 0) a = 0) <-> (forall p q : btp Cp, le_p_sub p (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r pf)) -> le_p_sub_r (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r pf)) q -> le (g p) (g' (exist _ _ (Add_intro2 _ (ba_p_ens Cp) r))) /\ le (g' (exist _ _ (Add_intro2 _ (ba_p_ens Cp) r))) (g q))). intros T Bp Cp r h1 h0 A g g' h2 h3 Cr g'' h4. destruct h1 as [h1a h1b h1c]. pose proof h1a as h1a'. split. intros h5 p q ha hb. destruct p as [p h6], q as [q h7]. assert (h8:Inn (ba_p_ens Bp) p). apply h1b. left. assumption. assert (h9:Inn (ba_p_ens Bp) q). apply h1b. left. assumption. assert (h10:Inn (ba_p_ens Bp) r). apply h1b. right. constructor. pose (Couple (exist _ _ h8, pls) (exist _ _ h10, mns)) as PR. pose (Couple (exist _ _ h10, pls) (exist _ _ h9, mns)) as RQ. assert (h11:Finite PR). apply finite_couple. assert (h12:Finite RQ). apply finite_couple. assert (h15:card_fun1 PR = card_fun1 (dom_rel PR)). unfold PR. assert (h16:p <> r). intro h17. pose proof h6 as h6'. rewrite h17 in h6'. contradiction. rewrite dom_rel_couple. simpl. rewrite card_fun1_couple. rewrite card_fun1_couple. reflexivity. intro h13. apply exist_injective in h13. rewrite h13 in h16. contradict h16. reflexivity. intro h17. inversion h17. assert (h16:card_fun1 RQ = card_fun1 (dom_rel RQ)). unfold RQ. assert (h16:q <> r). intro h17. pose proof h7 as h7'. rewrite h17 in h7'. contradiction. rewrite dom_rel_couple. simpl. rewrite card_fun1_couple. rewrite card_fun1_couple. reflexivity. intro h13. apply exist_injective in h13. rewrite h13 in h16. contradict h16. reflexivity. intro h17. inversion h17. assert (hincl:Included (ran_rel PR) signe). red. intros s h17. destruct s. left. right. assert (hincl':Included (ran_rel RQ) signe). red. intros s h17. destruct s. left. right. pose (fps_to_f_eq_card_dom_rel ba_eq_dec_p sign_eq_dec PR mns h11 h15) as apr. pose (fps_to_f_eq_card_dom_rel ba_eq_dec_p sign_eq_dec RQ mns h12 h16) as arq. pose (fin_map_new_ran apr signe_finite hincl) as apr'. pose (fin_map_new_ran arq signe_finite hincl') as arq'. assert (h18:Included (im_proj1_sig (dom_rel PR)) (Add (ba_p_ens Cp) r)). red. intros x h18. destruct h18 as [x h18 ? h19]. rewrite h19. destruct h18 as [h18]. destruct h18 as [s h18]. inversion h18 as [h20 | h21]. simpl. left. assumption. simpl. right. constructor. pose proof (h5 (dom_rel PR) (dom_rel_finite _ h11) h18 apr') as h17. assert (h19: el_prod_p T apr' = %0). unfold el_prod_p. assert (h20: times_set_p (Im (dom_rel PR) (fun i : btp Bp => eps_p i (apr' |-> i))) (finite_image (btp Bp) (btp Bp) (dom_rel PR) (fun i : btp Bp => eps_p i (apr' |-> i)) (fin_map_fin_dom apr')) = (sub_lift h1a (exist _ _ h6)) %* (%-exist _ _ (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)))). assert (h19: Im (dom_rel PR) (fun i : btp Bp => eps_p i (apr' |-> i)) = Couple (sub_lift h1a (exist _ _ h6)) (%- exist _ _ (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)))). apply Extensionality_Ensembles. red. split. red. intros x h20. destruct h20 as [x h20 ? h21]. rewrite h21. clear h21. destruct h20 as [h20]. destruct h20 as [s h20]. destruct s. inversion h20 as [h21 | h22]. destruct x as [x h22]. apply exist_injective in h21. assert (h23:eps_p (exist (Inn (ba_p_ens Bp)) p h8) (apr' |-> exist (Inn (ba_p_ens Bp)) p h8) = sub_lift h1a (exist (Inn (A_p T (Bc_p T Cp))) p h6)). unfold eps_p, sub_lift. simpl. unfold apr'. assert (h24:fin_map_new_ran apr signe_finite hincl |-> (exist _ _ h8) = pls). rewrite <- fin_map_new_ran_compat. unfold apr. assert (h25:exist _ _ h22 = exist _ _ h8). apply proj1_sig_injective; auto. rewrite h25 in h20. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec PR mns h11 h15 _ h20) as h23. simpl in h23. assumption. constructor. exists pls; auto. assert (he:exist _ _ h22 = exist _ _ h8). apply proj1_sig_injective; auto. rewrite <- he. assumption. rewrite h24. apply proj1_sig_injective. simpl. reflexivity. rewrite h23. left. inversion h20 as [h21 | h22]. destruct x as [x h22']. apply exist_injective in h22. assert (h23:eps_p (exist (Inn (ba_p_ens Bp)) r h10) (apr' |-> exist (Inn (ba_p_ens Bp)) r h10) = (%-exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)))). unfold eps_p. unfold apr'. assert (h24:fin_map_new_ran apr signe_finite hincl |-> (exist _ _ h10) = mns). rewrite <- fin_map_new_ran_compat. unfold apr. assert (h25:exist _ _ h22' = exist _ _ h10). apply proj1_sig_injective; auto. rewrite h25 in h20. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec PR mns h11 h15 _ h20) as h23. simpl in h23. assumption. constructor. exists mns; auto. assert (he:exist _ _ h22' = exist _ _ h10). apply proj1_sig_injective; auto. rewrite <- he. assumption. rewrite h24. f_equal. apply proj1_sig_injective. simpl. reflexivity. rewrite h23. right. red. intros x h19. destruct h19. apply Im_intro with (exist _ _ h8). constructor. exists pls. left. unfold sub_lift. simpl. assert (h23:eps_p (exist (Inn (ba_p_ens Bp)) p h8) (apr' |-> exist (Inn (ba_p_ens Bp)) p h8) = sub_lift h1a (exist (Inn (A_p T (Bc_p T Cp))) p h6)). unfold eps_p, sub_lift. simpl. unfold apr'. assert (h24:fin_map_new_ran apr signe_finite hincl |-> (exist _ _ h8) = pls). assert (h25:Inn PR (exist _ _ h8, pls)). left. rewrite <- fin_map_new_ran_compat. unfold apr. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec PR mns h11 h15 _ h25) as h23. simpl in h23. assumption. constructor. exists pls. assumption. rewrite h24. apply proj1_sig_injective; auto. rewrite h23. unfold sub_lift. apply proj1_sig_injective; auto. apply Im_intro with (exist _ _ h10). constructor. exists mns. right. unfold eps_p, apr'. assert (h19: fin_map_new_ran apr signe_finite hincl |-> exist (Inn (ba_p_ens Bp)) r h10 = mns). assert (h25:Inn PR (exist _ _ h10, mns)). right. unfold apr. rewrite <- fin_map_new_ran_compat. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec PR mns h11 h15 _ h25) as h23. simpl in h23. assumption. constructor. exists mns; auto. rewrite h19. f_equal. apply proj1_sig_injective; auto. pose proof (finite_couple (sub_lift h1a (exist (Inn (A_p T (Bc_p T Cp))) p h6)) (%-exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)))) as h20. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (btp Bp) (btp Bp) (dom_rel PR) (fun i : btp Bp => eps_p i (apr' |-> i)) (fin_map_fin_dom apr')) h20 h19) as h21. dependent rewrite -> h21. rewrite times_set_couple_p'. reflexivity. rewrite h20 at 1. clear h20. destruct ha as [ha1 ha2]. rewrite le_iff_p in ha2. rewrite <- ha2. f_equal. f_equal. apply proof_irrelevance. specialize (h17 h19). clear h19. unfold el_prod_compose_p1 in h17. assert (h19: Im (dom_rel PR) (fun i : btp Bp => eps (sig_fun_app g'' 0 i) (apr' |-> i)) = Couple (g (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) p h6)) (- g' (exist (Inn (Add (ba_p_ens Cp) r)) r (Add_intro2 T (ba_p_ens Cp) r)))). apply Extensionality_Ensembles. red. split. red. intros x h19. destruct h19 as [x h19 ? h20]. rewrite h20. clear h20. destruct h19 as [h19]. destruct h19 as [s h19]. inversion h19 as [h20 | h21]. destruct x as [x h22]. pose proof h20 as h20'. apply exist_injective in h20. assert (h23: eps (sig_fun_app g'' 0 (exist (Inn (ba_p_ens Bp)) p h8)) (apr' |-> exist (Inn (ba_p_ens Bp)) p h8) = g (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) p h6)). unfold eps. unfold apr'. assert (h23:fin_map_new_ran apr signe_finite hincl |-> exist (Inn (ba_p_ens Bp)) p h8 = apr |-> exist _ _ h8). rewrite <- fin_map_new_ran_compat. reflexivity. constructor. exists s. rewrite h20'. assumption. rewrite h23. unfold apr. assert (h24: fps_to_f_eq_card_dom_rel ba_eq_dec_p sign_eq_dec PR mns h11 h15 |-> exist (Inn (ba_p_ens Bp)) p h8 = pls). pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec PR mns h11 h15 _ h19) as h25. simpl in h25. rewrite h20'. rewrite H. assumption. rewrite h24. unfold sig_fun_app. destruct classic_dec as [h26 | h27]. pose proof (h4 (exist _ _ h26)) as h28. destruct h28 as [h28 h29]. rewrite h29 at 1. pose proof h3 as h3'. destruct h3' as [h30 h31]. specialize (h31 (exist _ _ h6)). rewrite h31 at 1. f_equal. apply proj1_sig_injective. reflexivity. contradict h27. unfold Cr, im_proj2_sig. assert (h26:Inn (Add (ba_p_ens Cp) r) p). left. assumption. apply Im_intro with (exist _ _ h26). constructor. apply proj1_sig_injective; auto. rewrite h23 at 1. left. assert (h23:eps (sig_fun_app g'' 0 (exist (Inn (ba_p_ens Bp)) r h10)) (apr' |-> exist (Inn (ba_p_ens Bp)) r h10) = - g' (exist _ _ (Add_intro2 _ (ba_p_ens Cp) r))). unfold eps, apr'. assert (h24:fin_map_new_ran apr signe_finite hincl |-> exist _ _ h10 = apr |-> exist _ _ h10). rewrite <- fin_map_new_ran_compat. reflexivity. constructor. exists s. rewrite h21; auto. rewrite h24. unfold apr. rewrite h21. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec PR mns h11 h15 _ h19) as h25. simpl in h25. rewrite h25 at 1. rewrite <- H. f_equal. unfold sig_fun_app. destruct classic_dec as [h26 | h27]. pose proof (h4 (exist _ _ h26)) as h27. destruct h27 as [h27 h28]. rewrite h28 at 1. f_equal. apply proj1_sig_injective. simpl. rewrite <- h21. simpl. reflexivity. contradict h27. rewrite <- h21. unfold Cr, im_proj2_sig. apply Im_intro with (exist _ _ (Add_intro2 _ (ba_p_ens Cp) r)). constructor. apply proj1_sig_injective; auto. rewrite h23 at 1. right. red. intros x h19. destruct h19. unfold PR. rewrite dom_rel_couple at 1. simpl. apply Im_intro with (exist _ _ h8). left. unfold eps, apr'. rewrite <- fin_map_new_ran_compat. unfold apr. assert (h19:Inn PR (exist _ _ h8, pls)). left. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec PR mns h11 h15 _ h19) as h20. rewrite h20 at 1. simpl. unfold sig_fun_app. destruct classic_dec as [h21 | h22]. pose proof (h4 (exist _ _ h21)) as h22. destruct h22 as [h22 h23]. rewrite h23 at 1. pose proof h3 as h3'. destruct h3' as [h3a h3b]. specialize (h3b (exist _ _ h6)). rewrite h3b at 1. f_equal. apply proj1_sig_injective; auto. contradict h22. unfold Cr, im_proj2_sig. assert (h23:Inn (Add (ba_p_ens Cp) r) p). left. assumption. apply Im_intro with (exist _ _ h23). constructor. apply proj1_sig_injective; auto. unfold PR. constructor. exists pls. left. apply Im_intro with (exist _ _ h10). constructor. exists mns. right. unfold eps. assert (h19: apr' |-> exist (Inn (ba_p_ens Bp)) r h10 = mns). unfold apr'. rewrite <- fin_map_new_ran_compat. unfold apr. assert (h19:Inn PR (exist _ _ h10, mns)). right. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec PR mns h11 h15 _ h19) as h20. rewrite h20 at 1. simpl. reflexivity. unfold PR. constructor. exists mns. right. rewrite h19 at 1. f_equal. unfold sig_fun_app. destruct classic_dec as [h21 | h22]. pose proof (h4 (exist _ _ h21)) as h25. destruct h25 as [h25 h26]. rewrite h26 at 1. f_equal. apply proj1_sig_injective. simpl. reflexivity. contradict h22. unfold Cr. assert (h20:Inn (Add (ba_p_ens Cp) r) r). auto with sets. apply Im_intro with (exist _ _ h20). constructor. apply proj1_sig_injective; auto. pose proof (finite_couple (g (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) p h6)) (- g' (exist (Inn (Add (ba_p_ens Cp) r)) r (Add_intro2 T (ba_p_ens Cp) r)))) as h20. assert (h21:times_set _ (finite_image (btp Bp) (bt A) (dom_rel PR) (fun i : btp Bp => eps (sig_fun_app g'' 0 i) (apr' |-> i)) (fin_map_fin_dom apr')) = times_set _ h20). apply times_set_functional. assumption. rewrite h21 in h17 at 1. clear h21. rewrite times_set_couple' in h17. split. rewrite le_iff. assumption. rewrite le_iff. pose proof hb as hb'. destruct hb' as [h21 h22]. assert (h24: Included (im_proj1_sig (dom_rel RQ)) (Add (ba_p_ens Cp) r)). red. intros x h24. destruct h24 as [x h24 ? h25]. rewrite h25. destruct h24 as [h24]. destruct h24 as [s h24]. inversion h24 as [h26 | h27]. simpl. right. constructor. simpl. left. assumption. pose proof (h5 (dom_rel RQ) (dom_rel_finite _ h12) h24 arq') as h25. assert (h19':el_prod_p T arq' = %0). unfold el_prod_p. assert (h20': times_set_p (Im (dom_rel RQ) (fun i : btp Bp => eps_p i (arq' |-> i))) (finite_image (btp Bp) (btp Bp) (dom_rel RQ) (fun i : btp Bp => eps_p i (arq' |-> i)) (fin_map_fin_dom arq')) = (exist _ _ (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c))) %* (%- sub_lift h21 (exist _ _ h7))). assert (h19'': Im (dom_rel RQ) (fun i : btp Bp => eps_p i (arq' |-> i)) = Couple (exist _ _ (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c))) (%- sub_lift h21 (exist _ _ h7))). apply Extensionality_Ensembles. red. split. red. intros x h20'. destruct h20' as [x h20' ? h21']. rewrite h21'. clear h21'. destruct h20' as [h20']. destruct h20' as [s h20']. destruct s. inversion h20' as [h21' | h22']. destruct x as [x h22']. apply exist_injective in h21'. assert (h23':eps_p (exist (Inn (ba_p_ens Bp)) r h10) (arq' |-> exist (Inn (ba_p_ens Bp)) r h10) = exist _ _ (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c))). unfold eps_p, sub_lift. simpl. unfold arq'. assert (h24':fin_map_new_ran arq signe_finite hincl' |-> (exist _ _ h10) = pls). rewrite <- fin_map_new_ran_compat. unfold apr. assert (h25':exist _ _ h22' = exist _ _ h10). apply proj1_sig_injective; auto. rewrite h25' in h20'. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec RQ mns h12 h16 _ h20') as h23'. simpl in h23'. assumption. unfold RQ. constructor. exists pls. left. rewrite h24'. apply proj1_sig_injective. simpl. reflexivity. rewrite h23'. left. inversion h20' as [h21' | h22']. destruct x as [x h22'']. apply exist_injective in h22'. assert (h23': eps_p (exist (Inn (ba_p_ens Bp)) q h9) (arq' |-> exist (Inn (ba_p_ens Bp)) q h9) = %-sub_lift h21 (exist (Inn (A_p T (Bc_p T Cp))) q h7)). unfold eps_p. unfold arq'. assert (h24':fin_map_new_ran arq signe_finite hincl' |-> (exist _ _ h9) = mns). rewrite <- fin_map_new_ran_compat. unfold arq. assert (h25':exist _ _ h22'' = exist _ _ h9). apply proj1_sig_injective; auto. rewrite h25' in h20'. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec RQ mns h12 h16 _ h20') as h23'. simpl in h23'. assumption. constructor. exists mns. assert (h26:exist _ _ h9 = exist _ _ h22''). apply proj1_sig_injective; auto. rewrite h26. assumption. rewrite h24'. f_equal. apply proj1_sig_injective. simpl. reflexivity. rewrite h23'. right. red. intros x h19'. destruct h19'. apply Im_intro with (exist _ _ h10). constructor. exists pls. left. assert (h23:eps_p (exist (Inn (ba_p_ens Bp)) r h10) (arq' |-> exist (Inn (ba_p_ens Bp)) r h10) = exist _ _ h10). unfold eps_p. simpl. unfold arq'. assert (h24':fin_map_new_ran arq signe_finite hincl' |-> (exist _ _ h10) = pls). assert (h25':Inn RQ (exist _ _ h10, pls)). left. rewrite <- fin_map_new_ran_compat. unfold arq. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec RQ mns h12 h16 _ h25') as h23'. simpl in h23'. assumption. unfold RQ. constructor. exists pls. left. rewrite h24'. apply proj1_sig_injective; auto. rewrite h23. apply proj1_sig_injective; auto. apply Im_intro with (exist _ _ h9). constructor. exists mns. right. unfold eps_p, arq'. assert (h19': fin_map_new_ran arq signe_finite hincl' |-> exist (Inn (ba_p_ens Bp)) q h9 = mns). assert (h25':Inn RQ (exist _ _ h9, mns)). right. unfold arq. rewrite <- fin_map_new_ran_compat. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec RQ mns h12 h16 _ h25') as h23'. simpl in h23'. assumption. unfold RQ. constructor. exists mns. right. rewrite h19'. f_equal. apply proj1_sig_injective; auto. pose proof (finite_couple (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c))) (%-sub_lift h21 (exist (Inn (A_p T (Bc_p T Cp))) q h7))) as h20'. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (btp Bp) (btp Bp) (dom_rel RQ) (fun i : btp Bp => eps_p i (arq' |-> i)) (fin_map_fin_dom arq')) h20' h19'') as h21'. dependent rewrite -> h21'. rewrite times_set_couple_p'. reflexivity. rewrite h20' at 1. clear h20'. destruct hb as [hb1 hb2]. rewrite le_iff_p in hb2. rewrite <- hb2. f_equal. f_equal. f_equal. apply proof_irrelevance. specialize (h25 h19'). clear h19'. unfold el_prod_compose_p1 in h25. assert (h19': Im (dom_rel RQ) (fun i : btp Bp => eps (sig_fun_app g'' 0 i) (arq' |-> i)) = Couple (g' (exist (Inn (Add (ba_p_ens Cp) r)) r (Add_intro2 T (ba_p_ens Cp) r))) (- g (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) q h7))). apply Extensionality_Ensembles. red. split. red. intros x h19'. destruct h19' as [x h19' ? h20']. rewrite h20'. clear h20'. destruct h19' as [h19']. destruct h19' as [s h19']. inversion h19' as [h20' | h21']. destruct x as [x h22']. assert (h23: (eps (sig_fun_app g'' 0 (exist (Inn (ba_p_ens Bp)) r h10)) (arq' |-> exist (Inn (ba_p_ens Bp)) r h10)) = (g' (exist (Inn (Add (ba_p_ens Cp) r)) r (Add_intro2 T (ba_p_ens Cp) r)))). unfold eps. unfold arq'. assert (h23:fin_map_new_ran arq signe_finite hincl' |-> exist (Inn (ba_p_ens Bp)) r h10 = arq |-> exist _ _ h10). rewrite <- fin_map_new_ran_compat. reflexivity. unfold dom_rel. constructor. exists s. rewrite h20'. assumption. rewrite h23. unfold arq. assert (h24': fps_to_f_eq_card_dom_rel ba_eq_dec_p sign_eq_dec RQ mns h12 h16 |-> exist (Inn (ba_p_ens Bp)) r h10 = pls). pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec RQ mns h12 h16 _ h19') as h25'. simpl in h25'. rewrite h20'. rewrite H. assumption. rewrite h24'. unfold sig_fun_app. destruct classic_dec as [h26' | h27']. pose proof (h4 (exist _ _ h26')) as h28'. destruct h28' as [h28' h29']. rewrite h29' at 1. f_equal. apply proj1_sig_injective; auto. contradict h27'. unfold Cr, im_proj2_sig. assert (h26:Inn (Add (ba_p_ens Cp) r) r). right. auto with sets. apply Im_intro with (exist _ _ h26). constructor. apply proj1_sig_injective; auto. rewrite h23 at 1. left. assert (h23: (eps (sig_fun_app g'' 0 (exist (Inn (ba_p_ens Bp)) q h9)) (arq' |-> exist (Inn (ba_p_ens Bp)) q h9)) = (- g (exist (fun x0 : T => Inn (A_p T (Bc_p T Cp)) x0) q h7))). unfold eps, arq'. assert (h24':fin_map_new_ran arq signe_finite hincl' |-> exist _ _ h9 = arq |-> exist _ _ h9). rewrite <- fin_map_new_ran_compat. reflexivity. unfold RQ. constructor. exists mns. right. rewrite h24'. unfold arq. rewrite h21'. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec RQ mns h12 h16 _ h19') as h25'. simpl in h25'. rewrite h25' at 1. rewrite <- H. f_equal. unfold sig_fun_app. destruct classic_dec as [h26 | h27]. pose proof (h4 (exist _ _ h26)) as h27. destruct h27 as [h27 h28]. rewrite h28 at 1. pose proof h3 as h3'. destruct h3' as [h3a h3b]. specialize (h3b (exist _ _ h7)). rewrite h3b at 1. f_equal. apply proj1_sig_injective. simpl. rewrite <- h21'. simpl. reflexivity. contradict h27. rewrite <- h21'. unfold Cr, im_proj2_sig. assert (hin:Inn (Add (ba_p_ens Cp) r) q). left. assumption. apply Im_intro with (exist _ _ hin). constructor. apply proj1_sig_injective; auto. rewrite h23 at 1. right. red. intros x h19'. destruct h19'. unfold RQ. rewrite dom_rel_couple at 1. simpl. apply Im_intro with (exist _ _ h10). left. unfold eps, arq'. rewrite <- fin_map_new_ran_compat. unfold arq. assert (h19':Inn RQ (exist _ _ h10, pls)). left. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec RQ mns h12 h16 _ h19') as h20'. rewrite h20' at 1. simpl. unfold sig_fun_app. destruct classic_dec as [h21' | h22']. pose proof (h4 (exist _ _ h21')) as h22'. destruct h22' as [h22' h23']. rewrite h23' at 1. pose proof h3 as h3'. destruct h3' as [h3a h3b]. specialize (h3b (exist _ _ h6)). f_equal. apply proj1_sig_injective; auto. contradict h22'. unfold Cr, im_proj2_sig. assert (h23:Inn (Add (ba_p_ens Cp) r) r). right. auto with sets. apply Im_intro with (exist _ _ h23). constructor. apply proj1_sig_injective; auto. unfold RQ. constructor. exists pls. left. apply Im_intro with (exist _ _ h9). constructor. exists mns. right. unfold eps. assert (h19':arq' |-> exist (Inn (ba_p_ens Bp)) q h9 = mns). unfold arq'. rewrite <- fin_map_new_ran_compat. unfold arq. assert (h19':Inn RQ (exist _ _ h9, mns)). right. pose proof (fps_to_f_eq_card_dom_rel_compat' ba_eq_dec_p sign_eq_dec RQ mns h12 h16 _ h19') as h20'. rewrite h20' at 1. simpl. reflexivity. unfold RQ. constructor. exists mns. right. rewrite h19' at 1. f_equal. unfold sig_fun_app. destruct classic_dec as [h21' | h22']. pose proof (h4 (exist _ _ h21')) as h25'. destruct h25' as [h25' h26']. rewrite h26' at 1. f_equal. pose proof h3 as h3'. destruct h3' as [h3a h3b]. specialize (h3b (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) q h7)). rewrite h3b at 1. f_equal. apply proj1_sig_injective. simpl. reflexivity. contradict h22'. unfold Cr. assert (h20':Inn (Add (ba_p_ens Cp) r) q). left. assumption. apply Im_intro with (exist _ _ h20'). constructor. apply proj1_sig_injective; auto. pose proof (finite_couple (g' (exist (Inn (Add (ba_p_ens Cp) r)) r (Add_intro2 T (ba_p_ens Cp) r))) (- g (exist (fun x : T => Inn (A_p T (Bc_p T Cp)) x) q h7))) as h20'. assert (h21':times_set _ (finite_image (btp Bp) (bt A) (dom_rel RQ) (fun i : btp Bp => eps (sig_fun_app g'' 0 i) (arq' |-> i)) (fin_map_fin_dom arq')) = times_set _ h20'). apply times_set_functional. assumption. rewrite h21' in h25 at 1. clear h21'. rewrite times_set_couple' in h25. assumption. (* <-*) intros h5 F h6 h7 a h8. unfold el_prod_p in h8. pose proof (@decompose_int_setminus). assert (h10:Inn (ba_p_ens Bp) r). apply h1b. right. constructor. pose proof (decompose_int_setminus F (Singleton (exist _ _ h10))) as h11. destruct (classic (Inn F (exist _ _ h10))) as [h12 | h13]. assert (h13: Intersection F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)) = Singleton (exist _ _ h10)). apply Extensionality_Ensembles. red. split. red. intros x h14. destruct h14 as [x h14 h15]. destruct h15. constructor. red. intros x h13. destruct h13. constructor; auto with sets. rewrite h13 in h11. clear h13. assert (h14:Im F (fun i : btp Bp => eps_p i (a |-> i)) = Union (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i:btp Bp => eps_p i (a |-> i))) (Singleton (eps_p (exist _ _ h10) (a |-> (exist _ _ h10))))). rewrite h11 at 1. rewrite im_union. rewrite im_singleton. rewrite Union_commutative. reflexivity. assert (hfa: Finite (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i)))). apply finite_image. apply finite_setminus; auto. assert (hfs: Finite (Singleton (eps_p (exist (Inn (ba_p_ens Bp)) r h10) (a |-> exist (Inn (ba_p_ens Bp)) r h10)))). apply Singleton_is_finite. assert (h15:Finite (Union (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) (Singleton (eps_p (exist (Inn (ba_p_ens Bp)) r h10) (a |-> exist (Inn (ba_p_ens Bp)) r h10))))). apply Union_preserves_Finite; auto. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (btp Bp) (btp Bp) F (fun i : btp Bp => eps_p i (a |-> i)) (fin_map_fin_dom a)) h15 h14) as h16. dependent rewrite -> h16 in h8. clear h16. pose proof (times_set_union_p' _ _ hfa hfs h15) as h16. rewrite h16 in h8. clear h16. assert (h28:Included (Singleton (exist (Inn (ba_p_ens Bp)) r h10)) F). red. intros x h29. destruct h29. assumption. assert (h23:Included (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (ba_p_ens Bp)). red. intros x h24. destruct h24 as [x h24]. rewrite H0. apply proj2_sig. assert (h34: times_set (Im (Singleton (exist (Inn (ba_p_ens Bp)) r h10)) (fun i : btp Bp => eps (sig_fun_app g'' 0 i) (restriction a h28 |-> i))) (finite_image (btp Bp) (bt A) (Singleton (exist (Inn (ba_p_ens Bp)) r h10)) (fun i : btp Bp => eps (sig_fun_app g'' 0 i) (restriction a h28 |-> i)) (fin_map_fin_dom (restriction a h28))) = eps (sig_fun_app g'' 0 (exist (Inn (ba_p_ens Bp)) r h10)) (restriction a h28 |-> exist (Inn (ba_p_ens Bp)) r h10)). rewrite <- times_set_sing. apply times_set_functional. rewrite im_singleton. reflexivity. assert (hs: Inn (Singleton (exist (Inn (ba_p_ens Bp)) r h10)) (exist (Inn (ba_p_ens Bp)) r h10)). constructor. rewrite restriction_compat in h34; auto. destruct (a |-> exist (Inn (ba_p_ens Bp)) r h10). simpl in h8. rewrite times_set_sing_p' in h8. assert (h16:Inn (ba_p_ens Cp) (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa))). pose proof (subtract_preserves_inclusion _ _ r h7) as h17. rewrite sub_add_same_nin in h17; auto. assert (h18:Included (Subtract F (exist _ _ h10)) F). auto with sets. pose (restriction a h18) as a'. pose proof (el_prod_in_gen_p a') as h19. assert (h20:Included (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (ba_p_ens Cp)). red. intros x h21. destruct h21 as [x h21 ? h22]. rewrite h22. clear h22. clear y. destruct h21 as [h21 h22]. rewrite <- (in_im_proj1_sig_iff (ba_p_ens Bp) F x) in h21 at 1. apply h7 in h21. inversion h21 as [p h23' | p h24]. assumption. contradict h22. inversion h24 as [h25]. assert (h26:exist _ _ h10 = exist _ _ (proj2_sig x)). apply proj1_sig_injective. simpl. assumption. rewrite h26 at 1. rewrite unfold_sig. constructor. assert (h21:Included (Gen_Ens_p (im_proj1_sig (Subtract F (exist (Inn (ba_p_ens Bp)) r h10))) (incl_ens_btp (Subtract F (exist (Inn (ba_p_ens Bp)) r h10)))) (ba_p_ens Cp)). pose proof (gen_ens_p_subalg_of_p h1a _ h20 (incl_ens_btp (Subtract F (exist (Inn (ba_p_ens Bp)) r h10)))) as h22. rewrite <- h22 at 1. apply incl_gen_ens_p. apply h21. unfold el_prod_p, a' in h19. assert (h24: proj1_sig (times_set_p (Im (Subtract F (exist (Inn (ba_p_ens Bp)) r h10)) (fun i : btp Bp => eps_p i (restriction a h18 |-> i))) (finite_image (btp Bp) (btp Bp) (Subtract F (exist (Inn (ba_p_ens Bp)) r h10)) (fun i : btp Bp => eps_p i (restriction a h18 |-> i)) (fin_map_fin_dom (restriction a h18)))) = proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)). f_equal. apply times_set_functional_p. apply im_ext_in. intros x h22. rewrite restriction_compat. reflexivity. assumption. rewrite <- h24. clear h24. assumption. pose proof (zero_min_sub h1a (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)))) as h18. assert (h19: le_p_sub_r (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c))) (%-exist (Inn (ba_p_ens Cp)) (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16)). assert (h21: le_p (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c))) (sub_lift h1a (%-exist (Inn (ba_p_ens Cp)) (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16))). rewrite sub_lift_comp_eq. rewrite le_iff_p'. rewrite <- h8. rewrite comm_prod_p. f_equal. unfold sub_lift. simpl. apply proj1_sig_injective. simpl. reflexivity. apply proj1_sig_injective. simpl. reflexivity. apply (le_p_sub_intro_r (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c))) (%-exist (Inn (ba_p_ens Cp)) (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16) h1a h21). pose proof (h5 %0 (%- exist _ _ h16) h18 h19) as h17. destruct h17 as [h17a h17b]. clear h17a. rewrite homo_comp_p1 in h17b; auto. rewrite le_iff' in h17b. pose proof h3 as h3'. red in h3'. destruct h3' as [h20 h21]. rewrite h21 in h17b. simpl in h17b. destruct h1a as [h1l [h1m h1r]]. assert (h22:Inn Cr (exist _ _ (h1l _ h16))). unfold Cr. apply Im_intro with (exist _ _ (h20 _ h16)). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. assert (h23':Inn Cr (exist _ _ h10)). unfold Cr. pose proof (Add_intro2 _ (ba_p_ens Cp) r) as h23'. apply Im_intro with (exist _ _ h23'). constructor. apply proj1_sig_injective. simpl. reflexivity. pose proof (h4 (exist _ _ h22)) as h24. pose proof (h4 (exist _ _ h23')) as h25. simpl in h24, h25. destruct h24 as [h24a h24b], h25 as [h25a h25b]. assert (h26:g' (exist _ _ (Add_intro2 _ (ba_p_ens Cp) r)) = g' (exist _ _ h25a)). f_equal. apply proj1_sig_injective; auto. rewrite <- h26 in h25b. clear h26. assert (h27:g' (exist (Inn (Add (ba_p_ens Cp) r)) (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) (h20 (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16)) = g' (exist _ _ h24a)). f_equal. apply proj1_sig_injective; auto. rewrite <- h27 in h24b. rewrite <- h24b in h17b at 1. rewrite <- h25b in h17b at 1. rewrite (el_prod_compose_decompose_restriction_p1 (sig_fun_app g'' 0) _ _ h28). unfold el_prod_compose_p1. rewrite h34 at 1. simpl. pose proof h3 as h3'. red in h3'. destruct h3' as [h3a h3b]. pose proof (h3b (exist _ _ h16)) as h40. simpl in h40. assert (h41: h3a (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16 = (h20 (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16)). apply proof_irrelevance. rewrite h41 in h40 at 1. rewrite <- h40 in h24b at 1. clear h41. assert (h42: sig_fun_app g'' 0 (exist (Inn (ba_p_ens Bp)) r h10) = g'' (exist (Inn Cr) (exist (Inn (ba_p_ens Bp)) r h10) h23')). unfold sig_fun_app. destruct classic_dec as [h43 | h44]. f_equal. apply proj1_sig_injective; auto. contradiction. rewrite h42 at 1. clear h42. rewrite h24b in h17b at 1. clear h24b. assert (h41:Included (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (ba_p_ens Cp)). red. intros x h42. destruct h42 as [x h42 ? h43]. rewrite h43. clear h43. destruct h42 as [h42 h43]. assert (h44:Inn (im_proj1_sig F) (proj1_sig x)). eapply Im_intro; auto. assumption. pose proof (h7 _ h44) as h45. inversion h45 as [c h46 h46'| c h47 h47']. assumption. inversion h47; clear h47. contradict h43. destruct x as [x h48]. simpl in h47'. simpl in H0. assert (h49 : exist (Inn (ba_p_ens Bp)) r h10 = exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h48). apply proj1_sig_injective. simpl. assumption. rewrite h49. constructor. assert (h41' : Included (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (ba_p_ens Bp)). auto with sets. pose (im_proj2_sig _ h41) as F'. assert (hfin:Finite F'). apply finite_image; auto. rewrite <- finite_full_sig_iff. apply finite_image. eapply Finite_downward_closed. apply h6. red; intros d h50. destruct h50; auto. pose (fun c:sig_set F' => (a |-> (exist _ _ (h1l _ (proj2_sig (proj1_sig c)))))) as fa. assert (h50: Included (Im (full_sig F') fa) signe). red. intros x ?. destruct x. left. right. pose (sig_fun_to_fin_map_ran ba_eq_dec_p fa hfin mns signe signe_finite h50) as a'. assert (h43: times_set (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps (sig_fun_app g'' 0 i) (restriction a (setminus_inc F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) |-> i))) (finite_image (btp Bp) (bt A) (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps (sig_fun_app g'' 0 i) (restriction a (setminus_inc F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) |-> i)) (fin_map_fin_dom (restriction a (setminus_inc F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))))) = g (exist (Inn (ba_p_ens Cp)) (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16)). assert (h42: exist (Inn (ba_p_ens Cp)) (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16 = times_set_p (Im F' (fun i : btp Cp => eps_p i (a' |-> i))) (finite_image _ _ _ _ hfin)). apply proj1_sig_injective. simpl. assert (h51:im_proj1_sig (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) = im_proj1_sig (Im F' (fun i : btp Cp => eps_p i (a' |-> i)))). unfold F',im_proj1_sig, im_proj2_sig. do 3 rewrite im_im. rewrite im_full_sig_im_eq. simpl. apply Extensionality_Ensembles. red; split; red. intros x h52. destruct h52 as [x h52 ? h53]. rewrite h53. clear h53 y. apply Im_intro with (exist _ _ h52). constructor. simpl. unfold a'. assert (h54: Inn F' (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h52 (proj1_sig x) eq_refl)))). unfold F'. unfold im_proj2_sig, im_proj1_sig. rewrite im_full_sig_im_eq. simpl. apply Im_intro with (exist _ _ h52). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. pose proof (sig_fun_to_fin_map_ran_compat ba_eq_dec_p fa hfin mns signe signe_finite h50 (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h52 (proj1_sig x) eq_refl))) h54) as h53. unfold F' in h53. unfold F'. rewrite h53 at 1. unfold fa. simpl. unfold eps_p. assert (h55:x = exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h1l (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h52 (proj1_sig x) eq_refl)))). apply proj1_sig_injective. simpl. reflexivity. rewrite <- h55. destruct (a |-> x). simpl. reflexivity. assert (h57: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h1l h1m)) (proj1_sig x)). rewrite <- h1r. apply h41. eapply Im_intro. apply h52. reflexivity. pose proof (ba_p_subst_comp _ _ h1r _ (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h52 (proj1_sig x) eq_refl)) h57) as h56. rewrite h56. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. intros x h51. destruct h51 as [x h51 ? h52]. rewrite h52. clear h52 y h51. destruct x as [x h51]. simpl. apply Im_intro with x. assumption. unfold a'. assert (h53: Inn F' (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h51 (proj1_sig x) eq_refl)))). unfold F', im_proj2_sig, im_proj1_sig. rewrite im_full_sig_im_eq. simpl. apply Im_intro with (exist _ _ h51). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. pose proof (sig_fun_to_fin_map_ran_compat ba_eq_dec_p fa hfin mns signe signe_finite h50 (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h51 (proj1_sig x) eq_refl))) h53) as h52. unfold F' in h52. unfold F'. rewrite h52 at 1. unfold fa, eps_p. simpl. assert (h54:x = exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h1l (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h51 (proj1_sig x) eq_refl)))). apply proj1_sig_injective. simpl. reflexivity. rewrite <- h54. destruct (a |-> x). simpl. reflexivity. assert (h56: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h1l h1m)) (proj1_sig x)). rewrite <- h1r. apply h41. eapply Im_intro. apply h51. reflexivity. pose proof (ba_p_subst_comp _ _ h1r _ (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h51 (proj1_sig x) eq_refl)) h56) as h55. rewrite h55 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply (times_set_sub_compat_p _ _ h1a' _ _ hfa (finite_image {x : T | Inn (ba_p_ens Cp) x} (btp Cp) F' (fun i : btp Cp => eps_p i (a' |-> i)) hfin) h51). rewrite h42 at 1. clear h42. rewrite homo_times_set_p1. pose proof (im_im F' (fun i : btp Cp => eps_p i (a' |-> i)) g) as h51. pose proof (finite_image _ _ _ (fun x : {x : T | Inn (ba_p_ens Cp) x} => g ((fun i : btp Cp => eps_p i (a' |-> i)) x)) hfin) as h52. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (btp Cp) (bt A) (Im F' (fun i : btp Cp => eps_p i (a' |-> i))) g (finite_image {x : T | Inn (ba_p_ens Cp) x} (btp Cp) F' (fun i : btp Cp => eps_p i (a' |-> i)) hfin)) h52 h51) as h53. dependent rewrite -> h53. clear h53. unfold F', im_proj2_sig, im_proj1_sig. simpl. pose proof (im_im (full_sig (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x)))) (fun p : sig_set (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x))) => exist (Inn (ba_p_ens Cp)) (proj1_sig p) (h41 (proj1_sig p) (proj2_sig p))) (fun x : {x : T | Inn (ba_p_ens Cp) x} => g (eps_p x (a' |-> x)))) as h53. pose proof (im_full_sig_im_eq (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x)) (fun x : sig_set (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x))) => (fun x0 : {x0 : T | Inn (ba_p_ens Cp) x0} => g (eps_p x0 (a' |-> x0))) ((fun p : sig_set (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0))) => exist (Inn (ba_p_ens Cp)) (proj1_sig p) (h41 (proj1_sig p) (proj2_sig p))) x))) as him. rewrite him in h53 at 1. clear him. simpl in h53. assert (h54: Finite (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))). apply finite_setminus; auto. assert (h55:Finite (Im (full_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (fun x : sig_set (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) => g (eps_p (exist (Inn (ba_p_ens Cp)) (proj1_sig (proj1_sig x)) (h41 (proj1_sig (proj1_sig x)) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (proj1_sig x) (proj2_sig x) (proj1_sig (proj1_sig x)) eq_refl))) (a' |-> exist (Inn (ba_p_ens Cp)) (proj1_sig (proj1_sig x)) (h41 (proj1_sig (proj1_sig x)) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (proj1_sig x) (proj2_sig x) (proj1_sig (proj1_sig x)) eq_refl))))))). rewrite <- h53. rewrite im_im. rewrite im_full_sig_im_eq. simpl. apply finite_image. rewrite <- finite_full_sig_iff. assumption. assert (h57:times_set _ h52 = times_set _ h55). apply times_set_functional. assumption. unfold F' in h57. rewrite h57 at 1. clear h57. apply times_set_functional. apply Extensionality_Ensembles; red; split; red. (* <= *) intros x h58. destruct h58 as [x h58 ? h59]. rewrite h59 at 1. clear h59 y. apply Im_intro with (exist _ _ h58). constructor. simpl. rewrite eps_p_eq. unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. unfold eps. rewrite restriction_compat. unfold a'. assert (h59:Inn F' (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h58 (proj1_sig x) eq_refl)))). unfold F', im_proj2_sig, im_proj1_sig. rewrite im_full_sig_im_eq. simpl. apply Im_intro with (exist _ _ h58). constructor. apply proj1_sig_injective; auto. pose proof (sig_fun_to_fin_map_ran_compat ba_eq_dec_p fa hfin mns signe signe_finite h50 (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h58 (proj1_sig x) eq_refl))) h59) as h60. unfold F' in h60. unfold F'. rewrite h60 at 1. unfold fa. simpl. assert (h61:x = exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h1l (proj1_sig x) (h41 (proj1_sig x) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h58 (proj1_sig x) eq_refl)))). apply proj1_sig_injective; auto. rewrite <- h61. destruct (a |-> x). unfold sig_fun_app. assert (h62:Inn Cr x). unfold Cr, im_proj2_sig. assert (h64:Inn (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (proj1_sig x)). eapply Im_intro. apply h58. reflexivity. pose proof (h41 (proj1_sig x) h64) as h63. pose proof (Add_intro1 _ _ r _ h63) as h65. apply Im_intro with (exist _ _ h65). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. destruct classic_dec as [h63 | h64]. pose proof (h4 (exist _ _ h63)) as h65. destruct h65 as [h65 h66]. rewrite h66 at 1. pose proof h3 as h3'. destruct h3' as [h3l h3r]. rewrite h3r. simpl. f_equal. apply proj1_sig_injective; auto. contradiction. rewrite homo_comp_p1; auto. f_equal. unfold sig_fun_app. assert (h62:Inn Cr x). unfold Cr, im_proj2_sig. assert (h64:Inn (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (proj1_sig x)). eapply Im_intro. apply h58. reflexivity. pose proof (h41 (proj1_sig x) h64) as h63. pose proof (Add_intro1 _ _ r _ h63) as h65. apply Im_intro with (exist _ _ h65). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. destruct classic_dec as [h63 | h64]. pose proof (h4 (exist _ _ h63)) as h65. destruct h65 as [h65 h66]. rewrite h66 at 1. pose proof h3 as h3'. destruct h3' as [h3l h3r]. rewrite h3r. f_equal. apply proj1_sig_injective; auto. contradiction. assumption. (* >= *) intros x h56. destruct h56 as [x h56 q h57]. rewrite h57. clear h57 q h56. destruct x as [x h56]. simpl. apply Im_intro with x. assumption. rewrite eps_p_eq. unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. unfold eps. rewrite restriction_compat. unfold a'. assert (h57:Inn F' (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h56 (proj1_sig x) eq_refl)))). unfold F', im_proj2_sig, im_proj1_sig. rewrite im_full_sig_im_eq. simpl. apply Im_intro with (exist _ _ h56). constructor. apply proj1_sig_injective; auto. pose proof (sig_fun_to_fin_map_ran_compat ba_eq_dec_p fa hfin mns signe signe_finite h50 _ h57) as h58. unfold F' in h58. unfold F'. rewrite h58 at 1. unfold fa. simpl. assert (h59:x = exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h1l (proj1_sig x) (h41 (proj1_sig x) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h56 (proj1_sig x) eq_refl)))). apply proj1_sig_injective; auto. rewrite <- h59. destruct (a |-> x). unfold sig_fun_app. assert (h62:Inn Cr x). unfold Cr, im_proj2_sig. assert (h64:Inn (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (proj1_sig x)). eapply Im_intro. apply h56. reflexivity. pose proof (h41 (proj1_sig x) h64) as h63. pose proof (Add_intro1 _ _ r _ h63) as h65. apply Im_intro with (exist _ _ h65). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. destruct classic_dec as [h63 | h64]. pose proof (h4 (exist _ _ h63)) as h65. destruct h65 as [h65 h66]. rewrite h66 at 1. pose proof h3 as h3'. destruct h3' as [h3l h3r]. rewrite h3r. simpl. f_equal. apply proj1_sig_injective; auto. contradiction. rewrite homo_comp_p1; auto. f_equal. unfold sig_fun_app. assert (h62:Inn Cr x). unfold Cr, im_proj2_sig. assert (h64:Inn (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (proj1_sig x)). eapply Im_intro. apply h56. reflexivity. pose proof (h41 (proj1_sig x) h64) as h63. pose proof (Add_intro1 _ _ r _ h63) as h65. apply Im_intro with (exist _ _ h65). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. destruct classic_dec as [h63 | h64]. pose proof (h4 (exist _ _ h63)) as h65. destruct h65 as [h65 h66]. rewrite h66 at 1. pose proof h3 as h3'. destruct h3' as [h3l h3r]. rewrite h3r. f_equal. apply proj1_sig_injective; auto. contradiction. assumption. assumption. rewrite h43 at 1. assumption. simpl in h8. rewrite times_set_sing_p' in h8. assert (h16:Inn (ba_p_ens Cp) (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa))). pose proof (subtract_preserves_inclusion _ _ r h7) as h17. rewrite sub_add_same_nin in h17; auto. assert (h18:Included (Subtract F (exist _ _ h10)) F). auto with sets. pose (restriction a h18) as a'. pose proof (el_prod_in_gen_p a') as h19. assert (h20:Included (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (ba_p_ens Cp)). red. intros x h21. destruct h21 as [x h21 ? h22]. rewrite h22. clear h22. clear y. destruct h21 as [h21 h22]. rewrite <- (in_im_proj1_sig_iff (ba_p_ens Bp) F x) in h21 at 1. apply h7 in h21. inversion h21 as [p h23' | p h24]. assumption. contradict h22. inversion h24 as [h25]. assert (h26:exist _ _ h10 = exist _ _ (proj2_sig x)). apply proj1_sig_injective. simpl. assumption. rewrite h26 at 1. rewrite unfold_sig. constructor. assert (h21:Included (Gen_Ens_p (im_proj1_sig (Subtract F (exist (Inn (ba_p_ens Bp)) r h10))) (incl_ens_btp (Subtract F (exist (Inn (ba_p_ens Bp)) r h10)))) (ba_p_ens Cp)). pose proof (gen_ens_p_subalg_of_p h1a _ h20 (incl_ens_btp (Subtract F (exist (Inn (ba_p_ens Bp)) r h10)))) as h22. rewrite <- h22 at 1. apply incl_gen_ens_p. apply h21. unfold el_prod_p, a' in h19. assert (h24: proj1_sig (times_set_p (Im (Subtract F (exist (Inn (ba_p_ens Bp)) r h10)) (fun i : btp Bp => eps_p i (restriction a h18 |-> i))) (finite_image (btp Bp) (btp Bp) (Subtract F (exist (Inn (ba_p_ens Bp)) r h10)) (fun i : btp Bp => eps_p i (restriction a h18 |-> i)) (fin_map_fin_dom (restriction a h18)))) = proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)). f_equal. apply times_set_functional_p. apply im_ext_in. intros x h22. rewrite restriction_compat. reflexivity. assumption. rewrite <- h24. clear h24. assumption. pose proof (one_max_sub_r h1a (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)))) as h18. assert (h19: le_p_sub (exist _ _ h16) (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)))). assert (h21: le_p (sub_lift h1a (exist _ _ h16)) (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)))). rewrite le_iff_p. rewrite <- h8. f_equal. unfold sub_lift. simpl. apply proj1_sig_injective. simpl. reflexivity. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply (le_p_sub_intro (exist _ _ h16) (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c))) h1a h21). pose proof (h5 (exist _ _ h16) %1 h19 h18) as h17. destruct h17 as [h17a h17b]. clear h17b. rewrite le_iff in h17a. pose proof h3 as h3'. red in h3'. destruct h3' as [h20 h21]. rewrite h21 in h17a. simpl in h17a. destruct h1a as [h1l [h1m h1r]]. assert (h22:Inn Cr (exist _ _ (h1l _ h16))). unfold Cr. apply Im_intro with (exist _ _ (h20 _ h16)). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. assert (h23':Inn Cr (exist _ _ h10)). unfold Cr. pose proof (Add_intro2 _ (ba_p_ens Cp) r) as h23'. apply Im_intro with (exist _ _ h23'). constructor. apply proj1_sig_injective. simpl. reflexivity. pose proof (h4 (exist _ _ h22)) as h24. pose proof (h4 (exist _ _ h23')) as h25. simpl in h24, h25. destruct h24 as [h24a h24b], h25 as [h25a h25b]. assert (h26:g' (exist _ _ (Add_intro2 _ (ba_p_ens Cp) r)) = g' (exist _ _ h25a)). f_equal. apply proj1_sig_injective; auto. rewrite <- h26 in h25b. clear h26. assert (h27:g' (exist (Inn (Add (ba_p_ens Cp) r)) (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) (h20 (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16)) = g' (exist _ _ h24a)). f_equal. apply proj1_sig_injective; auto. rewrite <- h27 in h24b. rewrite <- h24b in h17a at 1. rewrite <- h25b in h17a at 1. rewrite (el_prod_compose_decompose_restriction_p1 (sig_fun_app g'' 0) _ _ h28). unfold el_prod_compose_p1. rewrite h34 at 1. simpl. pose proof h3 as h3'. red in h3'. destruct h3' as [h3a h3b]. pose proof (h3b (exist _ _ h16)) as h40. simpl in h40. assert (h41: h3a (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16 = (h20 (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16)). apply proof_irrelevance. rewrite h41 in h40 at 1. rewrite <- h40 in h24b at 1. clear h41. assert (h42: sig_fun_app g'' 0 (exist (Inn (ba_p_ens Bp)) r h10) = g'' (exist (Inn Cr) (exist (Inn (ba_p_ens Bp)) r h10) h23')). unfold sig_fun_app. destruct classic_dec as [h43 | h44]. f_equal. apply proj1_sig_injective; auto. contradiction. rewrite h42 at 1. clear h42. rewrite h24b in h17a at 1. clear h24b. assert (h41:Included (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (ba_p_ens Cp)). red. intros x h42. destruct h42 as [x h42 ? h43]. rewrite h43. clear h43. destruct h42 as [h42 h43]. assert (h44:Inn (im_proj1_sig F) (proj1_sig x)). eapply Im_intro; auto. assumption. pose proof (h7 _ h44) as h45. inversion h45 as [c h46 h46'| c h47 h47']. assumption. inversion h47; clear h47. contradict h43. destruct x as [x h48]. simpl in h47'. simpl in H0. assert (h49 : exist (Inn (ba_p_ens Bp)) r h10 = exist (fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0) x h48). apply proj1_sig_injective. simpl. assumption. rewrite h49. constructor. assert (h41' : Included (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (ba_p_ens Bp)). auto with sets. pose (im_proj2_sig _ h41) as F'. assert (hfin:Finite F'). apply finite_image; auto. rewrite <- finite_full_sig_iff. apply finite_image. eapply Finite_downward_closed. apply h6. red; intros d h50. destruct h50; auto. pose (fun c:sig_set F' => (a |-> (exist _ _ (h1l _ (proj2_sig (proj1_sig c)))))) as fa. assert (h50: Included (Im (full_sig F') fa) signe). red. intros x ?. destruct x. left. right. pose (sig_fun_to_fin_map_ran ba_eq_dec_p fa hfin mns signe signe_finite h50) as a'. assert (h43: times_set (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps (sig_fun_app g'' 0 i) (restriction a (setminus_inc F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) |-> i))) (finite_image (btp Bp) (bt A) (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps (sig_fun_app g'' 0 i) (restriction a (setminus_inc F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) |-> i)) (fin_map_fin_dom (restriction a (setminus_inc F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))))) = g (exist (Inn (ba_p_ens Cp)) (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16)). assert (h42: exist (Inn (ba_p_ens Cp)) (proj1_sig (times_set_p (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) hfa)) h16 = times_set_p (Im F' (fun i : btp Cp => eps_p i (a' |-> i))) (finite_image _ _ _ _ hfin)). apply proj1_sig_injective. simpl. assert (h51:im_proj1_sig (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (fun i : btp Bp => eps_p i (a |-> i))) = im_proj1_sig (Im F' (fun i : btp Cp => eps_p i (a' |-> i)))). unfold F',im_proj1_sig, im_proj2_sig. do 3 rewrite im_im. rewrite im_full_sig_im_eq. simpl. apply Extensionality_Ensembles. red; split; red. intros x h52. destruct h52 as [x h52 ? h53]. rewrite h53. clear h53 y. apply Im_intro with (exist _ _ h52). constructor. simpl. unfold a'. assert (h54: Inn F' (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h52 (proj1_sig x) eq_refl)))). unfold F'. unfold im_proj2_sig, im_proj1_sig. rewrite im_full_sig_im_eq. simpl. apply Im_intro with (exist _ _ h52). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. pose proof (sig_fun_to_fin_map_ran_compat ba_eq_dec_p fa hfin mns signe signe_finite h50 (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h52 (proj1_sig x) eq_refl))) h54) as h53. unfold F' in h53. unfold F'. rewrite h53 at 1. unfold fa. simpl. unfold eps_p. assert (h55:x = exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h1l (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h52 (proj1_sig x) eq_refl)))). apply proj1_sig_injective. simpl. reflexivity. rewrite <- h55. destruct (a |-> x). simpl. reflexivity. assert (h57: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h1l h1m)) (proj1_sig x)). rewrite <- h1r. apply h41. eapply Im_intro. apply h52. reflexivity. pose proof (ba_p_subst_comp _ _ h1r _ (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h52 (proj1_sig x) eq_refl)) h57) as h56. rewrite h56. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. intros x h51. destruct h51 as [x h51 ? h52]. rewrite h52. clear h52 y h51. destruct x as [x h51]. simpl. apply Im_intro with x. assumption. unfold a'. assert (h53: Inn F' (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h51 (proj1_sig x) eq_refl)))). unfold F', im_proj2_sig, im_proj1_sig. rewrite im_full_sig_im_eq. simpl. apply Im_intro with (exist _ _ h51). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. pose proof (sig_fun_to_fin_map_ran_compat ba_eq_dec_p fa hfin mns signe signe_finite h50 (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h51 (proj1_sig x) eq_refl))) h53) as h52. unfold F' in h52. unfold F'. rewrite h52 at 1. unfold fa, eps_p. simpl. assert (h54:x = exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h1l (proj1_sig x) (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h51 (proj1_sig x) eq_refl)))). apply proj1_sig_injective. simpl. reflexivity. rewrite <- h54. destruct (a |-> x). simpl. reflexivity. assert (h56: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h1l h1m)) (proj1_sig x)). rewrite <- h1r. apply h41. eapply Im_intro. apply h51. reflexivity. pose proof (ba_p_subst_comp _ _ h1r _ (h41 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h51 (proj1_sig x) eq_refl)) h56) as h55. rewrite h55 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective. simpl. reflexivity. apply (times_set_sub_compat_p _ _ h1a' _ _ hfa (finite_image {x : T | Inn (ba_p_ens Cp) x} (btp Cp) F' (fun i : btp Cp => eps_p i (a' |-> i)) hfin) h51). rewrite h42 at 1. clear h42. rewrite homo_times_set_p1. pose proof (im_im F' (fun i : btp Cp => eps_p i (a' |-> i)) g) as h51. pose proof (finite_image _ _ _ (fun x : {x : T | Inn (ba_p_ens Cp) x} => g ((fun i : btp Cp => eps_p i (a' |-> i)) x)) hfin) as h52. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (btp Cp) (bt A) (Im F' (fun i : btp Cp => eps_p i (a' |-> i))) g (finite_image {x : T | Inn (ba_p_ens Cp) x} (btp Cp) F' (fun i : btp Cp => eps_p i (a' |-> i)) hfin)) h52 h51) as h53. dependent rewrite -> h53. clear h53. unfold F', im_proj2_sig, im_proj1_sig. simpl. pose proof (im_im (full_sig (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x)))) (fun p : sig_set (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x))) => exist (Inn (ba_p_ens Cp)) (proj1_sig p) (h41 (proj1_sig p) (proj2_sig p))) (fun x : {x : T | Inn (ba_p_ens Cp) x} => g (eps_p x (a' |-> x)))) as h53. pose proof (im_full_sig_im_eq (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x)) (fun x : sig_set (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x : T => Inn (A_p T (Bc_p T Bp)) x))) => (fun x0 : {x0 : T | Inn (ba_p_ens Cp) x0} => g (eps_p x0 (a' |-> x0))) ((fun p : sig_set (Im (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0))) => exist (Inn (ba_p_ens Cp)) (proj1_sig p) (h41 (proj1_sig p) (proj2_sig p))) x))) as him. rewrite him in h53 at 1. clear him. simpl in h53. assert (h54: Finite (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))). apply finite_setminus; auto. assert (h55:Finite (Im (full_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (fun x : sig_set (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) => g (eps_p (exist (Inn (ba_p_ens Cp)) (proj1_sig (proj1_sig x)) (h41 (proj1_sig (proj1_sig x)) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (proj1_sig x) (proj2_sig x) (proj1_sig (proj1_sig x)) eq_refl))) (a' |-> exist (Inn (ba_p_ens Cp)) (proj1_sig (proj1_sig x)) (h41 (proj1_sig (proj1_sig x)) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (proj1_sig x) (proj2_sig x) (proj1_sig (proj1_sig x)) eq_refl))))))). rewrite <- h53. rewrite im_im. rewrite im_full_sig_im_eq. simpl. apply finite_image. rewrite <- finite_full_sig_iff. assumption. assert (h57:times_set _ h52 = times_set _ h55). apply times_set_functional. assumption. unfold F' in h57. rewrite h57 at 1. clear h57. apply times_set_functional. apply Extensionality_Ensembles; red; split; red. intros x h58. destruct h58 as [x h58 ? h59]. rewrite h59 at 1. clear h59 y. apply Im_intro with (exist _ _ h58). constructor. simpl. rewrite eps_p_eq. unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. unfold eps. rewrite restriction_compat. unfold a'. assert (h59:Inn F' (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h58 (proj1_sig x) eq_refl)))). unfold F', im_proj2_sig, im_proj1_sig. rewrite im_full_sig_im_eq. simpl. apply Im_intro with (exist _ _ h58). constructor. apply proj1_sig_injective; auto. pose proof (sig_fun_to_fin_map_ran_compat ba_eq_dec_p fa hfin mns signe signe_finite h50 (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h58 (proj1_sig x) eq_refl))) h59) as h60. unfold F' in h60. unfold F'. rewrite h60 at 1. unfold fa. simpl. assert (h61:x = exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h1l (proj1_sig x) (h41 (proj1_sig x) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h58 (proj1_sig x) eq_refl)))). apply proj1_sig_injective; auto. rewrite <- h61. destruct (a |-> x). unfold sig_fun_app. assert (h62:Inn Cr x). unfold Cr, im_proj2_sig. assert (h64:Inn (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (proj1_sig x)). eapply Im_intro. apply h58. reflexivity. pose proof (h41 (proj1_sig x) h64) as h63. pose proof (Add_intro1 _ _ r _ h63) as h65. apply Im_intro with (exist _ _ h65). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. destruct classic_dec as [h63 | h64]. pose proof (h4 (exist _ _ h63)) as h65. destruct h65 as [h65 h66]. rewrite h66 at 1. pose proof h3 as h3'. destruct h3' as [h3l h3r]. rewrite h3r. simpl. f_equal. apply proj1_sig_injective; auto. contradiction. rewrite homo_comp_p1; auto. f_equal. unfold sig_fun_app. assert (h62:Inn Cr x). unfold Cr, im_proj2_sig. assert (h64:Inn (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (proj1_sig x)). eapply Im_intro. apply h58. reflexivity. pose proof (h41 (proj1_sig x) h64) as h63. pose proof (Add_intro1 _ _ r _ h63) as h65. apply Im_intro with (exist _ _ h65). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. destruct classic_dec as [h63 | h64]. pose proof (h4 (exist _ _ h63)) as h65. destruct h65 as [h65 h66]. rewrite h66 at 1. pose proof h3 as h3'. destruct h3' as [h3l h3r]. rewrite h3r. f_equal. apply proj1_sig_injective; auto. contradiction. assumption. intros x h56. destruct h56 as [x h56 q h57]. rewrite h57. clear h57 q h56. destruct x as [x h56]. simpl. apply Im_intro with x. assumption. rewrite eps_p_eq. unfold ba_conv_elt, ba_conv_type. rewrite transfer_eq_refl. unfold eps. rewrite restriction_compat. unfold a'. assert (h57:Inn F' (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h41 (proj1_sig x) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h56 (proj1_sig x) eq_refl)))). unfold F', im_proj2_sig, im_proj1_sig. rewrite im_full_sig_im_eq. simpl. apply Im_intro with (exist _ _ h56). constructor. apply proj1_sig_injective; auto. pose proof (sig_fun_to_fin_map_ran_compat ba_eq_dec_p fa hfin mns signe signe_finite h50 _ h57) as h58. unfold F' in h58. unfold F'. rewrite h58 at 1. unfold fa. simpl. assert (h59:x = exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h1l (proj1_sig x) (h41 (proj1_sig x) (Im_intro (btp Bp) T (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10))) (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h56 (proj1_sig x) eq_refl)))). apply proj1_sig_injective; auto. rewrite <- h59. destruct (a |-> x). unfold sig_fun_app. assert (h62:Inn Cr x). unfold Cr, im_proj2_sig. assert (h64:Inn (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (proj1_sig x)). eapply Im_intro. apply h56. reflexivity. pose proof (h41 (proj1_sig x) h64) as h63. pose proof (Add_intro1 _ _ r _ h63) as h65. apply Im_intro with (exist _ _ h65). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. destruct classic_dec as [h63 | h64]. pose proof (h4 (exist _ _ h63)) as h65. destruct h65 as [h65 h66]. rewrite h66 at 1. pose proof h3 as h3'. destruct h3' as [h3l h3r]. rewrite h3r. simpl. f_equal. apply proj1_sig_injective; auto. contradiction. rewrite homo_comp_p1; auto. f_equal. unfold sig_fun_app. assert (h62:Inn Cr x). unfold Cr, im_proj2_sig. assert (h64:Inn (im_proj1_sig (Setminus F (Singleton (exist (Inn (ba_p_ens Bp)) r h10)))) (proj1_sig x)). eapply Im_intro. apply h56. reflexivity. pose proof (h41 (proj1_sig x) h64) as h63. pose proof (Add_intro1 _ _ r _ h63) as h65. apply Im_intro with (exist _ _ h65). constructor. simpl. apply proj1_sig_injective. simpl. reflexivity. destruct classic_dec as [h63 | h64]. pose proof (h4 (exist _ _ h63)) as h65. destruct h65 as [h65 h66]. rewrite h66 at 1. pose proof h3 as h3'. destruct h3' as [h3l h3r]. rewrite h3r. f_equal. apply proj1_sig_injective; auto. contradiction. assumption. assumption. rewrite h43 at 1. rewrite comm_prod. assumption. pose proof (f_equal (sig_fun_app g'' 0) h8) as h14. pose proof h3 as h3'. destruct h3' as [h3a h3b]. assert (h15:sig_fun_app g'' 0 %0 = 0). unfold sig_fun_app. destruct classic_dec as [h15 | h16]. pose proof (h4 (exist _ _ h15)) as h17. destruct h17 as [h17 h18]. rewrite h18 at 1. simpl. specialize (h3b %0). simpl in h17. destruct h1a' as [h1l [h1m h1r]]. pose proof (ba_p_subst_zero _ _ h1r) as h19. rewrite homo_zero_p1 in h3b. rewrite h3b. f_equal. apply proj1_sig_injective. simpl. symmetry. assumption. assumption. reflexivity. rewrite h15 in h14. rewrite <- h14 at 2. unfold el_prod_compose_p1. assert (h16:Included (im_proj1_sig F) (ba_p_ens Cp)). eapply incl_add_nin_incl. apply h7. contradict h13. destruct h13 as [r h13 q h16]. assert (h17:r = exist _ _ h10). apply proj1_sig_injective. simpl. rewrite h16 at 1. reflexivity. rewrite h17 in h13. assumption. unfold sig_fun_app. pose proof (el_prod_in_gen_p a) as h17. pose proof (gen_ens_p_subalg_of_p h1a _ h16 (incl_ens_btp F)) as h18. rewrite <- h18 in h17. pose proof (incl_gen_ens_p (im_proj1_sig F) h16) as h19. apply h19 in h17. unfold el_prod_p in h17. destruct classic_dec as [h20 | h21]. pose proof (h4 (exist _ _ h20)) as h21. destruct h21 as [h21 h22]. rewrite h22 at 1. simpl. pose proof (h3b (exist _ _ h17)) as h23. simpl in h23. assert (h24: g' (exist (Inn (Add (ba_p_ens Cp) r)) (proj1_sig (times_set_p (Im F (fun i : btp Bp => eps_p i (a |-> i))) (finite_image (btp Bp) (btp Bp) F (fun i : btp Bp => eps_p i (a |-> i)) (fin_map_fin_dom a)))) (h3a (proj1_sig (times_set_p (Im F (fun i : btp Bp => eps_p i (a |-> i))) (finite_image (btp Bp) (btp Bp) F (fun i : btp Bp => eps_p i (a |-> i)) (fin_map_fin_dom a)))) h17)) = g' (exist (Inn (Add (ba_p_ens Cp) r)) (proj1_sig (times_set_p (Im F (fun i : btp Bp => eps_p i (a |-> i))) (finite_image (btp Bp) (btp Bp) F (fun i : btp Bp => eps_p i (a |-> i)) (fin_map_fin_dom a)))) h21)). f_equal. apply proj1_sig_injective. simpl. reflexivity. rewrite <- h24 at 1. rewrite <- h23 at 1. pose proof (times_set_sub_compat_p _ _ h1a). pose (im_proj2_sig _ h16) as F'. destruct h1a' as [h1l [h1m h1r]]. pose (fun x:sig_set F' => a |-> (exist _ _ (h1l _ (proj2_sig (proj1_sig x))))) as fa. assert (h25:Finite F'). unfold F'. apply finite_image. rewrite <- finite_full_sig_iff. apply finite_image; auto. assert (h26: Included (Im (full_sig F') fa) signe). red. intros s h26. destruct s; constructor. pose (sig_fun_to_fin_map_ran ba_eq_dec_p fa h25 mns signe signe_finite h26) as a'. assert (h27:im_proj1_sig (Im F (fun i : btp Bp => eps_p i (a |-> i))) = im_proj1_sig (Im F' (fun i:btp Cp => eps_p i (a' |-> i)))). unfold im_proj1_sig, F', im_proj2_sig. do 2 rewrite im_im. rewrite im_im. unfold im_proj1_sig. rewrite im_full_sig_im_eq. simpl. rewrite (im_full_sig_proj1_sig F) at 1. rewrite im_im. apply im_ext_in. intros x h27. destruct x as [x h27']. simpl. unfold a'. assert (h29: Inn F' (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h16 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T F (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h27' (proj1_sig x) eq_refl)))). unfold F', im_proj2_sig, im_proj1_sig. rewrite im_full_sig_im_eq. simpl. apply Im_intro with (exist _ _ h27'). constructor. apply proj1_sig_injective. simpl. reflexivity. pose proof (sig_fun_to_fin_map_ran_compat ba_eq_dec_p fa h25 mns signe signe_finite h26 (exist (Inn (ba_p_ens Cp)) (proj1_sig x) (h16 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T F (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h27' (proj1_sig x) eq_refl))) h29) as h28. unfold F' in h28. unfold F'. rewrite h28 at 1. unfold fa. simpl. unfold eps_p. assert (h30: exist (Inn (ba_p_ens Bp)) (proj1_sig x) (h1l (proj1_sig x) (h16 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T F (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h27' (proj1_sig x) eq_refl))) = x). apply proj1_sig_injective. simpl. reflexivity. rewrite h30. destruct (a |-> x). simpl. reflexivity. assert (h32: Inn (ba_p_ens (Subalg_p Bp (ba_p_ens Cp) h1l h1m)) (proj1_sig x)). rewrite <- h1r. apply h16. apply Im_intro with x; auto. pose proof (ba_p_subst_comp _ _ h1r _ (h16 (proj1_sig x) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T F (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) x h27' (proj1_sig x) eq_refl)) h32) as h31. rewrite h31 at 1. simpl. f_equal. f_equal. apply proj1_sig_injective; auto. assert (h28:Finite (Im F (fun i : btp Bp => eps_p i (a |-> i)))). apply finite_image; auto. assert (h29:Finite (Im F' (fun i : btp Cp => eps_p i (a' |-> i)))). apply finite_image; auto. pose proof (times_set_sub_compat_p _ _ h1a _ _ h28 h29 h27) as h30. assert (h31:exist _ _ h17 = exist _ _ (proj2_sig (times_set_p (Im F' (fun i : btp Cp => eps_p i (a' |-> i))) h29))). apply proj1_sig_injective. simpl. assert (h31:h28 = (finite_image (btp Bp) (btp Bp) F (fun i : btp Bp => eps_p i (a |-> i)) (fin_map_fin_dom a))). apply proof_irrelevance. rewrite <- h31. assumption. rewrite h31. rewrite <- unfold_sig. rewrite homo_times_set_p1. apply times_set_functional. rewrite im_im. unfold F', im_proj2_sig, im_proj1_sig. rewrite im_im. rewrite im_full_sig_im_eq. simpl. rewrite (im_full_sig_proj1_sig F) at 1. rewrite im_im. f_equal. apply functional_extensionality. intro x. destruct classic_dec as [h33 | h34]. unfold eps, eps_p. assert (h34: a' |-> exist (Inn (ba_p_ens Cp)) (proj1_sig (proj1_sig x)) (h16 (proj1_sig (proj1_sig x)) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T F (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (proj1_sig x) (proj2_sig x) (proj1_sig (proj1_sig x)) eq_refl)) = a |-> proj1_sig x). unfold a'. assert (h36: Inn F' (exist (Inn (ba_p_ens Cp)) (proj1_sig (proj1_sig x)) (h16 (proj1_sig (proj1_sig x)) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T F (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (proj1_sig x) (proj2_sig x) (proj1_sig (proj1_sig x)) eq_refl)))). unfold F', im_proj1_sig, im_proj2_sig. rewrite im_full_sig_im_eq. simpl. apply Im_intro with x. constructor. apply proj1_sig_injective; auto. pose proof (sig_fun_to_fin_map_ran_compat ba_eq_dec_p fa h25 mns signe signe_finite h26 (exist (Inn (ba_p_ens Cp)) (proj1_sig (proj1_sig x)) (h16 (proj1_sig (proj1_sig x)) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T F (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (proj1_sig x) (proj2_sig x) (proj1_sig (proj1_sig x)) eq_refl))) h36) as h35. unfold F' in h35. unfold F'. rewrite h35 at 1. unfold fa. simpl. f_equal. apply proj1_sig_injective. simpl. reflexivity. unfold F' in h34. unfold F'. rewrite h34 at 1. destruct (a |-> proj1_sig x). pose proof (h4 (exist _ _ h33)) as h35. destruct h35 as [h35 h36]. rewrite h36 at 1. pose proof (h3b (exist _ _ (h16 (proj1_sig (proj1_sig x)) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T F (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (proj1_sig x) (proj2_sig x) (proj1_sig (proj1_sig x)) eq_refl)))) as h37. simpl in h37. rewrite h37 at 1. f_equal. apply proj1_sig_injective; auto. rewrite homo_comp_p1. f_equal. pose proof (h4 (exist _ _ h33)) as h35. destruct h35 as [h35 h36]. rewrite h36 at 1. pose proof (h3b (exist _ _ (h16 (proj1_sig (proj1_sig x)) (Im_intro (sig_set (A_p T (Bc_p T Bp))) T F (proj1_sig (P:=fun x0 : T => Inn (A_p T (Bc_p T Bp)) x0)) (proj1_sig x) (proj2_sig x) (proj1_sig (proj1_sig x)) eq_refl)))) as h37. simpl in h37. rewrite h37 at 1. f_equal. apply proj1_sig_injective; auto. assumption. contradict h34. unfold Cr. destruct x as [x h32]. simpl. assert (h33:Inn (im_proj1_sig F) (proj1_sig x)). apply Im_intro with x; auto. apply h7 in h33. apply Im_intro with (exist _ _ h33). constructor. apply proj1_sig_injective. simpl. reflexivity. assumption. contradict h21. unfold Cr. rewrite h8 at 1. unfold im_proj2_sig. assert (h20:Inn (Add (ba_p_ens Cp) r) (proj1_sig (Bzero_p T (Bc_p T Cp)))). left. apply proj2_sig. apply Im_intro with (exist _ _ h20). constructor. simpl. apply proj1_sig_injective. simpl. destruct h1a' as [h1l [h1m h1r]]. pose proof (ba_p_subst_zero _ _ h1r) as h21. unfold ba_p_ens. rewrite h21. simpl. reflexivity. Qed. (*I'm taking the time to add the final condition involving a supremum, in this theorem, in order to then pass it to the next theorem so as to utilize [constructive_definite_description] as opposed to constructive_indefinite_description. A good technique to use is to add as many properties as you can to an existential condition to force unique existence, as you'll see with the next theorem.*) Lemma homo_into_complete_extends_to_one_step_extension : forall {T:Type} (Bp Cp:Bool_Alg_p T) (r:T) (pfo:one_step_extends Bp Cp r) (A:Bool_Alg) (pfc:set_complete A), forall (g:btp Cp->bt A), homomorphism_p1 _ g -> exists f:btp Bp->bt A, homomorphism_p1 _ f /\ extends_sig f g /\ let P := Im [p : btp Cp | le_p_sub p (exist _ _ (one_step_extends_in Bp Cp r pfo))] g in f (exist _ _ (one_step_extends_in Bp Cp r pfo)) = sup_set_complete P pfc. intros T Bp Cp r h1 A h2 g h3. destruct (classic (Inn (ba_p_ens Cp ) r)) as [hi | hni]. pose proof h1 as h1'. apply trivial_step_extends_eq in h1'; auto. subst. exists g. split; auto. split. red. exists (inclusion_reflexive _). intro x. f_equal. apply proj1_sig_injective. simpl. reflexivity. simpl. pose proof (sup_set_complete_compat (Im [p : btp Cp | le_p_sub p (exist (Inn (ba_p_ens Cp)) r (one_step_extends_in Cp Cp r h1))] g) h2) as ha. assert (hb:sup (Im [p : btp Cp | le_p_sub p (exist (Inn (ba_p_ens Cp)) r (one_step_extends_in Cp Cp r h1))] g) (g (exist _ _ (one_step_extends_in Cp Cp r h1)))). red. split. red. intros x h5. destruct h5 as [x h5]. subst. destruct h5 as [h5]. destruct h5 as [h5 h6]. apply homo_mono_p1; auto. unfold sub_lift in h6. assert (h7:proj2_sig x = (subalg_of_p_incl Cp Cp h5 (proj1_sig x) (proj2_sig x))). apply proof_irrelevance. rewrite <- h7 in h6. destruct x as [x h9]. simpl in h6. assumption. intros b h5. red in h5. apply h5. apply Im_intro with (exist _ _ (one_step_extends_in Cp Cp r h1)). constructor. pose proof (le_p_sub_intro (exist _ _ (one_step_extends_in Cp Cp r h1)) (exist _ _ (one_step_extends_in Cp Cp r h1)) (refl_subalg_of_p _ Cp)) as h6. apply h6. clear h6. unfold sub_lift. simpl. assert (h8: subalg_of_p_incl Cp Cp (refl_subalg_of_p T Cp) r (one_step_extends_in Cp Cp r h1) = one_step_extends_in Cp Cp r h1). apply proof_irrelevance. rewrite h8. apply refl_le_p. reflexivity. eapply sup_unq. apply hb. apply ha. pose proof h1 as h1'. pose proof (one_step_extends_in Bp Cp r h1) as h4. pose (Im [p:btp Cp | le_p_sub p (exist _ _ h4)] g) as P. pose (Im [q:btp Cp | le_p_sub_r (exist _ _ h4) q] g) as Q. assert (h5: forall p q:btp Cp, le_p_sub p (exist _ _ h4) -> le_p_sub_r (exist _ _ h4) q -> le (g p) (sup_set_complete P h2) /\ le (sup_set_complete P h2) (g q)). pose proof (sup_set_complete_compat P h2) as h7. pose proof (inf_set_complete_compat Q h2) as h8. assert (h9:forall p q:btp Cp, le_p_sub p (exist (Inn (ba_p_ens Bp)) r h4) -> le_p_sub_r (exist (Inn (ba_p_ens Bp)) r h4) q -> le (g p) (g q)). intros p q h9 h10. apply homo_mono_p1; auto. destruct h9 as [h9a h9b], h10 as [h10a h10b]. assert (h11:le_p (sub_lift h9a p) (sub_lift h10a q)). apply (trans_le_p (sub_lift h9a p) (exist (Inn (ba_p_ens Bp)) r h4) (sub_lift h10a q)); auto. assert (h9a = h10a). apply proof_irrelevance. subst. rewrite (le_p_sub_lift_iff h10a p q). assumption. assert (h10:lb Q (sup_set_complete P h2)). red in h7. destruct h7 as [h7a h7b]. assert (h10:forall q:bt A, Inn Q q -> ub P q). intros q h10. destruct h10 as [q h10]. subst. destruct h10 as [h10]. red. intros p h11. destruct h11 as [p h11]. subst. destruct h11 as [h11]. apply h9; auto. red. intros q h11. apply h7b; auto. assert (h11:le (sup_set_complete P h2) (inf_set_complete Q h2)). red in h8. destruct h8 as [h8a h8b]. specialize (h8b _ h10). assumption. intros p q h12 h13. split. specialize (h9 _ _ h12 h13). assert (h14:Inn P (g p)). unfold P. apply Im_intro with p. constructor. assumption. reflexivity. destruct h7 as [h7a h7b]. red in h7a. apply h7a; auto. destruct h8 as [h8a h8b]. red in h8a. assert (h14:Inn Q (g q)). apply Im_intro with q; auto. constructor. assumption. specialize (h8a _ h14). apply (trans_le (sup_set_complete P h2) (inf_set_complete Q h2) (g q)); auto. destruct h1 as [h1a h1b h1c]. pose (fun x:sig_set (Add (ba_p_ens Cp) r) => match (in_add_nin_dec _ _ hni _ (proj2_sig x)) with | left P => (g (exist _ _ P)) | _ => (sup_set_complete P h2) end) as g'. pose (im_proj2_sig (Add (ba_p_ens Cp) r) (one_step_extends_incl Bp Cp r h1')) as Cr. assert (h6:forall x:(btp Bp), Inn Cr x -> Inn (Add (ba_p_ens Cp) r) (proj1_sig x)). intros x h7. destruct h7. clear h1c. subst. simpl. apply proj2_sig. pose (fun x:sig_set Cr => g' (exist _ _ (h6 _ (proj2_sig x)))) as g''. assert (h7:im_proj1_sig Cr = Add (ba_p_ens Cp) r). unfold im_proj1_sig, Cr, im_proj2_sig. rewrite im_im. simpl. rewrite im_full_sig_proj1_sig. f_equal. assert (h8:extends_sig g' g). red. assert (h9: Included (A_p T (Bc_p T Cp)) (Add (ba_p_ens Cp) r)). auto with sets. exists h9. unfold g'. intro x. simpl. destruct (in_add_nin_dec (ba_p_ens Cp) r hni (proj1_sig x) (h9 (proj1_sig x) (proj2_sig x))) as [h10 | h11]. f_equal. apply proj1_sig_injective. simpl. reflexivity. destruct x as [x h12]. simpl in h11. clear h1c. subst. contradiction. assert (h9:forall x:sig_set Cr, g'' x = g' (exist _ _ (h6 _ (proj2_sig x)))). intros; auto. unfold Cr in g''. assert (h11:forall x : sig_set (im_proj2_sig (Add (ba_p_ens Cp) r) (one_step_extends_incl Bp Cp r h1')), exists pfin : Inn (Add (ba_p_ens Cp) r) (proj1_sig (proj1_sig x)), g'' x = g' (exist (Inn (Add (ba_p_ens Cp) r)) (proj1_sig (proj1_sig x)) pfin)). intro x. destruct x as [x h12]. simpl. destruct h12 as [x h12 q h13]. rewrite h13. clear h13 q. simpl. destruct x as [x h13]. simpl. exists h13. rewrite h9. simpl. f_equal. apply proj1_sig_injective; auto. pose proof (one_step_extends_homo_ext_crit_iff Bp Cp r h1' hni A g g' h3 h8 g'' h11) as h10. simpl in h10. assert (h12: forall p q : btp Cp, le_p_sub p (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r h1')) -> le_p_sub_r (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r h1')) q -> le (g p) (g' (exist (Inn (Add (ba_p_ens Cp) r)) r (Add_intro2 T (ba_p_ens Cp) r))) /\ le (g' (exist (Inn (Add (ba_p_ens Cp) r)) r (Add_intro2 T (ba_p_ens Cp) r))) (g q)). intros p q h13 h14. assert (h15:h4 = one_step_extends_in Bp Cp r h1'). apply proof_irrelevance. rewrite <- h15 in h13, h14. pose proof (h5 _ _ h13 h14) as h16. destruct h16 as [h16 h17]. unfold g'. split. simpl. destruct (in_add_nin_dec (ba_p_ens Cp) r hni r (Add_intro2 T (ba_p_ens Cp) r)) as [h18 | h19]. contradiction. assumption. simpl. destruct (in_add_nin_dec (ba_p_ens Cp) r hni r (Add_intro2 T (ba_p_ens Cp) r)) as [h18 | h19]. contradiction. assumption. rewrite <- h10 in h12. pose proof (f_equal (@ba_p_ens _) h1c) as h13. rewrite ba_p_ens_gen_p_eq in h13. assert (h16:h1b = one_step_extends_incl Bp Cp r h1'). apply proof_irrelevance. rewrite h16 in h13. assert (h14:Gen_Ens_p (im_proj1_sig (im_proj2_sig (Add (ba_p_ens Cp) r) (one_step_extends_incl Bp Cp r h1'))) (incl_im_proj1_sig_ba_p_ens (im_proj2_sig (Add (ba_p_ens Cp) r) (one_step_extends_incl Bp Cp r h1'))) = Gen_Ens_p (Add (ba_p_ens Cp) r) (one_step_extends_incl Bp Cp r h1')). assert (h14: im_proj1_sig (im_proj2_sig (Add (ba_p_ens Cp) r) (one_step_extends_incl Bp Cp r h1')) = Add (ba_p_ens Cp) r). rewrite <- im_proj1_sig_undoes_im_proj2_sig; auto. pose proof (subsetT_eq_compat _ (fun S=>Included S (ba_p_ens Bp)) _ _ (incl_im_proj1_sig_ba_p_ens (im_proj2_sig (Add (ba_p_ens Cp) r) (one_step_extends_incl Bp Cp r h1'))) (one_step_extends_incl Bp Cp r h1') h14) as h15. unfold ba_p_ens in h15. unfold ba_p_ens. dependent rewrite -> h15. reflexivity. rewrite <- h14 in h13. pose proof (homo_extension_criterion_p1 (A:=A) h13 g'') as h17. assert (h18: forall F : Ensemble (btp Bp), Finite F -> Included F (im_proj2_sig (Add (ba_p_ens Cp) r) (one_step_extends_incl Bp Cp r h1')) -> forall a : Fin_map F signe mns, el_prod_p T a = %0 -> el_prod_compose_p1 T (sig_fun_app g'' 0) a = 0). intros F h19 h20 a h21. assert (h22: Included (im_proj1_sig F) (Add (ba_p_ens Cp) r)). red. intros x h23. destruct h23 as [x h23 q h24]. rewrite h24. clear q h24. apply h20 in h23. destruct h23 as [x h23 q h24]. rewrite h24. clear q h24. simpl. apply proj2_sig. specialize (h12 _ h19 h22 a h21). assumption. rewrite <- h17 in h18. destruct h18 as [f [h18 h19]]. exists f. split. assumption. unfold g'', g' in h19. red in h19. split. red. destruct h1a as [h1l [h1m h1r]]. exists h1l. intro x. destruct x as [x h20]. simpl. assert (h21:Inn (Add (ba_p_ens Cp) r) x). auto with sets. assert (h22:Inn (im_proj2_sig (Add (ba_p_ens Cp) r) (one_step_extends_incl Bp Cp r h1')) (exist _ _ (h1b _ h21))). apply Im_intro with (exist _ _ h21). constructor. apply proj1_sig_injective; auto. specialize (h19 (exist _ _ h22)). simpl in h19. destruct ( in_add_nin_dec (ba_p_ens Cp) r hni x (h6 (exist (Inn (ba_p_ens Bp)) x (h1b x h21)) h22)) as [h23 | h24]. assert (h24:h1l x h20 = h1b x h21). apply proof_irrelevance. rewrite h24. rewrite <- h19 at 1. f_equal. apply proj1_sig_injective; auto. pose proof h20 as h20'. rewrite h24 in h20'. contradiction. simpl. assert (h20:Inn Cr (exist _ _ (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)))). unfold Cr. apply Im_intro with (exist _ _ (Add_intro2 _ _ r)). constructor. simpl. apply proj1_sig_injective; auto. specialize (h19 (exist _ _ h20)). simpl in h19. destruct (in_add_nin_dec (ba_p_ens Cp) r hni r (h6 (exist (Inn (ba_p_ens Bp)) r (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c))) h20)) as [h22 | h23]. contradiction. unfold P in h19. assert (h24:h4 = one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)). apply proof_irrelevance. rewrite <- h24. rewrite h19. f_equal. apply proj1_sig_injective. simpl. reflexivity. Qed. (*Uniuqe existence, the next step from last theorem is now proved, with again the supremum condition added to allow for construcive_definite_description in the future.*) Lemma homo_into_complete_extends_to_one_step_extension_unq : forall {T:Type} (Bp Cp:Bool_Alg_p T) (r:T) (pfo:one_step_extends Bp Cp r) (A:Bool_Alg) (pfc:set_complete A), forall (g:btp Cp->bt A), homomorphism_p1 _ g -> exists! f:btp Bp->bt A, homomorphism_p1 _ f /\ extends_sig f g /\ let P := Im [p : btp Cp | le_p_sub p (exist _ _ (one_step_extends_in Bp Cp r pfo))] g in f (exist _ _ (one_step_extends_in Bp Cp r pfo)) = sup_set_complete P pfc. intros T Bp Cp r h1 A h2 g h3. pose proof (homo_into_complete_extends_to_one_step_extension Bp Cp r h1 A h2 g h3) as h4. destruct h4 as [f h4]. exists f. red. split. assumption. intros f' h5. destruct h4 as [h4a [h4b h4c]]. destruct h5 as [h5a [h5b h5c]]. simpl in h4c. destruct h1 as [h1a h1b h1c]. assert (h7: Included (Add (ba_p_ens Cp) r) (ba_p_ens Bp)). rewrite h1c. rewrite ba_p_ens_gen_p_eq. apply gen_ens_includes_p. assert (h8: ba_p_ens Bp = Gen_Ens_p (Add (ba_p_ens Cp) r) h7). rewrite h1c at 1. rewrite ba_p_ens_gen_p_eq. f_equal. apply proof_irrelevance. assert (h9: agree_on f f' (im_proj2_sig (Add (ba_p_ens Cp) r) h7)). red. intros x h9. destruct x as [x h10]. destruct h4b as [h13 h14], h5b as [h15 h16]. inversion h9 as [q h11 ? h12 h17]. apply exist_injective in h12. destruct q as [q ha]. simpl. simpl in h10, h12. destruct y as [y hb]. apply exist_injective in h17. destruct ha as [q ha | q ha]. specialize (h14 (exist _ _ ha)). specialize (h16 (exist _ _ ha)). simpl in h14, h16. assert (h18:f (exist _ _ h10) = f (exist _ _ (h13 q ha))). f_equal. apply proj1_sig_injective. simpl. assumption. assert (h19:f' (exist _ _ h10) = f' (exist _ _ (h15 q ha))). f_equal. apply proj1_sig_injective. simpl. assumption. rewrite h18, h19 at 1. rewrite <- h14, h16 at 1. reflexivity. destruct ha. assert (h20:f (exist _ _ h10) = f (exist _ _ (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)))). f_equal. apply proj1_sig_injective. simpl. assumption. assert (h21:f' (exist _ _ h10) = f' (exist _ _ (one_step_extends_in Bp Cp r (one_step_extends_intro Bp Cp r h1a h1b h1c)))). f_equal. apply proj1_sig_injective. simpl. assumption. rewrite h20, h21 at 1. rewrite h4c, h5c at 1. reflexivity. pose proof (gen_ens_determines_homo_p1 _ f f' h4a h5a (Add (ba_p_ens Cp) r) h7 h8 h9) as h6. assumption. Qed. End HomoExtensionTheorem.bool2-v0-3/FiniteMaps.v0000644000175000017500000051174513575574055015605 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-2015 Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Import SetUtilities. Require Import FunctionalExtensionality. Require Import ProofIrrelevance. Require Import Setoid. Require Import Description. Require Import DecidableDec. Require Import ProofIrrelevance. Require Import FunctionProperties. Require Import TypeUtilities. Require Import LogicUtilities. Require Import Equality. Require Import EqDec. Require Import SetUtilities1_5. Require Import ArithUtilities. (*All of the [classic]s/LEMs that still exist here, are connected to a function [sig_fun_app], which in retrospect was too hasty and wasteful, since it seems that everything that [sig_fun_app] does can be done with [fun_in_ens], only without the [classic_dec]s. Replacing everything will not be ready for the Winter 2019 release, ([ExtensionsOfHomomorphisms] and [Subalgebras] are HUGE!), so next year, I'll hopefully eliminate [sig_fun_app], and keep this FiniteMaps library completely portable and only dependent on decidable equality on the underlying type of the domain set.*) (*This constructs an explicit function from the inductive definition of [functionally_paired].*) Definition fin_fps_to_f {T U:Type} {A:Ensemble T} {B:Ensemble U} (pfdt:type_eq_dec T) (pffin:Finite A) (S:Ensemble (T*U)) (pffp:functionally_paired A B S) (def:U) : T->U := let (f, _) := pffp in fun x:T => match (in_fin_ens_dec pfdt A pffin x) with | left pfi => proj1_sig ((constructive_definite_description _ (f _ pfi))) | right _ => def end. Lemma fin_fps_to_f_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} (pfdt:type_eq_dec T) (pff:Finite A) (S:Ensemble (T*U)) (pf:functionally_paired A B S) (def:U), let f := fin_fps_to_f pfdt pff S pf def in forall (x:T), (In A x -> In S (x, (f x))) /\ (~In A x -> f x = def). unfold fin_fps_to_f. intros T U A B h1 h2 S h3 def x. split. intro h6. destruct h3 as [h3 h4]. destruct (in_fin_ens_dec h1 A h2 x) as [h7 | h7]. destruct constructive_definite_description as [h9 h10]. simpl. destruct h10; auto. contradiction. intro h6. destruct h3 as [h3 h4]. destruct (in_fin_ens_dec h1 A h2 x) as [h7 | h7]; auto. contradiction. Qed. Lemma fin_fps_to_f_in_ran : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)} (pfdt:type_eq_dec T) (pff:Finite A) (pf:functionally_paired A B S) (def:U) (x:T), In A x -> In B (fin_fps_to_f pfdt pff _ pf def x). intros T U A B S hdt hf h1 def x h2. pose proof (fin_fps_to_f_compat hdt hf S h1 def x) as h3. destruct h3 as [h3 h4]. specialize (h3 h2). destruct h1 as [h1 h5]. pose proof (h5 _ h3) as h6. simpl in h6. destruct h6; auto. Qed. Lemma fin_fps_to_f_def : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)} (pfdt:type_eq_dec T) (pff:Finite A) (pf:functionally_paired A B S) (def:U) (x:T), ~ In A x -> fin_fps_to_f pfdt pff _ pf def x = def. intros T U A B S hdt hf h1 def x h2. pose proof (fin_fps_to_f_compat hdt hf S h1 def x) as h3. destruct h3 as [h3 h4]. apply h4; auto. Qed. Lemma fin_fps_to_f_compat' : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)} (pfdt:type_eq_dec T) (pff:Finite A) (pf:functionally_paired A B S) (def:U), let f := fin_fps_to_f pfdt pff _ pf def in S = [pr:T*U | snd pr = f (fst pr) /\ In A (fst pr)]. intros T U A B S hdt hf h1 def f. unfold f. apply Extensionality_Ensembles; red; split; red; intros pr h2; destruct pr as [x y]. constructor; simpl. pose proof (fin_fps_to_f_compat hdt hf S h1 def x) as h3. destruct h3 as [h3 h4]. destruct h1 as [h1 h5]. pose proof (h5 _ h2) as h5'. simpl in h5'. destruct h5' as [h7 h8]. split; auto. pose proof (h1 _ h7) as h9. destruct h9 as [y' h9]. red in h9. destruct h9 as [h9 h10]. pose proof (h10 _ (conj h8 h2)) as h11. subst. apply h10. split. apply fin_fps_to_f_in_ran; auto. apply h3; auto. destruct h2 as [h2]. simpl in h2. destruct h2 as [h2 h3]. subst. destruct h1 as [h1 h4]. pose proof (h1 _ h3) as h1'. destruct h1' as [y' h5]. red in h5. destruct h5 as [h5 h6]. destruct h5 as [h5 h7]. pose proof (fin_fps_to_f_compat hdt hf S (functional_pairs_intro A B S h1 h4) def x) as h8. destruct h8; auto. Qed. Lemma fp_ran_rel_im: forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)} (pfdt:type_eq_dec T) (pff:Finite A) (pf:functionally_paired A B S) (def:U), ran_rel S = Im (dom_rel S) (fin_fps_to_f pfdt pff _ pf def). intros T U A B S hdt hf h1 def. pose proof (fin_fps_to_f_compat hdt hf S h1 def) as h2. simpl in h2. destruct h1 as [h1a h1b]. apply Extensionality_Ensembles; red; split; red; intros y h3. destruct h3 as [h3]. destruct h3 as [x h3]. apply Im_intro with x. constructor; auto. exists y; auto. specialize (h2 x). destruct h2 as [h2 h4]. pose proof (h1b _ h3) as h5. simpl in h5. destruct h5 as [h5 h6]. pose proof (h1a _ h5) as h7. destruct h7 as [y' h8]. red in h8. destruct h8 as [h8 h9]. destruct h8 as [h8 h10]. pose proof (h9 _ (conj h6 h3)) as h9'. subst. apply h9. split. apply fin_fps_to_f_in_ran; auto. apply h2; auto. destruct h3 as [x h3]. subst. destruct h3 as [h3]. destruct h3 as [y h4]. constructor. exists x. specialize (h2 x). destruct h2 as [h2 h3]. apply h2. pose proof (h1b _ h4) as h5. simpl in h5. destruct h5; auto. Qed. Lemma fp_inj : forall {T U:Type} {A A':Ensemble T} {B B':Ensemble U} (pfdt:type_eq_dec T) (pfdt':type_eq_dec T) (pff:Finite A) (pff':Finite A') (S S':Ensemble (T*U)) (pf:functionally_paired A B S) (pf':functionally_paired A' B' S'), S = S' -> A = A'. intros T U A A' B B' hdt hdt' h1 h2 S S' h3 h4 h5. subst. pose proof (fp_dom_rel h3). subst. pose proof (fp_dom_rel h4). subst. reflexivity. Qed. Lemma fin_fps_to_f_functional : forall {T U:Type} {A A':Ensemble T} {B B':Ensemble U} (pfdt:type_eq_dec T) (pfdt':type_eq_dec T) (pff:Finite A) (pff':Finite A') (S S':Ensemble (T*U)) (pf:functionally_paired A B S) (pf':functionally_paired A' B' S') (def:U) (x:T), S = S' -> fin_fps_to_f pfdt pff S pf def x = fin_fps_to_f pfdt' pff' S' pf' def x. intros T U A A' B B' hdt hdt' h1 h2 S S' h3 h4 def x h5. subst. pose proof (fp_inj hdt hdt' h1 h2 S' S' h3 h4 eq_refl). subst. pose proof (fin_fps_to_f_compat hdt h1 S' h3 def x) as h5. pose proof (fin_fps_to_f_compat hdt' h2 S' h4 def x) as h6. destruct h5 as [h5 h7], h6 as [h6 h8]. destruct (in_fin_ens_dec hdt A' h2 x) as [h9 | h9]. specialize (h5 h9). specialize (h6 h9). eapply fp_functional. apply h3. apply h5. apply h6. specialize (h7 h9). specialize (h8 h9). congruence. Qed. Lemma fp_ran_rel_im' : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)} (pfdt:type_eq_dec T) (pff:Finite A) (pf:functionally_paired A B S) (def:U), ran_rel S = Im A (fin_fps_to_f pfdt pff _ pf def). intros T U A B S hdt hf h1 def. pose proof (fp_ran_rel_im hdt hf h1 def) as h2. rewrite (fp_dom_rel h1) in h2. assumption. Qed. Lemma fp_sub_fin_fps_to_f : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)} (pfdt:type_eq_dec T) (pff:Finite A) (pf:functionally_paired A B S) (x:T) (def:U), In A x -> functionally_paired (Subtract A x) B (Subtract S (x, (fin_fps_to_f pfdt pff _ pf def x))). intros T U A B S hdt hf h1 x def h2. pose proof h1 as h1'. destruct h1' as [h1a h1b]. pose proof (fin_fps_to_f_compat hdt hf _ h1 def x) as h3. destruct h3 as [h3 h4]. specialize (h3 h2). constructor. intros x' h5'. destruct h5' as [h5 h6]. exists (fin_fps_to_f hdt hf S h1 def x'). red. split. split. apply fin_fps_to_f_in_ran; auto. constructor. rewrite (fin_fps_to_f_compat' hdt hf h1 def) at 1. constructor. simpl; split; auto. intro h7. inversion h7. subst. apply h6. constructor. intros y h3'. destruct h3' as [h8 h9]. inversion h9 as [h10 h11]. clear h9. specialize (h1a _ h5). destruct h1a as [y' h13]. red in h13. destruct h13 as [h13 h14]. destruct h13 as [h13 h15]. pose proof (h14 _ (conj h8 h10)). subst. symmetry. pose proof (fin_fps_to_f_compat hdt hf S h1 def x') as h3'. destruct h3' as [h3a h3b]. apply h14. split. apply fin_fps_to_f_in_ran; auto. apply h3a; auto. intros pr h5. destruct h5 as [h5l h5r]. inversion h1 as [h1l h1r]. specialize (h1r _ h5l). destruct h1r as [h10 h11]. split. constructor; auto. intro h5. inversion h5. subst. rewrite (fin_fps_to_f_compat' hdt hf h1 def) in h5l. destruct h5l as [h5l]. destruct h5l as [h5a h5b]. rewrite <- h5a in h5r. contradict h5r. rewrite surjective_pairing. constructor. assumption. Qed. Lemma fp_sub_add' : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)} (pfdt:type_eq_dec T) (pff:Finite A) (pf:functionally_paired A B S), forall (x:T) (y:U) (def:U), In A x -> In B y -> ~ Inn S (x, y) -> functionally_paired A B (Subtract (Add S (x, y)) (x, (fin_fps_to_f pfdt pff _ pf def x))). intros T U A B S hdt hf h1 x y def h0' hb h0. assert (h2:y <> fin_fps_to_f hdt hf _ h1 def x). intro h2. subst. contradict h0. rewrite (fin_fps_to_f_compat' hdt hf h1 def) at 1. constructor. simpl. split; auto. rewrite sub_add_comm. pose proof (fp_sub_fin_fps_to_f hdt hf h1 _ def h0') as h3. assert (h4:~ In (Subtract A x) x). intro h4. destruct h4. contradict H0. constructor. assert (h5:~In (Subtract S (x,fin_fps_to_f hdt hf _ h1 def x)) (x, y)). intro h6. destruct h6. contradiction. pose proof (fp_add h3 _ _ h4 hb h5) as h6. rewrite add_sub_compat_in in h6. assumption. assumption. intro h3. inversion h3. contradiction. Qed. Lemma fin_fps_to_f_s_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {S:Ensemble (T*U)} (pfdt:type_eq_dec T) (pff:Finite A) (pf:functionally_paired A B S) (def:U), forall x:T, In A x -> In S (x, fin_fps_to_f pfdt pff _ pf def x). intros T U A B S hdt hf h1 def x h2. pose proof (fin_fps_to_f_compat' hdt hf h1 def) as h3. simpl in h3. assert (h4:In [pr : T * U | snd pr = fin_fps_to_f hdt hf _ h1 def (fst pr) /\ In A (fst pr)] (x, fin_fps_to_f hdt hf _ h1 def x)). constructor. simpl. split; auto. rewrite <- h3 in h4. assumption. Qed. Inductive Fin_map {T U:Type} (A:Ensemble T) (B:Ensemble U) (def:U) : Type := fin_map_intro : type_eq_dec T -> Finite A -> Finite B -> forall S:Ensemble (T*U), functionally_paired A B S -> Fin_map A B def. Lemma fin_map_fin_dom : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (f:Fin_map A B def), Finite A. intros T U A B def f. destruct f; assumption. Qed. Lemma fin_map_fin_ran : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (f:Fin_map A B def), Finite B. intros T U A B def f. destruct f; assumption. Qed. Lemma fin_map_type_eq_dec : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), type_eq_dec T. intros ? ? ? ? ? F. destruct F; auto. Qed. Definition fin_map_app {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:(Fin_map A B def)): T->U. intro x. destruct F as [hdt ha ? S h1]. refine ((fin_fps_to_f hdt ha _ h1 def) x). Defined. Notation "F |-> x" := (fin_map_app F x) (at level 20). Lemma fin_map_intro_compat : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (def:U) (pfa:Finite A) (pfb:Finite B) (S:Ensemble (T*U)) (pffp:functionally_paired A B S) (pr:T*U), In S pr -> fin_map_intro A B def pfdt pfa pfb S pffp |-> fst pr = snd pr. intros T U hdt A B def h1 h2 S h3 pr h4. simpl. pose proof (fin_fps_to_f_compat' hdt h1 h3 def) as h5. simpl in h5. rewrite h5 in h4. clear h5. destruct h4 as [h4]. destruct h4 as [h4 h5]. rewrite h4. reflexivity. Qed. Lemma fin_map_app_in : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (f:Fin_map A B def) (x:T), In A x -> In B (f |-> x). intros T U A B def f x h1. destruct f as [hdt h2 h3 S h4]. destruct h4 as [h4 h5]. unfold fin_map_app. pose proof (fin_fps_to_f_compat hdt h2 S (functional_pairs_intro A B S h4 h5) def x) as h6. destruct h6 as [h6 h7]. specialize (h6 h1). pose proof (h5 _ h6) as h8. simpl in h8. destruct h8; auto. Qed. Lemma fin_map_app_def : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (f:Fin_map A B def) (x:T), ~In A x -> f |-> x = def. intros T U A B def f x h1. destruct f as [hdt h2 h3 S h4]. unfold fin_map_app. pose proof (fin_fps_to_f_compat hdt h2 S h4 def x) as h5. destruct h5 as [h5 h6]. apply h6; auto. Qed. Lemma fin_map_app_s_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} (def:U) (pfdt:type_eq_dec T) (pfa:Finite A) (pfb:Finite B) (S1 S2:Ensemble (T*U)) (pf1:functionally_paired A B S1) (pf2:functionally_paired A B S2), S1 = S2 <-> fin_map_intro A B def pfdt pfa pfb S1 pf1 = fin_map_intro A B def pfdt pfa pfb S2 pf2. intros T U A B def hdt pfa pfb S1 S2 h1 h2. (* -> *) split. intro h3. subst. pose proof (proof_irrelevance _ h1 h2). subst. reflexivity. (* <- *) intro h3. apply NNPP. intro h4. destruct h1 as [h1l h1r]. destruct h2 as [h2l h2r]. assert (h5:(exists pr:(T*U), In S1 pr /\ ~In S2 pr) \/ (exists pr:(T*U), In S2 pr /\ ~In S1 pr)). apply NNPP. intro h6. pose proof (not_or_and _ _ h6) as h7. destruct h7 as [h7l h7r]. pose proof (not_ex_all_not _ _ h7l) as h8. pose proof (not_ex_all_not _ _ h7r) as h9. simpl in h8. simpl in h9. assert (h10:S1 = S2). apply Extensionality_Ensembles. red. split. (* <= *) red. intros pr h1. specialize (h8 pr). tauto. (* >= *) red. intros pr h1. specialize (h9 pr). tauto. contradiction. assert (h6:S1 = S2). injection h3. tauto. contradiction. Qed. Definition im_fin_map {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) := Im A (fin_map_app F). Lemma in_im_fin_map : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) (x:T), In A x -> In (im_fin_map F) (F|->x). intros T U A B def F x h1. apply Im_intro with x; auto. Qed. Lemma im_fin_map_finite : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), Finite (im_fin_map F). intros T U A B def F. unfold im_fin_map. pose proof (fin_map_fin_dom F) as h1. apply finite_image. assumption. Qed. Lemma im_fin_map_ran : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (f:Fin_map A B def), Included (im_fin_map f) B. intros T U A B def f. unfold im_fin_map. red. intros u h1. destruct h1 as [x h1]. subst. apply fin_map_app_in. assumption. Qed. Lemma fin_map_card_im_le_dom : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), (card_fun (im_fin_map F) (im_fin_map_finite F)) <= (card_fun A (fin_map_fin_dom F)). intros T U A B def F. pose proof (card_fun_compat A (fin_map_fin_dom F)) as h1. pose proof (card_fun_compat (im_fin_map F) (im_fin_map_finite F)) as h2. unfold im_fin_map. apply (cardinal_decreases _ _ _ (fin_map_app F) _ h1 _ h2). Qed. Lemma fin_map_card_im_le_dom' : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), (card_fun1 (im_fin_map F)) <= (card_fun1 A). intros T U A B def F. pose proof (card_fun1_compat A) as h1. pose proof (card_fun1_compat (im_fin_map F)) as h2. destruct h1 as [h1l h1r]. destruct h2 as [h2l h2r]. specialize (h1l (fin_map_fin_dom F)). specialize (h2l (im_fin_map_finite F)). unfold im_fin_map. eapply cardinal_decreases. apply h1l. apply h2l. Qed. Lemma fp_inv : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (S:Ensemble (T*U)) (def:U) (pff:Finite A) (pr:T*U) (pf:functionally_paired A B S), In S pr -> (fin_fps_to_f pfdt pff _ pf def) (fst pr) = snd pr. intros T U hdt A B S def hf pr h1 h2. pose proof (fin_fps_to_f_compat hdt hf S h1 def (fst pr)) as h3. destruct h3 as [h3 h4]. inversion h1 as [h1a h1b]. pose proof (h1b _ h2) as h1b'. destruct h1b' as [h5 h6]. pose proof (h1a _ h5) as h1a'. destruct h1a' as [y h7]. red in h7. destruct h7 as [h7l h7r]. specialize (h3 h5). symmetry. pose proof (fp_in_ran _ _ _ h1 _ h3) as h8. simpl in h8. pose proof (h7r _ (conj h8 h3)). subst. rewrite surjective_pairing in h2. symmetry. apply h7r. split; auto. Qed. Lemma fp_inv' : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (S:Ensemble (T*U)) (def:U) (pff:Finite A) (pf:functionally_paired A B S) (x:T) (y:U), In S (x, y) -> fin_fps_to_f pfdt pff _ pf def x= y. intros T U hdt A B S def hf h1 x y h2. assert (h3:x = fst (x, y)). auto. assert (h4:y = snd (x, y)). auto. rewrite h3, h4. apply fp_inv; auto. Qed. Lemma fin_fps_to_f_inj : forall {T U:Type} (pfdt:type_eq_dec T) {A:Ensemble T} {B:Ensemble U} (S1 S2:Ensemble (T*U)) (def:U) (pff:Finite A) (pf1:functionally_paired A B S1) (pf2:functionally_paired A B S2), fin_fps_to_f pfdt pff _ pf1 def = fin_fps_to_f pfdt pff _ pf2 def -> S1 = S2. intros T U hdt A B S1 S2 def hf h3 h4 h5. pose proof (fin_fps_to_f_compat hdt hf S1 h3 def) as h6. simpl in h6. pose proof (fin_fps_to_f_compat hdt hf S2 h4 def) as h7. simpl in h6, h7. inversion h3 as [h8 h9]. inversion h4 as [h10 h11]. apply Extensionality_Ensembles; red; split; red; intros pr h12; destruct pr as [x y]. pose proof (h9 _ h12) as h9'. simpl in h9'. destruct h9' as [h9' h13]. pose proof (h7 x) as h7'. destruct h7' as [h7' h14]. specialize (h7' h9'). rewrite <- h5 in h7'. pose proof (h8 _ h9') as h8'. destruct h8' as [y' h8']. red in h8'. destruct h8' as [h15 h16]. pose proof (h16 y (conj h13 h12)) as h18. subst. pose proof (fin_fps_to_f_in_ran hdt hf h3 def _ h9') as h17. specialize (h6 x). destruct h6 as [h6a h6b]. specialize (h6a h9'). specialize (h16 _ (conj h17 h6a)). subst. assumption. pose proof (h11 _ h12) as h11'. simpl in h11'. destruct h11' as [h11' h13]. pose proof (h6 x) as h6'. destruct h6' as [h6' h14]. specialize (h6' h11'). rewrite h5 in h6'. pose proof (h10 _ h11') as h10'. destruct h10' as [y' h10']. red in h10'. destruct h10' as [h15 h16]. pose proof (h16 y (conj h13 h12)) as h18. subst. pose proof (fin_fps_to_f_in_ran hdt hf h3 def _ h11') as h17. specialize (h7 x). destruct h7 as [h7a h7b]. specialize (h7a h11'). rewrite <- h5 in h7a. specialize (h16 _ (conj h17 h7a)). subst. rewrite h5. assumption. Qed. Lemma fin_fps_to_f_inj' : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B C:Ensemble U) (S1 S2:Ensemble (T*U)) (def:U) (pff:Finite A) (pf1:functionally_paired A B S1) (pf2:functionally_paired A C S2), fin_fps_to_f pfdt pff _ pf1 def= fin_fps_to_f pfdt pff _ pf2 def-> S1 = S2. intros T U hdt A B C S1 S2 def hf h1 h2 h3. assert (h4:forall (A':Ensemble T) (B' C':Ensemble U) (S1' S2':Ensemble (T*U)) (h1':functionally_paired A' B' S1') (h2':functionally_paired A' C' S2') (pff':Finite A') (h3':fin_fps_to_f hdt pff' _ h1' def = fin_fps_to_f hdt pff' _ h2' def), Included S1' S2'). intros A' B' C' S1' S2' h1' h2' hf' h3'. (* <= *) red. intros pr h4. inversion h1' as [h1a h1b]; inversion h2' as [h2a h2b]. pose (fin_fps_to_f hdt hf' _ h1' def (fst pr)). pose proof (fp_in_dom _ _ _ h1' _ h4) as h5. specialize (h2a _ h5). destruct h2a as [y h6]. red in h6. destruct h6 as [h6l h6r]. destruct h6l as [h6la h6lb]. pose proof (fp_inv hdt _ _ _ def hf' _ h2' h6lb) as h6. simpl in h6. rewrite <- h3' in h6. pose proof (fin_fps_to_f_compat hdt hf' S1' h1' def (fst pr)) as h7. destruct h7 as [h7a h7b]. specialize (h7a h5). specialize (h1a _ h5). specialize (h1b _ h7a). simpl in h1b. destruct h1b as [h1b h1c]. destruct h1a as [y' h7]. red in h7. destruct h7 as [h7c h7d]. subst. pose proof (h7d _ (conj h1c h7a)). subst. rewrite surjective_pairing in h4. pose proof (fp_functional h1' _ _ _ h7a h4) as h8. rewrite surjective_pairing. rewrite <- h8. assumption. apply Extensionality_Ensembles. red; split. eapply h4. apply h3. symmetry in h3. eapply h4. apply h3. Qed. Lemma fin_map_app_inj : forall {T U : Type} {A : Ensemble T} {B : Ensemble U} {def : U} (f1 f2:Fin_map A B def), fin_map_app f1 = fin_map_app f2 -> f1 = f2. intros T U A B def f1 f2 h1. unfold fin_map_app in h1. destruct f1 as [hdt h1a h1b S1 h2]. destruct f2 as [hdt' h2a h2b S2 h3]. pose proof (type_eq_dec_eq hdt hdt'). subst. pose proof (proof_irrelevance _ h1a h2a). pose proof (proof_irrelevance _ h1b h2b); subst. rewrite <- fin_map_app_s_compat. eapply fin_fps_to_f_inj. apply h1. Qed. Lemma fin_map_to_fps_lemma : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), let pfdt := fin_map_type_eq_dec F in let pff := fin_map_fin_dom F in exists! (S:Ensemble (T*U)), exists (pf:functionally_paired A B S), fin_fps_to_f pfdt pff _ pf def = fin_map_app F. intros T U A B def F. simpl. simpl. destruct F as [hdt' h1 h2 S h3]. exists S. red. split. exists h3. unfold fin_map_app. apply functional_extensionality. intro x. f_equal. apply type_eq_dec_eq. apply proof_irrelevance. intros S' h4. destruct h4 as [h4 h5]. unfold fin_map_app in h5. symmetry. eapply fin_fps_to_f_inj. rewrite h5 at 1. apply functional_extensionality. intro. f_equal. apply type_eq_dec_eq. apply proof_irrelevance. Qed. Definition fin_map_to_fps {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) : Ensemble (T*U) := proj1_sig (constructive_definite_description _ (fin_map_to_fps_lemma F)). Lemma fin_map_to_fps_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), exists (pf:functionally_paired A B (fin_map_to_fps F)), fin_fps_to_f (fin_map_type_eq_dec F) (fin_map_fin_dom F) _ pf def = fin_map_app F. intros T U A B def F. unfold fin_map_to_fps. destruct constructive_definite_description as [S h1]. simpl. assumption. Qed. Lemma fin_map_to_fps_inj : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (def:U), injective (@fin_map_to_fps T U A B def). intros T U A B def. red. intros F G h1. pose proof (fin_map_to_fps_compat F) as h2. pose proof (fin_map_to_fps_compat G) as h3. destruct h2 as [h2a h2b]. destruct h3 as [h3a h3b]. apply fin_map_app_inj. rewrite <- h2b. rewrite <- h3b. pose proof (subsetT_eq_compat _ _ _ _ h2a h3a h1) as h4. dependent rewrite -> h4. f_equal. apply type_eq_dec_eq. apply proof_irrelevance. Qed. Lemma fin_map_ext : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (f1 f2:Fin_map A B def), (forall x:T, In A x -> f1 |-> x = f2 |-> x) -> f1 = f2. intros T U A B def f1 f2 h1. destruct f1 as [hdt' h2 h3 S h4]. destruct f2 as [hdt'' h5 h6 S' h7]. pose proof (proof_irrelevance _ h2 h5). pose proof (proof_irrelevance _ h3 h6). pose proof (type_eq_dec_eq hdt' hdt''). subst. rewrite <- (fin_map_app_s_compat _ hdt'' h5 h6). inversion h4 as [h9 h10]. inversion h7 as [h11 h12]. apply Extensionality_Ensembles; red; split; red; intros pr h13. pose proof (fin_fps_to_f_compat hdt'' h5 S' h7 def (fst pr)) as h17. destruct h17 as [h17 h18]. specialize (h10 _ h13). destruct h10 as [h10 h14]. specialize (h17 h10). pose proof (fin_fps_to_f_compat hdt'' h5 S h4 def (fst pr)) as h19. destruct h19 as [h19 h20]. specialize (h19 h10). specialize (h11 _ h10). destruct h11 as [y h11]. red in h11. destruct h11 as [h11 h15]. destruct h11 as [h11 h16]. specialize (h1 _ h10). simpl in h1. rewrite h1 in h19. rewrite surjective_pairing in h13. pose proof (fp_functional h4 _ _ _ h13 h19) as h21. rewrite <- h21 in h17. rewrite surjective_pairing. assumption. pose proof (fin_fps_to_f_compat hdt'' h5 S' h7 def (fst pr)) as h17. destruct h17 as [h17 h18]. specialize (h12 _ h13). destruct h12 as [h12a h12b]. specialize (h1 _ h12a). pose proof (fin_fps_to_f_compat hdt'' h5 S h4 def (fst pr)) as h19. destruct h19 as [h19 h20]. specialize (h19 h12a). specialize (h17 h12a). rewrite surjective_pairing in h13. pose proof (fp_functional h7 _ _ _ h13 h17) as h21. simpl in h1. rewrite <- h1 in h21. rewrite <- h21 in h19. rewrite surjective_pairing. assumption. Qed. Lemma fin_map_ext_iff : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (f1 f2:Fin_map A B def), (forall x:T, In A x -> f1 |-> x = f2 |-> x) <-> f1 = f2. intros T U A B def f1 f2. split. apply fin_map_ext; auto. intro; subst; auto. Qed. Lemma fin_map_to_fps_s_compat : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (def:U) (pfa:Finite A) (pfb:Finite B) (S:Ensemble (T*U)) (pffp:functionally_paired A B S), S = fin_map_to_fps (fin_map_intro A B def pfdt pfa pfb S pffp). intros T U hdt A B def h1 h2 S h3. unfold fin_map_to_fps. destruct constructive_definite_description as [S' h4]. simpl. destruct h4 as [h4 h5]. rewrite (fin_map_app_s_compat def hdt h1 h2 _ _ h3 h4). apply fin_map_ext. intros x h6. simpl. pose proof (fin_fps_to_f_compat hdt h1 S h3 def x) as h7. pose proof (fin_fps_to_f_compat hdt h1 S' h4 def x) as h8. destruct h7 as [h7a h7b], h8 as [h8a h8b]. specialize (h7a h6). specialize (h8a h6). pose proof (proof_irrelevance _ h1 (fin_map_fin_dom (fin_map_intro A B def hdt h1 h2 S h3))) as h9. rewrite h9. simpl in h5. pose proof (type_eq_dec_eq hdt (fin_map_type_eq_dec (fin_map_intro A B def hdt h1 h2 S h3))) as he. rewrite <- he in h5. rewrite h5. simpl. f_equal. apply proof_irrelevance. Qed. Lemma empty_map1_eq : forall {T U:Type} {B:Ensemble U} {def:U} (f1 f2:Fin_map (Empty_set T) B def), f1 = f2. intros; apply fin_map_ext; intros; contradiction. Qed. Definition inj {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) : Prop := forall x y:T, In A x -> In A y -> F |-> x = F |-> y -> x = y. Definition surj {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) : Prop := forall b:U, In B b -> exists a:T, In A a /\ F |-> a = b. Definition bij {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) : Prop := inj F /\ surj F. Lemma inj_card_im_eq : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), inj F -> card_fun A (fin_map_fin_dom F) = card_fun (im_fin_map F) (im_fin_map_finite F). intros T U A B def F h0. pose proof (card_fun_compat (full_sig A) (iff1 (finite_full_sig_iff A) (fin_map_fin_dom F))) as h1. pose proof (card_fun_compat (full_sig (im_fin_map F)) (iff1 (finite_full_sig_iff (im_fin_map F)) (im_fin_map_finite F))) as h2. rewrite card_fun_full_sig_eq. rewrite (card_fun_full_sig_eq (im_fin_map F)). pose (fun x:(sig_set A) => exist _ (F |-> (proj1_sig x)) (in_im_fin_map F _ (proj2_sig x))) as f. pose proof (injective_preserves_cardinal). assert (h3:Image.injective _ _ f). red. intros x y. unfold f. intro h4. apply exist_injective in h4. apply h0 in h4. apply proj1_sig_injective. assumption. apply proj2_sig. apply proj2_sig. unfold sig_set in h1. unfold sig_set in h2. assert (h4:full_sig (im_fin_map F) = (Im (full_sig A) f)). apply Extensionality_Ensembles. red. split. red. intros y h4. unfold sig_set in y. destruct y as [y h5]. destruct h5 as [x h6 y h5]. subst. apply Im_intro with (exist _ x h6). constructor. unfold f. apply proj1_sig_injective. simpl. reflexivity. red. intros y h5. constructor. assert (h5:cardinal {x : U | In (im_fin_map F) x} (Im (full_sig A) f) (card_fun (full_sig (im_fin_map F)) (iff1 (finite_full_sig_iff (im_fin_map F)) (im_fin_map_finite F)))). rewrite <- h4. assumption. pose proof (injective_preserves_cardinal _ _ _ _ _ h3 h1 _ h5) as h6. symmetry. assumption. Qed. Lemma inj_card_im_eq' : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), inj F -> card_fun1 A = card_fun1 (im_fin_map F). intros T U A B def F h1. do 2 erewrite <- card_fun_card_fun1_compat. apply inj_card_im_eq; auto. Qed. (*See [FiniteMaps2] for the converse of inj_card_im_eq'.*) Lemma surj_im_ran_eq : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), surj F -> (im_fin_map F) = B. intros T U A B def F h1. red in h1. apply Extensionality_Ensembles. red. split. red. intros y h2. pose proof (im_fin_map_ran F) as h3. red in h3. apply h3; auto. red. intros y h2. specialize (h1 _ h2). destruct h1 as [x h1]. destruct h1; subst. apply in_im_fin_map; auto. Qed. Lemma im_ran_eq_surj : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), (im_fin_map F) = B -> surj F. intros T U A B def F h1. red. intros b h2. rewrite <- h1 in h2. destruct h2 as [b h2]. clear h1. subst. exists b. split; auto. Qed. Lemma surj_im_ran_card_eq : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), surj F -> card_fun (im_fin_map F) (im_fin_map_finite F) = card_fun B (fin_map_fin_ran F). intros T U A B def F h1. pose proof (surj_im_ran_eq F h1) as h2. pose proof (subsetT_eq_compat _ _ _ _ (im_fin_map_finite F) (fin_map_fin_ran F) h2) as h3. dependent rewrite -> h3. reflexivity. Qed. Lemma surj_im_ran_card_eq' : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), surj F -> card_fun1 (im_fin_map F) = card_fun1 B. intros T U A B def F h1. do 2 erewrite <- card_fun_card_fun1_compat. apply surj_im_ran_card_eq; auto. Qed. Lemma im_ran_card_eq_surj : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), card_fun1 (im_fin_map F) = card_fun1 B -> surj F. intros T U A B def F h1. pose proof (im_ran_eq_surj F) as h2. apply h2. apply incl_card_fun1_eq; auto. apply (fin_map_fin_ran F). apply im_fin_map_ran. Qed. Lemma bij_dom_ran_card_eq : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), bij F -> card_fun A (fin_map_fin_dom F) = card_fun B (fin_map_fin_ran F). intros T U A B def F h1. red in h1. destruct h1 as [h1l h1r]. rewrite inj_card_im_eq; auto. rewrite surj_im_ran_card_eq; auto. Qed. Lemma bij_dom_ran_card_eq' : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), bij F -> card_fun1 A = card_fun1 B. intros T U A B def F h1. do 2 erewrite <- card_fun_card_fun1_compat. eapply bij_dom_ran_card_eq; auto. apply h1. Qed. Definition fin_map_to_fun {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (f:Fin_map A B def) : T->U. destruct f as [hdt h1 h2 S h3]. refine (fin_fps_to_f hdt h1 _ h3 def). Defined. Lemma fin_map_to_fun_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (f:Fin_map A B def), forall x:T, f |-> x = (fin_map_to_fun f) x. intros T U A B def f x. unfold fin_map_app, fin_map_to_fun. destruct f. reflexivity. Qed. Lemma fin_map_new_ran_lemma : forall {T U:Type} {A:Ensemble T} {B C:Ensemble U} {def:U} (F:Fin_map A B def), Finite C -> Included B C -> exists! F':Fin_map A C def, forall x:T, In A x -> F |-> x = F' |-> x. intros T U A B C def F h1 h2. pose (fin_map_app F) as f. destruct F as [hdt h3 h4 S h6]. assert (h7:functionally_paired A C S). constructor. inversion h6 as [h6a h6b]. intros x h7. specialize (h6a x h7). destruct h6a as [y h6a]. red in h6a. destruct h6a as [h8 h9]. assert (h10:In C y /\ In S (x, y)). intuition auto with sets. exists y. red. split. assumption. intros u h13. destruct h13 as [h13l h13r]. pose proof (fp_in_ran _ _ _ h6 _ h13r) as h14. simpl in h14. specialize (h9 u). tauto. inversion h6 as [h6a h6b]. intros pr h13. specialize (h6b _ h13). intuition auto with sets. pose (fin_map_intro _ _ def hdt h3 h1 _ h7) as F'. exists F'. red. split. intros x ha. unfold F'. pose proof (fin_fps_to_f_s_compat hdt h3 h6 def _ ha) as h9. pose proof (fin_fps_to_f_s_compat hdt h3 h7 def _ ha) as h10. pose proof (fp_functional h6 _ _ _ h9 h10) as h11. assumption. intros G h8. unfold F'. apply fin_map_ext. intros x ha. simpl. pose proof (fin_fps_to_f_s_compat hdt h3 h6 def _ ha) as h9. pose proof (fin_fps_to_f_s_compat hdt h3 h7 def _ ha) as h10. pose proof (fp_functional h6 _ _ _ h9 h10) as h11. rewrite <- h8; auto. Qed. Definition fin_map_new_ran {T U:Type} {A:Ensemble T} {B C:Ensemble U} {def:U} (F:Fin_map A B def) (pfc:Finite C) (pfinc:Included B C) : Fin_map A C def. refine (proj1_sig (constructive_definite_description _ (fin_map_new_ran_lemma F pfc pfinc))). Defined. Lemma fin_map_new_ran_compat : forall {T U:Type} {A:Ensemble T} {B C:Ensemble U} {def:U} (F:Fin_map A B def) (pfc:Finite C) (pfinc:Included B C), let F' := fin_map_new_ran F pfc pfinc in forall x:T, In A x -> F |-> x = F' |-> x. intros T U A B C def F h1 h2 F'. unfold fin_map_new_ran in F'. destruct constructive_definite_description as [F'' h4]. simpl in F'. unfold F'. apply h4. Qed. Lemma fun_to_fin_map_lemma : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (def:U) (pfa:Finite A) (f:T->U), exists! F:Fin_map A (Im A f) def, forall x:T, In A x -> (fin_map_to_fun F) x = f x. intros T U hdt A def h1 f. pose [pr:T*U | snd pr = f (fst pr) /\ In A (fst pr)] as S. assert (h3:functionally_paired A (Im A f) S). constructor. intros x h3. exists (f x). red. repeat split. apply Im_intro with x. assumption. reflexivity. simpl. assumption. intros y h4. destruct h4 as [h4l h4r]. destruct h4l. subst. destruct h4r as [h5]. simpl in h5. symmetry. destruct h5; assumption. intros pr h3. destruct h3 as [h3]. destruct h3 as [h3l h3r]. split; try assumption. rewrite h3l. apply Im_intro with (fst pr). assumption. reflexivity. pose proof (finite_image _ _ _ f h1) as h4. pose (fin_map_intro _ _ def hdt h1 h4 S h3) as F. exists F. red. split. simpl. pose proof (fin_fps_to_f_compat' hdt h1 h3 def) as h5. intros x h6. assert (h7:In S (x, (f x))). constructor. simpl. split; auto. rewrite h5 in h7. destruct h7 as [h8]. simpl in h8. symmetry. destruct h8; assumption. intros F' h5. apply fin_map_ext. intros x ha. simpl. pose proof (fin_fps_to_f_compat hdt h1 S h3 def x) as h7. destruct h7 as [h7 h8]. specialize (h7 ha). specialize (h5 _ ha). pose proof (fin_map_to_fun_compat F' x) as h9. rewrite <- h9 in h5. rewrite h5. destruct h7 as [h7]. simpl in h7. destruct h7; auto. Qed. Definition fun_to_fin_map {T U:Type} (pfte: type_eq_dec T) (A:Ensemble T) (def:U) (pfa:Finite A) (f:T->U) : Fin_map A (Im A f) def := proj1_sig (constructive_definite_description _ (fun_to_fin_map_lemma pfte A def pfa f)). Lemma fun_to_fin_map_compat : forall {T U:Type} (pfte:type_eq_dec T) (A:Ensemble T) (def:U) (pfa:Finite A) (f:T->U), forall x:T, In A x -> (fun_to_fin_map pfte A def pfa f) |-> x = f x. intros T U he A def h1 f x h2. unfold fun_to_fin_map. destruct constructive_definite_description as [F h4]. simpl. specialize (h4 _ h2). rewrite <- h4. apply fin_map_to_fun_compat. Qed. Lemma fin_map_to_fun_undoes_fun_to_fin_map : forall {T U:Type} (pfte: type_eq_dec T) (A:Ensemble T) (pf:Finite A) (def:U) (p:T->U) (x:T), In A x -> fin_map_to_fun (fun_to_fin_map pfte A def pf p) x = p x. intros T U hte A h1 def p x h2. rewrite <- fin_map_to_fun_compat. rewrite fun_to_fin_map_compat;auto. Qed. Lemma im_fin_map_app_undoes_fun_to_fin_map : forall {T U:Type} (pfte:type_eq_dec T) (E:Ensemble T) (f:T->U) (def:U) (pf:Finite E), Im E (fin_map_app (fun_to_fin_map pfte E def pf f)) = Im E f. intros T U h0 E f def h1. apply Extensionality_Ensembles. red. split. red. intros y h2. destruct h2 as [y h2 x]. subst. apply Im_intro with y. assumption. apply fun_to_fin_map_compat. assumption. red. intros y h2. destruct h2 as [y h2 x]. subst. apply Im_intro with y. assumption. rewrite fun_to_fin_map_compat; auto. Qed. Lemma fin_map_to_fps_fun_to_fin_map_functional : forall {T U:Type} {A A':Ensemble T} {def:U} (pfa:Finite A) (pfa':Finite A') (pfet pfet':type_eq_dec T) (f f':T->U), A = A' -> f = f' -> fin_map_to_fps (fun_to_fin_map pfet A def pfa f) = fin_map_to_fps (fun_to_fin_map pfet' A' def pfa' f'). intros T U A A' def hfa hfa' het het' f f' h1 h2. subst. pose proof (proof_irrelevance _ hfa hfa'). pose proof (type_eq_dec_eq het het'); subst. reflexivity. Qed. Lemma fin_map_cart_prod_fin1 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def), Inhabited (cart_prod A B) -> Finite A /\ Finite B. intros T U V A B C def F h1. destruct F as [h4 h5 S h6]. apply cart_prod_fin_rev; assumption. Qed. Lemma fin_map_cart_prod_fin2 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def), A <> Empty_set _ -> Finite B. intros T U V A B C def F h1. inversion F as [h4 h5 S h6]. pose proof (not_empty_Inhabited _ _ h1) as h3. apply cart_prod_fin_rev2 with A; assumption. Qed. Definition fin_map_sing {T U:Type} (pfte:type_eq_dec T) (A:Ensemble T) (pfa:Finite A) (def val:U) := fin_map_intro _ _ def pfte pfa (Singleton_is_finite _ val) _ (fp_cart_prod_sing A val). (*Image of second operand set of domain, with point of first operand, [x], fixed.*) Inductive im1 {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (x:T): Ensemble V := | im1_intro : In A x -> forall y:U, In B y -> forall v:V, v = (F |-> (x, y)) -> In (im1 F x) v. (*Image of first operand set of domain, with point in second operand, [y], fixed.*) Inductive im2 {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (y:U): Ensemble V := | im2_intro : In B y -> forall x:T, In A x -> forall v:V, v = F |-> (x, y) -> In (im2 F y) v. Lemma incl_im1 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (x:T), Included (im1 F x) (im_fin_map F). intros T U V A B C def F x v h1. destruct h1 as [h1 y h2]. econstructor. eapply in_cart_prod. apply h1. apply h2. assumption. Qed. Lemma incl_im2 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (y:U), Included (im2 F y) (im_fin_map F). intros T U V A B C def F y v h1. destruct h1 as [h1 x h2]. econstructor. eapply in_cart_prod. apply h2. apply h1. assumption. Qed. Lemma finite_im1 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (x:T), Finite (im1 F x). intros. eapply Finite_downward_closed. eapply im_fin_map_finite. apply incl_im1. Qed. Lemma finite_im2 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (y:U), Finite (im2 F y). intros. eapply Finite_downward_closed. eapply im_fin_map_finite. apply incl_im2. Qed. Lemma fp_fin_map_to_fps : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), functionally_paired A B (fin_map_to_fps F). intros T U A B def F. pose proof (fin_map_to_fps_compat F) as h1. destruct h1; auto. Qed. Lemma in_fin_map_to_fps : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), forall x:T, In A x -> In (fin_map_to_fps F) (x, F |-> x). intros T U A B def F x h1. unfold fin_map_to_fps. destruct constructive_definite_description as [S h2]. simpl. destruct h2 as [h2 h3]. pose proof (fin_fps_to_f_s_compat (fin_map_type_eq_dec F) (fin_map_fin_dom F) h2 def _ h1) as h4. unfold fin_map_app. destruct F as [hdt h4' h5 S' h6]. simpl in h3. unfold fin_map_app in h3. simpl in h3. pose proof (type_eq_dec_eq hdt (fin_map_type_eq_dec (fin_map_intro A B def hdt h4' h5 S' h6))) as he. pose proof (proof_irrelevance _ (fin_map_fin_dom (fin_map_intro A B def hdt h4' h5 S' h6)) h4') as he'. rewrite <- he, he' in h3. pose proof (fin_fps_to_f_inj hdt _ _ def h4' h2 h6 h3). rewrite <- he,he' in h4. clear he he'. subst. assert (h7:h2 = h6). apply proof_irrelevance. subst; auto. Qed. Definition fin_map_eq {T U:Type} {A:Ensemble T} {B C:Ensemble U} {def:U} (F:Fin_map A B def) (G:Fin_map A C def) := exists pf:Included B C, fin_map_new_ran F (fin_map_fin_ran G) pf = G. Lemma fps_eq_fin_map_eq : forall {T U:Type} {A:Ensemble T} {B C:Ensemble U} {def:U} (F:Fin_map A B def) (G:Fin_map A C def), Included B C -> fin_map_to_fps F = fin_map_to_fps G -> fin_map_eq F G. intros T U A B C def F G h0 h1. unfold fin_map_to_fps in h1. destruct constructive_definite_description as [S h2]; destruct constructive_definite_description as [S' h3]. simpl in h1. destruct h2 as [h2 h2']. destruct h3 as [h3 h3']. red. exists h0. unfold fin_map_new_ran. destruct constructive_definite_description as [G' h4]. simpl. apply fin_map_ext. intros x h5. rewrite <- h4. rewrite <- h2'. rewrite <- h3'. pose proof (fin_fps_to_f_s_compat (fin_map_type_eq_dec F) (fin_map_fin_dom F) h2 def _ h5) as h6. pose proof (fin_fps_to_f_s_compat (fin_map_type_eq_dec G) (fin_map_fin_dom G) h3 def _ h5) as h7. rewrite h2', h3'. rewrite h2' in h6. rewrite h3' in h7. rewrite <- h1 in h7. destruct h2 as [h8 h9]. pose proof (h9 (x, G |-> x) h7) as h10. simpl in h10. destruct h10 as [? h11]. pose proof (fin_map_app_in F x h5) as h12. pose proof (h8 x h5) as h10. destruct h10 as [y h10]. red in h10. destruct h10 as [h10a h10b]. pose proof (h10b _ (conj h12 h6)) as h13. pose proof (h10b _ (conj h11 h7)) as h14. rewrite <- h13, h14. reflexivity. assumption. Qed. Definition extends {T U:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (F':Fin_map A' B def) (F:Fin_map A B def) := Included A A' /\ forall x:T, In A x -> F |-> x = F' |-> x. Lemma restriction_ex : forall {T U:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (F':Fin_map A' B def), Included A A' -> exists! F:Fin_map A B def, forall x:T, In A x -> F |-> x = F' |-> x. intros T U A A' B def F h1. destruct F as [hdt h2 h3 S' h4]. pose [pr:T*U | In S' pr /\ In A (fst pr)] as S. pose proof (functionally_paired_restriction _ _ B S' h1 h4) as h5. exists (fin_map_intro _ _ _ hdt (Finite_downward_closed _ _ h2 _ h1) h3 _ h5). red. split. intros x h6. simpl. pose (Finite_downward_closed _ _ h2 _ h1) as h1'. pose proof (fin_fps_to_f_compat hdt h1' [pr : T * U | In S' pr /\ In A (fst pr)] h5 def x) as hf. pose proof (fin_fps_to_f_compat hdt h2 S' h4 def x) as hf'. simpl in hf, hf'. destruct hf as [hfa hfb], hf' as [hfa' hfb']. specialize (hfa h6). pose proof h6 as h6'. apply h1 in h6'. specialize (hfa' h6'). destruct hfa as [hfa]. simpl in hfa. destruct hfa as [hfa1 hfa2]. eapply fp_functional. apply h4. apply hfa1. apply hfa'. intros F h6. apply fin_map_ext. intros x ha. simpl. specialize (h6 _ ha). rewrite h6. simpl. pose (Finite_downward_closed _ _ h2 _ h1) as h1'. pose proof (fin_fps_to_f_compat hdt h1' [pr : T * U | In S' pr /\ In A (fst pr)] h5 def x) as hf. pose proof (fin_fps_to_f_compat hdt h2 S' h4 def x) as hf'. simpl in hf, hf'. destruct hf as [hfa hfb], hf' as [hfa' hfb']. specialize (hfa ha). pose proof ha as ha'. apply h1 in ha'. specialize (hfa' ha'). destruct hfa as [hfa]. simpl in hfa. destruct hfa as [hfa1 hfa2]. eapply fp_functional. apply h4. apply hfa1. apply hfa'. Qed. Definition restriction {T U:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (F':Fin_map A' B def) (pf:Included A A') : Fin_map A B def. refine (proj1_sig (constructive_definite_description _ (restriction_ex F' pf))). Defined. Lemma restriction_compat : forall {T U:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (F':Fin_map A' B def) (pf:Included A A'), forall x:T, In A x -> (restriction F' pf) |-> x = F' |-> x. intros T U A A' B def F' h1 x h2. unfold restriction. destruct constructive_definite_description as [F h3]. simpl. apply h3; auto. Qed. Lemma restriction_extends_compat : forall {T U:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A' B def) (pf:Included A A'), extends F (restriction F pf). intros T U A A' B def F h1. red. split; auto. apply restriction_compat. Qed. Lemma restriction_extends_compat_iff : forall {T U:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (F':Fin_map A' B def) (F:Fin_map A B def) (pf:Included A A'), extends F' F <-> (restriction F' pf) = F. intros T U A A' B def F' F h1. split. intro h2. red in h2. destruct h2 as [h2l h2r]. apply fin_map_ext. intros x h3. rewrite restriction_compat. symmetry. apply h2r; auto. assumption. intro h2. subst. apply restriction_extends_compat. Qed. Lemma not_extends_compat : forall {T U:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (F':Fin_map A' B def) (F:Fin_map A B def) (pf:Included A A'), ~ extends F' F -> (restriction F' pf) <> F. intros T U A A' B def F' F pf h1. intro h2. contradict h1. rewrite <- restriction_extends_compat_iff in h2. assumption. Qed. Definition fin_map_from_add {T U} {A:Ensemble T} {B:Ensemble U} {def} {a} (F:Fin_map (Add A a) B def) : Fin_map A B def := restriction F (incl_add _ _). Lemma fin_map_from_add_compat : forall {T U} {A:Ensemble T} {B:Ensemble U} {def} {a:T} (F:Fin_map (Add A a) B def) (x:T) (pf:Inn A x), (fin_map_from_add F) |-> x = F |-> x. unfold fin_map_from_add; intros; apply restriction_compat; auto. Qed. Lemma fin_map_from_add_neq : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} {a:T} (pfdu:type_eq_dec U) (f1 f2:Fin_map (Add A a) B def), f1 <> f2 -> fin_map_from_add f1 <> fin_map_from_add f2 \/ f1 |-> a <> f2 |-> a. intros T U A B def a hdu f1 f2 h1. destruct (hdu (f1 |-> a) (f2 |-> a)) as [h4 | h4]. left. contradict h1. apply fin_map_ext. intros c h3. destruct h3 as [|c h3]; subst. erewrite <- fin_map_from_add_compat; auto. erewrite <- fin_map_from_add_compat; auto. rewrite h1; auto. destruct h3; auto. right; auto. Qed. Lemma distinct_fin_maps_differ_at_point : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfdu:type_eq_dec U) (f1 f2:Fin_map A B def), f1 <> f2 -> exists x:T, In A x /\ f1 |-> x <> f2 |-> x. intros T U A B def hdu f1. pose proof (fin_map_fin_dom f1) as h1. generalize dependent U. induction h1. intros U B def hdu f1 f2 h1. contradict h1. apply empty_map1_eq. intros U B def hdu f1 f2 h2. apply fin_map_from_add_neq in h2. destruct h2 as [h2|h2]. specialize (IHh1 U B def hdu (fin_map_from_add f1) (fin_map_from_add f2) h2). destruct IHh1 as [c [h4 h5]]. exists c. split. left; auto. rewrite fin_map_from_add_compat in h5; auto. rewrite fin_map_from_add_compat in h5; auto. exists x. split; auto. right; constructor. assumption. Qed. Lemma not_extends_differs_at_point : forall {T U:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (pfdu:type_eq_dec U) (F':Fin_map A' B def) (F:Fin_map A B def), Included A A' -> ~ extends F' F -> exists x:T, In A x /\ F |-> x <> F' |-> x. intros T U A A' B def hdu f1 f2 h1 h2. pose proof (not_extends_compat _ _ h1 h2) as h3. pose proof (distinct_fin_maps_differ_at_point hdu _ _ h3) as h4. destruct h4 as [x h4]. destruct h4 as [h4 h5]. exists x. split; auto. apply neq_sym. pose proof (restriction_compat f1 h1 x h4) as h6. rewrite <- h6. assumption. Qed. Lemma fin_map_to_fps_ext : forall {T U:Type} {A:Ensemble T} {B C:Ensemble U} {def:U} (F:Fin_map A B def) (G:Fin_map A C def), (forall x:T, F |-> x = G |-> x) -> fin_map_to_fps F = fin_map_to_fps G. intros T U A B C def F G h1. pose proof (fin_map_to_fps_compat F) as h2. pose proof (fin_map_to_fps_compat G) as h3. destruct h2 as [h2a h2b]. destruct h3 as [h3a h3b]. pose proof (fin_fps_to_f_inj' (fin_map_type_eq_dec F) A B C _ _ def (fin_map_fin_dom F) h2a h3a) as h4. apply h4. apply functional_extensionality. intro x. specialize (h1 x). destruct F, G. rewrite h2b. simpl. pose proof (proof_irrelevance _ (fin_map_fin_dom (fin_map_intro A B def t f f0 S f1)) f) as h5. rewrite h5. pose proof (type_eq_dec_eq (fin_map_type_eq_dec (fin_map_intro A B def t f f0 S f1)) t) as h9. rewrite h9. simpl in h1. rewrite h1. pose proof (fin_map_to_fps_s_compat t A C def f2 f3 S0 f4) as h10. apply fin_fps_to_f_functional. apply fin_map_to_fps_s_compat. Qed. Lemma fin_map_app_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:(Fin_map A B def)) (x:T), F |-> x = fin_fps_to_f (fin_map_type_eq_dec F) (fin_map_fin_dom F) _ (fp_fin_map_to_fps F) def x. intros T U A B def F x. unfold fin_map_app. pose proof (fin_map_type_eq_dec F) as hte. pose proof (fin_map_fin_dom F) as hf. destruct F as [hte' h1 h2 S h3]. apply fin_fps_to_f_functional. apply fin_map_to_fps_s_compat. Qed. Lemma in_fin_map_to_fps_iff : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) (pr:T*U), In A (fst pr) -> (In (fin_map_to_fps F) pr <-> snd pr = F |-> fst pr). intros T U A B def F pr h0. split. intro h1. rewrite fin_map_app_compat. pose proof (fin_map_to_fps_compat F) as h2. destruct h2 as [h2 h4]. inversion h2 as [h5 h6]. specialize (h6 _ h1). destruct h6 as [h6 h7]. pose proof (in_fin_map_to_fps F _ h6) as h8. eapply fp_functional. apply h2. rewrite surjective_pairing in h1. apply h1. rewrite <- h4 in h8. pose proof (proof_irrelevance _ h2 (fp_fin_map_to_fps F)). subst. assumption. intro h1. pose proof (in_fin_map_to_fps F _ h0) as h2. rewrite <- h1 in h2. rewrite surjective_pairing. assumption. Qed. Lemma fin_map_to_fps_ext_in : forall {T U:Type} {A:Ensemble T} {B C:Ensemble U} {def:U} (F:Fin_map A B def) (G:Fin_map A C def), (forall x:T, In A x -> F |-> x = G |-> x) -> fin_map_to_fps F = fin_map_to_fps G. intros T U A B C def F G h1. pose proof (fin_map_to_fps_compat F) as h2. pose proof (fin_map_to_fps_compat G) as h3. destruct h2 as [h2a h2b]. destruct h3 as [h3a h3b]. apply Extensionality_Ensembles; red; split; red; intros pr h4. pose proof (fin_fps_to_f_s_compat (fin_map_type_eq_dec F) (fin_map_fin_dom F) h2a def (fst pr)) as h5. inversion h2a as [h7 h8]. specialize (h8 _ h4). destruct h8 as [h8 h9]. specialize (h5 h8). rewrite in_fin_map_to_fps_iff; auto. rewrite <- h1; auto. symmetry. eapply fp_functional. apply h2a. pose proof (fin_map_to_fps_compat F) as h10. destruct h10 as [h10 h11]. rewrite <- h11. pose proof (proof_irrelevance _ h10 h2a); subst. apply h5. rewrite surjective_pairing in h4. assumption. pose proof (fin_fps_to_f_s_compat (fin_map_type_eq_dec G) (fin_map_fin_dom G) h3a def (fst pr)) as h5. inversion h3a as [h7 h8]. specialize (h8 _ h4). destruct h8 as [h8 h9]. specialize (h5 h8). rewrite in_fin_map_to_fps_iff; auto. rewrite h1; auto. symmetry. eapply fp_functional. apply h3a. pose proof (fin_map_to_fps_compat G) as h10. destruct h10 as [h10 h11]. rewrite <- h11. pose proof (proof_irrelevance _ h10 h3a); subst. apply h5. rewrite surjective_pairing in h4. assumption. Qed. Lemma fun_to_fin_map_sing_im : forall {T U:Type} (pfte: type_eq_dec T) (A:Ensemble T) (def val:U) (pfa:Finite A), fin_map_eq (fun_to_fin_map pfte A def pfa (fun x:T=>val)) (fin_map_sing pfte A pfa def val). intros T U hte A def val h1. red. ex_goal. red. intros u h2. destruct h2 as [x h2 u]. subst. constructor. exists hex. apply fin_map_to_fps_inj; auto. apply fin_map_to_fps_ext_in; auto. intros x h3. rewrite <- fin_map_new_ran_compat; auto. unfold fin_map_sing. simpl. pose proof (fun_to_fin_map_compat hte A def h1 (fun _ : T => val) x h3) as h4. rewrite h4. pose proof (fin_fps_to_f_compat hte h1 (cart_prod A (Singleton val)) (fp_cart_prod_sing A val) def x) as h5. destruct h5 as [h5 h6]. specialize (h5 h3). inversion h5 as [h7]. simpl in h7. destruct h7 as [h7 h8]. destruct h8. reflexivity. Qed. Lemma fin_map_sing_const : forall {T U:Type} (pfte:type_eq_dec T) (A:Ensemble T) (def val:U) (pfa:Finite A) (x:T), In A x -> (fin_map_sing pfte A pfa def val) |-> x = val. intros T U hte A def val h1 x h0. unfold fin_map_app. destruct (fin_map_sing hte A h1 def val) as [hte' h2 h3 S h4]. pose proof (fin_fps_to_f_compat hte h1 S h4 def x) as h5. destruct h5 as [h5 h6]. specialize (h5 h0). inversion h4 as [h7 h8]. specialize (h8 _ h5). simpl in h8. destruct h8 as [h8a h8b]. pose proof (type_eq_dec_eq hte hte'). pose proof (proof_irrelevance _ h1 h2). subst. inversion h8b; auto. Qed. Lemma im_fin_map_sing : forall {T U:Type} (pfte:type_eq_dec T) (A:Ensemble T) (def val:U) (pf:Finite A), A <> Empty_set _ -> im_fin_map (fin_map_sing pfte A pf def val) = Singleton val. intros T U hte A def val h1 h0. unfold im_fin_map. apply Extensionality_Ensembles. red. split. (* <= *) red. intros x h2. destruct h2 as [x h2 y h3]. rewrite h3. rewrite fin_map_sing_const. constructor. assumption. (* >= *) red. intros x h2. destruct h2; subst. pose proof h1 as h2. induction h2 as [|A h2 h3 x h4]. contradict h0. reflexivity. apply Im_intro with x. right. constructor. rewrite fin_map_sing_const. reflexivity. right. constructor. Qed. Lemma restriction_trivial_eq : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) (pf:Included A A), restriction F pf = F. intros T U A B def F h1. pose proof (restriction_compat F h1) as h2. apply fin_map_ext. assumption. Qed. Lemma fun_to_fin_map_ran_ex : forall {T U:Type} (pfdt: type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (def:U) (f:T->U), Finite A -> Finite B -> Included (Im A f) B -> exists! F:Fin_map A B def, forall (x:T) (pf:In A x), F |-> x = f x. intros T U hdt A B def f h1 h2 h3. pose (fun_to_fin_map hdt A def h1 f) as f'. exists (fin_map_new_ran f' h2 h3). red. split. intros x h4. pose proof (fin_map_new_ran_compat f' h2 h3 x) as h5. simpl in h5. specialize (h5 h4). rewrite <- h5; auto. unfold f'. rewrite fun_to_fin_map_compat; auto. intros f'' h4. apply fin_map_ext. intros x h5. specialize (h4 _ h5). rewrite h4. rewrite <- (fin_map_new_ran_compat f' h2 h3 x). unfold f'. rewrite fun_to_fin_map_compat; auto. assumption. Qed. Definition fun_to_fin_map_ran {T U:Type} (pfte:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (def:U) (f:T->U) (pfa:Finite A) (pfb:Finite B) (pfincl:Included (Im A f) B) : Fin_map A B def := proj1_sig (constructive_definite_description _ (fun_to_fin_map_ran_ex pfte A B def f pfa pfb pfincl)). Lemma fun_to_fin_map_ran_compat : forall {T U:Type} (pfte:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (def:U) (f:T->U) (pfa:Finite A) (pfb:Finite B) (pfincl:Included (Im A f) B), let f' := fun_to_fin_map_ran pfte A B def f pfa pfb pfincl in forall x:T, In A x -> f' |-> x = f x. unfold fun_to_fin_map_ran; intros; destruct constructive_definite_description; auto. Qed. Lemma sig_fun_to_fin_map_ex : forall {T U:Type} {A:Ensemble T} (pfte:type_eq_dec T) (f:sig_set A->U), Finite A -> forall def:U, exists! F:Fin_map A (Im (full_sig A) f) def, forall (x:T) (pf:In A x), F |-> x = f (exist _ _ pf). intros T U A hte f h0 def. pose [pr:T*U | exists pf:(In A (fst pr)), f (exist _ _ pf) = snd pr] as S. assert (h1:functionally_paired A (Im (full_sig A) f) S). constructor. intros x h1. exists (f (exist _ _ h1)). red. split. split. apply Im_intro with (exist _ _ h1). constructor. reflexivity. constructor. simpl. exists h1. reflexivity. intros x' h2. destruct h2 as [h2a h2b]. inversion h2b as [h3]. simpl in h3. destruct h3 as [h3 h4]. pose proof (proof_irrelevance _ h3 h1); subst; auto. intros pr h1. destruct h1 as [h1]. destruct h1 as [h1 h2]. split; auto. rewrite <- h2. apply Im_intro with (exist _ _ h1). constructor. reflexivity. pose proof h0 as h0'. rewrite finite_full_sig_iff in h0'. pose proof (finite_image _ _ _ f h0') as h2. exists (fin_map_intro _ _ _ hte h0 h2 S h1). red. split. intros x h3. simpl. pose proof (fin_fps_to_f_s_compat hte h0 h1 def _ h3) as h4. simpl in h4. inversion h4 as [h5]. destruct h5 as [h5 h6]. simpl in h6, h5. pose proof (proof_irrelevance _ h3 h5); subst. rewrite h6; auto. intros F h3. apply fin_map_ext. intros x h4. specialize (h3 _ h4). rewrite h3. simpl. pose proof (fin_fps_to_f_s_compat hte h0 h1 def _ h4) as h5. inversion h5 as [h6]. clear h5. simpl in h6. destruct h6 as [h6 h7]. pose proof (proof_irrelevance _ h4 h6); subst. rewrite h7. reflexivity. Qed. Definition sig_fun_to_fin_map {T U:Type} {A:Ensemble T} (pfte:type_eq_dec T) (f:sig_set A->U) (pf:Finite A) (def:U) := proj1_sig (constructive_definite_description _ (sig_fun_to_fin_map_ex pfte f pf def)). Lemma sig_fun_to_fin_map_compat : forall {T U:Type} {A:Ensemble T} (pfte:type_eq_dec T) (f:sig_set A->U) (pf:Finite A) (def:U), let F:= sig_fun_to_fin_map pfte f pf def in forall (x:T) (pf:In A x), F |-> x = f (exist _ _ pf). intros T U A hte f h1 def F. unfold F, sig_fun_to_fin_map. destruct constructive_definite_description as [F' h2]. simpl. assumption. Qed. Lemma sig_fun_to_fin_map_ran_ex : forall {T U:Type} {A:Ensemble T} (pfte:type_eq_dec T) (f:sig_set A->U), Finite A -> forall (def:U) (B:Ensemble U), Finite B -> Included (Im (full_sig A) f) B -> exists! F:Fin_map A B def, forall (x:T) (pf:In A x), F |-> x = f (exist _ _ pf). intros T U A hte f h1 def B h0 h2. pose (sig_fun_to_fin_map hte f h1 def) as F. pose (fin_map_new_ran F h0 h2) as F'. pose proof (fin_map_new_ran_compat F h0 h2) as h4. simpl in h4. pose proof (sig_fun_to_fin_map_compat hte f h1 def) as h5. exists F'. red. split. intros x h3. unfold F'. specialize (h4 x). rewrite <- h4. specialize (h5 _ h3). unfold F. assumption. assumption. intros G h3. apply fin_map_ext. intros x h6. specialize (h3 _ h6). rewrite h3. unfold F'. specialize (h5 _ h6). specialize (h4 x). rewrite <- h4. unfold F. assumption. assumption. Qed. Definition sig_fun_to_fin_map_ran {T U:Type} {A:Ensemble T} (pfte:type_eq_dec T) (f:sig_set A->U) (pfa:Finite A) (def:U) (B:Ensemble U) (pfb:Finite B) (pfi:Included (Im (full_sig A) f) B) := proj1_sig (constructive_definite_description _ (sig_fun_to_fin_map_ran_ex pfte f pfa def B pfb pfi)). Lemma sig_fun_to_fin_map_ran_compat : forall {T U:Type} {A:Ensemble T} (pfte:type_eq_dec T) (f:sig_set A->U) (pfa:Finite A) (def:U) (B:Ensemble U) (pfb:Finite B) (pfi:Included (Im (full_sig A) f) B), let F := sig_fun_to_fin_map_ran pfte f pfa def B pfb pfi in forall (x:T) (pf:In A x), F |-> x = f (exist _ _ pf). intros T U A hte f h1 def B h2 h3 F. unfold F, sig_fun_to_fin_map_ran. destruct constructive_definite_description as [F' h4]. simpl. assumption. Qed. Lemma fin_map_seg_to_full_ex : forall {U:Type} {m:nat} {def:U} {B:Ensemble U} (F:Fin_map (seg m) B def), exists! F':Fin_map (Full_set (m @)) B def, forall (i:m @), F' |-> i = F |-> (proj1_sig i). intros U m def B F. pose [pr: (m @) * U | snd pr = F |-> proj1_sig (fst pr)] as S. assert (hfp: functionally_paired (Full_set (m @)) B S). unfold S. constructor. intros x h2. exists (F |-> (proj1_sig x)). red; split. split. apply fin_map_app_in. constructor. apply proj2_sig. constructor; simpl; auto. intros y h3. destruct h3 as [h3 h4]. destruct h4 as [h4]. simpl in h4. subst; auto. intros y h1; constructor. constructor. destruct h1 as [h1]. rewrite h1. apply fin_map_app_in. constructor. apply proj2_sig. exists (fin_map_intro (Full_set (m @)) B def (segT_eq_dec m) (finite_full_set_segT _) (fin_map_fin_ran F) S hfp). red; split. intro i. assert (h1:Inn S (i, F |-> (proj1_sig i))). unfold S. constructor; simpl; auto. assert (h2:i = fst (i, F |-> (proj1_sig i))); auto. rewrite h2 at 1. rewrite fin_map_intro_compat. simpl; auto. assumption. intros F' h0. apply fin_map_ext. intros i h2. specialize (h0 i). rewrite h0. assert (h1:Inn S (i, F |-> (proj1_sig i))). unfold S. constructor; simpl; auto. assert (h3:i = fst (i, F |-> (proj1_sig i))); auto. rewrite h3 at 1. rewrite fin_map_intro_compat. simpl; auto. assumption. Qed. Definition fin_map_seg_to_full {U:Type} {m:nat} {def:U} {B:Ensemble U} (F:Fin_map (seg m) B def) : Fin_map (Full_set (m @)) B def := proj1_sig (constructive_definite_description _ (fin_map_seg_to_full_ex F)). Lemma fin_map_seg_to_full_compat : forall {U:Type} {m:nat} {def:U} {B:Ensemble U} (F:Fin_map (seg m) B def) (i:m @), let F' := fin_map_seg_to_full F in F' |-> i = F |-> (proj1_sig i). unfold fin_map_seg_to_full; intros; destruct constructive_definite_description; auto. Qed. Lemma fin_map_to_sig_fun_ex : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), exists! g:sig_set A->U, forall x:sig_set A, g x = F |-> (proj1_sig x). intros T U A B def F. exists (fun x:sig_set A => F |-> (proj1_sig x)). red. split; auto. intros F' h1. apply functional_extensionality. intro x'. rewrite h1. reflexivity. Qed. Definition fin_map_to_sig_fun {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) := proj1_sig (constructive_definite_description _ (fin_map_to_sig_fun_ex F)). Lemma fin_map_to_sig_fun_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), let g := fin_map_to_sig_fun F in forall x:sig_set A, g x = F |-> (proj1_sig x). intros T U A B def F g x. unfold g, fin_map_to_sig_fun. destruct constructive_definite_description as [g' h1]. simpl. apply h1. Qed. Lemma im_fin_map_to_sig_fun : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), Im (full_sig A) (fin_map_to_sig_fun F) = im_fin_map F. intros T U A B def F. pose proof (fin_map_to_sig_fun_compat F) as h1. simpl in h1. unfold im_fin_map. apply Extensionality_Ensembles. red. split. red. intros x h2. destruct h2 as [x h2]. subst. clear h2. apply Im_intro with (proj1_sig x). apply proj2_sig. apply h1. red. intros x h2. destruct h2 as [x h2]. subst. apply Im_intro with (exist _ _ h2). constructor. symmetry. apply h1. Qed. Lemma restriction_sig_fin_map_to_sig_fun : forall {T U:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) (pf:Included A' A), restriction_sig (fin_map_to_sig_fun F) _ pf = fin_map_to_sig_fun (restriction F pf). intros T U A A' B def F h1. apply functional_extensionality. intro x. pose proof (fin_map_to_sig_fun_compat F) as h2. simpl in h2. pose proof (fin_map_to_sig_fun_compat (restriction F h1)) as h2'. simpl in h2, h2'. unfold restriction_sig. rewrite h2, h2'. simpl. rewrite restriction_compat; auto. apply proj2_sig. Qed. Lemma fun_in_ens_to_fin_map_ex : forall {T U:Type} {A:Ensemble T} (pfte:type_eq_dec T) (f:fun_in_ens A U) (pf:Finite A) (def:U), exists! F:Fin_map A (im_in_ens f) def, forall (x:T) (pf:In A x), F |-> x = f x pf. intros T U A hte f h2 def. pose (sig_fun_to_fin_map hte (fun x:sig_set A => f (proj1_sig x) (proj2_sig x)) h2 def) as F. pose proof (im_in_ens_eq f) as h3. symmetry in h3. pose (fin_map_new_ran F (finite_im_in_ens f h2)) as F'. hfirst F'. rewrite h3; auto with sets. exists (F' hf). red; split. intros x h4. unfold F'. rewrite <- fin_map_new_ran_compat. unfold F. erewrite sig_fun_to_fin_map_compat. simpl. f_equal. assumption. intros G h4. apply fin_map_ext. intros. erewrite h4. unfold F'. rewrite <- fin_map_new_ran_compat. unfold F. erewrite sig_fun_to_fin_map_compat. f_equal. assumption. Grab Existential Variables. assumption. Qed. Definition fun_in_ens_to_fin_map {T U:Type} {A:Ensemble T} (pfte:type_eq_dec T) (f:fun_in_ens A U) (pf:Finite A) (def:U) := proj1_sig (constructive_definite_description _ (fun_in_ens_to_fin_map_ex pfte f pf def)). Lemma fun_in_ens_to_fin_map_compat : forall {T U:Type} {A:Ensemble T} (pfte:type_eq_dec T) (f:fun_in_ens A U) (pf:Finite A) (def:U), let F := fun_in_ens_to_fin_map pfte f pf def in forall (x:T) (pf:In A x), F |-> x = f x pf. unfold fun_in_ens_to_fin_map; intros; destruct constructive_definite_description; auto. Qed. Lemma fun_in_ens_to_fin_map_ran_ex : forall {T U:Type} (pfte:type_eq_dec T) {A:Ensemble T} (f:fun_in_ens A U) (pf:Finite A) (B:Ensemble U) (pfb:Finite B) (def:U), Included (im_in_ens f) B -> exists! F:Fin_map A B def, forall (x:T) (pf:In A x), F |-> x = f x pf. intros T U hte A f h1 B h2 def h3. pose (sig_fun_to_fin_map_ran hte (fun x:sig_set A => f (proj1_sig x) (proj2_sig x)) h1 def B h2) as F. hfirst F. red; intros x h4. apply h3. destruct h4 as [x h4]; subst. clear h4. destruct x as [x h4]. simpl. econstructor. reflexivity. pose (F hf) as F'. exists F'. red; split. intros x h4. unfold F', F. erewrite sig_fun_to_fin_map_ran_compat; simpl. reflexivity. intros G h4. apply fin_map_ext. intros x h5. specialize (h4 _ h5). rewrite h4. unfold F', F. erewrite sig_fun_to_fin_map_ran_compat; simpl. reflexivity. Qed. Definition fun_in_ens_to_fin_map_ran {T U:Type} {A:Ensemble T} (pfte:type_eq_dec T) (f:fun_in_ens A U) (pf:Finite A) (B:Ensemble U) (pfb:Finite B) (def:U) (pfincl:Included (im_in_ens f) B) : Fin_map A B def := proj1_sig (constructive_definite_description _ (fun_in_ens_to_fin_map_ran_ex pfte f pf B pfb def pfincl)). Lemma fun_in_ens_to_fin_map_ran_compat : forall {T U:Type} {A:Ensemble T} (pfte:type_eq_dec T) (f:fun_in_ens A U) (pf:Finite A) (B:Ensemble U) (pfb:Finite B) (def:U) (pfincl:Included (im_in_ens f) B), let F := fun_in_ens_to_fin_map_ran pfte f pf B pfb def pfincl in forall (x:T) (pf:In A x), F |-> x = f x pf. unfold fun_in_ens_to_fin_map_ran; intros; destruct constructive_definite_description; auto. Qed. Lemma fin_map_inj_iff : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), injective (fin_map_to_sig_fun F) <-> inj F. intros T U A B def F. pose proof (fin_map_to_sig_fun_compat F) as h1. simpl in h1. split. intro h2. red in h2. red. intros x y h3 h4 h5. specialize (h2 (exist _ _ h3) (exist _ _ h4)). do 2 rewrite h1 in h2. simpl in h2. specialize (h2 h5). apply exist_injective in h2. assumption. intro h2. red in h2. red. intros x y h3. destruct x as [x h4], y as [y h5]. do 2 rewrite h1 in h3. simpl in h3. specialize (h2 x y h4 h5 h3). apply proj1_sig_injective. simpl. assumption. Qed. Lemma fin_map_compose_inj_ex : forall {T U V:Type} {A:Ensemble T} {B:Ensemble V} {defv:V} (pfue : type_eq_dec U) (F:(Fin_map A B defv)) (deft:T) (g:T->U), injective g -> exists! F':Fin_map (Im A g) B defv, forall x:T, In A x -> F |-> x = F' |-> (g x). intros T U V A B defv hue F deft g h1. pose proof (incl_im_im_left_inverse_compose _ h1 A deft (fin_map_app F)) as h2. pose proof (im_fin_map_ran F) as h3. assert (h4:Included (Im (Im A g) (fun u : U => F |-> ((left_inverse g h1 A, deft) ||-> u))) B). auto with sets. exists (fin_map_new_ran (fun_to_fin_map hue (Im A g) defv (finite_image _ _ _ g (fin_map_fin_dom F)) (fun u => F |-> (((left_inverse _ h1 A), deft) ||-> u))) (fin_map_fin_ran F) h4). red. split. intros x h5. rewrite <- fin_map_new_ran_compat. rewrite fun_to_fin_map_compat. simpl. unfold sig_fun_app. destruct classic_dec as [h6 | h7]. pose proof (proof_irrelevance _ h6 (Im_intro T U A g x h5 (g x) eq_refl)); subst. rewrite left_inverse_compat. reflexivity. contradict h7. apply Im_intro with x; auto. apply Im_intro with x; auto. apply Im_intro with x; auto. intros F' h5. apply fin_map_ext. intros u h6. destruct h6 as [x h6]. subst. rewrite <- fin_map_new_ran_compat. rewrite fun_to_fin_map_compat. rewrite <- h5; auto. f_equal. simpl. unfold sig_fun_app. destruct classic_dec as [h7 | h8]. pose proof (proof_irrelevance _ h7 (Im_intro T U A g x h6 (g x) eq_refl)); subst. rewrite left_inverse_compat; auto. contradict h8. apply Im_intro with x; auto. apply Im_intro with x; auto. apply Im_intro with x; auto. Qed. Definition fin_map_compose_inj {T U V:Type} {A:Ensemble T} {B:Ensemble V} {defv:V} (pfue:type_eq_dec U) (F:(Fin_map A B defv)) (deft:T) (g:T->U) (pf:injective g) : Fin_map (Im A g) B defv := proj1_sig (constructive_definite_description _ (fin_map_compose_inj_ex pfue F deft g pf)). Lemma fin_map_compose_inj_compat : forall {T U V:Type} {A:Ensemble T} {B:Ensemble V} {defv:V} (pfue:type_eq_dec U) (F:(Fin_map A B defv)) (deft:T) (g:T->U) (pf:injective g), let F' := fin_map_compose_inj pfue F deft g pf in forall x:T, In A x -> F |-> x = F' |-> (g x). intros T U V A B defv hue F deft g pf F'. unfold F'. unfold fin_map_compose_inj. destruct constructive_definite_description as [K h2]. simpl. assumption. Qed. Lemma fin_map_left_inverse_ex : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfue:type_eq_dec U) (F:Fin_map A B def), inj F -> forall deft:T, exists! F':Fin_map (im_fin_map F) A deft, forall x:T, In A x -> F' |-> (F |-> x) = x. intros T U A B def hue F h1 deft. rewrite <- fin_map_inj_iff in h1. pose (left_inverse_full (fin_map_to_sig_fun F) h1) as f1. pose (fun x:sig_set (Im (Full_set (sig_set A)) (fin_map_to_sig_fun F)) => proj1_sig (f1 x)) as f2. pose proof (im_fin_map_to_sig_fun F) as h2. assert (h3:forall x:sig_set (im_fin_map F), In (Im (full_sig A) (fin_map_to_sig_fun F)) (proj1_sig x)). intro x. rewrite h2. apply proj2_sig. pose (fun x:sig_set (im_fin_map F) => f2 (exist _ _ (h3 x))) as f3. assert (h4: Included (Im (full_sig (im_fin_map F)) f3) A). red. intros x h4. destruct h4 as [x h4 ? h5]. subst. unfold f3, f2, f1. apply proj2_sig. exists (sig_fun_to_fin_map_ran hue f3 (im_fin_map_finite F) deft _ (fin_map_fin_dom F) h4). red. split. intros x h5. assert (h6:Inn (im_fin_map F) (F |-> x)). apply Im_intro with x; auto. rewrite (sig_fun_to_fin_map_ran_compat hue f3 (im_fin_map_finite F) deft A (fin_map_fin_dom F) h4 _ h6). unfold f3, f2, f1, left_inverse_full. simpl. pose proof (left_inverse_compat (fin_map_to_sig_fun F) h1 (Full_set (sig_set A)) (exist _ _ h5)) (Full_intro _ (exist _ _ h5)) as h7. pose proof (f_equal (@proj1_sig _ _) h7) as h8. simpl in h8. rewrite <- h8. f_equal. f_equal. apply proj1_sig_injective. simpl. rewrite fin_map_to_sig_fun_compat. simpl. reflexivity. intros F' h5. apply fin_map_ext. intros x h6. pose proof (sig_fun_to_fin_map_ran_compat hue f3 (im_fin_map_finite F) deft A (fin_map_fin_dom F) h4 _ h6) as h7. simpl in h7. rewrite h7. destruct h6 as [x h6]. subst. rewrite h5; auto. unfold f3, f2, f1, left_inverse_full. pose proof (left_inverse_compat (fin_map_to_sig_fun F) h1 (Full_set (sig_set A)) (exist _ _ h6)) (Full_intro _ (exist _ _ h6)) as h8. pose proof (f_equal (@proj1_sig _ _) h8) as h9. simpl in h9. rewrite <- h9. f_equal. f_equal. apply proj1_sig_injective; simpl. rewrite fin_map_to_sig_fun_compat. simpl. reflexivity. Qed. Definition fin_map_left_inverse {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfue:type_eq_dec U) (F:Fin_map A B def) (pf:inj F) (deft:T) := proj1_sig (constructive_definite_description _ (fin_map_left_inverse_ex pfue F pf deft)). Lemma fin_map_left_inverse_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfue:type_eq_dec U) (F:Fin_map A B def) (pf:inj F) (deft:T), let F':=fin_map_left_inverse pfue F pf deft in forall x:T, In A x -> F' |-> (F |-> x) = x. intros T U A B def hue F h1 deft F'. unfold F', fin_map_left_inverse. intros x h2. destruct constructive_definite_description as [F'' h3]. simpl. apply h3; auto. Qed. Lemma fin_map_left_inverse_inj : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfue:type_eq_dec U) (F:Fin_map A B def) (pf:inj F) (deft:T), inj (fin_map_left_inverse pfue F pf deft). intros T U A B def hue F h1 deft. red in h1. red. intros x y h2 h3 h4. destruct h2 as [x h2], h3 as [y h3]. subst. pose proof (fin_map_left_inverse_compat hue F h1 deft) as h5. simpl in h5. rewrite h5 in h4; auto. rewrite h5 in h4; auto. subst. reflexivity. Qed. Lemma fin_map_to_fps_dom_rel : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), dom_rel (fin_map_to_fps F) = A. intros T U A B def F. rewrite dom_rel_eq. destruct F as [hte h1 h2 S h3]. apply Extensionality_Ensembles. red. split. red. intros x h4. destruct h4 as [pr h4 x h5]. subst. rewrite <- fin_map_to_fps_s_compat in h4. destruct h3 as [h3l h3r]. specialize (h3r _ h4). destruct h3r; auto. red. intros x h4. destruct h3 as [h3l h3r]. pose proof h3l as h3l'. specialize (h3l' _ h4). destruct h3l' as [y h3l']. apply Im_intro with (x, y). rewrite <- fin_map_to_fps_s_compat. red in h3l'. destruct h3l' as [h3a h3b]. destruct h3a; auto. auto. Qed. Lemma fin_map_to_fps_ran_rel : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), ran_rel (fin_map_to_fps F) = Im A (fin_map_to_fun F). intros T U A B def F. unfold ran_rel. apply Extensionality_Ensembles. red. split. red. intros y h1. destruct h1 as [h1]. destruct h1 as [x h1]. apply Im_intro with x. pose proof (fp_fin_map_to_fps F) as h2. destruct h2 as [h2l h2r]. specialize (h2r _ h1). simpl in h2r. destruct h2r; auto. unfold fin_map_to_fun. destruct F as [hte h2 h3 S h4]. rewrite <- fin_map_to_fps_s_compat in h1. pose proof (fin_fps_to_f_compat' hte h2 h4 def) as h5. rewrite h5 in h1. destruct h1 as [h1]. simpl in h1. destruct h1; auto. red. intros y h1. constructor. destruct h1 as [x h1 y h2]. subst. exists x. unfold fin_map_to_fun. destruct F as [hte h2 h3 S h4]. pose proof (fin_fps_to_f_compat' hte h2 h4 def) as h5. rewrite <- fin_map_to_fps_s_compat. rewrite h5 at 1. constructor. simpl. split; auto. Qed. Lemma fp_fin_map_to_fps' : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), functionally_paired (dom_rel (fin_map_to_fps F)) (ran_rel (fin_map_to_fps F)) (fin_map_to_fps F). intros T U A B def F. destruct F as [hte h1 h2 S h3]. rewrite fin_map_to_fps_dom_rel. rewrite fin_map_to_fps_ran_rel. rewrite <- fin_map_to_fps_s_compat. constructor. intros x h4. inversion h3 as [h3l h3r]. pose proof (h3l _ h4) as h3l'. destruct h3l' as [y h3l']. red in h3l'. exists y. red. split. split. apply Im_intro with x. assumption. simpl. destruct h3l' as [h3a h3b]. destruct h3a as [h5 h6]. rewrite (fin_fps_to_f_compat' hte h1 h3 def) in h6. destruct h6 as [h6]. simpl in h6. destruct h6; auto. destruct h3l' as [[]]; auto. intros y' h5. destruct h5 as [h5l h5r]. specialize (h3r _ h5r). simpl in h3r. destruct h3r as [h6 h7]. destruct h3l' as [h3a h3b]. specialize (h3b _ (conj h7 h5r)). assumption. intros pr h4. inversion h3 as [h5 h6]. split. specialize (h6 _ h4). destruct h6; auto. apply Im_intro with (fst pr). specialize (h6 _ h4). destruct h6; auto. simpl. rewrite (fin_fps_to_f_compat' hte h1 h3 def) in h4. destruct h4 as [h4]. destruct h4; auto. Qed. Lemma fin_map_to_fps_fin_map_app_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) (pr:T*U), In (fin_map_to_fps F) pr -> F |-> (fst pr) = snd pr. intros T U A B ef F pr h1. pose proof (fin_map_to_fps_compat F) as h2. destruct h2 as [h2 h3]. unfold fin_map_app. destruct F. rewrite <- fin_map_to_fps_s_compat in h1. apply fp_inv. assumption. Qed. Lemma fin_map_app_incl_fps : forall {T U:Type} {A A':Ensemble T} {B B':Ensemble U} {def:U} (F:Fin_map A B def) (F':Fin_map A' B' def), Included (fin_map_to_fps F) (fin_map_to_fps F') -> forall x:T, Inn A x -> F |-> x = F' |-> x. intros T U A A' B B' def F F' h1 x h2. unfold fin_map_app. destruct F as [hte h3 h4 S h5]; destruct F' as [hte' h6 h7 S' h8]. do 2 rewrite <- fin_map_to_fps_s_compat in h1. rewrite (fin_fps_to_f_compat' hte h3 h5 def) in h1. rewrite (fin_fps_to_f_compat' hte' h6 h8 def) in h1. specialize (h1 (x, fin_fps_to_f hte h3 S h5 def x)). hfirst h1. constructor. simpl. split; auto. specialize (h1 hf). inversion h1 as [h10]. simpl in h10. destruct h10; auto. Qed. Lemma fin_map_app_fin_map_to_fps_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) (pr:T*U), In A (fst pr) -> F |-> (fst pr) = snd pr -> In (fin_map_to_fps F) pr. intros T U A B def F pr h0 h1. destruct F as [hte h2 h3 S h4]. rewrite <- fin_map_to_fps_s_compat. unfold fin_map_app in h1. rewrite (fin_fps_to_f_compat' hte h2 h4 def). constructor. symmetry in h1. split; auto. Qed. Lemma fin_map_to_fps_finite : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), Finite (fin_map_to_fps F). intros T U A B def F. destruct F as [hte h1 h2 S h3]. rewrite <- fin_map_to_fps_s_compat. pose proof (fps_inc_cart_prod h3) as h4. pose proof (cart_prod_fin _ _ h1 h2) as h5. apply Finite_downward_closed with (cart_prod A B); assumption. Qed. Lemma fin_fps_to_f_undoes_fin_map_to_fps : forall {T U:Type} {def:U} {A:Ensemble T} {B:Ensemble U} (F:Fin_map A B def), let S := fin_map_to_fps F in forall (pf:functionally_paired A B S) (x:T), fin_fps_to_f (fin_map_type_eq_dec F) (fin_map_fin_dom F) _ pf def x = F |-> x. intros T U def A B F S h0 x. pose proof (fin_map_to_fun_compat F) as h1. rewrite h1. unfold fin_map_to_fun. destruct F as [hte h2 h3 S' h4]. pose proof (fin_map_to_fps_s_compat hte A B def h2 h3 _ h4) as h5. unfold S. symmetry in h5. pose proof (subsetT_eq_compat _ _ _ _ h0 h4 h5) as h6. dependent rewrite -> h6. apply fin_fps_to_f_functional; auto. Qed. Lemma fin_map_to_fps_new_ran_compat : forall {T U:Type} (def:U) (A:Ensemble T) (B C:Ensemble U) (F:Fin_map A B def) (pfc:Finite C) (pfinc:Included B C), fin_map_to_fps (fin_map_new_ran F pfc pfinc) = fin_map_to_fps F. intros T U def A B C F h1 h2. pose proof (fin_map_to_fps_compat F) as h3. destruct h3 as [h3a h3b]. pose proof (fin_map_to_fps_compat (fin_map_new_ran F h1 h2)) as h4. destruct h4 as [h4a h4b]. symmetry. apply fin_map_to_fps_ext_in. intros x h5. apply fin_map_new_ran_compat; auto. Qed. Lemma fin_map_sing_ran_ex : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (pfa:Finite A) (pfb:Finite B) (def val:U) (pfin:Inn B val), exists! F:Fin_map A B def, forall x:T, Inn A x -> F |-> x = val. intros T U hte A B h1 h2 def val h3. pose (fin_map_sing hte _ h1 def val) as F. pose (fin_map_new_ran F h2) as F_. hfirst F_. red; intros x h4; destruct h4; subst. assumption. pose (F_ hf) as F'. exists F'. red. split. intros x h5. unfold F', F_. rewrite <- (fin_map_new_ran_compat F h2 hf). unfold F. rewrite fin_map_sing_const. reflexivity. assumption. assumption. intros G h5. apply fin_map_to_fps_inj. unfold F', F_. rewrite fin_map_to_fps_new_ran_compat. unfold F. unfold fin_map_sing. rewrite <- fin_map_to_fps_s_compat. apply Extensionality_Ensembles. red. split. red. intros pr h6. destruct h6 as [h6]. destruct h6 as [h6l h6r]. rewrite surjective_pairing. inversion h6r as [h9]. rewrite <- h9. rewrite <- (h5 _ h6l). apply in_fin_map_to_fps. assumption. red. intros pr h6. destruct G as [hte' h7 h8 S h9]. rewrite <- fin_map_to_fps_s_compat in h6. pose proof (fp_in_dom _ _ _ h9 _ h6) as h10. rewrite surjective_pairing. constructor. simpl; split. assumption. apply h5 in h10. simpl in h10. rewrite <- h10. rewrite fp_inv; auto. constructor. Qed. Definition fin_map_sing_ran {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (pfa:Finite A) (pfb:Finite B) (def val:U) (pfin:Inn B val) : Fin_map A B def := proj1_sig (constructive_definite_description _ (fin_map_sing_ran_ex pfdt A B pfa pfb def val pfin)). Lemma fin_map_sing_ran_compat : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (pfa:Finite A) (pfb:Finite B) (def val:U) (pfin:Inn B val), let F := fin_map_sing_ran pfdt A B pfa pfb def val pfin in forall x:T, Inn A x -> F |-> x = val. intros T U hdt A B h1 h2 def val h3 F x h4. unfold F. unfold fin_map_sing_ran. destruct constructive_definite_description as [F' h5]. simpl. apply h5. assumption. Qed. Lemma fin_map_im_full_sig_eq_ex : forall {T U V:Type} {A:Ensemble T} {g:sig_set A->U} {B:Ensemble V} {defv:V} (pte:type_eq_dec T) (F:Fin_map (Im (full_sig A) g) B defv), Finite A -> forall defu:U, exists! F':Fin_map A B defv, forall x:T, In A x -> F' |-> x = F |-> ((g, defu) ||-> x). intros T U V A g B defv hte F h1 defu. pose (fun x:T => (F|-> ((g, defu) ||-> x))) as f. pose (fun_to_fin_map hte A defv h1 f) as G. assert (h2:Included (Im A f) B). red. intros b h2. destruct h2 as [b h2]. subst. unfold f. simpl. unfold sig_fun_app. destruct (classic_dec (In A b)) as [h6 | h7]. apply fin_map_app_in. apply Im_intro with (exist _ _ h6). constructor. reflexivity. contradiction. pose proof (fin_map_fin_ran F) as h3. pose (fin_map_new_ran G h3 h2) as G'. exists G'. red. split. intros x h0. simpl. unfold sig_fun_app. destruct (classic_dec (In A x)) as [h4 | h5]. rewrite <- (fin_map_new_ran_compat G h3 h2). unfold G. unfold f. simpl. rewrite fun_to_fin_map_compat. unfold sig_fun_app. destruct (classic_dec (In A x)) as [h5 | h6]. pose proof (proof_irrelevance _ h4 h5); subst; auto. contradiction. assumption. assumption. contradiction. intros K h4. unfold G'. apply fin_map_ext. intros x h5. specialize (h4 _ h5). rewrite h4. rewrite <- (fin_map_new_ran_compat G h3 h2). unfold G, f. rewrite fun_to_fin_map_compat. reflexivity. assumption. assumption. Qed. Definition fin_map_im_full_sig_eq {T U V:Type} {A:Ensemble T} {g:sig_set A->U} {B:Ensemble V} {defv:V} (pfte:type_eq_dec T) (F:Fin_map (Im (full_sig A) g) B defv) (pf:Finite A) (defu:U) : Fin_map A B defv := proj1_sig (constructive_definite_description _ (fin_map_im_full_sig_eq_ex pfte F pf defu)). Lemma fin_map_im_full_sig_eq_compat : forall {T U V:Type} {A:Ensemble T} {g:sig_set A->U} {B:Ensemble V} {defv:V} (pfte:type_eq_dec T) (F:Fin_map (Im (full_sig A) g) B defv) (pf:Finite A) (defu:U), let F' := fin_map_im_full_sig_eq pfte F pf defu in forall x:T, In A x -> F' |-> x = F |-> ((g, defu) ||-> x). intros T U V A g B defv hte F h1 defu F'. unfold F'. unfold fin_map_im_full_sig_eq. destruct constructive_definite_description as [G h3]. simpl. assumption. Qed. Lemma fin_map_to_fin_map_im_full_sig_comp : forall {T U V:Type} {A:Ensemble T} {B:Ensemble V} {defv:V} (pfue:type_eq_dec U) (F:Fin_map A B defv) (g:sig_set A->U) (defu:U), injective g -> exists (F':(Fin_map (Im (full_sig A) g) B defv)), forall x:T, In A x -> F' |-> ((g, defu) ||-> x) = F |-> x. intros T U V A B defv hue F g defu h1. assert (h3:forall u:U, In (Im (full_sig A) g) u -> exists! x, u = g x). intros u h3. destruct h3 as [u h3]. subst. exists u; auto. red. split. reflexivity. intros x h4. apply h1. assumption. pose (fun x:sig_set (Im (full_sig A) g) => (proj1_sig (proj1_sig (constructive_definite_description _ (h3 _ (proj2_sig x)))))) as f. pose (fun x:sig_set (Im (full_sig A) g) => F |-> (f x)) as k. assert (h4:Included (Im (full_sig (Im (full_sig A) g)) k) B). red. intros x h4. destruct h4 as [x h4]. subst. destruct x as [x h5]. destruct h5 as [x h5]. subst. unfold k. apply fin_map_app_in. unfold f. simpl. destruct constructive_definite_description as [a h6]. simpl. apply proj2_sig. pose proof (fin_map_fin_dom F) as h5. pose proof (fin_map_fin_ran F) as h6. rewrite finite_full_sig_iff in h5. pose proof (finite_image _ _ _ g h5) as h7. pose (sig_fun_to_fin_map_ran hue k h7 defv B h6 h4) as F'. exists F'. intros x h8. unfold F'. pose proof (sig_fun_to_fin_map_ran_compat hue k h7 defv B h6 h4) as h9. simpl in h9. assert (h10:In (Im (full_sig A) g) ((g, defu) ||-> x)). apply Im_intro with (exist _ _ h8). constructor. simpl. unfold sig_fun_app. destruct (classic_dec (In A x)) as [hi | hni]. pose proof (proof_irrelevance _ h8 hi); subst; auto. contradiction. specialize (h9 _ h10). rewrite h9. unfold k. simpl. unfold f. f_equal. destruct constructive_definite_description as [x' h11]. simpl. simpl in h11. unfold sig_fun_app in h11. destruct (classic_dec (In A x)) as [h12 | h13]. apply h1 in h11. destruct x' as [x' h13]. simpl. apply exist_injective in h11. subst. reflexivity. contradiction. Qed. Lemma cart_prod_subsets_finite : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), Finite A -> Finite B -> FiniteT {S:Ensemble (T*U) | Included S (cart_prod A B)}. intros T U A B h1 h2. apply power_set_finitet. apply cart_prod_fin; assumption. Qed. Lemma finitet_fin_maps : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (def:U), Finite A -> Finite B -> FiniteT (Fin_map A B def). intros T U A B def h1 h2. pose (@fin_map_to_fps T U A B def) as f. assert (h4:forall F:(Fin_map A B def), Included (f F) (cart_prod A B)). intro F. apply fps_inc_cart_prod. pose proof (fin_map_to_fps_compat F) as h5. destruct h5. unfold f. assumption. pose (fun F:(Fin_map A B def) => exist (fun S:Ensemble (T*U) => Included S (cart_prod A B)) (f F) (h4 F)) as f'. pose proof (cart_prod_fin _ _ h1 h2) as h5. pose proof (power_set_finitet _ h5) as h6. apply (inj_finite _ _ f' h6). red. intros F1 F2 h8. unfold f' in h8. unfold f in h8. pose proof (exist_injective _ _ _ _ _ h8) as h9. apply fin_map_to_fps_inj. assumption. intros; tauto. Qed. Lemma finite_fin_maps : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (def:U), Finite A -> Finite B -> Finite (Full_set (Fin_map A B def)). intros T U A B def h1 h2. pose proof (finitet_fin_maps A B def h1 h2). apply FiniteT_Finite. assumption. Qed. Lemma finite_fin_maps_squared : forall {T U:Type} (A:Ensemble T) (B:Ensemble U) (def:U), Finite A -> Finite B -> Finite (Full_set ((Fin_map A B def) * (Fin_map A B def))). intros T U A B def h1 h2. pose proof (finite_fin_maps A B def h1 h2) as h3. assert (h4:Full_set (Fin_map A B def * Fin_map A B def) = cart_prod (Full_set (Fin_map A B def)) (Full_set (Fin_map A B def))). apply Extensionality_Ensembles. red. split. red. intros x ?. constructor. split; constructor. red. split. rewrite h4. apply cart_prod_fin; auto. Qed. Lemma finite_fin_map_ens : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (S:Ensemble (Fin_map A B def)), Finite A -> Finite B -> Finite S. intros T U A B def S h1 h2. assert (h0:Included S (Full_set (Fin_map A B def))). red. intros; constructor. apply Finite_downward_closed with (A:=Full_set (Fin_map A B def)). apply finite_fin_maps; auto. assumption. Qed. Lemma finite_fin_map_squared_ens : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (S:Ensemble ((Fin_map A B def) * (Fin_map A B def))), Finite A -> Finite B -> Finite S. intros T U A B def S h1 h2. assert (h0:Included S (Full_set ((Fin_map A B def) * (Fin_map A B def)))). red; intros; constructor. apply Finite_downward_closed with (A:=Full_set (Fin_map A B def * Fin_map A B def)). apply finite_fin_maps_squared; auto. assumption. Qed. Lemma fin_map_empty1 : forall {T U:Type} {B:Ensemble U} {def:U} (F:Fin_map (Empty_set T) B def), fin_map_to_fps F = Empty_set _. intros T U B def f. pose proof (fin_map_to_fps_compat f) as h1. destruct h1 as [h1 h2]. apply Extensionality_Ensembles. red; split; auto with sets. red. intros pr h3. destruct h1 as [h4 h5]. pose proof (h5 _ h3) as h6. destruct h6; contradiction. Qed. Lemma fin_map_cart_empty11 : forall {T U V:Type} {A:Ensemble U} {B:Ensemble V} {def:V} (F:Fin_map (cart_prod (Empty_set T) A) B def), fin_map_to_fps F = Empty_set _. intros T U V A B def F. pose proof (fin_map_to_fps_compat F) as h1. destruct h1 as [h1 h2]. apply Extensionality_Ensembles. red; split; auto with sets. red. intros pr h3. destruct h1 as [h4 h5]. pose proof (h5 _ h3) as h6. destruct h6 as [h6l h6r]. rewrite cart_prod_empty in h6l. contradiction. Qed. Lemma fin_map_cart_empty21 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble V} {def:V} (F:Fin_map (cart_prod A (Empty_set U)) B def), fin_map_to_fps F = Empty_set _. intros T U V A B def F. pose proof (fin_map_to_fps_compat F) as h1. destruct h1 as [h1 h2]. apply Extensionality_Ensembles. red; split; auto with sets. red. intros pr h3. destruct h1 as [h4 h5]. pose proof (h5 _ h3) as h6. destruct h6 as [h6l h6r]. rewrite cart_prod_empty' in h6l. contradiction. Qed. Lemma fin_map_empty2 : forall {T U:Type} {A:Ensemble T} {def:U} (F:Fin_map A (Empty_set U) def), fin_map_to_fps F = Empty_set _. intros T U A def f. pose proof (fin_map_to_fps_compat f) as h1. destruct h1 as [h1 h2]. inversion f as [hte h3 h4 S h5]. induction h3 as [|A h6 h7 x h8]. (*Empty*) destruct h1 as [h9 h10]. apply Extensionality_Ensembles. red. split. (* <= *) red. intros pr h11. pose proof (h10 _ h11) as h12. destruct h12; contradiction. (* >= *) auto with sets. (* Add *) destruct h1 as [h9 h10]. pose proof (h9 x (Add_intro2 _ A x)) as h11. destruct h11 as [y h12]. red in h12. destruct h12 as [h12]. destruct h12; contradiction. Qed. Lemma empty_map_fps : forall {T U:Type} {A:Ensemble T} {def:U} (F:(Fin_map A (Empty_set U) def)), functionally_paired A (Empty_set U) (Empty_set (T*U)). intros T U A def F. destruct F as [hte h1 h2 S h3]. pose proof (no_fp_empty2 _ _ h3) as h4. rewrite h4. apply fp_empty1. Qed. Lemma empty_map_unq1 : forall {T U:Type} {B:Ensemble U} {def:U} (F:Fin_map (Empty_set T) B def), F = fin_map_intro _ _ def (fin_map_type_eq_dec F) (Empty_is_finite T) (fin_map_fin_ran F) _ (fp_empty1 T U B). intros T U B def F. apply fin_map_app_inj. apply functional_extensionality. intro x. pose proof (Noone_in_empty _ x) as h1. pose proof (fin_map_app_def F _ h1) as h2. rewrite h2. simpl. pose proof (fin_fps_to_f_compat (fin_map_type_eq_dec F) (Empty_is_finite _) (Empty_set (T * U)) (fp_empty1 T U B) def x) as h3. destruct h3 as [h3 h4]. specialize (h4 h1). rewrite h4; auto. Qed. Lemma empty_map_unq1' : forall {T U:Type} {B:Ensemble U} {def:U} (F:Fin_map (Empty_set T) B def) (pft:Finite (Empty_set T)) (pfb:Finite B) (pffp:functionally_paired (Empty_set _) B (Empty_set _)), F = fin_map_intro _ _ def (fin_map_type_eq_dec F) pft pfb _ pffp. intros T U B def F h1 h2 h3. pose proof (empty_map_unq1 F) as h4. rewrite h4. pose proof (proof_irrelevance _ h1 (Empty_is_finite _)) as h5. pose proof (proof_irrelevance _ h2 (fin_map_fin_ran F)) as h6. pose proof (proof_irrelevance _ h3 (fp_empty1 T U B)) as h7. rewrite h5, h6, h7. f_equal. apply type_eq_dec_eq. Qed. Lemma cart_empty_map_unq11 : forall {T U V:Type} {A:Ensemble U} {B:Ensemble V} {def:V} (F:Fin_map (cart_prod (Empty_set T) A) B def), F = fin_map_intro _ _ def (fin_map_type_eq_dec F) (fin_map_fin_dom F) (fin_map_fin_ran F) (Empty_set _) (fp_cart_empty11 A B). intros T U V A B def F. apply fin_map_app_inj. apply functional_extensionality. intro pr. pose proof (Noone_in_empty _ pr) as h1. rewrite <- (cart_prod_empty A) in h1. pose proof (fin_map_app_def F _ h1) as h2. rewrite h2. pose proof (fin_map_fin_dom F) as h5. pose proof (fin_fps_to_f_compat (fin_map_type_eq_dec F) h5 (Empty_set (T * U * V)) (fp_cart_empty11 A B) def pr) as h3. simpl. destruct h3 as [h3 h4]. specialize (h4 h1). rewrite <- h4 at 1. apply fin_fps_to_f_functional; auto. Qed. Lemma cart_empty_map_unq21 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble V} {def:V} (F:Fin_map (cart_prod A (Empty_set U)) B def), F = fin_map_intro _ _ def (fin_map_type_eq_dec F) (fin_map_fin_dom F) (fin_map_fin_ran F) (Empty_set _) (fp_cart_empty21 A B). intros T U V A B def F. apply fin_map_app_inj. apply functional_extensionality. intro pr. pose proof (Noone_in_empty _ pr) as h1. rewrite <- (cart_prod_empty' A) in h1. pose proof (fin_map_app_def F _ h1) as h2. rewrite h2. simpl. pose proof (fin_fps_to_f_compat (fin_map_type_eq_dec F) (fin_map_fin_dom F) _ (fp_cart_empty21 A B) def pr) as h3. destruct h3 as [h3 h4]. specialize (h4 h1). rewrite h4; auto. Qed. Lemma cart_empty_map_unq11' : forall {T U V:Type} {A:Ensemble U} {B:Ensemble V} {def:V} (F:Fin_map (cart_prod (Empty_set T) A) B def) (pfa:Finite (cart_prod (Empty_set T) A)) (pfb:Finite B) (pffp:functionally_paired (cart_prod (Empty_set T) A) B (Empty_set _)), F = fin_map_intro _ _ def (fin_map_type_eq_dec F) pfa pfb _ pffp. intros T U V A B def F h1 h2 h3. pose proof (cart_empty_map_unq11 F) as h4. rewrite h4. pose proof (proof_irrelevance _ h1 (fin_map_fin_dom F)) as h5. pose proof (proof_irrelevance _ h2 (fin_map_fin_ran F)) as h6. pose proof (proof_irrelevance _ h3 (fp_cart_empty11 A B)) as h7. rewrite h5, h6, h7. f_equal. apply type_eq_dec_eq. Qed. Lemma cart_empty_map_unq21' : forall {T U V:Type} {A:Ensemble T} {B:Ensemble V} {def:V} (F:Fin_map (cart_prod A (Empty_set U)) B def) (pfa:Finite (cart_prod A (Empty_set U))) (pfb:Finite B) (pffp:functionally_paired (cart_prod A (Empty_set U)) B (Empty_set _)), F = fin_map_intro _ _ def (fin_map_type_eq_dec F) pfa pfb _ pffp. intros T U V A B def F h1 h2 h3. pose proof (cart_empty_map_unq21 F) as h4. rewrite h4. pose proof (proof_irrelevance _ h1 (fin_map_fin_dom F)) as h5. pose proof (proof_irrelevance _ h2 (fin_map_fin_ran F)) as h6. pose proof (proof_irrelevance _ h3 (fp_cart_empty21 A B)) as h7. rewrite h5, h6, h7. f_equal. apply type_eq_dec_eq. Qed. Lemma empty_fin_map_ex1 : forall (T U:Type) (pf:type_eq_dec T) (def:U) (B:Ensemble U), Finite B -> exists ! F:Fin_map (Empty_set T) B def, True. intros T U hte def B h1. exists (fin_map_intro _ _ def hte (Empty_is_finite T) h1 (Empty_set _) (fp_empty1 T U B)). red. split; auto. intros F ?. symmetry. pose proof (empty_map_unq1' F (Empty_is_finite _) h1 (fp_empty1 T U B)) as h2. rewrite h2. f_equal. apply type_eq_dec_eq. Qed. Lemma cart_empty_fin_map_ex21 : forall (T U V:Type) (pfte:type_eq_dec T) (pfue:type_eq_dec U) (def:V) (A:Ensemble T) (B:Ensemble V), Finite A -> Finite B -> exists ! F:Fin_map (cart_prod A (Empty_set U)) B def, True. intros T U V hte hue def A B h1 h2. exists (fin_map_intro _ _ def (impl_type_eq_dec_prod hte hue) (cart_prod_fin _ _ h1 (Empty_is_finite U)) h2 _ (fp_cart_empty21 A B)). red. split; auto. intros F ?. symmetry. erewrite cart_empty_map_unq21' at 1. f_equal. apply type_eq_dec_eq. Qed. Lemma cart_empty_fin_map_ex11 : forall (T U V:Type) (def:V) (pfte:type_eq_dec T) (pfue:type_eq_dec U) (A:Ensemble U) (B:Ensemble V), Finite A -> Finite B -> exists ! F:Fin_map (cart_prod (Empty_set T) A) B def, True. intros T U V def hte hue A B h1 h2. exists (fin_map_intro _ _ def (impl_type_eq_dec_prod hte hue) (cart_prod_fin _ _ (Empty_is_finite T) h1) h2 (Empty_set _) (fp_cart_empty11 A B)). red. split; auto. intros F ?. symmetry. erewrite cart_empty_map_unq11' at 1. f_equal. apply type_eq_dec_eq. Qed. Definition empty_map1 (T U:Type) (pfte:type_eq_dec T) (def:U) (B:Ensemble U) (pf:Finite B) : Fin_map (Empty_set T) B def := proj1_sig (constructive_definite_description _ (empty_fin_map_ex1 T U pfte def B pf)). Definition cart_empty_map11 (T U V:Type) (pfte:type_eq_dec T) (pfue:type_eq_dec U) (def:V) (A:Ensemble U) (B:Ensemble V) (pfa:Finite A) (pfb:Finite B) : Fin_map (cart_prod (Empty_set T) A) B def := proj1_sig (constructive_definite_description _ (cart_empty_fin_map_ex11 T U V def pfte pfue A B pfa pfb)). Definition cart_empty_map21 (T U V:Type) (pfte:type_eq_dec T) (pfue:type_eq_dec U) (def:V) (A:Ensemble T) (B:Ensemble V) (pfa:Finite A) (pfb:Finite B) : Fin_map (cart_prod A (Empty_set U)) B def := proj1_sig (constructive_definite_description _ (cart_empty_fin_map_ex21 T U V pfte pfue def A B pfa pfb)). Lemma empty_map1_compat : forall {T U:Type} {def:U} {B:Ensemble U} (F:Fin_map (Empty_set T) B def) (pf:Finite B), F = empty_map1 T U (fin_map_type_eq_dec F) def B pf. intros T U def B F pf. unfold empty_map1. destruct constructive_definite_description as [F']. simpl. destruct F' as [the h1 h2 S h3]. pose proof (fp_empty1_s _ _ h3) as h4. pose proof h3 as h5. rewrite h4 in h5. pose proof (empty_map_unq1' F h1 pf h5) as h6. rewrite h6. pose proof (proof_irrelevance _ h2 pf) as h7. rewrite h7. pose proof (subsetT_eq_compat _ _ _ _ h3 h5 h4) as h8. dependent rewrite -> h8. f_equal. apply type_eq_dec_eq. Qed. Lemma empty_map1_def : forall (T U:Type) (pfte:type_eq_dec T) (def:U) (B:Ensemble U) (pf:Finite B) (x:T), (empty_map1 T U pfte def B pf) |-> x = def. intros T U hte def B h1 x. unfold fin_map_app. destruct (empty_map1 T U hte def B h1) as [hte' h2 h3 S h4]. pose proof (fin_fps_to_f_compat hte' h2 S h4 def x) as h5. destruct h5 as [? h5]. apply h5. intro; contradiction. Qed. Lemma full_set_empty_map_sing1: forall (T U:Type) (pfte:type_eq_dec T) (def:U) (B:Ensemble U) (pf:Finite B), Full_set (Fin_map (Empty_set T) B def) = Singleton (empty_map1 T U pfte def B pf). intros T U hte def B pf. apply Extensionality_Ensembles; red; split. red. intros F ?. pose proof (empty_map1_compat F pf) as h2. rewrite h2. rewrite in_sing_iff. f_equal. apply type_eq_dec_eq. red. intros; constructor. Qed. Lemma cart_empty_map11_compat : forall {T U V:Type} {def:V} {A:Ensemble U} {B:Ensemble V} (pfte:type_eq_dec T) (pftu:type_eq_dec U) (F:Fin_map (cart_prod (Empty_set _) A) B def) (pfa:Finite A), F = cart_empty_map11 T U V pfte pftu def A B pfa (fin_map_fin_ran F). intros T U V def A B hdt hdu F h1. unfold cart_empty_map11. destruct constructive_definite_description as [F']. simpl. destruct F' as [hdtu h3 h4 S h5]. pose proof (@cart_prod_empty T U A) as h6. pose proof h5 as h7. rewrite h6 in h7. pose proof (fp_empty1_s _ _ h7) as h8. rewrite h8 in h7. rewrite <- h6 in h7. pose proof (cart_prod_fin _ _ (Empty_is_finite T) h1) as h9. pose proof (cart_empty_map_unq11' F h9 h4 h7) as h10. rewrite h10. pose proof (proof_irrelevance _ h9 h3) as h11. pose proof (subsetT_eq_compat _ _ _ _ h5 h7 h8) as h12. dependent rewrite -> h12. rewrite h11. f_equal. apply type_eq_dec_eq. Qed. Lemma cart_empty_map11_def : forall (T U V:Type) (pfte:type_eq_dec T) (pfue:type_eq_dec U) (def:V) (A:Ensemble U) (B:Ensemble V) (pfa:Finite A) (pfb:Finite B) (pr:(T*U)), (cart_empty_map11 T U V pfte pfue def A B pfa pfb) |-> pr = def. intros T U V hte hue def A B h1 h2 pr. unfold fin_map_app. destruct (cart_empty_map11 T U V hte hue def A B h1 h2) as [hdt h3 h4 S h5]. pose proof h5 as h5'. rewrite cart_prod_empty in h5'. apply fp_empty1_s in h5'. subst. pose proof (fin_fps_to_f_compat hdt h3 _ h5 def pr) as h9. destruct h9 as [? h9]. apply h9. intro h10. destruct h10 as [h10]. destruct h10; contradiction. Qed. Lemma cart_empty_map21_compat : forall {T U V:Type} {def:V} {A:Ensemble T} {B:Ensemble V} (pfte:type_eq_dec T) (pfue:type_eq_dec U) (F:Fin_map (cart_prod A (Empty_set _)) B def) (pfa:Finite A), F = cart_empty_map21 T U V pfte pfue def A B pfa (fin_map_fin_ran F). intros T U V def A B hte hue F h1. unfold cart_empty_map21. destruct constructive_definite_description as [F']. simpl. destruct F' as [htue h3 h4 S h5]. pose proof (@cart_prod_empty' T U A) as h6. pose proof h5 as h7. rewrite h6 in h7. pose proof (fp_empty1_s _ _ h7) as h8. rewrite h8 in h7. rewrite <- h6 in h7. pose proof (cart_prod_fin _ _ h1 (Empty_is_finite U) ) as h9. pose proof (cart_empty_map_unq21' F h9 h4 h7) as h10. rewrite h10. pose proof (proof_irrelevance _ h9 h3) as h11. pose proof (subsetT_eq_compat _ _ _ _ h5 h7 h8) as h12. dependent rewrite -> h12. rewrite h11. f_equal. apply type_eq_dec_eq. Qed. Lemma cart_empty_map21_def : forall (T U V:Type) (hte:type_eq_dec T) (hue:type_eq_dec U) (def:V) (A:Ensemble T) (C:Ensemble V) (pfa:Finite A) (pfc:Finite C) (pr:(T*U)), (cart_empty_map21 T U V hte hue def A C pfa pfc) |-> pr = def. intros T U V hte hue def A B h1 h2 pr. unfold fin_map_app. destruct (cart_empty_map21 T U V hte hue def A B h1 h2) as [htue h3 h4 S h5]. pose proof h5 as h5'. rewrite cart_prod_empty' in h5'. apply fp_empty1_s in h5'; subst. pose proof (fin_fps_to_f_compat htue h3 _ h5 def pr) as h6. destruct h6 as [? h6]; apply h6; clear h6; intro h6; destruct h6 as [h6]; destruct h6; contradiction. Qed. Lemma im2_empty1 : forall {T U V:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (D:Ensemble U) (C:Ensemble V) (def:V) (pfd:Finite D) (pfc:Finite C) (y:U), im2 (cart_empty_map11 T U V pfdt pfdu def D C pfd pfc) y = Empty_set _. intros T U V hdt hdu D C def hd hc y. apply Extensionality_Ensembles; red; split; auto with sets. red; intros x h2. destruct h2; contradiction. Qed. Lemma im2_empty2 : forall {T U V:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (A:Ensemble T) (C:Ensemble V) (def:V) (pfa:Finite A) (pfc:Finite C) (y:U), im2 (cart_empty_map21 T U V pfdt pfdu def A C pfa pfc) y = Empty_set _. intros T U V hdt hdu D C def hd hc y. apply Extensionality_Ensembles; red; split; auto with sets. red; intros x h2. destruct h2; contradiction. Qed. Lemma empty_map_unq2: forall {T U:Type} {A:Ensemble T} {def:U} (F:(Fin_map A (Empty_set U) def)), F = fin_map_intro _ _ def (fin_map_type_eq_dec F) (fin_map_fin_dom F) (Empty_is_finite U) _ (empty_map_fps F). intros T U A def F. apply fin_map_ext. intros x h1. simpl. pose proof (fin_map_app_in F _ h1). contradiction. Qed. Lemma empty_map_unq2' : forall {T U:Type} {A:Ensemble T} {def:U} (F:(Fin_map A (Empty_set U) def)) (pfa:Finite A) (pfe:Finite (Empty_set U)) (pffp:functionally_paired A (Empty_set _) (Empty_set _)), F = fin_map_intro _ _ _ (fin_map_type_eq_dec F) pfa pfe _ pffp. intros T U A def F h1 h2 h3. pose proof (empty_map_unq2 F) as h4. rewrite h4. pose proof (proof_irrelevance _ h1 (fin_map_fin_dom F)) as h5. pose proof (proof_irrelevance _ h2 (Empty_is_finite U)) as h6. pose proof (proof_irrelevance _ h3 (empty_map_fps F)) as h7. rewrite h5, h6, h7; f_equal; apply type_eq_dec_eq. Qed. Lemma empty_fin_map_ex2 : forall (T U:Type) (pfte:type_eq_dec T) (def:U), exists ! F:Fin_map (Empty_set T) (Empty_set U) def, True. intros T U hte def. exists (fin_map_intro _ _ _ hte (Empty_is_finite T) (Empty_is_finite U) (Empty_set _) (fp_empty1 T U _)). red. split; auto. intros F ?. symmetry. rewrite empty_map_unq2' at 1. f_equal. apply type_eq_dec_eq. Qed. Definition empty_map (T U:Type) (pfte:type_eq_dec T) (def:U) : Fin_map (Empty_set T) (Empty_set U) def. refine (proj1_sig (constructive_definite_description _ (empty_fin_map_ex2 T U pfte def))). Defined. Lemma empty_map_compat : forall {T U:Type} {def:U} (F:Fin_map (Empty_set T) (Empty_set U) def), F = empty_map T U (fin_map_type_eq_dec F) def. intros T U def F. unfold empty_map. destruct constructive_definite_description as [F']. simpl. destruct F' as [hdt h1 h2 S h3]. pose proof (no_fp_empty2' _ _ h3). subst. rewrite empty_map_unq2' at 1. f_equal. apply type_eq_dec_eq. Qed. Lemma empty_map_def : forall (T U:Type) (pfte:type_eq_dec T) (def:U), forall x:T, (empty_map T U pfte def) |-> x = def. intros T U hte def x. destruct (empty_map T U hte def) as [hdtu h1 h2 S h3]. simpl. pose proof (fp_empty1_s _ _ h3). subst. pose proof (fin_fps_to_f_compat hdtu h1 (Empty_set (T * U)) h3 def x) as h4. apply h4. intro. contradiction. Qed. Lemma full_set_empty_map_sing_empty: forall (T U:Type) (pfdt:type_eq_dec T) (def:U), Full_set (Fin_map (Empty_set T) (Empty_set U) def) = Singleton (empty_map T U pfdt def). intros T U hdt def. apply Extensionality_Ensembles. red. split. red. intros F ?. pose proof (empty_map_compat F) as h1. rewrite h1. rewrite in_sing_iff. f_equal. apply type_eq_dec_eq. red. intros; constructor. Qed. Lemma empty_map_ex2_fin : forall {T U:Type} (A:Ensemble T) (def:U), Finite A -> A <> Empty_set T -> forall F:Fin_map A (Empty_set U) def, False. intros T U A def h1. induction h1. intro h2. contradict h2. reflexivity. intros ? F. pose proof (fin_map_app_in F x (Add_intro2 _ A x)). contradiction. Qed. Lemma full_set_non_empty_map_empty : forall (T U:Type) (def:U) (A:Ensemble T), A <> Empty_set T -> Full_set (Fin_map A (Empty_set U) def) = Empty_set _. intros T U def A h1. apply Extensionality_Ensembles; red;split; auto with sets. red. intros F ?. inversion F as [hdt h2 h3 S h4]. pose proof (empty_map_ex2_fin _ _ h2 h1 F). contradiction. Qed. Lemma fun_to_fin_map_empty_set1 : forall {T U:Type} (pfte:type_eq_dec T) (def:U) (pfe:Finite (Empty_set _)) (f:T->U), fun_to_fin_map pfte _ def pfe f = empty_map1 T U pfte def (Im (Empty_set _) f) (finite_image _ _ _ f pfe). intros; rewrite empty_map1_compat at 1. f_equal. apply type_eq_dec_eq. Qed. Lemma fun_to_fin_map_empty_set11 : forall {T U V:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:V) (B:Ensemble U) (pfb:Finite B) (pfce:Finite (cart_prod (Empty_set _) B)) (f:T*U->V), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in fun_to_fin_map pfdtu _ def pfce f = cart_empty_map11 T U V pfdt pfdu def _ _ pfb (finite_image _ _ _ f pfce). intros T U V hdt hdu def B h1 h3 f. simpl. apply fin_map_ext. intros pr h4. destruct h4 as [h4]. destruct h4; contradiction. Qed. Lemma fun_to_fin_map_empty_set21 : forall {T U V:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:V) (A:Ensemble T) (pfa:Finite A) (pfce:Finite (cart_prod A (Empty_set _))) (f:T*U->V), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in fun_to_fin_map pfdtu _ def pfce f = cart_empty_map21 T U V pfdt pfdu def _ _ pfa (finite_image _ _ _ f pfce). intros T U V hdt hdu def A h1 h3 f. apply fin_map_ext. intros pr h4. destruct h4 as [h4]. destruct h4; contradiction. Qed. Lemma im1_fun_to_fin_map : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} (pfte:type_eq_dec T) (pfue:type_eq_dec U) (f:(T*U)->V) (def:V) (pf:Finite (cart_prod A B)) (x:T), Finite A -> In A x -> let pfdtu := impl_type_eq_dec_prod pfte pfue in im1 (fun_to_fin_map pfdtu (cart_prod A B) def pf f) x = Im B (fun y => f (x, y)). intros. apply Extensionality_Ensembles; red; split; red; intros y h4. destruct h4 as [h4 y h6]; subst. rewrite fun_to_fin_map_compat. econstructor. apply h6. reflexivity. constructor; auto. destruct h4 as [y h4]; subst. econstructor; auto. apply h4. rewrite fun_to_fin_map_compat; auto. constructor; auto. Qed. Lemma im2_fun_to_fin_map : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} (pfte:type_eq_dec T) (pfue:type_eq_dec U) (f:(T*U)->V) (def:V) (pf:Finite (cart_prod A B)) (y:U), Finite B -> In B y -> let pfdtu := impl_type_eq_dec_prod pfte pfue in im2 (fun_to_fin_map pfdtu (cart_prod A B) def pf f) y = Im A (fun x => f (x, y)). intros. apply Extensionality_Ensembles; red; split; red; intros x h4. destruct h4 as [h4 x h6]; subst. rewrite fun_to_fin_map_compat. econstructor. apply h6. reflexivity. constructor; auto. destruct h4 as [x h4]; subst. econstructor; auto. apply h4. rewrite fun_to_fin_map_compat; auto. constructor; auto. Qed. Lemma fin_map_im_eq_ex : forall {T U V:Type} {A:Ensemble T} {g:T->U} {B:Ensemble V} {defv:V} (pfdt:type_eq_dec T) (F:Fin_map (Im A g) B defv), Finite A -> exists! F':Fin_map A B defv, forall x:T, In A x -> F' |-> x = F |-> (g x). intros T U V A g B defv hte F h1. pose proof (finite_inh_or_empty _ h1) as h2. destruct h2 as [hi | hni]. destruct hi as [a hi]. pose (fun x:sig_set A => g (proj1_sig x)) as g'. assert (h2:Im A g = Im (full_sig A) g'). apply Extensionality_Ensembles. red. split. red. intros u h2. destruct h2 as [u h2]. subst. apply Im_intro with (exist _ _ h2). constructor. unfold g'. simpl. reflexivity. red. intros x h2. destruct h2 as [x h2]. subst. unfold g'. clear h2. destruct x as [x h2]. simpl. apply Im_intro with x; auto. assert (h3:Included (Im (full_sig A) g') (Im A g)). rewrite h2. auto with sets. pose (restriction F h3) as F'. pose (@fin_map_im_full_sig_eq). exists (fin_map_im_full_sig_eq hte F' h1 (g a)). red. split. intros x h4. rewrite fin_map_im_full_sig_eq_compat. unfold F'. rewrite restriction_compat. f_equal. unfold g'. simpl. unfold sig_fun_app. destruct classic_dec as [h5 | h6]. simpl. reflexivity. contradiction. apply Im_intro with (exist _ _ h4). constructor. simpl. unfold sig_fun_app. destruct classic_dec as [h5 | h6]. pose proof (proof_irrelevance _ h4 h5); subst; auto. contradiction. assumption. intros F'' h4. apply fin_map_ext. intros x h5. rewrite fin_map_im_full_sig_eq_compat. unfold F'. rewrite restriction_compat. rewrite h4. f_equal. simpl. unfold sig_fun_app. destruct classic_dec as [h6 | h7]. unfold g'. simpl. reflexivity. contradiction. assumption. apply Im_intro with (exist _ _ h5). constructor. simpl. unfold sig_fun_app. destruct classic_dec as [h6 | h7]. f_equal. apply proj1_sig_injective. reflexivity. contradiction. assumption. subst. exists (empty_map1 T V hte defv B (fin_map_fin_ran F)). red. split. intros; contradiction. intros. pose proof (empty_fin_map_ex1) as h2. rewrite (empty_map1_compat x') at 1. f_equal. apply type_eq_dec_eq. Qed. Definition fin_map_im_eq {T U V:Type} {A:Ensemble T} {g:T->U} {B:Ensemble V} {defv:V} (pfdt:type_eq_dec T) (F:Fin_map (Im A g) B defv) (pf:Finite A) := proj1_sig (constructive_definite_description _ (fin_map_im_eq_ex pfdt F pf)). Lemma fin_map_im_eq_compat : forall {T U V:Type} {A:Ensemble T} {g:T->U} {B:Ensemble V} {defv:V} (pfdt:type_eq_dec T) (F:Fin_map (Im A g) B defv) (pf:Finite A), let F' := fin_map_im_eq pfdt F pf in forall x:T, In A x -> F' |-> x = F |-> (g x). unfold fin_map_im_eq; intros; destruct constructive_definite_description; auto. Qed. Lemma fin_map_to_fin_map_im_eq : forall {T U V:Type} {A:Ensemble T} {B:Ensemble V} {defv:V} (pfue:type_eq_dec U) (F : Fin_map A B defv) (g : T -> U), injective g -> exists F' : Fin_map (Im A g) B defv, forall x : T, In A x -> F' |-> (g x) = F |-> x. intros T U V A B defv hue F g h1. pose proof (finite_inh_or_empty _ (fin_map_fin_dom F)) as hf. destruct hf as [h2 | h3]. destruct h2 as [a h2]. pose (fun x:sig_set A => g (proj1_sig x)) as g'. assert (h3:injective g'). red. intros x y h3. unfold g' in h3. apply h1 in h3. apply proj1_sig_injective. assumption. pose proof (fin_map_to_fin_map_im_full_sig_comp hue F g' (g' (exist _ _ h2)) h3) as h4. destruct h4 as [F' h4]. assert (h5:Included (Im A g) (Im (full_sig A) g')). red. intros u h5. destruct h5 as [u h5]. subst. apply Im_intro with (exist _ _ h5). constructor. unfold g'. simpl. reflexivity. exists (restriction F' h5). intros x h6. rewrite restriction_compat. rewrite <- h4. f_equal. simpl. unfold sig_fun_app. destruct classic_dec as [h7 | h8]. unfold g'. simpl. reflexivity. contradiction. assumption. apply Im_intro with x; auto. subst. pose (empty_map1 U V hue defv B (fin_map_fin_ran F)) as f. pose proof (image_empty T U g) as h2. assert (h3:Included (Im (Empty_set T) g) (Empty_set U)). rewrite h2. auto with sets. exists (restriction f h3). intros; contradiction. Qed. Lemma im_full_set_fin_map_im_inj : forall {T U V W:Type} {A:Ensemble T} {B:Ensemble V} {defv:V} (deft:T) (pfte:type_eq_dec T) (pfue:type_eq_dec U), Finite A -> forall (g:T->U) (pfi:injective g) (k:Fin_map (Im A g) B defv->W), Im (Full_set (Fin_map (Im A g) B defv)) k = Im (Full_set (Fin_map A B defv)) (fun f:Fin_map A B defv => (k (fin_map_compose_inj pfue f deft g pfi))). intros T U V W A B defv deft hte hue h1 g h2 k. apply Extensionality_Ensembles. red. split. red. intros v h3. destruct h3 as [v h3]. subst. clear h3. pose (fin_map_im_eq hte v h1) as F. apply Im_intro with F. constructor. f_equal. apply fin_map_ext. intros x h3. destruct h3 as [x h3]. subst. rewrite <- fin_map_compose_inj_compat. unfold F. rewrite fin_map_im_eq_compat. reflexivity. assumption. assumption. red. intros v h3. destruct h3 as [v h3]. subst. clear h3. pose proof (fin_map_to_fin_map_im_eq hue v _ h2) as h3. destruct h3 as [F h3]. apply Im_intro with F. constructor. f_equal. apply fin_map_ext. intros x h4. destruct h4 as [x h4]. subst. rewrite h3; auto. rewrite <- fin_map_compose_inj_compat. reflexivity. assumption. Qed. Definition fin_map_dom_subst {T U:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (pf:A = A') (F:Fin_map A B def) : Fin_map A' B def. subst. refine F. Defined. Lemma fin_map_dom_subst_eq_refl : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def), fin_map_dom_subst eq_refl F = F. intros T U A B def F. apply fin_map_ext. intros x h1. unfold fin_map_dom_subst, eq_rect_r, eq_rect. simpl. reflexivity. Qed. Lemma fin_map_dom_subst_compat : forall {T U:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (pf:A = A') (F:Fin_map A B def) (x:T), let F' := fin_map_dom_subst pf F in In A x -> F |-> x = F' |-> x. intros T U A A' B def h1 F x. subst. intros F' h2. unfold F'. rewrite fin_map_dom_subst_eq_refl. reflexivity. Qed. Definition fin_map_dom_subst_fun {T U V:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (pf:A = A') (k:Fin_map A B def->V) : Fin_map A' B def->V. intro f. subst. refine (k f). Defined. Lemma im_full_set_fin_map_dom_subst : forall {T U V:Type} {A A':Ensemble T} {B:Ensemble U} {def:U} (pf:A = A') (k:Fin_map A B def->V), Im (Full_set (Fin_map A B def)) k = Im (Full_set (Fin_map A' B def)) (fin_map_dom_subst_fun pf k). intros T U V A A' B def h1 k. subst. unfold fin_map_dom_subst_fun. unfold eq_rect_r. simpl. f_equal. Qed. (*This changes the value of an element in the domain of the finite map*) Lemma fin_map_switch_ex : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def : U} (pfdt:type_eq_dec T) (F:Fin_map A B def) (key:T) (val:U) (ink:In A key) (inv:In B val), exists! F':Fin_map A B def, F' |-> key = val /\ (forall x:T, In A x -> x <> key -> F' |-> x = F |-> x). intros T U A B def h0 F key val h1 h2. destruct F as [hdt h3 h4 S h5]. pose (Setminus S [pr:T*U | fst pr = key]) as S'. pose (Add S' (key, val)) as S''. assert (h6:functionally_paired A B S''). constructor. intros x h6. destruct (h0 x key) as [| h7]; subst. exists val. red. split. split; auto. unfold S'', S'. right. constructor. intros x h7. destruct h7 as [h7 h8]. inversion h8 as [pr h9 | pr h9]. subst. destruct h9 as [h9 h10]. contradict h10. constructor; auto. subst. inversion h9; auto. exists (fin_fps_to_f hdt h3 _ h5 val x). red. split. split. apply fin_fps_to_f_in_ran; auto. left. constructor. apply fin_fps_to_f_s_compat; auto. intro h8. destruct h8 as [h8]. simpl in h8. contradiction. intros y h9. destruct h9 as [h9 h10]. inversion h10 as [pr h11 | pr h11 h13]. subst. destruct h11 as [h11 h12]. apply fp_inv'; auto. inversion h11. contradict h7; auto. intros pr h6. destruct h6 as [pr h6 | pr h6]. destruct h6 as [h6 h7]. rewrite (fin_fps_to_f_compat' hdt h3 h5 def) in h6. destruct h6 as [h6]. destruct h6 as [h6 h8]. split; auto. rewrite h6. apply fin_fps_to_f_in_ran; auto. destruct h6. simpl. split; auto. exists (fin_map_intro A B def hdt h3 h4 S'' h6). red. split. split. simpl. apply fp_inv'; auto. right. constructor. intros x h7 h8. simpl. apply fp_inv'. left. constructor. apply fin_fps_to_f_s_compat. assumption. intro h9. destruct h9 as [h9]. simpl in h9. contradiction. intros F' h7. destruct h7 as [h7 h8]. apply fin_map_ext. intros x h9. specialize (h8 _ h9). simpl. simpl in h8. destruct (h0 x key) as [h10 | h10]. rewrite h10. rewrite h7. apply fp_inv'; auto. right. constructor. rewrite h8; auto. pose proof (fin_fps_to_f_compat hdt h3 S'' h6 def x) as h11. pose proof (fin_fps_to_f_compat hdt h3 S h5 def x) as h12. destruct h11 as [h11 h13]. destruct h12 as [h12 h14]. specialize (h11 h9). specialize (h12 h9). inversion h11 as [pr h15a h15b | pr h15a h15b]. destruct h15a as [h15a h15c]. eapply fp_functional. apply h5. apply h15a. assumption. inversion h15a as [h18]. contradict h10; auto. Qed. Definition fin_map_switch {T U:Type} {A:Ensemble T} {B:Ensemble U} {def : U} (pfdt:type_eq_dec T) (F:Fin_map A B def) (key:T) (val:U) (ink:In A key) (inv:In B val) : Fin_map A B def := proj1_sig (constructive_definite_description _ (fin_map_switch_ex pfdt F key val ink inv)). Lemma fin_map_switch_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def : U} (pfdt:type_eq_dec T) (F:Fin_map A B def) (key:T) (val:U) (ink:In A key) (inv:In B val), let F' := fin_map_switch pfdt F key val ink inv in F' |-> key = val /\ (forall x:T, In A x -> x <> key -> F' |-> x = F |-> x). unfold fin_map_switch; intros; destruct constructive_definite_description; auto. Qed. Require Import ListUtilities. Lemma fin_map_seg_to_im_in_ens_ex : forall {T U:Type} {n:nat} {B:Ensemble U} {defu:U} (pfdt:type_eq_dec T) (F:Fin_map (seg n) B defu) (l:list T) (pfn:n <= length l), exists! F':Fin_map (im_im_lind_singularize pfdt l n pfn) B defu, forall x (pf:Inn (im_im_lind_singularize pfdt l n pfn) x), F' |-> x = F |-> inv_inv_im_im_lind_singularize pfdt l n pfn x pf. intros T U n B defu hdt F l hn. pose (fun (x:sig_set (im_im_lind_singularize hdt l n hn)) => F |-> inv_inv_im_im_lind_singularize hdt l n hn _ (proj2_sig x)) as f. pose (sig_fun_to_fin_map_ran hdt f (finite_im_im_lind_singularize hdt _ _ _) defu B (fin_map_fin_ran F)) as F'. hfirst F'. red; intros x h2. destruct h2 as [x h2]; subst. unfold f. apply fin_map_app_in. constructor. apply lt_inv_inv_im_im_lind_singularize. exists (F' hf). red; split. intros x h1. unfold F'. erewrite sig_fun_to_fin_map_ran_compat. unfold f; simpl. reflexivity. intros G h1. apply fin_map_ext. intros x h2. unfold F'. erewrite sig_fun_to_fin_map_ran_compat. unfold f. simpl. symmetry. apply h1. Grab Existential Variables. assumption. Qed. Definition fin_map_seg_to_im_in_ens {T U:Type} {n:nat} {B:Ensemble U} {defu:U} (pfdt:type_eq_dec T) (F:Fin_map (seg n) B defu) (l:list T) (pfn:n <= length l) : Fin_map (im_im_lind_singularize pfdt l n pfn) B defu := proj1_sig (constructive_definite_description _ (fin_map_seg_to_im_in_ens_ex pfdt F l pfn)). Lemma fin_map_seg_to_im_in_ens_compat : forall {T U:Type} {n:nat} {B:Ensemble U} {defu:U} (pfdt:type_eq_dec T) (F:Fin_map (seg n) B defu) (l:list T) (pfn:n <= length l), let F' := fin_map_seg_to_im_in_ens pfdt F l pfn in forall x (pf:Inn (im_im_lind_singularize pfdt l n pfn) x), F' |-> x = F |-> inv_inv_im_im_lind_singularize pfdt l n pfn x pf. unfold fin_map_seg_to_im_in_ens; intros; destruct constructive_definite_description; auto. Qed. Lemma fin_map_to_fps_list_power_compat : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (A:Ensemble T) (B:Ensemble U) (def:U) (F:Fin_map A B def) (la:list T) (lb:list U) (lp:list (T*U)), list_to_set la = A -> list_to_set lb = B -> list_to_set lp = fin_map_to_fps F -> NoDup la -> NoDup lp -> synced la lp -> In lp (list_power la lb). intros T U hdt hdu A B def F la lb lp h1 h2 h3 h4 h5 h6. pose proof (fin_map_to_fps_compat F) as h7. destruct h7 as [h8 h9]. clear h9. rewrite (fp_fpl_compat _ _ _ _ _ _ h1 h2 h3) in h8. rewrite fpl_in_list_power in h8; auto. Qed. Lemma fin_map_to_list_pairs_lemma : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (F:Fin_map A B def) (la:list T) (lb:list U), NoDup la -> list_to_set la = A -> list_to_set lb = B -> exists (lp:list (T*U)), list_to_set lp = fin_map_to_fps F /\ NoDup lp /\ In lp (list_power la lb). intros T U A B def hdt hdu F la lb h0 h1 h2. pose proof (fin_map_to_fps_compat F) as h3. destruct h3 as [h4 h5]. clear h5. pose proof (fin_map_to_fps_finite F) as h5. pose proof (finite_set_list _ h5) as h6. destruct h6 as [lp h7]. symmetry in h7. rewrite (fp_fpl_compat _ _ _ _ _ _ h1 h2 h7) in h4. pose proof (sync_fpl h4) as h6. destruct h6 as [lp' h6]. destruct h6 as [h6a [h6b h6c]]. exists lp'. repeat split. rewrite <- h7. rewrite h6c. reflexivity. apply (synced_no_dup _ _ h0 h6b). pose proof (synced_no_dup _ _ h0 h6b) as h8. rewrite h6c in h7. apply (fin_map_to_fps_list_power_compat hdt hdu _ _ _ _ _ _ _ h1 h2 h7); assumption. Qed. Record nice_map_lists {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) := {n_la:list T; n_lb:list U; n_lp:list (T*U); n_im := map (@snd T U) n_lp; la_compat : list_to_set n_la = A; lb_compat : list_to_set n_lb = B; lp_compat : list_to_set n_lp = fin_map_to_fps F; nda : NoDup n_la; ndb : NoDup n_lb; ndp: NoDup n_lp; in_lp: In n_lp (list_power n_la n_lb); fpl : functionally_paired_l n_la n_lb n_lp}. (*The complications relating to testing if the cartesian product is empty accounts for such aberrations as the cartesian product of an infinite set and the empty set, which, empty, would qualify as the domain of a finite map according to my definition. This is why aberrations like [lab_empty2] and [lab_empty2'] appear.*) Record nice_map_lists2 {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) := {n_la2:list T; n_lb2:list U; n_lab2:list (T*U) := list_prod n_la2 n_lb2; n_lc2:list V; n_lp2:list (T*U*V); n_im2 := map (@snd (T*U) V) n_lp2; la_lb_compat2 : Finite A -> Finite B -> list_to_set n_la2 = A /\ list_to_set n_lb2 = B; lab_compat2 : list_to_set n_lab2 = cart_prod A B; lab_empty2 : if (nil_dec' n_lab2) then (n_la2 = nil /\ A = Empty_set _) \/ (n_lb2 = nil /\ B = Empty_set _) else (list_to_set n_la2 = A /\ list_to_set n_lb2 = B); lab_empty2' : if (nil_dec' n_lab2) then (A = Empty_set _ -> n_la2 = nil) /\ (B = Empty_set _ -> n_lb2 = nil) else True; lc_compat2 : list_to_set n_lc2 = C; lp_compat2 : list_to_set n_lp2 = fin_map_to_fps F; nda2 : NoDup n_la2; ndb2 : NoDup n_lb2; ndab2 : NoDup n_lab2; ndc : NoDup n_lc2; ndp2 : NoDup n_lp2; in_lp2 : In n_lp2 (list_power n_lab2 n_lc2); fpl2 : functionally_paired_l n_lab2 n_lc2 n_lp2}. Lemma fin_map_ex_nice_map_lists : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (F:Fin_map A B def), exists _:(nice_map_lists F), True. intros T U A B def hdt hdu F. pose proof (fin_map_fin_dom F) as h1. pose proof (fin_map_fin_ran F) as h2. pose proof (finite_set_list_no_dup _ h1) as h3. pose proof (finite_set_list_no_dup _ h2) as h4. destruct h3 as [la h3]. destruct h3 as [h3l h3r]. destruct h4 as [lb h4]. destruct h4 as [h4l h4r]. symmetry in h3l. symmetry in h4l. pose proof (fin_map_to_list_pairs_lemma hdt hdu F _ _ h3r h3l h4l) as h5. destruct h5 as [lp h5]. destruct h5 as [h5a [h5b h5c]]. pose proof (in_list_power_synced _ _ _ h5c) as h6. pose proof (fpl_in_list_power hdt hdu _ lb _ h6 h3r h5b) as h7. pose proof h5c as h8. rewrite <- h7 in h8. exists (Build_nice_map_lists _ _ _ _ _ F _ _ _ h3l h4l h5a h3r h4r h5b h5c h8). constructor. Qed. Lemma fin_map_ex_nice_map_lists_list_to_set : forall {T U:Type} {la:list T} {lb:list U} {def:U} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (F:Fin_map (list_to_set la) (list_to_set lb) def), NoDup la -> NoDup lb -> exists nml:(nice_map_lists F), la = n_la _ nml /\ lb = n_lb _ nml. intros T U la lb def hdt hdu F h1 h2. pose proof (fin_map_to_list_pairs_lemma hdt hdu F la lb h1 (eq_refl _) (eq_refl _)) as h3. destruct h3 as [lp h3]. destruct h3 as [h3a [h3b h3c]]. pose proof (in_list_power_synced _ _ _ h3c) as h4. pose proof (fpl_in_list_power hdt hdu _ lb _ h4 h1 h3b) as h5. pose proof h3c as h8. rewrite <- h5 in h8. exists (Build_nice_map_lists _ _ _ _ _ F la lb lp (eq_refl _) (eq_refl _) h3a h1 h2 h3b h3c h8). simpl. split; auto. Qed. Lemma fin_map_ex_nice_map_lists_list_to_set' : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (F:Fin_map A B def) (la:list T) (lb:list U), list_to_set la = A -> list_to_set lb = B -> NoDup la -> NoDup lb -> exists nml:(nice_map_lists F), la = n_la _ nml /\ lb = n_lb _ nml. intros T U A B def hdt hdu F la lb h1 h2 h3 h4. pose proof (fin_map_to_list_pairs_lemma hdt hdu F la lb h3 h1 h2) as h5. destruct h5 as [lp h5]. destruct h5 as [h5a [h5b h5c]]. pose proof (in_list_power_synced _ _ _ h5c) as h4'. pose proof (fpl_in_list_power hdt hdu _ lb _ h4' h3 h5b) as h5'. pose proof h5c as h8. rewrite <- h5' in h8. exists (Build_nice_map_lists _ _ _ _ _ F la lb lp h1 h2 h5a h3 h4 h5b h5c h8). simpl. split; auto. Qed. Lemma fin_map_ex_nice_map_lists_list_to_set_dom : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (F:Fin_map A B def) (la:list T), list_to_set la = A -> NoDup la -> exists nml:(nice_map_lists F), la = n_la _ nml. intros T U A B def hdt hdu F la h1 h2. pose proof (fin_map_fin_ran F) as h3. pose proof (finite_set_list_no_dup _ h3) as h4. destruct h4 as [lb h4]. destruct h4 as [h4l h4r]. symmetry in h4l. pose proof (fin_map_to_list_pairs_lemma hdt hdu F la lb h2 h1 h4l) as h5. destruct h5 as [lp h5]. destruct h5 as [h5a [h5b h5c]]. pose proof (in_list_power_synced _ _ _ h5c) as h4'. pose proof (fpl_in_list_power hdt hdu _ lb _ h4' h2 h5b) as h5'. pose proof h5c as h8. rewrite <- h5' in h8. exists (Build_nice_map_lists _ _ _ _ _ F la lb lp h1 h4l h5a h2 h4r h5b h5c h8). simpl. reflexivity. Qed. Lemma fin_map_ex_nice_map_lists_intro : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (A:Ensemble T) (B:Ensemble U) (la:list T) (lb:list U) (lp:list (T*U)) (def:U) (pfa:Finite A) (pfb:Finite B) (pfin:In lp (list_power la lb)) (pffp:functionally_paired A B (list_to_set lp)), list_to_set la = A -> list_to_set lb = B -> NoDup la -> NoDup lb -> exists nml:(nice_map_lists (fin_map_intro A B def pfdt pfa pfb (list_to_set lp) pffp)), la = n_la _ nml /\ lb = n_lb _ nml /\ lp = n_lp _ nml. intros T U hdt hdu A B la lb lp def h1 h2 h3 h4 h1' h2' h3' h4'. pose proof (in_list_power_synced _ _ _ h3) as h4''. pose proof (list_power_no_dup _ _ _ h3' h4' h3) as h5. pose proof (fpl_in_list_power hdt hdu _ lb _ h4'' h3' h5) as h5'. pose proof (fp_fpl_compat _ _ _ _ _ _ h1' h2' (eq_refl (list_to_set lp))) as h6. pose proof h4 as h8. rewrite h6 in h8. pose proof h8 as h9. rewrite h5' in h9. pose proof (fin_map_to_fps_s_compat hdt _ _ def h1 h2 _ h4) as h10. exists (Build_nice_map_lists _ _ _ _ _ _ la lb lp h1' h2' h10 h3' h4' h5 h9 h8). simpl. split; auto. Qed. Lemma fin_map_ex_nice_map_lists2 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (F:Fin_map (cart_prod A B) C def), exists _:(nice_map_lists2 F), True. intros T U V A B C def hdt hdu hdv F. pose proof (fin_map_fin_dom F) as h1. dependent induction h1. pose proof (fin_map_fin_ran F) as h1. pose proof (finite_set_list_no_dup _ h1) as h3. destruct h3 as [lc h3]. destruct h3 as [h3l h3r]. symmetry in h3l. pose proof (fin_map_to_fps_finite F) as h4. pose proof (finite_set_list_no_dup _ h4) as h5. destruct h5 as [lp h5]. destruct h5 as [h5l h5r]. symmetry in h5l. destruct (classic_dec (Finite A /\ Finite B)) as [hf | nhf]. destruct hf as [hfa hfb]. pose proof (finite_set_list_no_dup _ hfa) as h0. pose proof (finite_set_list_no_dup _ hfb) as h0'. destruct h0 as [la h0]. destruct h0' as [lb h0']. destruct h0 as [h0l h0r]. destruct h0' as [h0l' h0r']. assert (hn: Finite A -> Finite B -> list_to_set la = A /\ list_to_set lb = B). symmetry in h0l. symmetry in h0l'. tauto. assert (h6':list_to_set (list_prod la lb) = cart_prod A B). apply list_prod_cart_prod_compat; auto. assert (h9':NoDup (list_prod la lb)). apply no_dup_list_prod; assumption. assert (h10':lp = nil). symmetry in x. pose proof (cart_prod_empty_rev A B x) as h10. destruct h10 as [h10l | h10r]. generalize dependent F. rewrite h10l. intros F h11 h12. rewrite fin_map_cart_empty11 in h12. pose proof (empty_set_nil _ h12). subst. reflexivity. generalize dependent F. rewrite h10r. intros F h11 h12. rewrite fin_map_cart_empty21 in h12. pose proof (empty_set_nil _ h12). subst. reflexivity. assert (h11':In lp (list_power (list_prod la lb) lc)). rewrite <- x in h6'. pose proof (empty_set_nil _ h6') as h12. rewrite h12. simpl. left. rewrite h10'. reflexivity. assert (h12':functionally_paired_l (list_prod la lb) lc lp). rewrite <- x in h6'. pose proof (empty_set_nil _ h6') as h12. rewrite h12. rewrite h10'. apply fpl_empty1. assert (h13': if nil_dec' (list_prod la lb) then (la = nil /\ A = Empty_set _) \/ (lb = nil /\ B = Empty_set _) else (list_to_set la = A /\ list_to_set lb = B)). destruct (nil_dec' (list_prod la lb)) as [h14 | h15]. symmetry in x. pose proof (@cart_prod_empty_rev _ _ _ _ x) as h15. destruct h15 as [h15l | h15r]. left. rewrite h15l in h0l. symmetry in h0l. pose proof (empty_set_nil _ h0l). split; auto. right. rewrite h15r in h0l'. symmetry in h0l'. pose proof (empty_set_nil _ h0l'). split; auto. contradict h15. rewrite <- x in h6'. apply empty_set_nil. assumption. assert (h13'': if (nil_dec' (list_prod la lb)) then (A = Empty_set _ -> la = nil) /\ (B = Empty_set _ -> lb = nil) else True). destruct (nil_dec' (list_prod la lb)) as [h14 | h15]. split. intro h16. rewrite h16 in h0l. symmetry in h0l. apply empty_set_nil. assumption. intro h16. rewrite h16 in h0l'. symmetry in h0l'. apply empty_set_nil. assumption. constructor. exists (Build_nice_map_lists2 _ _ _ _ _ _ _ F la lb lc lp hn h6' h13' h13'' h3l h5l h0r h0r' h9' h3r h5r h11' h12'). constructor. assert (h6:list_to_set (list_prod nil nil) = cart_prod A B). simpl. assumption. pose proof (NoDup_nil T) as h7. pose proof (NoDup_nil U) as h8. assert (h9:NoDup (list_prod (@nil T) (@nil U))). simpl. constructor. assert (h10:lp = nil). symmetry in x. pose proof (cart_prod_empty_rev A B x) as h10. destruct h10 as [h10l | h10r]. generalize dependent F. rewrite h10l. intros F h11 h12. rewrite fin_map_cart_empty11 in h12. pose proof (empty_set_nil _ h12). subst. reflexivity. generalize dependent F. rewrite h10r. intros F h11 h12. rewrite fin_map_cart_empty21 in h12. pose proof (empty_set_nil _ h12). subst. reflexivity. assert (h11:In lp (list_power (list_prod nil nil) lc)). simpl. left. subst. reflexivity. assert (h12: functionally_paired_l (list_prod nil nil) lc lp). simpl. subst. apply fpl_empty1. assert (h13: if nil_dec' (list_prod (@nil T) (@nil U)) then ((@nil T) = nil /\ A = Empty_set _) \/ ((@nil U) = nil /\ B = Empty_set _) else (list_to_set nil = A /\ list_to_set nil = B)). simpl. destruct (nil_dec' nil) as [h13 | h14]. symmetry in x. pose proof (cart_prod_empty_rev _ _ x) as h14. destruct h14. left. split; auto. right. split; auto. contradict h14. reflexivity. assert (h13': if (nil_dec' (list_prod (@nil T) (@nil U))) then (A = Empty_set _ -> (@nil T) = nil) /\ (B = Empty_set _ -> (@nil U) = nil) else True). simpl. destruct (nil_dec' nil) as [h13' | h14']. tauto. constructor. assert (hn: Finite A -> Finite B -> list_to_set nil = A /\ list_to_set nil = B). tauto. exists (Build_nice_map_lists2 _ _ _ _ _ _ _ F nil nil lc lp hn h6 h13 h13' h3l h5l h7 h8 h9 h3r h5r h11 h12). constructor. (* Add *) pose proof (fin_map_fin_dom F) as h2. pose proof (fin_map_fin_ran F) as h3. pose proof (Add_intro2 _ A0 x0) as h4. rewrite x in h4. pose proof (Inhabited_intro _ _ _ h4) as h5. pose proof (cart_prod_fin_rev _ _ h2 h5) as h6. destruct h6 as [h6l h6r]. pose proof (finite_set_list_no_dup _ h6l) as h7. pose proof (finite_set_list_no_dup _ h6r) as h8. pose proof (finite_set_list_no_dup _ h3) as h9. destruct h7 as [la h7]. destruct h8 as [lb h8]. destruct h9 as [lc h9]. destruct h7 as [h7l h7r]. destruct h8 as [h8l h8r]. destruct h9 as [h9l h9r]. symmetry in h7l. symmetry in h8l. symmetry in h9l. pose proof (list_prod_cart_prod_compat A B _ _ h7l h8l) as h10. pose proof (no_dup_list_prod _ _ h7r h8r) as h11. pose proof (fin_map_to_list_pairs_lemma (impl_type_eq_dec_prod hdt hdu) hdv F _ lc h11 h10 h9l) as h12. destruct h12 as [lp h12]. destruct h12 as [h12a [h12b h12c]]. pose proof (in_list_power_synced _ _ _ h12c) as h13. pose proof (fpl_in_list_power (impl_type_eq_dec_prod hdt hdu) hdv _ lc _ h13 h11 h12b) as h7. pose proof h12c as h14. rewrite <- h7 in h14. assert (h15: if (nil_dec' (list_prod la lb)) then (la = nil /\ A = Empty_set _) \/ (lb = nil /\ B = Empty_set _) else list_to_set la = A /\ list_to_set lb = B). destruct (nil_dec' (list_prod la lb)) as [h16 | h17]. pose proof h5 as h15. rewrite <- h10 in h15. rewrite h16 in h15. simpl in h15. inversion h15. contradiction. split; assumption. assert (h15': if (nil_dec' (list_prod la lb)) then (A = Empty_set _ -> la = nil) /\ (B = Empty_set _ -> lb = nil) else True). destruct (nil_dec' (list_prod la lb)) as [h16 | h17]. pose proof h5 as h15'. rewrite <- h10 in h15'. rewrite h16 in h15'. simpl in h15'. inversion h15'. contradiction. constructor. assert (hn: Finite A -> Finite B -> list_to_set la = A /\ list_to_set lb = B). tauto. exists (Build_nice_map_lists2 _ _ _ _ _ _ _ F la lb lc lp hn h10 h15 h15' h9l h12a h7r h8r h11 h9r h12b h12c h14). constructor. Qed. Lemma fin_map_ex_nice_map_lists2_list_to_set_dom : forall {T U V:Type} {la:list T} {lb:list U} {C:Ensemble V} {def:V} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (F:Fin_map (cart_prod (list_to_set la) (list_to_set lb)) C def), NoDup la -> NoDup lb -> exists nml2:(nice_map_lists2 F), la = n_la2 _ nml2 /\ lb = n_lb2 _ nml2. intros T U V la lb C def hdt hdu hdv F ha hb. pose proof (fin_map_fin_ran F) as h1. pose proof (finite_set_list_no_dup _ h1) as h3. destruct h3 as [lc h3]. destruct h3 as [h3l h3r]. symmetry in h3l. assert (h4:list_to_set (list_prod la lb) = cart_prod (list_to_set la) (list_to_set lb)). apply list_prod_cart_prod_compat; auto. assert (h5: if nil_dec' (list_prod la lb) then la = nil /\ list_to_set la = Empty_set T \/ lb = nil /\ list_to_set lb = Empty_set U else list_to_set la = list_to_set la /\ list_to_set lb = list_to_set lb). destruct (nil_dec' (list_prod la lb)) as [h5 | h6]. assert (h6:list_to_set (list_prod la lb) = Empty_set _). rewrite h5. simpl. reflexivity. rewrite h4 in h6. pose proof (cart_prod_empty_rev _ _ h6) as h7. destruct h7 as [h7l | h7r]. left. split; auto. apply empty_set_nil. assumption. right. split; auto. apply empty_set_nil. assumption. split; auto. assert (h6:if nil_dec' (list_prod la lb) then (list_to_set la = Empty_set T -> la = nil) /\ (list_to_set lb = Empty_set U -> lb = nil) else True). destruct (nil_dec' (list_prod la lb)). split. intro h6. pose proof (empty_set_nil _ h6). assumption. intro h7. pose proof (empty_set_nil _ h7). assumption. constructor. pose proof (no_dup_list_prod _ _ ha hb) as h7. pose proof (fin_map_to_list_pairs_lemma (impl_type_eq_dec_prod hdt hdu) hdv F (list_prod la lb) _ h7 h4 h3l) as h3. destruct h3 as [lp h3]. destruct h3 as [h3a [h3b h3c]]. pose proof (in_list_power_synced _ _ _ h3c) as h4'. pose proof (fpl_in_list_power (impl_type_eq_dec_prod hdt hdu) hdv _ lc _ h4' h7 h3b) as h5'. pose proof h3c as h8. rewrite <- h5' in h8. assert (h9:Finite (list_to_set la) -> Finite (list_to_set lb) -> list_to_set la = list_to_set la /\ list_to_set lb = list_to_set lb). auto. exists (Build_nice_map_lists2 _ _ _ _ _ _ _ F la lb lc lp h9 h4 h5 h6 h3l h3a ha hb h7 h3r h3b h3c h8). simpl. split; auto. Qed. Lemma fin_map_ex_nice_map_lists2_dom : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (F:Fin_map (cart_prod A B) C def) (la:list T) (lb:list U), list_to_set la = A -> list_to_set lb = B -> NoDup la -> NoDup lb -> exists nml2:(nice_map_lists2 F), la = n_la2 _ nml2 /\ lb = n_lb2 _ nml2. intros T U V A B C def hdt hdu hdv F la lb hla hlb ha hb. pose proof (fin_map_fin_ran F) as h1. pose proof (finite_set_list_no_dup _ h1) as h3. destruct h3 as [lc h3]. destruct h3 as [h3l h3r]. symmetry in h3l. assert (h4:list_to_set (list_prod la lb) = cart_prod A B). apply list_prod_cart_prod_compat; auto. assert (h5: if nil_dec' (list_prod la lb) then la = nil /\ A = Empty_set T \/ lb = nil /\ B = Empty_set U else list_to_set la = A /\ list_to_set lb = B). destruct (nil_dec' (list_prod la lb)) as [h5 | h6]. assert (h6:list_to_set (list_prod la lb) = Empty_set _). rewrite h5. simpl. reflexivity. rewrite h4 in h6. pose proof (cart_prod_empty_rev _ _ h6) as h7. destruct h7 as [h7l | h7r]. left. split; auto. apply empty_set_nil. rewrite <- h7l. assumption. right. split; auto. apply empty_set_nil. rewrite <- h7r. assumption. split; auto. assert (h6:if nil_dec' (list_prod la lb) then (A = Empty_set T -> la = nil) /\ (B = Empty_set U -> lb = nil) else True). destruct (nil_dec' (list_prod la lb)). split. intro h6. rewrite h6 in hla. apply empty_set_nil. assumption. intro h6. rewrite h6 in hlb. apply empty_set_nil. assumption. constructor. pose proof (no_dup_list_prod _ _ ha hb) as h7. pose proof (fin_map_to_list_pairs_lemma (impl_type_eq_dec_prod hdt hdu) hdv F (list_prod la lb) _ h7 h4 h3l) as h3. destruct h3 as [lp h3]. destruct h3 as [h3a [h3b h3c]]. pose proof (in_list_power_synced _ _ _ h3c) as h4'. pose proof (fpl_in_list_power (impl_type_eq_dec_prod hdt hdu) hdv _ lc _ h4' h7 h3b) as h5'. pose proof h3c as h8. rewrite <- h5' in h8. assert (h9:Finite A -> Finite B -> list_to_set la = A /\ list_to_set lb = B). auto. exists (Build_nice_map_lists2 _ _ _ _ _ _ _ F la lb lc lp h9 h4 h5 h6 h3l h3a ha hb h7 h3r h3b h3c h8). simpl. split; auto. Qed. Lemma n_im_im_fin_map_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) (nml:nice_map_lists F), list_to_set (n_im F nml) = im_fin_map F. intros T U A B def F nml. unfold n_im. unfold im_fin_map. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h1. rewrite <- list_to_set_in_iff in h1. rewrite in_map_iff in h1. destruct h1 as [pr h1]. destruct h1 as [h1l h1r]. rewrite list_to_set_in_iff in h1r. rewrite (lp_compat F nml) in h1r. pose proof (fin_map_to_fps_compat F) as h2. subst. apply Im_intro with (fst pr). destruct h2 as [h1 h2]. destruct h1 as [h3 h4]. pose proof (h4 _ h1r) as h5. destruct h5 as [h5l h5r]; assumption. symmetry. apply fin_map_to_fps_fin_map_app_compat. assumption. (* >= *) red. intros y h1. destruct h1 as [x h2 y]. subst. rewrite <- list_to_set_in_iff. rewrite in_map_iff. exists (x, F |-> x). simpl. split. reflexivity. pose proof (lp_compat F nml) as h3. rewrite list_to_set_in_iff. rewrite h3. pose proof (fin_map_to_fps_compat F) as h4. destruct h4 as [h4 h5]. rewrite (fin_fps_to_f_compat' (fin_map_type_eq_dec F) (fin_map_fin_dom F) h4 def). constructor. simpl. split. rewrite h5. reflexivity. assumption. Qed. Lemma fpl_f_compat : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} {def:U} (F:Fin_map A B def) (nml:nice_map_lists F) (x:T) (pf:In x (n_la F nml)), let f := fpl_to_fun_in_dep (fpl F nml) in f x pf = F |-> x. intros T U A B def F nml x h5'. pose proof (fpl_to_fun_in_dep_compat (fpl F nml) x h5') as h4. destruct F as [hdt h1 h2 S h3]. simpl. pose proof (la_compat _ nml) as h6. pose proof h5' as h5. rewrite list_to_set_in_iff in h5. rewrite h6 in h5. pose proof (lp_compat _ nml) as h7. rewrite list_to_set_in_iff in h4. rewrite h7 in h4. rewrite <- fin_map_to_fps_s_compat in h4. rewrite (fin_fps_to_f_compat' hdt h1 h3 def) in h4 at 1. destruct h4 as [h4l h4r]. simpl in h4l. destruct h4l; assumption. Qed. Lemma fpl_f_compat_list_to_set : forall {T U:Type} (hdu:type_eq_dec U) (li:list T) (lj:list U) (def:U) (F:Fin_map (list_to_set li) (list_to_set lj) def) (nml:nice_map_lists F) (pf:In (n_lp F nml) (list_power li lj)) (pfi:NoDup li) (pfj:NoDup lj) (x:T) (pfin:In x li), let pffp := in_list_power_fpl (fin_map_type_eq_dec F) hdu li lj (n_lp F nml) pfi (list_power_no_dup li lj (n_lp F nml) pfi pfj pf) pf in let f := fpl_to_fun_in_dep pffp in f x pfin = F |-> x. intros T U hdu li lj def F nml h1 h2 h3 x hin hffp f. pose proof (fpl_to_fun_in_dep_compat hffp x hin) as h4. destruct F as [hdt h9 h10 S h11]. pose proof (lp_compat _ nml) as h6. rewrite list_to_set_in_iff in h4. rewrite h6 in h4. rewrite <- fin_map_to_fps_s_compat in h4. rewrite (fin_fps_to_f_compat' hdt h9 h11 def) in h4 at 1. destruct h4 as [h4l h4r]. simpl in h4l. destruct h4l; auto. Qed. Lemma fpl_f_compat_pseudo_list_to_set : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (F:Fin_map I J def) (li:list T) (lj:list U) (lp:list (T*U)) (pf:In lp (list_power li lj)) (pfi:NoDup li) (pfj:NoDup lj) (x:T) (pfin:In x li), list_to_set li = I -> list_to_set lj = J -> list_to_set lp = fin_map_to_fps F -> let pffp := in_list_power_fpl (fin_map_type_eq_dec F) pfdu li lj lp pfi (list_power_no_dup li lj lp pfi pfj pf) pf in let f := fpl_to_fun_in_dep pffp in f x pfin = F |-> x. intros T U I J def hdu F li lj lp h1 h2 h3 x hin h2' h3' h4' hfp f. pose proof (fpl_to_fun_in_dep_compat hfp x hin) as h4. destruct F. unfold fin_map_app. rewrite list_to_set_in_iff in h4 at 1. rewrite h4' in h4. rewrite <- fin_map_to_fps_s_compat in h4. rewrite (fin_fps_to_f_compat' t f0 f2 def) in h4. destruct h4 as [h4]. destruct h4 as [h4l h4r]. simpl in h4l. assumption. Qed. Lemma fpl_f_compat_pseudo_list_to_set' : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (F:Fin_map I J def) (li:list T) (lj:list U) (nml:nice_map_lists F) (pf:In (n_lp F nml) (list_power li lj)) (pfi:NoDup li) (pfj:NoDup lj) (x:T) (pfin:In x li), list_to_set li = I -> list_to_set lj = J -> let pffp := in_list_power_fpl (fin_map_type_eq_dec F) pfdu li lj (n_lp F nml) pfi (list_power_no_dup li lj (n_lp F nml) pfi pfj pf) pf in let f := fpl_to_fun_in_dep pffp in f x pfin = F |-> x. intros T U I J def hdu F li lj nml h1 h2 h3 x hin h2' h3' hfp f. pose proof (fpl_to_fun_in_dep_compat hfp x hin) as h4. destruct F. unfold fin_map_app. pose proof (lp_compat _ nml) as h6. rewrite list_to_set_in_iff in h4. rewrite h6 in h4. rewrite <- fin_map_to_fps_s_compat in h4. rewrite (fin_fps_to_f_compat' t f0 f2 def) in h4 at 1. destruct h4 as [h4]. destruct h4 as [h4l h4r]. simpl in h4l. destruct h4l; auto. Qed. Lemma list_to_set_n_la2 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (nml2:nice_map_lists2 F), Finite A -> list_to_set (n_la2 F nml2) = A. intros ? ? ? ? ? ? ? F nml hf. destruct F, nml; auto; simpl. destruct (nil_dec' n_lab3) as [hnil | hnil]. subst. apply list_prod_eq_nil in hnil. destruct hnil as [hnil | hnil]; subst. simpl in in_lp3. destruct in_lp3 as [h4 | h4]; try contradiction. subst. simpl. destruct lab_empty3 as [h4 | h4]. destruct h4; auto. destruct h4; subst. specialize (la_lb_compat3 hf (Empty_is_finite _)). simpl in la_lb_compat3. destruct la_lb_compat3; auto. simpl in in_lp3. rewrite list_prod_nil in fpl3, in_lp3. simpl in in_lp3. destruct in_lp3 as [h1 | h1]. subst. rewrite list_prod_nil in lab_compat3. simpl in lab_compat3. destruct lab_empty2'0 as [h4 h5]. pose proof (finite_inh_or_empty _ hf) as ha. destruct ha as [ha | ha]; subst. destruct lab_empty3 as [h10 | h10]. destruct h10; subst. destruct ha; contradiction. pose proof (cart_prod_fin_rev2 _ _ f ha) as hb. specialize (la_lb_compat3 hf hb). destruct la_lb_compat3; auto. specialize (h4 eq_refl). subst. reflexivity. contradiction. destruct lab_empty3; auto. Qed. Lemma list_to_set_n_lb2 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (nml2:nice_map_lists2 F), Finite B -> list_to_set (n_lb2 F nml2) = B. intros ? ? ? ? ? ? ? F nml hf. destruct F, nml; auto; simpl. destruct (nil_dec' n_lab3) as [hnil | hnil]. subst. apply list_prod_eq_nil in hnil. destruct hnil as [hnil | hnil]; subst. simpl in in_lp3. destruct in_lp3 as [h4 | h4]; try contradiction. subst. simpl. destruct lab_empty3 as [h4 | h4]. destruct h4; auto. subst. specialize (la_lb_compat3 (Empty_is_finite _) hf). simpl in la_lb_compat3. destruct la_lb_compat3; auto. destruct h4; subst; auto. simpl in in_lp3. rewrite list_prod_nil in fpl3, in_lp3. simpl in in_lp3. destruct in_lp3 as [h1 | h1]. subst. rewrite list_prod_nil in lab_compat3. simpl in lab_compat3. destruct lab_empty2'0 as [h4 h5]. destruct lab_empty3 as [h8 | h8]. destruct h8; auto. subst. specialize (la_lb_compat3 (Empty_is_finite _) hf). destruct la_lb_compat3; auto. pose proof (finite_inh_or_empty _ hf) as hb. destruct hb as [hb | hb]; subst. pose proof (cart_prod_fin_rev1 _ _ f hb) as hc. specialize (la_lb_compat3 hc hf). destruct la_lb_compat3; auto. reflexivity. contradiction. destruct lab_empty3; auto. Qed. Lemma list_to_set_n_lp2_eq : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (nml2:nice_map_lists2 F), list_to_set (n_lp2 F nml2) = fin_map_to_fps F. intros ? ? ? ? ? ? ? F nml. destruct F, nml; auto. Qed. Lemma impl_in_n_la2 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (nml2:nice_map_lists2 F), Finite A -> forall (x:T), Inn A x -> In x (n_la2 F nml2). intros; rewrite list_to_set_in_iff; rewrite list_to_set_n_la2; auto. Qed. Lemma impl_in_n_lb2 : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (nml2:nice_map_lists2 F), Finite B -> forall (y:U), Inn B y -> In y (n_lb2 F nml2). intros; rewrite list_to_set_in_iff; rewrite list_to_set_n_lb2; auto. Qed. Lemma fpl2_f_compat : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V} (F:Fin_map (cart_prod A B) C def) (nml:nice_map_lists2 F) (pr:T*U) (pfin:In pr (n_lab2 F nml)), let pffp := fpl2 F nml in let f := fpl_to_fun_in_dep pffp in f pr pfin = F |-> pr. intros T U V A B C def F nml pr hin hfp f. unfold fin_map_app. pose proof (fpl_to_fun_in_dep_compat hfp pr hin) as h0. rewrite list_to_set_in_iff in h0. pose proof (list_to_set_n_lp2_eq F nml) as heq. dependent destruction F. rewrite <- fin_map_to_fps_s_compat in heq at 1. rewrite heq in h0. pose proof @fp_inv. pose proof (fp_inv t _ _ _ def f _ f1 h0) as h2. simpl in h2. rewrite h2; auto. Qed. Lemma im1_im1l_compat : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V}, Finite A -> forall (F:Fin_map (cart_prod A B) C def) (nml:(nice_map_lists2 F)) (x:T) (pfin:In x (n_la2 F nml)), im1 F x = list_to_set (im1l (fpl2 F nml) x pfin). intros T U V A B C def h1 F nml x h2. apply Extensionality_Ensembles; red; split; red; intros y h3. destruct h3 as [h4 b h3]. subst. rewrite <- list_to_set_in_iff. unfold im1l. rewrite in_map_in_dep_iff. exists b. ex_goal. rewrite list_to_set_in_iff. rewrite list_to_set_n_lb2; auto. eapply cart_prod_fin_rev2. apply (fin_map_fin_dom F). econstructor. apply h4. exists hex. apply (fpl2_f_compat F nml). rewrite <- list_to_set_in_iff in h3. unfold im1l in h3. rewrite in_map_in_dep_iff in h3. destruct h3 as [b [h3 h4]]; subst. pose proof h2 as h2'. rewrite list_to_set_in_iff in h2'. rewrite list_to_set_n_la2 in h2'; auto. pose proof (Inhabited_intro _ _ _ h2') as hia. pose proof (cart_prod_fin_rev2 _ _ (fin_map_fin_dom F) hia) as hfb. pose proof h3 as h3'. rewrite list_to_set_in_iff in h3'. rewrite list_to_set_n_lb2 in h3'; auto. pose proof (Inhabited_intro _ _ _ h3') as hib. econstructor; auto. apply h3'. apply (fpl2_f_compat F nml). Qed. Lemma im2_im2l_compat : forall {T U V:Type} {A:Ensemble T} {B:Ensemble U} {C:Ensemble V} {def:V}, Finite B -> forall (F:Fin_map (cart_prod A B) C def) (nml:(nice_map_lists2 F)) (y:U) (pfin:In y (n_lb2 F nml)), im2 F y = list_to_set (im2l (fpl2 F nml) y pfin). intros T U V A B C def h1 F nml x h2. apply Extensionality_Ensembles; red; split; red; intros y h3. destruct h3 as [h4 b h3]. subst. rewrite <- list_to_set_in_iff. unfold im2l. rewrite in_map_in_dep_iff. exists b. ex_goal. rewrite list_to_set_in_iff. rewrite list_to_set_n_la2; auto. eapply cart_prod_fin_rev1. apply (fin_map_fin_dom F). econstructor. apply h4. exists hex. apply (fpl2_f_compat F nml). rewrite <- list_to_set_in_iff in h3. unfold im2l in h3. rewrite in_map_in_dep_iff in h3. destruct h3 as [a [h3 h4]]; subst. pose proof h2 as h2'. rewrite list_to_set_in_iff in h2'. rewrite list_to_set_n_lb2 in h2'; auto. pose proof (Inhabited_intro _ _ _ h2') as hib. pose proof (cart_prod_fin_rev1 _ _ (fin_map_fin_dom F) hib) as hfa. pose proof h3 as h3'. rewrite list_to_set_in_iff in h3'. rewrite list_to_set_n_la2 in h3'; auto. pose proof (Inhabited_intro _ _ _ h3') as hia. econstructor; auto. apply h3'. apply (fpl2_f_compat F nml). Qed. (*This is analogous to [fin_map_compose_inj], but it doesn't require injectivity. That is, it takes a Finite map over a [list_to_set l], and composes it with a given function in such a way that respects the injectivity/duplicate issues. It gets around this, by letting the user specify a "mask" in [linds_prod] that corresponds to the specific "pre-image choices" from [Im A g]. Specifically, each element in [linds_prod] is a list of indices that will serve as pre-images for a given element in the domain (an [Im] set).*) Lemma fin_map_list_compose_linds_prod_ex : forall {T U V:Type} {l:list T} {B:Ensemble V} {defv:V} (pfdu:type_eq_dec U) (F:Fin_map (list_to_set l) B defv) (g:T->U) (ln:list nat) (pfln:In ln (linds_prod pfdu (map g l))), exists! F':Fin_map (Im (list_to_set l) g) B defv, forall (x:T) (pfx:In x l), let pflt := lt_eq_trans _ _ _ (in_l_domain_linds_prod_ind_lt pfdu _ _ (in_map g l x pfx) ln pfln) (map_length _ _) in let x' := nth_lt _ _ pflt in F' |-> g x = F |-> x'. intros T U V l B defV h2 F g ln h4. assert (h5:forall y:U, Inn (Im (list_to_set l) g) y -> In y (map g l)). intros y h6. rewrite <- map_im_compat in h6. rewrite list_to_set_in_iff; auto. pose (fun (y:sig_set (Im (list_to_set l) g)) => (in_l_domain_linds_prod_ind_in h2 _ (proj1_sig y) (h5 _ (proj2_sig y)) ln h4)) as h7. pose (fun (y:sig_set (Im (list_to_set l) g)) => let pflt := lt_eq_trans _ _ _ (in_l_domain_linds_prod_ind_lt h2 _ _ (h5 _ (proj2_sig y)) ln h4) (map_length _ _) in let x := nth_lt _ _ pflt in F |-> x) as f. pose proof (finite_image _ _ _ g (list_to_set_finite l)) as h8. pose proof @sig_fun_to_fin_map_ran. pose (sig_fun_to_fin_map_ran h2 f h8 defV B (fin_map_fin_ran F)) as h9. hfirst h9. red; intros x h10. destruct h10 as [x h10]; subst. unfold f. apply fin_map_app_in. rewrite <- list_to_set_in_iff. apply in_nth_lt. pose (h9 hf) as F'. exists F'. red. split. intros x h10. unfold F', h9. erewrite sig_fun_to_fin_map_ran_compat. unfold f. simpl. repeat f_equal. apply nth_lt_functional2. repeat f_equal. apply proof_irrelevance. intros G h10. apply fin_map_ext. intros y h11. destruct h11 as [x h11]; subst. rewrite <- list_to_set_in_iff in h11. rewrite (h10 _ h11). unfold F', h9. erewrite sig_fun_to_fin_map_ran_compat. unfold f. repeat f_equal. apply nth_lt_functional2. repeat f_equal. simpl. apply proof_irrelevance. Grab Existential Variables. econstructor. rewrite <- list_to_set_in_iff. apply h11. reflexivity. econstructor. rewrite <- list_to_set_in_iff. apply h10. reflexivity. Qed. Definition fin_map_list_compose_linds_prod {T U V:Type} {l:list T} {B:Ensemble V} {defv:V} (pfdu:type_eq_dec U) (F:Fin_map (list_to_set l) B defv) (g:T->U) (ln:list nat) (pfln:In ln (linds_prod pfdu (map g l))) := proj1_sig (constructive_definite_description _ (fin_map_list_compose_linds_prod_ex pfdu F g ln pfln)). Lemma fin_map_list_compose_linds_prod_compat : forall {T U V:Type} {l:list T} {B:Ensemble V} {defv:V} (pfdu:type_eq_dec U) (F:Fin_map (list_to_set l) B defv) (g:T->U) (ln:list nat) (pfln:In ln (linds_prod pfdu (map g l))) (x:T) (pfx:In x l), let pflt := lt_eq_trans _ _ _ (in_l_domain_linds_prod_ind_lt pfdu _ _ (in_map g l x pfx) ln pfln) (map_length _ _) in let x' := nth_lt _ _ pflt in let F' := fin_map_list_compose_linds_prod pfdu F g ln pfln in F' |-> g x = F |-> x'. unfold fin_map_list_compose_linds_prod; intros; destruct constructive_definite_description; simpl; auto. Qed. (*This is a more specific version of [fin_map_list_compose_linds_prod_compat].*) Definition linds_prod_fin_map_seg {T U:Type} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (E:Ensemble T) (i_E:nat -> sig_set E) (n:nat) (F:Fin_map (seg n) B def) (ln:list nat) (pf:In ln (linds_prod pfdt (map (fun c => proj1_sig (i_E c)) (seg_list n)))) : Fin_map (Im (seg n) (fun i => proj1_sig (i_E i))) B def := let pfe := eq_sym (list_to_set_seg_list_eq n) in let F0 := fin_map_dom_subst pfe F in let F1 := fin_map_list_compose_linds_prod pfdt F0 (fun c => proj1_sig (i_E c)) ln pf in let pfe' := eq_sym (im_functional1 pfe (fun i => proj1_sig (i_E i))) in fin_map_dom_subst pfe' F1. Lemma linds_prod_fin_map_seg_compat : forall {T U:Type} {B:Ensemble U} {def:U} (pfdt:type_eq_dec T) (E:Ensemble T) (i_E:nat -> sig_set E) (n:nat) (F:Fin_map (seg n) B def) (ln:list nat) (pfln:In ln (linds_prod pfdt (map (fun c => proj1_sig (i_E c)) (seg_list n)))) (i:nat) (pflt:i < n), let F' := linds_prod_fin_map_seg pfdt E i_E n F ln pfln in let i' := in_l_domain_linds_prod_ind pfdt (map (fun c : nat => proj1_sig (i_E c)) (seg_list n)) (proj1_sig (i_E i)) (in_map (fun c : nat => proj1_sig (i_E c)) (seg_list n) i (iff2 (in_seg_list_iff i n) pflt)) ln pfln in F' |-> proj1_sig (i_E i) = F |-> i'. Proof. intros T U B def h1 E i_E n F ln h2 i h3. simpl. unfold linds_prod_fin_map_seg. erewrite <- fin_map_dom_subst_compat. pose proof (fin_map_list_compose_linds_prod_compat h1 (fin_map_dom_subst (eq_sym (list_to_set_seg_list_eq n)) F) (fun c : nat => proj1_sig (i_E c)) ln h2 i) as h4. hfirst h4. rewrite in_seg_list_iff; auto. specialize (h4 hf). simpl in h4. rewrite h4. rewrite nth_lt_seg_list. erewrite <- fin_map_dom_subst_compat. repeat f_equal. apply proof_irrelevance. constructor. eapply lt_eq_trans. apply in_l_domain_linds_prod_ind_lt. rewrite map_length. rewrite length_seg_list; auto. econstructor. rewrite <- list_to_set_in_iff. rewrite in_seg_list_iff. apply h3. reflexivity. Qed. bool2-v0-3/Permutations.v0000644000175000017500000024614313575574055016235 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2019, Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (* This file correspdons to Chapter 2 in Part II of Dietlinde Lau's "Function Algebras." The subject is "The Galois Connection between function- and relation- algebras." *) Require Import ProofIrrelevance. Require Import FunctionalExtensionality. Require Import NPeano. Require Import EqDec. Require Import FunctionProperties. Require Import FunctionProperties2. Require Import ArithUtilities. Require Import LogicUtilities. Require Import SetUtilities. Require Import ListUtilities. (*Returns the function that maps j to k, k to j, and fixes all else. A very complicated function for a simple transposition. The more obvious way to do it by destructing on boolean equality requires two passes, and perhaps is less efficient, although I include that afterwards too. *) Fixpoint transpose_seg_fun {n j k} : j < n -> k < n -> seg_fun n nat := match n with | 0 => fun pf => False_rect _ (lt_n_0 _ pf) | S n' => match j with | 0 => match k with | 0 => fun _ _ i _ => i | S k' => fun _ pfk i => let pfk' := lt_S_n _ _ pfk in match i with | 0 => fun _ => S k' | S i' => match i' with | 0 => fun _ => match k' with | 0 => 0 | S _ => 1 end | S i'' => fun pfi => let pfi' := lt_S_n _ _ pfi in let pf0 := Olt _ _ pfi' in let t := transpose_seg_fun pf0 pfk' (S i'') pfi' in match t with | 0 => 0 | S _ => S (S i'') end end end end | S j' => match k with | 0 => fun pfj _ i => let pfj' := lt_S_n _ _ pfj in match i with | 0 => fun _ => S j' | S i' => match i' with | 0 => fun _ => match j' with | 0 => 0 | S _ => 1 end | S i'' => fun pfi => let pfi' := lt_S_n _ _ pfi in let pf0 := Olt _ _ pfi' in let t := transpose_seg_fun pfj' pf0 (S i'') pfi' in match t with | 0 => 0 | S _ => S (S i'') end end end | S k' => fun pfj pfk i => let pfj' := lt_S_n _ _ pfj in let pfk' := lt_S_n _ _ pfk in match i with | 0 => fun _ => 0 | S i' => fun pfi => let pfi' := lt_S_n _ _ pfi in let pf0 := Olt _ _ pfi' in let t := transpose_seg_fun pfj' pfk' i' pfi' in S t end end end end. (*Seems to be less efficient than above in regards to simplification.*) Definition transpose_seg_fun' n r s : seg_fun n nat := fun i (pfi:i < n) => if (beq_nat i r) then s else if (beq_nat i s) then r else i. Lemma transpose_seg_fun_functional : forall {n r r' s s'} (pfr:r < n) (pfr':r' < n) (pfs:s < n) (pfs':s' < n), r = r' -> s = s' -> transpose_seg_fun pfr pfs = transpose_seg_fun pfr' pfs'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma transpose_seg_fun_functional' : forall {n r r' s s' i i'} (pfr:r < n) (pfr':r' < n) (pfs:s < n) (pfs':s' < n) (pfi: i < n) (pfi':i' < n), r = r' -> s = s' -> i = i' -> transpose_seg_fun pfr pfs i pfi = transpose_seg_fun pfr' pfs' i' pfi'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma transpose_seg_fun_comm : forall {n r s} (pfr:r < n) (pfs:s < n), transpose_seg_fun pfr pfs = transpose_seg_fun pfs pfr. intro n. induction n as [|n ih]. intros ? ? h1; contradict (lt_n_0 _ h1). intro r. destruct r as [|r]. intro s. destruct s as [|s]. intros hr hs. assert (hr = hs). apply proof_irrelevance. subst; auto. intros hr hs. simpl. apply functional_extensionality_dep. intro i. destruct i as [|i]; auto. destruct i as [|i]; auto. apply functional_extensionality. intro hi. rewrite ih; auto. intro s. destruct s as [|s]. intros hr hs. simpl. apply functional_extensionality_dep. intro i. destruct i as [|i]; auto. destruct i as [|i]; auto. apply functional_extensionality. intro; rewrite ih; auto. intros hr hs. simpl. apply functional_extensionality_dep. intro i. destruct i; auto. apply functional_extensionality. intro hi. pose proof (lt_S_n _ _ hi) as hi'. specialize (ih r s (lt_S_n _ _ hr) (lt_S_n _ _ hs)). rewrite ih; auto. Qed. Lemma transpose_seg_fun_same : forall {n r} (pfr pfr':r < n), transpose_seg_fun pfr pfr' = id_seg n. intro n. induction n as [|n ih]; simpl. intros ? h1; contradict (lt_n_0 _ h1). intros r hr. destruct r; auto. intro hr'. assert (hr = hr'). apply proof_irrelevance. subst. apply functional_extensionality_dep. intro i. destruct i; auto. apply functional_extensionality. intro hi. unfold id_seg. f_equal. rewrite ih. unfold id_seg; auto. Qed. Lemma transpose_seg_fun_l : forall {n j k} (pfj pfj':j < n) (pfk:k < n), transpose_seg_fun pfj pfk j pfj' = k. intro n. induction n as [|n ih]. intros ? ? h1. contradict (lt_n_0 _ h1). intro j. destruct j as [|j]. intro k. destruct k as [|k]; auto. intro k; destruct k as [|k]; simpl. intros hj hk. simpl in ih. intro h0. destruct j as [|j]; auto. rewrite ih; auto. intros hj hk. f_equal; auto. Qed. Lemma transpose_seg_fun_r : forall {n j k} (pfj:j < n) (pfk pfk':k < n), transpose_seg_fun pfj pfk k pfk' = j. intros; rewrite transpose_seg_fun_comm, transpose_seg_fun_l; auto. Qed. Lemma transpose_seg_fun_neq : forall {n j k} (pfj:j < n) (pfk pfk':k < n), forall i (pfi:i < n), i <> j -> i <> k -> transpose_seg_fun pfj pfk i pfi = i. intro n. induction n as [|n ih]; simpl. intros; omega. intro j. destruct j as [|j]. intro k. destruct k as [|k]. intros; auto. intros h0 hk hk' i hi. destruct i as [|i]. intros; omega. intros he he1. destruct i as [|i]. destruct k as [|k]. contradict he1; auto. reflexivity. rewrite ih; auto with arith. intro k. destruct k as [|k]. intros hk hk' h1 i hi h2 h3. destruct i as [|i]. contradict h3; auto. destruct i as [|i]. destruct j as [|j]. contradict h2; auto. reflexivity. rewrite ih; auto. omega. intros hj hk hk' i h1 h2 h3. destruct i as [|i]; auto. f_equal; auto. rewrite ih; auto. apply lt_S_n in hk'; auto. Qed. Lemma transpose_seg_fun_compat : forall {n j k} (pfj:j < n) (pfk:k < n), let f := transpose_seg_fun pfj pfk in f j pfj = k /\ f k pfk = j /\ forall i (pfi:i < n), i <> j -> i <> k -> f i pfi = i. intros n j k hj hk. simpl. split. rewrite transpose_seg_fun_l; auto. split. rewrite transpose_seg_fun_r; auto. apply transpose_seg_fun_neq; auto. Qed. Lemma tranpose_seg_fun_compat' : forall {n r s} (pfr:r < n) (pfs:s < n) i (pfi: i < n), transpose_seg_fun pfr pfs i pfi = transpose_seg_fun' n r s i pfi. intros n r s hr hs i hi. unfold transpose_seg_fun'. destruct (nat_eq_dec i r) as [h1 | h1]; subst. rewrite beq_nat_same, transpose_seg_fun_l; auto. pose proof h1 as h1'. rewrite <- neq_nat_compat in h1. rewrite h1. destruct (nat_eq_dec i s) as [h2 | h2]; subst. rewrite beq_nat_same, transpose_seg_fun_r; auto. pose proof h2 as h2'. rewrite <- neq_nat_compat in h2. rewrite h2. rewrite transpose_seg_fun_neq; auto. Qed. Lemma transpose_same_seg_fun_iff : forall {n r i m} (pfr:r < n) (pfi:i < n), transpose_seg_fun pfr pfr i pfi = m <-> i = m. intro n. induction n as [|n ih]. intros ? ? ? h1. contradict (lt_n_O _ h1). intros r i m. destruct r as [|r], i as [|i], m as [|m]; split; intro h1; auto with arith; simpl in h1; try discriminate; try apply S_inj in h1; f_equal; auto; try rewrite ih in h1; auto; subst. pose proof (lt_S_n _ _ pfr) as hr. pose proof (lt_S_n _ _ pfi) as hi. simpl; f_equal. rewrite ih; auto. Qed. Lemma lt_transpose_seg_fun : forall {n r s i} (pfr:r < n) (pfs:s < n) (pf:i < n), transpose_seg_fun pfr pfs i pf < n. intros n r s i h1 h2. pose proof (transpose_seg_fun_compat h1 h2) as h4. simpl in h4. destruct h4 as [h4 [h5 h6]]. destruct (eq_nat_dec i r) as [h7 | h7]. rewrite h7. intro h1'. assert (h8:h1 = h1'). apply proof_irrelevance. rewrite <- h8, h4; auto. destruct (eq_nat_dec i s) as [h8 | h8]. rewrite h8. intro h2'. assert (h9:h2 = h2'). apply proof_irrelevance. rewrite <- h9, h5; auto. intros; rewrite h6; auto. Qed. Lemma transpose_seg_fun_S : forall {n r s i} (pfr:r < n) (pfr':r < S n) (pfs:s < n) (pfs':s < S n) (pfi:i < n) (pfi':i < S n), transpose_seg_fun pfr' pfs' i pfi' = transpose_seg_fun pfr pfs i pfi. intro n; induction n as [|n ih]. intros; omega. intros r s i hr hs hi hi' hr' hs'. destruct r as [|r], s as [|s], i as [|i]; auto. pose proof hs' as h1. apply lt_S_n in h1. destruct (nat_eq_dec i s) as [|h2]; subst. rewrite transpose_seg_fun_r. rewrite transpose_seg_fun_r; auto. rewrite transpose_seg_fun_neq; auto. rewrite transpose_seg_fun_neq; auto. destruct (nat_eq_dec i r) as [|h2]; subst. rewrite transpose_seg_fun_l. rewrite transpose_seg_fun_l; auto. rewrite transpose_seg_fun_neq; auto. rewrite transpose_seg_fun_neq; auto. destruct (nat_eq_dec i r) as [|h2]; subst. rewrite transpose_seg_fun_l. rewrite transpose_seg_fun_l; auto. destruct (nat_eq_dec i s) as [|h2']; subst. rewrite transpose_seg_fun_r. rewrite transpose_seg_fun_r; auto. rewrite transpose_seg_fun_neq; auto. rewrite transpose_seg_fun_neq; auto. Qed. Lemma transpose_seg_fun_Sn : forall {n s} (pfn pfn':n < S n) (pfs:s < S n), transpose_seg_fun pfn pfs n pfn' = s. intros; rewrite transpose_seg_fun_l; auto. Qed. Lemma transpose_seg_fun_Sn' : forall {n s} (pfn pfn':n < S n) (pfr:s < S n), transpose_seg_fun pfr pfn' n pfn' = s. intros; rewrite transpose_seg_fun_r; auto. Qed. Lemma transpose_seg_fun0_eq_iff : forall {n s m a} (pf0:0 < n) (pfm:m < n) (pfs:s < n), transpose_seg_fun pf0 pfs m pfm = a <-> (a = 0 /\ s = m) \/ (a = s /\ m = 0) \/ (a = m /\ m <> 0 /\ m <> s). intro n; induction n as [|n ih]. intros ? ? ? h1; contradict (lt_irrefl _ h1). intros s m a; destruct s as [|s], m as [|m], a as [|a]; intros h1 h2 h3; split; intro h4; subst; try tauto; try rewrite tranpose_seg_fun_l in h4; try rewrite transpose_seg_fun_r in h4; try discriminate; try destruct h4 as [[? ?] | [[? ?] | [? ?]]]; try discriminate. right; right; auto with arith. apply S_inj in H; subst. rewrite transpose_seg_fun_same; auto. destruct H0 as [h4]; contradict h4; auto. rewrite transpose_seg_fun_l in h4. apply S_inj in h4; subst. right; left; auto. apply S_inj in H; subst. rewrite transpose_seg_fun_l; auto. simpl in h4. destruct m as [|m], s as [|s]. left; auto. discriminate. rewrite transpose_seg_fun_same in h4. simpl in h4; discriminate. pose proof (lt_S_n _ _ h2) as h5. pose proof (lt_S_n _ _ h3) as h6. destruct (zerop (transpose_seg_fun (Olt _ _ (lt_S_n _ _ h2)) (lt_S_n _ _ h3) (S m) (lt_S_n _ _ h2))) as [h7 | h7]. rewrite h7 in h4. rewrite ih in h7. destruct h7 as [h7 | [h7 | h7]]. destruct h7 as [? h7]. apply S_inj in h7; subst. left; auto. destruct h7; discriminate. destruct h7; discriminate. rewrite (S_pred _ _ h7) in h4. discriminate. apply S_inj in H0; subst. simpl. destruct m as [|m]; auto. rewrite transpose_seg_fun_r; auto. simpl in h4. destruct m as [|m], s as [|s]. discriminate. apply S_inj in h4; subst. right; right; auto with arith. destruct (zerop (transpose_seg_fun (Olt (S m) n (lt_S_n (S m) n h2)) (lt_S_n 0 n h3) (S m) (lt_S_n (S m) n h2))) as [h5 | h5]. rewrite h5 in h4. discriminate. rewrite (S_pred _ _ h5) in h4. apply S_inj in h4; subst. right; right; auto. destruct (zerop (transpose_seg_fun (Olt (S m) n (lt_S_n (S m) n h2)) (lt_S_n (S s) n h3) (S m) (lt_S_n (S m) n h2))) as [h5 | h5]. rewrite h5 in h4. discriminate. rewrite (S_pred _ _ h5) in h4. apply S_inj in h4; subst. destruct (nat_eq_dec s m) as [|h6]; subst. rewrite transpose_seg_fun_r in h5. contradict (lt_irrefl _ h5). right; right; auto. apply S_inj in H; subst. rewrite transpose_seg_fun_neq; auto. destruct H0; auto. Qed. (*Wyckoff*) Lemma inj_transpose_seg_fun : forall {n r s} (pfr:r < n) (pfs:s < n), inj_seg (transpose_seg_fun pfr pfs). intros n r s hr hs; red; intros i j h1 h2 h3. destruct (eq_nat_dec i r) as [h4 | h4], (eq_nat_dec j r) as [h5 | h5]. subst; auto. destruct (eq_nat_dec j s) as [h6 | h6]. subst; auto. rewrite transpose_seg_fun_l, transpose_seg_fun_r in h3; auto. subst. rewrite transpose_seg_fun_l, transpose_seg_fun_neq in h3; auto. contradict (h6 (eq_sym h3)). subst. rewrite transpose_seg_fun_l in h3. destruct (eq_nat_dec i s) as [h5 | h5]. revert h3. generalize h1. rewrite h5. intros hs' h6. rewrite transpose_seg_fun_r in h6; subst. contradict h4; auto. rewrite transpose_seg_fun_neq in h3; try subst; auto. contradict h5; auto. destruct (nat_eq_dec j s) as [h6 | h6]; subst. rewrite transpose_seg_fun_r in h3. destruct (eq_nat_dec i s) as [h6 | h6]; auto. rewrite transpose_seg_fun_neq in h3; try subst; auto. contradict h4; auto. destruct (eq_nat_dec i s) as [h9 | h9]; subst. rewrite transpose_seg_fun_r, transpose_seg_fun_neq in h3; try subst; auto. contradict h5; auto. rewrite transpose_seg_fun_neq, transpose_seg_fun_neq in h3; subst; auto. Qed. Lemma transpose_seg_fun_endo : forall {n r s} (pfr:r < n) (pfs:s < n), seg_fun_endo (transpose_seg_fun pfr pfs) n. intros; constructor; intros; apply lt_transpose_seg_fun. Qed. Definition perm_seg {n} (f:seg_fun n nat) : Prop := inj_seg_fun_endo f n. Lemma transpose_seg_fun_perm : forall {n r s} (pfr:r < n) (pfs:s < n), perm_seg (transpose_seg_fun pfr pfs). intros; constructor; intros. apply transpose_seg_fun_endo. eapply inj_transpose_seg_fun. apply H. Qed. Lemma perm_seg_to_S : forall {n} (f:seg_fun (S n) nat) (pf0:0 < S n), f 0 pf0 = n -> perm_seg f -> perm_seg (seg_fun_to_S f). intros n f h1 h2 h3. red in h3; red. inversion h3 as [? ? h5 h6]; revert h1 h2; subst. inversion h5 as [? ? h7]; subst. intros h1 h2. constructor. constructor. intros i hi. specialize (h7 (S i) (lt_n_S _ _ hi)). apply lt_S_dec in h7. destruct h7 as [|h7]; auto. rewrite <- h7 in h2. specialize (h6 0 (S i) h1 (lt_n_S _ _ hi) h2); discriminate. intros i j hi hj h8. unfold seg_fun_to_S in h8. apply h6 in h8; auto with arith. Qed. Lemma perm_seg_to_S' : forall {n} (f:seg_fun (S n) nat) (pf0:0 < S n), f 0 pf0 = n -> perm_seg (seg_fun_to_S f) -> perm_seg f. intros n f h0 h1 h2. red in h2; red. inversion h2 as [? ? h3 h4 h5 h6]. eapply inj_seg_fun_endo_to_S' in h2. apply h2. apply le_n_Sn. erewrite f_equal_dep. rewrite h1 at 1. apply lt_n_Sn. reflexivity. revert h1; revert h0; inversion h3 as [? ? h7 h8]; subst; intros h0 h1. intro h5. rewrite in_seg_fun_iff in h5. destruct h5 as [i [h5 h6]]. specialize (h7 i h5). rewrite h6 in h7. erewrite f_equal_dep in h7. rewrite h1 in h7 at 1. contradict (lt_irrefl _ h7). reflexivity. Qed. Lemma perm_seg_fun1_eq : forall (f:seg_fun 1 nat) i, perm_seg f -> forall (pf:i < 1), f i pf = 0. intros ? ? h1 ?; inversion h1; apply seg_fun_endo1_eq; auto. Qed. Lemma perm_seg_left_shift : forall n, perm_seg (left_shift_seg_fun n). intros; apply inj_seg_fun_endo_shift. Qed. Lemma perm_seg_right_shift : forall n, perm_seg (right_shift_seg_fun n). intros; apply inj_seg_fun_endo_shift. Qed. (*Wyckoff*) (*This transposes values of a given [seg_fun] at [j] and [k] -- and leaves all else fixed.*) Definition transpose_functor {U n j k} (f:seg_fun n U) (pfj:j < n) (pfk:k < n) : seg_fun n U := fun i pfi => if (beq_nat i j) then f k pfk else if (beq_nat i k) then f j pfj else f i pfi. Lemma transpose_functor_functional : forall {U} {n j j' k k'} (f:seg_fun n U) (pfj : j < n) (pfj': j' < n) (pfk : k < n) (pfk': k' < n), j = j' -> k = k' -> transpose_functor f pfj pfk = transpose_functor f pfj' pfk'. intros; subst; f_equal; try apply proof_irrelevance. Qed. Lemma transpose_functor_functional' : forall {U} {n j j' k k' i i'} (f:seg_fun n U) (pfj : j < n) (pfj': j' < n) (pfk : k < n) (pfk': k' < n) (pfi : i < n) (pfi': i' < n), j = j' -> k = k' -> i = i' -> transpose_functor f pfj pfk i pfi = transpose_functor f pfj' pfk' i' pfi'. intros; subst; f_equal; try apply proof_irrelevance. Qed. (*Wyckoff*) Lemma transpose_functor_same : forall {U:Type} {n:nat} (f:seg_fun n U) (r:nat) (pfr pfr':r < n), transpose_functor f pfr pfr' = f. unfold transpose_functor; intros; apply functional_extensionality_dep; intro x; apply functional_extensionality_dep; intro h1. destruct (eq_nat_dec x r) as [h2 | h2]; subst. rewrite beq_nat_same. f_equal; apply proof_irrelevance. rewrite <- neq_nat_compat in h2. rewrite h2; auto. Qed. (*Wyckoff*) Lemma inj_transpose_functor : forall {U:Type} {n:nat} (f:seg_fun n U) (r s:nat) (pfr:r < n) (pfs:s < n), inj_seg f -> inj_seg (transpose_functor f pfr pfs). intros U n f r s h1 h2 h3. red in h3; red. intros i j h4 h5 h6. unfold transpose_functor in h6. destruct (eq_nat_dec i r) as [h7|h7]; subst. rewrite beq_nat_same in h6. destruct (eq_nat_dec j r) as [h8|h8]; subst; auto. rewrite <- neq_nat_compat in h8. rewrite h8 in h6. destruct (eq_nat_dec j s) as [h9|h9]; subst; auto. rewrite beq_nat_same in h6. apply h3 in h6; auto. rewrite <- neq_nat_compat in h9. rewrite h9 in h6. apply h3 in h6; subst. rewrite beq_nat_same in h9; discriminate. rewrite <- neq_nat_compat in h7. rewrite h7 in h6. destruct (eq_nat_dec i s) as [h9|h9]; subst; auto. rewrite beq_nat_same in h6. destruct (eq_nat_dec j r) as [h10|h10]; subst; auto. rewrite beq_nat_same in h6. apply h3 in h6; auto. rewrite <- neq_nat_compat in h10. rewrite h10 in h6. destruct (eq_nat_dec j s) as [h11|h11]; subst; auto. rewrite <- neq_nat_compat in h11. rewrite h11 in h6. apply h3 in h6; subst. rewrite beq_nat_same in h10; discriminate. rewrite <- neq_nat_compat in h9. rewrite h9 in h6. destruct (eq_nat_dec j r) as [h11|h11]; subst; auto. rewrite beq_nat_same in h6. apply h3 in h6; subst. rewrite beq_nat_same in h9; discriminate. rewrite <- neq_nat_compat in h11. rewrite h11 in h6. destruct (eq_nat_dec j s) as [h12|h12]; subst; auto. rewrite beq_nat_same in h6. apply h3 in h6; subst. rewrite beq_nat_same in h7; discriminate. rewrite <- neq_nat_compat in h12. rewrite h12 in h6. apply h3 in h6; auto. Qed. (*Wyckoff*) Lemma inj_transpose_functor' : forall {U:Type} {n:nat} (f:seg_fun n U) (r s:nat) (pfr:r < n) (pfs:s < n), inj_seg (transpose_functor f pfr pfs) -> inj_seg f. intros U n f r s h1 h2 h3; red in h3; red; intros x y h4 h5 h6. destruct (eq_nat_dec x r) as [|h7], (eq_nat_dec y s) as [|h8]; subst. eapply h3. unfold transpose_functor. do 2 rewrite beq_nat_same. destruct (eq_nat_dec s r) as [h7 | h7]; subst. rewrite beq_nat_same; auto. rewrite <- neq_nat_compat in h7. rewrite h7. erewrite f_equal_dep. rewrite <- h6 at 1. f_equal. apply proof_irrelevance. reflexivity. destruct (eq_nat_dec r y) as [|h9]; subst; auto. apply neq_sym in h9. rewrite <- neq_nat_compat in h9. specialize (h3 s y h2 h5). unfold transpose_functor in h3. rewrite h9, beq_nat_same in h3. rewrite <- neq_nat_compat in h8. rewrite h8 in h3. destruct (nat_eq_dec s r) as [|hsr]; subst. rewrite beq_nat_same in h3. apply h3. erewrite f_equal_dep. rewrite h6 at 1; auto. reflexivity. rewrite <- neq_nat_compat in hsr. rewrite hsr in h3. erewrite f_equal_dep in h6. apply h3 in h6; subst. rewrite beq_nat_same in h8; discriminate. reflexivity. rewrite <- neq_nat_compat in h7. specialize (h3 x r h4 h1). unfold transpose_functor in h3. rewrite h7, beq_nat_same in h3. destruct (eq_nat_dec x s) as [|hxs]; subst; auto. rewrite <- neq_nat_compat in hxs. rewrite hxs in h3. hfirst h3. rewrite h6. f_equal; apply proof_irrelevance. apply h3 in hf; subst. rewrite beq_nat_same in h7; discriminate. rewrite <- neq_nat_compat in h7, h8. destruct (eq_nat_dec y r) as [|hyr], (eq_nat_dec x s) as [|hxs]; subst; auto. specialize (h3 r s h1 h4). unfold transpose_functor in h3. rewrite beq_nat_same, h7, beq_nat_same in h3. rewrite h3; auto. erewrite f_equal_dep. rewrite h6 at 1. f_equal; apply proof_irrelevance. reflexivity. rewrite <- neq_nat_compat in hxs. specialize (h3 x s h4 h2). unfold transpose_functor in h3. rewrite h7, hxs, beq_nat_same in h3. rewrite beq_nat_sym in h8. rewrite h8 in h3. rewrite h6 in h3. hfirst h3. f_equal; apply proof_irrelevance. specialize (h3 hf); subst. rewrite beq_nat_same in hxs; discriminate. rewrite <- neq_nat_compat in hyr. specialize (h3 r y h1 h5). unfold transpose_functor in h3. rewrite beq_nat_sym in h7. hfirst h3. rewrite beq_nat_same, hyr, h8. rewrite <- h6. f_equal; apply proof_irrelevance. specialize (h3 hf); subst. rewrite beq_nat_same in hyr; discriminate. rewrite <- neq_nat_compat in hyr. rewrite <- neq_nat_compat in hxs. specialize (h3 x y h4 h5). unfold transpose_functor in h3. hfirst h3. rewrite h7, h8, hyr, hxs; auto. auto. Grab Existential Variables. assumption. assumption. Qed. Definition transposed_seg_list {n r s} (pfr:r < n) (pfs:s < n) : list nat := map_seg (transpose_seg_fun pfr pfs). Lemma in_transposed_seg_list_iff : forall {n r s i} (pfr:r < n) (pfs:s < n), (In i (transposed_seg_list pfr pfs) <-> i < n). unfold transposed_seg_list; intros n r s i hr hs; split; intro h1. rewrite in_map_seg_iff in h1. destruct h1 as [m [h1 ]]; subst. apply lt_transpose_seg_fun. rewrite in_map_seg_iff. destruct (nat_eq_dec i r) as [|hir], (nat_eq_dec i s) as [|his]; subst; subst. exists _, hs. rewrite transpose_seg_fun_r; auto. exists _, hs. rewrite transpose_seg_fun_r; auto. exists _, hr. rewrite transpose_seg_fun_l; auto. exists _, h1. rewrite transpose_seg_fun_neq; auto. Qed. Lemma length_transposed_seg_list : forall {n r s} (pfr:r < n) (pfs:s < n), length (transposed_seg_list pfr pfs) = n. unfold transposed_seg_list; intros; rewrite length_map_seg; auto. Qed. (*Shortened proof than less efficient version using old [transpose_seg_fun]. Proof of efficiency in applications on constants!*) Lemma transposed_seg_list_01 : transposed_seg_list lt_0_2 lt_1_2 = 1::0::nil. unfold transposed_seg_list, transpose_seg_fun. rewrite map_seg_eq_iff. simpl. exists eq_refl. intros m h1. destruct m as [|m]; auto. destruct m as [|m]; auto. destruct m as [|m]; omega. Qed. Lemma transposed_seg_list_0n : forall n (pf:0 < n), transposed_seg_list (lt_0_Sn n) (lt_n_Sn n) = n::seq 1 (pred n) ++ 0::nil. intros n hn. unfold transposed_seg_list. rewrite map_seg_eq_iff. ex_goal. simpl; rewrite app_length, seq_length; simpl. omega. exists hex. intros m hm h1. unfold h1. clear h1. destruct (zerop m) as [|h2]; subst. rewrite transpose_seg_fun_l. symmetry. gen0. rewrite cons_app. intro h1. erewrite nth_lt_app1. simpl; auto. pose proof hm as h3. apply le_lt_eq_dec in h3. destruct h3 as [h3 | h3]. apply lt_S_n in h3. symmetry. gen0. rewrite cons_app. intro. erewrite nth_lt_app1. rewrite transpose_seg_fun_neq; auto with arith. simpl. destruct m as [|m]. contradict (lt_n_0 _ h2). erewrite nth_indep. rewrite <- nth_lt_compat. rewrite nth_lt_seq; auto with arith. rewrite seq_length; auto with arith. intro; subst; contradict (lt_irrefl _ h2). intro; subst; contradict (lt_irrefl _ h3). apply S_inj in h3; subst. rewrite transpose_seg_fun_r. symmetry. gen0. rewrite cons_app. intro. erewrite nth_lt_app2; simpl. rewrite seq_length. rewrite <- (S_pred _ _ h2), minus_same; auto. Grab Existential Variables. simpl; rewrite seq_length, <- (S_pred _ _ h2); auto with arith. rewrite seq_length. omega. simpl; rewrite seq_length, <- (S_pred _ _ h3); auto. simpl; rewrite seq_length, <- (S_pred _ _ hn); auto. Qed. Lemma transposed_seg_list_app_sing : forall (n r s:nat) (pfr:r < n) (pfs:s < n), let pfr' := lt_trans _ _ _ pfr (lt_n_Sn _) in let pfs' := lt_trans _ _ _ pfs (lt_n_Sn _) in transposed_seg_list pfr' pfs' = transposed_seg_list pfr pfs ++ n :: nil. unfold transposed_seg_list; intros n r s h1 h2. rewrite map_seg_eq_iff. ex_goal. rewrite app_length, length_map_seg. simpl. rewrite S_compat; auto. exists hex. intros m h3. intro h4. unfold h4. clear h4. pose proof h3 as ha. apply le_lt_eq_dec in ha. destruct ha as [ha | ha]. apply lt_S_n in ha. erewrite nth_lt_app1. rewrite nth_lt_map_seg. simpl. destruct r as [|r], s as [|s]. rewrite transpose_seg_fun_same; auto. destruct m as [|m]. rewrite transpose_seg_fun_l; auto. destruct m as [|m], s as [|s]. rewrite transpose_seg_fun_r; auto. rewrite transpose_seg_fun_neq; auto. rewrite transpose_seg_fun_same; simpl. rewrite transpose_seg_fun_neq; auto. destruct (nat_eq_dec m s) as [|h4]; subst. do 2 rewrite transpose_seg_fun_r; auto. rewrite transpose_seg_fun_neq, transpose_seg_fun_neq; auto with arith. destruct m as [|m]. rewrite transpose_seg_fun_r; auto. destruct m as [|m], r as [|r]. rewrite transpose_seg_fun_l; auto. rewrite transpose_seg_fun_neq; auto. rewrite transpose_seg_fun_same; simpl. rewrite transpose_seg_fun_neq; auto. destruct (nat_eq_dec m r) as [|h4]; subst. rewrite transpose_seg_fun_l, transpose_seg_fun_l; auto. rewrite transpose_seg_fun_neq, transpose_seg_fun_neq; auto. destruct m as [|m]. rewrite transpose_seg_fun_neq; auto with arith. destruct (nat_eq_dec m r) as [h4 | h4], (nat_eq_dec m s) as [h5 | h5]; subst; subst. do 2 rewrite transpose_seg_fun_same; auto. do 2 rewrite transpose_seg_fun_l; auto. do 2 rewrite transpose_seg_fun_r; auto. rewrite transpose_seg_fun_neq, transpose_seg_fun_neq; auto. omega. apply S_inj in ha; subst. erewrite transpose_seg_fun_neq, nth_lt_app2; simpl. rewrite length_map_seg, minus_same; auto with arith. auto. intro; subst; contradict (lt_irrefl _ h1). intro; subst; contradict (lt_irrefl _ h2). Grab Existential Variables. rewrite length_map_seg; auto with arith. auto. Qed. Definition transposed_list_nat (l:list nat) (r s:nat) (pfr:r < length l) (pfs:s < length l) : list nat := map_seg (transpose_functor (nth_lt l) pfr pfs). Lemma transposed_list_nat_eq_nil_iff : forall (l:list nat) (r s:nat) (pfr:r < length l) (pfs:s < length l), transposed_list_nat l r s pfr pfs = nil <-> l = nil. unfold transposed_list_nat; intros l r s h1 h2; split; intro h3. apply map_seg_eq_nil in h3. apply length_eq0 in h3; auto. subst; simpl; auto. Qed. Lemma transposed_list_nat_same : forall (l:list nat) (r:nat) (pfr pfr':r < length l), transposed_list_nat l r r pfr pfr' = l. intros l r h1 h2. unfold transposed_list_nat. rewrite transpose_functor_same. apply map_seg_nth_lt. Qed. Lemma length_transposed_list_nat : forall (l:list nat) (r s:nat) (pfr:r < length l) (pfs:s < length l), length (transposed_list_nat l r s pfr pfs) = length l. intros l r s h1 h2; unfold transposed_list_nat. rewrite length_map_seg; auto. Qed. Lemma transposed_list_nat_functional : forall (l l':list nat) (pfe:l = l') (r s:nat) (pfr:r < length l) (pfs:s < length l), transposed_list_nat l r s pfr pfs = transposed_list_nat l' r s (lt_length_subst _ _ r pfe pfr) (lt_length_subst _ _ s pfe pfs). intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma transposed_list_nat_functional' : forall (l:list nat) (r s r' s':nat) (pfr:r < length l) (pfs:s < length l) (pfer:r = r') (pfes:s = s'), transposed_list_nat l r s pfr pfs = transposed_list_nat l r' s' (eq_lt_trans _ _ _ (eq_sym pfer) pfr) (eq_lt_trans _ _ _ (eq_sym pfes) pfs). intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma transposed_list_nat_undoes_itself : forall l r s (pfr: r < length l) (pfs: s < length l) (pfr':r < length (transposed_list_nat l r s pfr pfs)) (pfs':s < length (transposed_list_nat l r s pfr pfs)), transposed_list_nat (transposed_list_nat l r s pfr pfs) r s pfr' pfs' = l. intros l r s hr hs hr' hs'. unfold transposed_list_nat. rewrite map_seg_eq_iff. ex_goal. rewrite length_map_seg; auto. exists hex. intros m hlt hlt'. unfold transpose_functor. unfold hlt'. clear hlt'. destruct (eq_nat_dec m r); subst. rewrite beq_nat_same. rewrite nth_lt_map_seg'. rewrite beq_nat_same. destruct (eq_nat_dec s r); subst. rewrite beq_nat_same. f_equal; apply proof_irrelevance. rewrite <- neq_nat_compat in n. rewrite n. f_equal; apply proof_irrelevance. rewrite <- neq_nat_compat in n. rewrite n. destruct (eq_nat_dec m s); subst. rewrite beq_nat_same. rewrite nth_lt_map_seg'. rewrite beq_nat_same. apply nth_lt_functional3. rewrite <- neq_nat_compat in n0. rewrite n0. rewrite nth_lt_map_seg'. rewrite n, n0. f_equal; apply proof_irrelevance. Qed. Lemma transposed_list_nat_eq_iff : forall (l l':list nat) (r s:nat) (pfr:r < length l) (pfs:s < length l), transposed_list_nat l r s pfr pfs = l' <-> exists (pf:length l = length l'), let pfr' := lt_eq_trans _ _ _ pfr pf in let pfs' := lt_eq_trans _ _ _ pfs pf in l = transposed_list_nat l' r s pfr' pfs'. intros l l' r s h1 h2. split; intro h3. subst. ex_goal. rewrite length_transposed_list_nat; auto. exists hex. intros h3 h4. rewrite transposed_list_nat_undoes_itself; auto. destruct h3 as [h3 h4]. simpl in h4. generalize h1 h2. rewrite h4. intros h5 h6. rewrite transposed_list_nat_undoes_itself; auto. Qed. Lemma transposed_list_nat_comm : forall (l:list nat) (r s:nat) (pfr:r < length l) (pfs:s < length l), transposed_list_nat l r s pfr pfs = transposed_list_nat l s r pfs pfr. intros l r s hr hs. unfold transposed_list_nat. rewrite map_seg_eq_iff. ex_goal. rewrite length_map_seg; auto. exists hex. intros m hlt h1. rewrite nth_lt_map_seg'. unfold transpose_functor. unfold h1. clear h1. destruct (eq_nat_dec m r) as [h2 | h2], (eq_nat_dec m s) as [h3 | h3]; subst; try repeat rewrite beq_nat_same; subst; try repeat rewrite beq_nat_same; try (f_equal; apply proof_irrelevance). rewrite <- neq_nat_compat in h3. rewrite h3; auto. rewrite <- neq_nat_compat in h2. rewrite h2; auto. rewrite <- neq_nat_compat in h2, h3. rewrite h2, h3; auto. f_equal; apply proof_irrelevance. Qed. Lemma transposed_list_nat1 : forall (l:list nat) (r s:nat) (pfr:r < length l) (pfs:s < length l), length l = 1 -> transposed_list_nat l r s pfr pfs = l. intros l r s hr hs h1. unfold transposed_list_nat. rewrite map_seg_eq_iff. exists (eq_refl _). intros m h2. simpl. unfold transpose_functor. destruct (eq_nat_dec m r); subst. rewrite beq_nat_same. pose proof hr as h3. pose proof hs as h4. rewrite h1 in h3, h4. apply lt1 in h3. apply lt1 in h4; subst. apply nth_lt_functional3. pose proof hr as h3. pose proof hs as h4. pose proof h2 as h2'. rewrite h1 in h2', h3, h4. apply lt1 in h2'; apply lt1 in h3; apply lt1 in h4; subst. contradict n; auto. Qed. Lemma transposed_seg_list_compat : forall (n r s:nat) (pfr:r < n) (pfs:s < n), transposed_seg_list pfr pfs = transposed_list_nat (seg_list n) r s (lt_eq_trans _ _ _ pfr (eq_sym (length_seg_list _))) (lt_eq_trans _ _ _ pfs (eq_sym (length_seg_list _))). intros n r s h1 h2. unfold transposed_seg_list, transposed_list_nat. rewrite map_seg_eq_iff. ex_goal. rewrite length_map_seg, length_seg_list; auto. exists hex. intros i h3. simpl. rewrite nth_lt_map_seg'. symmetry. gen0. intro h4. clear hex. unfold transpose_functor. destruct (eq_nat_dec i r) as [|h5], (eq_nat_dec i s) as [|h6]; subst; subst; try repeat rewrite beq_nat_same; try rewrite nth_lt_seg_list; try rewrite transpose_seg_fun_r; try rewrite transpose_seg_fun_l; auto. rewrite <- neq_nat_compat in h5. rewrite h5, nth_lt_seg_list; auto. pose proof h5 as h5'. pose proof h6 as h6'. rewrite <- neq_nat_compat in h5, h6. rewrite h5, h6, nth_lt_seg_list. rewrite transpose_seg_fun_neq; auto. Qed. Lemma in_transposed_list_nat_iff : forall (l:list nat) (r s:nat) (pfr:r < length l) (pfs:s < length l) (i:nat), In i (transposed_list_nat l r s pfr pfs) <-> In i l. unfold transposed_list_nat, transpose_functor; intros l r s h1 h2 i; split; intro h3. rewrite in_map_seg_iff in h3. destruct h3 as [j [h3 h4]]; subst. destruct (eq_nat_dec j r) as [h4 | h4]; subst. rewrite beq_nat_same. apply in_nth_lt. destruct (eq_nat_dec j s) as [h5 | h5]; subst. rewrite <- neq_nat_compat in h4. rewrite beq_nat_same, h4. apply in_nth_lt. rewrite <- neq_nat_compat in h4, h5. rewrite h4, h5. apply in_nth_lt. rewrite in_map_seg_iff. destruct (eq_nat_dec i (nth_lt l s h2)) as [|h4]; subst. exists r, h1. rewrite beq_nat_same; auto. destruct (eq_nat_dec i (nth_lt l r h1)) as [|h5]; subst. exists s, h2. destruct (eq_nat_dec s r) as [|h5]; subst; auto. contradict h4; auto. f_equal; apply proof_irrelevance. rewrite <- neq_nat_compat in h5. rewrite beq_nat_same, h5; auto. exists (lind nat_eq_dec _ _ h3). exists (lt_lind nat_eq_dec _ _ _). destruct (eq_nat_dec (lind nat_eq_dec _ _ h3) r) as [|h6], (eq_nat_dec (lind nat_eq_dec _ _ h3) s) as [|h7]; subst; try rewrite beq_nat_same. erewrite nth_lt_functional3. rewrite lind_compat; auto. contradict h5. erewrite nth_lt_functional3. rewrite lind_compat; auto. contradict h4. erewrite nth_lt_functional3. rewrite lind_compat; auto. rewrite <- neq_nat_compat in h6, h7. rewrite h6, h7. erewrite nth_lt_functional3. rewrite lind_compat; auto. Qed. Lemma transposed_list_nat_app_sing : forall (l:list nat) (r s:nat) (pfr:r < length l) (pfs:s < length l), transposed_list_nat (l++((length l)::nil)) r s (lt_app _ _ _ pfr) (lt_app _ _ _ pfs) = transposed_list_nat l r s pfr pfs ++ (length l) :: nil. intros l r s h1 h2. unfold transposed_list_nat. rewrite map_seg_eq_iff. ex_goal. do 2 rewrite app_length; simpl. f_equal. rewrite length_map_seg; auto. exists hex. unfold transpose_functor; intros m h3. symmetry. pose proof h1 as h1'. pose proof h2 as h2'. apply lt_neq in h1'. apply lt_neq in h2'. rewrite <- neq_nat_compat in h1', h2'. rewrite beq_nat_sym in h1', h2'. destruct (nat_eq_dec m (length l)) as [h5|h5], (nat_eq_dec r s) as [ha|ha]; subst. rewrite h1'. erewrite nth_lt_app2. simpl. erewrite nth_lt_app2; simpl. rewrite minus_same, length_map_seg, minus_same; auto. erewrite nth_lt_app2. simpl. rewrite length_map_seg, minus_same; auto. rewrite h1', h2'. erewrite nth_lt_app2; simpl. rewrite minus_same; auto. destruct (eq_nat_dec m s) as [|hms]; subst. rewrite beq_nat_same. erewrite nth_lt_app1. rewrite nth_lt_map_seg'. rewrite beq_nat_same. erewrite nth_lt_app1. reflexivity. rewrite <- neq_nat_compat in hms. rewrite hms. erewrite nth_lt_app1. rewrite nth_lt_map_seg', hms. erewrite nth_lt_app1. reflexivity. destruct (nat_eq_dec m r) as [|hmr], (nat_eq_dec m s) as [|hms]; subst; try rewrite beq_nat_same; subst. erewrite nth_lt_app1. rewrite nth_lt_map_seg'. rewrite beq_nat_same. erewrite nth_lt_app1. reflexivity. erewrite nth_lt_app1. rewrite nth_lt_map_seg'. rewrite beq_nat_same. erewrite nth_lt_app1. reflexivity. rewrite <- neq_nat_compat in hmr. rewrite hmr. erewrite nth_lt_app1. rewrite nth_lt_map_seg'. rewrite hmr, beq_nat_same. erewrite nth_lt_app1. reflexivity. rewrite <- neq_nat_compat in hmr, hms. rewrite hmr, hms. erewrite nth_lt_app1. rewrite nth_lt_map_seg'. rewrite hmr, hms. erewrite nth_lt_app1. reflexivity. Grab Existential Variables. rewrite length_map_seg; auto. pose proof h3 as h6. rewrite app_length in h6. simpl in h6. omega. rewrite length_map_seg. omega. rewrite length_map_seg. omega. rewrite length_map_seg. omega. rewrite length_map_seg; rewrite app_length in h3; simpl in h3. omega. rewrite length_map_seg. omega. auto with arith. rewrite length_map_seg; auto. auto. rewrite length_map_seg; auto. Qed. Lemma transposed_list_nat_app_sing' : forall (l:list nat) (r s:nat) (pfr:r < length l) (pfs:s < length l) (pfr':r < length (l++((length l)::nil))) (pfs':s < length (l++((length l)::nil))), transposed_list_nat (l++((length l)::nil)) r s pfr' pfs' = transposed_list_nat l r s pfr pfs ++ (length l) :: nil. intros l r s h1 h2 h3 h4. assert (h3 = lt_app _ _ _ h1). apply proof_irrelevance. assert (h4 = lt_app _ _ _ h2). apply proof_irrelevance. subst. apply transposed_list_nat_app_sing. Qed. Lemma transposed_list_nat_01 : forall l a b (pfl0:0 < length (a::b::l)) (pfl1:1 < length (a::b::l)), transposed_list_nat (a::b::l) 0 1 pfl0 pfl1 = b::a::l. intros l a b h1 h2. unfold transposed_list_nat. rewrite map_seg_eq_iff. simpl. exists eq_refl. intros m h3. unfold transpose_functor. destruct m as [|m]; subst. rewrite beq_nat_same; simpl; auto. simpl. destruct m as [|m]. rewrite beq_nat_same; auto. pose proof (lt_0_Sn m) as h4. apply lt_neq in h4. rewrite <- neq_nat_compat, beq_nat_sym in h4. rewrite h4. apply nth_indep. omega. Qed. Lemma transposed_list_nat_01' : forall l (pfl0:0 < length l) (pfl1:1 < length l), transposed_list_nat l 0 1 pfl0 pfl1 = hd2 l pfl1 :: hd' l (lt_length _ _ pfl0) :: tl2 l. intros l h1. pose proof (lt_length _ _ h1) as h2. pose proof (hd_compat' _ h2) as h3. simpl in h3. revert h1. rewrite h3. intros h1 h4; simpl. simpl in h1, h4. pose proof h4 as h4'. apply lt_S_n in h4'. pose proof (lt_length _ _ h4') as h5. pose proof (hd_compat' _ h5) as h6. simpl in h6. generalize h1, h4. rewrite h6. intros; rewrite transposed_list_nat_01. unfold hd2, tl2. simpl; auto. Qed. Lemma transposed_list_nat_S : forall (l:list nat) (r s:nat) (pfr:r < length l) (pfs:s < length l) (pfsr:S r < S (length (transposed_list_nat l r s pfr pfs))) (pfss:S s < S (length (transposed_list_nat l r s pfr pfs))), transposed_list_nat (length l :: transposed_list_nat l r s pfr pfs) (S r) (S s) pfsr pfss = length l :: l. intros l r s hr hs hsr hss. unfold transposed_list_nat. rewrite map_seg_eq_iff. ex_goal. simpl. f_equal. rewrite length_map_seg; auto. exists hex. intros m h1 h2. unfold transpose_functor, h2. clear h2. destruct (eq_nat_dec m (S r)) as [h2 | h2], (eq_nat_dec m (S s)) as [h3 | h3]; subst. apply S_inj in h3; subst. rewrite beq_nat_same, nth_lt_cons, nth_lt_map_seg', beq_nat_same. rewrite nth_lt_cons. apply nth_lt_functional3. rewrite beq_nat_same, nth_lt_cons, nth_lt_map_seg', beq_nat_same. apply S_inj_neq in h3. apply neq_sym in h3. rewrite <- neq_nat_compat in h3. rewrite h3. simpl. rewrite nth_lt_compat. apply nth_indep. assumption. rewrite <- neq_nat_compat in h2. rewrite h2, beq_nat_same. rewrite nth_lt_cons, nth_lt_map_seg', beq_nat_same; simpl. rewrite nth_lt_compat. apply nth_indep. assumption. pose proof h2 as h2'. pose proof h3 as h3'. rewrite <- neq_nat_compat in h2, h3. rewrite h2, h3. simpl. simpl in h1; rewrite length_map_seg in h1. destruct m as [|m]; auto. erewrite nth_indep. erewrite <- nth_lt_compat. rewrite nth_lt_map_seg. apply S_inj_neq in h2'. apply S_inj_neq in h3'. rewrite <- neq_nat_compat in h2', h3'. rewrite h2', h3'. rewrite nth_lt_compat. apply nth_indep. auto with arith. rewrite length_map_seg; auto with arith. Grab Existential Variables. auto with arith. Qed. Lemma transposed_list_nat_cons_0_S_lind : forall (l:list nat) (a:nat) (pfi:In (length l) l) (pf0 : 0 < length (a::l)) (pfs : S (lind nat_eq_dec l (length l) pfi) < length (a::l)), let pfs' := lt_S_n _ _ pfs in transposed_list_nat (a :: l) 0 (S (lind nat_eq_dec l (length l) pfi)) pf0 pfs = (length l) :: (list_subst l (lind nat_eq_dec l (length l) pfi) pfs' a). intros l a h1 h2 h3. unfold transposed_list_nat. rewrite map_seg_eq_iff. ex_goal. simpl. f_equal. rewrite length_list_subst; auto. exists hex. intros m h4; simpl. unfold transpose_functor. destruct (eq_nat_dec m 0) as [|hnm]; subst. rewrite beq_nat_same. rewrite nth_lt_cons. erewrite nth_lt_functional3. rewrite lind_compat; auto. pose proof hnm as hnm'. rewrite <- neq_nat_compat in hnm. rewrite hnm. destruct (eq_nat_dec m (S (lind nat_eq_dec _ _ h1))); subst. rewrite beq_nat_same. rewrite nth_lt_0'. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite nth_lt_list_subst_same; auto. rewrite length_list_subst. apply lt_lind. destruct m. contradict n; auto. contradict hnm'; auto. pose proof n as h5. rewrite <- neq_nat_compat in h5. rewrite h5. rewrite nth_lt_cons. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite nth_lt_list_subst_neq. apply nth_lt_functional3. omega. rewrite length_list_subst; auto with arith. Grab Existential Variables. rewrite length_list_subst; auto with arith. rewrite length_list_subst; auto with arith. Qed. Lemma transposed_list_nat_0_pred : forall l (pfl0:0 < length l) (pfl1:1 < length l) (pfp:pred (length l) < length l), transposed_list_nat l 0 (pred (length l)) pfl0 pfp = nth_lt _ (pred (length l)) pfp :: seq_lind l 1 (pred (pred (length l))) pfl1 ++ hd' l (lt_length _ _ pfl0) ::nil. intros l h1 h2 h3. unfold transposed_list_nat, transpose_functor. rewrite map_seg_eq_iff. ex_goal. simpl. rewrite app_length, length_seq_lind. rewrite S_compat; auto. omega. omega. exists hex. intros m h4 h5. destruct m as [|m]; simpl; auto. destruct (zerop (pred (length l))) as [hz | hz]. rewrite hz at 1. erewrite nth_indep. rewrite <- nth_lt_compat. erewrite nth_lt_app2. simpl. rewrite length_seq_lind. rewrite hz; simpl; rewrite <- minus_n_O; simpl. destruct m as [|m]; simpl. destruct l as [|l]; simpl. omega. simpl in hz. apply length_eq0 in hz; subst; auto. rewrite <- length_tl in hz. apply length_eq0 in hz; subst; auto. rewrite tl_eq_nil_iff in hz. destruct hz as [|h7]; subst. simpl in h1; omega. destruct h7; subst. simpl. destruct m; auto. rewrite <- pred_of_minus. auto with arith. (* destruct (nat_eq_dec m (pred (pred (length l)))) as [h6 | h6]; subst. rewrite minus_same, nth_lt0'; auto. pose proof (le_minus_iff m (pred (pred (length l)))) as hl. hl hl. pose proof (lt_pred _ _ h4) as ha. red in ha. pose proof (le_pred _ _ ha) as h7. simpl in h7; auto. rewrite hl in hl0. rewrite hl0. rewrite nth_lt0'; auto. rewrite <- pred_of_minus. apply le_pred. apply le_pred_n.*) rewrite app_length, length_seq_lind. simpl. omega. omega. rewrite (S_pred _ _ hz) at 1. pose proof (lt_length _ _ h1) as h6. pose proof (hd_compat' _ h6) as h7. simpl in h7. pose proof (nth_lt_functional1 _ _ (S m) h7 h4) as h8. hfirst h8. simpl. rewrite length_tl. omega. specialize (h8 hf). rewrite h8. pose proof hf as h9. simpl in h9. apply lt_S_n in h9. rewrite length_tl in h9. red in h9. apply mono_pred_le in h9. simpl in h9. apply le_lt_eq_dec in h9. destruct h9 as [h9 | h9]. apply lt_neq in h9. pose proof h9 as h9'. rewrite <- neq_nat_compat in h9. rewrite h9; auto. clear h5 h7. simpl. symmetry. erewrite nth_indep. erewrite <- nth_lt_compat. (* rewrite nth_lt_seq_lind. apply lt_S_n in h9. rewrite <- (S_pred _ _ h9); auto. simpl in hf. rewrite <- pred_of_minus. auto with arith. subst. rewrite beq_nat_same. simpl. simpl in hex. erewrite nth_indep. rewrite <- nth_lt_compat.*) erewrite nth_lt_app1. erewrite nth_lt_seq_lind. simpl. rewrite h8 at 1. simpl; auto. apply nat_eq_dec. rewrite app_length, length_seq_lind; simpl. rewrite S_compat, <- (S_pred _ _ hz). omega. rewrite <- pred_of_minus; auto with arith. rewrite h9 at 1. rewrite beq_nat_same. erewrite nth_indep. rewrite <- nth_lt_compat. erewrite nth_lt_app2; simpl. rewrite length_seq_lind, h9, minus_same at 1. rewrite nth_lt0'; auto. rewrite <- pred_of_minus; auto with arith. rewrite app_length; simpl; rewrite length_seq_lind, h9, S_compat; auto with arith. rewrite <- pred_of_minus; auto with arith. Grab Existential Variables. clear h5 h7. subst. rewrite length_seq_lind; auto with arith. rewrite <- pred_of_minus; auto with arith. rewrite app_length; simpl; rewrite S_compat, length_seq_lind, <- (S_pred _ _ hz), h9; auto with arith. rewrite <- pred_of_minus; auto with arith. rewrite length_seq_lind. pose proof (lt_pred _ _ h4) as h10. apply lt_dec in h10. destruct h10; try contradiction; auto. rewrite <- pred_of_minus; auto with arith. rewrite app_length, length_seq_lind, S_compat; simpl. rewrite <- (S_pred _ _ hz); auto. pose proof (lt_pred _ _ h4); auto. rewrite <- pred_of_minus; auto with arith. rewrite length_seq_lind; auto with arith. rewrite hz; omega. rewrite <- pred_of_minus; auto with arith. rewrite app_length, length_seq_lind, S_compat, hz; simpl. omega. rewrite hz; auto with arith. Qed. Lemma last_transposed_list_nat_pred : forall l r (pfr:r < length l) (pfp:pred (length l) < length l) (d:nat), last (transposed_list_nat l r (pred (length l)) pfr pfp) d = nth_lt l r pfr. intros l r h1 h2 d. rewrite (last_compat _ _ (lt_length _ _ (lt_eq_trans _ _ _ h1 (eq_sym (length_transposed_list_nat _ _ _ _ _))))). rewrite lst_eq. unfold transposed_list_nat. rewrite nth_lt_map_seg'. gen0. rewrite length_map_seg. intro h3. unfold transpose_functor. destruct (eq_nat_dec (pred (length l)) r) as [| h4]; subst. rewrite beq_nat_same. apply nth_lt_functional3. rewrite beq_nat_same. rewrite <- neq_nat_compat in h4. rewrite h4; auto. Qed. Lemma lst_transposed_list_nat_pred : forall l r (pfr:r < length l) (pfp:pred (length l) < length l) (pfne:transposed_list_nat l r (pred (length l)) pfr pfp <> nil), lst (transposed_list_nat l r (pred (length l)) pfr pfp) pfne = nth_lt l r pfr. intros; rewrite lst_compat', last_transposed_list_nat_pred; auto. Qed. Lemma lst_transposed_list_nat : forall l r s (pfr:r < length l) (pfs:s < length l) (pfne:transposed_list_nat l r s pfr pfs <> nil), r < pred (length l) -> s < pred (length l) -> lst (transposed_list_nat l r s pfr pfs) pfne = lst l (iffn1 (transposed_list_nat_eq_nil_iff l r s pfr pfs) pfne). intros l r s hr hs hne h1 h2. pose proof (lt_length _ _ hr) as hnn. unfold transposed_list_nat. rewrite lst_map_seg. unfold transpose_functor. destruct (eq_nat_dec (pred (length l)) r); subst. rewrite beq_nat_same. apply lt_irrefl in h1; contradiction. rewrite <- neq_nat_compat in n. rewrite n. destruct (eq_nat_dec (pred (length l)) s); subst. apply lt_irrefl in h2; contradiction. rewrite <- neq_nat_compat in n0. rewrite n0. rewrite nth_lt_eq_iff. exists (in_lst _ _). pose proof (lt_lind nat_eq_dec _ _ (in_lst l hnn)) as h3. exists (pred (count_occ nat_eq_dec l (lst l (iffn1 (transposed_list_nat_eq_nil_iff l r s hr hs) hne)))). split. eapply lt_pred_n. apply gt_lt. rewrite <- (count_occ_In nat_eq_dec). apply in_lst. gen0. do 2 rewrite lst_eq. intro h4. apply lind_occ_pred_nth_lt_pred. Qed. Lemma no_dup_transposed_list_nat : forall l r s (pfr:r < length l) (pfs:s < length l), NoDup l -> NoDup (transposed_list_nat l r s pfr pfs). intros l r s h1 h2 h3. unfold transposed_list_nat. apply no_dup_map_seg_inj. apply inj_transpose_functor. apply inj_seg_nth_lt; auto. Qed. Lemma no_dup_transposed_list_nat' : forall l r s (pfr:r < length l) (pfs:s < length l), NoDup (transposed_list_nat l r s pfr pfs) -> NoDup l. intros l r s h1 h2 h3. unfold transposed_list_nat in h3. pose proof (no_dup_map_seg_inj' _ h3) as h4. pose proof (inj_transpose_functor' _ _ _ _ _ h4) as h5. apply inj_seg_nth_lt' in h5; auto. apply nat_eq_dec. Qed. Lemma list_to_set_transposed_list_nat : forall l r s (pfr:r < length l) (pfs:s < length l), list_to_set (transposed_list_nat l r s pfr pfs) = list_to_set l. intros l r s h1 h2; apply Extensionality_Ensembles; red; split; red; intros x h3; rewrite <- list_to_set_in_iff in h3; rewrite <- list_to_set_in_iff. rewrite in_transposed_list_nat_iff in h3; auto. rewrite in_transposed_list_nat_iff; auto. Qed. Lemma transposed_list_nat_eq : forall (l l':list nat) (r s:nat) (pfr:r < length l) (pfs:s < length l) (pfe:transposed_list_nat l r s pfr pfs = l'), let pfr' := lt_eq_trans _ _ _ pfr (eq_trans (eq_sym (length_transposed_list_nat _ _ _ _ _)) (length_subst _ _ pfe)) in let pfs' := lt_eq_trans _ _ _ pfs (eq_trans (eq_sym (length_transposed_list_nat _ _ _ _ _)) (length_subst _ _ pfe)) in l = transposed_list_nat l' r s pfr' pfs'. intros. pose proof (transposed_list_nat_functional _ _ pfe r s) as h1. hfirst h1. rewrite length_transposed_list_nat; auto. specialize (h1 hf). hfirst h1. rewrite length_transposed_list_nat; auto. specialize (h1 hf0). rewrite transposed_list_nat_undoes_itself in h1. rewrite h1. f_equal; apply proof_irrelevance. Qed. Definition cycled_seg_list n k := map_seg (shift_seg_fun n k). Lemma in_cycled_seg_list_iff : forall n k i, In i (cycled_seg_list n k) <-> i < n. unfold cycled_seg_list; intros n k i; split; intro h1. rewrite in_map_seg_iff in h1. destruct h1 as [m [h1 h2]]; subst. apply lt_shift_seg_fun. rewrite in_map_seg_iff. unfold shift_seg_fun. assert (h2:0 < n). omega. destruct (lt_eq_lt_dec i k) as [[h3 | h3] | h3]. pose proof (minus_mod_le). destruct (le_lt_dec (k - i) n) as [h4 | h4]. apply le_lt_eq_dec in h4. destruct h4 as [h4 | h4]. exists (n - (k - i)). ex_goal. apply lt_minus; auto. apply lt_le_weak; auto. rewrite lt_iff in h3; auto. exists hex. rewrite minus_minus_plus_same_le; auto with arith. rewrite plus_mod, mod_same, plus_0_l. rewrite mod_mod. apply lt_mod_eq; auto. subst. exists 0, h2. rewrite plus_0_l. rewrite <- (lt_mod_eq i (k - i)) at 2; auto. symmetry. rewrite mod_eq_iff. apply dv_refl. assumption. auto with arith. exists ((k*(n - 1) + i) mod n), (lt_mod _ _ h2). rewrite plus_mod. rewrite mod_mod at 1. rewrite <- minus_dist_l, <- plus_mod, mult_1_r, minus_plus_plus_same_le, mult_comm, plus_mod, times_mod_l, plus_0_l, mod_mod, lt_mod_eq; auto. pose proof (mult_O_le k n) as h5. destruct h5 as [|h5]; subst. apply lt_irrefl in h2; contradiction. rewrite mult_comm; auto. subst. exists _, h2. rewrite plus_0_l, lt_mod_eq; auto. exists ((k*(n - 1) + i) mod n), (lt_mod _ _ h2). pose proof (lt_trans _ _ _ h3 h1) as h4. rewrite <- (lt_mod_eq _ _ h4) at 2. rewrite <- minus_dist_l, <- plus_mod, mult_1_r, minus_plus_plus_same_le, plus_mod, times_mod_r, plus_0_l, mod_mod, lt_mod_eq; auto. pose proof (mult_O_le k n) as h5. destruct h5 as [|h5]; subst. apply lt_irrefl in h2; contradiction. rewrite mult_comm; auto. Qed. Lemma length_cycled_seg_list : forall k n, length (cycled_seg_list n k) = n. intros k n. unfold cycled_seg_list. rewrite length_map_seg; auto. Qed. Lemma list_to_set_cycled_seg_list : forall k n, list_to_set (cycled_seg_list n k) = seg n. intros k n; apply Extensionality_Ensembles; red; split; red; intros i h1. rewrite <- list_to_set_in_iff, in_cycled_seg_list_iff in h1; constructor; auto. destruct h1. rewrite <- list_to_set_in_iff. rewrite in_cycled_seg_list_iff; auto. Qed. Lemma shift_left_list : forall n, n > 0 -> cycled_seg_list n 1 = seq 1 (pred n) ++ 0 :: nil. intros n h1. unfold cycled_seg_list. rewrite map_seg_eq_iff. ex_goal. rewrite app_length, seq_length; simpl. omega. exists hex. intros m h2. simpl. pose proof h2 as h3. apply lt_dec in h3. destruct h3 as [h3 | h3]. erewrite nth_lt_app1. erewrite nth_lt_seq. unfold shift_seg_fun. rewrite plus_comm. rewrite lt_mod_eq; auto. omega. erewrite nth_lt_app2. simpl. rewrite seq_length. rewrite <- h3, minus_same. unfold shift_seg_fun. rewrite h3. rewrite S_compat. rewrite <- (S_pred _ _ h2), mod_same; auto. Grab Existential Variables. rewrite seq_length, h3; auto. rewrite seq_length; auto. Qed. Lemma shift_right_list : forall n, n > 0 -> cycled_seg_list n (pred n) = (pred n) :: seq 0 (pred n). intros n h1. unfold cycled_seg_list. rewrite map_seg_eq_iff. ex_goal. simpl; rewrite seq_length; auto. rewrite <- (S_pred _ _ h1); auto. exists hex. intros m h2 h3. pose proof h2 as h4. destruct m as [|m]. simpl. unfold shift_seg_fun. rewrite plus_0_l, lt_mod_eq; auto. eapply lt_pred_n. apply h4. rewrite nth_lt_cons, nth_lt_seq, plus_0_l. unfold shift_seg_fun. rewrite plus_S_shift_l, <- plus_S_shift_r. rewrite <- (S_pred _ _ h1), plus_mod, mod_same, plus_0_r. rewrite mod_mod, lt_mod_eq; auto with arith. Qed. Inductive permuted_seg_list : list nat -> Prop := | permuted_seg_list_id : forall n, permuted_seg_list (seg_list n) | permuted_seg_list_transpose : forall (l':list nat) r s (pfr:r < length l') (pfs:s < length l'), permuted_seg_list l' -> permuted_seg_list (transposed_list_nat l' r s pfr pfs). Lemma permuted_seg_list_nil : permuted_seg_list nil. change (permuted_seg_list (seg_list 0)). constructor. Qed. Lemma permuted_seg_list_0 : permuted_seg_list (0::nil). change (permuted_seg_list (seg_list 1)). constructor. Qed. Lemma lt_permuted_seg_list : forall (l:list nat) (i:nat), permuted_seg_list l -> In i l -> i < length l. intros l i h1. revert i. induction h1 as [i|l r s h1 h2]. intros j h4. rewrite in_seg_list_iff in h4; auto. rewrite length_seg_list; auto. intros i h4. rewrite in_transposed_list_nat_iff in h4; auto. rewrite length_transposed_list_nat; auto. Qed. Lemma permuted_seg_list_app_sing : forall (l:list nat), permuted_seg_list l -> permuted_seg_list (l++((length l)::nil)). intros l h1. induction h1. rewrite length_seg_list. rewrite <- seg_list_S. constructor. rewrite length_transposed_list_nat. rewrite <- (transposed_list_nat_app_sing _ _ _ pfr pfs). constructor. assumption. Qed. Lemma no_dup_permuted_seg_list : forall (l:list nat), permuted_seg_list l -> NoDup l. intros l h1. induction h1. apply no_dup_seg_list. unfold transposed_list_nat. apply no_dup_map_seg_inj. apply inj_transpose_functor. apply inj_seg_nth_lt. assumption. Qed. (* Lemma nth_lt_permuted_seg_list_preserve : forall (l:list nat), permuted_seg_list l -> seg_fun_lt_preserve (nth_lt l). intros l h1. induction h1. apply nth_lt_seg_list_preserve. red in IHh1; red. intros i h2. unfold transposed_list_nat. rewrite nth_lt_map_seg', length_map_seg. unfold transpose_functor. repeat destruct eq_nat_dec; auto. Qed. *) Lemma list_to_set_permuted_seg_list : forall l, permuted_seg_list l -> list_to_set l = seg (length l). intros l h1. induction h1. rewrite list_to_set_seg_list_eq, length_seg_list; auto. rewrite length_transposed_list_nat. rewrite <- IHh1. rewrite list_to_set_transposed_list_nat; auto. Qed. Lemma transposed_list_nat_cons_app_sing : forall (l:list nat) (m n:nat) (pf0 : 0 < length (m :: l ++ n::nil)) (pfS : S (length l) < length (m :: l ++ n::nil)), transposed_list_nat (m :: l ++ n::nil) 0 (S (length l)) pf0 pfS = n :: l ++ m::nil. intros l m n h0 hS. unfold transposed_list_nat. rewrite map_seg_eq_iff. ex_goal. simpl; f_equal. simpl; do 2 rewrite app_length; auto. exists hex. intros i hi. unfold transpose_functor. destruct (zerop i) as [|hz]; subst. rewrite beq_nat_same. rewrite nth_lt_S; simpl. erewrite nth_lt_app2. simpl. rewrite minus_same; auto. pose proof hz as hz'. apply lt_neq in hz. rewrite <- neq_nat_compat, beq_nat_sym in hz. rewrite hz. destruct (eq_nat_dec i (S (length l))) as [|hni]; subst. rewrite beq_nat_same. rewrite nth_lt0'. symmetry. gen0. rewrite app_comm_cons. intro h1. erewrite nth_lt_app2. simpl. rewrite minus_same; auto. pose proof hni as hni'. rewrite <- neq_nat_compat in hni. rewrite hni. assert (h1:0 < i). omega. symmetry. gen0. generalize hi. rewrite (S_pred _ _ h1). intros h2 h3. do 2 rewrite nth_lt_S. simpl. do 2 erewrite nth_lt_app1. f_equal. Grab Existential Variables. simpl in h3. apply lt_S_n in h3. rewrite app_length, S_compat in h3. simpl in h2. omega. simpl; auto. auto. Qed. Lemma permuted_seg_list_S : forall (l:list nat), permuted_seg_list l -> permuted_seg_list (l ++ length l::nil). intros l h1. induction h1; simpl. rewrite length_seg_list, <- seg_list_S. constructor. rewrite length_transposed_list_nat. rewrite <- transposed_list_nat_app_sing. constructor. assumption. Qed. Lemma permuted_seg_list_n_seg_list : forall n, permuted_seg_list (n::seg_list n). intro n. induction n as [|n ih]. rewrite seg_list0. change (permuted_seg_list (seg_list 1)). constructor. rewrite seg_list_S. pose proof (transposed_list_nat_undoes_itself (S n::seg_list n++n::nil) 0 (S n) (lt_0_Sn _)) as h1. hfirst h1; simpl. apply lt_n_S. rewrite app_length; simpl; auto with arith. rewrite length_seg_list, S_compat; auto with arith. specialize (h1 hf). hfirst h1. rewrite length_transposed_list_nat; simpl. rewrite app_length, S_compat. auto with arith. specialize (h1 hf0). hfirst h1. rewrite length_transposed_list_nat; simpl. rewrite app_length, S_compat, length_seg_list; auto with arith. specialize (h1 hf1). rewrite <- h1. constructor. pose proof (transposed_list_nat_cons_app_sing (seg_list n) (S n) n (lt_0_Sn _)) as h2. hfirst h2. simpl. apply lt_n_S. rewrite app_length; simpl; rewrite S_compat; auto with arith. specialize (h2 hf2). match goal with |- permuted_seg_list (transposed_list_nat _ _ _ ?pf _) => generalize pf end. generalize hf. revert h2. simpl. gen0. rewrite <- (length_seg_list n). rewrite length_seg_list. intros h3 h4 h5 h6. erewrite f_equal_dep in h4. rewrite h4 at 1. rewrite length_seg_list. rewrite app_comm_cons. pose proof (permuted_seg_list_S (n::seg_list n) ih) as h7. simpl in h7. rewrite length_seg_list in h7. rewrite <- app_comm_cons; auto. apply proof_irrelevance. Qed. Lemma permuted_seg_list_cons : forall (l:list nat), permuted_seg_list (length l :: seg_list (length l)). intros; apply permuted_seg_list_n_seg_list. Qed. Lemma permuted_seg_list_cons' : forall (l:list nat), permuted_seg_list l -> permuted_seg_list (length l::l). intros l h1. induction h1; simpl. rewrite length_seg_list. apply permuted_seg_list_n_seg_list. rewrite length_transposed_list_nat. pose proof (transposed_list_nat_undoes_itself (length l'::transposed_list_nat l' r s pfr pfs) (S r) (S s)) as h2. simpl in h2. hfirst h2. rewrite length_transposed_list_nat; auto with arith. specialize (h2 hf). hfirst h2. rewrite length_transposed_list_nat; auto with arith. specialize (h2 hf0). hfirst h2. rewrite length_transposed_list_nat; auto with arith. specialize (h2 hf1). hfirst h2. rewrite length_transposed_list_nat; auto with arith. specialize (h2 hf2). rewrite <- h2. constructor. rewrite transposed_list_nat_S; auto. Qed. Lemma map_seg_shift_fun_0n : forall n, map_seg (shift_seg_fun 0 n) = nil. intro; simpl; auto. Qed. Definition cycled_list_nat l k (pf:0 < length l) : list nat := map_seg (shift_functor (nth_lt l) k pf). Definition left_shift_list_nat l (pf:0 < length l) : list nat := cycled_list_nat l (pred (length l)) pf . Definition right_shift_list_nat l (pf:0 < length l) : list nat := cycled_list_nat l 1 pf. Lemma length_cycled_list_nat : forall l k (pf:0 < length l), length (cycled_list_nat l k pf) = length l. intros l k h1. unfold cycled_list_nat. rewrite length_map_seg; auto. Qed. Lemma length_left_shift_list_nat : forall l (pf:0 < length l), length (left_shift_list_nat l pf) = length l. unfold left_shift_list_nat; intros; apply length_cycled_list_nat. Qed. Lemma length_right_shift_list_nat : forall l (pf:0 < length l), length (right_shift_list_nat l pf) = length l. unfold right_shift_list_nat; intros; apply length_cycled_list_nat. Qed. Lemma nth_lt_cycled_list_nat : forall l k i (pf0:0 < length l) (pfi:i < length (cycled_list_nat l k pf0)), nth_lt (cycled_list_nat l k pf0) i pfi = nth_lt l ((i + k) mod (length l)) (lt_mod _ _ pf0). intros l k i h1 h2. unfold cycled_list_nat. rewrite nth_lt_map_seg'. unfold shift_functor; auto. Qed. Lemma nth_lt_left_shift_list_nat : forall l i (pf0:0 < length l) (pfi:i < length (left_shift_list_nat l pf0)), nth_lt (left_shift_list_nat l pf0) i pfi = nth_lt l ((i + (pred (length l))) mod (length l)) (lt_mod _ _ pf0). intros l k i h1. unfold left_shift_list_nat. rewrite nth_lt_cycled_list_nat. unfold shift_functor; auto. Qed. Lemma nth_lt_right_shift_list_nat : forall l i (pf0:0 < length l) (pfi:i < length (right_shift_list_nat l pf0)), nth_lt (right_shift_list_nat l pf0) i pfi = nth_lt l ((S i) mod (length l)) (lt_mod _ _ pf0). intros l k i h1. unfold right_shift_list_nat. rewrite nth_lt_cycled_list_nat. unfold shift_functor; auto. gen0. rewrite S_compat. intros; apply nth_lt_functional3. Qed. Lemma right_shift_list_nat_eq : forall l (pf:0 < length l), let pfn := lt_length _ _ pf in right_shift_list_nat l pf = (tl l) ++ hd' l pfn :: nil. intros l h1 h2. unfold right_shift_list_nat. unfold cycled_list_nat. rewrite map_seg_eq_iff. simpl. unfold h2. clear h2. ex_goal. rewrite app_length; simpl. rewrite length_tl; omega. exists hex. intros m h3. pose proof h3 as h5. apply lt_dec in h5. destruct h5 as [h5 | h5]. unfold shift_functor. pose proof (lt_n_S _ _ h5) as h6. rewrite <- (S_pred _ _ h3) in h6. gen0. rewrite S_compat, (lt_mod_eq _ _ h6). intro h7. symmetry. gen0. intro. rewrite <- length_tl in h5. rewrite (nth_lt_app1 _ _ _ h5). symmetry. generalize h7. apply lt_length in h3. generalize h5. rewrite (hd_compat' _ h3). intros h8 h9. rewrite nth_lt_cons. simpl. apply nth_lt_functional3. subst. unfold shift_functor. gen0. rewrite S_compat, <- (S_pred _ _ h1). intro h4. erewrite nth_lt_app2. symmetry. gen0. rewrite length_tl, <- minus_n_n. intro h5. simpl. symmetry. generalize h4. rewrite mod_same. intro h6. rewrite nth_lt0'. apply hd_functional'. Grab Existential Variables. rewrite length_tl; auto with arith. Qed. Lemma right_shift_cons : forall l a (pf0:0 < length (a::l)), right_shift_list_nat (a::l) pf0 = l ++ a::nil. intros l a h1. unfold right_shift_list_nat, cycled_list_nat. rewrite map_seg_eq_iff. ex_goal. simpl; rewrite app_length; simpl; rewrite S_compat; auto. exists hex. intros m h2. simpl. unfold shift_functor. gen0. rewrite S_compat, S_mod. destruct (nil_dec' l) as [|h3]; subst. simpl; simpl in h2. intro h3. destruct m as [|m]; auto. destruct m as [|m]; auto. apply not_nil_lt in h3. pose proof (lt_n_S _ _ h3) as h4. rewrite (lt_mod_eq _ _ h4). simpl in h2. rewrite (lt_mod_eq _ _ h2), S_compat. pose proof h2 as h2'. apply lt_dec in h2'. simpl in h2'. destruct h2' as [h2'|h2']; subst. pose proof (lt_n_S _ _ h2') as h5. rewrite (lt_mod_eq _ _ h5). intro. rewrite nth_lt_cons. erewrite nth_lt_app1. apply nth_lt_functional3. rewrite mod_same. intros h5. rewrite nth_lt_0'. symmetry. erewrite nth_lt_app2. simpl. rewrite minus_same; auto. Grab Existential Variables. apply le_refl. assumption. Qed. Lemma cycled_seg_list_0n : forall n, cycled_seg_list 0 n = nil. intro; unfold cycled_seg_list; apply map_seg_shift_fun_0n. Qed. Lemma cycled_seg_list_n0 : forall n, cycled_seg_list n 0 = seg_list n. unfold cycled_seg_list. intro n. rewrite (shift_seg_fun0 n). apply map_seg_id_seg. Qed. Lemma cycled_seg_list_nS : forall n k (pf:0 < length (cycled_seg_list n k)), cycled_seg_list n (S k) = right_shift_list_nat (cycled_seg_list n k) pf. intros n k h1. unfold cycled_seg_list. rewrite map_seg_eq_iff. ex_goal. rewrite length_right_shift_list_nat, length_map_seg; auto. exists hex. intros m h2 h3. rewrite nth_lt_right_shift_list_nat, nth_lt_map_seg'. symmetry. gen0. rewrite length_map_seg. intro h4. unfold shift_seg_fun. pose proof h2 as h5. apply lt_dec in h5. destruct h5 as [h5 | h5]. rewrite (lt_mod_eq (S m)) in h4. rewrite (lt_mod_eq (S m)). rewrite plus_S_shift_r, <- plus_S_shift_l; auto. assumption. rewrite (S_pred _ _ h2). apply lt_n_S; auto. rewrite h5. rewrite <- (S_pred _ _ h2), mod_same, plus_0_l. rewrite plus_S_shift_r, <- plus_S_shift_l; auto. rewrite <- (S_pred _ _ h2). rewrite plus_mod, mod_same, plus_0_l, mod_mod; auto. Qed. Lemma right_shift_seg_list : forall n (pf:0 < length (seg_list n)), right_shift_list_nat (seg_list n) pf = seq 1 (pred n) ++ 0 :: nil. intros n h1. pose proof h1 as h1'. rewrite length_seg_list in h1'. unfold right_shift_list_nat, cycled_list_nat. rewrite map_seg_eq_iff. ex_goal. rewrite length_seg_list, app_length, seq_length; simpl. rewrite S_compat. rewrite <- (S_pred _ _ h1'); auto. exists hex. intros m h2 h3. unfold shift_functor. rewrite S_compat. rewrite (nth_lt_eq_iff nat_eq_dec). ex_goal. pose proof h2 as h4. rewrite length_seg_list in h4. apply lt_dec in h4. destruct h4 as [h4 | h4]. pose proof @nth_lt_app1. pose proof (nth_lt_app1 (seq 1 (pred n)) (0::nil) m) as h5. hfirst h5. rewrite seq_length; auto. specialize (h5 hf h3). rewrite h5. rewrite nth_lt_seq. rewrite plus_comm, S_compat. rewrite in_seg_list_iff. omega. generalize h3. rewrite h4. intro h5. pose proof (nth_lt_app2 (seq 1 (pred n)) (0::nil) (pred n)) as h6. hfirst h6. rewrite seq_length; auto. specialize (h6 hf h5). rewrite h6. simpl. rewrite seq_length, minus_same. rewrite in_seg_list_iff; auto. exists hex0. exists 0. split. rewrite count_occ_no_dup; auto. apply no_dup_seg_list. change (lind nat_eq_dec (seg_list n) (nth_lt (seq 1 (pred n) ++ 0 :: nil) m h3) hex0 = S m mod length (seg_list n)). pose proof h2 as h4. rewrite length_seg_list in h4. apply lt_dec in h4. destruct h4 as [h4 | h4]. gen0. pose proof (nth_lt_app1 (seq 1 (pred n)) (0::nil) m) as h6. hfirst h6. rewrite seq_length; auto. specialize (h6 hf h3). rewrite h6. rewrite nth_lt_seq, plus_comm, S_compat. intro h5. rewrite lind_seg_list, lt_mod_eq; auto. rewrite length_seg_list; auto. rewrite in_seg_list_iff in h5; auto. unfold h3. unfold h3 in hex0. clear h3. subst. gen0. erewrite nth_lt_app2. intro h5. rewrite lind_eq_iff. split. rewrite nth_lt_seg_list. symmetry. gen0. rewrite seq_length, <- minus_n_n. intro; simpl. rewrite <- (S_pred _ _ h1'), length_seg_list, mod_same; auto. intro h6. rewrite in_firstn_iff in h6. destruct h6 as [i [h6 h7]]. rewrite nth_lt_seg_list in h7. symmetry in h7. revert h7. gen0. rewrite seq_length, <- minus_n_n. intro h7. simpl. intro; subst. rewrite <- (S_pred _ _ h1'), length_seg_list, mod_same in h6. apply lt_irrefl in h6; auto. Grab Existential Variables. rewrite <- (S_pred _ _ h1'), length_seg_list, mod_same; auto with arith. rewrite <- (S_pred _ _ h1'), length_seg_list, mod_same; auto. rewrite seq_length; auto. Qed. Lemma transposed_list_nat_seq_app_0n : forall n (pf:0 < pred n) (pfp:pred n < length (seq 1 (pred n) ++ 0 :: n :: nil)) (pfn: n < length (seq 1 (pred n) ++ 0 :: n :: nil)), transposed_list_nat (seq 1 (pred n) ++ 0 :: n :: nil) (pred n) n pfp pfn = seq 1 n ++ 0::nil. intros n h1 h2 h3. unfold transposed_list_nat. rewrite map_seg_eq_iff. ex_goal. do 2 rewrite app_length; simpl. do 2 rewrite seq_length. omega. exists hex. intros m h4; simpl. pose proof h4 as h4'. rewrite app_length, seq_length in h4'. simpl in h4'. pose proof (O_lt_pred_n_lt_n _ h1) as hc. rewrite plus_S_shift_r, S_compat, <- (S_pred _ _ hc) in h4'. unfold transpose_functor. destruct (eq_nat_dec m (pred n)) as [|hr]; subst. rewrite beq_nat_same. erewrite nth_lt_app2. erewrite nth_lt_app1, nth_lt_seq. simpl. rewrite seq_length. rewrite n_minus_pred_n; auto. rewrite <- (S_pred _ _ hc); auto. pose proof hr as hr'. rewrite <- neq_nat_compat in hr. rewrite hr. destruct (eq_nat_dec m (pred n)) as [|hr0]; subst. pose proof (lt_pred_n _ _ hc) as hp. apply lt_neq in hp. rewrite <- neq_nat_compat in hp. rewrite hp. do 2 erewrite nth_lt_app2. gen0. rewrite seq_length, <- minus_n_n. intro. symmetry. simpl. rewrite seq_length. pose proof hc as h5. apply lt_le_weak in h5. rewrite le_minus_iff in h5. rewrite h5; auto. destruct (eq_nat_dec m n) as [|h5]; subst. rewrite beq_nat_same. erewrite nth_lt_app2. erewrite nth_lt_app2. simpl. rewrite seq_length, minus_same, seq_length, minus_same; auto. pose proof h5 as h5'. rewrite <- neq_nat_compat in h5. rewrite h5. erewrite nth_lt_app1. erewrite nth_lt_app1. do 2 rewrite nth_lt_seq; auto. Grab Existential Variables. rewrite seq_length. apply lt_dec in h4'. simpl in h4'. destruct h4' as [he | he]; subst; auto. contradict hr0; auto. contradict h5'; auto. rewrite seq_length. apply lt_dec in h4'. simpl in h4'. destruct h4' as [he | he]; subst; auto. apply lt_dec in he. destruct he; subst; auto. contradict hr; auto. contradict h5'; auto. rewrite seq_length; auto. rewrite seq_length; auto. rewrite seq_length; auto. contradict hr'; auto. contradict hr'; auto. rewrite seq_length; auto. rewrite seq_length; auto with arith. Qed. (*Perhaps erase this, the model for the above, if it's not needed!*) Lemma transposed_list_nat_seq_app0 : forall n (pf:0 < (pred (pred n))) (pfpp:pred (pred n) < (length (seq 1 (pred n) ++ 0 :: nil))) (pfp: pred n < (length (seq 1 (pred n) ++ 0 :: nil))), transposed_list_nat (seq 1 (pred n) ++ 0 :: nil) (pred (pred n)) (pred n) pfpp pfp = seq 1 (pred (pred n)) ++ 0::pred n::nil. intros n h1 h2 h3. unfold transposed_list_nat. rewrite map_seg_eq_iff. ex_goal. do 2 rewrite app_length; simpl. do 2 rewrite seq_length. omega. (*maybe speed up!*) exists hex. intros m h4; simpl. pose proof h4 as h4'. rewrite app_length, seq_length in h4'. simpl in h4'. rewrite S_compat in h4'. pose proof (O_lt_pred_n_lt_n _ h1) as ha. pose proof (lt_trans _ _ _ h1 ha) as hb. pose proof (O_lt_pred_n_lt_n _ hb) as hc. pose proof (lt_trans _ _ _ hb hc) as hd. rewrite <- (S_pred _ _ hd) in h4'. unfold transpose_functor. destruct (eq_nat_dec m (pred (pred n))) as [|hm]; subst. rewrite beq_nat_same. do 2 erewrite nth_lt_app2. gen0. rewrite seq_length, <- minus_n_n. intro h5. symmetry. gen0. rewrite seq_length, <- minus_n_n. intro; simpl; auto. pose proof hm as hm'. rewrite <- neq_nat_compat in hm'. rewrite hm'. destruct (eq_nat_dec m (pred n)) as [|hr0]; subst. rewrite beq_nat_same. erewrite nth_lt_app1, nth_lt_seq. erewrite nth_lt_app2. symmetry. gen0. rewrite seq_length. pose proof (n_minus_pred_n (pred n) h1) as h8. rewrite h8. intro h9. simpl. pose proof (O_lt_pred_n_lt_n _ h1) as h6. pose proof h6 as h7. apply lt_le_weak in h6. rewrite le_minus_iff in h6. rewrite <- (S_pred _ _ h7); auto. pose proof hr0 as hr0'. rewrite <- neq_nat_compat in hr0'. rewrite hr0'. do 2 erewrite nth_lt_app1, nth_lt_seq; auto. Grab Existential Variables. rewrite seq_length. apply lt_dec in h4'. destruct h4' as [h5 | h5]; subst. apply lt_dec in h5. destruct h5 as [h5 | h5]; subst; auto. contradict hm; auto. contradict hr0; auto. rewrite seq_length. apply lt_dec in h4'. destruct h4' as [h5 | h5]; subst; auto. contradict hr0; auto. rewrite seq_length; auto with arith. rewrite seq_length; auto. rewrite seq_length; auto. rewrite seq_length; auto. Qed. Lemma permuted_seg_list_cons'' : forall (l:list nat) (a:nat) (pfi:In (length l) l) (pf:a < length l), ~In a l -> permuted_seg_list (list_subst l _ (lt_lind nat_eq_dec _ _ pfi) a) -> permuted_seg_list (a::l). intros l a h1 h2 h3 h4. pose proof (transposed_list_nat_undoes_itself (a::l) 0 (S (lind nat_eq_dec _ _ h1)) (lt_0_Sn _)) as h5. hfirst h5. simpl. apply lt_n_S. apply lt_lind. specialize (h5 hf). hfirst h5. rewrite length_transposed_list_nat. apply lt_0_Sn. specialize (h5 hf0). hfirst h5. rewrite length_transposed_list_nat. simpl. apply lt_n_S. apply lt_lind. specialize (h5 hf1). rewrite <- h5. constructor. rewrite transposed_list_nat_cons_0_S_lind. gterm0. redterm0 c. redterm1 c0. rewrite <- (length_list_subst l _ c1) at 1. apply permuted_seg_list_cons'. gterm0. redterm1 c2. fold c3. assert (h6:c3 = lt_lind _ _ _ _). apply proof_irrelevance. rewrite h6; auto. Qed. (*Transposes the maximum number and the maximally-indexed element, then removes the maximum numbe, i.e. removes [pred n] with [lind _ (pred n) _ _] .*) Definition removelast_list_to_set_seg (l:list nat) (pfnd:NoDup l) (n:nat) (pfn:0 < n) (pfe:list_to_set l = seg n) : list nat := let pfi := iff2 (list_to_set_in_iff l _) (in_eq_set _ _ (eq_sym pfe) _ (in_pred_seg _ pfn)) in removelast (transposed_list_nat l (lind nat_eq_dec l (pred n) pfi) (pred n) (lt_lind nat_eq_dec l (pred n) pfi) (list_to_set_eq_seg_lt' _ _ _ pfnd pfe pfi)). Lemma removelast_list_to_set_seg_functional : forall (l l':list nat) (pfnd: NoDup l) (pfnd':NoDup l') (n:nat) (pfn:0 < n) (pfe: list_to_set l = seg n) (pfe':list_to_set l' = seg n), l = l' -> removelast_list_to_set_seg l pfnd n pfn pfe = removelast_list_to_set_seg l' pfnd' n pfn pfe'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma length_removelast_list_to_set_seg : forall (l:list nat) (pfnd:NoDup l) (n:nat) (pfn:0 < n) (pfe:list_to_set l = seg n), length (removelast_list_to_set_seg l pfnd n pfn pfe) = pred n. unfold removelast_list_to_set_seg; intros. rewrite length_removelast, length_transposed_list_nat. pose proof (list_to_set_eq_seg _ _ pfnd pfe); subst; auto. Qed. Lemma no_dup_removelast_list_to_set_seg : forall l (pfnd:NoDup l) n (pfn:0 < n) (pfe:list_to_set l = seg n), NoDup (removelast_list_to_set_seg l pfnd n pfn pfe). intros. unfold removelast_list_to_set_seg. apply no_dup_removelast. apply nat_eq_dec. apply no_dup_transposed_list_nat; auto. Qed. Lemma seg_list_pred_eq_removelast_to_set_seg : forall (l:list nat) (pfnd:NoDup l) (pfl:0 < length l) (pfe:list_to_set l = seg (length l)), seg_list (pred (length l)) = removelast_list_to_set_seg l pfnd (length l) pfl pfe -> let pfll := iff2 (list_to_set_in_iff l (pred (length l))) (in_eq_set (seg (length l)) (list_to_set l) (eq_sym pfe) (pred (length l)) (in_pred_seg (length l) pfl)) in l = transposed_list_nat (seg_list (length l)) (lind nat_eq_dec l (pred (length l)) pfll) (pred (length l)) (lt_eq_trans _ _ _ (lt_lind nat_eq_dec _ _ pfll) (eq_sym (length_seg_list _))) (lt_eq_trans _ _ _ (lt_pred_n _ _ pfl) (eq_sym (length_seg_list _))). intros l hnd h1 h2 h3. unfold removelast_list_to_set_seg in h3. symmetry in h3. rewrite removelast_eq_iff in h3. erewrite lst_compat' in h3 at 1. rewrite last_transposed_list_nat_pred in h3. rewrite lind_compat in h3. rewrite <- seg_list_S in h3. pose proof (transposed_list_nat_functional _ _ h3 (lind nat_eq_dec l (pred (length l)) (iff2 (list_to_set_in_iff l (pred (length l))) (in_eq_set (seg (length l)) (list_to_set l) (eq_sym h2) (pred (length l)) (in_pred_seg (length l) h1)))) (pred (length l))) as h4. hfirst h4. rewrite length_transposed_list_nat. apply lt_lind. specialize (h4 hf). hfirst h4. rewrite length_transposed_list_nat. eapply lt_pred_n. apply h1. specialize (h4 hf0). rewrite transposed_list_nat_undoes_itself in h4. simpl. rewrite h4 at 1. erewrite transposed_list_nat_functional. f_equal. apply proof_irrelevance. apply proof_irrelevance. Grab Existential Variables. rewrite <- (S_pred _ _ h1); auto. intro h4. rewrite transposed_list_nat_eq_nil_iff in h4. pose proof (lt_length _ _ h1); contradiction. Qed. Lemma nth_lt_removelast_list_to_set_seg : forall l (pfnd:NoDup l) (pflt:0 < length l) (pfe:list_to_set l = seg (length l)) (r:nat) (pfr:r < length (removelast_list_to_set_seg l pfnd (length l) pflt pfe)), r = lind nat_eq_dec l (pred (length l)) (iff2 (list_to_set_in_iff l (pred (length l))) (in_eq_set _ _ (eq_sym pfe) _ (in_seg _ _ (lt_pred_n _ (length l) pflt)))) \/ nth_lt (removelast_list_to_set_seg l pfnd (length l) pflt pfe) r pfr = nth_lt l r (lt_trans _ _ _ (lt_eq_trans _ _ _ pfr (length_removelast_list_to_set_seg l pfnd (length l) pflt pfe)) (lt_pred_n _ (length l) pflt)). intros l hnd hlt he r hr. unfold removelast_list_to_set_seg. erewrite nth_lt_removelast. unfold transposed_list_nat. rewrite nth_lt_map_seg'. unfold transpose_functor. (*One day clean this up with a match goal or tactic!!!*) destruct (eq_nat_dec r (lind nat_eq_dec _ _ (iff2 (list_to_set_in_iff l (pred (length l))) (in_eq_set (seg (length l)) (list_to_set l) (eq_sym he) (pred (length l)) (in_pred_seg (length l) hlt))))) as [ha | ha]; subst. rewrite beq_nat_same. left. apply lind_functional3. pose proof ha as ha'. rewrite <- neq_nat_compat in ha'. rewrite ha'. destruct (eq_nat_dec r (pred (length l))) as [hc | hc]; subst. rewrite beq_nat_same. right. pose proof hr as h2. rewrite length_removelast_list_to_set_seg in h2. omega. pose proof hc as hc'. rewrite <- neq_nat_compat in hc'. rewrite hc'. right. apply nth_lt_functional3. apply nat_eq_dec. Grab Existential Variables. rewrite length_transposed_list_nat. rewrite length_removelast_list_to_set_seg in hr; omega. Qed. Inductive permuted_lists {T:Type} : list T -> list T -> Prop := | permuted_lists_nil : permuted_lists nil nil | permuted_lists_intro : forall (l l':list T) (l0:list nat) r s (pfr:r < length l0) (pfs:s < length l0) (pfp:permuted_seg_list l0) (pfe:length l0 = length l'), permuted_lists l' (map_in_dep (fun i (pfi:In i l0) => let pfilt := lt_permuted_seg_list _ _ pfp pfi in let pflt := lt_eq_trans _ _ _ pfilt pfe in nth_lt l' i pflt)). bool2-v0-3/SetUtilities2.v0000644000175000017500000027035313575574055016254 0ustar dwyckoff76dwyckoff76(*Copyright (C) 2015-2019 Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import Setoid. Require Import Basics. Require Import NaryFunctions. Require Import Infinite_sets. Require Import NPeano. Require Import ProofIrrelevance. Require Import Description. Require Import SetUtilities. Require Import SetUtilities1_5. Require Import ArithUtilities. Require Import TypeUtilities. Require Import FiniteMaps. Require Import ListUtilities. Require Import DecidableDec. Require Import LogicUtilities. Require Import Infinite_sets. Require Import FunctionalExtensionality. Require Import FunctionProperties. Require Import WellOrders. Require Import Classical_Wf. Require Import FiniteBags. Require Import RelationClasses. Require Import WellOrders. Require Import EqDec. (*Version 0.3 bool2 -- certified -- Version 8.4 pl4 Coq.*) Section Transfer. Lemma extends_sig1_transfer_fun_sig_iff : forall {T U:Type} {A A':Ensemble T} (pf:A = A') (f:T->U) (g:sig_set A->U), extends_sig1 f g <-> extends_sig1 f (transfer_fun_sig pf g). intros T U A A' h1 f g. subst. rewrite transfer_fun_sig_eq_refl. tauto. Qed. Lemma extends_sig_sig_transfer_fun_compat : forall {T U:Type} {A B B':Ensemble T} {E:Ensemble (sig_set A)} (f:sig_set B->U) (g:sig_set E->U) (pf:B = B') (pfs: sig_set B = sig_set B'), extends_sig_sig f g <-> extends_sig_sig (transfer_fun pfs f) g. intros T U A B B' E f h1 h2 h3. subst. pose proof (proof_irrelevance _ h3 (eq_refl _)); subst. rewrite transfer_fun_eq_refl. tauto. Qed. Lemma extends_sig_sig_transfer_fun_sig1_iff : forall {T U:Type} {A B B':Ensemble T} {E:Ensemble (sig_set A)} (pf: B = B') (f:sig_set B->U) (g:sig_set E->U), extends_sig_sig f g <-> extends_sig_sig (transfer_fun_sig pf f) g. intros T U A B B' E h1 f g. subst. rewrite transfer_fun_sig_eq_refl. tauto. Qed. Lemma extends_sig_sig_transfer_fun_sig2_iff : forall {T U:Type} {A B:Ensemble T} {E E':Ensemble (sig_set A)} (pf: E = E') (f:sig_set B->U) (g:sig_set E->U), extends_sig_sig f g <-> extends_sig_sig f (transfer_fun_sig pf g). intros T U A B B' E h1 f g. subst. rewrite transfer_fun_sig_eq_refl. tauto. Qed. Lemma proj1_sig_snd_transfer_dep_incl : forall {T:Type} {A A':Ensemble T} (pf:A = A') (f:T*{C:Ensemble T | Included C A}-> T*{C:Ensemble T | Included C A}) (pr:T*{C:Ensemble T | Included C A'}), proj1_sig (snd (transfer_dep (U:=fun D => T*{C:Ensemble T | Included C D}-> T*{C:Ensemble T | Included C D}) pf f pr)) = proj1_sig (snd (f (transfer_dep_r (U := (fun D => (T*{C:Ensemble T | Included C D})%type)) pf pr))). intros; subst; rewrite transfer_dep_eq_refl, transfer_dep_r_eq_refl; auto. Qed. Lemma proj1_sig_snd_transfer_dep_incl' : forall {T:Type} {A A':Ensemble T} (pf:A = A') (f:T*{C:Ensemble T | Included C A}-> T*{C:Ensemble T | Included C A}) (pr:T*{C:Ensemble T | Included C A}), proj1_sig (snd (transfer_dep (U := (fun D => (T*{C:Ensemble T | Included C D})%type)) pf (f pr))) = proj1_sig (snd (f pr)). intros; subst; rewrite transfer_dep_eq_refl; auto. Qed. Lemma proj1_sig_snd_transfer_dep_incl_r : forall {T:Type} {A A':Ensemble T} (pf:A' = A) (f:T*{C:Ensemble T | Included C A}-> T*{C:Ensemble T | Included C A}) (pr:T*{C:Ensemble T | Included C A'}), proj1_sig (snd (transfer_dep_r (U:=fun D => T*{C:Ensemble T | Included C D}-> T*{C:Ensemble T | Included C D}) pf f pr)) = proj1_sig (snd (f (transfer_dep (U := (fun D => (T*{C:Ensemble T | Included C D})%type)) pf pr))). intros; subst; rewrite transfer_dep_eq_refl, transfer_dep_r_eq_refl; auto. Qed. Lemma proj1_sig_snd_transfer_dep_incl_r' : forall {T:Type} {A A':Ensemble T} (pf:A' = A) (f:T*{C:Ensemble T | Included C A}-> T*{C:Ensemble T | Included C A}) (pr:T*{C:Ensemble T | Included C A}), proj1_sig (snd (transfer_dep_r (U := (fun D => (T*{C:Ensemble T | Included C D})%type)) pf (f pr))) = proj1_sig (snd (f pr)). intros; subst; rewrite transfer_dep_r_eq_refl; auto. Qed. Lemma im_full_sig_functional : forall {T U:Type} {A B:Ensemble T} (f:sig_set A -> U) (pf:A = B), Im (full_sig A) f = Im (full_sig B) (transfer_fun_sig pf f). intros T U A B f h1. subst. rewrite transfer_fun_sig_eq_refl; auto. Qed. Lemma im_ext_full_hetero : forall {T U V:Type} (pf:T = U) (f:T->V) (g:U->V), (forall x:T, f x = g (transfer pf x)) -> Im (Full_set T) f = Im (Full_set U) g. intros T U V h1 f g h2; subst. apply im_ext_in. intros x h3. rewrite h2. rewrite transfer_eq_refl; auto. Qed. Lemma im_ext_full_hetero_set1 : forall {T U:Set} {V:Type} (pf:T = U) (f:T->V) (g:U->V), (forall x:T, f x = g (s_transfer pf x)) -> Im (Full_set T) f = Im (Full_set U) g. intros T U V h1 f g h2; subst. apply im_ext_in. intros x h3. rewrite h2. rewrite s_transfer_eq_refl; auto. Qed. Lemma segT_functional_id : forall {m n:nat} (pf:m = n) (x:m @), let pfeq := segT_functional _ _ pf in proj1_sig x = proj1_sig (s_transfer pfeq x). intros m n h1 x; subst; simpl. destruct x as [x hx]. simpl. rewrite segT_functional_eq_refl. rewrite s_transfer_eq_refl; auto. Qed. Lemma segT_functional_id_r : forall {m n:nat} (pf:n = m) (x:m @), let pfeq := segT_functional _ _ pf in proj1_sig x = proj1_sig (s_transfer_r pfeq x). intros m n h1 x; subst; simpl. destruct x as [x hx]. simpl. rewrite segT_functional_eq_refl. rewrite s_transfer_r_eq_refl; auto. Qed. Lemma proj1_segT_transfer : forall {m n:nat} (pf:m = n) (x:m @) (pfeq:m @ = n @), proj1_sig x = proj1_sig (s_transfer pfeq x). intros m n h1 x h2; subst; simpl. rewrite (segT_functional_id (eq_refl)). repeat f_equal. apply proof_irrelevance. Qed. Lemma proj1_segT_transfer_r : forall {m n:nat} (pf:n = m) (x:m @) (pfeq:n @ = m @), proj1_sig x = proj1_sig (s_transfer_r pfeq x). intros m n h1 x h2; subst; simpl. rewrite (segT_functional_id_r (eq_refl)). repeat f_equal. apply proof_irrelevance. Qed. Lemma seg_one_t_functional_id : forall {m n:nat} (pf:m = n) (x:seg_one_t m), let pfeq := seg_one_t_functional _ _ pf in proj1_sig x = proj1_sig (s_transfer pfeq x). intros m n h1 x; subst; simpl. rewrite seg_one_t_functional_eq_refl; auto. Qed. Lemma proj1_seg_one_t_transfer : forall {m n:nat} (pf:m = n) (x:seg_one_t m) (pfeq:seg_one_t m = seg_one_t n), proj1_sig x = proj1_sig (s_transfer pfeq x). intros m n h1 x h2; subst; simpl. rewrite (seg_one_t_functional_id (eq_refl)). repeat f_equal. apply proof_irrelevance. Qed. End Transfer. Section Card. Lemma card_fun1_full_sig : forall {T:Type} (S:Ensemble T), Finite S -> card_fun1 (full_sig S) = card_fun1 S. intros T S h0. pose proof (im_proj1_sig_full_sig S) as h2. rewrite h2 at 3. eapply injective_preserves_card_fun1. red; intros; apply proj1_sig_injective; auto. rewrite <- finite_full_sig_iff. assumption. Qed. Lemma card_fun1_im_proj2_sig : forall {T:Type} {P:T->Prop} (A:Ensemble T) (pfx:(forall x:T, Inn A x -> P x)), Finite A -> card_fun1 A = card_fun1 (im_proj2_sig A pfx). intros T P A h1 h2. rewrite <- card_fun1_full_sig; auto. apply injective_preserves_card_fun1. red. intros x y h3. apply exist_injective in h3. apply proj1_sig_injective; auto. rewrite <- finite_full_sig_iff; auto. Qed. Lemma le_im_card' : forall {T U:Type} {A:Ensemble T} (f:sig_set A-> U), Finite A -> card_fun1 (Im (full_sig A) f) <= card_fun1 A. intros T U A f h1 . pose proof h1 as h1'. rewrite finite_full_sig_iff in h1'. pose proof (le_im_card _ f h1') as h3. rewrite card_fun1_full_sig in h3; auto. Qed. Lemma preserves_card_fun1_injective : forall {T U:Type} {A:Ensemble T} (f:sig_set A -> U), Finite A -> card_fun1 A = card_fun1 (Im (full_sig A) f) -> injective f. intros T U A f h0 h1. red; intros x y h4. apply NNPP. intro ha. destruct x as [x h5], y as [y h6]. assert (h7 : Included (Couple x y) A). red; intros a h7; destruct h7; auto. pose proof (decompose_setminus_inc _ _ h7) as h8. pose proof (setminus_inc A (Couple x y)) as h9. pose proof (f_equal card_fun1 h8) as h8'. rewrite card_disj_union'' in h8'. rewrite card_fun1_couple in h8'. pose proof (im_full_sig_union_eq' f h7 h9 h8) as h10. termr h10. redterm0 y0. rename c into S. redterm1 y0. rename c into R. assert (h11 : R = Singleton (f (exist _ _ h5))). unfold R, restriction_sig. apply Extensionality_Ensembles; red; split; red; intros a h12. destruct h12 as [a h12]. clear h8. subst. clear h12. destruct a as [a h11]. destruct h11; simpl. pose proof (proof_irrelevance _ h5 (h7 x (Couple_l T x y)));subst. constructor. rewrite h4 at 1. pose proof (proof_irrelevance _ h6 (h7 y (Couple_r T x y))); subst. constructor. destruct h12. apply Im_intro with (exist _ _ (Couple_l _ x y)). constructor. f_equal. apply proj1_sig_injective; auto. unfold R in h11. rewrite h11 in h10. destruct (classic_dec (Inn S (f (exist (Inn A) x h5)))) as [h13 | h13]. pose proof (inclusion_iff_union (Singleton (f (exist (Inn A) x h5))) S) as h14. assert (h15 : Included (Singleton (f (exist (Inn A) x h5))) S). red; intros a h16. destruct h16. apply h13. rewrite h14 in h15. unfold S in h15. rewrite h15 in h10. pose proof (f_equal card_fun1 h10) as h16. rewrite h16 in h1. rewrite h8' in h1. symmetry in h1. pose proof (finite_setminus A (Couple x y) h0) as h17. redterm0 S. rename S into W. redterm0 W. rename c0 into g. pose proof (le_im_card' g h17) as h18. unfold g in h18. omega. pose proof (f_equal card_fun1 h10) as h16. rewrite card_disj_union'' in h16. rewrite card_fun1_sing in h16. rewrite h16 in h1. rewrite h8' in h1. pose proof (finite_setminus A (Couple x y) h0) as h17. rename S into W. redterm0 W. redterm0 W. rename c0 into g. pose proof (le_im_card' g h17) as h19. unfold g in h19. omega. apply Singleton_is_finite. apply finite_image. rewrite <- finite_full_sig_iff. apply finite_setminus; auto. apply Extensionality_Ensembles; red; split; red; intros c h17. destruct h17 as [c h17 h18]. destruct h17. clear h8. inversion h18 as [c' h19 ? h20]. subst. destruct c' as [c' h21]. unfold restriction_sig in h20. simpl in h20. contradiction. destruct h17. intro. clear h8. subst. contradict ha. apply proj1_sig_injective; auto. apply finite_couple. apply finite_setminus; auto. apply Extensionality_Ensembles; red; split; red; intros c h17. destruct h17 as [c h17 h18]. destruct h17. destruct h18 as [h18 h19]. contradict h19. left. destruct h18 as [h18 h19]. contradict h19. right. destruct h17. Qed. Lemma surj_impl_inj_sig_set_eq_card : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} (f:sig_set A -> sig_set B), Finite A -> Finite B -> card_fun1 A = card_fun1 B -> surjective f -> injective f. intros T U A B f h1 h2 h3 h4. pose proof h4 as h4'. rewrite surj_iff' in h4'. unfold im_proj1_sig in h4'. rewrite im_im in h4'. rewrite <- h4' in h3. red; intros x y h5. destruct x as [x h6]. destruct y as [y h7]. pose proof h2 as h2'. rewrite <- h4' in h2'. pose proof (proj2_sig (f (exist _ _ h6))) as h8; simpl in h8. rewrite <- h4' in h8 at 1. pose proof (proj2_sig (f (exist _ _ h7))) as h9; simpl in h9. rewrite <- h4' in h9 at 1. pose proof (card_fun1_eq_sub _ _ h1 h2' h3 x (proj1_sig (f (exist _ _ h6))) h6 h8) as h10. pose proof (preserves_card_fun1_injective (fun x : sig_set A => proj1_sig (f x)) h1 h3) as h11. red in h11. pose proof (f_equal (@proj1_sig _ _) h5) as h5'. specialize (h11 (exist _ _ h6) (exist _ _ h7) h5'). assumption. Qed. Lemma surj_impl_bij_sig_set_eq_card : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} (f:sig_set A -> sig_set B), Finite A -> Finite B -> card_fun1 A = card_fun1 B -> surjective f -> bijective f. intros; split; try apply surj_impl_inj_sig_set_eq_card; auto. Qed. Lemma inj_impl_surj_sig_set_eq_card : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} (f:sig_set A -> sig_set B), Finite A -> Finite B -> card_fun1 A = card_fun1 B -> injective f -> surjective f. intros T U A B f h1 h2 h3 h4. rewrite surj_iff'. unfold im_proj1_sig. rewrite im_im. pose proof (injective_preserves_card_fun1 (fun x => proj1_sig (f x))) as hi. hfirst hi. red; intros x y h5. destruct x as [x h6], y as [y h7]. simpl in h5. apply proj1_sig_injective in h5. red in h4. apply h4 in h5. assumption. apply incl_card_fun1_eq; auto. red; intros y h6. destruct h6 as [y h6]. subst. apply proj2_sig. pose proof h1 as h1'. rewrite finite_full_sig_iff in h1'. specialize (hi hf). pose proof (injective_preserves_card_fun1 _ hf (full_sig A) h1') as h6. rewrite <- h6. rewrite card_fun1_full_sig; auto. Qed. Lemma inj_impl_bij_sig_set_eq_card : forall {T U:Type} {A:Ensemble T} {B:Ensemble U} (f:sig_set A -> sig_set B), Finite A -> Finite B -> card_fun1 A = card_fun1 B -> injective f -> bijective f. intros; split; try apply inj_impl_surj_sig_set_eq_card; auto. Qed. Lemma card_pairwise_disjoint : forall {T:Type} (F:Family T) (pff:Finite F), (forall S:Ensemble T, Inn F S -> Finite S) -> pairwise_disjoint F -> card_fun1 (FamilyUnion F) = plus_bag_nat (im_set_bag nat_eq_dec F pff card_fun1). intros T F h1. pose proof (finite_set_list_no_dup _ h1) as h2. destruct h2 as [l h2]. destruct h2 as [h2l h2r]. generalize dependent F. revert h2r. induction l as [|A l h1]; simpl. intros h1 F h2 h3 h4 h5. subst. pose proof (im_set_bag_empty nat_eq_dec (@card_fun1 T)) as him. pose proof (proof_irrelevance _ h2 (Empty_is_finite (Ensemble T))); subst. rewrite im_set_bag_empty. rewrite plus_bag_nat_empty. rewrite empty_family_union. apply card_fun1_empty. intros h2 F h4 h5 hin h6. pose proof (no_dup_cons_nin _ _ h2) as h3. pose proof (f_equal (fun S=>Subtract S A) h5) as h7. simpl in h7. rewrite list_to_set_in_iff in h3. rewrite sub_add_same_nin in h7; auto. pose proof (no_dup_cons _ _ h2) as h8. pose proof (subtract_preserves_finite F A h4) as h9. pose proof (incl_subtract F A) as h11. specialize (h1 h8 _ h9 h7). hfirst h1. intros S h10. specialize (hin _ (h11 _ h10)); auto. specialize (h1 hf (pairwise_disjoint_incl _ _ _ h6 h11)). subst. rewrite family_union_add. pose proof (sub_add_same_nin _ _ h3) as heq. rewrite <- heq at 1; auto. pose proof (hin _ (Add_intro2 _ (list_to_set l) A)) as h13. pose proof (Finite_Finite_Union (list_to_set l)) as h16. hfirst h16. intros S h15. apply hin; auto. left; auto. specialize (h16 hf0). hfirst h16. eapply Finite_downward_closed. apply h4. auto with sets. specialize (h16 hf1). pose proof (sub_add_same_nin _ _ h3) as h17. rewrite h17. pose proof (proof_irrelevance _ h4 (Add_preserves_Finite _ _ _ hf1)); subst. rewrite plus_bag_nat_im_set_bag_add. pose proof (pairwise_disjoint_add_family_union _ _ h3 h6) as h18. red in h18. rewrite Intersection_commutative in h18. pose proof (card_disj_union' _ _ h13 h16 h18) as h19. pose proof (card_fun_card_fun1_compat _ (Union_preserves_Finite _ _ _ h13 h16)) as h21. rewrite <- h21. rewrite <- h19. simpl in h17. simpl in h1. rewrite h17 in h1 at 1. pose proof (subsetT_eq_compat _ _ _ _ h9 hf1 h17) as h22. dependent rewrite -> h22 in h1. rewrite <- h1. rewrite (card_fun_card_fun1_compat _ h13). rewrite (card_fun_card_fun1_compat _ h16). ring. assumption. Qed. (*See subsequent lemma for easing the finitude to just one set.*) Lemma bij_ex_impl_eq_card : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (pfa:Finite A) (pfb:Finite B) (P:T->U->Prop), (forall x:T, Inn A x -> exists! y:U, Inn B y /\ P x y) -> (forall y:U, Inn B y -> exists! x:T, Inn A x /\ P x y) -> card_fun1 A = card_fun1 B. intros T U hdt A B h1 h2 P h3 h4. destruct (finite_inh_or_empty_dec _ h2) as [hinh | h5]. destruct hinh as [d hinh]. pose [pr:T*U | exists (pfa:Inn A (fst pr)), exists (pfb:Inn B (snd pr)), proj1_sig (constructive_definite_description _ (h3 _ pfa)) = snd pr /\ proj1_sig (constructive_definite_description _ (h4 _ pfb)) = fst pr] as S. assert (h5:fun_well_defined S). red. split. intros x h5. destruct h5 as [h5]. destruct h5 as [y h5]. exists y. red. split. split; auto. constructor. exists x. assumption. intros y' h6. destruct h6 as [h6l h6r]. destruct h5 as [h5]. destruct h5 as [h5l h5r]. destruct h5r as [h5a h5b]. simpl in h5l. simpl in h5b. simpl in h5a. destruct constructive_definite_description as [b h8]. destruct constructive_definite_description as [a h9]. simpl in h5b. destruct h5b; subst. destruct h9 as [h9l h9r]. destruct h8 as [h10l h10r]. destruct h6r as [h6r]. destruct h6r as [h11 h12]. destruct h12 as [h12 h13]. destruct constructive_definite_description as [b' h8']. destruct constructive_definite_description as [a' h9']. simpl in h13. destruct h13; subst. simpl in h9'. simpl in h8'. destruct h9' as [h9a' h9b']. destruct h8' as [h8a' h8b']. pose proof (h3 _ h5l) as h10. destruct h10 as [y'' h10]. red in h10. destruct h10 as [h10a h10b]. destruct h10a as [h10a' h10b']. pose proof (h10b _ (conj h8a' h8b')). subst. pose proof (h10b _ (conj h10l h10r)). subst. reflexivity. intros pr h5. destruct h5 as [h5]. destruct h5 as [h5 h6]. destruct h6 as [h6 h7]. destruct constructive_definite_description as [b h8]. destruct constructive_definite_description as [a h9]. simpl in h7. destruct h7; subst. split. constructor. exists (snd pr). constructor. simpl. exists h5. exists h6. split. destruct constructive_definite_description as [y' h10]. simpl. destruct h10 as [h10l h10r]. destruct h9 as [h9l h9r]. pose proof (h3 _ h5) as h11. destruct h11 as [y h11]. red in h11. destruct h11 as [h11l h11r]. destruct h11l as [h11a h11b]. pose proof (h11r _ (conj h10l h10r)). subst. pose proof (h11r _ h8). assumption. destruct constructive_definite_description as [a h10]. simpl. pose proof (h4 _ h6) as h11. destruct h11 as [a' h11]. red in h11. destruct h11 as [h11a h11b]. pose proof (h11b _ h10). subst. pose proof (h11b _ h9). assumption. constructor. exists (fst pr). constructor. simpl. exists h5. exists h6. destruct constructive_definite_description as [b h10]. destruct constructive_definite_description as [a h11]. simpl. split. pose proof (h3 _ h5) as h12. destruct h12 as [b' h12r]. red in h12r. destruct h12r as [h12a h12b]. pose proof (h12b _ h10); subst. pose proof (h12b _ h8); subst. reflexivity. pose proof (h4 _ h6) as h12. destruct h12 as [a' h12r]. red in h12r. destruct h12r as [h12a h12b]. pose proof (h12b _ h11); subst. pose proof (h12b _ h9); subst. reflexivity. red in h5. assert (h6:dom_rel S = A). apply Extensionality_Ensembles. red. split. red. intros x h6. destruct h6 as [h6]. destruct h6 as [y h6]. destruct h6 as [h6]. simpl in h6. destruct h6 as [h6]. assumption. red. intros x h6. constructor. exists (proj1_sig (constructive_definite_description (fun y : U => Inn B y /\ P x y) (h3 _ h6))). constructor. simpl. exists h6. destruct constructive_definite_description as [y h7]. simpl. destruct h7 as [h7l h7r]. exists h7l. split; auto. destruct constructive_definite_description as [x' h8]. simpl. destruct h8 as [h8l h8r]. pose proof (h4 _ h7l) as h8. destruct h8 as [x'' h8]. red in h8. destruct h8 as [h8a h8b]. pose proof (h8b _ (conj h8l h8r)); subst. pose proof (h8b _ (conj h6 h7r)); subst. reflexivity. rewrite h6 in h5. assert (h7:ran_rel S = B). apply Extensionality_Ensembles. red. split. red. intros y h7. destruct h7 as [h7]. destruct h7 as [x h7]. destruct h7 as [h7]. destruct h7 as [h7 h8]. destruct h8 as [h8 h9]. simpl in h8. assumption. red. intros y h7. constructor. exists (proj1_sig (constructive_definite_description (fun x : T => Inn A x /\ P x y) (h4 _ h7))). constructor. simpl. destruct constructive_definite_description as [x h8]. simpl. destruct h8 as [h8l h8r]. exists h8l. exists h7. split. destruct constructive_definite_description as [y' h9]. simpl. pose proof (h3 _ h8l) as h10. destruct h10 as [y'' h10]. red in h10. destruct h10 as [h10l h10r]. pose proof (h10r _ (conj h7 h8r)). pose proof (h10r _ h9). congruence. destruct constructive_definite_description as [x' h9]. simpl. pose proof (h4 _ h7) as h10. destruct h10 as [x'' h10]. red in h10. destruct h10 as [h10l h10r]. pose proof (h10r _ h9). pose proof (h10r _ (conj h8l h8r)). congruence. rewrite h7 in h5. pose (fin_map_intro _ _ d hdt h1 h2 _ h5) as F. apply bij_dom_ran_card_eq' with F. red. split. red. intros x x' h8 h9 h10. pose proof (in_fin_map_to_fps F x h8) as h11. pose proof (in_fin_map_to_fps F x' h9) as h12. unfold F in h11. unfold F in h12. rewrite <- fin_map_to_fps_s_compat in h11. rewrite <- fin_map_to_fps_s_compat in h12. unfold F in h10. rewrite h10 in h11. destruct h11 as [h11]. simpl in h11. destruct h11 as [h11a h11b]. destruct h11b as [h11b h11c]. destruct constructive_definite_description as [b h13]. destruct (constructive_definite_description _ (h4 (fin_fps_to_f hdt h1 S h5 d x') h11b)) as [a h14]. simpl in h11c. destruct h11c as [h11d h11e]. destruct h13 as [h13l h13r]. pose proof (h4 _ h13l) as h15. destruct h15 as [x'' h15]. red in h15. destruct h15 as [h15l h15r]. pose proof (h15r _ (conj h8 h13r)). rewrite <- h11d in h14. pose proof (h15r _ h14). inversion h12 as [h16]. clear h12. simpl in h16. destruct h16 as [h16 h17]. destruct h17 as [h17 h18]. destruct constructive_definite_description as [b' h19]. destruct (constructive_definite_description _ (h4 (fin_fps_to_f hdt h1 S h5 d x') h17)) as [a' h20]. simpl in h18. destruct h18 as [h18l h18r]. rewrite <- h11d in h20. pose proof (h15r _ h20). congruence. red. intros b h8. exists (proj1_sig (constructive_definite_description _ (h4 _ h8))). split. destruct constructive_definite_description as [a h9]. simpl. destruct h9; auto. destruct constructive_definite_description as [a h9]. simpl. destruct h9 as [h9l h9r]. pose proof (fin_fps_to_f_s_compat hdt h1 h5 d _ h9l) as h10. destruct h10 as [h10]. simpl in h10. destruct h10 as [h10 h11]. destruct h11 as [h11 h12]. destruct constructive_definite_description as [y h13]. destruct (constructive_definite_description _ (h4 (fin_fps_to_f hdt h1 S h5 d a) h11)) as [x h14]. simpl in h12. destruct h12 as [h12l h12r]. pose proof (h3 _ h10) as h15. destruct h15 as [y' h15]. red in h15. destruct h15 as [h15l h15r]. rewrite <- h12r in h15r. destruct h14 as [h14l h14r]. pose proof (h15r _ (conj h11 h14r)). rewrite <- h12r in h13. pose proof (h15r _ h13). rewrite <- h12r in h9r. pose proof (h15r _ (conj h8 h9r)). congruence. subst. pose proof (eq_empty A) as h7. hfirst h7. red; intros a h8. specialize (h3 _ h8). destruct h3 as [y h3]. red in h3; destruct h3 as [h3 h9]. destruct h3; contradiction. apply h7 in hf; subst. do 2 rewrite card_fun1_empty; auto. Qed. (*Uses LEM*) (*Like above function, but with only one required finitude.*) Lemma bij_ex_impl_eq_card' : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (pfa:Finite A) (P:T->U->Prop), (forall x:T, Inn A x -> exists! y:U, Inn B y /\ P x y) -> (forall y:U, Inn B y -> exists! x:T, Inn A x /\ P x y) -> card_fun1 A = card_fun1 B. intros T U hdt A B h1 P h2 h3. assert (h4:Finite B). apply NNPP. intro h4. assert (h5:~FiniteT (sig_set B)). intro h5. apply FiniteT_sig_Finite in h5. contradiction. pose (fun x:sig_set B => (proj1_sig (constructive_definite_description _ (h3 _ (proj2_sig x))))) as f. rewrite <- Finite_FiniteT_iff in h5. assert (h6:Finite (Im (Full_set (sig_set B)) f)). eapply Finite_downward_closed. apply h1. red. intros x h6. destruct h6 as [x h6]. subst. unfold f. destruct constructive_definite_description as [a h7]. simpl. destruct h7; assumption. pose proof (Pigeonhole_bis _ _ _ _ h5 h6) as h7. contradict h7. red. intros x y h7. destruct x as [x h8]. destruct y as [y h9]. unfold f in h7. destruct constructive_definite_description as [a h10]. destruct constructive_definite_description as [a' h11]. simpl in h7. subst. simpl in h10. simpl in h11. pose proof (h3 _ h9) as h12. destruct h12 as [a h12]. red in h12. destruct h12 as [h12l h12r]. pose proof (h12r _ h11). subst. destruct h10 as [h10l h10r]. specialize (h2 _ h10l). destruct h2 as [b h2]. red in h2. destruct h2 as [h2l h2r]. destruct h12l as [h12a h12b]. pose proof (h2r _ (conj h8 h10r)). subst. pose proof (h2r _ (conj h9 h12b)). subst. apply proj1_sig_injective. simpl. reflexivity. apply (bij_ex_impl_eq_card hdt _ _ h1 h4 _ h2 h3). Qed. Lemma bij_eq_card : forall {T U:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (B:Ensemble U) (pfa:Finite A) (f:sig_set A -> sig_set B), bijective f -> card_fun1 A = card_fun1 B. intros T U h1 A B h2 f h3. red in h3. destruct h3 as [h3 h4]. red in h3, h4. pose (fun (x:T) (y:U) => exists pf:Inn A x, y = proj1_sig (f (exist _ _ pf))) as P. apply (bij_ex_impl_eq_card' h1 A B h2 P) ; auto. intros x h5. exists (proj1_sig (f (exist _ _ h5))). red. split. split. apply proj2_sig. unfold P. exists h5. reflexivity. intros y h6. destruct h6 as [h6 h7]. unfold P in h7. destruct h7 as [h7 h8]; subst. f_equal. f_equal. f_equal. apply proof_irrelevance. intros y h5. specialize (h4 (exist _ _ h5)). destruct h4 as [x h4]. pose proof (f_equal (@proj1_sig _ _) h4) as h4'; simpl in h4'; subst. destruct x as [x hx]. exists x. red; split. split. assumption. unfold P. exists hx. reflexivity. intros x' h6. destruct h6 as [h6 h7]. unfold P in h7. destruct h7 as [h7 h8]. apply proj1_sig_injective in h8. apply h3 in h8. apply exist_injective in h8; auto. Qed. End Card. Section Sort. (*These [pop_ens] lemmas are kind of deprecated, since you can achieve the same effect with [can_set_nat_choice].*) Lemma pop_ens_iter_aux_functional : forall {T:Type} (S S':Ensemble T) (pff:Finite S) (pff':Finite S') (pfi:Inhabited S) (pfi':Inhabited S') (R:T->T->Prop) (pfso:sortable R) (def:T) (pfeq:S = S'), pop_ens_iter_aux S pff pfi R pfso def = transfer_dep_r (U:=fun C => T*{A:Ensemble T | Included A C}-> T*{A:Ensemble T | Included A C}) pfeq (pop_ens_iter_aux S' pff' pfi' R pfso def). intros T S S' h1 h2 h3 h4 R h5 h6 h7. subst. pose proof (proof_irrelevance _ h1 h2); subst. rewrite transfer_dep_r_eq_refl; auto. Qed. Lemma fst_pop_ens_iter_aux_functional : forall {T:Type} (A A':Ensemble T) (pff:Finite A) (pff':Finite A') (pfi:Inhabited A) (pfi':Inhabited A') (R:T->T->Prop) (pfso:sortable R) (pft:Transitive R) (pfir:irrefl R) (def:T) (pfeq:A = A') (pr:(T*{C:Ensemble T | Included C A})), fst (pop_ens_iter_aux A pff pfi R pfso def pr) = fst (pop_ens_iter_aux A' pff' pfi' R pfso def (transfer_dep (U:=fun C:Ensemble T => (T*{B:Ensemble T | Included B C})%type) pfeq pr)). intros T A A' h1 h2 h3 h4 R h5 h6 h7 def h8 pr. subst. rewrite transfer_dep_eq_refl. pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. Definition pop_l_iter_aux {T:Type} (l:list T) (pf:l <> nil) (R:T->T->Prop) (pfso:sortable R) (def:T) := pop_ens_iter_aux (list_to_set l) (list_to_set_finite _) (not_nil_inhabited _ pf) R pfso def. Lemma pop_l_iter_aux_eq : forall {T:Type} (l:list T) (pf:l <> nil) (R:T->T->Prop) (pfso:sortable R) (pft:Transitive R) (pfir:irrefl R) (def:T), pop_l_iter_aux l pf R pfso def = pop_ens_iter_aux (list_to_set l) (list_to_set_finite _) (not_nil_inhabited _ pf) R pfso def. intros? l; induction l; try reflexivity; unfold pop_l_iter_aux; auto. Qed. Lemma pop_ens_iter_sing_S : forall {T:Type} (a:T) (pff:Finite (Singleton a)) (pfi:Inhabited (Singleton a)) (R:T->T->Prop) (pfso:sortable R) (def:T) (n:nat), pop_ens_iter (Singleton a) pff pfi R pfso def (S n) = def. intros T a h1 h0 R h3 def n. unfold pop_ens_iter. rewrite pop_ens_iter_aux_sing. gterml. redterm0 x. redterm1 c. redterm0 c0. simpl. lr_match_goal'; simpl; auto. destruct n as [|n]. simpl in hr. rewrite min_set_sortable_sing, subtract_sing in hr. pose proof (In_singleton _ a) as h6. rewrite <- hr in h6. contradiction. terml hr. redterm0 x0. redterm0 c2. redterm2 c3. redterm1 c3. pose (snd c5) as E. pose proof (snd_iter_uo_const c4 E) as h4; unfold c4, E in h4. hfirst h4. intro pr. lr_match_goal'; apply proj1_sig_injective; simpl. pose proof (min_set_sortable_compat (Singleton a) h1 h0 _ h3) as h5; simpl in h5. destruct h5 as [h5 h6]. inversion h5 as [h7]. rewrite h7 at 1. rewrite subtract_sing; auto. pose proof (min_set_sortable_compat (Singleton a) h1 h0 _ h3) as h5; simpl in h5. destruct h5 as [h5 h6]. inversion h5 as [h7]. rewrite h7 at 1. rewrite subtract_sing; auto. specialize (h4 hf c5 n). unfold c5 in h4. rewrite h4 in hr; simpl in hr. rewrite min_set_sortable_sing, subtract_sing in hr. symmetry in hr. apply sing_not_empty in hr; contradiction. Qed. (*The below is usually better than the above!*) (*This is to choose a canonical (least) finite set of natural numbers from a finite family.*) Definition can_set_nat_choice (F:Family nat) (pff:Finite F) (pfi:Inhabited F) (pfs:forall S:Ensemble nat, Inn F S -> Finite S) : Ensemble nat := let f := (fun S => sort_set_nat (proj1_sig S) (pfs _ (proj2_sig S))) in list_to_set (min_set_sortable (Im (full_sig F) f) (finite_image _ _ _ f (iff1 (finite_full_sig_iff F) pff)) (inh_im (full_sig F) f (inh_full_sig F pfi)) lt_list_nat lt_list_nat_sortable). Lemma in_can_set_nat_choice : forall (F:Family nat) (pff:Finite F) (pfi:Inhabited F) (pfs:forall S:Ensemble nat, Inn F S -> Finite S), Inn F (can_set_nat_choice F pff pfi pfs). intros F h1 h2 h3. unfold can_set_nat_choice. gterm0. redterm0 c. redterm4 c0. redterm3 c0. redterm2 c0. redterm1 c0. redterm0 c0. pose proof (min_set_sortable_compat c1 c2 c3 c4 c5) as h4. unfold c1, c2, c3, c4, c5 in h4. simpl in h4. destruct h4 as [h4 h5]. inversion h4 as [F' h8 h6 h7]; subst. clear h8. destruct F' as [F' h8]. simpl in h7. pose proof (in_list_reps_sort_set_nat F' (h3 F' h8)) as h9. destruct h9 as [h9]. destruct h9 as [h9 h10]. rewrite <- h7 in h9. subst. assumption. Qed. Lemma can_set_nat_choice_sing : forall (A:Ensemble nat) (pff:Finite (Singleton A)) (pfi:Inhabited (Singleton A)) (pfs:forall S:Ensemble nat, Inn (Singleton A) S -> Finite S), can_set_nat_choice (Singleton A) pff pfi pfs = A. intros A h1 h2 h3. unfold can_set_nat_choice. match goal with |- list_to_set (min_set_sortable _ ?pf1 ?pf2 _ _ ) = _ => generalize pf1 pf2 end. rewrite full_sig_sing. match goal with |- forall _ _, list_to_set (min_set_sortable ?C _ _ _ _) = _ => assert (h4:C = Singleton (sort_set_nat A (h3 A (In_singleton (Ensemble nat) A)))) end. rewrite im_singleton; auto. intros h5 h6. rewrite (min_set_sortable_functional _ _ h5 (Singleton_is_finite _ _) h6 (inh_sing _) _ _ lt_list_nat_sortable h4). rewrite min_set_sortable_sing. apply sort_set_nat_compat. Qed. Definition min_ens_list_nat (F:Ensemble (list nat)) (pff:Finite F) (pfi:Inhabited F) : list nat := min_set_sortable F pff pfi lt_list_nat lt_list_nat_sortable. Lemma ens_ex_dec : forall {T:Type} (P:T->Prop) (R:Relation T) (A:Ensemble T), sort_dec R -> Finite A -> pred_dec P -> {exists x:T, Inn A x /\ P x} + {~exists x:T, Inn A x/\ P x}. intros T P R A hr hfa h1. destruct (l_ex_dec (sort_dec_eq_dec _ hr) P (sort_set R hr _ hfa) h1) as [h3 | h3]. left. destruct h3 as [x [h3 h4]]. rewrite <- in_sort_set_iff in h3. exists x; split; auto. right. contradict h3. destruct h3 as [x [h3 h4]]. exists x. rewrite <- in_sort_set_iff. split; auto. Qed. Definition inv_im_min {T U:Type} (R:Relation T) (pfwo:well_order R) (f:T->U) (y:U) (pfs:surjective f) : T := min_set _ _ pfwo (inv_im1 f y) (inh_surj f y pfs). Definition inv_im_min_nat {U:Type} (f:nat->U) (y:U) (pf:surjective f) : nat := inv_im_min lt wo_nat f y pf. Lemma inv_im_min_compat : forall {T U:Type} (R:Relation T) (pfwo:well_order R) (f:T->U) (y:U) (pfs:surjective f) (x:T) (pf:Inn (inv_im1 f y) x), f x = y /\ Inn (inv_im1 f y) (inv_im_min R pfwo f y pfs) /\ ~ R x (inv_im_min R pfwo f y pfs). unfold inv_im_min; intros T U R h1 f y h2 x h3. pose proof (min_set_compat _ R h1 _ (inh_surj f y h2)) as h4. simpl in h4. destruct h4 as [h4 h5]. split; auto. pose proof (min_set_compat T R h1 (inv_im1 f y) (inh_surj f y h2)) as h6. simpl in h6. destruct h6 as [h6 h7]. inversion h6. destruct h3; auto. Qed. Lemma inv_im_min_compat' : forall {T U:Type} (R:Relation T) (pfwo:well_order R) (f:T->U) (y:U) (pfs:surjective f), f (inv_im_min R pfwo f y pfs) = y. intros. unfold inv_im_min. pose proof (min_set_compat T R pfwo (inv_im1 f y) (inh_surj f y pfs)) as h1. simpl in h1. destruct h1 as [h1 h2]. inversion h1; auto. Qed. Lemma in_inv_im_min : forall {T U:Type} (R:Relation T) (pfwo:well_order R) (f:T->U) (y:U) (pfs:surjective f), Inn (inv_im1 f y) (inv_im_min R pfwo f y pfs). intros T U R h1 f y h2. econstructor. apply inv_im_min_compat'. Qed. Lemma inv_im_min_nat_compat : forall {U:Type} (f:nat->U) (y:U) (pfs:surjective f) (n:nat) (pf:Inn (inv_im1 f y) n), f n = y /\ Inn (inv_im1 f y) (inv_im_min_nat f y pfs) /\ ~ n < (inv_im_min_nat f y pfs). unfold inv_im_min; intros; apply inv_im_min_compat; auto. Qed. Lemma inv_im_min_nat_compat' : forall {U:Type} (f:nat->U) (pf:surjective f) (y:U), f (inv_im_min_nat f y pf) = y. unfold inv_im_min_nat; intros; apply inv_im_min_compat'. Qed. Lemma in_inv_im_min_nat : forall {U:Type} (f:nat->U) (pf:surjective f) (y:U), Inn (inv_im1 f y) (inv_im_min_nat f y pf). intros; constructor; apply inv_im_min_nat_compat'. Qed. Lemma le_inv_im_min : forall {T U:Type} (R:Relation T) (pfwo:well_order R) (f:T->U) (pf:surjective f) (x:T), Rw R (inv_im_min R pfwo f (f x) pf) x. intros T U R h1 f h2 x. pose proof (inv_im_min_compat R h1 f (f x) h2 x) as h3. hfirst h3. econstructor; auto. specialize (h3 hf). destruct h3 as [h3 [h4 h5]]. destruct h1 as [h1 h6]. red in h6. pose proof (h6 (inv_im_min R {| wo_well_founded := h1; wo_total_strict_order := h6 |} f (f x) h2) x) as h7. destruct h7 as [h7 | [h7 | h7]]. left; auto. right; auto. contradiction. Qed. Lemma le_inv_im_min_nat : forall {U:Type} (f:nat->U) (pf:surjective f) (i:nat), inv_im_min_nat f (f i) pf <= i. unfold inv_im_min_nat; intros. rewrite <- Rw_le_iff. eapply le_inv_im_min. Qed. Lemma lind_map_seg_fun_to_seg_inv_im_min_nat : forall {U:Type} (pfdu:type_eq_dec U) (f:nat->U) (pfs:surjective f) (i:nat) (pf:In (f (inv_im_min_nat f (f i) pfs)) (map_seg (fun_to_seg f (S i)))), lind pfdu (map_seg (fun_to_seg f (S i))) (f (inv_im_min_nat f (f i) pfs)) pf = inv_im_min_nat f (f i) pfs. intros U hdu f h1 i h2. rewrite lind_eq_iff. split. erewrite nth_lt_functional3. rewrite nth_lt_map_seg. rewrite fun_to_seg_compat. reflexivity. intro h3. rewrite in_firstn_iff in h3. destruct h3 as [m [h3 h4]]. erewrite nth_lt_functional3 in h4. rewrite nth_lt_map_seg, fun_to_seg_compat in h4. rewrite inv_im_min_nat_compat' in h4. pose proof (inv_im_min_nat_compat f (f i) h1 m) as h5. hfirst h5. econstructor; auto. specialize (h5 hf). destruct h5 as [? [? ?]]; omega. Grab Existential Variables. eapply lt_le_trans. apply h3. eapply le_trans. apply le_inv_im_min_nat. auto with arith. rewrite length_map_seg. eapply le_trans. apply le_inv_im_min_nat. auto with arith. red. apply le_n_S. apply le_inv_im_min_nat. rewrite length_map_seg. red. apply le_n_S. apply le_inv_im_min_nat. Qed. Definition surj_nat_to_fun {U V:Type} (f:nat->U) (g:nat->V) (pfs:surjective f) : U -> V := fun y => g (inv_im_min_nat f y pfs). Lemma surj_nat_to_fun_compat : forall {U V:Type} (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (f:nat->U) (g:nat->V) (pfs:surjective f) (pfs':surjective g), fun_linds_eq_nat pfdu pfdv f g -> forall i:nat, surj_nat_to_fun f g pfs (f i) = g i. intros U V h1 h2 f g h3 h3' h4 i. unfold surj_nat_to_fun. pose proof (lind_map_seg_fun_to_seg_inv_im_min_nat h1 f h3 i) as h5. hfirst h5. rewrite in_map_seg_iff. exists (inv_im_min_nat f (f i) h3). ex_goal. red. apply le_n_S. apply le_inv_im_min_nat. exists hex. unfold fun_to_seg; auto. specialize (h5 hf). pose proof (lind_map_seg_fun_to_seg_inv_im_min_nat h2 g h3' i) as h5'. hfirst h5'. rewrite in_map_seg_iff. exists (inv_im_min_nat g (g i) h3'). ex_goal. red. apply le_n_S. apply le_inv_im_min_nat. exists hex. unfold fun_to_seg; auto. specialize (h5' hf0). red in h4. specialize (h4 (S i)). pose proof (linds_map_seg_eq_lind_occ h1 h2 f g i 0 (S i) (lt_n_Sn _) h4) as h6. hfirst h6. apply gt_lt. rewrite <- (count_occ_In h1). rewrite in_map_seg_iff. exists i, (lt_n_Sn _). rewrite fun_to_seg_compat; auto. specialize (h6 hf1). hfirst h6. apply in_map_seg. specialize (h6 hf2). hfirst h6. apply in_map_seg. specialize (h6 hf3). rewrite <- h5. pose proof (inv_im_min_nat_compat f (f i) h3 (inv_im_min_nat f (f i) h3)) as h7. hfirst h7. apply in_inv_im_min_nat. specialize (h7 hf4). destruct h7 as [h7 [h8 h9]]. generalize hf. rewrite h7. intro h10. pose proof (proof_irrelevance _ h10 hf2); subst. unfold lind. rewrite h6. unfold lind in h5'. pose proof (inv_im_min_nat_compat g (g i) h3' (inv_im_min_nat g (g i) h3')) as h7'. hfirst h7'. apply in_inv_im_min_nat. specialize (h7' hf5). destruct h7' as [h7' [h8' h9']]. pose proof (lind_functional h2 (map_seg (fun_to_seg g (S i))) (g i) (g (inv_im_min_nat g (g i) h3')) hf3 hf0) as h10. hfirst h10. rewrite inv_im_min_nat_compat'; auto. rewrite <- h7'. f_equal. erewrite <- lind_map_seg_fun_to_seg_inv_im_min_nat. unfold lind. apply lind_occ_functional; auto. Grab Existential Variables. rewrite in_map_seg_iff. exists (inv_im_min_nat g (g i) h3'). ex_goal. red. apply le_n_S. apply le_inv_im_min_nat. exists hex. rewrite fun_to_seg_compat; auto. Qed. Definition surj_nat_to_fun_sig {U V:Type} {D:Ensemble U} {E:Ensemble V} (f:nat->sig_set D) (g:nat->sig_set E) (pfs:surjective f) : sig_set D -> V := fun y => proj1_sig (surj_nat_to_fun f g pfs y). Lemma surj_nat_to_fun_sig_compat : forall {U V:Type} {D:Ensemble U} {E:Ensemble V} (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (f:nat->sig_set D) (g:nat->sig_set E) (pfs:surjective f) (pfs':surjective g), let pfdu' := sig_set_eq_dec pfdu D in let pfdv' := sig_set_eq_dec pfdv E in fun_linds_eq_nat pfdu' pfdv' f g -> forall i:nat, surj_nat_to_fun_sig f g pfs (f i) = proj1_sig (g i). intros U V D E hdu hdv hf g hs hs' hdu' hdv' h2 i. unfold surj_nat_to_fun_sig. f_equal. eapply surj_nat_to_fun_compat; auto. apply h2. Qed. Lemma lind_map_seg_fun_to_seg_inv_im_seg_min : forall {U:Type} (pfdu:type_eq_dec U) (f:nat->U) (n:nat) (y:sig_set (Im (seg n) f)) (pf:In (f (inv_im_seg_min f n y)) (map_seg (fun_to_seg f n))), lind pfdu (map_seg (fun_to_seg f n)) (f (inv_im_seg_min f n y)) pf = inv_im_seg_min f n y. intros U hdu f n y h1. rewrite lind_eq_iff. split. rewrite nth_lt_map_seg'. rewrite fun_to_seg_compat; auto. intro h2. rewrite in_firstn_iff in h2. destruct h2 as [m [h2 h3]]. rewrite nth_lt_map_seg' in h3. rewrite fun_to_seg_compat in h3. unfold inv_im_seg_min in h2. let P := type of h2 in match P with m < SetUtilities.min_set_nat ?x1 ?x2 ?x3 => pose proof (SetUtilities.min_set_nat_compat x1 x2 x3) as h4 end. simpl in h4. destruct h4 as [h4 h5]. destruct h4 as [h4]. destruct h4 as [h4 h6]. specialize (h5 m). hfirst h5. constructor. constructor. ex_goal. omega. pose proof (inv_im_seg_min_compat f n y) as h7. exists hex. rewrite fun_to_sig_seg_compat. apply proj1_sig_injective. rewrite <- h7 at 1. rewrite <- h3. simpl; auto. rewrite in_sing_iff. intro; omega. specialize (h5 hf). omega. Grab Existential Variables. rewrite length_map_seg. apply lt_impl_le. apply lt_inv_im_seg_min. rewrite length_map_seg. apply lt_inv_im_seg_min. Qed. Lemma nat_sig_to_fun_compat : forall {U V:Type} {D:Ensemble U} {E:Ensemble V} (pfdu:type_eq_dec U) (pfdv:type_eq_dec V) (f:nat->sig_set D) (g:nat->sig_set E) (i n:nat) (pf:i < n), linds pfdu (map_seg (fun_to_seg (proj1_sig_fun f) n)) = linds pfdv (map_seg (fun_to_seg (proj1_sig_fun g) n)) -> nat_sig_to_fun f g n (exist _ _ (Im_intro _ _ (seg n) (proj1_sig_fun f) i (in_seg _ _ pf) _ eq_refl)) = proj1_sig (g i). intros U V D E hdu hdv f g i n hlt hle. unfold nat_sig_to_fun. f_equal. pose proof (lind_map_seg_fun_to_seg_inv_im_seg_min hdu (proj1_sig_fun f) n) as h5. pose (Im_intro _ _ (seg n) (proj1_sig_fun f) _ (in_seg _ _ hlt) _ eq_refl) as h6. specialize (h5 (exist _ _ h6)). hfirst h5. rewrite in_map_seg_iff. exists (inv_im_seg_min (proj1_sig_fun f) n (exist (Inn (Im (seg n) (proj1_sig_fun f))) (proj1_sig_fun f i) h6)). ex_goal. apply lt_inv_im_seg_min. exists hex. unfold fun_to_seg; auto. specialize (h5 hf). pose proof (lind_map_seg_fun_to_seg_inv_im_seg_min hdv (proj1_sig_fun g) n) as h5'. pose proof (Im_intro _ _ (seg n) (proj1_sig_fun g) _ (in_seg _ _ hlt) _ eq_refl) as h6'. specialize (h5' (exist _ _ h6')). hfirst h5'. rewrite in_map_seg_iff. exists (inv_im_seg_min (proj1_sig_fun g) n (exist (Inn (Im (seg n) (proj1_sig_fun g))) (proj1_sig_fun g i) h6')). ex_goal. apply lt_inv_im_seg_min. exists hex. unfold fun_to_seg; auto. specialize (h5' hf0). fold h6. rewrite <- h5. pose proof (linds_map_seg_eq_lind_occ hdu hdv (proj1_sig_fun f) (proj1_sig_fun g) i 0 n hlt hle) as h7. hfirst h7. apply gt_lt. rewrite <- (count_occ_In hdu). rewrite in_map_seg_iff. exists i, hlt. rewrite fun_to_seg_compat; auto. specialize (h7 hf1). hfirst h7. apply in_map_seg. specialize (h7 hf2). hfirst h7. apply in_map_seg. specialize (h7 hf3). pose proof h6 as h8. apply in_proj1_sig_fun in h8. pose proof (inv_im_seg_min_sig_compat f n (exist _ _ h6)) as h9. pose proof (f_equal (@proj1_sig _ _) h9) as h10. generalize hf. rewrite proj1_sig_fun_compat, h10. intro h11; simpl. unfold lind. pose proof (proof_irrelevance _ hf2 h11); subst. rewrite h7. pose proof (inv_im_seg_min_sig_compat g n (exist _ _ h6')) as h9'. pose proof (f_equal (@proj1_sig _ _) h9') as h10'. simpl in h10'. rewrite proj1_sig_fun_compat in h10'. apply proj1_sig_injective. unfold proj1_sig_fun. rewrite <- h10'. repeat f_equal; auto. rewrite <- h5'. unfold lind; apply lind_occ_functional. reflexivity. reflexivity. auto. Qed. End Sort. Section Finite. Lemma Finite_Finite_Union_rev : forall {T:Type} (F:Family T), Finite (FamilyUnion F) -> Finite F. intros T F h2. apply NNPP. intro h3. assert (h4:Approximant _ F [S:Ensemble T | Inn F S /\ Included S (FamilyUnion F)]). constructor. eapply Finite_downward_closed. eapply power_set_finite. apply h2. red. intros A h4. destruct h4 as [h4]. destruct h4 as [h4l h4r]. constructor. assumption. red. intros S h4. destruct h4 as [h4]. destruct h4; auto. pose proof (make_new_approximant _ _ _ h3 h4) as h5. assert (h6: (Setminus F [S : Ensemble T | Inn F S /\ Included S (FamilyUnion F)]) = Empty_set _). apply Extensionality_Ensembles; red; split; auto with sets. red. intros A h6. destruct h6 as [h6 h7]. contradict h7. constructor. split; auto. red. intros x h7. apply family_union_intro with A; auto. rewrite h6 in h5. destruct h5. contradiction. Qed. Lemma card_cart_prod : forall {T U:Type} (A:Ensemble T) (B:Ensemble U), Finite A -> Finite B -> card_fun1 (cart_prod A B) = (card_fun1 A) * (card_fun1 B). intros T U A B h1 h2. destruct (eq_dec A (Empty_set _)) as [hae | hnae]. subst. rewrite cart_prod_empty. do 2 rewrite card_fun1_empty. ring. destruct (eq_dec B (Empty_set _)) as [hbe | hnbe]. subst. rewrite cart_prod_empty'. do 2 rewrite card_fun1_empty. ring. pose proof (not_empty_Inhabited _ _ hnbe) as h11. destruct h11 as [b h11]. pose proof (not_empty_Inhabited _ _ hnae) as h12. destruct h12 as [a h12]. pose proof (cart_prod_fin _ _ h1 h2) as h3. rewrite cart_prod_eq. rewrite cart_prod_eq in h3. pose proof (Finite_Finite_Union_rev _ h3) as h4. term0 h4. rename c into F. assert (h6: forall S, Inn F S -> Finite S). intros S h7. destruct h7 as [h7]. destruct h7 as [y h7]. destruct h7 as [h7l h7r]. subst. apply cart_prod_sing_fin; auto. assert (hpw: pairwise_disjoint F). red. intros X Y h7 h8 h9. destruct h8 as [h8]. destruct h8 as [y h8]. destruct h8 as [h8l h8r]. destruct h9 as [h9]. destruct h9 as [y' h9]. destruct h9 as [h9l h9r]. subst. red. apply Extensionality_Ensembles; red; split; auto with sets. red. intros pr h10. destruct h10 as [pr h10l h10r]. destruct h10l as [h10l]. destruct h10l as [h10a h10b]. destruct h10r as [h10r]. destruct h10r as [h10c h10d]. destruct h10b; destruct h10d; subst. contradict h7. reflexivity. pose proof (card_pairwise_disjoint _ h4 h6 hpw) as h5. rewrite h5. assert (h7:card_fun1 F = card_fun1 B). pose [pr:(Ensemble (T*U))*U | Inn B (snd pr) /\ fst pr = cart_prod A (Singleton (snd pr))] as P. assert (h8:dom_rel P = [S : Ensemble (T * U) | exists y : U, Inn B y /\ S = cart_prod A (Singleton y)]). apply Extensionality_Ensembles. red. split. red. intros S h9. destruct h9 as [h9]. destruct h9 as [y h9]. destruct h9 as [h9]. simpl in h9. constructor. exists y. assumption. red. intros S h9. destruct h9 as [h9]. constructor. destruct h9 as [y h9]. exists y. constructor. simpl. assumption. assert (h9:ran_rel P = B). apply Extensionality_Ensembles. red. split. red. intros y h9. destruct h9 as [h9]. destruct h9 as [x h9]. destruct h9 as [h9]. simpl in h9. destruct h9; assumption. red. intros y h9. constructor. exists (cart_prod A (Singleton y)). constructor. simpl. split; auto. assert (h10:fun_well_defined P). red. rewrite h8. rewrite h9. constructor. intros S h10. destruct h10 as [h10]. destruct h10 as [y h10]. destruct h10 as [h10l h10r]. exists y. red. split. split; auto. constructor. simpl. split; auto. intros y' h11'. destruct h11' as [h11l h11r]. destruct h11r as [h11']. destruct h11' as [h11a h11b]. simpl in h11b. rewrite h10r in h11b. assert (h13:Inn (cart_prod A (Singleton y)) (a, y)). constructor; simpl. split; auto; try constructor. rewrite h11b in h13. inversion h13 as [h14]. clear h13. simpl in h14. destruct h14 as [? h14]. destruct h14; auto. intros pr h10. destruct h10 as [h10]. destruct h10 as [h10l h10r]. split. constructor. exists (snd pr). split; auto. assumption. red in h10. rewrite h8 in h10. rewrite h9 in h10. rename F into S. pose (fin_map_intro _ _ b (ens_eq_dec (T*U)) h4 h2 P h10) as F. assert (h12':bij F). red. split. red. intros X Y h12' h13 h14. do 2 rewrite fin_map_app_compat in h14. pose proof (fin_fps_to_f_s_compat (ens_eq_dec (T*U)) h4 h10 b _ h12') as h15. pose proof (fin_fps_to_f_s_compat (ens_eq_dec (T*U)) h4 h10 b _ h13) as h16. assert (h17:P = fin_map_to_fps F). pose proof (fin_map_to_fps_s_compat (ens_eq_dec (T*U)) _ _ b h4 h2 P h10) as h17. rewrite h17. f_equal. assert (h18:fin_fps_to_f (ens_eq_dec (T*U)) h4 P h10 b X = fin_fps_to_f (ens_eq_dec (T*U)) h4 (fin_map_to_fps F) (fp_fin_map_to_fps F) b X). pose proof (subsetT_eq_compat _ _ _ _ h10 (fp_fin_map_to_fps F) h17) as h19. dependent rewrite -> h19. reflexivity. erewrite fin_fps_to_f_functional in h18. symmetry in h18. rewrite h14 in h18 at 1. erewrite fin_fps_to_f_functional in h15. rewrite h18 in h15 at 1. inversion h15 as [h19]. clear h15. simpl in h19. destruct h19 as [h19l h19r]. inversion h16 as [h20]. clear h16. simpl in h20. destruct h20 as [h20l h20r]. pose proof (subsetT_eq_compat _ _ _ _ h10 (fp_fin_map_to_fps F) h17) as h19. assert (h21:(fin_fps_to_f (ens_eq_dec (T*U)) h4 (fin_map_to_fps F) (fp_fin_map_to_fps F) b Y) = (fin_fps_to_f (ens_eq_dec (T*U)) h4 P h10 b Y)). dependent rewrite -> h19. reflexivity. erewrite fin_fps_to_f_functional in h19r. rewrite h21 in h19r at 1. rewrite <- h19r in h20r. rewrite h20r. reflexivity. reflexivity. assumption. assumption. red. intros y h12'. exists (cart_prod A (Singleton y)). constructor. constructor. exists y. split; auto. assert (h14:Inn (fin_map_to_fps F) (cart_prod A (Singleton y), y)). unfold F. rewrite <- fin_map_to_fps_s_compat. constructor. simpl. split; auto. pose proof (fin_map_to_fps_fin_map_app_compat F ((cart_prod A (Singleton y)), y) h14) as h13. simpl in h13. simpl. assumption. eapply bij_dom_ran_card_eq'. apply h12'. rewrite <- h7. assert (h8:forall S, Inn F S -> card_fun1 S = card_fun1 A). intros S h9. destruct h9 as [h9]. destruct h9 as [y h9]. destruct h9 as [h9l h9r]. subst. apply card_cart_prod_sing'; auto. pose proof (plus_bag_nat_im_set_bag_eq_f _ h4 card_fun1 (card_fun1 A) h8) as h9. rewrite h9. unfold F. ring. Qed. End Finite. Section Random. Definition tuple_prod_ens {T U n} (A:T->Ensemble U) (x:T^n) : Ensemble (U^n) := [y:U^n | forall i (pfi:i < n), Inn (A (nth_prod i pfi x)) (nth_prod i pfi y)]. Definition tuple_prod_ens_dep {T n} {U:T->Type} (A:forall x:T, Ensemble (U x)) (x:T^n) : Ensemble (tuple_prod U x) := [y:tuple_prod U x | (forall i (pfi:i < n), Inn (A (nth_prod i pfi x)) (nth_tuple_prod y i pfi))]. Lemma in_nth_tuple_prod_ens_dep : forall {T n} {U:T->Type} {A:forall t:T, Ensemble (U t)} {w:T^n} {t:tuple_prod U w} (pfi:Inn (tuple_prod_ens_dep A w) t) i (pfi:i < n), Inn (A (nth_prod i pfi w)) (nth_tuple_prod t i pfi). intros ? ? ? ? ? ? h1; destruct h1; auto. Qed. Lemma in_transfer_tuple_prod1 : forall {T} {U:T->Type} {A:forall t:T, Ensemble (U t)} {x:T} (a:U x), Inn (A x) a -> Inn (tuple_prod_ens_dep A (x, tt) (n := 1)) (transfer_to_tuple_prod1 a). intros T U A x a h1. constructor. intros i hi. pose proof (lt1 _ hi); subst. simpl; auto. Qed. Lemma in_transfer_to_tuple_prod1_iff : forall {T} {U:T->Type} {A:forall t, Ensemble (U t)} {s:T} {t:T^1} (pfe:t = (s, tt)) (u:U s), Inn (tuple_prod_ens_dep A (s, tt) (n := 1)) (transfer_to_tuple_prod1 u) <-> Inn (tuple_prod_ens_dep A t) (transfer_dep_r pfe (transfer_to_tuple_prod1 u)). intros; subst; rewrite transfer_dep_r_eq_refl; tauto. Qed. Definition tuple_prod_sys {T U X} (A:T->Ensemble U) (v:X->T) : X -> Ensemble U := fun x => A (v x). Definition tuple_prod_ens_sys {T U X n} (A:T->Ensemble U) (v:X->T) (x:X^n) : Ensemble (U^n) := [y:U^n | forall i (pfi:i < n), Inn (tuple_prod_sys A v (nth_prod i pfi x)) (nth_prod i pfi y)]. Lemma in_tuple_prod_ens_compat : forall {T U n} (A:T->Ensemble U) (x:T^n) (y:U^n), Inn (tuple_prod_ens A x) y -> Inn (tuple_prod_ens_dep A x) (transfer_to_tuple_prod_const _ y). intros T U n A x y h1; destruct h1 as [h1]; constructor; intros i hi; specialize (h1 i hi); rewrite nth_tuple_prod_transfer_to_tuple_prod_const; auto. Qed. Lemma in_tuple_prod_ens_dep_const_compat : forall {T U n} {x:T^n} (A:T->Ensemble U) (p:tuple_prod (fun _:T => U) x), Inn (tuple_prod_ens_dep A x) p -> Inn (tuple_prod_ens A x) (transfer_tuple_prod_const p). intros T U n x A p h1; destruct h1 as [h1]; constructor; intros i hi; specialize (h1 i hi); rewrite nth_prod_transfer_tuple_prod_const; auto. Qed. Lemma in_tuple_prod_ens_nth : forall {T U n} {A:T->Ensemble U} {x:T^n} {y:U^n}, Inn (tuple_prod_ens A x) y -> forall i (pfi:i < n), Inn (A (nth_prod i pfi x)) (nth_prod i pfi y). intros ? ? ? ? ? ? h1; destruct h1; auto. Qed. Lemma in_tuple_prod_ens_dep_nth_iff : forall {T n} {U:T->Type} {x:T^n} (A:forall x:T, Ensemble (U x)) (y:tuple_prod U x), Inn (tuple_prod_ens_dep A x) y <-> forall i (pfi:i < n), Inn (A (nth_prod i pfi x)) (nth_tuple_prod y i pfi). intros ? ? ? ? ? ?; split; intro h1. intros i hi. destruct h1; auto. constructor; auto. Qed. Definition sig_tuple_prod_ens_dep_const_to {T U n} (A:T->Ensemble U) (w:T^n) : sig_set (tuple_prod_ens_dep A w) -> sig_set (tuple_prod_ens A w) := fun p' => let (p, pfp) := p' in let pfp' := in_tuple_prod_ens_dep_const_compat A p pfp in exist (fun c => Inn (tuple_prod_ens A w) c) (transfer_tuple_prod_const p) pfp'. Lemma inj_sig_tuple_prod_ens_dep_const_to : forall {T U n} (A:T->Ensemble U) (w:T^n) (x x':sig_set (tuple_prod_ens_dep A w)), sig_tuple_prod_ens_dep_const_to A w x = sig_tuple_prod_ens_dep_const_to A w x' -> x = x'. unfold sig_tuple_prod_ens_dep_const_to. intros ? ? ? ? ? x x' h1. destruct x, x'. apply exist_injective in h1. apply inj_transfer_tuple_prod_const in h1; subst; f_equal; apply proof_irrelevance. Qed. Definition transfer_tuple_prod_ens {T U n} (A:T->Ensemble U) (x:T^n) : Ensemble (tuple_prod (fun _:T => U) x) := Im (tuple_prod_ens A x) (transfer_to_tuple_prod_const x). Definition transfer_tuple_prod_ens_dep {T U n} (A:T->Ensemble U) (x:T^n) : Ensemble (U^n) := Im (tuple_prod_ens_dep A x) (transfer_tuple_prod_const (x:=x)). Lemma transfer_tuple_prod_ens_compat : forall {T U n} (A:T->Ensemble U) (x:T^n), transfer_tuple_prod_ens A x = tuple_prod_ens_dep A x. intros T U n A x; apply Extensionality_Ensembles; red; split; red; intros y hy; destruct hy as [hy]; subst; [destruct H; constructor; intros i hi; rewrite nth_tuple_prod_transfer_to_tuple_prod_const; auto| apply Im_intro with (transfer_tuple_prod_const y); [constructor; intros i hi; specialize (hy i hi); rewrite nth_prod_transfer_tuple_prod_const; auto| rewrite transfer_to_tuple_prod_const_compat; auto]]. Qed. Lemma transfer_tuple_prod_ens_dep_compat : forall {T U n} (A:T->Ensemble U) (x:T^n), tuple_prod_ens A x = transfer_tuple_prod_ens_dep A x. unfold transfer_tuple_prod_ens_dep; intros T U n A x; apply Extensionality_Ensembles; red; split; red; intros y hy; destruct hy as [hy]; [apply Im_intro with (transfer_to_tuple_prod_const _ y); try constructor; [intros i hi; specialize (hy i hi); rewrite nth_tuple_prod_transfer_to_tuple_prod_const; auto| rewrite transfer_tuple_prod_const_to_compat; auto]| subst; destruct H; constructor; intros i hi; specialize (H i hi); rewrite nth_prod_transfer_tuple_prod_const; auto]. Qed. Lemma tuple_prod_ens0 : forall {T U} (A:T->Ensemble U) (x:T^0), tuple_prod_ens A x = Singleton tt. intros T U A x. destruct x; simpl. apply Extensionality_Ensembles; red; split; red; intros y h1. destruct h1. destruct y. constructor. destruct y. constructor. intros; omega. Qed. Lemma tuple_prod_ens_dep0 : forall {T} (U:T->Type) (A:forall x:T, Ensemble (U x)) (x:T^0), tuple_prod_ens_dep A x = Singleton tt. intros T U A x; destruct x; simpl. apply Extensionality_Ensembles; red; split; red; intros x h1; destruct h1. destruct x. constructor. constructor. intros; omega. Qed. Lemma tuple_prod_ens_n0 : forall {T U n} (A:T->Ensemble U) (x:T^n) (pfeq:n = 0), tuple_prod_ens A x = Singleton (transfer_tt pfeq U). intros; subst; rewrite transfer_tt_eq_refl; apply tuple_prod_ens0. Qed. Lemma tuple_prod_ens_dep_n0 : forall {T n} (U:T->Type) (A:forall x:T, Ensemble (U x)) (x:T^n) (pfeq:n = 0), tuple_prod_ens_dep A x = Singleton (transfer_tt_tuple_prod U pfeq x). intros; subst; rewrite transfer_tt_tuple_prod_eq_refl; apply tuple_prod_ens_dep0. Qed. Lemma in_tt_tuple_prod_ens_dep : forall {T} (U:T->Type) (A:forall x:T, Ensemble (U x)) (x:T^0), Inn (tuple_prod_ens_dep A x) tt. intros; rewrite tuple_prod_ens_dep0; constructor; auto. Qed. Lemma in_tt_tuple_prod_ens_dep' : forall {T} (U:T->Type) (A:forall x:T, Ensemble (U x)) (x:T^0), Inn (tuple_prod_ens_dep A x) x. intros; rewrite tuple_prod_ens_dep0; simpl in x; destruct x; constructor; auto. Qed. Lemma in_transfer_tt_tuple_prod_ens_dep : forall {T n} (U:T->Type) (A:forall x:T, Ensemble (U x)) (x:T^n) (pfeq:n = 0), Inn (tuple_prod_ens_dep A x) (transfer_tt_tuple_prod U pfeq x). intros; subst; simpl; rewrite tuple_prod_ens_dep0; simpl in x; destruct x; constructor; auto. Qed. Lemma tuple_prod_ens_S : forall {T U n} (A:T->Ensemble U) (x:T^S n), tuple_prod_ens A x = [t:U^(S n) | Inn (A (fst x)) (fst t) /\ Inn (tuple_prod_ens A (snd x)) (snd t)]. intros T U n A x; apply Extensionality_Ensembles; red; split; red; intros y h1; destruct h1 as [h1]; constructor. split. specialize (h1 0 (lt_0_Sn _)). destruct x, y; simpl; simpl in h1; auto. destruct x as [? x], y as [? y]; simpl. constructor. intros i hi. specialize (h1 (S i) (lt_n_S _ _ hi)); simpl in h1; auto. pose proof (proof_irrelevance _ hi (lt_S_n _ _ (lt_n_S _ _ hi))) as h2. rewrite h2; auto. destruct h1 as [h1 h2]. destruct x as [? x], y as [? y]; simpl in h1, h2. destruct h2. intros i hi. destruct i as [|i]. simpl; auto. specialize (H i (lt_S_n _ _ hi)). simpl; auto. Qed. Lemma tuple_prod_ens_dep_S : forall {T n} {U:T->Type} (A:forall t:T, Ensemble (U t)) (x:T^S n), tuple_prod_ens_dep A x = [t:tuple_prod U x | Inn (A (fst x)) (fst t) /\ Inn (tuple_prod_ens_dep A (snd x)) (snd t)]. intros T n U A x; apply Extensionality_Ensembles; red; split; red; intros y h1; destruct h1 as [h1]; constructor. destruct x as [? x]; simpl. destruct y as [u y]; simpl in u; simpl; simpl in y. split. specialize (h1 0 (lt_0_Sn _)); simpl in h1; auto. constructor. intros i hi. specialize (h1 (S i) (lt_n_S _ _ hi)); simpl in h1; auto. pose proof (proof_irrelevance _ hi (lt_S_n _ _ (lt_n_S _ _ hi))) as h2. rewrite h2; auto. destruct h1 as [h1 h2]. destruct x as [? x], y as [? y]; simpl in h1, h2. destruct h2. intros i hi. destruct i as [|i]. simpl; auto. specialize (H i (lt_S_n _ _ hi)). simpl; auto. Qed. Lemma finite_mem_rel_classes_im_rel_sig_fun : forall {T U:Type} {A:Ensemble T} (f:sig_set A->U), Finite A -> finite_mem (rel_classes_im_rel_sig_fun f). intros T U A f h1. red; intros S h2. destruct h2 as [S h2]. subst. destruct h2 as [x h2]. subst. clear h2. apply finite_image. eapply Finite_downward_closed. rewrite <- finite_full_sig_iff; auto. red; intros; constructor. Qed. Lemma couple_eq_add_iff : forall {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (a b c:T), Couple a b = Add A c <-> ((a = b /\ a = c /\ (A = Singleton a \/ A = Empty_set _)) \/ (a <> b /\ ((a = c /\ (A = Singleton b \/ A = Couple a b)) \/ (b = c /\ (A = Singleton a \/ A = Couple a b))))). intros T hdt A a b c; split; intro h1. pose proof (finite_couple a b) as hf. rewrite h1 in hf. pose proof (add_downward_closed _ _ hf) as hf'. pose proof (Add_intro2 _ A c) as h2. rewrite <- h1 in h2. destruct h2. pose proof (Couple_r _ a b) as h2. rewrite h1 in h2. destruct h2 as [x h2 | x h2]; subst. destruct (hdt x a) as [|h3]; subst. rewrite in_add_eq, couple_singleton in h1; auto. right. split; auto. left. split; auto. destruct (in_fin_ens_dec hdt A hf' a) as [hi | hi]. right. rewrite in_add_eq in h1; auto. left. rewrite couple_comm in h1. rewrite couple_add_add in h1. apply add_nin_inj_eq in h1; subst. rewrite add_empty_sing; auto. rewrite add_empty_sing. rewrite in_sing_iff; auto. assumption. destruct h2. rewrite couple_singleton in h1. apply sing_eq_add in h1. destruct h1 as [? h2]. left; tauto. destruct (hdt a b) as [|h3]; subst. rewrite couple_singleton in h1. apply sing_eq_add in h1. destruct h1 as [? h1]. left; tauto. right. split; auto. right. split; auto. pose proof (Couple_l _ a b) as h4. rewrite h1 in h4. destruct h4 as [a h4 | a h4]; subst. destruct (in_fin_ens_dec hdt A hf' b) as [hi | hi]. right. rewrite in_add_eq in h1; auto. rewrite couple_add_add in h1. apply add_nin_inj_eq in h1; subst. left. rewrite add_empty_sing; auto. rewrite add_empty_sing. rewrite in_sing_iff; auto. assumption. destruct h4. contradict h3; auto. destruct h1 as [h1 | h1]. destruct h1 as [h1 [h2 h3]]; subst. subst. rewrite couple_singleton. destruct h3; subst. rewrite in_add_eq; auto. constructor. rewrite add_empty_sing; auto. destruct h1 as [h1 h2]. destruct h2 as [h2 | h2]. destruct h2 as [h2 h3]; subst. destruct h3; subst. rewrite couple_comm, couple_add_sing; auto. rewrite in_add_eq; auto. left. destruct h2 as [? h2]; subst. destruct h2; subst. rewrite couple_add_sing; auto. rewrite in_add_eq; auto. right. Qed. Lemma couple_eq_add_add_iff : forall {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (a b x y:T), Couple a b = Add (Add A x) y <-> (a = b /\ x = y /\ a = x /\ (A = Empty_set _ \/ A = Singleton a)) \/ (a <> b /\ ((x = y /\ (((x = a /\ (A = Couple a b \/ A = Singleton b)) \/ (x = b /\ (A = Couple a b \/ A = Singleton a))))) \/ (x <> y /\ (x = a /\ y = b \/ x = b /\ y = a ) /\ (A = Empty_set _ \/ A = Singleton a \/ A = Singleton b \/ A = Couple a b)))). intros T hdt A a b x y. pose proof (finite_couple a b) as hf. split; intro h1. rewrite h1 in hf. pose proof (add_downward_closed _ _ hf) as hf'. pose proof (add_downward_closed _ _ hf') as hf''. pose proof h1 as h1'. rewrite couple_eq_add_iff in h1'; auto. destruct h1' as [h1' | h1']. destruct h1' as [h1' [h2 [h3 | h3]]]; subst; subst. symmetry in h3. apply sing_eq_add in h3. destruct h3 as [? [? | ?]]; subst; tauto. apply add_not_empty in h3; contradiction. destruct h1' as [h1' h2]. destruct h2 as [h2 | h2]. destruct h2 as [? h2]; subst. destruct h2 as [h2 | h2]. symmetry in h2. apply sing_eq_add in h2. destruct h2 as [? [? | ?]]; subst. right. split; auto. right. split; auto. right. split; auto. tauto. right. split; auto. destruct (hdt x y) as [h3 | h3]; subst. rewrite in_add_eq in h1. left. split; auto. pose proof (Couple_r _ y b) as h3. rewrite h1 in h3. destruct h3 as [b h3 | b h3]. destruct (in_fin_ens_dec hdt A hf'' y) as [h4 | h4]. rewrite in_add_eq in h2; auto. rewrite couple_comm, couple_add_add in h1. apply add_nin_inj_eq in h1; auto; subst. rewrite add_empty_sing; tauto. rewrite add_empty_sing, in_sing_iff; auto. destruct h3. tauto. right; constructor. pose proof (Couple_l _ y b) as h4. right. split; auto. rewrite <- h2 in h4. destruct h4 as [y h4 |y h4]; subst. rewrite in_add_eq in h1. pose proof (Add_intro2 _ A x) as h5. rewrite <- h1 in h5. destruct h5. contradict h3; auto. split; auto. destruct (in_fin_ens_dec hdt _ hf'' b) as [h5 | h5]. rewrite in_add_eq in h1; subst; auto. rewrite couple_add_add in h1. apply add_nin_inj_eq in h1; auto. rewrite add_empty_sing in h1; subst; tauto. rewrite add_empty_sing, in_sing_iff; auto. left; auto. destruct h4. contradict h3; auto. destruct h2 as [? [h3 | h4]]; subst. right. symmetry in h3. apply sing_eq_add in h3. destruct h3 as [? [? | ?]]; subst. tauto. tauto. pose proof (Add_intro2 _ A x) as h5. rewrite h4 in h5. destruct h5. pose proof (Couple_r _ a y) as h5. rewrite <- h4 in h5. destruct h5 as [y h5 | y h5]. rewrite in_add_eq in h1. destruct (in_fin_ens_dec hdt _ hf'' a) as [h6 | h6]. rewrite in_add_eq in h4; subst; tauto. rewrite couple_comm, couple_add_add in h1. apply add_nin_inj_eq in h1; auto. rewrite add_empty_sing in h1; subst; tauto. rewrite add_empty_sing, in_sing_iff; auto. left; auto. destruct h5. contradict h1'; auto. rewrite in_add_eq in h1. pose proof (Couple_l _ a y) as h5. rewrite <- h4 in h5. destruct h5 as [a h5 | a h5]; subst. destruct (in_fin_ens_dec hdt _ hf'' y) as [h6 | h6]. rewrite in_add_eq in h1; subst; tauto. rewrite couple_add_add in h1. apply add_nin_inj_eq in h1; auto. rewrite add_empty_sing in h1; subst; tauto. rewrite add_empty_sing, in_sing_iff; auto. destruct h5. contradict h1'; auto. right; constructor. destruct h1 as [h1 | h1]. destruct h1 as [? [? [? [? | ?]]]]; subst; subst. rewrite add_empty_sing, in_add_eq. rewrite couple_singleton; auto. constructor. do 2 rewrite in_add_eq. rewrite couple_singleton; auto. constructor. constructor. constructor. destruct h1 as [h1 h2]. destruct h2 as [h2 | h2]. destruct h2 as [h2 h3]; subst. destruct h3 as [h3 | h3]. destruct h3 as [? [? | ?]]; subst. rewrite in_add_eq. rewrite in_add_eq; auto. left. right. constructor. rewrite in_add_eq, <- couple_add_sing, couple_comm; auto. right; constructor. destruct h3 as [? [? | ?]]; subst. rewrite in_add_eq. rewrite in_add_eq; auto. right. right. constructor. rewrite in_add_eq, couple_add_sing; auto. right. constructor. destruct h2 as [h2 h3]. destruct h3 as [h3 h4]. destruct h3 as [[? ?] | [? ?]], h4 as [? | [? | [? | ?]]]; subst. rewrite couple_add_add; auto. rewrite (in_add_eq (Singleton a) a), couple_add_add; auto. rewrite <- couple_add_add, <- couple_add_sing; auto. constructor. rewrite in_add_eq, couple_comm, <- couple_add_sing; auto. left; constructor. do 2 rewrite in_add_eq; auto. left. right. left. rewrite couple_comm, couple_add_add; auto. rewrite in_add_eq, <- couple_add_sing; auto. left; constructor. rewrite (in_add_eq (Singleton b) b), <- couple_add_sing, couple_comm; auto. constructor. rewrite in_add_eq, in_add_eq; auto. right. left. left. Qed. End Random. Lemma infinite_nat : ~Finite (Full_set nat). pose proof nat_infinite as h1. rewrite Finite_FiniteT_iff. assumption. Qed. Lemma infinite_pos_evens : ~Finite pos_evens. intro h1. rewrite pos_evens_eq in h1. pose proof (Pigeonhole_bis _ _ _ _ infinite_nat h1) as h2. contradict h2. red. intros; omega. Qed. Lemma infinite_pos_odds : ~Finite pos_odds. intro h1. rewrite pos_odds_eq in h1. pose proof (Pigeonhole_bis _ _ _ _ infinite_nat h1) as h2. contradict h2. red. intros; omega. Qed. Definition segT_transfer {m n:nat} (pf:m = n) (a:m @) : n @ := s_transfer_dep pf a. Definition segT_transfer_r {m n:nat} (pf:n = m) (a:m @) : n @ := s_transfer_dep_r pf a. Lemma segT_transfer_transfer_r_compat : forall {m n:nat} (pf:m = n) (a:m @), segT_transfer pf a = segT_transfer_r (eq_sym pf) a. intros m n h1 a. subst. unfold segT_transfer, segT_transfer_r. rewrite s_transfer_dep_eq_refl. pose proof (proof_irrelevance _ (@eq_sym nat n n (@eq_refl nat n)) (eq_refl)) as h1. rewrite h1. rewrite s_transfer_dep_r_eq_refl. reflexivity. Qed. Lemma segT_transfer_r_transfer_compat : forall {m n:nat} (pf:n = m) (a:m @), segT_transfer_r pf a = segT_transfer (eq_sym pf) a. intros m n h1 a. subst. unfold segT_transfer, segT_transfer_r. rewrite s_transfer_dep_r_eq_refl. pose proof (proof_irrelevance _ (@eq_sym nat m m (@eq_refl nat m)) (eq_refl _)); subst. rewrite s_transfer_dep_eq_refl; auto. Qed. Lemma segT_transfer_eq_refl : forall {m:nat} (pf:m = m) (a:m @), segT_transfer pf a = a. intros m h1 a. unfold segT_transfer. pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. rewrite s_transfer_dep_eq_refl; auto. Qed. Lemma segT_transfer_r_eq_refl : forall {m:nat} (pf:m = m) (a:m @), segT_transfer_r pf a = a. intros m h1 a. unfold segT_transfer_r. pose proof (proof_irrelevance _ h1 (eq_refl _)); subst. rewrite s_transfer_dep_r_eq_refl; auto. Qed. Lemma segT_transfer_undoes_segT_transfer_r : forall {m n:nat} (pf:m = n) (a:n @), segT_transfer pf (segT_transfer_r pf a) = a. intros m n h1 a. subst. rewrite segT_transfer_eq_refl, segT_transfer_r_eq_refl. reflexivity. Qed. Lemma segT_transfer_r_undoes_segT_transfer : forall {m n:nat} (pf:m = n) (a:m @), segT_transfer_r pf (segT_transfer pf a) = a. intros m n h1 a. subst. rewrite segT_transfer_eq_refl, segT_transfer_r_eq_refl. reflexivity. Qed. Lemma segT_transfer_cond : forall {m n:nat} (pf:m = n) (a:m @), proj1_sig a < n. intros m n h1 a. subst. apply proj2_sig. Qed. Lemma segT_transfer_compat : forall {m n:nat} (pf:m = n) (a:m @), segT_transfer pf a = exist (fun a => a < n) _ (segT_transfer_cond pf a). intros m n h1 a. unfold segT_transfer. destruct h1. rewrite transfer_dep_eq_refl. apply proj1_sig_injective; auto. Qed. Lemma segT_transfer_r_compat : forall {m n:nat} (pf:m = n) (a:n @), segT_transfer_r pf a = exist (fun a => a < m) _ (segT_transfer_cond (eq_sym pf) a). intros m n h1 a. unfold segT_transfer_r. destruct h1. rewrite transfer_dep_r_eq_refl. apply proj1_sig_injective; auto. Qed. Definition transfer_segT_fun1 {n m U} (pf:n = m) (f:n @->U) : m @ -> U := fun x => f (segT_transfer (eq_sym pf) x). Definition transfer_segT_fun2 {n m T} (pf:n = m) (f:T -> n @) : T -> m @ := fun x => segT_transfer pf (f x). Lemma transfer_segT_fun1_eq_refl : forall {U n} (pf:n = n) (f:n @ -> U), transfer_segT_fun1 pf f = f. intros U n h1 f; apply functional_extensionality; unfold transfer_segT_fun1, eq_rect_r, eq_rect; intro. pose proof (proof_irrelevance _ (eq_sym h1) (eq_refl _)) as h2; rewrite h2; auto. Qed. Lemma transfer_segT_fun2_eq_refl : forall {T n} (pf:n = n) (f:T -> n @), transfer_segT_fun2 pf f = f. intros U n h1 f; apply functional_extensionality; unfold transfer_segT_fun1, eq_rect_r, eq_rect; intro. pose proof (proof_irrelevance _ h1 (eq_refl _)) as h2; rewrite h2; auto. Qed. Lemma transfer_segT_fun1_compat : forall {U n m} (pf:n = m) (f:n @ -> U) (x:m @), transfer_segT_fun1 pf f x = f (segT_transfer (eq_sym pf) x). intros; subst; rewrite transfer_segT_fun1_eq_refl, segT_transfer_eq_refl; auto. Qed. Lemma transfer_segT_fun2_compat : forall {T n m} (pf:n = m) (f:T -> n @) (x:T), transfer_segT_fun2 pf f x = segT_transfer pf (f x). intros; subst; rewrite transfer_segT_fun2_eq_refl,segT_transfer_eq_refl; auto. Qed. Definition segT_transfer_ens {m n:nat} (pf:m = n) (S:Ensemble (m @)) : Ensemble (n @) := transfer_dep (U := (fun a:nat => Ensemble (segT a))) pf S. Definition segT_transfer_ens_r {m n:nat} (pf:m = n) (S:Ensemble (n @)) : Ensemble (m @) := transfer_dep_r (U := (fun a:nat => Ensemble (segT a))) pf S. Definition segT_transfer_fam {m n:nat} (pf:m = n) (S:Family (m @)) : Family (n @) := transfer_dep (U := (fun a:nat => Family (segT a))) pf S. Definition segT_transfer_fam_r {m n:nat} (pf:m = n) (S:Family (n @)) : Family (m @) := transfer_dep_r (U := (fun a:nat => Family (a @))) pf S. Lemma segT_transfer_ens_eq_refl : forall (m:nat) (S:Ensemble (m @)), segT_transfer_ens eq_refl S = S. intros. unfold segT_transfer_ens. rewrite transfer_dep_eq_refl. reflexivity. Qed. Lemma segT_transfer_ens_r_eq_refl : forall (m:nat) (S:Ensemble (m @)), segT_transfer_ens_r eq_refl S = S. intros. unfold segT_transfer_ens_r. rewrite transfer_dep_r_eq_refl. reflexivity. Qed. Lemma segT_transfer_fam_eq_refl : forall (m:nat) (S:Family (m @)), segT_transfer_fam eq_refl S = S. intros. unfold segT_transfer_fam. rewrite transfer_dep_eq_refl. reflexivity. Qed. Lemma segT_transfer_fam_r_eq_refl : forall (m:nat) (S:Family (m @)), segT_transfer_fam_r eq_refl S = S. intros. unfold segT_transfer_fam_r. rewrite transfer_dep_r_eq_refl. reflexivity. Qed. Lemma in_segT_transfer_ens : forall {m n:nat} (pf:m = n) (S:Ensemble (m @)) (x:m @), Inn S x <-> Inn (segT_transfer_ens pf S) (segT_transfer pf x). intros m n h1 S x. subst. rewrite segT_transfer_ens_eq_refl, segT_transfer_eq_refl. tauto. Qed. Lemma in_segT_transfer_ens_r : forall {m n:nat} (pf:m = n) (S:Ensemble (n @)) (x:n @), Inn S x <-> Inn (segT_transfer_ens_r pf S) (segT_transfer_r pf x). intros m n h1 S x. subst. rewrite segT_transfer_ens_r_eq_refl. rewrite segT_transfer_r_eq_refl. tauto. Qed. Lemma in_segT_transfer_ens_hetero : forall {m n:nat} (pf:m = n) (S:Ensemble (m @)) (x:n @), Inn (segT_transfer_ens pf S) x <-> Inn S (segT_transfer_r pf x). intros m n h1 S x. subst. rewrite segT_transfer_ens_eq_refl. rewrite segT_transfer_r_eq_refl. tauto. Qed. Lemma in_segT_transfer_ens_r_hetero : forall {m n:nat} (pf:m = n) (S:Ensemble (n @)) (x:m @), Inn (segT_transfer_ens_r pf S) x <-> Inn S (segT_transfer pf x). intros m n h1 S x. subst. rewrite segT_transfer_ens_r_eq_refl. rewrite segT_transfer_eq_refl. tauto. Qed. Lemma in_segT_transfer_fam : forall {m n:nat} (pf:m = n) (F:Family (m @)) (S:Ensemble (m @)), Inn F S <-> Inn (segT_transfer_fam pf F) (segT_transfer_ens pf S). intros m n h1 F S. subst. rewrite segT_transfer_fam_eq_refl, segT_transfer_ens_eq_refl. tauto. Qed. Lemma in_segT_transfer_fam_r : forall {m n:nat} (pf:m = n) (F:Family (n @)) (S:Ensemble (n @)), Inn F S <-> Inn (segT_transfer_fam_r pf F) (segT_transfer_ens_r pf S). intros m n h1 F S. subst. rewrite segT_transfer_fam_r_eq_refl, segT_transfer_ens_r_eq_refl. tauto. Qed. Lemma in_segT_transfer_fam_hetero : forall {m n:nat} (pf:m = n) (F:Family (m @)) (S:Ensemble (n @)), Inn (segT_transfer_fam pf F) S <-> Inn F (segT_transfer_ens_r pf S). intros m n h1 F S. subst. rewrite segT_transfer_fam_eq_refl, segT_transfer_ens_r_eq_refl. tauto. Qed. Lemma in_segT_transfer_fam_r_hetero : forall {m n:nat} (pf:m = n) (F:Family (n @)) (S:Ensemble (m @)), Inn (segT_transfer_fam_r pf F) S <-> Inn F (segT_transfer_ens pf S). intros m n h1 F S. subst. rewrite segT_transfer_fam_r_eq_refl. rewrite segT_transfer_ens_eq_refl. tauto. Qed. (*Leftover from seg_set_one_t days. Might phase out eventually.*) Definition sot_transfer {m n:nat} (pf:m = n) (a:seg_one_t m) : seg_one_t n := s_transfer_dep pf a. Definition sot_transfer_r {m n:nat} (pf:n = m) (a:seg_one_t m) : seg_one_t n := s_transfer_dep_r pf a. Lemma sot_transfer_transfer_r_compat : forall {m n:nat} (pf:m = n) (a:seg_one_t m), sot_transfer pf a = sot_transfer_r (eq_sym pf) a. intros m n h1 a. subst. unfold sot_transfer, sot_transfer_r. rewrite s_transfer_dep_eq_refl. pose proof (proof_irrelevance _ (@eq_sym nat n n (@eq_refl nat n)) (eq_refl _)) as h1. rewrite h1; auto. Qed. Lemma sot_transfer_r_transfer_compat : forall {m n:nat} (pf:n = m) (a:seg_one_t m), sot_transfer_r pf a = sot_transfer (eq_sym pf) a. intros m n h1 a. subst. unfold sot_transfer, sot_transfer_r. rewrite s_transfer_dep_r_eq_refl. pose proof (proof_irrelevance _ (@eq_sym nat m m (@eq_refl nat m)) (eq_refl _)) as h1. rewrite h1; auto. Qed. Lemma sot_transfer_eq_refl : forall {m:nat} (a:seg_one_t m), sot_transfer eq_refl a = a. intros m a. unfold sot_transfer. rewrite s_transfer_dep_eq_refl. reflexivity. Qed. Lemma sot_transfer_r_eq_refl : forall {m:nat} (a:seg_one_t m), sot_transfer_r eq_refl a = a. intros m a. unfold sot_transfer_r. rewrite s_transfer_dep_r_eq_refl. reflexivity. Qed. Lemma sot_transfer_undoes_sot_transfer_r : forall {m n:nat} (pf:m = n) (a:seg_one_t n), sot_transfer pf (sot_transfer_r pf a) = a. intros m n h1 a. subst. rewrite sot_transfer_eq_refl, sot_transfer_r_eq_refl. reflexivity. Qed. Lemma sot_transfer_r_undoes_sot_transfer : forall {m n:nat} (pf:m = n) (a:seg_one_t m), sot_transfer_r pf (sot_transfer pf a) = a. intros m n h1 a. subst. rewrite sot_transfer_eq_refl, sot_transfer_r_eq_refl. reflexivity. Qed. Lemma sot_transfer_cond : forall {m n:nat} (pf:m = n) (a:seg_one_t m), 1 <= (proj1_sig a) <= n. intros m n h1 a. subst. apply proj2_sig. Qed. Lemma sot_transfer_eq : forall {m n:nat} (pf:m = n) (a:seg_one_t m), sot_transfer pf a = exist _ _ (sot_transfer_cond pf a). intros m n h1 a. unfold sot_transfer. destruct h1. rewrite transfer_dep_eq_refl. apply proj1_sig_injective; auto. Qed. Lemma sot_transfer_r_eq : forall {m n:nat} (pf:m = n) (a:seg_one_t n), sot_transfer_r pf a = exist _ _ (sot_transfer_cond (eq_sym pf) a). intros m n h1 a. unfold sot_transfer_r. destruct h1. rewrite transfer_dep_r_eq_refl. apply proj1_sig_injective; auto. Qed. Section FiniteSets. (*We'll be using lists as sets directly, transparently, with set equality the following predicate, and some [Setoid] functionality.*) Inductive set_eq {T} (l l':list T) : Prop := | set_eq_intro : incl l l' -> incl l' l -> set_eq l l'. Lemma refl_set_eq : forall T (l:list T), set_eq l l. intros; constructor; auto; apply incl_refl. Qed. Lemma sym_set_eq : forall T, Symmetric (set_eq (T:=T)). intro; red; intros; inversion H; constructor; auto. Qed. Lemma trans_set_eq : forall T, Transitive (set_eq (T:=T)). intro; red; intros ? ? ? h1 h2; destruct h1, h2; constructor; auto; [eapply trans_incl; try apply H; try assumption | eapply trans_incl; try apply H2; try assumption]. Qed. End FiniteSets. Add Parametric Relation (T:Type) : (list T) (@set_eq T) reflexivity proved by (@refl_set_eq T) symmetry proved by (@sym_set_eq T) transitivity proved by (@trans_set_eq T) as list_eq. Infix "==" := set_eq (at level 70). Notation "x <=> y" := (not (x == y)) (at level 70). Section FiniteSets'. Lemma not_set_eq_neq : forall {T} (A B:list T), A <=> B -> A <> B. intros ? ? ? h1; contradict h1; subst. reflexivity. Qed. Add Parametric Relation (T:Type) : (list T) (@incl T) reflexivity proved by (@incl_refl T) transitivity proved by (@trans_incl T) as Fincl. Lemma list_ext : forall {T} (E F:list T), incl E F -> incl F E -> E == F. intros; constructor; auto. Qed. Inductive full_fin_set {T} (A:list T) : Prop := | full_fin_set_intro : (forall x:T, In x A) -> full_fin_set A. Add Parametric Morphism {T} : full_fin_set with signature (@set_eq T ==> iff) as full_fin_set_eq. Proof. intros ? ? h1; destruct h1. split; intro h1; destruct h1; constructor; intros; [eapply H; auto | eapply H0; auto]. Qed. Add Parametric Morphism {T} : full_fin_set with signature (@incl T ++> impl) as full_fin_set_incl. Proof. intros l l' h1; red; intro h2. constructor. intro x. destruct h2 as [h2]. specialize (h2 x). apply h1 in h2; auto. Qed. Add Parametric Morphism {T} : full_fin_set with signature (flip (@incl T) --> impl) as full_fin_set_incl'. Proof. intros l l' h1; red; intro h2; destruct h2; constructor; intro x; specialize (H x); apply h1 in H; auto. Qed. Lemma in_addl_eq : forall {T} (A:list T) (a:T), In a A -> a :: A == A. intros T A a h1; constructor; red; intros x h2. destruct h2 as [|h2]; subst; auto. right; auto. Qed. Lemma set_eq_nil : forall {T} (A:list T), A == nil -> A = nil. intros T A h1. inversion h1 as [h2]. apply incl_nil' in h2; auto. Qed. Lemma disc_set_eq : forall {T} (A:list T) (a:T), a :: A <=> nil. intros T A a h1. inversion h1. apply incl_nil' in H; discriminate. Qed. Add Parametric Morphism {T} (a:T) : (In a) with signature (@incl T ++> impl) as in_set_incl_impl. intros A B h1; red; intro h2; apply h1 in h2; auto. Qed. Add Parametric Morphism {T} (a:T) : (In a) with signature (set_eq ==> impl) as in_set_impl. intros A B h1; inversion h1 as [h2 h3]; red; intro h4. apply h2 in h4; auto. Qed. Add Parametric Morphism {T} (a:T) : (In a) with signature (set_eq ==> flip impl) as in_set_flip_impl. intros A B h1; inversion h1 as [h2 h3]; red; intro h4. apply h3 in h4; auto. Qed. Add Parametric Morphism {T} (a:T) : (In a) with signature (set_eq ==> iff) as in_set_iff. intros l l' h1; destruct h1 as [h2 h3]; split; intro h4; [apply h2 in h4; auto | apply h3 in h4; auto]. Qed. Add Parametric Morphism {T} (a:T) : (cons a) with signature (set_eq ==> set_eq) as cons_set_eq. intros A B h1. constructor; red; intros c h2; destruct h2 as [|h2]; subst. left; auto. rewrite h1 in h2. right; auto. left; auto. rewrite <- h1 in h2. right; auto. Qed. Add Parametric Morphism {T} (pfdt:type_eq_dec T) (a:T) : (remove pfdt a) with signature (@incl T ==> @incl T) as remove_incl. intros A B h1; red in h1; red; intros x hx; rewrite in_remove_iff in hx; destruct hx as [h2 h3]; subst. rewrite in_remove_iff. split; auto. Qed. Add Parametric Morphism {T} (pfdt:type_eq_dec T) (a:T) : (remove pfdt a) with signature (set_eq ==> set_eq) as remove_set_eq. intros ? ? h1. destruct h1; constructor; apply remove_incl; auto. Qed. Lemma remove_sig_set_eq : forall {T} (pfdt:type_eq_dec T) (l:list T) (a:T) (pfn : ~In a l), map (@proj1_sig _ _) (remove_sig_list pfdt pfn) == remove pfdt a l. intros T hdt l a h1. constructor; red; intros x h2; [rewrite in_map_iff in h2; rewrite in_remove_iff; destruct h2 as [x' [h2 h3]]; destruct x' as [x' hx']; subst; unfold remove_sig_list in h3; rewrite in_map_in_dep_iff in h3; destruct h3 as [x [h3 h4]]; lr_match' h4; subst; try contradiction; try destruct x' as [x' h5]; apply exist_injective in h4; subst; simpl; auto | rewrite in_remove_iff in h2; rewrite in_map_iff; destruct h2 as [h2 h3]; exists (exist (fun c => c <> a) _ h3); split; auto; unfold remove_sig_list; rewrite in_map_in_dep_iff; exists x, h2; lr_match_goal'; try contradiction; f_equal; apply proof_irrelevance]. Qed. Definition cardl {T} (pfdt:type_eq_dec T) (E:list T) := length (list_singularize pfdt E). Lemma addl_eq_iff : forall {T} (pfdt:type_eq_dec T) (A B:list T) a, ~In a A -> (B == a :: A <-> A == (remove pfdt a B) /\ In a B). intros T hdt A B a h1; split; intro h2. split. constructor; red; intros c h3. rewrite in_remove_iff. pose proof (in_cons a _ _ h3) as h4. rewrite <- h2 in h4. split; auto; intro; subst; contradiction. rewrite in_remove_iff in h3. destruct h3 as [h3 h4]. rewrite h2 in h3. destruct h3 as [|h3]; subst. contradict h4; auto. assumption. rewrite h2. left; auto. destruct h2 as [h2 hi]. constructor; red; intros c h3. destruct (hdt c a) as [|h5]; subst. left; auto. pose proof (in_remove_iff hdt B c a) as h6. hr h6. split; auto. rewrite <- h6 in hr. rewrite <- h2 in hr. right; auto. destruct (hdt c a) as [|h4]; subst; auto. destruct h3 as [|h3]; subst. contradict h4; auto. rewrite h2 in h3. rewrite in_remove_iff in h3. destruct h3; auto. Qed. Lemma addl_inj_iff : forall {T} (pfdt:type_eq_dec T) (A B:list T) a b, ~In a A -> ~In a B -> ~In b A -> ~In b B -> (b :: B == a :: A <-> (a = b /\ A == B )). intros T hdt A B a b; intros; split; intro h1. inversion h1 as [h2 h3]. red in h2, h3. destruct (hdt a b) as [|h6]; subst. split; auto. constructor; red; intros x h4. destruct (hdt x b) as [|h5]; subst. contradiction. specialize (h3 _ (in_cons _ _ _ h4)). destruct h3; subst; auto; contradict h5; auto. pose proof (in_cons b _ _ h4) as h5. rewrite h1 in h5. destruct h5 as [|h5]; subst. contradiction. assumption. pose proof (in_eq b B) as h7. rewrite h1 in h7. destruct h7 as [|h7]; subst. contradict h6; auto. contradiction. destruct h1 as [? h1]; subst. rewrite h1; reflexivity. Qed. Lemma addl_inj_iff' : forall {T} (pfdt:type_eq_dec T) (A B:list T) a b, a <> b -> ~In a A -> ~In b B -> In a B -> In b A -> (b :: B == a :: A <-> (remove pfdt b A == remove pfdt a B)). intros T hdt A B a b; intros; split; intro h1. inversion h1 as [h2 h3]. red in h2, h3. constructor; red; intros x h4. rewrite in_remove_iff in h4. destruct h4 as [h4 h5]. destruct (hdt x a) as [|h6]; subst. contradiction. specialize (h3 _ (in_cons a _ _ h4)). destruct h3 as [|h3]; subst. contradict h5; auto. rewrite in_remove_iff. split; auto. rewrite in_remove_iff in h4. destruct h4 as [h4 h5]. destruct (hdt x b) as [|h6]; subst. contradiction. specialize (h2 _ (in_cons b _ _ h4)). destruct h2 as [|h2]; subst. contradict h5; auto. rewrite in_remove_iff; split; auto. inversion h1 as [h2 h3]. red in h2, h3. constructor; red; intros x h4; destruct (hdt x a) as [|he], (hdt x b) as [|he']; subst. left; auto. left; auto. right; auto. destruct h4 as [|h4]; subst. contradict he'; auto. right. specialize (h3 x). hfirst h3. rewrite in_remove_iff; split; auto. specialize (h3 hf). rewrite in_remove_iff in h3. destruct h3; auto. subst. left; auto. right; auto. left; auto. right. destruct h4 as [|h4]; subst. contradict he; auto. specialize (h2 x). hfirst h2. rewrite in_remove_iff; split; auto. specialize (h2 hf). rewrite in_remove_iff in h2; destruct h2; auto. Qed. Add Parametric Morphism {T} (pfdt:type_eq_dec T) : (list_singularize pfdt) with signature (@incl T ++> @incl T) as list_singularize_incl. intros A B h1; red in h1; red; intros a h2. rewrite <- list_singularize_in_iff in h2. apply h1 in h2. rewrite <- list_singularize_in_iff; auto. Qed. Add Parametric Morphism {T} (pfdt:type_eq_dec T) : (list_singularize pfdt) with signature (@set_eq T ==> @set_eq T) as list_singularize_set_eq. intros ? ? h1; destruct h1 as [h1 h2]; red in h1, h2; constructor; intros c h3; rewrite <- list_singularize_in_iff in h3; rewrite <- list_singularize_in_iff; [apply h1 in h3; auto | apply h2 in h3; auto]. Qed. Add Parametric Morphism {T} (pfdt:type_eq_dec T) : (fun l => length (list_singularize pfdt l)) with signature (@set_eq T ==> @eq nat) as length_list_singularize_set_eq. intros l l' h1. apply list_to_sets_eq_length. destruct h1. apply anti_sym_incl_list_to_set; auto. Qed. Add Parametric Morphism {T} (pfdt:type_eq_dec T) : (cardl pfdt) with signature (@set_eq T ==> @eq nat) as cardl_eq. unfold cardl; intro l. induction l as [|a l ih]; simpl. intros ? h2. symmetry in h2. apply set_eq_nil in h2; subst; simpl; auto. intro l'. destruct l' as [|a' l']. intro h2. apply set_eq_nil in h2; discriminate. intro h2. destruct (pfdt a a') as [|h3]; subst. destruct (in_dec pfdt a' l) as [h4 | h4], (in_dec pfdt a' l') as [h5 | h5]. rewrite list_singularize_in_cons; auto. rewrite list_singularize_in_cons; auto. apply ih. do 2 rewrite in_addl_eq in h2; auto. rewrite in_addl_eq in h2. pose proof (ih (a'::l') h2) as h6. simpl in h6. rewrite list_singularize_in_cons; auto. assumption. rewrite list_singularize_nin; auto. rewrite list_singularize_in_cons; auto. simpl. rewrite (in_addl_eq l') in h2. symmetry in h2. rewrite (addl_eq_iff pfdt) in h2. destruct h2 as [h2 h6]. specialize (ih (remove pfdt a' l')). hfirst ih. assumption. specialize (ih hf). rewrite list_singularize_remove, length_remove_no_dup in ih. pose proof h5 as h5'. rewrite (list_singularize_in_iff pfdt) in h5'. apply in_not_nil in h5'. apply not_nil_lt in h5'. omega. apply no_dup_list_singularize. rewrite <- list_singularize_in_iff; auto. assumption. assumption. assumption. rewrite list_singularize_nin; auto. rewrite list_singularize_nin; auto; simpl. f_equal; apply ih. rewrite addl_inj_iff in h2; auto. destruct h2; symmetry; auto. destruct (in_dec pfdt a l) as [h4|h4], (in_dec pfdt a' l') as [h5|h5]. rewrite list_singularize_in_cons. rewrite list_singularize_in_cons. do 2 rewrite in_addl_eq in h2; auto. assumption. assumption. rewrite list_singularize_in_cons; auto. rewrite list_singularize_nin; auto. rewrite in_addl_eq in h2. specialize (ih (a'::l')). apply ih in h2. rewrite list_singularize_nin in h2; auto. assumption. rewrite list_singularize_nin; auto. rewrite list_singularize_in_cons; auto. simpl. rewrite (in_addl_eq l') in h2. symmetry in h2. rewrite (addl_eq_iff pfdt) in h2. destruct h2 as [h2 h6]. specialize (ih (remove pfdt a l') h2). rewrite ih, list_singularize_remove, length_remove_no_dup. rewrite (list_singularize_in_iff pfdt) in h5. apply in_not_nil in h5. apply not_nil_lt in h5. omega. apply no_dup_list_singularize. rewrite <- list_singularize_in_iff; auto. assumption. assumption. assumption. rewrite list_singularize_nin; auto. rewrite list_singularize_nin; auto. simpl; f_equal. pose proof (in_eq a l) as h6. rewrite h2 in h6. destruct h6 as [|h6]; subst. contradict h3; auto. pose proof (in_eq a' l') as h6'. rewrite <- h2 in h6'. destruct h6' as [|h6']; subst. contradiction. specialize (ih (remove pfdt a (a'::l'))). hfirst ih. rewrite <- h2. rewrite remove_cons_eq, nin_remove_eq; auto. reflexivity. specialize (ih hf). rewrite list_singularize_remove, list_singularize_nin, length_remove_no_dup in ih; auto. constructor. rewrite <- list_singularize_in_iff; auto. apply no_dup_list_singularize. right. rewrite <- list_singularize_in_iff; auto. right; auto. Qed. Definition proj1_sig_set_eq {T} {P Q:T->Prop} (l: list {x:T | P x}) (l':list {x:T | Q x}) : Prop := map (@proj1_sig _ P) l == map (@proj1_sig _ Q) l'. Lemma proj1_sig_set_nil1 : forall {T} (P Q:T->Prop) (l': list {x:T | P x}), proj1_sig_set_eq nil l' (P:=P) -> l' = nil. unfold proj1_sig_set_eq; intros ? ? ? l. destruct l; auto; simpl. intro h1. symmetry in h1. apply disc_set_eq in h1; contradiction. Qed. Lemma proj1_sig_set_nil2 : forall {T} (P Q:T->Prop) (l: list {x:T | P x}), proj1_sig_set_eq l nil (Q:=Q) -> l = nil. unfold proj1_sig_set_eq; intros ? ? ? l. destruct l; auto; simpl. intro h1. apply disc_set_eq in h1; contradiction. Qed. End FiniteSets'. bool2-v0-3/README0000644000175000017500000000537613575600217014223 0ustar dwyckoff76dwyckoff76Version 0.3 Alpha bool2. Heretofore, every effort will be made to include a release at the end of each year, before Winter Solstice. This consistency means however, that sometimes releases won't be complete, and this latest release of BooleanAlgebrasIntro2, Version 0.3 Alpha, 2019, defers one last step in the Recursion Theorem (subterm injectivity) to next release. There is still some sloppy old code from Boolean algebras that needs to be cleaned, although much has been cleaned. Version 0.3 finds us switching from Boolean agebras altogether, into, first function algebras, then partial algebras. The handful of definitions and theorems used from function algebras are taken from Dietlinde Lau's "Function Algebras on Finite Sets." Eventually, we are driven to our final course, translating Horst Reichel's "Initial Computability, Algebraic Specifications, and Partial Algebras," into Coq. This huge task will hopefully occupy my main efforts for the next forty or so years! The reason for the switch is complicated and philosophical. The last version (0.2) culminated in Sikorski's Extension Criterion, and this version includes many corollaries and variants of that, like the monomorphism criterion, and finitary/countable versions. We also dabble a bit with atomless Boolean algebras, and a few experiments in original definitions. There are also many new utilities files, and the finite maps library has been made more resistant to the law of excluded middle, and cleaned up. It's still far from fully clean, though. The big news metaphysically, though, is that this software project succeeds so well in interacting with the Sun Spirit that it's no longer safe to use during the day! During peak magnetic fluctuations from the sun, the computer displays signs of digital possession, by glitchy typing, and arbitrary repetitions of keyboard stokes in the shell. It's intermittent, some days working some days not, but it always works at night. Whether or not I'm right in my conclusion is not really the point, but rather the point is code, hidden digital interfaces, and tracing broken symmetries. Hopefully the manner of possession will change as we progress since, this metaphysical divide is the whole point of bool2. The most practical use of bool2 is hardware certification, but the whole fun and excitement comes from the digital spaces we are also certifying, and the mysterious aspect! Digital and financial alchemy, other possibilities, I know little about, but have seen glimpses of. The plan from now on is to upgrade versions of Coq after each release. This version 0.3 still uses Version 8.4pl4 of Coq. An effort will be made to keep the version consistent with the current LTS Ubuntu release. -- Daniel Wyckoff, dwyckoff76@remove.the.excess.around.hotmail.com.kakabool2-v0-3/_CoqProject0000644000175000017500000000146513575574040015475 0ustar dwyckoff76dwyckoff76-R . BooleanAlgebrasIntro2 BoolAlgBasics.v InfiniteOperations.v Subalgebras.v TypeUtilities.v EnsemblesImplicit.v LogicUtilities.v FieldsOfSets.v FunctionProperties.v SetUtilities.v Topology.v ImageImplicit.v DecidableDec.v CSB.v ArithUtilities.v RegularOpenSets.v ListUtilities.v FiniteMaps.v FiniteOperations.v FiniteBags.v SetUtilities2.v TypeUtilities2.v FiniteMaps2.v SetUtilities3.v Homomorphisms.v ExtensionsOfHomomorphisms.v Ordinals.v WellOrders.v Classical_Wf.v ZornsLemma.v Quotients.v Relation_Definitions_Implicit.v AtomlessBooleanAlgebras.v ListUtilities2.v SetUtilities1_5.v ListUtilities3.v SetUtilities2_5.v FunctionProperties2.v FunctionAlgebrasBasics.v GaloisConnection.v Permutations.v PartialAlgebrasBasics.v PartialFunctionAlgebras.v FunctionProperties3.vbool2-v0-3/CSB.v0000644000175000017500000001170213575574055014141 0ustar dwyckoff76dwyckoff76(* This entire file was copied and pasted from Daniel Schepler's "Zorn's Lemma."*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*From Daniel Schepler's "Zorn's Lemma"*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4*) Require Export FunctionProperties. Require Import DecidableDec. Require Import Description. Require Import Classical. (* CSB = Cantor-Schroeder-Bernstein theorem *) Section CSB. Variable X Y:Type. Variable f:X->Y. Variable g:Y->X. Hypothesis f_inj: injective f. Hypothesis g_inj: injective g. Inductive X_even: X->Prop := | not_g_img: forall x:X, (forall y:Y, g y <> x) -> X_even x | g_Y_odd: forall y:Y, Y_odd y -> X_even (g y) with Y_odd: Y->Prop := | f_X_even: forall x:X, X_even x -> Y_odd (f x). Inductive X_odd: X->Prop := | g_Y_even: forall y:Y, Y_even y -> X_odd (g y) with Y_even: Y->Prop := | not_f_img: forall y:Y, (forall x:X, f x <> y) -> Y_even y | f_X_odd: forall x:X, X_odd x -> Y_even (f x). Scheme X_even_coind := Minimality for X_even Sort Prop with Y_odd_coind := Minimality for Y_odd Sort Prop. Scheme X_odd_coind := Minimality for X_odd Sort Prop with Y_even_coind := Minimality for Y_even Sort Prop. Lemma even_odd_excl: forall x:X, ~(X_even x /\ X_odd x). Proof. intro. assert (X_even x -> ~ X_odd x). 2:tauto. pose proof (X_even_coind (fun x:X => ~ X_odd x) (fun y:Y => ~ Y_even y)). apply H. intuition. destruct H1. apply H0 with y. reflexivity. intuition. inversion H2. apply g_inj in H3. apply H1. rewrite <- H3. assumption. intuition. inversion H2. apply H3 with x0. reflexivity. apply f_inj in H3. apply H1. rewrite <- H3. assumption. Qed. Lemma even_odd_excl2: forall y:Y, ~(Y_even y /\ Y_odd y). Proof. intro. assert (Y_odd y -> ~ Y_even y). 2:tauto. pose proof (Y_odd_coind (fun x:X => ~ X_odd x) (fun y:Y => ~ Y_even y)). apply H. intuition. destruct H1. apply H0 with y0. reflexivity. intuition. inversion H2. apply g_inj in H3. apply H1. rewrite <- H3. assumption. intuition. inversion H2. apply H3 with x. reflexivity. apply f_inj in H3. apply H1. rewrite <- H3. assumption. Qed. Definition finv: forall y:Y, (exists x:X, f x = y) -> { x:X | f x = y }. intros. apply constructive_definite_description. destruct H. exists x. red; split. assumption. intros. apply f_inj. transitivity y; trivial. symmetry; trivial. Defined. Definition ginv: forall x:X, (exists y:Y, g y = x) -> { y:Y | g y = x }. intros. apply constructive_definite_description. destruct H. exists x0. red; split. assumption. intros. apply g_inj. transitivity x; trivial; symmetry; trivial. Defined. Definition ginv_odd: forall x:X, X_odd x -> { y:Y | g y = x }. intros. apply ginv. destruct H. exists y. reflexivity. Defined. Definition finv_noteven: forall y:Y, ~ Y_even y -> { x:X | f x = y }. intros. apply finv. apply NNPP. red; intro. contradict H. constructor 1. intro; red; intro. apply H0. exists x. assumption. Defined. Definition CSB_bijection (x:X) : Y := match (classic_dec (X_odd x)) with | left o => proj1_sig (ginv_odd x o) | right _ => f x end. Definition CSB_bijection2 (y:Y) : X := match (classic_dec (Y_even y)) with | left _ => g y | right ne => proj1_sig (finv_noteven y ne) end. Lemma CSB_comp1: forall x:X, CSB_bijection2 (CSB_bijection x) = x. Proof. intro. unfold CSB_bijection; case (classic_dec (X_odd x)). intro. destruct ginv_odd. simpl. unfold CSB_bijection2; case (classic_dec (Y_even x1)). intro. assumption. intro. destruct x0. contradict n. apply g_inj in e. rewrite e. assumption. intro. unfold CSB_bijection2; case (classic_dec (Y_even (f x))). intro. contradict n. inversion y. pose proof (H x). contradict H1; reflexivity. apply f_inj in H. rewrite <- H. assumption. intro. destruct finv_noteven. simpl. apply f_inj. assumption. Qed. Lemma CSB_comp2: forall y:Y, CSB_bijection (CSB_bijection2 y) = y. Proof. intro. unfold CSB_bijection2; case (classic_dec (Y_even y)). intro. unfold CSB_bijection; case (classic_dec (X_odd (g y))). intro. destruct ginv_odd. simpl. apply g_inj. assumption. intro. contradict n. constructor. assumption. intro. destruct finv_noteven. simpl. unfold CSB_bijection; case (classic_dec (X_odd x)). intro. contradict n. rewrite <- e. constructor 2. assumption. trivial. Qed. Theorem CSB: exists h:X->Y, bijective h. Proof. exists CSB_bijection. apply invertible_impl_bijective. exists CSB_bijection2. exact CSB_comp1. exact CSB_comp2. Qed. End CSB. bool2-v0-3/TypeUtilities2.v0000644000175000017500000034163213575574055016441 0ustar dwyckoff76dwyckoff76(*Copyright (C) 2014-2019 Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import FunctionalExtensionality. Require Import ProofIrrelevance. Require Import Description. Require Import NaryFunctions. Require Import Setoid. Require Import NPeano. Require Import Basics. Require Import EqDec. Require Import DecidableDec. Require Import TypeUtilities. Require Import ListUtilities. Require Import ListUtilities2. Require Import FunctionProperties. Require Import FunctionProperties2. Require Import SetUtilities. Require Import SetUtilities2. Require Import ArithUtilities. Require Import LogicUtilities. (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Lemma FiniteT_sig_list : forall {T} (pfdt:type_eq_dec T) (l:list T), FiniteT {x:T | In x l}. intros T hdt l. induction l as [|a l ih]; simpl. pose (fun y:{x:T | False} => proj2_sig y) as f. simpl in f. assert (h1:invertible f). apply bijective_impl_invertible. red; split; red; intros; contradiction. pose proof (invertible_impl_inv_invertible _ h1) as h2. eapply bij_finite. apply empty_finite. apply h2. destruct (in_dec hdt a l) as [h1 | h1]. assert (h2:forall x, In x l <-> In x (a::l)). intro x; split; intro h2. right; auto. destruct h2 as [|h2]; subst; auto. pose (fun (y:{x:T | In x l}) => exist (fun c => In c (a::l)) _ (iff1 (h2 (proj1_sig y)) (proj2_sig y))) as f. assert (h3:invertible f). unfold f. apply bijective_impl_invertible. red; split; red. intros x y h3. apply exist_injective in h3. apply proj1_sig_injective in h3; auto. intro y. destruct y as [x h3]. destruct h3 as [|h3]; subst. exists (exist (fun c => In c l) _ h1). apply proj1_sig_injective; simpl; auto. exists (exist (fun c => In c l) x h3). apply proj1_sig_injective; simpl; auto. eapply bij_finite. apply ih. apply h3. pose proof (add_finite _ ih) as h2. pose (fun (x:option {x:T | In x l}) => match x with | Some y => exist _ _ (or_intror _ (proj2_sig y)) | None => exist (fun c => a = c \/ In c l) a (or_introl _ (eq_refl a)) end) as f. assert (h3:invertible f). unfold f. apply bijective_impl_invertible. red; split; red. intros x y h4. destruct x, y. apply exist_injective in h4. f_equal. apply proj1_sig_injective in h4; auto. apply exist_injective in h4. pose proof (proj2_sig s) as h3; simpl in h3. rewrite h4 in h3. contradiction. apply exist_injective in h4. pose proof (proj2_sig s) as h3; simpl in h3. rewrite <- h4 in h3. contradiction. reflexivity. intro y. destruct y as [y hy]. destruct hy as [|hy]; subst. exists None. auto. exists (Some (exist (fun c => In c l) _ hy)). apply proj1_sig_injective; simpl; auto. eapply bij_finite. apply h2. apply h3. Qed. Lemma FiniteT_card_prod : forall {T U:Type} (pft:FiniteT T) (pfu:FiniteT U), FiniteT_card _ (finite_prod _ _ pft pfu) = (FiniteT_card _ pft) * (FiniteT_card _ pfu). intros T U h1 h2. apply FiniteT_card_cond. rewrite full_set_prod. pose proof (card_fun1_compat (cart_prod (Full_set T) (Full_set U))) as h3. destruct h3 as [h3l h3r]. hfirst h3l. apply cart_prod_fin. rewrite Finite_FiniteT_iff. assumption. rewrite Finite_FiniteT_iff. assumption. specialize (h3l hf). rewrite card_cart_prod in h3l. do 2 rewrite FiniteT_card_fun1_compat. assumption. rewrite Finite_FiniteT_iff. assumption. rewrite Finite_FiniteT_iff. assumption. Qed. Lemma card_nprod_in_l : forall {T:Type} (l:list T), NoDup l -> forall n:nat, FiniteT_card _ (finite_nprod_in_l l n) = (length l)^n. intros T l h0 n. revert l h0. induction n as [|n h1]; simpl; auto. intros l h0. pose proof FiniteT_card_unit as h1. rewrite <- h1. f_equal. apply proof_irrelevance. intros l h0. specialize (h1 l). pose proof (FiniteT_card_prod (list_to_set_finitet l) (finite_nprod_in_l l n)) as h2. terml h2. redterm0 x. fold c in h2. gterml. redterm0 x0. fold c0. pose proof (proof_irrelevance _ c c0) as h5. rewrite <- h5. rewrite h2. f_equal. rewrite card_in_l; auto. specialize (h1 h0); auto. Qed. Lemma false_prod : forall T, equi (False * T) False. intro T. apply (equi_intro (False*T) False (fun i:False*T => False_rect False (fst i))). econstructor; intro x. contradict (fst x). contradiction. Grab Existential Variables. intro; contradiction. Qed. Lemma false_n : forall n:nat, equi (False ^ S n) False. intro; simpl; apply false_prod. Qed. Lemma nth_tuple_prod_seg_fun_depT_nth_to_tuple_prod : forall {T n} {U:T->Type} (x:T^n) (f:seg_fun_depT_nth U x) i (pfi:i < n), nth_tuple_prod (seg_fun_depT_nth_to_tuple_prod x f) i pfi = f i pfi. intros T n; induction n as [|n ih]; simpl. intros; omega. intros U x f i. destruct i as [|i]. intros; f_equal; apply proof_irrelevance. intros hi. specialize (ih U (snd x) (seg_fun_depT_nth_to_S f) i (lt_S_n _ _ hi)). rewrite ih at 1. apply seg_fun_depT_nth_to_S_compat. Qed. Section Transfer. (*Wyckoff*) Lemma in_transfer_dep_r_nprod_to_list_const_iff : forall {T U} {n} {x y:T} (pfe:x = y) (u:(fun _ => U) x) (tpl:T^n) (p:tuple_prod (fun _ => U) tpl), In (transfer_dep_r pfe u (U := fun _ => U)) (nprod_to_list _ _ (transfer_tuple_prod_const p)) <-> in_nprod u (transfer_tuple_prod_const p). intros; subst; split; intro h2; [rewrite transfer_dep_r_eq_refl, <- in_nprod_iff' in h2 | rewrite in_nprod_iff' in h2]; auto. Qed. End Transfer. Section Nprod. (*If [i >= n], then returns [tt]*) Fixpoint nprod_subst {T:Type} {n} i : T^n -> T -> T^n := match n with | 0 => fun _ _ => tt | S n' => match i with | 0 => fun x y => (y, snd x) | S i' => fun x y => (fst x, nprod_subst i' (snd x) y) end end. (*This is a "reverse-directioned" version of [nprod]. It glues [y] to the right side of [x:T^n].*) Fixpoint glue_nprod_elt {T n} : T ^n -> T -> T ^ S n := match n with | 0 => fun _ y => (y, tt) | S n' => fun x y => let (a, x') := x in let z := glue_nprod_elt x' y in (a, z) end. Fixpoint glue_nprod {T m n} : T^m -> T^n -> T^(m + n) := match m with | 0 => fun x y => transfer_nprod (eq_sym (plus_O_n n)) y | S m' => fun x y => (fst x, glue_nprod (snd x) y) end. Lemma glue_nprod_elt_tt : forall {T} (x:T), glue_nprod_elt tt (n := 0) x = (x, tt). auto. Qed. Lemma glue_nprod_tt_l : forall {T n} (x:T^n), glue_nprod tt (m := 0) x = transfer_nprod (eq_sym (plus_0_l _)) x. auto. Qed. Lemma glue_nprod_tt_r : forall {T n} (x:T^n), glue_nprod x tt (n := 0) = transfer_nprod (eq_sym (plus_0_r _)) x. intros T n. induction n as [|n ih]; simpl; intro x; destruct x; simpl; try do 2 f_equal; try apply proof_irrelevance. rewrite ih. gterml. redterm0 x. redterm1 c. fold c0. generalize c0. intro h1. symmetry. gen1. revert h1. simpl. rewrite <- plus_n_O. intros; do 2 rewrite transfer_nprod_eq_refl; auto. Qed. Lemma glue_nprod_S_l : forall {T m n} (x:T^m) (y:T^n) (a:T), glue_nprod (a, x) y (m := S m) = (a, glue_nprod x y). auto. Qed. Lemma nprod_to_list_glue : forall {T m n} (x:T^m) (y:T^n), nprod_to_list _ _ (glue_nprod x y) = nprod_to_list _ _ x ++ nprod_to_list _ _ y. intros T m. induction m as [|m ih]; simpl. intros. rewrite transfer_nprod_eq_refl; auto. intros n x y. destruct x; simpl; f_equal; auto. Qed. Lemma nprod_to_list_glue_nprod_of_list : forall (l l':list nat), nprod_to_list _ _ (glue_nprod (nprod_of_list _ l) (nprod_of_list _ l')) = l ++ l'. intro l. induction l as [|a l ih]; simpl; intro l'. rewrite transfer_nprod_eq_refl, nprod_to_list_nprod_of_list_eq; auto. f_equal; auto. Qed. Lemma nprod_of_list_app_nprod_to_list : forall {T m n} (x:T^m) (y:T^n), nprod_of_list _ (nprod_to_list _ _ x ++ nprod_to_list _ _ y) = transfer_nprod_r (length_app_nprod_to_list _ _) (glue_nprod x y). intros T m. induction m as [|m ih]; simpl; intros. rewrite nprod_of_list_nprod_to_list_eq; f_equal. apply proof_irrelevance. rewrite transfer_nprod_eq_refl; auto. destruct x as [t x]; simpl. rewrite transfer_nprod_r_S; simpl; f_equal. specialize (ih n x y). rewrite ih. f_equal; apply proof_irrelevance. Qed. Lemma glue_nprod_of_list : forall {T} (l l':list T), glue_nprod (nprod_of_list _ l) (nprod_of_list _ l') = transfer_nprod (app_length _ _) (nprod_of_list _ (l ++ l')). intros T l. induction l as [|a l ih]; simpl. intro; f_equal; apply proof_irrelevance. intro l'. rewrite transfer_nprod_S; simpl. f_equal. rewrite ih. repeat f_equal; apply proof_irrelevance. Qed. Lemma plus_nprod_eq0_iff : forall {n} (x:nat^n), plus_nprod x = 0 <-> all_sing_nprod x 0. intros; rewrite plus_nprod_eq, plusl_eq0_iff, all_sing_nprod_to_list_iff; tauto. Qed. Lemma plus_nprod_eq1_iff : forall {n} (x:nat^n), plus_nprod x = 1 <-> exists! r s (y:nat^r) (z:nat^s) (pfe:n = r + (S s)), x = transfer_nprod_r pfe (glue_nprod y (1, z) (n := S s)) /\ all_sing_nprod y 0 /\ all_sing_nprod z 0. intros n x. rewrite plus_nprod_eq, plusl_eq1_iff; split; intro h1; destruct h1 as [c [[d h3] h2]]; red in h3; destruct h3 as [[h3 [h4 h6]] h5]. exists (length c); red; split. exists (length d); red; split. exists (nprod_of_list _ c); red; split. exists (nprod_of_list _ d); red; split. pose proof (f_equal (@length _) h3) as h7. rewrite length_nprod_to_list, app_length in h7; simpl in h7. subst. exists (eq_refl _). red; split. rewrite transfer_nprod_r_eq_refl. split. rewrite <- nprod_of_list_cons. rewrite (glue_nprod_of_list c (1::d)) at 1; simpl. simpl in h3. symmetry. gen1. rewrite <- h3; simpl. intro h7. rewrite nprod_of_list_nprod_to_list_eq. apply transfer_nprod_undoes_transfer_nprod_r. split; rewrite all_sing_nprod_of_list_iff; auto. intros; apply proof_irrelevance. intros y h7. destruct h7 as [h8 [[h9 [h10 h12]] h11]]; subst. apply (all_sing_nprod_eq (nprod_of_list _ d) y 0); auto. rewrite all_sing_nprod_of_list_iff; auto. intros y h7. destruct h7 as [z [[h9 [[h10 [h14 h13]] h12]] h11]]; subst. apply (all_sing_nprod_eq (nprod_of_list _ c) y 0); auto. rewrite all_sing_nprod_of_list_iff; auto. intros m h7. destruct h7 as [y [[z [[h10 [h14 h13]] h12]] h11]]; subst. pose proof (f_equal (@length _) h3) as h15. rewrite app_length, length_nprod_to_list in h15; simpl in h15. apply plus_inj_l in h15. apply S_inj in h15; auto. intros m h7. destruct h7 as [r [[y [[q [[h14 h15] h13]] h12]] h11]]; subst. destruct h15 as [[h15 [h17 h19]] h18]; subst. rewrite transfer_nprod_r_eq_refl in h3. rewrite nprod_to_list_glue in h3; simpl in h3. apply (all_sing_app_eq _ _ _ _ 0 1) in h3; auto. destruct h3; subst. apply length_nprod_to_list. rewrite all_sing_nprod_to_list_iff; auto. rename h3 into y. destruct h4 as [z [[h4 h7] h8]]; subst. red in h7. destruct h7 as [[h7 [h9 h11]] h10]; subst. exists (nprod_to_list _ _ y). red; split. exists (nprod_to_list _ _ z). red; split. simpl. rewrite transfer_nprod_r_eq_refl. split. rewrite nprod_to_list_glue; simpl; auto. split; rewrite all_sing_nprod_to_list_iff; auto. intros l' h12. destruct h12 as [h12 h13]. rewrite transfer_nprod_r_eq_refl, nprod_to_list_glue in h12; simpl in h12. apply (all_sing_app_eq _ _ _ _ 0 1) in h12; auto with arith. destruct h12; auto. destruct h13; auto. destruct h13; auto. intros l h12. destruct h12 as [l' h12]. red in h12. destruct h12 as [[h12 h13] h14]. rewrite transfer_nprod_r_eq_refl, nprod_to_list_glue in h12; simpl in h12. destruct h13 as [h13 h15]. apply (all_sing_app_eq _ _ _ _ 0 1) in h12; destruct h12; auto. rewrite all_sing_nprod_to_list_iff; auto. Qed. Lemma times_nprod_eq0_iff : forall {n} (x:nat^n), times_nprod x = 0 <-> in_nprod 0 x. intros; rewrite times_nprod_l, timesl_eq0_iff, in_nprod_iff'; tauto. Qed. Lemma times_nprod_eq1_iff : forall {n} (x:nat^n), times_nprod x = 1 <-> all_sing_nprod x 1. intros; rewrite times_nprod_l, timesl_eq1_iff, <- all_sing_nprod_to_list_iff; tauto. Qed. (*Splits an nprod like this : split_prod (a,b,c,d,e,f,g,h,i,j,k) 8 = (a,b,c,d,e,f,g,h), (i,j,k)*) Fixpoint split_nprod {T n} : T^n -> forall i, i <= n -> T^i * T^(n-i) := match n with | 0 => fun _ i pfi => let pfi' := eq_sym (le0 _ pfi) in let pfm := zero_minus i in (transfer_nprod pfi' tt, transfer_nprod pfm tt) | S n' => fun (x:T^(S n')) i => match i with | 0 => fun _ => (transfer_nprod (eq_refl 0) tt, transfer_nprod (minus_n_O (S n')) x) | S i' => fun pfi => let pfi' := le_S_n _ _ pfi in let z := split_nprod (snd x) i' pfi' in ((fst x, fst z), snd z) end end. Lemma glue_nprod_elt_pr : forall {T n} (x:T^n) (a y:T), glue_nprod_elt (a, x) y (n := S n) = (a, glue_nprod_elt x y). auto. Qed. Lemma split_nprod0 : forall {T n} (x:T^n) (pf:0 <= n), split_nprod x 0 pf = (tt, transfer_nprod (minus_n_O _) x). intros T n. destruct n as [|n]; simpl. intros x h1. rewrite transfer_nprod_eq_refl; do 2 f_equal. destruct x; auto. Qed. Lemma split_nprod_l : forall {T n} (x:T^n) (pf:n <= n), split_nprod x n pf = (x, transfer_nprod (eq_sym (minus_same _)) tt). intros T n. induction n as [|n ih]; simpl. intros x h0. destruct x. rewrite transfer_nprod_eq_refl. do 2 f_equal; apply proof_irrelevance. intros x h0. destruct x as [a x']. specialize (ih x'). simpl. rewrite ih; simpl. do 2 f_equal; apply proof_irrelevance. Qed. Lemma split_nprod_S : forall {T n} (x:T^n) (a:T) i (pfi:S i <= S n), let pfi' := le_S_n _ _ pfi in split_nprod (a, x) (S i) pfi (n:=S n) = ((a, fst (split_nprod x i pfi')), snd (split_nprod x i pfi')). intros ? n; induction n; simpl; intros; auto. Qed. Lemma glue_nprod_elt_compat : forall {T n} (x:T^n) (y:T) (pfn:n <= S n), split_nprod (glue_nprod_elt x y) n pfn = (x, transfer_nprod (eq_sym (S_minus _)) (y, tt)). intros T n. induction n as [|n ih]. intros ? y h1. simpl; rewrite transfer_nprod_eq_refl. destruct x; do 2 f_equal; apply proof_irrelevance. intros x y h1. destruct x as [a x']. specialize (ih x' y (le_n_Sn _)). rewrite glue_nprod_elt_pr. rewrite split_nprod_S. pose proof (proof_irrelevance _ (le_n_Sn n) (le_S_n _ _ h1)) as he. rewrite <- he, ih. simpl; do 2 f_equal; apply proof_irrelevance. Qed. Lemma glue_nprod_compat : forall {T m n} (x:T^m) (y:T^n) (pf:m <= m + n), split_nprod (glue_nprod x y) m pf = (x, transfer_nprod (eq_sym (plus_minus_same_r _ _)) y). intros T m. induction m as [|m ih]. intros n x y h1; simpl. destruct x. rewrite transfer_nprod_eq_refl, split_nprod0. repeat f_equal; apply proof_irrelevance. intros n x y h1. destruct x as [a x]; simpl. specialize (ih n x y). specialize (ih (le_S_n m (m + n) h1)). rewrite ih. simpl; do 2 f_equal; apply proof_irrelevance. Qed. (*This is used directly in the star superposition operation in FunctionAlgebrasBasics.*) Definition star_split {T m n i} (x:T^(m + (pred n))) (pf0:0 < i) (pfi:i < n) : T^(pred i) * (T^m * T^(n-i)) := let pfi' := le_trans _ _ _ (le_pred_n i) (le_pred_sum _ _ _ pfi) in let pfm := le_pred_sum_minus _ _ _ pfi in let pfe := sum_pred_minus _ _ _ pf0 pfi in let pr := split_nprod x (pred i) pfi' in let qr := split_nprod (snd pr) m pfm in (fst pr, (fst qr, transfer_nprod pfe (snd qr))). End Nprod. Section FinT. (*This is intended for developments in which you need to know a bit more about the exact elements of a given primitive finite type, than you can with [FiniteT] -- for instance, canonical indices and lists associated with the finite type, etc. [FinT] is ideal for switching back and forth between the type-set [A] and the initial segment [0 .. |A| - 1]. The inner details of [FinT] are quite cumbersome, and the hope is that the interface will obviate the need to deal with the depth directly.*) Record FinT (T:Type) : Type := {finU : Type; finU_eq_dec : type_eq_dec finU; fin_lU:list finU; fin_fU : fun_in_dep fin_lU T; fin_inv : inv_dep fin_fU }. Lemma FinT_FiniteT : forall {T}, FinT T -> FiniteT T. intros T h1. induction h1. pose proof (inv_dep_sig _ fin_inv0) as h1. eapply bij_finite. apply (FiniteT_sig_list finU_eq_dec0). apply h1. Qed. Lemma bij_fin_fU : forall {T} (FT:FinT T), bij_dep (fin_fU _ FT). intros; apply inv_impl_bij_dep; apply fin_inv. Qed. Lemma surj_fin_fU : forall {T} (FT:FinT T), surj_dep (fin_fU _ FT). intros; apply bij_impl_surj_dep; apply bij_fin_fU. Qed. Lemma inj_fin_fU : forall {T} (FT:FinT T), inj_dep (fin_fU _ FT). intros; apply bij_impl_inj_dep; apply bij_fin_fU. Qed. Definition listT {T} (FT:FinT T) : list T := map_in_dep (fin_fU _ FT). Definition listU {T} (FT:FinT T) : list (finU _ FT) := fin_lU _ FT. Lemma listT_eq_nil_iff : forall {T} (FT:FinT T), listT FT = nil <-> listU FT = nil. unfold listT; intros. rewrite map_in_dep_eq_nil_iff; tauto. Qed. Lemma in_listT : forall {T} (F:FinT T) x, In x (listT F). unfold listT; intros T F. destruct F; simpl. intro x. rewrite in_map_in_dep_iff. pose proof (inv_impl_bij_dep _ fin_inv0) as h1. red in h1. destruct h1 as [h1 h2]. red in h1, h2. specialize (h2 x); auto. Qed. Definition fin_fT {T} (FT:FinT T) (x:T) : {y:finU _ FT | In y (listU FT)} := let f := inverse_dep (fin_fU _ FT) (fin_inv _ FT) in exist (fun c => In c (listU FT)) (projT1 f x) (projT2 f _). Lemma fin_fT_compat : forall {T} (FT:FinT T), (forall (x:finU _ FT) (pfi:In x (listU FT)), fin_fT FT (fin_fU _ FT x pfi) = exist _ x pfi) /\ (forall x:T, fin_fU _ FT (proj1_sig (fin_fT FT x)) (proj2_sig (fin_fT FT x)) = x). unfold fin_fT; intros T FT. pose proof (inverse_dep_compat (fin_fU _ FT) (fin_inv _ FT)) as h1; simpl in h1. destruct h1 as [h1 h2]. split; auto; intros; apply proj1_sig_injective; simpl; auto. Qed. Definition fin_fT' {T} (FT:FinT T) (x:T) : finU _ FT := proj1_sig (fin_fT FT x). Lemma in_fin_fT' : forall {T} (FT:FinT T) (x:T), In (fin_fT' FT x) (listU FT). intros; unfold fin_fT'; apply proj2_sig. Qed. Lemma fin_fT_compat' : forall {T} (FT:FinT T), (forall (x:finU _ FT) (pfi:In x (listU FT)), fin_fT' FT (fin_fU _ FT x pfi) = x) /\ (forall x:T, fin_fU _ FT (fin_fT' FT x) (in_fin_fT' _ _) = x). intros T FT. unfold fin_fT'. pose proof (inverse_dep_compat (fin_fU T FT) (fin_inv T FT)) as h1. simpl in h1. destruct h1 as [h1 h2]. split; intro x. intro hi. specialize (h1 x hi); simpl in h1; simpl. destruct (inverse_dep (fin_fU T FT) (fin_inv T FT)); simpl; auto. specialize (h2 x); simpl in h1; simpl. clear h1. destruct FT; simpl; simpl in h2. rewrite <- h2. apply f_equal_dep. symmetry. gterml. redterm1 x0. fold c; fold c in h2. destruct c; simpl; simpl in h2. rewrite h2; auto. Qed. Lemma inv_fin_fT : forall {T} (FT:FinT T), invertible (fin_fT FT). intros T FT. pose proof (fin_fT_compat FT) as h1. destruct h1 as [h1 h2]. pose proof (intro_invertible (fin_fT FT) (fun x => fin_fU _ FT (proj1_sig x) (proj2_sig x))) as h3. apply h3; auto. intro y; destruct y; auto. Qed. Lemma in_listU : forall {T} (FT:FinT T) (x:T), In (fin_fT' FT x) (listU FT). unfold fin_fT'; intros; apply proj2_sig. Qed. Lemma in_listU_iff : forall {T} (FT:FinT T) (y:finU _ FT), In y (listU FT) <-> exists x:T, y = fin_fT' FT x. intros T FT y; split; intro h1. exists (fin_fU _ FT y h1). pose proof (fin_fT_compat FT) as h2; destruct h2 as [h2 h3]. specialize (h2 y h1). pose proof (f_equal (@proj1_sig _ _) h2) as h4. simpl in h4. rewrite <- h4 at 1. unfold fin_fT'. rewrite h2; simpl; auto. destruct h1 as [x ?]; subst; apply in_listU. Qed. Lemma in_listT' : forall {T} (FT:FinT T) (x:finU _ FT) (pf:In x (fin_lU _ FT)), In (fin_fU _ FT x pf) (listT FT). intros; apply in_listT. Qed. Definition map_lT {T} (FT:FinT T) (l:list T) : list {x:finU _ FT | In x (listU FT)} := map (fin_fT FT) l. Definition map_lT' {T} (FT:FinT T) (l:list T) : list (finU _ FT) := map (fin_fT' FT) l. Definition map_lT_cons {T} (FT:FinT T) (l:list T) (a:finU _ FT) : list {x:finU _ FT | In x (a::listU FT)} := map_in_dep (fun x (pf:In x l) => exist (fun c => In c (a::listU FT)) (fin_fT' FT x) (in_cons _ _ _ (in_listU _ _))). Lemma in_map_lT_iff : forall {T} (FT:FinT T) (l:list T) (y:{x:finU _ FT | In x (listU FT)}), In y (map_lT FT l) <-> In (fin_fU _ FT (proj1_sig y) (proj2_sig y)) l. unfold map_lT; intros T FT l y. pose proof (fin_fT_compat FT) as hf; destruct hf as [ha hb]. rewrite in_map_iff. split; intro h1. destruct h1 as [x [h1 h2]]; subst. rewrite hb; auto. destruct y as [y hy]. exists (fin_fU _ FT y hy). split. rewrite ha; auto. simpl in h1; auto. Qed. Lemma in_map_lT_iff' : forall {T} (FT:FinT T) (l:list T) (y:finU _ FT), In y (map_lT' FT l) <-> exists (pfy:In y (listU FT)), In (fin_fU _ FT y pfy) l. unfold map_lT'; intros T FT l y. pose proof (fin_fT_compat' FT) as h0; destruct h0 as [ha hb]. rewrite in_map_iff. split; intro h1. destruct h1 as [x [h1 ?]]; subst. exists (in_listU _ _). erewrite f_equal_dep. rewrite hb at 1. apply H. reflexivity. destruct h1 as [h1 h2]. exists (fin_fU _ FT y h1). split; auto. Qed. Lemma in_map_lT_cons_iff : forall {T} (FT:FinT T) (l:list T) (a:finU _ FT) (y:{x:finU _ FT | In x (a::listU FT)}), In y (map_lT_cons FT l a) <-> exists (pfi:In (proj1_sig y) (listU FT)), In (fin_fU _ FT (proj1_sig y) pfi) l. unfold map_lT_cons; intros T FT l a y. pose proof (fin_fT_compat' FT) as h1; destruct h1 as [h1 h2]. rewrite in_map_in_dep_iff. split; intro h3. destruct h3 as [x [h3 ?]]; subst; simpl. exists (in_listU _ _). erewrite f_equal_dep. rewrite h2 at 1. apply h3. reflexivity. destruct h3 as [h3 h4]. exists _, h4. apply proj1_sig_injective; simpl; auto. Qed. Definition map_lU {T} (FT:FinT T) (l:list (finU _ FT)) (pfi:incl l (listU FT)) : list T := map_in_dep (fun x (pf:In x l) => fin_fU _ FT x (pfi _ pf)). Lemma in_map_lU_iff : forall {T} (FT:FinT T) (l:list (finU _ FT)) (pfi:incl l (listU FT)) (x:T), In x (map_lU FT l pfi) <-> In (fin_fT' FT x) l. unfold map_lU; intros T FT l hi x. pose proof (fin_fT_compat' FT) as h2. destruct h2 as [h2 h3]. rewrite in_map_in_dep_iff. split; intro h1. destruct h1 as [y [h1 ?]]; subst. rewrite h2; auto. exists _, h1. erewrite f_equal_dep. rewrite h3 at 1. reflexivity. reflexivity. Qed. Definition map_lU_cons {T} (FT:FinT T) a : list {x:finU _ FT | In x (a::(listU FT))} := map_sig_cons (listU FT) a. Lemma length_map_lU_cons : forall {T} (FT:FinT T) a, length (map_lU_cons FT a) = length (listU FT). unfold map_lU_cons; intros; apply length_map_sig_cons. Qed. Lemma in_map_lU_cons_iff : forall {T} (FT:FinT T) a x, In x (map_lU_cons FT a) <-> In (proj1_sig x) (listU FT). unfold map_lU_cons; intros; rewrite in_map_sig_cons_iff; tauto. Qed. Definition map_lU_sig {T} (FT:FinT T) (l:list {x:finU _ FT | In x (listU FT)}) : list T := map (fun x => fin_fU _ FT (proj1_sig x) (proj2_sig x)) l. Lemma in_map_lU_sig_iff : forall {T} (FT:FinT T) (l:list {x:finU _ FT | In x (listU FT)}) (x:T), In x (map_lU_sig FT l) <-> In (fin_fT FT x) l. unfold map_lU_sig; intros T FT l x; rewrite in_map_iff. pose proof (fin_fT_compat FT) as h3; destruct h3 as [h3 h4]. split; intro h1. destruct h1 as [y [h1 h2]]; subst. rewrite h3. destruct y as [y hy]; simpl; auto. exists (fin_fT FT x). rewrite h4. split; auto. Qed. Lemma in_map_lU_sig_iff' : forall {T} (FT:FinT T) (l:list {x:finU _ FT | In x (listU FT)}) (x:T), In x (map_lU_sig FT l) <-> In (fin_fT' FT x) (map_proj1_sig l). intros; rewrite in_map_lU_sig_iff, (in_map_proj1_sig_iff (P := fun c => In c (listU FT))); split; intro h1. exists (in_fin_fT' _ _). eapply in_subst1. apply proj1_sig_injective; simpl. unfold fin_fT'. reflexivity. assumption. destruct h1 as [h1 h2]. unfold fin_fT; simpl. pose proof (inverse_dep_compat (fin_fU _ FT) (fin_inv _ FT)) as h3. simpl in h3. destruct h3 as [h3 h4]. destruct (inverse_dep (fin_fU _ FT) (fin_inv _ FT)); simpl; simpl in h4, h3. erewrite f_equal_dep. apply h2. specialize (h3 _ h1). rewrite <- h3. f_equal. pose proof (fin_fT_compat' FT) as h5. destruct h5 as [h5 h6]. erewrite f_equal_dep. rewrite h6. reflexivity. reflexivity. Qed. Definition cardU {T} (FT:FinT T) : nat := cardl (finU_eq_dec _ FT) (listU FT). Definition fin_eqU {T U} (F:FinT T) (F':FinT U) : Prop := cardU F = cardU F'. Lemma refl_fin_eqU : forall T, Reflexive _ (@fin_eqU T T). intros; red; intro; red; auto. Qed. Lemma sym_fin_eqU : forall {T U} {FT: FinT T} {FU:FinT U}, fin_eqU FT FU-> fin_eqU FU FT. intros ? ? ? ? h1; red in h1; red; auto. Qed. Lemma trans_fin_eqU : forall {T U V} {FT:FinT T} {FU:FinT U} {FV:FinT V}, fin_eqU FT FU -> fin_eqU FU FV -> fin_eqU FT FV. intros ? ? ? ? ? ? h1 h2; red in h1; red in h2; red; congruence. Qed. Definition fin_eqU' {T} (F:FinT T) (G:FinT T) : Prop := fin_eqU F G. Lemma refl_fin_eqU' : forall T, Reflexive _ (@fin_eqU' T). unfold fin_eqU; intros; red; intro F. destruct F as [U hU lU]; split; try apply equi_refl; auto. Qed. Lemma symm_fin_eqU' : forall T, Symmetric _ (@fin_eqU' T). unfold fin_eqU', fin_eqU; intro; red; intros F F'; destruct F, F'; f_equal; simpl; intro; auto. Qed. Lemma trans_fin_eqU' : forall T, Transitive _ (@fin_eqU' T). unfold fin_eqU; intro; red; intros F G H; destruct F, G, H; intros hfg hgh. simpl in hfg, hgh; simpl; congruence. Qed. Add Parametric Relation (T:Type) : (FinT T) (@fin_eqU' T) reflexivity proved by (refl_fin_eqU' T) symmetry proved by (symm_fin_eqU' T) transitivity proved by (trans_fin_eqU' T) as FinT_eq'. Add Parametric Morphism T : (@listT T) with signature (@fin_eqU' T ==> @set_eq T) as FinT_listT. intros l l'; destruct l, l'; simpl; intro; constructor; red; intros; apply in_listT. Qed. Lemma in_listT_iff : forall {T} (F:FinT T) y, In y (listT F) <-> exists (x : (finU _ F)) (pfx:In x (listU F)), fin_fU _ F x pfx = y. unfold listT; intros; apply in_map_in_dep_iff. Qed. Definition finT_eq_dec {T} (FT:FinT T) : type_eq_dec T. red; intros x y. pose proof (in_listT FT x) as hx. pose proof (in_listT FT y) as hy. rewrite in_listT_iff in hx, hy. pose proof (inverse_dep_compat _ (fin_inv _ FT)) as h1. simpl in h1. destruct h1 as [h1 h2]. destruct (finU_eq_dec _ FT (projT1 (inverse_dep (fin_fU _ FT) (fin_inv _ FT)) x) (projT1 (inverse_dep (fin_fU _ FT) (fin_inv _ FT)) y)) as [h3 | h3]. left. destruct hx as [x' [hx' ?]], hy as [y' [hy' ?]]; subst. do 2 rewrite h1 in h3; subst; f_equal; apply proof_irrelevance. right. destruct hx as [x' [hx' ?]], hy as [y' [hy' ?]]; subst. do 2 rewrite h1 in h3; subst; contradict h3. pose proof (inv_impl_bij_dep _ (fin_inv _ FT)) as h4. red in h4. destruct h4 as [h4 h5]. apply h4 in h3; auto. Defined. Definition cardT {T} (FT:FinT T) : nat := cardl (finT_eq_dec FT) (listT FT). Lemma card_eq : forall {T} (FT:FinT T), cardT FT = cardU FT. unfold cardT, cardU, cardl, listT, listU; intros; erewrite length_list_singularize_map_in_dep_inj. reflexivity. apply inv_impl_inj_dep. apply fin_inv. Qed. Lemma length_listT : forall {T} (FT:FinT T), length (listT FT) = length (listU FT). unfold listT, listU; intros; rewrite length_map_in_dep; auto. Qed. Lemma length_listU : forall {T} (FT:FinT T), length (listU FT) = length (listT FT). unfold listT, listU; intros; rewrite length_map_in_dep; auto. Qed. Definition FinT_False : FinT False := Build_FinT False False type_eq_dec_false nil (fun_in_dep_nil _ _) (inv_dep_nil _). Lemma cardT_false : cardT FinT_False = 0. auto. Qed. Lemma cardU_false : cardU FinT_False = 0. auto. Qed. Definition list_to_FinT {T} (pfdt:type_eq_dec T) (l:list T) : FinT {x:T | In x l} := Build_FinT {x:T | In x l} T pfdt l (id_in_dep_sig' l) (inv_id_in_dep_sig' _). Definition FinU_cons {T:Type} (FT:FinT T) (a:finU _ FT) : FinT {x:finU _ FT| In x (a::listU FT)} := Build_FinT _ (finU _ FT) (finU_eq_dec _ FT) (a::listU FT) (id_in_dep_sig' (a::listU FT)) (inv_id_in_dep_sig' (a::listU FT)). Definition FinT_inv {T U:Type} (FT:FinT T) (f:T->U) (pfinv:invertible f) : FinT U := Build_FinT U (finU _ FT) (finU_eq_dec _ FT) (listU FT) (fun_in_dep_fun_compose (fin_fU _ FT) f) (inv_fun_in_dep_fun_compose (fin_fU _ FT) f (fin_inv _ FT) pfinv). Lemma in_FinU_cons : forall {T} (FT:FinT T) (a:finU _ FT), In a (listU (FinU_cons FT a)). unfold listU, FinU_cons, fin_lU; intros; left; auto. Qed. Lemma in_FinU_cons' : forall {T} (FT:FinT T) (a x:finU _ FT), In x (listU FT) -> In x (listU (FinU_cons FT a)). unfold listU, FinU_cons, fin_lU; intros; right; auto. Qed. Lemma FinT_induction : forall (P:forall T, FinT T -> Prop), P False (FinT_False) -> (forall T (FT:FinT T), P T FT -> forall a, P _ (FinU_cons FT a)) -> (forall T U (FT:FinT T) (FU:FinT U), fin_eqU FT FU -> P T FT -> P U FU) -> forall T (FT:FinT T), P T FT. intros P h1 h2 h3 T FT. revert h1 h2 h3. destruct FT. revert fin_inv0; revert fin_fU0; revert finU_eq_dec0; revert T. induction fin_lU0. intros T h1 f h2 h3 h4 h5. pose proof (false_dep_eq _ h2 ) as h6. do 2 red in f; simpl in f. gterm0. specialize (h5 False T FinT_False c). hfirst h5. constructor. specialize (h5 hf h3); auto. intros T hdu f h1 h2 h3 h4. destruct (in_dec hdu a fin_lU0) as [h5 | h5]. pose (Build_FinT T finU0 hdu fin_lU0 (fun_from_cons f) (inv_from_cons_in f h5 h1)) as FT. specialize (IHfin_lU0 T hdu (fun_from_cons f) (inv_from_cons_in f h5 h1) h2 h3 h4). gterm0. specialize (h4 T T FT c). hfirst h4. red; simpl. pose proof (in_addl_eq _ _ h5) as h7. symmetry in h7. apply cardl_eq; auto. specialize (h4 hf IHfin_lU0); auto. pose proof (bij_impl_inj_dep _ (inv_impl_bij_dep _ h1)) as h6. pose proof (inv_from_cons_ran f h5 h1) as h7. simpl in h7. specialize (IHfin_lU0 {x:T | x <> f a (in_eq _ _)} hdu (fun_from_cons_ran _ (inv_impl_inj_dep f h1) h5) h7 h2 h3 h4). specialize (h3 _ _ IHfin_lU0). simpl in h3. specialize (h3 a). term0 h3. gterm0. specialize (h4 _ T c c0). apply h4. red; simpl; auto. assumption. Qed. Definition list_singularizeT {T} (FT:FinT T) : list T := list_singularize (finT_eq_dec FT) (listT FT). Definition list_singularizeU {T} (FT:FinT T) : list (finU _ FT) := list_singularize (finU_eq_dec _ FT) (listU FT). Lemma length_list_singularizeT : forall {T} (FT:FinT T), length (list_singularizeT FT) = length (list_singularize (finU_eq_dec _ FT) (listU FT)). unfold list_singularizeT, listT; intros. erewrite length_list_singularize_map_in_dep_inj. reflexivity. apply inj_fin_fU. Qed. Lemma in_list_singularizeT : forall {T} (F:FinT T) x, In x (list_singularizeT F). unfold list_singularizeT. intros; rewrite <- list_singularize_in_iff; apply in_listT. Qed. Lemma length_list_singularizeU : forall {T} (FT:FinT T), length (list_singularizeU FT) = length (list_singularize (finU_eq_dec _ FT) (listU FT)). auto. Qed. Lemma in_list_singularizeU : forall {T} (FT:FinT T) x, In (fin_fT' FT x) (list_singularizeU FT). unfold list_singularizeU, fin_fT'; intros; rewrite <- list_singularize_in_iff; apply proj2_sig. Qed. Lemma in_listU_singularize : forall {T} (FT:FinT T) (x:finU _ FT), In x (listU FT) -> In x (list_singularizeU FT). unfold list_singularizeU; intros; rewrite <- list_singularize_in_iff; auto. Qed. Add Parametric Morphism T : (@list_singularizeT T) with signature (@fin_eqU' T ==> @set_eq T) as FinT_list_singularize. unfold list_singularizeT, listT, finT_eq_dec; intros F F';simpl; intro hcrd; constructor; red; intros x hx; pose proof (in_listT F x) as h1; destruct F, F'; rewrite <- list_singularize_in_iff, in_map_in_dep_iff; rewrite <- list_singularize_in_iff, in_map_in_dep_iff in hx; destruct hx as [y [hx ?]]; subst; unfold listT in h1; simpl in h1; simpl in hx; simpl; simpl in hcrd; [pose proof (inverse_dep_compat _ fin_inv1) as h2| pose proof (inverse_dep_compat _ fin_inv0) as h2]; simpl in h2; destruct h2 as [h2 h3]; [specialize (h3 (fin_fU0 y hx)) | specialize (h3 (fin_fU1 y hx))]; terml h3; redterm0 x; exists _, c; rewrite <- h3; auto. Qed. Definition fin_eqT {T U} (FT:FinT T) (FT':FinT U) : Prop := cardT FT = cardT FT'. Lemma refl_fin_eqT : forall T, Reflexive _ (@fin_eqT T T). intros; red; intro; red; auto. Qed. Lemma sym_fin_eqT : forall {T U} (FT: FinT T) (FU:FinT U), fin_eqT FT FU-> fin_eqT FU FT. intros ? ? ? ? h1; red in h1; red; auto. Qed. Lemma trans_fin_eqT : forall {T U V} (FT:FinT T) (FU:FinT U) (FV:FinT V), fin_eqT FT FU -> fin_eqT FU FV -> fin_eqT FT FV. intros ? ? ? ? ? ? h1 h2; red in h1; red in h2; red; congruence. Qed. Lemma fin_eqTU : forall {T U} (FT:FinT T) (FT':FinT U), fin_eqT FT FT' -> fin_eqU FT FT'. unfold fin_eqT, fin_eqU; intros ? ? FT FT' h1; erewrite card_eq in h1; rewrite h1; apply card_eq. Qed. Lemma fin_eqUT : forall {T U} (FT:FinT T) (FT':FinT U), fin_eqU FT FT' -> fin_eqT FT FT'. unfold fin_eqT, fin_eqU; intros ? ? FT FT' h1; erewrite card_eq; rewrite h1; rewrite card_eq; auto. Qed. Lemma in_fin_eqT : forall {T} (FT:FinT T) (a:finU _ FT), In a (listU FT) -> fin_eqT (FinU_cons FT a) FT. unfold FinU_cons; intros; red; unfold cardT, cardl, listT, id_in_dep_sig', id_prop_dep_sig'; symmetry. rewrite (length_list_singularize_map_in_dep_inj (finU_eq_dec _ FT)). symmetry. rewrite (length_list_singularize_map_in_dep_inj (finU_eq_dec _ FT)). simpl. rewrite list_singularize_in_cons. reflexivity. assumption. apply inv_impl_inj_dep. simpl. apply bij_impl_inv_dep. red; split; red. intros x y h1 h2 h3. apply exist_injective in h3; auto. intro y. exists (proj1_sig y). ex_goal. destruct y; auto. exists hex. apply proj1_sig_injective; auto. apply inv_impl_inj_dep. apply (fin_inv _ FT). Qed. Definition fin_eqT' {T} (F:FinT T) (G:FinT T) : Prop := fin_eqT F G. Lemma refl_fin_eqT' : forall T, Reflexive _ (@fin_eqT' T). unfold fin_eqT; intros; red; intro F. destruct F as [U hU lU]; split; try apply equi_refl; auto. Qed. Lemma symm_fin_eqT' : forall T, Symmetric _ (@fin_eqT' T). unfold fin_eqT', fin_eqT; intro; red; intros F F'; destruct F, F'; f_equal; simpl; intro; auto; simpl in H; simpl. Qed. Lemma trans_fin_eqT' : forall T, Transitive _ (@fin_eqT' T). unfold fin_eqT', fin_eqT; intro; red; intros F G H; destruct F, G, H; intros hfg hgh. simpl in hfg, hgh; simpl; congruence. Qed. Add Parametric Relation (T:Type) : (FinT T) (@fin_eqT' T) reflexivity proved by (refl_fin_eqT' T) symmetry proved by (symm_fin_eqT' T) transitivity proved by (trans_fin_eqT' T) as FinT_eqT'. Lemma cardT_eq : forall {T} (FT:FinT T), cardT FT = length (list_singularizeT FT). unfold list_singularizeT, listT, cardT, cardl; intros T FT; pose proof (inv_impl_bij_dep _ (fin_inv _ FT)) as h1; red in h1; destruct h1 as [h1 h2]; unfold finT_eq_dec, listT; auto; destruct FT; auto; simpl. Qed. Lemma cardT_set_eq : forall {T} (FT FT':FinT T) (a:T), listT FT' == a :: listT FT -> cardT FT = cardT FT'. intros; do 2 rewrite cardT_eq. unfold list_singularizeT. pose proof (type_eq_dec_eq (finT_eq_dec FT) (finT_eq_dec FT')) as h1. rewrite h1. pose proof (in_listT FT a) as h2. rewrite in_addl_eq in H; auto. generalize (finT_eq_dec FT'). intro h3. symmetry. erewrite length_list_singularize_set_eq. reflexivity. assumption. Qed. Lemma lt_cardT_not_nil : forall {T} (FT:FinT T) (pf0:0 < cardT FT), listU FT <> nil. intros T FT h1. unfold cardT, cardl in h1. apply lt_length in h1. rewrite list_singularize_nil_iff in h1. unfold listT in h1. rewrite map_in_dep_eq_nil_iff in h1; auto. Qed. Lemma lt_cardU_not_nil : forall {T} (FT:FinT T) (pf0:0 < cardU FT), listU FT <> nil. intros; apply lt_cardT_not_nil. rewrite card_eq; auto. Qed. Lemma in_fin_eqU : forall {T} (FT:FinT T) (a:finU _ FT), In a (listU FT) -> fin_eqU (FinU_cons FT a) FT. intros; red; unfold cardT, cardU, cardl, FinU_cons, listT; intros; simpl; unfold id_in_dep_sig', id_prop_dep_sig'; simpl. rewrite list_singularize_in_cons; simpl; auto. Qed. Lemma list_singularizeU_in_cons : forall {T} (FT:FinT T) (a:finU _ FT), In a (listU FT) -> list_singularizeU (FinU_cons FT a) = list_singularizeU FT. unfold list_singularizeU; intros; simpl. rewrite leq_iff. ex_goal. rewrite list_singularize_in_cons; auto. exists hex. intros i hi; simpl. symmetry. gen0. generalize hi. rewrite list_singularize_in_cons; auto; intros; apply nth_lt_functional3. apply finU_eq_dec. Qed. Lemma length_list_singularizeT_in_cons : forall {T} (FT:FinT T) (a:finU _ FT), In a (listU FT) -> length (list_singularizeT (FinU_cons FT a)) = length (list_singularizeT FT). intros; do 2 rewrite length_list_singularizeT; simpl; rewrite list_singularize_in_cons; auto. Qed. Lemma length_list_singularizeT_nin_cons : forall {T} (FT:FinT T) (a:finU _ FT), ~In a (listU FT) -> length (list_singularizeT (FinU_cons FT a)) = S (length (list_singularizeT FT)). intros; do 2 rewrite length_list_singularizeT; simpl. rewrite list_singularize_nin; auto. Qed. Lemma list_singularizeT_in_cons : forall {T} (FT:FinT T) (a:finU _ FT), In a (listU FT) -> map_proj1_sig (list_singularizeT (FinU_cons FT a)) = list_singularizeU FT. intros T FT a h1. unfold map_proj1_sig. rewrite map_eq_iff. ex_goal. rewrite length_list_singularizeT_in_cons, length_list_singularizeU. unfold list_singularizeT. unfold listT. rewrite (length_list_singularize_map_in_dep_inj (finU_eq_dec T FT) (finT_eq_dec FT)); auto. apply inj_fin_fU. assumption. exists hex. intros n h2. unfold list_singularizeT, listT. erewrite nth_lt_functional3. rewrite nth_lt_list_singularize_map_in_dep. simpl. gen0. simpl. intro h3. symmetry. gen0. revert h3. rewrite list_singularize_in_cons; auto. intros; apply nth_lt_functional3. Grab Existential Variables. apply inj_fin_fU. rewrite length_list_singularizeT in h2; simpl in h2. simpl; auto. Qed. Lemma cardT_in_cons : forall {T} (FT:FinT T) (a:finU _ FT), In a (listU FT) -> cardT (FinU_cons FT a) = cardT FT. intros; do 2 rewrite cardT_eq; apply length_list_singularizeT_in_cons; auto. Qed. Lemma cardU_in_cons : forall {T} (FT:FinT T) (a:finU _ FT), In a (listU FT) -> cardU (FinU_cons FT a) = cardU FT. intros; do 2 rewrite <- card_eq; apply cardT_in_cons; auto. Qed. (*This is the index in the [NoDup]-ed defining list for a [FinT] type.*) Definition lindT {T} (FT:FinT T) (x:T) : nat := lind (finT_eq_dec FT) (list_singularizeT FT) x (iff1 (list_singularize_in_iff _ _ _) (in_listT _ _)). (*This amounts to the same, but for the type underlying [T].*) Definition lindU {T} (FT:FinT T) (x:finU _ FT) (pfi:In x (listU FT)): nat := lind (finU_eq_dec _ FT) (list_singularizeU FT) x (iff1 (list_singularize_in_iff _ _ _) pfi). Lemma lt_lindT : forall {T} (FT:FinT T) x, lindT FT x < cardT FT. unfold cardT, cardl, lindT; intros; apply lt_lind. Qed. Lemma inj_lindT : forall {T} (FT:FinT T), Injective (lindT FT). intros; red; intros ? ? h1. unfold lindT, list_singularizeT in h1. rewrite lind_inj_iff in h1; auto. Qed. Lemma lt_lindU : forall {T} (FT:FinT T) x (pfi:In x (listU FT)), lindU FT x pfi < cardU FT. unfold cardU, cardl, lindT; intros; apply lt_lind. Qed. Lemma inj_lindU : forall {T} (FT:FinT T), inj_dep (lindU FT). intros; red; intros ? ? h1 h2 h3. unfold lindU, list_singularizeU in h3. rewrite lind_inj_iff in h3; auto. Qed. (* This is the index in the defining list, which can presumably have duplicates, and is therefore subject to vague indices. It's not injective.*) Definition lindT' {T} (FT:FinT T) (x:T) : nat := lind (finT_eq_dec FT) (listT FT) x (in_listT _ _). Lemma lt_lindT' : forall {T} (FT:FinT T) (x:T), lindT' FT x < length (listT FT). unfold listT, lindT'; intros; apply lt_lind. Qed. Definition nthT {T} (FT:FinT T) i (pfi:i < cardT FT) : T := let pfi' := lt_eq_trans _ _ _ pfi (cardT_eq FT) in nth_lt (list_singularizeT FT) i pfi. Definition nthU {T} (FT:FinT T) i (pfi:i < cardU FT) : finU _ FT := nth_lt (list_singularizeU FT) i pfi. Definition nthT' {T} (FT:FinT T) i (pfi:i < length (listT FT)) : T := nth_lt (listT FT) i pfi. Definition nthU' {T} (FT:FinT T) i (pfi:i < length (listU FT)) : finU _ FT := nth_lt (listU FT) i pfi. Lemma nthT_functional : forall {T} (FT:FinT T) i i' (pfi : i < cardT FT) (pfi': i' < cardT FT), i = i' -> nthT FT i pfi = nthT FT i' pfi'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma nthU_functional : forall {T} (FT:FinT T) i i' (pfi : i < cardU FT) (pfi': i' < cardU FT), i = i' -> nthU FT i pfi = nthU FT i' pfi'. intros; subst; f_equal; apply proof_irrelevance. Qed. Lemma in_nthU : forall {T} (FT:FinT T) i (pfi:i < cardU FT), In (nthU FT i pfi) (listU FT). unfold nthU; intros. pose proof (in_nth_lt _ _ pfi) as h1. rewrite <- list_singularize_in_iff in h1; auto. Qed. Lemma in_nthU' : forall {T} (FT:FinT T) i (pfi:i < length (listU FT)), In (nthU' FT i pfi) (listU FT). unfold nthU'; intros; apply in_nth_lt. Qed. Lemma fin_fU_nthU : forall {T} (FT:FinT T) i (pfi:i < cardU FT) (pfin:In (nthU FT i pfi) (listU FT)), let pfi' := lt_eq_trans _ _ _ pfi (eq_sym (card_eq _)) in fin_fU _ FT (nthU FT i pfi) pfin = nthT FT i pfi'. intros T FT i hi hin; simpl. unfold nthU, list_singularizeU, nthT, list_singularizeT, listT. erewrite nth_lt_functional3. rewrite nth_lt_list_singularize_map_in_dep. repeat f_equal; apply proof_irrelevance. Grab Existential Variables. apply inj_fin_fU. Qed. Lemma fin_fU_nthU' : forall {T} (FT:FinT T) i (pfi:i < length (listU FT)) (pfin:In (nthU' FT i pfi) (listU FT)), let pfi' := lt_eq_trans _ _ _ pfi (length_listU _) in fin_fU _ FT (nthU' FT i pfi) pfin = nthT' FT i pfi'. unfold nthU', nthT', listU, listT; intros T FT i hi hin; simpl. rewrite nth_lt_map_in_dep'. apply f_equal_dep. f_equal. apply proof_irrelevance. Qed. Lemma fin_fT_nthT : forall {T} (FT:FinT T) i (pfi:i < cardT FT), let pfi' := lt_eq_trans _ _ _ pfi (card_eq FT) in fin_fT FT (nthT FT i pfi) = exist (fun c => In c (listU FT)) (nthU FT i pfi') (in_nthU _ _ pfi'). intros T FT i hi; simpl. unfold fin_fT, nthU. apply proj1_sig_injective; simpl. pose proof (fin_fT_compat FT) as h1. destruct h1 as [h1 h2]. specialize (h1 (nthU FT i (lt_eq_trans _ _ _ hi (card_eq FT))) (in_nthU _ _ _)). rewrite fin_fU_nthU in h1. simpl in h1. unfold fin_fT in h1. apply exist_injective in h1. erewrite nthT_functional. rewrite h1 at 1. unfold nthU. auto. auto. Qed. Lemma fin_fT_nthT' : forall {T} (FT:FinT T) i (pfi:i < cardT FT), let pfi' := lt_eq_trans _ _ _ pfi (card_eq FT) in fin_fT' FT (nthT FT i pfi) = nthU FT i pfi'. unfold fin_fT'; intros; rewrite fin_fT_nthT; simpl; auto. Qed. Lemma fin_fT_nthT'' : forall {T} (FT:FinT T) i (pfi:i < length (listT FT)), let pfi' := lt_eq_trans _ _ _ pfi (length_listT FT) in fin_fT FT (nthT' FT i pfi) = exist (fun c => In c (listU FT)) (nthU' FT i pfi') (in_nthU' _ _ _). unfold fin_fT; intros T FT i hi; apply proj1_sig_injective; simpl. pose proof (inverse_dep_compat (fin_fU _ FT) (fin_inv _ FT)) as h2. simpl in h2; destruct h2 as [h2 h3]. gtermr. specialize (h2 y (in_nthU' _ _ _)). fold y. rewrite <- h2. f_equal. unfold nthT', listT. rewrite nth_lt_map_in_dep'. apply f_equal_dep. unfold y, nthU'. f_equal. apply proof_irrelevance. Qed. Lemma fin_fT_nthT''' : forall {T} (FT:FinT T) i (pfi:i < length (listT FT)), let pfi' := lt_eq_trans _ _ _ pfi (length_listT FT) in fin_fT' FT (nthT' FT i pfi) = nthU' FT i pfi'. unfold fin_fT'; intros; rewrite fin_fT_nthT''; simpl; auto. Qed. Lemma nth_lindT : forall {T} (FT:FinT T) (y:T) (pflt:lindT FT y < cardT FT), nthT FT (lindT FT y) pflt = y. intros T FT y hhlt. unfold nthT,list_singularizeT, lindT, listT. erewrite nth_lt_functional3. rewrite nth_lt_lind_eq; auto. Qed. Lemma nth_lindU : forall {T} (FT:FinT T) (y:finU _ FT) (pfi:In y (listU FT)) (pflt:lindU FT y pfi < cardU FT), nthU FT (lindU FT y pfi) pflt = y. intros T FT y hi hhlt. unfold nthU,list_singularizeU, lindU, listU, list_singularizeU. erewrite nth_lt_functional3. rewrite lind_compat. reflexivity. Qed. Lemma lind_nthT : forall {T} (FT:FinT T) i (pfi:i < cardT FT), lindT FT (nthT FT i pfi) = i. unfold lindT, nthT; intros. rewrite lind_nth_lt_no_dup; auto. apply no_dup_list_singularize. Qed. Lemma lind_nthU : forall {T} (FT:FinT T) i (pfi:i < cardU FT), lindU FT (nthU FT i pfi) (in_nthU _ _ _) = i. unfold lindU, nthU; intros. rewrite lind_nth_lt_no_dup; auto. apply no_dup_list_singularize. Qed. Lemma surj_seg_nthT : forall {T} (FT:FinT T), surj_seg (nthT FT). intros T FT; red; intro y. exists (lindT FT y), (lt_lindT _ _). apply nth_lindT. Qed. Lemma inj_seg_nthT : forall {T} (FT:FinT T), inj_seg (nthT FT). unfold nthT; intros T FT; red; intros x y hx hy h1. apply nth_lt_inj in h1; auto. apply no_dup_list_singularize. Qed. Lemma listT_eq_iff : forall {T} (FT:FinT T) (l:list T), listT FT = l <-> exists pfe:length (listT FT) = length l, forall i (pfi:i < length (listT FT)), let pfi' := lt_eq_trans _ _ _ pfi pfe in nthT' FT i pfi = nth_lt l i pfi'. intros T FT l; split; intro h1; subst. exists eq_refl. intros i hi hi'. apply nth_lt_functional3. destruct h1 as [h1 h2]; simpl in h2. unfold listT. rewrite map_in_dep_eq_iff. pose proof h1 as h1'. rewrite length_listT in h1'. exists h1'. intros i hi. pose proof (fin_fU_nthU' FT i hi (in_nth_lt _ _ _)) as h3. simpl in h3. unfold listU in h3. rewrite h3 at 1. rewrite h2. f_equal; apply proof_irrelevance. Qed. Lemma listT_set_eq_iff : forall {T} (FT:FinT T) (l:list T), listT FT == l <-> forall x, In x l. intros T FT l; split; intro h1. intro. rewrite <- in_set_iff. apply in_listT. apply h1. constructor; red; intros; auto. apply in_listT. Qed. Lemma no_dup_list_singularizeT : forall {T} (FT:FinT T), NoDup (list_singularizeT FT). unfold list_singularizeT; intros; apply no_dup_list_singularize. Qed. Lemma lt_cardT : forall {T} (FT:FinT T) i, i < cardT FT -> listU FT <> nil. unfold cardT, cardl, listT; intros ? ? h1. pose proof (map_in_dep_eq_nil_iff (fin_fU T FT)) as h2. rewrite <- h2. pose proof (list_singularize_nil_iff (finT_eq_dec FT) (map_in_dep (fin_fU T FT))) as h3. rewrite <- h3. eapply lt_length. Qed. Lemma lt_cardT_length : forall {T} (FT:FinT T) i, i < cardT FT -> 0 < length (listT FT). unfold cardT, cardl; intros ? ? ? h1. apply lt_length in h1. apply not_nil_lt. unfold listT in h1. rewrite list_singularize_nil_iff in h1. rewrite map_in_dep_eq_nil_iff in h1; auto. rewrite listT_eq_nil_iff; auto. Qed. Lemma lt_length_cardT : forall {T} (FT:FinT T) i, i < length (listT FT) -> 0 < cardT FT. unfold cardT, cardl; intros ? ? ? h1. apply lt_length in h1. apply not_nil_lt. unfold listT. rewrite list_singularize_nil_iff. rewrite map_in_dep_eq_nil_iff; auto. rewrite listT_eq_nil_iff in h1; auto. Qed. Lemma lt_cardT' : forall {T} (FT:FinT T), T -> 0 < cardT FT. intros T FT x; eapply lt_length_cardT. pose proof (lt_lindT FT x) as h1. apply lt_cardT_length in h1. apply h1. Qed. Lemma cardT_nin_cons : forall {T} (FT:FinT T) (a:finU _ FT), ~ In a (listU FT) -> cardT (FinU_cons FT a) = S (cardT FT). intros T FT a h1; do 2 rewrite cardT_eq. apply length_list_singularizeT_nin_cons; auto. Qed. Lemma lt_cardU : forall {T} (FT:FinT T) i, i < cardU FT -> listU FT <> nil. intros; eapply lt_cardT; rewrite card_eq; apply H. Qed. Lemma lt_cardU_length : forall {T} (FT:FinT T) i, i < cardU FT -> 0 < length (listU FT). intros; rewrite <- card_eq in H; rewrite length_listU; eapply lt_cardT_length; apply H. Qed. Lemma lt_length_cardU : forall {T} (FT:FinT T) i, i < length (listU FT) -> 0 < cardU FT. intros; rewrite <- card_eq; rewrite length_listU in H. apply lt_length_cardT in H; auto. Qed. Lemma lt_cardU' : forall {T} (FT:FinT T), T -> 0 < cardU FT. intros; rewrite <- card_eq; apply lt_cardT'; auto. Qed. Lemma lt_cardT_cons : forall {T} (FT:FinT T) (a:finU _ FT), 0 < cardT (FinU_cons FT a). intros T FT a. pose proof (lt_length_cardT (FinU_cons FT a) 0) as h1. hfirst h1. rewrite length_listT. apply not_nil_lt. pose proof (in_FinU_cons FT a) as h2. eapply in_not_nil. apply h2. auto. Qed. Lemma lt_cardU_cons : forall {T} (FT:FinT T) (a:finU _ FT), 0 < cardU (FinU_cons FT a). intros; rewrite <- card_eq; apply lt_cardT_cons. Qed. Lemma cardU_nin_cons : forall {T} (FT:FinT T) (a:finU _ FT), ~ In a (listU FT) -> cardU (FinU_cons FT a) = S (cardU FT). intros; rewrite <- card_eq, cardT_nin_cons, card_eq; auto. Qed. Definition hdU {T} (FT:FinT T) (pf0:0 < cardU FT) : finU _ FT := hd' (listU FT) (lt_cardU FT _ pf0). Lemma in_hdU : forall {T} (FT:FinT T) (pf0:0 < cardU FT), In (hdU FT pf0) (listU FT). unfold hdU, listU; intros; apply in_hd'. Qed. Definition hdT {T} (FT:FinT T) (pf0:0 < cardT FT) : T := fin_fU _ FT (hdU _ (lt_eq_trans _ _ _ pf0 (card_eq _))) (in_hdU _ _). Lemma hd_fin_lU_eq' : forall {T} (FT:FinT T) (pfn:listU FT <> nil) (pf0:0 < cardU FT), let pf0' := lt_cardU_length _ _ pf0 in hd' (listU FT) pfn = nth_lt (listU FT) 0 pf0'. intros T FT hn h0 h1. rewrite nth_lt0'. f_equal. apply proof_irrelevance. Qed. Lemma hdT_eq: forall {T} (FT:FinT T) (pf0:0 < cardT FT) (pfn:listU FT <> nil), hdT FT pf0 = fin_fU T FT (hd' (listU FT) pfn) (in_hd' _ _). unfold hdT, hdU; intros; apply f_equal_dep; f_equal. apply proof_irrelevance. Qed. Definition FinT_in_tl {T:Type} (FT:FinT T) (pf0:0 < cardU FT) (pfin:In (hdU FT pf0) (tl (listU FT))) : FinT T := Build_FinT T (finU _ FT) (finU_eq_dec _ FT) (tl (listU FT)) (tl_in_dep (fin_fU _ FT)) (inv_tl _ (lt_cardU FT 0 pf0) pfin (fin_inv _ FT)). Definition FinT_nin_tl {T:Type} (FT:FinT T) (pf0:0 < cardU FT) (pfnin:~In (hdU FT pf0) (tl (listU FT))) : FinT {x:T | x <> hdT FT (lt_eq_trans _ _ _ pf0 (eq_sym (card_eq _)))} := let pflt := lt_eq_trans _ _ _ pf0 (eq_sym (card_eq _)) in let pfn := lt_cardU_not_nil _ pf0 in Build_FinT _ (finU _ FT) (finU_eq_dec _ FT) (tl (listU FT)) (transfer_sig_fun_dep' (iff_pred_neq_r (eq_sym (hdT_eq FT pflt (lt_cardU _ _ pf0)))) (tl_in_dep_ran (fin_fU _ FT) (lt_cardU _ _ _) (inj_fin_fU _) pfnin)) (inv_transfer_sig_fun_dep' _ _ (inv_tl_ran _ _ _ _ (fin_inv _ FT))). Definition minT {T} (FT:FinT T) (pf:0 < cardT FT) := nthT FT 0 pf. Definition maxT {T} (FT:FinT T) (pf:0 < cardT FT) := nthT FT (pred (cardT FT)) (lt_pred_n _ _ pf). Definition minU {T} (FT:FinT T) (pf:0 < cardU FT) := nthU FT 0 pf. Definition maxU {T} (FT:FinT T) (pf:0 < cardU FT) := nthU FT (pred (cardU FT)) (lt_pred_n _ _ pf). Lemma minT_functional : forall {T} (FT:FinT T) (pf pf':0 < cardT FT), minT FT pf = minT FT pf'. intros; f_equal; apply proof_irrelevance. Qed. Lemma maxT_functional : forall {T} (FT:FinT T) (pf pf':0 < cardT FT), maxT FT pf = maxT FT pf'. intros; f_equal; apply proof_irrelevance. Qed. Lemma minU_functional : forall {T} (FT:FinT T) (pf pf':0 < cardU FT), minU FT pf = minU FT pf'. intros; f_equal; apply proof_irrelevance. Qed. Lemma maxU_functional : forall {T} (FT:FinT T) (pf pf':0 < cardU FT), maxU FT pf = maxU FT pf'. intros; f_equal; apply proof_irrelevance. Qed. Lemma in_minT : forall {T} (FT:FinT T) (pf:0 < cardT FT), In (minT FT pf) (listT FT). unfold minT; intros; apply in_listT. Qed. Lemma in_maxT : forall {T} (FT:FinT T) (pf:0 < cardT FT), In (maxT FT pf) (listT FT). unfold minT; intros; apply in_listT. Qed. Lemma in_minU : forall {T} (FT:FinT T) (pf:0 < cardU FT), In (minU FT pf) (listU FT). unfold minU; intros; apply in_nthU. Qed. Lemma in_maxU : forall {T} (FT:FinT T) (pf:0 < cardU FT), In (maxU FT pf) (listU FT). unfold maxU; intros; apply in_nthU. Qed. Lemma fin_fU_minU : forall {T} (FT:FinT T) (pf0:0 < cardU FT), let pf0' := lt_eq_trans _ _ _ pf0 (eq_sym (card_eq FT)) in fin_fU _ FT (minU FT pf0) (in_minU _ pf0) = minT FT pf0'. unfold minT, minU; intros; apply fin_fU_nthU. Qed. Lemma fin_fU_maxU : forall {T} (FT:FinT T) (pf0:0 < cardU FT), let pf0' := lt_eq_trans _ _ _ pf0 (eq_sym (card_eq FT)) in fin_fU _ FT (maxU FT pf0) (in_maxU _ _) = maxT FT pf0'. unfold maxT, maxU; intros; rewrite fin_fU_nthU; apply f_equal_dep; rewrite card_eq; auto. Qed. Lemma fin_fT_minT : forall {T} (FT:FinT T) (pf0:0 < cardT FT), let pf0' := lt_eq_trans _ _ _ pf0 (card_eq FT) in fin_fT FT (minT FT pf0) = exist _ (minU FT pf0') (in_minU FT _). unfold fin_fT, minT, minU; intros; apply proj1_sig_injective; simpl. pose proof (inverse_dep_compat (fin_fU T FT) (fin_inv T FT)) as h1. simpl in h1. destruct h1 as [h1 h2]; auto. specialize (h1 (nthU FT 0 (lt_eq_trans _ _ _ pf0 (card_eq FT))) (in_nthU _ _ _)). rewrite fin_fU_nthU in h1. erewrite nthT_functional. rewrite h1 at 1; auto. reflexivity. Qed. Lemma fin_fT_maxT : forall {T} (FT:FinT T) (pf0:0 < cardT FT), let pf0' := lt_eq_trans _ _ _ pf0 (card_eq FT) in fin_fT FT (maxT FT pf0) = exist _ (maxU FT pf0') (in_maxU FT _). unfold fin_fT, maxT, minT, minU; intros; apply proj1_sig_injective; simpl. pose proof (inverse_dep_compat (fin_fU T FT) (fin_inv T FT)) as h1. simpl in h1. destruct h1 as [h1 h2]. specialize (h1 (nthU FT (pred (cardU FT)) (lt_pred_n _ _ (lt_eq_trans _ _ _ pf0 (card_eq _)))) (in_nthU _ _ _)). rewrite fin_fU_nthU in h1. erewrite nthT_functional. rewrite h1 at 1. unfold maxU; apply nthU_functional. destruct h1; auto. rewrite card_eq; auto. Qed. Lemma fin_fT_minT' : forall {T} (FT:FinT T) (pf0:0 < cardT FT), let pf0' := lt_eq_trans _ _ _ pf0 (card_eq FT) in fin_fT' FT (minT FT pf0) = minU FT pf0'. unfold fin_fT'; intros; rewrite fin_fT_minT; simpl; auto. Qed. Lemma fin_fT_maxT' : forall {T} (FT:FinT T) (pf0:0 < cardT FT), let pf0' := lt_eq_trans _ _ _ pf0 (card_eq FT) in fin_fT' FT (maxT FT pf0) = maxU FT pf0'. unfold fin_fT'; intros; rewrite fin_fT_maxT; simpl; auto. Qed. Lemma hd_list_singularizeT' : forall {T} (FT:FinT T) (pfn:list_singularizeT FT <> nil), hd' (list_singularizeT FT) pfn = minT FT (not_nil_lt _ pfn). unfold list_singularizeT, minT, nthT, list_singularizeT; intros; erewrite <- nth_lt0; auto. Qed. Lemma hdT_cons : forall {T} (FT:FinT T) (a:finU _ FT) (pf0:0 < cardT (FinU_cons FT a)), hdT (FinU_cons FT a) pf0 = fin_fU _ (FinU_cons FT a) a (in_eq _ _). unfold hdT, nthT; simpl; unfold id_in_dep_sig', id_prop_dep_sig', list_singularizeT; intros; f_equal; apply proof_irrelevance. Qed. Lemma nthT_in_cons : forall {T} (FT:FinT T) (a:finU _ FT) i (pfi:i < cardT (FinU_cons FT a)) (pfin:In a (listU FT)), let pfi' := lt_eq_trans _ _ _ pfi (in_fin_eqT _ _ pfin) in nthT (FinU_cons FT a) i pfi = exist (fun c => In c (a::listU FT)) (fin_fT' FT (nthT FT i pfi')) (in_cons _ _ _ (in_fin_fT' _ _)). intros T FT a i hi hin h2. unfold minT, FinU_cons, nthT, list_singularizeT, listT, id_in_dep_sig', id_prop_dep_sig'. gterml. rename x into b. redterm2 b. rename c into l. redterm0 l. rename c into l'. redterm1 l. rename c into hdu. fold l' hdu. redterm0 l'. rename c into f. pose proof (nth_lt_functional3 l i) as h3. pose proof (nth_lt_list_singularize_map_in_dep (finU_eq_dec _ FT) hdu f i) as h4. hfirst h4. apply inv_impl_inj_dep. unfold f. simpl. apply bij_impl_inv_dep; red; split; red. intros x y hx hy h1. apply exist_injective in h1; auto. intro y. destruct y as [y hy]. exists y, hy; auto. specialize (h4 hf). hfirst h4. simpl; simpl in hi; auto. eapply lt_eq_trans. apply hi. unfold cardT, cardl. rewrite list_singularize_in_cons. unfold listT. gterml. redterm0 x. redterm0 c. redterm0 c0. rename c1 into g. pose proof (length_list_singularize_map_in_dep_inj (finU_eq_dec T FT) (finT_eq_dec (FinU_cons FT a)) g) as h5. hfirst h5. unfold g. simpl. unfold id_in_dep_sig', id_prop_dep_sig'. simpl. red. intros ? ? ? ? h6. apply exist_injective in h6; auto. specialize (h5 hf0). fold g. simpl in h5; simpl. rewrite h5. rewrite list_singularize_in_cons; auto. assumption. specialize (h4 hf0). simpl in h4. unfold l'. simpl. unfold f in h4. simpl in h4. apply proj1_sig_injective; simpl. pose proof (f_equal (@proj1_sig _ _) h4) as h5. simpl in h5. erewrite nth_lt_functional3. rewrite h5 at 1. gtermr. redterm0 y. pose proof (fin_fT_compat' FT) as h6. destruct h6 as [h6 h7]. generalize hf0 h2. simpl. rewrite list_singularize_in_cons. intros h8 h9. symmetry. erewrite nth_lt_functional3. rewrite nth_lt_list_singularize_map_in_dep at 1. rewrite h6. reflexivity. assumption. Grab Existential Variables. apply inv_impl_inj_dep. apply fin_inv. Qed. Lemma nthT_nin_cons : forall {T} (FT:FinT T) (a:finU _ FT) (pfn: ~In a (listU FT)) (pf0:0 < cardT (FinU_cons FT a)) i (pfi:i < cardT (FinU_cons FT a)) (pflt0:i > 0), let pfi' := lt_eq_trans _ _ _ (mono_pred_lt _ _ (lt_eq_trans _ _ _ pfi (cardT_nin_cons FT _ pfn)) pflt0) (card_eq FT) in proj1_sig (nthT (FinU_cons FT a) i pfi) = nthU FT (pred i) pfi'. intros T FT a hn h0 i hi h1 hi'. unfold nthT, nthU, fin_fT', fin_fT; simpl. unfold list_singularizeT, listT. erewrite nth_lt_functional3. rewrite (nth_lt_list_singularize_map_in_dep (finU_eq_dec _ FT)). simpl. gen0. simpl. rewrite list_singularize_nin; auto. simpl. intro h2. destruct i as [|i]. omega. simpl. unfold list_singularizeU. rewrite nth_lt_compat. apply nth_indep. auto with arith. Grab Existential Variables. apply inj_fin_fU. simpl. rewrite list_singularize_nin. simpl. pose proof hi as h2. rewrite cardT_nin_cons in h2. rewrite card_eq in h2; auto. assumption. assumption. Qed. Lemma nthU_in_cons : forall {T} (FT:FinT T) (a:finU _ FT) i (pfi:i < cardU (FinU_cons FT a)) (pfin:In a (listU FT)), let pfi' := lt_eq_trans _ _ _ pfi (in_fin_eqU _ _ pfin) in nthU (FinU_cons FT a) i pfi = nthU FT i pfi'. intros T FT a i hi hin; simpl. unfold nthU, fin_fT', fin_fT; simpl. symmetry. gen0. generalize hi. unfold list_singularizeU, cardU, cardl; simpl. erewrite list_singularize_in_cons; intros; f_equal. apply proof_irrelevance. assumption. Qed. Lemma nthU_nin_cons : forall {T} (FT:FinT T) (a:finU _ FT) (pfin:~In a (listU FT)) i (pfpos:0 In c (a::listU FT)) (fin_fT' FT (minT FT pf0')) (in_cons _ _ _ (in_fin_fT' _ _)). unfold minT; intros; apply nthT_in_cons. Qed. Lemma minT_nin_cons : forall {T} (FT:FinT T) (a:finU _ FT) (pf0:0 < cardT (FinU_cons FT a)), ~In a (listU FT) -> minT (FinU_cons FT a) pf0 = fin_fU _ (FinU_cons FT a) a (in_eq _ _). unfold minT, cardT, cardl, hdT, nthT; simpl; unfold id_in_dep_sig', id_prop_dep_sig'; simpl; intros; unfold listT, list_singularizeT, listT. erewrite f_equal_dep. rewrite nth_lt_list_singularize_map_in_dep. simpl. unfold id_in_dep_sig', id_prop_dep_sig'. apply proj1_sig_injective; simpl. gen0. simpl. rewrite list_singularize_nin; auto. intros; simpl; auto. instantiate (2 := 0); simpl; auto. reflexivity. Grab Existential Variables. apply bij_impl_inj_dep. apply inv_impl_bij_dep. apply fin_inv. apply finU_eq_dec. simpl. rewrite list_singularize_nin; auto. simpl; auto with arith. Qed. Lemma minU_in_cons : forall {T} (FT:FinT T) (a:finU _ FT) (pf0:0 < cardU (FinU_cons FT a)) (pfi:In a (listU FT)), let pf0' := lt_eq_trans _ _ _ pf0 (in_fin_eqU _ _ pfi) in minU (FinU_cons FT a) pf0 = minU FT pf0'. unfold minU; intros; apply nthU_in_cons; auto. Qed. Lemma minU_nin_cons : forall {T} (FT:FinT T) (a:finU _ FT) (pf0:0 < cardU (FinU_cons FT a)), ~In a (listU FT) -> minU (FinU_cons FT a) pf0 = a. intros; unfold minU; simpl; unfold nthU, list_singularizeU; simpl. gen0. unfold cardU, cardl; simpl. rewrite list_singularize_nin; auto. Qed. Lemma maxU_in_cons : forall {T} (FT:FinT T) (a:finU _ FT) (pf0:0 < cardU (FinU_cons FT a)) (pfi:In a (listU FT)), let pf0' := lt_eq_trans _ _ _ pf0 (in_fin_eqU _ _ pfi) in maxU (FinU_cons FT a) pf0 = maxU FT pf0'. unfold maxU; intros; simpl. gterml. redterm0 x. pose proof (nthU_in_cons FT a _ c pfi) as h1. simpl in h1. fold c. rewrite h1. apply f_equal_dep. rewrite (cardU_in_cons _ _ pfi) at 1; auto. Qed. Lemma list_singularizeT_nin_cons : forall {T} (FT:FinT T) (a:finU _ FT) (pfin: ~In a (listU FT)), map_proj1_sig (list_singularizeT (FinU_cons FT a)) = a :: list_singularizeU FT. intros T FT a h1; unfold map_proj1_sig; rewrite map_eq_iff. ex_goal. rewrite length_list_singularizeT; simpl. rewrite list_singularize_nin; simpl; auto. exists hex. intros n h2. destruct n as [|n]. simpl. pose proof (minT_nin_cons FT a h2 h1) as h3. unfold minT, nthT in h3. rewrite h3 at 1. simpl; auto; unfold list_singularizeT, listT; simpl. simpl. generalize h2. unfold list_singularizeT, listT; simpl. rewrite list_singularize_nin. simpl. intro h3. erewrite nth_indep. rewrite <- nth_lt_compat. rewrite nth_lt_list_singularize_map_in_dep. simpl. rewrite nth_lt_compat. apply nth_indep. apply lt_S_n in h3. rewrite (length_list_singularize_map_in_dep_inj (finU_eq_dec _ FT)) in h3. assumption. unfold fun_from_cons, id_in_dep_sig', id_prop_dep_sig'; simpl. red. intros ? ? ? ? h4; apply exist_injective in h4; auto. rewrite (length_list_singularize_map_in_dep_inj (finU_eq_dec _ FT)). rewrite length_list_singularizeT_nin_cons in h2; auto. apply lt_S_n in h2. rewrite length_list_singularizeT in h2; auto. unfold fun_from_cons, id_in_dep_sig', id_prop_dep_sig'; simpl. red. intros ? ? ? ? h4. apply exist_injective in h4; auto. unfold id_in_dep_sig', id_prop_dep_sig'; simpl. intro h3. rewrite in_map_in_dep_iff in h3. destruct h3 as [x [h3 h4]]. unfold fun_from_cons in h4. apply exist_injective in h4. subst. contradiction. Grab Existential Variables. unfold fun_from_cons, id_in_dep_sig', id_prop_dep_sig'; simpl. red. intros ? ? ? ? h4. apply exist_injective in h4; auto. rewrite length_list_singularizeT_nin_cons in h2; auto. apply lt_S_n in h2. rewrite length_list_singularizeT in h2; auto. Qed. Lemma listT_cons : forall {T:Type} (FT:FinT T) (a:finU _ FT), listT (FinU_cons FT a) = hdT (FinU_cons FT a) (lt_cardT_cons _ _) :: (map_sig_incl (incl_cons _ _)). intros T FT a. rewrite hdT_cons. rewrite listT_eq_iff. ex_goal. unfold listT. rewrite length_map_in_dep. f_equal. pose proof (length_map_sig_incl (incl_cons (listU FT) a)) as h1. simpl. rewrite h1 at 1; auto. exists hex. intros i hi hi'. unfold nthT'. simpl. destruct i as [|i]; auto. gterml; gtermr. redterm1 x. redterm1 y. fold c c0. assert (hc:c = c0). unfold c, c0. rewrite map_in_dep_eq_iff. ex_goal. pose proof (length_map_sig_incl (incl_cons (listU FT) a)) as h1. rewrite h1 at 1; auto. exists hex0. unfold fun_from_cons, id_in_dep_sig', id_prop_dep_sig', map_sig_incl; intros. rewrite nth_lt_map_in_dep'. repeat apply f_equal_dep; auto. rewrite hc. apply nth_indep. unfold c0. pose proof (length_map_sig_incl (incl_cons (listU FT) a)) as h1. rewrite h1 at 1; auto. simpl in hi. pose proof (lt_S_n _ _ hi) as h2. rewrite length_map_in_dep in h2; auto. Qed. Definition fin_transferT {T U} {FT:FinT T} {FU:FinT U} (pfe:fin_eqT FT FU) (x:T) : U := let pflt := lt_eq_trans _ _ _ (lt_lindT FT x) pfe in nthT FU (lindT FT x) pflt. Definition fin_transferU {T U} {FT:FinT T} {FU:FinT U} (pfe:fin_eqU FT FU) (y:finU _ FT) (pfi:In y (listU FT)) : finU _ FU := let pflt := lt_eq_trans _ _ _ (lt_lindU FT y pfi) pfe in nthU FU (lindU FT y pfi) pflt. Lemma fin_transferT_undoes : forall {T U} {FT:FinT T} {FU:FinT U} (pfe:fin_eqT FT FU) (pfe':fin_eqT FU FT) (x:T), fin_transferT pfe' (fin_transferT pfe x) = x. unfold fin_transferT, nthT, lindT; intros. rewrite nth_lt_eq_iff. ex_goal. apply in_list_singularizeT. exists hex. exists 0. split. apply gt_lt. rewrite <- (count_occ_In (finT_eq_dec FT)); auto. rewrite lind_nth_lt_no_dup. unfold lind. f_equal. apply proof_irrelevance. apply no_dup_list_singularize. Qed. Lemma fin_transferU_undoes : forall {T U} {FT:FinT T} {FU:FinT U} (pfe: fin_eqU FT FU) (pfe':fin_eqU FU FT) (x:finU _ FT) (pfx:In x (listU FT)) (pfx':In (fin_transferU pfe x pfx) (listU FU)), fin_transferU pfe' (fin_transferU pfe x pfx) pfx' = x. unfold fin_transferU, nthU, lindU; intros. rewrite nth_lt_eq_iff. ex_goal. apply in_listU_singularize; auto. exists hex. exists 0. split. apply gt_lt. rewrite <- (count_occ_In (finU_eq_dec _ FT)); auto. rewrite lind_nth_lt_no_dup. unfold lind. f_equal. apply proof_irrelevance. apply no_dup_list_singularize. Qed. Lemma fin_transfer_nthT : forall {T U} (FT:FinT T) (FU:FinT U) (pfe:fin_eqT FT FU) i (pfi:i < cardT FT), let pfi' := lt_eq_trans _ _ _ pfi pfe in fin_transferT pfe (nthT FT i pfi) = nthT FU i pfi'. unfold fin_transferT, nthT, lindT, list_singularizeT; intros. rewrite (nth_lt_eq_iff (finT_eq_dec FU)). ex_goal. rewrite <- (list_singularize_in_iff (finT_eq_dec FU)). apply in_listT. exists hex. exists 0. split. apply gt_lt. rewrite <- count_occ_In. rewrite <- list_singularize_in_iff. apply in_listT. rewrite lind_nth_lt_no_dup. gterml. redterm1 x. redterm0 c. pose proof (lind_nth_lt_no_dup (finT_eq_dec FU) _ _ c0 hex (no_dup_list_singularize _ _)) as h1. rewrite <- h1. unfold lind. reflexivity. apply no_dup_list_singularize. Qed. Lemma fin_transfer_nthU : forall {T U} (FT:FinT T) (FU:FinT U) (pfe:fin_eqU FT FU) i (pfi:i < cardU FT), let pfi' := lt_eq_trans _ _ _ pfi pfe in fin_transferU pfe (nthU FT i pfi) (in_nthU _ _ _) = nthU FU i pfi'. intros T U FT FU he i hi; simpl. unfold fin_transferU, nthU, lindU, list_singularizeU. rewrite (nth_lt_eq_iff (finU_eq_dec _ FU)). ex_goal. apply in_nth_lt. exists hex. exists 0. split. apply gt_lt. rewrite <- count_occ_In. apply in_nth_lt. rewrite lind_nth_lt_no_dup. gterml. redterm1 x. redterm0 c. pose proof (lind_nth_lt_no_dup (finU_eq_dec _ FU) _ _ c0 hex (no_dup_list_singularize _ _)) as h1. rewrite <- h1. unfold lind. reflexivity. apply no_dup_list_singularize. Qed. Lemma fin_transfer_minT : forall {T U} (FT:FinT T) (FU:FinT U) (pfe:fin_eqT FT FU) (pf0:0 < cardT FT), let pf0' := lt_eq_trans _ _ _ pf0 pfe in fin_transferT pfe (minT FT pf0) = minT FU pf0'. unfold minT; intros; apply fin_transfer_nthT. Qed. Lemma fin_transfer_maxT : forall {T U} (FT:FinT T) (FU:FinT U) (pfe:fin_eqT FT FU) (pf0:0 < cardT FT), let pf0' := lt_eq_trans _ _ _ pf0 pfe in fin_transferT pfe (maxT FT pf0) = maxT FU pf0'. unfold maxT; intros; rewrite fin_transfer_nthT; apply f_equal_dep; f_equal; apply pfe. Qed. Lemma fin_transfer_minU : forall {T U} (FT:FinT T) (FU:FinT U) (pfe:fin_eqU FT FU) (pf0:0 < cardU FT), let pf0' := lt_eq_trans _ _ _ pf0 pfe in fin_transferU pfe (minU FT pf0) (in_minU _ _) = minU FU pf0'. unfold minU; intros; erewrite f_equal_dep; auto. rewrite fin_transfer_nthU; auto. Qed. Lemma fin_transfer_maxU : forall {T U} (FT:FinT T) (FU:FinT U) (pfe:fin_eqU FT FU) (pf0:0 < cardU FT), let pf0' := lt_eq_trans _ _ _ pf0 pfe in fin_transferU pfe (maxU FT pf0) (in_maxU _ _) = maxU FU pf0'. unfold maxU; intros; erewrite f_equal_dep; auto. rewrite fin_transfer_nthU; auto. apply f_equal_dep; f_equal; apply pfe. Qed. Definition fin_transfer_lT {T U} {FT:FinT T} {FU:FinT U} (pfe:fin_eqT FT FU) (l:list T) : list U := map_seg (fun i (pfi:i < length l) => fin_transferT pfe (nth_lt l i pfi)). Definition fin_transfer_lU {T U} {FT:FinT T} {FU:FinT U} (pfe:fin_eqU FT FU) (l:list (finU _ FT)) (pfinc:incl l (listU FT)) : list (finU _ FU) := map_seg (fun i (pfi:i < length l) => fin_transferU pfe (nth_lt l i pfi) (pfinc _ (in_nth_lt _ _ _))). Lemma in_fin_transfer_lT_iff : forall {T U} {FT:FinT T} {FU:FinT U} (pfe:fin_eqT FT FU) (l:list T) (y:U), In y (fin_transfer_lT pfe l) <-> In (fin_transferT (eq_sym pfe) y) l. unfold fin_transfer_lT, fin_transferT; intros T U FT FU he l y; split; intro h1. rewrite in_map_seg_iff in h1. destruct h1 as [m [h2 ?]]; subst. gterm1. redterm0 c. fold c0. generalize c0. rewrite lind_nthT. intro; rewrite nth_lindT. apply in_nth_lt. rewrite in_map_seg_iff. exists _, (lt_lind (finT_eq_dec FT) _ _ h1). gterm0. unfold nthT, lindT, list_singularizeT. rewrite nth_lt_eq_iff. ex_goal. rewrite <- list_singularize_in_iff. apply in_listT. exists hex, 0. split. apply gt_lt. rewrite <- (count_occ_In _ _). rewrite <- list_singularize_in_iff. apply in_listT. symmetry. gen0. rewrite lind_compat. intro. rewrite lind_nth_lt_no_dup. unfold lind. apply lind_occ_functional; auto. apply no_dup_list_singularize. Qed. Lemma in_fin_transfer_lU_iff : forall {T U} {FT:FinT T} {FU:FinT U} (pfe:fin_eqU FT FU) (l:list (finU _ FT)) (pfinc:incl l (listU FT)), forall (y:finU _ FU), In y (fin_transfer_lU pfe l pfinc) <-> exists pfy:In y (listU FU), In (fin_transferU (eq_sym pfe) y pfy) l. unfold fin_transfer_lU, fin_transferU; intros T U FT FU he l hinc y; split; intro h1. rewrite in_map_seg_iff in h1. destruct h1 as [i [hi ?]]; subst. exists (in_nthU _ _ _). match goal with |- In ?y' _ => pose y' as y end. redterm0 y. fold c. generalize c. rewrite lind_nthU. intro h1. rewrite nth_lindU. apply in_nth_lt. destruct h1 as [h1 h2]. rewrite in_map_seg_iff. exists _, (lt_lind (finU_eq_dec _ FT) _ _ h2). unfold nthU, lindU. rewrite nth_lt_eq_iff. exists (in_listU_singularize _ _ h1). exists 0. split. apply gt_lt. rewrite <- (count_occ_In (finU_eq_dec _ FU) _ _) at 1. apply in_listU_singularize; auto. symmetry. gen0. rewrite lind_compat. intro. rewrite lind_nth_lt_no_dup. unfold lind. f_equal; apply proof_irrelevance. apply no_dup_list_singularize. Qed. Definition P_FinTU {T U} {FT:FinT T} {FU:FinT U} (pfe:fin_eqT FT FU) (P:T->Prop) : U->Prop := fun y => P (fin_transferT (sym_fin_eqT FT FU pfe) y). Definition P_FinTU' {T U} {FT:FinT T} {FU:FinT U} (pfe:fin_eqU FT FU) (P:T->Prop) : U->Prop := let pfe' := fin_eqUT _ _ pfe in P_FinTU pfe' P. Definition P_FinUT {T U} {FT:FinT T} {FU:FinT U} (pfe:fin_eqT FT FU) (P:U->Prop) : T->Prop := fun y => P (fin_transferT pfe y). Definition P_FinUT' {T U} {FT:FinT T} {FU:FinT U} (pfe:fin_eqU FT FU) (P:U->Prop) : T->Prop := let pfe' := fin_eqUT _ _ pfe in P_FinUT pfe' P. Definition fun_in_dep_to_funT {T U} (FT:FinT T) (f:fun_in_dep (listU FT) U) : T->U := fun x => f (fin_fT' FT x) (in_fin_fT' _ _). Definition ST {T} (FT:FinT T) (x:T) : T := right_shift_functor (fun i (pfi:i < cardT FT) => nthT FT i pfi) (Olt _ _ (lt_lindT FT x)) (lindT FT x) (lt_lindT FT x). Definition predT {T} (FT:FinT T) (x:T) : T := left_shift_functor (fun i (pfi:i < cardT FT) => nthT FT i pfi) (Olt _ _ (lt_lindT FT x)) (lindT FT x) (lt_lindT FT x). Fixpoint conj_in_dep {T} {l:list T} : forall (Q:prop_in_dep l), Prop := match l with | nil => fun _ => True | a::l' => fun Q => let Q' := fun_from_cons Q in Q a (in_eq _ _) /\ conj_in_dep Q' end. Lemma conj_in_dep1 : forall {T U} {l:list T} (f:fun_in_dep l U) (P:prop_in_dep (map_in_dep f)), conj_in_dep P -> forall (x:T) (pf:In x l), P (f x pf) (in_map_in_dep _ _ _). intros T U l. induction l as [|a l ih]; simpl. intros; contradiction. intros f P h1 b h3. destruct h1 as [h1 h2]. destruct h3 as [|h3]; subst. erewrite f_equal_dep. apply h1. f_equal; apply proof_irrelevance. specialize (ih (fun_from_cons f) (fun_from_cons P) h2 b h3). unfold fun_from_cons in ih. erewrite f_equal_dep. apply ih. f_equal. apply proof_irrelevance. Qed. Lemma conj_in_dep_functional : forall {T} (P:T->Prop) (l l':list T), l = l' -> conj_in_dep (fun (x:T) (pf:In x l) => P x) -> conj_in_dep (fun (x:T) (pf:In x l') => P x). intros; subst; auto. Qed. Lemma conj_in_dep_functional' : forall {T} (P Q R:T->Prop) (l :list {y:T | P y}) (l':list {y:T | Q y}), proj1_sig_leq l l' -> P = Q -> conj_in_dep (fun (x:{y:T | P y}) (pf:In x l) => R (proj1_sig x)) -> conj_in_dep (fun (x:{y:T | Q y}) (pf:In x l') => R (proj1_sig x)). intros T P Q R l l' h1 h2 h3; subst. revert h3 h1; revert l'. induction l as [|a l ih]; simpl. intros ? ? h1. apply proj1_sig_leq_nil1 in h1; subst; simpl; auto. intro l'. destruct l' as [|a' l']. intros ? h2. apply proj1_sig_leq_nil2 in h2; discriminate. intro h1. destruct h1 as [h1 h2]. intro h3. simpl. apply proj1_sig_leq_cons' in h3. destruct h3 as [h3 h4]. split. rewrite <- h3; auto. auto. Qed. Lemma conj_in_dep_ran : forall {U V} {l:list V} {a} (_:type_eq_dec V) (f:fun_in_dep (a::l) U) (P:U->Prop) (pfinj:inj_dep f) (pfnin : ~In a l), P (f a (in_eq _ _)) -> conj_in_dep (fun (x : {y : U | y <> f a (in_eq _ _)}) (_ : In x (map_in_dep (fun_from_cons_ran f pfinj pfnin))) => P (proj1_sig x)) -> conj_in_dep (fun (x : U) (_ : In x (map_in_dep (fun (x0 : V) (pf : In x0 l) => f x0 (in_cons a x0 l pf)))) => P x). intros U V l; induction l as [|a l ih]; simpl; auto. intros b hdv f P hinj' hnin' hpfb hpfa. unfold fun_from_cons. destruct hpfa as [hpfa h1]. split; auto. destruct (in_dec hdv a l) as [h5|h5]. pose proof (inj_from_cons2 _ hinj') as h6. specialize (ih b hdv (fun_from_cons2 f) P (inj_from_cons2 _ hinj')). hfirst ih. intro hb; pose proof (hnin' (or_intror hb)); auto. specialize (ih hf). hfirst ih. unfold fun_from_cons2. erewrite f_equal_dep; auto. apply hpfb. specialize (ih hf0). unfold fun_from_cons2 in ih. gterm0. let P := type of c in match P with forall x, In x (map_in_dep ?g') -> _ => pose g' as g end. fold g. term0 ih. let P := type of c0 in match P with forall x, In x (map_in_dep ?h') -> _ => pose h' as h end. fold h in ih. assert (h7:g = h). unfold g, h. apply functional_extensionality_dep; intro d; apply functional_extensionality; intro hd. f_equal. apply proof_irrelevance. rewrite h7; apply ih. unfold fun_from_cons, fun_from_cons_ran in h1; simpl in h1. let P := type of h1 in match P with conj_in_dep (fun x (_:In x ?lm') => _) => pose lm' as lm end. match goal with |- conj_in_dep (fun x (_:In x ?lm'') => _) => pose lm'' as lm' end. assert (h8:proj1_sig_leq lm lm'). red. unfold lm, lm'. unfold map_proj1_sig; do 2 rewrite map_map_in_dep; simpl. f_equal. apply functional_extensionality_dep. intro. apply functional_extensionality. intro; f_equal. apply proof_irrelevance. eapply conj_in_dep_functional'. apply h8. apply functional_extensionality_dep. intro. apply prop_ext. erewrite f_equal_dep. apply iff_refl. reflexivity. assumption. specialize (ih b hdv (fun_from_cons2 f) P (inj_from_cons2 _ hinj')). hfirst ih. intro h6. pose proof (hnin' (or_intror _ h6)); auto. specialize (ih hf). hfirst ih. unfold fun_from_cons2. erewrite f_equal_dep. apply hpfb. reflexivity. specialize (ih hf0). hfirst ih. unfold fun_from_cons in h1. let P := type of h1 in match P with conj_in_dep (fun x (_:In x ?lm') => _) => pose lm' as lm end. match goal with |- conj_in_dep (fun x (_:In x ?lm'') => _) => pose lm'' as lm' end. assert (h8:proj1_sig_leq lm lm'). red. unfold lm, lm'. unfold map_proj1_sig; do 2 rewrite map_map_in_dep; simpl. f_equal. apply functional_extensionality_dep. intro. unfold fun_from_cons2. apply functional_extensionality. intro; f_equal; apply proof_irrelevance. eapply conj_in_dep_functional'. apply h8. apply functional_extensionality_dep. intro. apply prop_ext. erewrite f_equal_dep. apply iff_refl. reflexivity. assumption. specialize (ih hf1). let P := type of ih in match P with conj_in_dep (fun x (_:In x ?lm') => _) => pose lm' as lm end. match goal with |- conj_in_dep (fun x (_:In x ?lm'') => _) => pose lm'' as lm' end. assert (h2:lm = lm'). unfold lm, lm'. f_equal. apply functional_extensionality_dep. intro. unfold fun_from_cons2. apply functional_extensionality. intro; f_equal; apply proof_irrelevance. eapply conj_in_dep_functional. apply h2. assumption. Qed. Definition finite_conj_prop {T} (FT:FinT T) (P:T->Prop) : Prop := conj_in_dep (fun x (pf:In x (listT FT)) => P x). Lemma finite_conj_prop_iff : forall {T} (FT:FinT T) (P:T->Prop), finite_conj_prop FT P <-> (forall x:T, P x). unfold finite_conj_prop, listT; intros T FT. destruct FT. revert fin_inv0; revert fin_fU0; revert T. induction fin_lU0; simpl. intros; split; intros. inversion fin_inv0. contradict (pfp x). auto. intros fin_U f hi P; split; intro h1. destruct h1 as [h1 h2]. intro x. destruct (finU_eq_dec0 (projT1 (inverse_dep f hi) x) a) as [h3|h3]. pose proof (inverse_dep_compat f hi) as h4; simpl in h4. destruct h4 as [h4 h5]. specialize (h5 x). revert h5. pose proof (f_equal_dep f (projT1 (inverse_dep f hi) x) a (projT2 (inverse_dep f hi) x) (in_eq _ _) h3) as h5. simpl in h5. rewrite h5. intro h6. rewrite <- h6; auto. pose proof (projT2 (inverse_dep f hi) x) as h4. simpl in h4. destruct h4 as [h4|h4]. rewrite h4 in h3 at 1. contradict h3; auto. pose proof (conj_in_dep1 _ _ h2 (projT1 (inverse_dep f hi) x) h4) as h5. unfold fun_from_cons in h5. pose proof (inverse_dep_compat f hi) as h6; simpl in h6. destruct h6 as [h6 h7]. specialize (h7 x). erewrite f_equal_dep in h7. rewrite h7 in h5 at 1; auto. reflexivity. simpl in IHfin_lU0. split; auto. destruct (in_dec finU_eq_dec0 a fin_lU0) as [h2|h2]. pose proof (inv_from_cons_in _ h2 hi) as h3. gterm0. specialize (IHfin_lU0 _ (fun_from_cons f) h3 P). pose proof h1 as h1'. rewrite <- IHfin_lU0 in h1. erewrite f_equal. apply h1. apply functional_extensionality_dep; intro x; apply functional_extensionality; intro h4; unfold fun_from_cons; auto. pose proof (inv_from_cons_ran _ h2 hi) as h3; simpl in h3. specialize (IHfin_lU0 _ _ h3 (fun x => P (proj1_sig x))). simpl in IHfin_lU0. hr IHfin_lU0; auto. rewrite <- IHfin_lU0 in hr. unfold fun_from_cons. eapply conj_in_dep_ran. apply finU_eq_dec0. apply h1. apply hr. Qed. Lemma finite_conj_prop_cons : forall {T} (FT:FinT T) (a : finU T FT) (pf0 : 0 < cardT (FinU_cons FT a)) (R:{x : finU T FT | In x (a :: fin_lU T FT)} -> Prop), finite_conj_prop (FinU_cons FT a) R -> R (hdT _ pf0). intros T FT a h0 R h1. red in h1. simpl in h1. destruct h1 as [h1]. unfold id_in_dep_sig', id_prop_dep_sig' in h1. unfold hdT, fin_fU; simpl; unfold id_in_dep_sig', id_prop_dep_sig'; auto. erewrite f_equal. apply h1. apply proj1_sig_injective; auto. Qed. Lemma finite_conj_prop_iff' : forall {T} (FT:FinT T) (R:T->Prop) (pf0:0 < cardT FT), finite_conj_prop FT R <-> (R (minT FT pf0) /\ (forall (i : nat) (pfi : i < pred (cardT FT)), R (nthT FT i (lt_trans i (pred (cardT FT)) (cardT FT) pfi (lt_pred_n 0 (cardT FT) pf0))) -> R (nthT FT (S i) (lt_eq_trans (S i) (S (pred (cardT FT))) (cardT FT) (lt_n_S i (pred (cardT FT)) pfi) (eq_sym (S_pred (cardT FT) 0 pf0)))))). intros T FT R h0; rewrite finite_conj_prop_iff. split; intro h1. intros; auto. revert h1; revert h0 R; revert FT; revert T. apply (FinT_induction (fun T (FT:FinT T) => forall (h0 : 0 < cardT FT) (R : T -> Prop), R (minT FT h0) /\ (forall (i : nat) (pfi : i < pred (cardT FT)), R (nthT FT i (lt_trans i (pred (cardT FT)) (cardT FT) pfi (lt_pred_n 0 (cardT FT) h0))) -> R (nthT FT (S i) (lt_eq_trans (S i) (S (pred (cardT FT))) (cardT FT) (lt_n_S i (pred (cardT FT)) pfi) (eq_sym (S_pred (cardT FT) 0 h0))))) -> forall x : T, R x)). intros; contradiction. intros T FT h1 a h3 R h5 x. destruct h5 as [h5 h6]. destruct x as [x hx]. destruct hx as [|hx]; subst. destruct (in_dec (finU_eq_dec _ FT) x (listU FT)) as [h7|h7]. hfirst h1. apply lt_cardT'. apply (fin_fU _ FT x h7). specialize (h1 hf). pose (compose_sig_prop (fin_fT FT) (sigl_from_cons R)) as R'. specialize (h1 R'). hfirst h1. split. repeat red. erewrite f_equal. apply h5. apply proj1_sig_injective; simpl. rewrite (minT_in_cons _ _ h3 h7) at 1; simpl. rewrite fin_fT_minT'. pose proof (fin_fT_compat FT) as h8. destruct h8 as [h8 h9]. specialize (h8 _ (in_minU FT (lt_eq_trans _ _ _ hf (card_eq FT)))). unfold fin_fT in h8. apply exist_injective in h8. destruct (inverse_dep (fin_fU T FT) (fin_inv T FT)) as [f h11]; simpl; simpl in h8. symmetry. erewrite minU_functional. rewrite <- h8 at 1. f_equal. rewrite fin_fU_minU; auto. apply f_equal; apply proof_irrelevance. intros i hi. specialize (h6 i). hfirst h6. rewrite cardT_in_cons; auto. specialize (h6 hf0). intro h8. hfirst h6. unfold R', compose_sig_prop, compose, sigl_from_cons in h8. pose proof (nthT_in_cons FT x i) as h9. simpl in h9. gterm0. redterm0 c. fold c0. specialize (h9 c0 h7). rewrite h9 at 1; auto. erewrite f_equal. apply h8. apply proj1_sig_injective; simpl. gterml. redterm0 x0. redterm0 c1. pose proof (fin_fT_compat FT) as h10. destruct h10 as [h10 h11]. rewrite fin_fT_nthT'. gtermr. pose proof (fin_fT_nthT' FT i c2) as h12. fold c2. rewrite <- h12. unfold fin_fT', fin_fT; simpl. repeat f_equal; apply proof_irrelevance. specialize (h6 hf1). unfold R', compose_sig_prop, compose, sigl_from_cons. erewrite f_equal. apply h6. apply proj1_sig_injective; simpl. pose proof (fin_fT_nthT FT (S i)) as h12. gterml. redterm0 x0. redterm0 c. specialize (h12 c0). unfold fin_fT in h12. apply exist_injective in h12. fold c0. rewrite h12. gtermr. redterm0 y. redterm0 c1. pose proof (nthT_in_cons FT x (S i) c2 h7) as h13. simpl in h13. fold c2. rewrite h13. simpl. rewrite fin_fT_nthT'. f_equal; apply proof_irrelevance. specialize (h1 hf0). specialize (h1 (fin_fU _ FT x h7)). unfold R', compose_sig_prop, compose, sigl_from_cons in h1. erewrite f_equal. apply h1. apply proj1_sig_injective; simpl. pose proof (fin_fT_compat FT) as h8. destruct h8 as [h8 h9]. specialize (h8 _ h7). unfold fin_fT in h8. apply exist_injective in h8. auto. rewrite (minT_nin_cons _ _ h3 h7) in h5. simpl in h5. unfold id_in_dep_sig', id_prop_dep_sig' in h5. erewrite f_equal. apply h5. f_equal; apply proof_irrelevance. destruct (finU_eq_dec _ FT x a) as [|he]; subst. hfirst h1. apply lt_cardT'. apply (fin_fU _ FT _ hx). specialize (h1 hf). pose (compose_sig_prop (fin_fT FT) (sigl_from_cons R)) as R'. specialize (h1 R'). hfirst h1. unfold R', compose_sig_prop, sigl_from_cons, compose; simpl. split. rewrite (minT_in_cons _ _ _ hx) in h5. erewrite f_equal. apply h5. apply proj1_sig_injective; simpl. rewrite fin_fT_minT'. pose proof (fin_fT_compat FT) as h7. destruct h7 as [h7 h8]. specialize (h7 (minU FT (lt_eq_trans _ _ _ hf (card_eq FT))) (in_minU _ _)). unfold fin_fT in h7. apply exist_injective in h7. unfold minT. erewrite minU_functional. rewrite <- h7 at 1. repeat f_equal. symmetry. erewrite f_equal_dep. rewrite fin_fU_minU. unfold minT. f_equal; apply proof_irrelevance. reflexivity. intros i hi h7. specialize (h6 i). hfirst h6. rewrite cardT_in_cons; auto. specialize (h6 hf0). hfirst h6. erewrite f_equal. apply h7. pose proof (nthT_in_cons FT a i) as h8. simpl in h8. gterml. redterm0 x. specialize (h8 c hx). fold c. rewrite h8 at 1. apply proj1_sig_injective; simpl. unfold fin_fT', fin_fT; simpl. repeat f_equal; apply proof_irrelevance. specialize (h6 hf1). erewrite f_equal. apply h6. apply proj1_sig_injective; simpl; repeat f_equal. symmetry. pose proof (nthT_in_cons FT a (S i)) as h8. simpl in h8. gterml. redterm0 x. redterm0 c. specialize (h8 c0 hx). fold c0. rewrite h8; simpl. unfold fin_fT', fin_fT; simpl. repeat f_equal; apply proof_irrelevance. specialize (h1 hf0). gterm0. specialize (h1 (fin_fU _ FT a hx)). unfold R', compose_sig_prop, compose, sigl_from_cons in h1. erewrite f_equal. apply h1. apply proj1_sig_injective; simpl. pose proof (fin_fT_compat FT) as h7. destruct h7 as [h7 h8]. specialize (h7 a hx). unfold fin_fT in h7. apply exist_injective in h7. rewrite <- h7 at 1; auto. hfirst h1. apply lt_cardT'. apply (fin_fU _ FT _ hx). destruct (in_dec (finU_eq_dec _ FT) a (listU FT)) as [h8 | h8]. specialize (h1 hf). pose (compose_sig_prop (fin_fT FT) (sigl_from_cons R)) as R'. specialize (h1 R'). hfirst h1. split. unfold R', compose_sig_prop, compose, sigl_from_cons. erewrite f_equal. apply h5. apply proj1_sig_injective; simpl. pose proof (inverse_dep_compat (fin_fU _ FT) (fin_inv _ FT)) as h9. simpl in h9. destruct h9 as [h9 h10]. specialize (h9 (minU FT (lt_eq_trans _ _ _ hf (card_eq _))) (in_minU _ _)). rewrite fin_fU_minU in h9. erewrite minT_functional. rewrite h9 at 1. pose proof (minT_in_cons FT a h3 h8) as h11. simpl in h11. rewrite h11; simpl. rewrite fin_fT_minT'. f_equal. apply proof_irrelevance. intros i hi. specialize (h6 i). hfirst h6. rewrite cardT_in_cons; auto. specialize (h6 hf0). intro h9. hfirst h6. unfold R', compose_sig_prop, compose, sigl_from_cons in h9. erewrite f_equal. apply h9. apply proj1_sig_injective; simpl. gterml. redterm0 x0. redterm0 c. pose proof (nthT_in_cons FT a i c0 h8) as h10. simpl in h10. fold c0. rewrite h10. simpl. unfold fin_fT', fin_fT; simpl. repeat f_equal; apply proof_irrelevance. specialize (h6 hf1). unfold R', compose_sig_prop, sigl_from_cons, compose. erewrite f_equal. apply h6. symmetry. gterml. redterm0 x0. pose proof (nthT_in_cons FT a (S i) c h8) as h10; simpl in h10. fold c; simpl. rewrite h10 at 1. apply proj1_sig_injective; simpl. unfold fin_fT', fin_fT; simpl. repeat f_equal; apply proof_irrelevance. specialize (h1 hf0). specialize (h1 (fin_fU _ FT x hx)). unfold R', compose_sig_prop, sigl_from_cons, compose in h1. erewrite f_equal. apply h1. apply proj1_sig_injective; simpl. pose proof (fin_fT_compat FT) as h11. destruct h11 as [h11 h12]. specialize (h11 x hx). unfold fin_fT in h11. apply exist_injective in h11; auto. specialize (h1 hf). pose (compose_sig_prop (fin_fT FT) (sigl_from_cons R)) as R'. specialize (h1 R'). hfirst h1. unfold R', compose_sig_prop, sigl_from_cons, compose. split. specialize (h6 0). hfirst h6. rewrite cardT_nin_cons; simpl; auto. specialize (h6 hf0). pose proof (minT_nin_cons FT a h3 h8) as h9. gterm0. redterm0 c. fold c0. generalize c0. rewrite fin_fT_minT; simpl. intro h10. hfirst h6. gterm0. redterm0 c1. pose proof (minT_nin_cons _ _ h3 h8) as h11. unfold minT in h11. erewrite nthT_functional. rewrite h11 at 1. simpl. erewrite f_equal. apply h5. rewrite minT_nin_cons; simpl; auto. reflexivity. specialize (h6 hf1). erewrite f_equal. apply h6. apply proj1_sig_injective; simpl. symmetry. gterml. redterm0 x0. redterm0 c1. pose proof (nthT_nin_cons FT a h8 h3 1 c2) as h11. hfirst h11; auto with arith. specialize (h11 hf2). simpl in h11. fold c2. rewrite h11 at 1. unfold minU. f_equal. apply proof_irrelevance. intros i hi h9. specialize (h6 (S i)). hfirst h6. rewrite cardT_nin_cons; simpl. omega. assumption. specialize (h6 hf0). hfirst h6. erewrite f_equal. apply h9. apply proj1_sig_injective; simpl. pose proof (nthT_nin_cons FT a h8 h3 (S i)) as h10. simpl in h10. gterml. redterm0 x0. redterm0 c. pose proof (cardT_nin_cons FT a h8) as h11. specialize (h10 c0 (lt_0_Sn _)). fold c0. rewrite h10 at 1. gtermr. redterm0 y. redterm0 c1. pose proof (fin_fT_nthT FT i c2) as h12. unfold fin_fT in h12. apply exist_injective in h12. fold c2. rewrite h12 at 1. f_equal; apply proof_irrelevance. specialize (h6 hf1). erewrite f_equal. apply h6. apply proj1_sig_injective; simpl. gterml. redterm0 x0. redterm0 c. pose proof (fin_fT_nthT FT (S i) c0) as h12. unfold fin_fT in h12. apply exist_injective in h12. fold c0. rewrite h12 at 1. gtermr. redterm0 y. redterm0 c1. pose proof (nthT_nin_cons FT a h8 h3 (S (S i)) c2 (lt_0_Sn _)) as h10. simpl in h10. fold c2. rewrite h10 at 1. apply f_equal_dep; auto. specialize (h1 hf0). specialize (h1 (fin_fU _ FT x hx)). unfold R', compose_sig_prop, compose, sigl_from_cons in h1. erewrite f_equal. apply h1. apply proj1_sig_injective; simpl. pose proof (fin_fT_compat FT) as h9. destruct h9 as [h9 h10]. specialize (h9 x hx). unfold fin_fT in h9. apply exist_injective in h9; auto. intros T U FT FU h1 h2 h3 R h4 x. destruct h4 as [h4 h5]. hfirst h2. rewrite card_eq. unfold cardU, listU; unfold cardT, listT, cardl in h3. pose proof (length_list_singularize_map_in_dep_inj (finU_eq_dec _ FU) (finT_eq_dec FU) (fin_fU _ FU)) as h6. hfirst h6. apply inv_impl_inj_dep. apply fin_inv. specialize (h6 hf). pose proof h3 as h3'. rewrite h6 in h3'. red in h1. unfold cardU in h1. unfold listU in h1. rewrite h1; auto. specialize (h2 hf (P_FinUT' h1 R)). hfirst h2. split. unfold P_FinUT', P_FinUT. rewrite fin_transfer_minT. erewrite f_equal. apply h4. f_equal; apply proof_irrelevance. intros i hi h6. specialize (h5 i). hfirst h5. erewrite f_equal. apply hi. f_equal; auto. do 2 rewrite card_eq; auto. specialize (h5 hf0). unfold P_FinUT', P_FinUT in h6. rewrite fin_transfer_nthT in h6. hfirst h5. erewrite f_equal. apply h6. apply f_equal_dep; auto. specialize (h5 hf1). unfold P_FinUT', P_FinUT. rewrite fin_transfer_nthT. erewrite f_equal. apply h5. apply f_equal_dep. reflexivity. specialize (h2 hf0). pose proof h1 as h6. apply fin_eqUT in h6. specialize (h2 (fin_transferT (eq_sym h6) x)). unfold P_FinUT', P_FinUT in h2. erewrite f_equal. apply h2. rewrite fin_transferT_undoes; auto. Qed. Lemma FinT_induction_weak : forall {T} (FT:FinT T) (pf0:0 < cardT FT) (P:T->Prop), P (minT FT pf0) -> (forall i (pfi:i < pred (cardT FT)), P (nthT FT i (lt_trans _ _ _ pfi (lt_pred_n _ _ pf0))) -> P (nthT FT (S i) (lt_eq_trans _ _ _ (lt_n_S _ _ pfi) (eq_sym (S_pred _ _ pf0))))) -> forall x:T, P x. intros T FT h0 P h1 h2. pose proof (finite_conj_prop_iff' FT P h0) as h3. hr h3. split; auto. rewrite <- h3 in hr. rewrite finite_conj_prop_iff in hr; auto. Qed. Lemma finite_conj_prop_in_cons_iff : forall {T} (FT:FinT T) (a:finU _ FT) (P:{x:finU _ FT | In x (a::listU FT)}->Prop) (pf:In a (listU FT)), finite_conj_prop (FinU_cons FT a) P <-> finite_conj_prop FT (fun x => P (exist (fun c => In c (a::listU FT)) _ (in_cons _ _ _ (in_listU FT x)))). intros; do 2 rewrite finite_conj_prop_iff; split; intros h1 x. specialize (h1 (exist (fun c => In c (a::listU FT)) (proj1_sig (fin_fT FT x)) (in_cons _ _ _ (proj2_sig (fin_fT FT x))))). erewrite f_equal. apply h1. f_equal; apply proof_irrelevance. destruct x as [x hx]. destruct hx as [|hx]; subst. specialize (h1 (fin_fU _ FT x pf)). erewrite f_equal. apply h1. apply f_equal_dep. pose proof (fin_fT_compat FT) as h2. destruct h2 as [h2 h3]. unfold fin_fT'. specialize (h2 x pf). rewrite h2; auto. specialize (h1 (fin_fU _ FT x hx)). erewrite f_equal. apply h1. apply proj1_sig_injective; simpl. pose proof (fin_fT_compat FT) as h2. destruct h2 as [h2 h3]. unfold fin_fT'. specialize (h2 x hx). rewrite h2; auto. Qed. Lemma finite_conj_prop_nin_cons_iff : forall {T} (FT:FinT T) (a:finU _ FT) (P:{x:finU _ FT | In x (a::listU FT)}->Prop) (pf: ~In a (listU FT)), finite_conj_prop (FinU_cons FT a) P <-> P (exist (fun c => In c (a::listU FT)) a (in_eq _ _)) /\ finite_conj_prop FT (fun x => P (exist (fun c => In c (a::listU FT)) _ (in_cons _ _ _ (in_listU FT x)))). intros; do 2 rewrite finite_conj_prop_iff; split; intro h1. split; auto. destruct h1 as [h1 h2]. intro x; destruct x as [x hx]. destruct hx as [|hx]; subst. erewrite f_equal. apply h1. f_equal; apply proof_irrelevance. specialize (h2 (fin_fU _ FT x hx)). erewrite f_equal. apply h2. apply f_equal_dep. pose proof (fin_fT_compat FT) as h3. unfold fin_fT'. destruct h3 as [h3 h4]. specialize (h3 x hx). rewrite h3; auto. Qed. Lemma finite_conj_prop_cons_iff' : forall {T} (FT:FinT T) (P:T->Prop) (pf0: 0 < cardU FT), let h := hdU FT pf0 in let h' := hdT FT (lt_eq_trans _ _ _ pf0 (eq_sym (card_eq _))) in finite_conj_prop FT P <-> P h' /\ (forall (pfi: In h (tl (listU FT))), finite_conj_prop (FinT_in_tl _ _ pfi) P) /\ (forall (pfn:~In h (tl (listU FT))), finite_conj_prop (FinT_nin_tl _ _ pfn) (P_to_dep P _)). intros T FT P h0 h h'; split; intro h1; try rewrite finite_conj_prop_iff in h1; try rewrite finite_conj_prop_iff; try esplit; auto. split; intros; try rewrite finite_conj_prop_iff; auto. intro x. destruct x as [x hx]. unfold P_to_dep; auto. destruct h1 as [h1 [h2 h3]]. intro x. destruct (finT_eq_dec FT x h') as [|h4]; subst; auto. unfold h' in h4. destruct (in_dec (finU_eq_dec _ FT) h (tl (listU FT))) as [h5 | h5]. specialize (h2 h5). rewrite finite_conj_prop_iff in h2; auto. specialize (h3 h5). rewrite finite_conj_prop_iff in h3; auto. specialize (h3 (exist (fun c => c <> hdT FT (lt_eq_trans 0 (cardU FT) (cardT FT) h0 (eq_sym (card_eq FT)))) _ h4)). unfold P_to_dep in h3; auto. Qed. Lemma finite_conj_prop_incl_iff : forall {T} (FT:FinT T) (P:T->Prop), finite_conj_prop FT P <-> (forall (l:list T), (forall x (pf:In x l), P x) -> forall a, ~In a l -> P a). pose proof (FinT_induction (fun T (FT:FinT T) => forall P : T -> Prop, finite_conj_prop FT P <-> (forall l : list T, (forall x : T, In x l -> P x) -> forall a : T, ~ In a l -> P a))) as h1; simpl in h1; apply h1; clear h1. intro; split; intro h1. intros; contradiction. red; red; simpl; auto . intros T FT h1 a P. pose proof (fin_fT_compat' FT) as h11; destruct h11 as [h11 h12]. destruct (in_dec (finU_eq_dec _ FT) a (listU FT)) as [h3 | h3]. pose proof (finite_conj_prop_in_cons_iff FT a P h3) as h4. rewrite h4; clear h4. rewrite h1; split; intros h4 l h5 x h6. pose proof (incl_from_cons _ _ _ (map_proj1_sig_incl (a::listU FT) l) h3) as h7. pose (map_lU FT (map_proj1_sig l) h7) as l'. specialize (h4 l'). hfirst h4. intros y hy. unfold l' in hy. rewrite in_map_lU_iff in hy. rewrite (in_map_proj1_sig_iff (P := fun c => In c (a::listU FT))) in hy. destruct hy as [hy h8]. specialize (h5 _ h8). erewrite f_equal. apply h5. f_equal. apply proof_irrelevance. specialize (h4 hf). destruct x as [x hx]. unfold l' in h4, hf; clear l'. pose proof (in_from_cons _ _ _ h3 hx) as h8. specialize (h4 (fin_fU _ FT x h8)). erewrite f_equal_dep. apply h4. intro h9. rewrite in_map_lU_iff in h9. rewrite (in_map_proj1_sig_iff (P := fun c => In c (a::listU FT))) in h9. destruct h9 as [h9 h10]. contradict h6. erewrite f_equal_dep. apply h10. rewrite h11; auto. rewrite h11; auto. specialize (h4 (map_lT_cons FT l a)). hfirst h4. intros y h7. rewrite (in_map_lT_cons_iff FT l a y) in h7. destruct h7 as [h7 h8]. specialize (h5 _ h8). erewrite f_equal. apply h5. apply proj1_sig_injective; simpl. rewrite h11; auto. specialize (h4 hf (exist (fun c => In c (a::listU FT)) (fin_fT' FT x) (in_cons _ _ _ (in_listU _ _)))). hfirst h4. rewrite (in_map_lT_cons_iff FT l a (exist (fun c => In c (a::listU FT)) _ (in_cons a (fin_fT' FT x) (listU FT) (in_listU FT x)))). simpl. contradict h6. destruct h6 as [h6 h7]. erewrite f_equal_dep in h7. rewrite h12 in h7 at 1. apply h7. reflexivity. specialize (h4 hf0). erewrite f_equal. apply h4. reflexivity. rewrite (finite_conj_prop_nin_cons_iff FT a _ h3). split; intro h4. destruct h4 as [h4 h5]. intros l h6 y h7. rewrite h1 in h5. pose (map_lU_sig FT (list_from_sig_cons_remove (finU_eq_dec _ FT) l)) as l'. specialize (h5 l'). hfirst h5. intros x hx. unfold l' in hx. rewrite in_map_lU_sig_iff, in_list_from_sig_cons_remove_iff in hx. destruct hx as [h8 h9]. specialize (h6 _ h8). erewrite f_equal. apply h6. apply proj1_sig_injective; simpl. unfold fin_fT', fin_fT; simpl; auto. specialize (h5 hf). destruct y as [y hy]. unfold l' in h5; clear hf l'. destruct hy as [|hy]; subst. erewrite f_equal. apply h4. f_equal; apply proof_irrelevance. specialize (h5 (fin_fU _ _ _ hy)). hfirst h5. intro h8. rewrite in_map_lU_sig_iff, in_list_from_sig_cons_remove_iff in h8. destruct h8 as [h8 h9]. contradict h7. erewrite f_equal_dep. apply h8. pose proof (fin_fT_compat FT) as h10; destruct h10 as [h10 h13]. rewrite h10; simpl; auto. specialize (h5 hf). erewrite f_equal_dep. apply h5. rewrite h11; auto. split. specialize (h4 nil). hfirst h4. intros; contradiction. specialize (h4 hf). apply h4. intro; contradiction. rewrite h1. intros l h5 x h6. specialize (h4 (map_lT_cons _ l a)). hfirst h4. intros y h7. rewrite (in_map_lT_cons_iff _ _ _ y) in h7. destruct h7 as [hi h7]. specialize (h5 _ h7). erewrite f_equal. apply h5. apply proj1_sig_injective; simpl. rewrite h11; auto. specialize (h4 hf). specialize (h4 (exist (fun c => In c (a::listU FT)) (fin_fT' FT x) (in_cons _ _ _ (in_listU _ _)))). hfirst h4. intro h7. let P := type of h7 in match P with In ?y' _ => pose y' as y end. fold y in h7. rewrite (in_map_lT_cons_iff _ _ _ y) in h7. destruct h7 as [h7 h8]. unfold y in h8; simpl in h8. erewrite f_equal_dep in h8. rewrite h12 in h8. contradict h6. apply h8. reflexivity. specialize (h4 hf0). erewrite f_equal_dep. apply h4. reflexivity. intros T U FT FU h1 h2 P. rewrite finite_conj_prop_iff. specialize (h2 (P_FinUT' h1 P)). rewrite finite_conj_prop_iff in h2. split; intro h3. intros l h4 a h5. hl h2. intro x. unfold P_FinUT', P_FinUT, fin_transferT; auto. auto. intro x. hr h2. intros l h4 a h5. unfold P_FinUT', P_FinUT, fin_transferT. pose proof (fin_eqUT _ _ h1) as h6. specialize (h3 (fin_transfer_lT h6 l)). hfirst h3. intros y h7. rewrite in_fin_transfer_lT_iff in h7. specialize (h4 _ h7). unfold P_FinUT', P_FinUT in h4. rewrite fin_transferT_undoes in h4; auto. specialize (h3 hf). specialize (h3 (fin_transferT h6 a)). hfirst h3. contradict h5. rewrite in_fin_transfer_lT_iff in h5. rewrite fin_transferT_undoes in h5; auto. specialize (h3 hf0). erewrite f_equal. apply h3. unfold fin_transferT. f_equal. apply proof_irrelevance. rewrite <- h2 in hr. pose proof (fin_eqUT _ _ h1) as h4. specialize (hr (fin_transferT (eq_sym h4) x)). unfold P_FinUT', P_FinUT in hr. rewrite fin_transferT_undoes in hr; auto. Qed. Lemma FinT_induction_strong : forall {T} (P:T->Prop), FinT T -> (forall (l:list T), (forall x (pf:In x l), P x) -> forall a, ~In a l -> P a) -> forall x, P x. intros T P FT h1. rewrite <- (finite_conj_prop_incl_iff FT) in h1 at 1. rewrite finite_conj_prop_iff in h1; auto. Qed. Definition tuple_prod_all_full {S} (FS:FinT S) (A:S->Type) : Type := seg_fun_prod_type (fun i (pf:i < cardT FS) => A (nthT FS i pf)). End FinT. bool2-v0-3/Classical_Wf.v0000644000175000017500000001077413575574055016074 0ustar dwyckoff76dwyckoff76(*Copyright (C) 2016, Daniel Wyckoff, except for the portions labeled "Schepler" which were copied and pasted from Daniel Schepler's "Zorn's Lemma" package.*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Import Classical. Require Import SetUtilities. Require Import Relation_Definitions. Require Import Relation_Definitions_Implicit. Require Import ClassicalChoice. Require Import Arith. Section MinimalElements. Variable T:Type. Variable R:relation T. (*Schepler*) (* R is well-founded if and only if every nonempty subset of T has a minimal element *) Definition minimal_element_property : Prop := forall S:Ensemble T, Inhabited S -> exists x:T, In S x /\ forall y:T, In S y -> ~ R y x. (*Schepler*) Lemma WF_implies_MEP: well_founded R -> minimal_element_property. Proof. unfold well_founded. unfold minimal_element_property. intros WF S Hinh. destruct Hinh. revert x H. apply (@well_founded_ind T R WF (fun x:T => In S x -> exists y:T, In S y /\ (forall z:T, In S z -> ~ R z y))). intros. case (classic (forall y:T, In S y -> ~ R y x)). exists x. split. assumption. assumption. intro. apply not_all_ex_not in H1. destruct H1. apply imply_to_and in H1. destruct H1. apply H with x0. apply NNPP. assumption. assumption. Qed. (*Schepler*) Lemma MEP_implies_WF: minimal_element_property -> well_founded R. Proof. unfold well_founded. unfold minimal_element_property. intro MEP. apply NNPP. intuition. apply not_all_ex_not in H. destruct H. assert (Inhabited [x:T | ~ Acc R x]). exists x. constructor; assumption. apply MEP in H0. destruct H0. destruct H0. destruct H0. contradict H0. constructor. intros. apply NNPP. intuition. apply H1 with y. constructor; assumption. assumption. Qed. End MinimalElements. (*Wyckoff*) Lemma mep_nat : minimal_element_property nat lt. apply WF_implies_MEP. apply lt_wf. Qed. Section DecreasingSequences. (* R is well-founded if and only if there is no infinite strictly decreasing sequence of elements of T *) Variable T:Type. Variable R:relation T. (*Schepler*) Definition decreasing_sequence_property := forall a:nat->T, exists n:nat, ~ R (a (S n)) (a n). (*Schepler*) Lemma WF_implies_DSP: well_founded R -> decreasing_sequence_property. Proof. unfold decreasing_sequence_property. intros WF a. remember (a 0) as a0. revert a0 a Heqa0. apply (well_founded_ind WF (fun x:T => forall a:nat->T, x = a 0 -> exists n:nat, ~ R (a (S n)) (a n))). intros. case (classic (R (a 1) (a 0))). intro. pose (b := fun n:nat => a (S n)). assert (exists n:nat, ~ R (b (S n)) (b n)). apply H with (a 1). rewrite H0. assumption. trivial. destruct H2. exists (S x0). unfold b in H2. assumption. exists 0. assumption. Qed. (*Schepler*) Lemma DSP_implies_WF: decreasing_sequence_property -> well_founded R. Proof. unfold decreasing_sequence_property. intro DSP. apply MEP_implies_WF. unfold minimal_element_property. intro S0. intros. apply NNPP. intuition. assert (forall x:T, In S0 x -> exists y:T, In S0 y /\ R y x). intros. apply NNPP. intuition. assert (forall y:T, ~(In S0 y /\ R y x)). apply not_ex_all_not. assumption. apply H0. exists x. split. assumption. intros. apply H3 with y. tauto. pose (S_type := {x:T | In S0 x}). assert (exists f:S_type -> S_type, forall x:S_type, R (proj1_sig (f x)) (proj1_sig x)). apply choice with (R:=fun x y:S_type => R (proj1_sig y) (proj1_sig x)). intro. destruct x. simpl. pose proof (H1 x i). destruct H2. destruct H2. exists (exist (fun x:T => In S0 x) x0 H2). simpl. assumption. destruct H2 as [f Hf]. destruct H. pose (b := nat_rect (fun n:nat => S_type) (exist (fun x:T => In S0 x) x H) (fun (n:nat) (x:S_type) => f x)). simpl in b. pose (a := fun n:nat => (proj1_sig (b n))). assert (forall n:nat, R (a (S n)) (a n)). unfold a. intro. simpl. apply Hf. contradict DSP. apply ex_not_not_all. exists a. apply all_not_not_ex. auto. Qed. End DecreasingSequences. bool2-v0-3/LogicUtilities.v0000644000175000017500000127415613575574055016502 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-2019, Daniel Wyckoff, except for the portions so labeleled which I got from Daniel Schepler, in his Zorn's Lemma package.*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (* Some of these lemmas are copied and pasted from Daniel Schepler's "Zorns Lemma" package. Such lemmas are notated "Schepler." Some tactics come from Benjamen Peirce's "Software Foundations" and are marked with "SF." All others are labeled "Wyckoff". *) (*Version 0.3 bool2 -- certified -- Coq 8.4pl4*) Require Import Specif. Require Import ProofIrrelevance. Require Import FunctionalExtensionality. Require Import Description. Require Import ClassicalChoice. Require Import EqDec. Require Import FunctionProperties. Require Import Equality. Lemma eq_JMeq : forall {T} {x y:T}, x = y -> x ~= y. intros; subst; constructor. Qed. (*Wyckoff*) Lemma neq_sym : forall {T:Type} (a b:T), a <> b -> b <> a. intros ? ? ? h1. intro. subst. contradict h1. apply eq_refl. Qed. (*Wyckoff*) Lemma neq_sym_iff : forall {T:Type} (a b:T), a <> b <-> b <>a. intros; split; apply neq_sym. Qed. (*Wyckoff*) Lemma eq_sym_iff : forall {T:Type} (a b:T), a = b <-> b = a. intros; split; auto. Qed. (*Wyckoff*) Lemma eq_iff : forall {T:Type} {P:T->Prop} (a b:T), a = b -> (P a <-> P b). intros; subst. tauto. Qed. (*Wyckoff*) Lemma s_eq_iff : forall {T:Set} {P:T->Prop} (a b:T), a = b -> (P a <-> P b). intros; subst. tauto. Qed. (*Wyckoff*) Lemma P_subst : forall {T:Type} {P:T->Prop} {x y:T} (pf:P x), x = y -> P y. intros; subst; auto. Qed. (*Wyckoff*) Lemma eq2 : forall {T:Type} {x y y':T}, (x = y) -> (x = y') -> (y = y'). intros; subst; auto. Defined. (*Wyckoff*) Lemma eq2' : forall {T:Type} {x x' y:T}, (x = y) -> (x' = y) -> (x = x'). intros; subst; auto. Defined. (*Wyckoff*) Lemma eq_trans : forall {T:Type} {x y z:T}, (x = y) -> (y = z) -> (x = z). intros; subst; auto. Defined. (*Wyckoff*) Lemma eq_trans' : forall {T:Type} {x y z:T}, (y = x) -> (z = y) -> (z = x). intros; subst; auto. Defined. (*Wyckoff*) Lemma neq_eq_trans : forall {T} {x y z:T}, x <> y -> y = z -> x <> z. intros; intro; subst; contradict H; auto. Qed. (*Wyckoff*) Lemma eq_neq_trans : forall {T} (a b c:T), a = b -> b <> c -> a <> c. intros; subst; intro; subst; contradict H0; auto. Qed. (*Wyckoff*) Lemma eq_dec_eq : forall {T:Type} (pfdt:type_eq_dec T) (a:T) (pf:{a = a} + {a <> a}), pf = left (eq_refl a). intros ? ? ? h1; destruct h1; f_equal. apply proof_irrelevance. contradict n; auto. Qed. (*Wyckoff*) Lemma eq_dec_neq_r : forall {T:Type} (pfdt:type_eq_dec T) (a a':T) (pfneq:a <> a') (pf: {a' = a} + {a' <> a}), pf = right (neq_sym _ _ pfneq). intros T h1 a a' h2 h3. destruct h3 as [|h3]; subst. contradict h2; auto. f_equal. apply proof_irrelevance. Qed. (*These switch functions are a bit out of place, but have to be in this file for dependency reasons.*) (*Wyckoff*) Definition switch_eq_dec {T:Type} (a b:T) (pf:{a = b} + {a <> b}) : {b = a} + {b <> a} := match pf with | left pf => left (eq_sym pf) | right pf => right (neq_sym _ _ pf) end. (*Wyckoff*) Lemma switch_eq_dec_eq : forall {T:Type} (pf:type_eq_dec T) (a b:T), pf a b = switch_eq_dec _ _ (pf b a). intros T hdt a b; unfold switch_eq_dec. destruct (hdt a b) as [|h1], (hdt b a) as [|h2]; f_equal; try apply proof_irrelevance; try contradict h1; try contradict h2; auto. Qed. (*Wyckoff*) Lemma dec_not_l : forall {P Q:Prop}, {P} + {Q} -> ~P -> Q. tauto. Qed. (*Wyckoff*) Lemma dec_not_r : forall {P Q:Prop}, {P} + {Q} -> ~Q -> P. tauto. Qed. (*Wyckoff*) Lemma eq_eq_refl : forall {T} {x:T} (pf:x = x), pf = eq_refl _. intros; apply proof_irrelevance. Qed. Section utilities. Variable P : Prop. Variable T : Type. Variable Q : T -> Prop. (*Wyckoff*) Lemma prop_dis_forall : (forall (x:T), P \/ (Q x)) -> (P \/ (forall (y:T), (Q y))). intro h1. assert (h2: P \/ ~P). tauto. case h2. tauto. intro h3. right. intro y. assert (h4: P \/ Q y). apply h1. tauto. Qed. End utilities. Section exist_fact. (*Wyckoff*) Lemma existTexist : forall (A:Type) (P:A->Prop) (x y:A) (p:(P x)) (q:(P y)), existT P x p = existT P y q -> exist P x p = exist P y q. intros A P x y p q h1. assert (h2: sig_of_sigT (existT P x p) = sig_of_sigT (existT P y q)). rewrite h1. reflexivity. apply h2. Qed. (*Wyckoff*) Lemma exist_injective : forall {U:Type} (P:U->Prop) (x y:U) (p:P x) (q:P y), exist _ x p = exist _ y q -> x = y. intros U P x y p q h1. pose proof (f_equal (@proj1_sig _ _) h1) as h2. simpl in h2. assumption. Qed. (*Schepler*) Lemma proj1_sig_injective: forall {A:Type} (P:A->Prop) (a1 a2:{x:A | P x}), proj1_sig a1 = proj1_sig a2 -> a1 = a2. Proof. intros. destruct a1. destruct a2. simpl in H. apply existTexist. apply subsetT_eq_compat; trivial. Qed. (*Schepler*) Lemma subset_eq_compatT: forall (U:Type) (P:U->Prop) (x y:U) (p:P x) (q:P y), x = y -> exist P x p = exist P y q. Proof. intros. destruct H. f_equal. apply proof_irrelevance. Qed. (*Wyckoff*) Lemma proj1_sig_eq_iff : forall {T} {P:T->Prop} (x:{t:T | P t}) (y:T), proj1_sig x = y <-> exists pf:P y, x = exist _ y pf. intros T P x y; split; intro h1; subst. exists (proj2_sig _). apply proj1_sig_injective; simpl; auto. destruct h1; subst; auto. Qed. (*Schepler*) Lemma choice_on_dependent_type: forall {A:Type} {B:A->Type} (R:forall a:A, B a -> Prop), (forall a:A, exists b:B a, R a b) -> exists f:(forall a:A, B a), forall a:A, R a (f a). Proof. intros. destruct (choice (fun (a:A) (b:{a:A & B a}) => match b with existT a' b0 => a=a' /\ R a' b0 end)) as [choice_fun]. intro a. destruct (H a) as [b]. exists (existT (fun a:A => B a) a b). split; trivial. assert (f0:forall a:A, {b:B a | R a b}). intro. pose proof (H0 a). destruct (choice_fun a) as [a' b]. destruct H1. destruct H1. exists b; trivial. exists (fun a:A => proj1_sig (f0 a)). intro. destruct (f0 a) as [b]. exact r. Qed. (*Wyckoff*) Lemma unfold_sig : forall {T:Type} (P:T->Prop) (x:{t:T | P t}), x = exist _ (proj1_sig x) (proj2_sig x). intros T P x. destruct x. simpl. reflexivity. Qed. (*Wyckoff*) Lemma proj1_sig_of_exist : forall {T:Type} {P:T->Prop} (x:T) (pf:P x), x = proj1_sig (exist P x pf). intros; auto. Qed. (*Wyckoff*) Lemma propositional_ex_unq : forall (P:Prop) (R:P->Prop), (exists pf:P, R pf) -> exists! pf:P, R pf. intros P R h1. destruct h1 as [h2 h3]. exists h2. red; split; auto. intros; apply proof_irrelevance. Qed. (*Wyckoff*) Definition constructive_definite_proof {P:Prop} {R:P->Prop} (pfe:exists pfp:P, R pfp) : {pfp:P | R pfp} := let x := constructive_definite_description _ (propositional_ex_unq _ _ pfe) in exist _ (proj1_sig x) (proj2_sig x). (*Wyckoff*) Lemma destruct_exists_pf_set_aux : forall {T:Type} (P:T->Prop) (x:T) (Q:P x->Prop), (exists pf:P x, Q pf) -> P x. intros T P x Q h1. destruct h1; auto. Qed. (*Wyckoff*) Lemma destruct_exists_pf_set : forall {T:Type} (P:T->Prop) (x:T) (Q:P x->Prop) (pf:exists pf:P x, Q pf), Q (destruct_exists_pf_set_aux P x Q pf). intros T P x Q h1. destruct h1 as [h1 h2]. pose proof (proof_irrelevance _ h1 (destruct_exists_pf_set_aux P x Q (ex_intro (fun pf : P x => Q pf) h1 h2))) as h3. rewrite <- h3. assumption. Qed. (*Wyckoff*) Lemma unq_ex_impl_ex : forall {T:Type} (P:T->Prop), (exists! x:T, P x) -> exists x:T, P x. intros T P h1. destruct h1 as [x h1]. red in h1. destruct h1 as [h1 h2]. exists x; auto. Qed. End exist_fact. (*Wyckoff*) Lemma conditional_correspondence : forall (P Q:Prop) (h1:{P}+{~P}) (h2:{Q}+{~Q}), (if h1 then true else false) = (if h2 then true else false) -> (P<->Q). intros P Q h1 h2 h3. assert (h4:forall (P' Q':Prop) (h1:{P'}+{~P'}) (h2:{Q'}+{~Q'}), (if h1 then true else false) = (if h2 then true else false) -> (P'->Q')). intros P' Q' h1' h2' h3'. destruct h1' as [h1a | h1b]; destruct h2' as [h2a | h2b]. auto. discriminate. discriminate. tauto. red. split. apply (h4 _ _ h1 h2 h3). symmetry in h3. apply (h4 _ _ h2 h1 h3). Qed. (*Wyckoff*) Definition iff1 {P Q:Prop} (pf:P <-> Q) (pfp:P) : Q. rewrite pf in pfp. refine pfp. Defined. (*Wyckoff*) Definition iff2 {P Q:Prop} (pf:P <-> Q) (pfq:Q) : P. rewrite <- pf in pfq. refine pfq. Defined. (*Wyckoff*) Definition iffn1 {P Q:Prop} (pf:P <-> Q) (pfp: ~P) : ~Q. rewrite pf in pfp. refine pfp. Defined. (*Wyckoff*) Definition iffn2 {P Q:Prop} (pf:P <-> Q) (pfq: ~Q) :~ P. rewrite <- pf in pfq. refine pfq. Defined. (*Wyckoff*) Lemma contrapos : forall (P Q:Prop), (P -> Q) -> (~Q -> ~P). intros; tauto. Qed. (*Wyckoff*) Definition simpl_prop (P:Prop) : Prop. simpl P. refine P. Defined. (*Wyckoff*) Definition simpl_pf {P:Prop} (pf:P) : (simpl_prop P). simpl pf. refine pf. Defined. (*Section Tactics. Top Level.*) (*Wyckoff*) Ltac hfirst H := let h := fresh "hf" in let P := type of H in match P with ?hyp -> _ => assert (h:hyp) end. (*Wyckoff*) Ltac hfirst_dec H := let h := fresh "hfd" in let P := type of H in match P with ?hyp -> _ => assert (hfd:{hyp} + {~hyp}) end. (*Wyckoff*) Ltac nhfirst_dec H := let h := fresh "hfd" in let P := type of H in match P with ~ ?hyp -> _ => assert (hfd:{hyp} + {~hyp}) end. (*Wyckoff*) Ltac hneg H := let h := fresh "hn" in let hyp := type of H in assert (h:~hyp). (*Wyckoff*) Ltac negh H := let h := fresh "nh" in let P := type of H in match P with ~?hyp => assert (h:hyp) end. (*Wyckoff*) Ltac hl H := let h := fresh "hl" in let P := type of H in match P with ?hyp <-> _ => assert (h:hyp) end. (*Wyckoff*) Ltac hr H := let h := fresh "hr" in let P := type of H in match P with _ <-> ?hyp => assert (h:hyp) end. (*Wyckoff*) Ltac nhl H := let P := type of H in match P with ?hyp <-> _ => let h := fresh "hl" in assert (h: ~hyp) end. (*Wyckoff*) Ltac nhr H := let h := fresh "hr" in let P := type of H in match P with _ <-> ?hyp => assert (h: ~hyp) end. (*Wyckoff*) Ltac ghl := let h := fresh "hl" in match goal with |- ?hyp <-> _ => assert (h:hyp) end. (*Wyckoff*) Ltac ghr := let h := fresh "hr" in match goal with |- _ <-> ?hyp => assert (h:hyp) end. (*Wyckoff*) Ltac gnhl := let h := fresh "hl" in match goal with |- ?hyp <-> _ => assert (h: ~hyp) end. (*Wyckoff*) Ltac gnhr := let h := fresh "hr" in match goal with |- _ <-> ?hyp => assert (h: ~hyp) end. (*Wyckoff*) Ltac ex_goal := match goal with |- exists h:?pf, _ => let x := fresh "hex" in assert (x:pf) end. (*Wyckoff*) Ltac terml h := let P := type of h in match P with ?x = _ => let x' := fresh "x" in pose x as x' end. (*Wyckoff*) Ltac termr h := let P := type of h in match P with _ = ?y => let y' := fresh "y" in pose y as y' end. (*Wyckoff*) Ltac gterml := match goal with |- ?x = _ => let x' := fresh "x" in pose x as x' end. (*Wyckoff*) Ltac gtermr := match goal with |- _ = ?y => let y' := fresh "y" in pose y as y' end. (*Wyckoff*) Ltac nterml h := let P := type of h in match P with ?x <> _ => let x' := fresh "x" in pose x as x' end. (*Wyckoff*) Ltac ntermr h := let P := type of h in match P with _ <> ?y => let y' := fresh "y" in pose y as y' end. (*Wyckoff*) Ltac ngterml := match goal with |- ?x <> _ => let x' := fresh "x" in pose x as x' end. (*Wyckoff*) Ltac ngtermr := match goal with |- _ <> ?y => let y' := fresh "y" in pose y as y' end. (*Poses the offset term in a predicate or the left hand side of an equality. (Do a [symmetry] for right-hand match. As Below So Above.*) Ltac gtermO0 := let c := fresh "c" in match goal with | [|- _ ?c'] => pose c' as c | [|- _ _ ?c'] => pose c' as c | [|- _ _ _ ?c'] => pose c' as c | [|- _ _ _ _ ?c'] => pose c' as c | [|- _ _ _ _ _ ?c'] => pose c' as c | [|- _ _ _ _ _ _ ?c']=> pose c' as c | [|- _ _ _ _ _ _ _ ?c']=> pose c' as c | [|- _ _ _ _ _ _ _ _ ?c'] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ ?c'] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ _ ?c'] => pose c' as c end. (*Wyckoff*) Ltac gtermO1 := let c := fresh "c" in match goal with | [|- _ ?c' _] => pose c' as c | [|- _ _ ?c' _] => pose c' as c | [|- _ _ _ ?c' _] => pose c' as c | [|- _ _ _ _ ?c' _] => pose c' as c | [|- _ _ _ _ _ ?c' _] => pose c' as c | [|- _ _ _ _ _ _ ?c' _]=> pose c' as c | [|- _ _ _ _ _ _ _ ?c' _]=> pose c' as c | [|- _ _ _ _ _ _ _ _ ?c' _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ ?c' _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ _ ?c' _] => pose c' as c end. (*Wyckoff*) Ltac gtermO2 := let c := fresh "c" in match goal with | [|- _ ?c' _ _] => pose c' as c | [|- _ _ ?c' _ _] => pose c' as c | [|- _ _ _ ?c' _ _] => pose c' as c | [|- _ _ _ _ ?c' _ _] => pose c' as c | [|- _ _ _ _ _ ?c' _ _] => pose c' as c | [|- _ _ _ _ _ _ ?c' _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ ?c' _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ _ ?c' _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ ?c' _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ _ ?c' _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermO3 := let c := fresh "c" in match goal with | [|- _ ?c' _ _ _] => pose c' as c | [|- _ _ ?c' _ _ _] => pose c' as c | [|- _ _ _ ?c' _ _ _] => pose c' as c | [|- _ _ _ _ ?c' _ _ _] => pose c' as c | [|- _ _ _ _ _ ?c' _ _ _] => pose c' as c | [|- _ _ _ _ _ _ ?c' _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ ?c' _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ _ ?c' _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ ?c' _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ _ ?c' _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermO4 := let c := fresh "c" in match goal with | [|- _ ?c' _ _ _ _] => pose c' as c | [|- _ _ ?c' _ _ _ _] => pose c' as c | [|- _ _ _ ?c' _ _ _ _] => pose c' as c | [|- _ _ _ _ ?c' _ _ _ _] => pose c' as c | [|- _ _ _ _ _ ?c' _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ ?c' _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ ?c' _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ _ ?c' _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ ?c' _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermO5 := let c := fresh "c" in match goal with | [|- _ ?c' _ _ _ _ _] => pose c' as c | [|- _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ _ _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ ?c' _ _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ ?c' _ _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermO6 := let c := fresh "c" in match goal with | [|- _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ ?c' _ _ _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ ?c' _ _ _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermO7 := let c := fresh "c" in match goal with | [|- _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ ?c' _ _ _ _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermO8 := let c := fresh "c" in match goal with | [|- _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermO9 := let c := fresh "c" in match goal with | [|- _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _]=> pose c' as c | [|- _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermA0 := let c := fresh "c" in match goal with | [|- _ -> _ ?c'] => pose c' as c | [|- _ -> _ _ ?c'] => pose c' as c | [|- _ -> _ _ _ ?c'] => pose c' as c | [|- _ -> _ _ _ _ ?c'] => pose c' as c | [|- _ -> _ _ _ _ _ ?c'] => pose c' as c | [|- _ -> _ _ _ _ _ _ ?c']=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ ?c']=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ ?c'] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ ?c'] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?c'] => pose c' as c end. (*Wyckoff*) Ltac gtermA1 := let c := fresh "c" in match goal with | [|- _ -> _ ?c' _] => pose c' as c | [|- _ -> _ _ ?c' _] => pose c' as c | [|- _ -> _ _ _ ?c' _] => pose c' as c | [|- _ -> _ _ _ _ ?c' _] => pose c' as c | [|- _ -> _ _ _ _ _ ?c' _] => pose c' as c | [|- _ -> _ _ _ _ _ _ ?c' _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ ?c' _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ ?c' _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ ?c' _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?c' _] => pose c' as c end. (*Wyckoff*) Ltac gtermA2 := let c := fresh "c" in match goal with | [|- _ -> _ ?c' _ _] => pose c' as c | [|- _ -> _ _ ?c' _ _] => pose c' as c | [|- _ -> _ _ _ ?c' _ _] => pose c' as c | [|- _ -> _ _ _ _ ?c' _ _] => pose c' as c | [|- _ -> _ _ _ _ _ ?c' _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ ?c' _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ ?c' _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ ?c' _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ ?c' _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermA3 := let c := fresh "c" in match goal with | [|- _ -> _ ?c' _ _ _] => pose c' as c | [|- _ -> _ _ ?c' _ _ _] => pose c' as c | [|- _ -> _ _ _ ?c' _ _ _] => pose c' as c | [|- _ -> _ _ _ _ ?c' _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ ?c' _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ ?c' _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ ?c' _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ ?c' _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermA4 := let c := fresh "c" in match goal with | [|- _ -> _ ?c' _ _ _ _] => pose c' as c | [|- _ -> _ _ ?c' _ _ _ _] => pose c' as c | [|- _ -> _ _ _ ?c' _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ ?c' _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ ?c' _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ ?c' _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ ?c' _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermA5 := let c := fresh "c" in match goal with | [|- _ -> _ ?c' _ _ _ _ _] => pose c' as c | [|- _ -> _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ ?c' _ _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ ?c' _ _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermA6 := let c := fresh "c" in match goal with | [|- _ -> _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ ?c' _ _ _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ ?c' _ _ _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermA7 := let c := fresh "c" in match goal with | [|- _ -> _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ ?c' _ _ _ _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermA8 := let c := fresh "c" in match goal with | [|- _ -> _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gtermA9 := let c := fresh "c" in match goal with | [|- _ -> _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _]=> pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _] => pose c' as c end. (*Wyckoff*) Ltac gterm0 := first [gtermO0 | gtermA0]. (*Wyckoff*) Ltac gterm1 := first [gtermO1 | gtermA1]. (*Wyckoff*) Ltac gterm2 := first [gtermO2 | gtermA2]. (*Wyckoff*) Ltac gterm3 := first [gtermO3 | gtermA3]. (*Wyckoff*) Ltac gterm4 := first [gtermO4 | gtermA4]. (*Wyckoff*) Ltac gterm5 := first [gtermO5 | gtermA5]. (*Wyckoff*) Ltac gterm6 := first [gtermO6 | gtermA6]. (*Wyckoff*) Ltac gterm7 := first [gtermO7 | gtermA7]. (*Wyckoff*) Ltac gterm8 := first [gtermO8 | gtermA8]. (*Wyckoff*) Ltac gterm9 := first [gtermO0 | gtermA9]. (*Wyckoff*) Ltac termO0 k := let c := fresh "c" in let P := type of k in match P with | _ ?c' => pose c' as c | _ _ ?c' => pose c' as c | _ _ _ ?c' => pose c' as c | _ _ _ _ ?c' => pose c' as c | _ _ _ _ _ ?c' => pose c' as c | _ _ _ _ _ _ ?c' => pose c' as c | _ _ _ _ _ _ _ ?c'=> pose c' as c | _ _ _ _ _ _ _ _ ?c' => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' => pose c' as c end. (*Wyckoff*) Ltac termO1 k := let c := fresh "c" in let P := type of k in match P with | _ ?c' _ => pose c' as c | _ _ ?c' _ => pose c' as c | _ _ _ ?c' _ => pose c' as c | _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ => pose c' as c end. (*Wyckoff*) Ltac termO2 k := let c := fresh "c" in let P := type of k in match P with | _ ?c' _ _ => pose c' as c | _ _ ?c' _ _ => pose c' as c | _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ => pose c' as c end. (*Wyckoff*) Ltac termO3 k := let c := fresh "c" in let P := type of k in match P with | _ ?c' _ _ _ => pose c' as c | _ _ ?c' _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termO4 k := let c := fresh "c" in let P := type of k in match P with | _ ?c' _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termO5 k := let c := fresh "c" in let P := type of k in match P with | _ ?c' _ _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termO6 k := let c := fresh "c" in let P := type of k in match P with | _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termO7 k := let c := fresh "c" in let P := type of k in match P with | _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termO8 k := let c := fresh "c" in let P := type of k in match P with | _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termO9 k := let c := fresh "c" in let P := type of k in match P with | _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termA0 k := let c := fresh "c" in let P := type of k in match P with | _ -> _ ?c' => pose c' as c | _ -> _ _ ?c' => pose c' as c | _ -> _ _ _ ?c' => pose c' as c | _ -> _ _ _ _ ?c' => pose c' as c | _ -> _ _ _ _ _ ?c' => pose c' as c | _ -> _ _ _ _ _ _ ?c' => pose c' as c | _ -> _ _ _ _ _ _ _ ?c'=> pose c' as c | _ -> _ _ _ _ _ _ _ _ ?c' => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ ?c' => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ _ ?c' => pose c' as c end. (*Wyckoff*) Ltac termA1 k := let c := fresh "c" in let P := type of k in match P with | _ -> _ ?c' _ => pose c' as c | _ -> _ _ ?c' _ => pose c' as c | _ -> _ _ _ ?c' _ => pose c' as c | _ -> _ _ _ _ ?c' _ => pose c' as c | _ -> _ _ _ _ _ ?c' _ => pose c' as c | _ -> _ _ _ _ _ _ ?c' _ => pose c' as c | _ -> _ _ _ _ _ _ _ ?c' _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ ?c' _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ ?c' _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ => pose c' as c end. (*Wyckoff*) Ltac termA2 k := let c := fresh "c" in let P := type of k in match P with | _ -> _ ?c' _ _ => pose c' as c | _ -> _ _ ?c' _ _ => pose c' as c | _ -> _ _ _ ?c' _ _ => pose c' as c | _ -> _ _ _ _ ?c' _ _ => pose c' as c | _ -> _ _ _ _ _ ?c' _ _ => pose c' as c | _ -> _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ => pose c' as c end. (*Wyckoff*) Ltac termA3 k := let c := fresh "c" in let P := type of k in match P with | _ -> _ ?c' _ _ _ => pose c' as c | _ -> _ _ ?c' _ _ _ => pose c' as c | _ -> _ _ _ ?c' _ _ _ => pose c' as c | _ -> _ _ _ _ ?c' _ _ _ => pose c' as c | _ -> _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termA4 k := let c := fresh "c" in let P := type of k in match P with | _ -> _ ?c' _ _ _ _ => pose c' as c | _ -> _ _ ?c' _ _ _ _ => pose c' as c | _ -> _ _ _ ?c' _ _ _ _ => pose c' as c | _ -> _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termA5 k := let c := fresh "c" in let P := type of k in match P with | _ -> _ ?c' _ _ _ _ _ => pose c' as c | _ -> _ _ ?c' _ _ _ _ _ => pose c' as c | _ -> _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termA6 k := let c := fresh "c" in let P := type of k in match P with | _ -> _ ?c' _ _ _ _ _ _ => pose c' as c | _ -> _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termA7 k := let c := fresh "c" in let P := type of k in match k with | _ -> _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termA8 k := let c := fresh "c" in let P := type of k in match P with | _ -> _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac termA9 k := let c := fresh "c" in let P := type of k in match P with | _ -> _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ -> _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac term0 k := first [termO0 k| termA0 k]. (*Wyckoff*) Ltac term1 k := first [termO1 k| termA1 k]. (*Wyckoff*) Ltac term2 k := first [termO2 k| termA2 k]. (*Wyckoff*) Ltac term3 k := first [termO3 k| termA3 k]. (*Wyckoff*) Ltac term4 k := first [termO4 k| termA4 k]. (*Wyckoff*) Ltac term5 k := first [termO5 k| termA5 k]. (*Wyckoff*) Ltac term6 k := first [termO6 k| termA6 k]. (*Wyckoff*) Ltac term7 k := first [termO7 k| termA7 k]. (*Wyckoff*) Ltac term8 k := first [termO8 k| termA8 k]. (*Wyckoff*) Ltac term9 k := first [termO0 k| gtermA9 k]. (*Wyckoff*) Ltac redterm0 k := let c := fresh "c" in let k' := eval red in k in match k' with | _ ?c' => pose c' as c | _ _ ?c' => pose c' as c | _ _ _ ?c' => pose c' as c | _ _ _ _ ?c' => pose c' as c | _ _ _ _ _ ?c' => pose c' as c | _ _ _ _ _ _ ?c' => pose c' as c | _ _ _ _ _ _ _ ?c'=> pose c' as c | _ _ _ _ _ _ _ _ ?c' => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' => pose c' as c end. (*Wyckoff*) Ltac redterm1 k := let c := fresh "c" in let k' := eval red in k in match k' with | _ ?c' _ => pose c' as c | _ _ ?c' _ => pose c' as c | _ _ _ ?c' _ => pose c' as c | _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ => pose c' as c end. (*Wyckoff*) Ltac redterm2 k := let c := fresh "c" in let k' := eval red in k in match k' with | _ ?c' _ _ => pose c' as c | _ _ ?c' _ _ => pose c' as c | _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ => pose c' as c end. (*Wyckoff*) Ltac redterm3 k := let c := fresh "c" in let k' := eval red in k in match k' with | _ ?c' _ _ _ => pose c' as c | _ _ ?c' _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ => pose c' as c end. (*Wyckoff*) Ltac redterm4 k := let c := fresh "c" in let k' := eval red in k in match k' with | _ ?c' _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac redterm5 k := let c := fresh "c" in let k' := eval red in k in match k' with | _ ?c' _ _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac redterm6 k := let c := fresh "c" in let k' := eval red in k in match k' with | _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac redterm7 k := let c := fresh "c" in let k' := eval red in k in match k' with | _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac redterm8 k := let c := fresh "c" in let k' := eval red in k in match k' with | _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac redterm9 k := let c := fresh "c" in let k' := eval red in k in match k' with | _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c | _ _ _ _ _ _ _ _ _ _ ?c' _ _ _ _ _ _ _ _ _ => pose c' as c end. (*Wyckoff*) Ltac ltl k := let c := fresh "xlt" in let P := type of k in match P with | ?c' < _ => pose c' as c end. (*Wyckoff*) Ltac ltr k := let c := fresh "xlt" in let P := type of k in match P with | _ < ?c' => pose c' as c end. (*Wyckoff*) Ltac gtl k := let c := fresh "xgt" in let P := type of k in match P with | ?c' > _ => pose c' as c end. (*Wyckoff*) Ltac gtr k := let c := fresh "xgt" in let P := type of k in match P with | _ > ?c' => pose c' as c end. (*Wyckoff*) Ltac gltl := let c := fresh "xlt" in match goal with |- ?c' < _ => pose c' as c end. (*Wyckoff*) Ltac gltr := let c := fresh "xlt" in match goal with |- _ < ?c' => pose c' as c end. (*Wyckoff*) Ltac ggtl := let c := fresh "xgt" in match goal with |- ?c' > _ => pose c' as c end. (*Wyckoff*) Ltac ggtr := let c := fresh "xgt" in match goal with |- _ > ?c' => pose c' as c end. (*Wyckoff*) Ltac lel k := let c := fresh "xle" in let P := type of k in match P with | ?c' <= _ => pose c' as c end. (*Wyckoff*) Ltac ler k := let c := fresh "xle" in let P := type of k in match P with | _ <= ?c' => pose c' as c end. (*Wyckoff*) Ltac gel k := let c := fresh "xge" in let P := type of k in match P with | ?c' >= _ => pose c' as c end. (*Wyckoff*) Ltac ger k := let c := fresh "xge" in let P := type of k in match P with | _ >= ?c' => pose c' as c end. (*Wyckoff*) Ltac glel := let c := fresh "xle" in match goal with |- ?c' <= _ => pose c' as c end. (*Wyckoff*) Ltac gler := let c := fresh "xle" in match goal with |- _ <= ?c' => pose c' as c end. (*Wyckoff*) Ltac ggel := let c := fresh "xge" in match goal with |- ?c' >= _ => pose c' as c end. (*Wyckoff*) Ltac gger := let c := fresh "xge" in match goal with |- _ >= ?c' => pose c' as c end. (*Wyckoff*) (*This is used to eliminate dependencies of [let]-posed propositions by posing such as a proof and clearing it.*) Ltac redterm_cp h := let pf := eval red in h in let h0 := fresh "hc" in pose proof h as h0; clear h. (*Wyckoff*) Ltac match_match_goal := match goal with |- context [match ?mtch' with | 0 => _ | S _ => _ end] => let y := fresh "mtch" in pose mtch' as y end. (*Wyckoff*) Ltac match_match_match_goal := match goal with |- context [match ?mtch' with | 0 => _ | S _ => match _ with | 0 => _ | S _ => _ end end] => let y := fresh "mtch" in pose mtch' as y end. (*Wyckoff*) Ltac minusl pf := match pf with ?a' - _ => let y := fresh "ma" in pose a' as y end. (*Wyckoff*) Ltac minusr pf := match pf with _ - ?b' => let y := fresh "mb" in pose b' as y end. (*This isolates the discriminee in a sumbool match in the given hypotheseis [h]*) (*An analogous version of SF's [desruct_head_match] but for hypotheses, and easier for me to understand!*) (*Wyckoff*) Ltac lr_match h := let P := type of h in let y := fresh "hlr" in match P with context [match ?x with | left _ => _ | right _ => _ end] => pose x as y end. (*Wyckoff*) Ltac lr_match_goal := let y := fresh "hlr" in match goal with |- context [match ?x with | left _ => _ | right _ => _ end] => pose x as y end. (*Wyckoff*) Ltac lr_match_sum h := let y := fresh "hlr" in let P := type of h in match P with context [match ?x with | inleft _ => _ | inright _ => _ end] => pose x as y end. (*Wyckoff*) Ltac lr_match_sum_goal := let y := fresh "hlr" in match goal with |- context [match ?x with | inleft _ => _ | inright _ => _ end] => pose x as y end. (*Sometimes lr_match doesn't work on ifs, and so . . . *) (*Wyckoff*) Ltac lr_if h := let y := fresh "hlr" in let P := type of h in match P with context [if ?x then _ else _] => pose x as y end. (*Wyckoff*) Ltac lr_if_goal := let y := fresh "hlr" in match goal with |- context [if ?x then _ else _] => pose x as y end. (*Wyckoff*) Ltac match_eq_refl h := let pfef := fresh "hef" in let P := type of h in match P with context [match ?e with | eq_refl => _ end] => pose e as pfef end. (*Wyckoff*) Ltac match_eq_refl_goal := let pfef := fresh "hef" in match goal with |- context [match ?e with | eq_refl => _ end] => pose e as pfef end. (*Wyckoff*) (*The [_x]-ed versions of the above 2 generalize and clear matched term.*) Ltac match_eq_refl_x h := let pfef := fresh "hef" in let P := type of h in match P with context [match ?e with | eq_refl => _ end] => pose e as pfef; fold pfef in h; generalize h pfef; clear pfef; intros h pfef end. Ltac match_eq_refl_goal_x := let pfef := fresh "hef" in match goal with |- context [match ?e with | eq_refl => _ end] => pose e as pfef; fold pfef; generalize pfef; clear pfef; intros pfef end. (*These primed versions are for when you want the default behavior, and a one-step solution. The unprimed versions are more useful when you want more control (like folding a matched expression in other expressions).*) Ltac lr_match' h := let y := fresh "hlr" in let hl' := fresh "hl" in let hr' := fresh "hr" in let P := type of h in match P with context [match ?x with | left _ => _ | right _ => _ end] => pose x as y; fold y in h; destruct y as [hl'|hr'] end. (*Wyckoff*) Ltac lr_match_goal' := let y := fresh "hlr" in let hl' := fresh "hl" in let hr' := fresh "hr" in match goal with |- context [match ?x with | left _ => _ | right _ => _ end] => pose x as y; fold y; destruct y as [hl'|hr'] end. (*Wyckoff*) Ltac lr_if' h := let y := fresh "hlr" in let hl' := fresh "hl" in let hr' := fresh "hr" in let P := type of h in match P with context [if ?x then _ else _] => pose x as y; fold y in h; destruct y as [hl'|hr'] end. (*Wyckoff*) Ltac lr_if_goal' := let y := fresh "hlr" in let hl' := fresh "hl" in let hr' := fresh "hr" in match goal with |- context [if ?x then _ else _] => pose x as y; fold y; destruct y as [hl'|hr'] end. (*Wyckoff*) Ltac lr_match_sum' h := let y := fresh "hlr" in let hl' := fresh "hl" in let hr' := fresh "hr" in let P := type of h in match P with context [match ?x with | inleft _ => _ | inright _ => _ end] => pose x as y; fold y in h; destruct y as [hl'|hr'] end. (*Wyckoff*) Ltac lr_match_sum_goal' := let y := fresh "hlr" in let hl' := fresh "hl" in let hr' := fresh "hr" in match goal with |- context [match ?x with | inleft _ => _ | inright _ => _ end] => pose x as y; fold y; destruct y as [hl'|hr'] end. (*Wyckoff*) Ltac match_eq_refl' h := let pfef := fresh "hef" in let pfef' := fresh "hef" in let P := type of h in match P with context [match ?e in (_ = _) return (_ _) with | eq_refl => _ end] => pose e as pfef; fold pfef; destruct pfef as [pfef'] end. (*Wyckoff*) Ltac match_eq_refl_imp' h := let pfef := fresh "hef" in let pfef' := fresh "hef" in let P := type of h in match P with context [match ?e in (_ = _) return (_ -> _) with | eq_refl => _ end] => pose e as pfef; fold pfef; destruct pfef as [pfef'] end. (*Wyckoff*) Ltac match_eq_refl_goal' := let pfef := fresh "hef" in let pfef' := fresh "hef" in match goal with |- context [match ?e in (_ = _) return (_ _) with | eq_refl => _ end] => pose e as pfef; fold pfef; destruct pfef as [pfef'] end. (*Wyckoff*) Ltac match_eq_refl_imp_goal' := let pfef := fresh "hef" in let pfef' := fresh "hef" in match goal with |- context [match ?e in (_ = _) return (_ -> _) with | eq_refl => _ end] => pose e as pfef; fold pfef; destruct pfef as [pfef'] end. (*The pattern for these (eventually?) 1000 generalization tactics is to match 0-9 hypotheses with gen[O-I][n][O-I]. *) (*Adding too many generalizations is of marginal utility, so it's not quite 1000 generalizations, although I prepared incomplete templates for them in the Preview file (the general file for half-worked stuff that won't make the next release).*) (*Wyckoff*) Ltac genO0O := match goal with | [|- _ _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ _ _ _ _ _ ?h = _] => generalize h | [|- _ _ _ _ _ ?h = _] => generalize h | [|- _ _ _ _ ?h = _] => generalize h | [|- _ _ _ ?h = _] => generalize h | [|- _ _ ?h = _] => generalize h | [|- _ ?h = _] => generalize h | [|- _ _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ _ _ _ _ _ _ ?h] => generalize h | [|- _ _ _ _ _ _ ?h] => generalize h | [|- _ _ _ _ _ ?h] => generalize h | [|- _ _ _ _ ?h] => generalize h | [|- _ _ _ ?h] => generalize h | [|- _ _ ?h] => generalize h | [|- _ ?h] => generalize h end. (*Wyckoff*) Ltac genO1O := match goal with | [|- _ _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ _ _ _ _ ?h _ = _] => generalize h | [|- _ _ _ _ ?h _ = _] => generalize h | [|- _ _ _ ?h _ = _] => generalize h | [|- _ _ ?h _ = _] => generalize h | [|- _ ?h _ = _] => generalize h | [|- _ _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ _ _ _ _ _ ?h _] => generalize h | [|- _ _ _ _ _ ?h _] => generalize h | [|- _ _ _ _ ?h _] => generalize h | [|- _ _ _ ?h _] => generalize h | [|- _ _ ?h _] => generalize h | [|- _ ?h _] => generalize h end. (*Wyckoff*) Ltac genO2O := match goal with | [|- _ _ _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ _ _ _ ?h _ _ = _] => generalize h | [|- _ _ _ ?h _ _ = _] => generalize h | [|- _ _ ?h _ _ = _] => generalize h | [|- _ ?h _ _ = _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ _ _ _ _ ?h _ _] => generalize h | [|- _ _ _ _ ?h _ _] => generalize h | [|- _ _ _ ?h _ _] => generalize h | [|- _ _ ?h _ _] => generalize h | [|- _ ?h _ _] => generalize h end. (*Wyckoff*) Ltac genO3O := match goal with | [|- _ _ _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ _ _ ?h _ _ _ = _] => generalize h | [|- _ _ ?h _ _ _ = _] => generalize h | [|- _ ?h _ _ _ = _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ _ _ _ ?h _ _ _] => generalize h | [|- _ _ _ ?h _ _ _] => generalize h | [|- _ _ ?h _ _ _] => generalize h | [|- _ ?h _ _ _] => generalize h end. (*Wyckoff*) Ltac genO4O := match goal with | [|- _ _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ _ ?h _ _ _ _ = _] => generalize h | [|- _ ?h _ _ _ _ = _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ _ _ ?h _ _ _ _] => generalize h | [|- _ _ ?h _ _ _ _] => generalize h | [|- _ ?h _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genO5O := match goal with | [|- _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ _ ?h _ _ _ _ = _] => generalize h | [|- _ ?h _ _ _ _ = _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ _ _ ?h _ _ _ _] => generalize h | [|- _ _ ?h _ _ _ _] => generalize h | [|- _ ?h _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genO6O := match goal with | [|- _ _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ ?h _ _ _ _ _ = _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ _ ?h _ _ _ _ _] => generalize h | [|- _ ?h _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genO7O := match goal with | [|- _ _ _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ ?h _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genO8O := match goal with | [|- _ _ _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genO9O := match goal with | [|- _ _ ?h _ _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _ _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genA0O := match goal with | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _?h = _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _?h = _] => generalize h | [|- _ -> _ _ _ _ _ _ _?h = _] => generalize h | [|- _ -> _ _ _ _ _ _?h = _] => generalize h | [|- _ -> _ _ _ _ _?h = _] => generalize h | [|- _ -> _ _ _ _?h = _] => generalize h | [|- _ -> _ _ _ ?h = _] => generalize h | [|- _ -> _ _ ?h = _] => generalize h | [|- _ -> _ ?h = _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _?h] => generalize h | [|- _ -> _ _ _ _ _ _ _ _?h]=> generalize h | [|- _ -> _ _ _ _ _ _ _?h]=> generalize h | [|- _ -> _ _ _ _ _ _?h] => generalize h | [|- _ -> _ _ _ _ _?h] => generalize h | [|- _ -> _ _ _ _?h] => generalize h | [|- _ -> _ _ _ ?h] => generalize h | [|- _ -> _ _ ?h] => generalize h | [|- _ -> _ ?h] => generalize h end. (*Wyckoff*) Ltac genA1O := match goal with | [|- _ -> _ _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ _ _ ?h _ = _] => generalize h | [|- _ -> _ _ ?h _ = _] => generalize h | [|- _ -> _ ?h _ = _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ _ _ _ ?h _] => generalize h | [|- _ -> _ _ _ ?h _] => generalize h | [|- _ -> _ _ ?h _] => generalize h | [|- _ -> _ ?h _] => generalize h end. (*Wyckoff*) Ltac genA2O := match goal with | [|- _ -> _ _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ _ ?h _ _ = _] => generalize h | [|- _ -> _ ?h _ _ = _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ _ _ ?h _ _] => generalize h | [|- _ -> _ _ ?h _ _] => generalize h | [|- _ -> _ ?h _ _] => generalize h end. (*Wyckoff*) Ltac genA3O := match goal with | [|- _ -> _ _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ ?h _ _ _ = _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ _ ?h _ _ _] => generalize h | [|- _ -> _ ?h _ _ _] => generalize h end. (*Wyckoff*) Ltac genA4O := match goal with | [|- _ -> _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ ?h _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genA5O := match goal with | [|- _ -> _ _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ ?h _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genA6O := match goal with | [|- _ -> _ _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genA7O := match goal with | [|- _ -> _ _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genA8O := match goal with | [|- _ -> _ _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genA9O := match goal with | [|- _ -> _ ?h _ _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genB0O := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ _ ?h = _] => generalize h | [|- _ -> _ -> _ ?h = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ _ _ ?h] => generalize h | [|- _ -> _ -> _ _ ?h] => generalize h | [|- _ -> _ -> _ ?h] => generalize h end. (*Wyckoff*) Ltac genB1O := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ ?h _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ _ ?h _] => generalize h | [|- _ -> _ -> _ ?h _] => generalize h end. (*Wyckoff*) Ltac genB2O := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ ?h _ _] => generalize h end. (*Wyckoff*) Ltac genB3O := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ ?h _ _ _] => generalize h end. (*Wyckoff*) Ltac genB4O := match goal with | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genB5O := match goal with | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genB6O := match goal with | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genB7O := match goal with | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genB8O := match goal with | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genB9O := match goal with | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genC0O := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ ?h] => generalize h end. (*Wyckoff*) Ltac genC1O := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ ?h _] => generalize h end. (*Wyckoff*) Ltac genC2O := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _] => generalize h end. (*Wyckoff*) Ltac genC3O := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _] => generalize h end. (*Wyckoff*) Ltac genC4O := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genC5O := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genC6O := match goal with | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genC7O := match goal with | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genC8O := match goal with | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genC9O := match goal with | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genD0O := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h] => generalize h end. (*Wyckoff*) Ltac genD1O := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _]=> generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _] => generalize h end. (*Wyckoff*) Ltac genD2O := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _] => generalize h end. (*Wyckoff*) Ltac genD3O := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _] => generalize h end. (*Wyckoff*) Ltac genD4O := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genD5O := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genD6O := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genD7O := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genD8O := match goal with | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genD9O := match goal with | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genE0O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h] => generalize h end. (*Wyckoff*) Ltac genE1O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _] => generalize h end. (*Wyckoff*) Ltac genE2O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _] => generalize h end. (*Wyckoff*) Ltac genE3O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _] => generalize h end. (*Wyckoff*) Ltac genE4O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genE5O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genE6O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genE7O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genE8O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genE9O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genF0O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h] => generalize h end. (*Wyckoff*) Ltac genF1O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _] => generalize h end. (*Wyckoff*) Ltac genF2O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _] => generalize h end. (*Wyckoff*) Ltac genF3O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _] => generalize h end. (*Wyckoff*) Ltac genF4O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genF5O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genF6O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genF7O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genF8O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genF9O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genG0O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h] => generalize h end. (*Wyckoff*) Ltac genG1O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _] => generalize h end. (*Wyckoff*) Ltac genG2O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _] => generalize h end. (*Wyckoff*) Ltac genG3O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _] => generalize h end. (*Wyckoff*) Ltac genG4O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genG5O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genG6O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genG7O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genG8O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genG9O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genH0O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h] => generalize h end. (*Wyckoff*) Ltac genH1O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _] => generalize h end. (*Wyckoff*) Ltac genH2O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _] => generalize h end. (*Wyckoff*) Ltac genH3O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _] => generalize h end. (*Wyckoff*) Ltac genH4O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genH5O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genH6O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genH7O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genH8O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genH9O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genI0O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h] => generalize h end. (*Wyckoff*) Ltac genI1O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _] => generalize h end. (*Wyckoff*) Ltac genI2O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _]=> generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _] => generalize h end. (*Wyckoff*) Ltac genI3O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _] => generalize h end. (*Wyckoff*) Ltac genI4O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genI5O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genI6O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genI7O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genI8O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genI9O := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _] => generalize h end. (*Wyckoff*) Ltac genO0A := match goal with | [|- _ _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ _ _ _ ?h = _ -> _] => generalize h | [|- _ _ _ ?h = _ -> _] => generalize h | [|- _ _ ?h = _ -> _] => generalize h | [|- _ ?h = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ _ _ _ _ ?h -> _] => generalize h | [|- _ _ _ _ ?h -> _] => generalize h | [|- _ _ _ ?h -> _] => generalize h | [|- _ _ ?h -> _] => generalize h | [|- _ ?h -> _] => generalize h end. (*Wyckoff*) Ltac genO1A := match goal with | [|- _ _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ _ _ ?h _ = _ -> _] => generalize h | [|- _ _ ?h _ = _ -> _] => generalize h | [|- _ ?h _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ _ _ _ ?h _ -> _] => generalize h | [|- _ _ _ ?h _ -> _] => generalize h | [|- _ _ ?h _ -> _] => generalize h | [|- _ ?h _ -> _] => generalize h end. (*Wyckoff*) Ltac genO2A := match goal with | [|- _ _ _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ _ ?h _ _ = _ -> _] => generalize h | [|- _ ?h _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ _ _ ?h _ _ -> _] => generalize h | [|- _ _ ?h _ _ -> _] => generalize h | [|- _ ?h _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genO3A := match goal with | [|- _ _ _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ ?h _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ _ ?h _ _ _ -> _] => generalize h | [|- _ ?h _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genO4A := match goal with | [|- _ _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ ?h _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genO5A := match goal with | [|- _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ ?h _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genO6A := match goal with | [|- _ _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genO7A := match goal with | [|- _ _ _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genO8A := match goal with | [|- _ _ _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genO9A := match goal with | [|- _ _ ?h _ _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genA0A := match goal with | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ _ ?h = _ -> _] => generalize h | [|- _ -> _ ?h = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ _ _ ?h -> _] => generalize h | [|- _ -> _ _ ?h -> _] => generalize h | [|- _ -> _ ?h -> _] => generalize h end. (*Wyckoff*) Ltac genA1A := match goal with | [|- _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ ?h _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ _ ?h _ -> _] => generalize h | [|- _ -> _ ?h _ -> _] => generalize h end. (*Wyckoff*) Ltac genA2A := match goal with | [|- _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ ?h _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genA3A := match goal with | [|- _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genA4A := match goal with | [|- _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genA5A := match goal with | [|- _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genA6A := match goal with | [|- _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genA7A := match goal with | [|- _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genA8A := match goal with | [|- _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genA9A := match goal with | [|- _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genB0A := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ ?h -> _] => generalize h end. (*Wyckoff*) Ltac genB1A := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ -> _] => generalize h end. (*Wyckoff*) Ltac genB2A := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genB3A := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genB4A := match goal with | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genB5A := match goal with | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genB6A := match goal with | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genB7A := match goal with | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genB8A := match goal with | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genB9A := match goal with | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genC0A := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h -> _] => generalize h end. (*Wyckoff*) Ltac genC1A := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ -> _] => generalize h end. (*Wyckoff*) Ltac genC2A := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genC3A := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genC4A := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genC5A := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genC6A := match goal with | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genC7A := match goal with | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genC8A := match goal with | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genC9A := match goal with | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genD0A := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h -> _] => generalize h end. (*Wyckoff*) Ltac genD1A := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _]=> generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ -> _] => generalize h end. (*Wyckoff*) Ltac genD2A := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genD3A := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genD4A := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genD5A := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genD6A := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genD7A := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genD8A := match goal with | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genD9A := match goal with | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genE0A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h -> _] => generalize h end. (*Wyckoff*) Ltac genE1A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ -> _] => generalize h end. (*Wyckoff*) Ltac genE2A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genE3A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genE4A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genE5A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genE6A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genE7A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genE8A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genE9A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genF0A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h -> _] => generalize h end. (*Wyckoff*) Ltac genF1A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ -> _] => generalize h end. (*Wyckoff*) Ltac genF2A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genF3A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genF4A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genF5A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genF6A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genF7A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genF8A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genF9A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genG0A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h -> _] => generalize h end. (*Wyckoff*) Ltac genG1A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ -> _] => generalize h end. (*Wyckoff*) Ltac genG2A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genG3A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genG4A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genG5A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genG6A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genG7A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genG8A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genG9A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genH0A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h -> _] => generalize h end. (*Wyckoff*) Ltac genH1A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ -> _] => generalize h end. (*Wyckoff*) Ltac genH2A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genH3A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genH4A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genH5A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genH6A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genH7A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genH8A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genH9A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genI0A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h -> _] => generalize h end. (*Wyckoff*) Ltac genI1A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ -> _] => generalize h end. (*Wyckoff*) Ltac genI2A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genI3A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genI4A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genI5A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genI6A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genI7A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genI8A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genI9A := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _] => generalize h end. (*Wyckoff*) Ltac genO0B := match goal with | [|- _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ _ ?h = _ -> _ -> _] => generalize h | [|- _ ?h = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ _ _ ?h -> _ -> _] => generalize h | [|- _ _ ?h -> _ -> _] => generalize h | [|- _ ?h -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO1B := match goal with | [|- _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ ?h _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ _ ?h _ -> _ -> _] => generalize h | [|- _ ?h _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO2B := match goal with | [|- _ _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ ?h _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO3B := match goal with | [|- _ _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO4B := match goal with | [|- _ _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO5B := match goal with | [|- _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO6B := match goal with | [|- _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO7B := match goal with | [|- _ _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO8B := match goal with | [|- _ _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO9B := match goal with | [|- _ _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA0B := match goal with | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ ?h -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA1B := match goal with | [|- _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA2B := match goal with | [|- _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA3B := match goal with | [|- _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA4B := match goal with | [|- _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA5B := match goal with | [|- _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA6B := match goal with | [|- _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA7B := match goal with | [|- _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA8B := match goal with | [|- _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA9B := match goal with | [|- _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB0B := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB1B := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB2B := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB3B := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB4B := match goal with | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB5B := match goal with | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB6B := match goal with | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB7B := match goal with | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB8B := match goal with | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB9B := match goal with | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genC0B := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genC1B:= match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genC2B := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genC3B := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genC4B := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genC5B := match goal with | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genC6B := match goal with | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genC7B := match goal with | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genC8B := match goal with | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genC9B := match goal with | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genD0B := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genD1B := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _ -> _]=> generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genD2B := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genD3B := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genD4B := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genD5B := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genD6B := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genD7B := match goal with | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genD8B := match goal with | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genD9B := match goal with | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genE0B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genE1B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genE2B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genE3B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genE4B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genE5B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genE6B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genE7B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genE8B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genE9B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genF0B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genF1B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genF2B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genF3B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genF4B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genF5B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genF6B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genF7B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genF8B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genF9B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genG0B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genG1B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genG2B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genG3B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _ -> _] => generalize h end. (*The lower-case [genNX]s traverse a string of implications from the right-most implication to the left-most, in accord with their eponymous indices, wherease the upper-case [GenNX]s traverses name indices from left to right.*) Ltac gen0O := first [genO0O | genA0O | genB0O | genC0O | genD0O | genE0O | genF0O | genG0O | genH0O | genI0O]. Ltac gen0A := first [genO0A | genA0A | genB0A | genC0A | genD0A | genE0A | genF0A | genG0A | genH0A | genI0A]. Ltac gen0B := first [genO0B | genA0B | genB0B | genC0B | genD0B | genE0B | genF0B | genG0B]. Ltac gen1O := first [genO1O | genA1O | genB1O | genC1O | genD1O | genE1O | genF1O | genG1O | genH1O | genI1O]. Ltac gen1A := first [genO1A | genA1A | genB1A | genC1A | genD1A | genE1A | genF1A | genG1A | genH1A | genI1A]. Ltac gen1B := first [genO1B | genA1B | genB1B | genC1B | genD1B | genE1B | genF1B | genG1B]. Ltac gen2O := first [genO2O | genA2O | genB2O | genC2O | genD2O | genE2O | genF2O | genG2O | genH2O | genI2O]. Ltac gen2A := first [genO2A | genA2A | genB2A | genC2A | genD2A | genE2A | genF2A | genG2A | genH2A | genI2A]. Ltac gen2B := first [genO2B | genA2B | genB2B | genC2B | genD2B | genE2B | genF2B | genG2B]. Ltac gen3O := first [genO3O | genA3O | genB3O | genC3O | genD3O | genE3O | genF3O | genG3O | genH3O | genI3O]. Ltac gen3A := first [genO3A | genA3A | genB3A | genC3A | genD3A | genE3A | genF3A | genG3A | genH3A | genI3A]. Ltac gen3B := first [genO3B | genA3B | genB3B | genC3B | genD3B | genE3B | genF3B | genG3B]. Ltac gen4O := first [genO4O | genA4O | genB4O | genC4O | genD4O | genE4O | genF4O | genG4O | genH4O | genI4O]. Ltac gen4A := first [genO4A | genA4A | genB4A | genC4A | genD4A | genE4A | genF4A | genG4A | genH4A | genI4A]. Ltac gen4B := first [genO4B | genA4B | genB4B | genC4B | genD4B | genE4B | genF4B]. Ltac gen5O := first [genO5O | genA5O | genB5O | genC5O | genD5O | genE5O | genF5O | genG5O | genH5O | genI5O]. Ltac gen5A := first [genO5A | genA5A | genB5A | genC5A | genD5A | genE5A | genF5A | genG5A | genH5A | genI5A]. Ltac gen5B := first [genO5B | genA5B | genB5B | genC5B | genD5B | genE5B | genF5B]. Ltac gen6O := first [genO6O | genA6O | genB6O | genC6O | genD6O | genE6O | genF6O | genG6O | genH6O | genI6O]. Ltac gen6A := first [genO6A | genA6A | genB6A | genC6A | genD6A | genE6A | genF6A | genG6A | genH6A | genI6A]. Ltac gen6B := first [genO6B | genA6B | genB6B | genC6B | genD6B | genE6B | genF6B]. Ltac gen7O := first [genO7O | genA7O | genB7O | genC7O | genD7O | genE7O | genF7O | genG7O | genH7O | genI7O]. Ltac gen7A := first [genO7A | genA7A | genB7A | genC7A | genD7A | genE7A | genF7A | genG7A | genH7A | genI7A]. Ltac gen7B := first [genO7B | genA7B | genB7B | genC7B | genD7B | genE7B | genF7B]. Ltac gen8O := first [genO8O | genA8O | genB8O | genC8O | genD8O | genE8O | genF8O | genG8O | genH8O | genI8O]. Ltac gen8A := first [genO8A | genA8A | genB8A | genC8A | genD8A | genE8A | genF8A | genG8A | genH8A | genI8A]. Ltac gen8B := first [genO8B | genA8B | genB8B | genC8B | genD8B | genE8B | genF8B]. Ltac gen9O := first [genO9O | genA9O | genB9O | genC9O | genD9O | genE9O | genF9O | genG9O | genH9O | genI9O]. Ltac gen9A := first [genO9A | genA9A | genB9A | genC9A | genD9A | genE9A | genF9A | genG9A | genH9A | genI9A]. Ltac gen9B := first [genO9B | genA9B | genB9B | genC9B | genD9B | genE9B | genF9B]. Ltac gen0 := first [gen0O | gen0A | gen0B]. Ltac gen1 := first [gen1O | gen1A | gen1B]. Ltac gen2 := first [gen2O | gen2A | gen2B]. Ltac gen3 := first [gen3O | gen3A | gen3B]. Ltac gen4 := first [gen4O | gen4A | gen4B]. Ltac gen5 := first [gen5O | gen5A | gen5B]. Ltac gen6 := first [gen6O | gen6A | gen6B]. Ltac gen7 := first [gen7O | gen7A | gen7B]. Ltac gen8 := first [gen8O | gen8A | gen8B]. Ltac gen9 := first [gen9O | gen9A | gen9B]. Ltac GenO0 := first [genO0O | genO0A |genO0B]. Ltac GenO1 := first [genO1O | genO1A |genO1B]. Ltac GenO2 := first [genO2O | genO2A |genO2B]. Ltac GenO3 := first [genO3O | genO3A |genO3B]. Ltac GenO4 := first [genO4O | genO4A |genO4B]. Ltac GenO5 := first [genO5O | genO5A |genO5B]. Ltac GenO6 := first [genO6O | genO6A |genO6B]. Ltac GenO7 := first [genO7O | genO7A |genO7B]. Ltac GenO8 := first [genO8O | genO8A |genO8B]. Ltac GenO9 := first [genO9O | genO9A |genO9B]. Ltac GenA0 := first [genA0O | genA0A |genA0B]. Ltac GenA1 := first [genA1O | genA1A |genA1B]. Ltac GenA2 := first [genA2O | genA2A |genA2B]. Ltac GenA3 := first [genA3O | genA3A |genA3B]. Ltac GenA4 := first [genA4O | genA4A |genA4B]. Ltac GenA5 := first [genA5O | genA5A |genA5B]. Ltac GenA6 := first [genA6O | genA6A |genA6B]. Ltac GenA7 := first [genA7O | genA7A |genA7B]. Ltac GenA8 := first [genA8O | genA8A |genA8B]. Ltac GenA9 := first [genA9O | genA9A |genA9B]. Ltac GenB0 := first [genB0O | genB0A |genB0B]. Ltac GenB1 := first [genB1O | genB1A |genB1B]. Ltac GenB2 := first [genB2O | genB2A |genB2B]. Ltac GenB3 := first [genB3O | genB3A |genB3B]. Ltac GenB4 := first [genB4O | genB4A |genB4B]. Ltac GenB5 := first [genB5O | genB5A |genB5B]. Ltac GenB6 := first [genB6O | genB6A |genB6B]. Ltac GenB7 := first [genB7O | genB7A |genB7B]. Ltac GenB8 := first [genB8O | genB8A |genB8B]. Ltac GenB9 := first [genB9O | genB9A |genB9B]. Ltac GenC0 := first [genC0O | genC0A |genC0B]. Ltac GenC1 := first [genC1O | genC1A |genC1B]. Ltac GenC2 := first [genC2O | genC2A |genC2B]. Ltac GenC3 := first [genC3O | genC3A |genC3B]. Ltac GenC4 := first [genC4O | genC4A |genC4B]. Ltac GenC5 := first [genC5O | genC5A |genC5B]. Ltac GenC6 := first [genC6O | genC6A |genC6B]. Ltac GenC7 := first [genC7O | genC7A |genC7B]. Ltac GenC8 := first [genC8O | genC8A |genC8B]. Ltac GenC9 := first [genC9O | genC9A |genC9B]. Ltac GenD0 := first [genD0O | genD0A |genD0B]. Ltac GenD1 := first [genD1O | genD1A |genD1B]. Ltac GenD2 := first [genD2O | genD2A |genD2B]. Ltac GenD3 := first [genD3O | genD3A |genD3B]. Ltac GenD4 := first [genD4O | genD4A |genD4B]. Ltac GenD5 := first [genD5O | genD5A |genD5B]. Ltac GenD6 := first [genD6O | genD6A |genD6B]. Ltac GenD7 := first [genD7O | genD7A |genD7B]. Ltac GenD8 := first [genD8O | genD8A |genD8B]. Ltac GenD9 := first [genD9O | genD9A |genD9B]. Ltac GenE0 := first [genE0O | genE0A |genE0B]. Ltac GenE1 := first [genE1O | genE1A |genE1B]. Ltac GenE2 := first [genE2O | genE2A |genE2B]. Ltac GenE3 := first [genE3O | genE3A |genE3B]. Ltac GenE4 := first [genE4O | genE4A |genE4B]. Ltac GenE5 := first [genE5O | genE5A |genE5B]. Ltac GenE6 := first [genE6O | genE6A |genE6B]. Ltac GenE7 := first [genE7O | genE7A |genE7B]. Ltac GenE8 := first [genE8O | genE8A |genE8B]. Ltac GenE9 := first [genE9O | genE9A |genE9B]. Ltac GenF0 := first [genF0O | genF0A |genF0B]. Ltac GenF1 := first [genF1O | genF1A |genF1B]. Ltac GenF2 := first [genF2O | genF2A |genF2B]. Ltac GenF3 := first [genF3O | genF3A |genF3B]. Ltac GenF4 := first [genF4O | genF4A |genF4B]. Ltac GenF5 := first [genF5O | genF5A |genF5B]. Ltac GenF6 := first [genF6O | genF6A |genF6B]. Ltac GenF7 := first [genF7O | genF7A |genF7B]. Ltac GenF8 := first [genF8O | genF8A |genF8B]. Ltac GenF9 := first [genF9O | genF9A |genF9B]. Ltac GenG0 := first [genG0O | genG0A |genG0B]. Ltac GenG1 := first [genG1O | genG1A |genG1B]. Ltac GenG2 := first [genG2O | genG2A |genG2B]. Ltac GenG3 := first [genG3O | genG3A |genG3B]. Ltac GenG4 := first [genG4O | genG4A]. Ltac GenG5 := first [genG5O | genG5A]. Ltac GenG6 := first [genG6O | genG6A]. Ltac GenG7 := first [genG7O | genG7A]. Ltac GenG8 := first [genG8O | genG8A]. Ltac GenG9 := first [genG9O | genG9A]. Ltac GenH0 := first [genH0O | genH0A]. Ltac GenH1 := first [genH1O | genH1A]. Ltac GenH2 := first [genH2O | genH2A]. Ltac GenH3 := first [genH3O | genH3A]. Ltac GenH4 := first [genH4O | genH4A]. Ltac GenH5 := first [genH5O | genH5A]. Ltac GenH6 := first [genH6O | genH6A]. Ltac GenH7 := first [genH7O | genH7A]. Ltac GenH8 := first [genH8O | genH8A]. Ltac GenH9 := first [genH9O | genH9A]. Ltac GenI0 := first [genI0O | genI0A]. Ltac GenI1 := first [genI1O | genI1A]. Ltac GenI2 := first [genI2O | genI2A]. Ltac GenI3 := first [genI3O | genI3A]. Ltac GenI4 := first [genI4O | genI4A]. Ltac GenI5 := first [genI5O | genI5A]. Ltac GenI6 := first [genI6O | genI6A]. Ltac GenI7 := first [genI7O | genI7A]. Ltac GenI8 := first [genI8O | genI8A]. Ltac GenI9 := first [genI9O | genI9A]. Ltac Gen0 := first [GenO0 | GenA0 | GenB0 | GenC0 | GenD0 | GenE0 | GenF0 | GenG0 | GenH0 | GenI0]. Ltac Gen1 := first [GenO1 | GenA1 | GenB1 | GenC1 | GenD1 | GenE1 | GenF1 | GenG1 | GenH1 | GenI1]. Ltac Gen2 := first [GenO2 | GenA2 | GenB2 | GenC2 | GenD2 | GenE2 | GenF2 | GenG2 | GenH2 | GenI2]. Ltac Gen3 := first [GenO3 | GenA3 | GenB3 | GenC3 | GenD3 | GenE3 | GenF3 | GenG3 | GenH3 | GenI3]. Ltac Gen4 := first [GenO4 | GenA4 | GenB4 | GenC4 | GenD4 | GenE4 | GenF4 | GenG4 | GenH4 | GenI4]. Ltac Gen5 := first [GenO5 | GenA5 | GenB5 | GenC5 | GenD5 | GenE5 | GenF5 | GenG5 | GenH5 | GenI5]. Ltac Gen6 := first [GenO6 | GenA6 | GenB6 | GenC6 | GenD6 | GenE6 | GenF6 | GenG6 | GenH6 | GenI6]. Ltac Gen7 := first [GenO7 | GenA7 | GenB7 | GenC7 | GenD7 | GenE7 | GenF7 | GenG7 | GenH7 | GenI7]. Ltac Gen8 := first [GenO8 | GenA8 | GenB8 | GenC8 | GenD8 | GenE8 | GenF8 | GenG8 | GenH8 | GenI8]. Ltac Gen9 := first [GenO9 | GenA9 | GenB9 | GenC9 | GenD9 | GenE9 | GenF9 | GenG9 | GenH9 | GenI9]. (*You get the idea, the next part of adding the appropriate amount of implications to these templates to corresopnd with the last character will be left as a ritual of sleep-deprivation and eager commits! I don't want to write a script to do this.*) Ltac genG4B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genG5B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genG6B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genG7B:= match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genG8B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genG9B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genH0B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genH1B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genH2B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genH3B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genH4B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genH5B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genH6B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genH7B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genH8B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genH9B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genI0B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genI1B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genI2B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genI3B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genI4B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genI5B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genI6B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genI7B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genI8B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genI9B := match goal with | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _] => generalize h | [|- _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO0C := match goal with | [|- _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ ?h -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO1C := match goal with | [|- _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO2C := match goal with | [|- _ _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO3C := match goal with | [|- _ _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO4C := match goal with | [|- _ _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO5C := match goal with | [|- _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO6C := match goal with | [|- _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO7C := match goal with | [|- _ _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO8C := match goal with | [|- _ _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genO9C := match goal with | [|- _ _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA0C := match goal with | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA1C := match goal with | [|- _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA2C := match goal with | [|- _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA3C := match goal with | [|- _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA4C := match goal with | [|- _ -> _ _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA5C := match goal with | [|- _ -> _ _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ _ ?h _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA6C := match goal with | [|- _ -> _ _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ _ ?h _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA7C := match goal with | [|- _ -> _ _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ _ ?h _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA8C := match goal with | [|- _ -> _ _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ _ ?h _ _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genA9C := match goal with | [|- _ -> _ ?h _ _ _ _ _ _ _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ ?h _ _ _ _ _ _ _ _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB0C := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB1C := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB2C := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ -> _ -> _ -> _] => generalize h end. (*Wyckoff*) Ltac genB3C := match goal with | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ = _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ _ ?h _ _ _ -> _ -> _ -> _] => generalize h | [|- _ -> _ -> _ ?h _ _ _ -> _ -> _ -> _] => generalize h end. (*SF*) Inductive Boxer : Type := | boxer : forall {A:Type}, A -> Boxer. (*Wyckoff*) Definition unbox (b:Boxer) : match b with | boxer T _ => T end:= match b with | boxer _ x => x end. Require Import List. (*SF*) Notation "'>>'" := (@nil Boxer) (at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1" := ((boxer v1)::nil) (at level 0, v1 at level 0). (*e : ltac_scope.*) (*SF*) Notation "'>>' v1 v2" := ((boxer v1)::(boxer v2)::nil) (at level 0, v1 at level 0, v2 at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1 v2 v3" := ((boxer v1)::(boxer v2)::(boxer v3)::nil) (at level 0, v1 at level 0, v2 at level 0, v3 at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1 v2 v3 v4" := ((boxer v1)::(boxer v2)::(boxer v3)::(boxer v4)::nil) (at level 0, v1 at level 0, v2 at level 0, v3 at level 0, v4 at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1 v2 v3 v4 v5" := ((boxer v1)::(boxer v2)::(boxer v3)::(boxer v4)::(boxer v5)::nil) (at level 0, v1 at level 0, v2 at level 0, v3 at level 0, v4 at level 0, v5 at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1 v2 v3 v4 v5 v6" := ((boxer v1)::(boxer v2)::(boxer v3)::(boxer v4)::(boxer v5) ::(boxer v6)::nil) (at level 0, v1 at level 0, v2 at level 0, v3 at level 0, v4 at level 0, v5 at level 0, v6 at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1 v2 v3 v4 v5 v6 v7" := ((boxer v1)::(boxer v2)::(boxer v3)::(boxer v4)::(boxer v5) ::(boxer v6)::(boxer v7)::nil) (at level 0, v1 at level 0, v2 at level 0, v3 at level 0, v4 at level 0, v5 at level 0, v6 at level 0, v7 at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1 v2 v3 v4 v5 v6 v7 v8" := ((boxer v1)::(boxer v2)::(boxer v3)::(boxer v4)::(boxer v5) ::(boxer v6)::(boxer v7)::(boxer v8)::nil) (at level 0, v1 at level 0, v2 at level 0, v3 at level 0, v4 at level 0, v5 at level 0, v6 at level 0, v7 at level 0, v8 at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1 v2 v3 v4 v5 v6 v7 v8 v9" := ((boxer v1)::(boxer v2)::(boxer v3)::(boxer v4)::(boxer v5) ::(boxer v6)::(boxer v7)::(boxer v8)::(boxer v9)::nil) (at level 0, v1 at level 0, v2 at level 0, v3 at level 0, v4 at level 0, v5 at level 0, v6 at level 0, v7 at level 0, v8 at level 0, v9 at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1 v2 v3 v4 v5 v6 v7 v8 v9 v10" := ((boxer v1)::(boxer v2)::(boxer v3)::(boxer v4)::(boxer v5) ::(boxer v6)::(boxer v7)::(boxer v8)::(boxer v9)::(boxer v10)::nil) (at level 0, v1 at level 0, v2 at level 0, v3 at level 0, v4 at level 0, v5 at level 0, v6 at level 0, v7 at level 0, v8 at level 0, v9 at level 0, v10 at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1 v2 v3 v4 v5 v6 v7 v8 v9 v10 v11" := ((boxer v1)::(boxer v2)::(boxer v3)::(boxer v4)::(boxer v5) ::(boxer v6)::(boxer v7)::(boxer v8)::(boxer v9)::(boxer v10) ::(boxer v11)::nil) (at level 0, v1 at level 0, v2 at level 0, v3 at level 0, v4 at level 0, v5 at level 0, v6 at level 0, v7 at level 0, v8 at level 0, v9 at level 0, v10 at level 0, v11 at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1 v2 v3 v4 v5 v6 v7 v8 v9 v10 v11 v12" := ((boxer v1)::(boxer v2)::(boxer v3)::(boxer v4)::(boxer v5) ::(boxer v6)::(boxer v7)::(boxer v8)::(boxer v9)::(boxer v10) ::(boxer v11)::(boxer v12)::nil) (at level 0, v1 at level 0, v2 at level 0, v3 at level 0, v4 at level 0, v5 at level 0, v6 at level 0, v7 at level 0, v8 at level 0, v9 at level 0, v10 at level 0, v11 at level 0, v12 at level 0). (* : ltac_scope.*) (*SF*) Notation "'>>' v1 v2 v3 v4 v5 v6 v7 v8 v9 v10 v11 v12 v13" := ((boxer v1)::(boxer v2)::(boxer v3)::(boxer v4)::(boxer v5) ::(boxer v6)::(boxer v7)::(boxer v8)::(boxer v9)::(boxer v10) ::(boxer v11)::(boxer v12)::(boxer v13)::nil) (at level 0, v1 at level 0, v2 at level 0, v3 at level 0, v4 at level 0, v5 at level 0, v6 at level 0, v7 at level 0, v8 at level 0, v9 at level 0, v10 at level 0, v11 at level 0, v12 at level 0, v13 at level 0). (* : ltac_scope.*) (*Wyckoff*) Ltac larg_list := match goal with | |- ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 ?x12 ?x13 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) | |- ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 ?x12 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) | |- ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11) | |- ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10) | |- ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9) | |- ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8) | |- ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7) | |- ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 = _ => constr:(>> x1 x2 x3 x4 x5 x6) | |- ?f ?x1 ?x2 ?x3 ?x4 ?x5 = _ => constr:(>> x1 x2 x3 x4 x5) | |- ?f ?x1 ?x2 ?x3 ?x4 = _ => constr:(>> x1 x2 x3 x4) | |- ?f ?x1 ?x2 ?x3 = _ => constr:(>> x1 x2 x3) | |- ?f ?x1 ?x2 = _ => constr:(>> x1 x2) | |- ?f ?x1 = _ => constr:(>> x1) | |- _ => idtac end. (*Wyckoff*) Ltac rarg_list := match goal with | |- _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 ?x12 ?x13 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) | |- _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 ?x12 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) | |- _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11) | |- _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10) | |- _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9) | |- _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8) | |- _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 => constr:(>> x1 x2 x3 x4 x5 x6 x7) | |- _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 => constr:(>> x1 x2 x3 x4 x5 x6) | |- _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 => constr:(>> x1 x2 x3 x4 x5) | |- _ = ?f ?x1 ?x2 ?x3 ?x4 => constr:(>> x1 x2 x3 x4) | |- _ = ?f ?x1 ?x2 ?x3 => constr:(>> x1 x2 x3) | |- _ = ?f ?x1 ?x2 => constr:(>> x1 x2) | |- _ = ?f ?x1 => constr:(>> x1) | |- _ => idtac end. (*Wyckoff*) Ltac larg_list_h h := let P := type of h in match P with | ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 ?x12 ?x13 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) | ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 ?x12 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) | ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11) | ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10) | ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9) | ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8) | ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 = _ => constr:(>> x1 x2 x3 x4 x5 x6 x7) | ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 = _ => constr:(>> x1 x2 x3 x4 x5 x6) | ?f ?x1 ?x2 ?x3 ?x4 ?x5 = _ => constr:(>> x1 x2 x3 x4 x5) | ?f ?x1 ?x2 ?x3 ?x4 = _ => constr:(>> x1 x2 x3 x4) | ?f ?x1 ?x2 ?x3 = _ => constr:(>> x1 x2 x3) | ?f ?x1 ?x2 = _ => constr:(>> x1 x2) | ?f ?x1 = _ => constr:(>> x1) | _ => idtac end. (*Wyckoff*) Ltac rarg_list_h h := let P := type of h in match P with | _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 ?x12 ?x13 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) | _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 ?x12 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12) | _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 ?x11 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11) | _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 ?x10 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9 x10) | _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 ?x9 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8 x9) | _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 ?x8 => constr:(>> x1 x2 x3 x4 x5 x6 x7 x8) | _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 ?x7 => constr:(>> x1 x2 x3 x4 x5 x6 x7) | _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 ?x6 => constr:(>> x1 x2 x3 x4 x5 x6) | _ = ?f ?x1 ?x2 ?x3 ?x4 ?x5 => constr:(>> x1 x2 x3 x4 x5) | _ = ?f ?x1 ?x2 ?x3 ?x4 => constr:(>> x1 x2 x3 x4) | _ = ?f ?x1 ?x2 ?x3 => constr:(>> x1 x2 x3) | _ = ?f ?x1 ?x2 => constr:(>> x1 x2) | _ = ?f ?x1 => constr:(>> x1) | _ => idtac end. (*SF*) Ltac get_fun_arg E := match E with | ?X1 ?X2 ?X3 ?X4 ?X5 ?X6 ?X7 ?X8 ?X9 ?X10 ?X11 ?X => constr:((X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11,X)) | ?X1 ?X2 ?X3 ?X4 ?X5 ?X6 ?X7 ?X8 ?X9 ?X10 ?X => constr:((X1 X2 X3 X4 X5 X6 X7 X8 X9 X10,X)) | ?X1 ?X2 ?X3 ?X4 ?X5 ?X6 ?X7 ?X8 ?X9 ?X => constr:((X1 X2 X3 X4 X5 X6 X7 X8 X9,X)) | ?X1 ?X2 ?X3 ?X4 ?X5 ?X6 ?X7 ?X8 ?X => constr:((X1 X2 X3 X4 X5 X6 X7 X8,X)) | ?X1 ?X2 ?X3 ?X4 ?X5 ?X6 ?X7 ?X => constr:((X1 X2 X3 X4 X5 X6 X7,X)) | ?X1 ?X2 ?X3 ?X4 ?X5 ?X6 ?X => constr:((X1 X2 X3 X4 X5 X6,X)) | ?X1 ?X2 ?X3 ?X4 ?X5 ?X => constr:((X1 X2 X3 X4 X5,X)) | ?X1 ?X2 ?X3 ?X4 ?X => constr:((X1 X2 X3 X4,X)) | ?X1 ?X2 ?X3 ?X => constr:((X1 X2 X3,X)) | ?X1 ?X2 ?X => constr:((X1 X2,X)) | ?X1 ?X => constr:((X1,X)) end. (*Wyckoff*) (*This function just "pose proof"s a term with itself. Used in conjunction with the one next below . . . *) Ltac eq_refl_t x := let h := fresh "he" in pose proof (eq_refl x) as h. (*Wyckoff*) (*. . . which unfolds a global constant in an equality proof like [he] (derived from the above function, so that Ltac can return a list of arguments in a function application folded globally*) Ltac unfold_larg h x := let la' := fresh "la" in unfold x in h; let larg := larg_list_h h in pose larg as la'. (*Example:say you have a term [x := P a1 a2 a3 a4]. Enter > eq_refl_t x. to generate a term [he:x = x] Then type [unfold_larg he x], to get the Boxer list >> a1 a2 a3 a4. *) (*SF*) Ltac find_head_match T := match T with context [?E] => match T with | E => fail 1 | _ => constr:(E) end end. (*SF*) Ltac false_post := solve [ assumption | discriminate | congruence ]. (*SF*) Tactic Notation "false_goal" := elimtype False. (*SF*) Tactic Notation "false" := false_goal; try false_post. (*SF*) Tactic Notation "tryfalse" := try solve [ false ]. (*SF*) Ltac destruct_if_post := tryfalse. (*SF*) Ltac destruct_head_match_core cont := match goal with | |- ?T1 = ?T2 => first [ let E := find_head_match T1 in cont E | let E := find_head_match T2 in cont E ] | |- ?T1 => let E := find_head_match T1 in cont E end; destruct_if_post. (*SF*) Tactic Notation "destruct_head_match" "as" simple_intropattern(I) := destruct_head_match_core ltac:(fun E => destruct E as I). (*SF*) Tactic Notation "destruct_head_match" := destruct_head_match_core ltac:(fun E => destruct E). (*End Tactics.*) (*Top Level*)bool2-v0-3/FieldsOfSets.v0000644000175000017500000024256113575574055016075 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014, Daniel Wyckoff*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (* Version 0.3 bool2 -- certified -- Coq 8.4 pl4 *) Require Export BoolAlgBasics. Require Import InfiniteOperations. Require Import SetUtilities. Require Import TypeUtilities. Require Export Powerset. Require Import ProofIrrelevanceFacts. Require Import ProofIrrelevance. Require Import Description. Require Import QArith. Require Import ArithUtilities. Require Import Equality. Require Import DecidableDec. Require Import LogicUtilities. Require Import FunctionProperties. Require Import NaryFunctions. Section power_set_sec. Variable Xt : Type. Definition X := Full_set Xt. Definition PX := Power_set _ X. (* PSA = Power Set Algebra. *) (* This definition just builds the Bconst component of the Boolean Algebra.*) Definition bc_psa := Build_Bconst (Ensemble Xt) PX (@Union Xt) (@Intersection Xt) X (Empty_set Xt) (@Comp Xt). Lemma full_power_compat : Full_set (Ensemble Xt) = PX. assert (h1:forall Y:(Ensemble Xt), Included Y X -> In PX Y). unfold PX. apply Definition_of_Power_set. assert (h2:forall Y:(Ensemble Xt), Included Y X). intro Y. unfold X. unfold Included. intros x h. apply (Full_intro Xt). assert (h3:forall Y:Ensemble Xt, In PX Y). intro Y. apply h1. apply h2. apply Extensionality_Ensembles. unfold Same_set. split. unfold Included. intros S h4. apply (h3 S). unfold Included. intros S h5. apply Full_intro. Qed. Lemma full_power_compat' : BS bc_psa = Full_set (Btype bc_psa). assert (h1: Btype bc_psa = Ensemble Xt). unfold bc_psa. reflexivity. assert (h2: BS bc_psa = PX). unfold bc_psa. reflexivity. rewrite h2. assert (h3:Full_set (Btype bc_psa) = Full_set (Ensemble Xt)). unfold bc_psa. reflexivity. rewrite h3. rewrite full_power_compat. reflexivity. Qed. Definition psa := Build_Bool_Alg bc_psa full_power_compat' (@Union_associative Xt) (@assoc_prod_psa Xt) (@Union_commutative Xt) (@Intersection_commutative Xt) (@abs_sum_psa Xt) (@abs_prod_psa Xt) (@Distributivity Xt) (@dist_prod_psa Xt) (@comp_sum_psa Xt) (@comp_prod_psa Xt). Lemma btype_bc_psa : bt psa = Ensemble Xt. simpl. reflexivity. Qed. Definition transfer_psa (P:bt psa) : Ensemble Xt := transfer btype_bc_psa P. End power_set_sec. Section field_of_sets_sec. Record Field_of_Sets := { Xt:Type; F:(Ensemble (Ensemble Xt)); non_empty_F : (exists S:(Ensemble Xt), In F S); Union_closed : forall S1 S2:(Ensemble Xt), In F S1 -> In F S2 -> In F (Union S1 S2); Int_closed : forall S1 S2:(Ensemble Xt), In F S1 -> In F S2 -> In F (Intersection S1 S2); Comp_closed : forall S :(Ensemble Xt), In F S -> In F (Comp S); Ft := {S:(Ensemble Xt) | In F S}; Union_fos := fun (S1 S2:Ft) => (let S1' := proj1_sig S1 in let S2' := proj1_sig S2 in exist (fun (S':(Ensemble Xt)) => (In F S')) (Union S1' S2') (Union_closed S1' S2' (proj2_sig S1) (proj2_sig S2))); Int_fos := fun (S1 S2:Ft) => (let S1' := proj1_sig S1 in let S2' := proj1_sig S2 in exist (fun (S':(Ensemble Xt)) => (In F S')) (Intersection S1' S2') (Int_closed S1' S2' (proj2_sig S1) (proj2_sig S2))); Comp_fos := fun (S:Ft) => (let S' := proj1_sig S in exist (fun (S'':(Ensemble Xt)) => (In F S'')) (Comp S') (Comp_closed S' (proj2_sig S))) }. Lemma f_fos_inj : forall (fos fos':Field_of_Sets) (pf:Xt fos = Xt fos'), (@transfer_dep _ (fun T=>Ensemble (Ensemble T)) _ _ pf (F fos)) = F fos' -> fos = fos'. intros fos fos' h1 h2. destruct fos; destruct fos'. simpl in h1. subst. simpl in h2. subst. f_equal; apply proof_irrelevance. Qed. Variable fos:Field_of_Sets. Lemma full_in_F : In (F fos) (Full_set (Xt fos)). elim (non_empty_F fos). intros S h1. assert (h2: In (F fos) (Comp S)). apply ((Comp_closed fos) S h1). assert (h3: In (F fos) (Union S (Comp S))). apply ((Union_closed fos) S (Comp S) h1 h2). assert (h4: (Union S (Comp S)) = (Full_set (Xt fos))). apply comp_sum_psa. rewrite h4 in h3. assumption. Qed. Lemma empty_in_F : In (F fos) (Empty_set (Xt fos)). elim (non_empty_F fos). intros S h1. assert (h2: In (F fos) (Comp S)). apply ((Comp_closed fos) S h1). assert (h3: In (F fos) (Intersection S (Comp S))). apply ((Int_closed fos) S (Comp S) h1 h2). assert (h4: (Intersection S (Comp S)) = (Empty_set (Xt fos))). apply comp_prod_psa. rewrite h4 in h3. assumption. Qed. Lemma assoc_sum_fos : forall N M P:(Ft fos), (Union_fos fos N (Union_fos fos M P)) = (Union_fos fos (Union_fos fos N M) P). intros N M P. unfold Union_fos; simpl. assert (h:Union (proj1_sig N) (Union (proj1_sig M) (proj1_sig P)) = Union (Union (proj1_sig N) (proj1_sig M)) (proj1_sig P)). apply Union_associative. apply existTexist. apply subsetT_eq_compat. apply h. Qed. Lemma assoc_prod_fos : forall N M P:(Ft fos), (Int_fos fos N (Int_fos fos M P)) = (Int_fos fos (Int_fos fos N M) P). intros N M P. unfold Int_fos; simpl. assert (h:(Intersection (proj1_sig N) (Intersection (proj1_sig M) (proj1_sig P))) = (Intersection (Intersection (proj1_sig N) (proj1_sig M)) (proj1_sig P))). apply assoc_prod_psa. apply existTexist. apply subsetT_eq_compat. apply h. Qed. Lemma comm_sum_fos : forall N M:(Ft fos), Union_fos fos N M = Union_fos fos M N. intros N M. unfold Union_fos; simpl. assert (h: (Union (proj1_sig N) (proj1_sig M)) = (Union (proj1_sig M) (proj1_sig N))). apply Union_commutative. apply existTexist. apply subsetT_eq_compat. apply h. Qed. Lemma comm_prod_fos : forall N M:(Ft fos), Int_fos fos N M = Int_fos fos M N. intros N M. unfold Int_fos; simpl. assert (h: (Intersection (proj1_sig N) (proj1_sig M)) = (Intersection (proj1_sig M) (proj1_sig N))). apply Intersection_commutative. apply existTexist. apply subsetT_eq_compat. apply h. Qed. Lemma abs_sum_fos : forall N M:(Ft fos), Union_fos fos N (Int_fos fos N M) = N. intros N M. unfold Union_fos; unfold Int_fos; simpl. assert (h1: (Union (proj1_sig N) (Intersection (proj1_sig N) (proj1_sig M))) = (proj1_sig N)). apply abs_sum_psa. assert (h2: N = exist (fun S' : Ensemble (Xt fos) => In (F fos) S') (proj1_sig N) (proj2_sig N)). destruct N; simpl. reflexivity. rewrite h2 at 8. apply existTexist. apply subsetT_eq_compat. apply h1. Qed. Lemma abs_prod_fos : forall N M:(Ft fos), Int_fos fos N (Union_fos fos N M) = N. intros N M. unfold Union_fos; unfold Int_fos; simpl. assert (h1: (Intersection (proj1_sig N) (Union (proj1_sig N) (proj1_sig M))) = (proj1_sig N)). apply abs_prod_psa. assert (h2: N = exist (fun S' : Ensemble (Xt fos) => In (F fos) S') (proj1_sig N) (proj2_sig N)). destruct N; simpl. reflexivity. rewrite h2 at 8. apply existTexist. apply subsetT_eq_compat. apply h1. Qed. Lemma dist_sum_fos : forall N M P:(Ft fos), Int_fos fos P (Union_fos fos N M) = Union_fos fos (Int_fos fos P N) (Int_fos fos P M). intros N M P. unfold Union_fos; unfold Int_fos; simpl. assert (h: (Intersection (proj1_sig P) (Union (proj1_sig N) (proj1_sig M))) = (Union (Intersection (proj1_sig P) (proj1_sig N)) (Intersection (proj1_sig P) (proj1_sig M)))). apply Distributivity. apply existTexist. apply subsetT_eq_compat. apply h. Qed. Lemma dist_prod_fos : forall N M P:(Ft fos), Union_fos fos P (Int_fos fos N M) = Int_fos fos (Union_fos fos P N) (Union_fos fos P M). intros N M P. unfold Union_fos; unfold Int_fos; simpl. assert (h: (Union (proj1_sig P) (Intersection (proj1_sig N) (proj1_sig M))) = (Intersection (Union (proj1_sig P) (proj1_sig N)) (Union (proj1_sig P) (proj1_sig M)))). apply dist_prod_psa. apply existTexist. apply subsetT_eq_compat. apply h. Qed. Definition Full_fos := exist (fun S':(Ensemble (Xt fos)) => (In (F fos) S')) (Full_set (Xt fos)) full_in_F. Lemma comp_sum_fos: forall N:(Ft fos), Union_fos fos N (Comp_fos fos N) = Full_fos. intro N. unfold Union_fos; unfold Comp_fos; unfold Full_fos; simpl. assert (h1: (Union (proj1_sig N) (Comp (proj1_sig N))) = (Full_set (Xt fos))). apply comp_sum_psa. apply existTexist. apply subsetT_eq_compat. assumption. Qed. Definition Empty_fos := exist (fun S':(Ensemble (Xt fos)) => (In (F fos) S')) (Empty_set (Xt fos)) empty_in_F. Lemma comp_prod_fos: forall N:(Ft fos), Int_fos fos N (Comp_fos fos N) = Empty_fos. intro N. unfold Int_fos; unfold Comp_fos; unfold Empty_fos; simpl. assert (h1: (Intersection (proj1_sig N) (Comp (proj1_sig N))) = (Empty_set (Xt fos))). apply comp_prod_psa. apply existTexist. apply subsetT_eq_compat. assumption. Qed. Definition Bc_fos := Build_Bconst (Ft fos) (Full_set (Ft fos)) (Union_fos fos) (Int_fos fos) Full_fos Empty_fos (Comp_fos fos). Lemma fos_full_compat : BS Bc_fos = Full_set (Btype Bc_fos). unfold Bc_fos; simpl. reflexivity. Qed. Definition fos_ba := Build_Bool_Alg Bc_fos fos_full_compat assoc_sum_fos assoc_prod_fos comm_sum_fos comm_prod_fos abs_sum_fos abs_prod_fos dist_sum_fos dist_prod_fos comp_sum_fos comp_prod_fos. Lemma fos_psa_ex : forall (T:Type), exists! fos:Field_of_Sets, (Xt fos) = T /\ (F fos) = power_set (Full_set (Xt fos)). intros T. pose (power_set (Full_set T)) as F'. assert (my_non_empty_F : exists S : Ensemble T, In F' S). exists (Empty_set _). constructor. auto with sets. assert (my_Union_closed : forall S1 S2 : Ensemble T, In F' S1 -> In F' S2 -> In F' (Union S1 S2)). intros. constructor. red. intros. constructor. assert (my_Int_closed : forall S1 S2 : Ensemble T, In F' S1 -> In F' S2 -> In F' (Intersection S1 S2)). intros. constructor. red. intros. constructor. assert (my_Comp_closed : forall S : Ensemble T, In F' S -> In F' (Comp S)). intros. constructor. red. intros. constructor. pose (Build_Field_of_Sets T F' my_non_empty_F my_Union_closed my_Int_closed my_Comp_closed) as fos'. exists fos'. simpl in fos'. red; split; try split; simpl; try reflexivity. intros B h1. destruct h1 as [h1 h2]. simpl. assert (h3:T = Xt fos'). reflexivity. rewrite h3 in h1. symmetry in h1. assert (h5:@transfer_dep _ (fun T => Ensemble (Ensemble T)) (Xt fos') (Xt B) h1 (F fos') = F B). simpl. unfold F'. pose (fun t:{T:Type & (Ensemble (Ensemble T))} => projT2 t = power_set (Full_set (projT1 t))) as P. pose proof (transfer_dep_prop _ _ h1 (F fos') P) as h4. unfold P in h4. simpl in h4. rewrite <- h2 in h4. unfold F' in h4. rewrite <- h4. reflexivity. eapply f_fos_inj. apply h5. Qed. Definition fos_psa (T:Type) := proj1_sig (constructive_definite_description _ (fos_psa_ex T)). Lemma fos_psa_compat : forall (T:Type), let fos:=fos_psa T in (Xt fos) = T /\ (F fos) = power_set (Full_set (Xt fos)). intros T Fos. unfold Fos. unfold fos_psa. destruct constructive_definite_description as [Fos' h1]. simpl. assumption. Qed. Definition proj1_sig_fos_psa {T:Type} (A:{S:Ensemble (Xt (fos_psa T)) | In (F (fos_psa T)) S}) : Ensemble T. destruct A as [A h1]. pose proof (fos_psa_compat T) as [h2 h3]. clear h1. rewrite h2 in A. refine A. Defined. Definition proj1_sig_fos_psa_compat : forall {T:Type} (A:{S:Ensemble (Xt (fos_psa T)) | In (F (fos_psa T)) S}), let A' := proj1_sig_fos_psa A in A' = transfer_dep match (fos_psa_compat T) with | conj P _ => P end (proj1_sig A). intros T A A'. unfold A'. symmetry. rewrite <- transfer_dep_eq_iff. unfold proj1_sig_fos_psa. destruct (fos_psa_compat T) as [h1 h2]. destruct A as [A h3]. simpl. destruct h1. simpl. reflexivity. Qed. Lemma proj1_sig_fos_psa_compat' : forall {T:Type} (A:{S:Ensemble (Xt (fos_psa T)) | In (F (fos_psa T)) S}), existT Ensemble (Xt (fos_psa T)) (proj1_sig A) = existT Ensemble T (proj1_sig_fos_psa A). intros T A. pose proof (proj1_sig_fos_psa_compat A) as h1. symmetry in h1. rewrite <- transfer_dep_eq_iff in h1. assumption. Qed. Definition exist_fos_psa {T:Type} (A:Ensemble T) : {S:Ensemble (Xt (fos_psa T)) |In (F (fos_psa T)) S}. pose proof (fos_psa_compat T) as [h2 h3]. rewrite <- h2 in A. pose proof (full_inclusion A) as h4. rewrite h3. assert (h5:In (power_set (Full_set (Xt (fos_psa T)))) A). constructor. red; intros; constructor. refine (exist _ A h5). Defined. Lemma exist_fos_psa_compat : forall {T:Type} (A:Ensemble T), let A' := exist_fos_psa A in proj1_sig A' = transfer_dep_r match (fos_psa_compat T) with | conj P _ => P end A. intros T A A'. destruct (fos_psa_compat T) as [h1 h2]. pose proof h1 as h1'. symmetry in h1'. rewrite <- (transfer_dep_transfer_dep_r_compat h1' h1). symmetry. rewrite <- transfer_dep_eq_iff. unfold A'. unfold exist_fos_psa. destruct (fos_psa_compat T) as [h3 h4]. unfold eq_rect_r. rewrite h4. simpl. rewrite h3. simpl. reflexivity. Qed. Lemma exist_fos_psa_compat' : forall {T:Type} (A:Ensemble T), existT Ensemble (Xt (fos_psa T)) (proj1_sig (exist_fos_psa A)) = existT Ensemble T A. intros T A. unfold exist_fos_psa. destruct (fos_psa_compat T) as [h1 h2]. rewrite h2, h1. unfold eq_rect_r. rewrite <- eq_rect_eq. simpl. reflexivity. Qed. Lemma proj1_sig_exist_fos_psa_int : forall {T:Type} (A B:Ensemble T), proj1_sig (exist_fos_psa (Intersection A B)) = Intersection (proj1_sig (exist_fos_psa A)) (proj1_sig (exist_fos_psa B)). intros T A B. unfold exist_fos_psa. destruct (fos_psa_compat T) as [h1 h2]. unfold eq_rect_r. rewrite h2. do 3 rewrite <- eq_rect_eq. simpl. rewrite h1. do 3 rewrite <- eq_rect_eq. reflexivity. Qed. Lemma proj1_sig_exist_fos_psa_union : forall {T:Type} (A B:Ensemble T), proj1_sig (exist_fos_psa (Union A B)) = Union (proj1_sig (exist_fos_psa A)) (proj1_sig (exist_fos_psa B)). intros T A B. unfold exist_fos_psa. destruct (fos_psa_compat T) as [h1 h2]. unfold eq_rect_r. rewrite h2. do 3 rewrite <- eq_rect_eq. simpl. rewrite h1. do 3 rewrite <- eq_rect_eq. reflexivity. Qed. Lemma proj1_sig_exist_fos_psa_comp : forall {T:Type} (A:Ensemble T), proj1_sig (exist_fos_psa (Comp A)) = Comp (proj1_sig (exist_fos_psa A)). intros T A. unfold exist_fos_psa. destruct (fos_psa_compat T) as [h1 h2]. unfold eq_rect_r. rewrite h2. do 2 rewrite <- eq_rect_eq. simpl. rewrite h1. do 2 rewrite <- eq_rect_eq. reflexivity. Qed. Lemma in_proj1_sig_exist_fos_psa_iff : forall {T:Type} (A:Ensemble T) (x:Xt (fos_psa T)), In (proj1_sig (exist_fos_psa A)) x <-> In A (transfer (match (fos_psa_compat T) with | conj pf _ => pf end) x). intros T A x. destruct (fos_psa_compat T) as [h1l h1r]. rewrite <- (transfer_dep_undoes_transfer_dep_r h1l A) at 2. rewrite <- transfer_in. rewrite exist_fos_psa_compat. destruct (fos_psa_compat T) as [h2 h3]. pose proof (proof_irrelevance _ h2 h1l) as h4; rewrite h4; tauto. Qed. Lemma in_proj1_sig_exist_fos_psa_iff_r : forall {T:Type} (A:Ensemble T) (x:T), In (proj1_sig (exist_fos_psa A)) (transfer_r (match (fos_psa_compat T) with | conj pf _ => pf end) x) <-> In A x. intros T A x. pose proof (transfer_in_r (match (fos_psa_compat T) with |conj pf _ => pf end) A x) as h1. rewrite h1. pose proof (exist_fos_psa_compat A) as h3. rewrite h3. reflexivity. Qed. End field_of_sets_sec. Section finite_cofinite_sec. Variable Xt : Type. Definition F_fc := [S:(Ensemble Xt) | Finite S \/ Finite (Comp S)]. Lemma Union_closed_fc : forall S1 S2:(Ensemble Xt), In F_fc S1 -> In F_fc S2 -> In F_fc (Union S1 S2). intros S1 S2 h1 h2. unfold F_fc. unfold F_fc in h1. unfold F_fc in h2. inversion h1 as [h3]. inversion h2 as [h4]. destruct h3 as [h5 | h6]. (*left*) destruct h4 as [h7 | h8]. (*left*) constructor. left. apply Union_preserves_Finite; assumption. (*right*) pose proof (comp_union _ S1 S2) as h9. assert (h10: Included (Intersection (Comp S1) (Comp S2)) (Comp S2)). auto with sets. pose proof (Finite_downward_closed _ _ h8 _ h10) as h11. rewrite <- h9 in h11. constructor. right; assumption. (*right*) destruct h4 as [h12 | h13]. (*left*) pose proof (comp_union _ S1 S2) as h14. assert (h15: Included (Intersection (Comp S1) (Comp S2)) (Comp S1)). auto with sets. pose proof (Finite_downward_closed _ _ h6 _ h15) as h16. rewrite <- h14 in h16. constructor. right; assumption. (*right*) constructor. right. pose proof (comp_union _ S1 S2) as h17. assert (h18: Included (Intersection (Comp S1) (Comp S2)) (Comp S1)). auto with sets. rewrite <- h17 in h18. apply Finite_downward_closed with (Comp S1); assumption. Qed. Lemma Comp_closed_fc : forall S :(Ensemble Xt), In F_fc S -> In F_fc (Comp S). intros S h1. unfold F_fc in h1. inversion h1 as [h2]. destruct h2 as [h3 | h4]. (*left*) unfold F_fc. constructor. right. pose proof (Complement_Complement _ S) as h5. rewrite h5. assumption. (*right*) unfold F_fc. constructor. left. assumption. Qed. Lemma Int_closed_fc : forall S1 S2:(Ensemble Xt), In F_fc S1 -> In F_fc S2 -> In F_fc (Intersection S1 S2). intros S1 S2 h1 h2. pose proof (Complement_Complement _ S1) as h3. pose proof (Complement_Complement _ S2) as h4. rewrite <- h3. rewrite <- h4. pose proof (comp_union _ (Comp S1) (Comp S2)) as h5. rewrite <- h5. pose proof (Union_closed_fc (Comp S1) (Comp S2)) as h6. pose proof (Comp_closed_fc S1 h1) as h7. pose proof (Comp_closed_fc S2 h2) as h8. pose proof (h6 h7 h8) as h9. apply Comp_closed_fc. assumption. Qed. Lemma non_empty_F_fc : (exists S:(Ensemble Xt), In F_fc S). unfold F_fc. exists (Empty_set Xt). constructor. left. constructor. Qed. Lemma full_in_F_fc : In F_fc (Full_set _). pose proof non_empty_F_fc as h1. elim h1. intros S h2. pose proof (Comp_closed_fc _ h2) as h3. pose proof (Union_closed_fc _ _ h2 h3) as h4. pose proof (comp_sum_psa S) as h5. rewrite <- h5; assumption. Qed. Lemma empty_in_F_fc : In F_fc (Empty_set _). pose proof non_empty_F_fc as h1. elim h1. intros S h2. pose proof (Comp_closed_fc _ h2) as h3. pose proof (Int_closed_fc _ _ h2 h3) as h4. pose proof (comp_prod_psa S) as h5. rewrite <- h5; assumption. Qed. Definition fin_cof_fos := Build_Field_of_Sets Xt F_fc non_empty_F_fc Union_closed_fc Int_closed_fc Comp_closed_fc. Definition fin_cof_ba := fos_ba fin_cof_fos. End finite_cofinite_sec. Section Count_Cocount. Variable Xt:Type. Definition F_cc := [S:(Ensemble Xt) | Countable S \/ Countable (Comp S)]. Lemma Union_closed_cc : forall S1 S2:(Ensemble Xt), In F_cc S1 -> In F_cc S2 -> In F_cc (Union S1 S2). intros S1 S2 h1 h2. inversion h1 as [h3]. inversion h2 as [h4]. unfold F_cc. constructor. destruct h3 as [h5 | h6]. (*left*) destruct h4 as [h7|h8]. (*left*) left. apply countable_union2; trivial. (*right*) right. pose proof (comp_union _ S1 S2) as h9. rewrite h9. apply (countable_downward_closed (Intersection (Comp S1) (Comp S2)) (Comp S2)). assumption. auto with sets. (*right*) destruct h4 as [h10|h11]. (*left*) right. pose proof (comp_union _ S1 S2) as h12. rewrite h12. apply (countable_downward_closed (Intersection (Comp S1) (Comp S2)) (Comp S1)). assumption. auto with sets. (*right*) right. pose proof (comp_union _ S1 S2) as h13. rewrite h13. apply countable_intersection; trivial. Qed. Lemma Comp_closed_cc : forall S :(Ensemble Xt), In F_cc S -> In F_cc (Comp S). unfold F_cc. intros S h1. inversion h1 as [h2]. constructor. pose proof (Complement_Complement _ S) as h3. rewrite h3. tauto. Qed. Lemma Int_closed_cc : forall S1 S2:(Ensemble Xt), In F_cc S1 -> In F_cc S2 -> In F_cc (Intersection S1 S2). intros S1 S2 h1 h2. pose proof (Complement_Complement _ S1) as h3. pose proof (Complement_Complement _ S2) as h4. rewrite <- h3. rewrite <- h4. pose proof (comp_union _ (Comp S1) (Comp S2)) as h5. rewrite <- h5. pose proof (Union_closed_cc (Comp S1) (Comp S2)) as h6. pose proof (Comp_closed_cc S1 h1) as h7. pose proof (Comp_closed_cc S2 h2) as h8. pose proof (h6 h7 h8) as h9. apply Comp_closed_cc. assumption. Qed. Lemma non_empty_F_cc : (exists S:(Ensemble Xt), In F_cc S). unfold F_cc. exists (Empty_set Xt). constructor. left. apply Finite_impl_Countable. constructor. Qed. Definition cnt_ccnt_fos := Build_Field_of_Sets Xt F_cc non_empty_F_cc Union_closed_cc Int_closed_cc Comp_closed_cc. Definition cnt_ccnt_ba := fos_ba cnt_ccnt_fos. End Count_Cocount. (*For now, these will only be done with rational points and intervals, which should be sufficient for my purposes.*) Section IntervalAlgebras. Require Import ListUtilities. Open Scope Q. (*Extended rationals, i.e. with infinity endpoints.*) Inductive Qx : Type := |neg_inf : Qx |pos_inf : Qx |q_qx : Q->Qx. Inductive qx_lt : Qx->Qx->Prop := |inf2_lt : qx_lt neg_inf pos_inf |lt_inf : forall a:Q, qx_lt (q_qx a) pos_inf |inf_lt : forall a:Q, qx_lt neg_inf (q_qx a) |q_qx_lt : forall a b:Q, a < b -> qx_lt (q_qx a) (q_qx b). Inductive qx_eq : Qx->Qx->Prop := |q_qx_eq:forall a b:Q, a == b -> qx_eq (q_qx a) (q_qx b) |neg_inf_eq :qx_eq neg_inf neg_inf |pos_inf_eq: qx_eq pos_inf pos_inf. Delimit Scope qx_scope with Qx. Open Scope qx_scope. Notation "x < y" := (qx_lt x y) : qx_scope. Notation "x == y" := (qx_eq x y) : qx_scope. Notation "x <= y" := (x < y \/ x == y) : qx_scope. Inductive qx_is_q : Qx -> Prop := | qx_is_q_intro : forall a:Q, qx_is_q (q_qx a). Lemma q_qx_fst : forall (p:Q*Q), q_qx (fst p) == fst (q_qx (fst p), q_qx (snd p))%Q. intros; constructor; apply Qeq_refl. Qed. Lemma q_qx_snd : forall (p:Q*Q), q_qx (snd p) == snd (q_qx (fst p), q_qx (snd p))%Q. intros; constructor; apply Qeq_refl. Qed. Lemma qx_lt_irrefl : forall (x:Qx), ~ x < x. intros x h1; dependent induction h1. apply Qlt_irrefl in H; auto. Qed. Lemma qx_eq_refl : forall (x:Qx), x == x. intro x. destruct x; try constructor; auto. constructor. Qed. Lemma sym_qx_eq : forall (x y:Qx), x == y -> y == x. intros x y h1. destruct x, y; try induction h1; try constructor; try (rewrite H; try constructor); try constructor. Qed. Lemma trans_qx_eq : forall (x y z:Qx), x == y -> y == z -> x == z. intros x y z h1 h2. destruct x, y; try inversion h1; try inversion h2; try constructor. subst. inversion h2. subst. rewrite H1. assumption. Qed. Lemma qx_le_refl : forall (x:Qx), x <= x. intro x. right. apply qx_eq_refl. Qed. Lemma trans_qx_lt : Transitive _ qx_lt. red. intros x y z h1 h2. destruct x, y, z; try dependent induction h1; try dependent induction h2; try constructor. eapply Qlt_trans; auto. apply H. apply H0. Qed. Lemma trans_qx_le_lt : forall x y z:Qx, x<=y -> y x y<=z -> x y <= z -> x <= z. intros x y z h1 h2. destruct x, y, z; try apply qx_le_refl; try inversion h1; try inversion h2. left. assumption. left. constructor. inversion H. subst. left. constructor. left. constructor. inversion H. inversion H. left. assumption. inversion H. inversion H2. inversion H2. left. constructor. left. constructor. subst. left. constructor. inversion H2. subst. rewrite H1. assumption. subst. inversion H2. subst. right. constructor. rewrite H1. assumption. Qed. Lemma trans_qx_le_eq : forall x y z:Qx, x <= y -> y == z -> x <= z. intros x y z h1 h2. destruct x, y, z; try apply qx_le_refl; try inversion h1; try inversion h2. left. assumption. left. constructor. left. constructor. left. constructor. inversion H. destruct h1; inversion H0. subst. destruct h1; inversion H0. subst. inversion H. inversion H. inversion H. left. constructor. left. constructor. subst. left. constructor. rewrite <- H2. inversion H. assumption. subst. right. constructor. inversion H. subst. rewrite <- H2. assumption. Qed. Lemma trans_qx_eq_lt : forall x y z:Qx, x == y -> y < z -> x < z. intros x y z h1 h2. destruct x, y, z; try inversion h2; try constructor; try inversion h1. subst. rewrite H4. assumption. Qed. Lemma trans_qx_lt_eq : forall x y z:Qx, x < y -> y == z -> x < z. intros x y z h1 h2. destruct x, y, z; try inversion h1; try inversion h2; try constructor. subst. rewrite <- H4. assumption. Qed. Lemma trans_qx_le : forall x y z:Qx, x <= y -> y <= z -> x <= z. intros x y z h1 h2. destruct h2 as [h2 | h2]. left. eapply trans_qx_le_lt. apply h1. apply h2. eapply trans_qx_le_eq. apply h1. apply h2. Qed. Lemma q_qx_eq_iff : forall (p q:Q), (q_qx p) == (q_qx q) <-> (p == q)%Q. intros; split; intro h1; try constructor; auto; try inversion h1; subst; auto. Qed. Lemma q_qx_le_iff : forall (p q:Q), (q_qx p) <= (q_qx q) <-> (p <= q)%Q. intros p q. split. intro h1. destruct h1 as [h1|h1]. inversion h1. subst. apply Qlt_impl_Qle. assumption. inversion h1. subst. rewrite H1. apply Qle_refl. intro h1. apply Qle_lt_or_eq in h1. destruct h1 as [h1 | h1]. left. constructor; auto. right. constructor. assumption. Qed. Lemma q_qx_lt_iff : forall (p q:Q), (q_qx p) < (q_qx q) <-> (p < q)%Q. intros p q; split; intro h1. inversion h1; auto. constructor; auto. Qed. Lemma q_qx_dec : forall (p q:Q), {q_qx p < q_qx q} + {q_qx p == q_qx q} + {q_qx q < q_qx p}. intros p q. pose proof (Q_dec p q) as h1. destruct h1 as [[h1 | h1]|h1]. left; constructor; constructor; auto. right; constructor; auto. left; right; constructor; auto. Qed. Lemma qx_dec : forall (p q:Qx), {p < q} + {p == q} + {q < p}. intros p q. destruct p, q; try apply q_qx_dec; try (left; right;constructor); try (left; left; constructor); try (right; constructor). Qed. Lemma qx_dec' : forall (p q:Qx), {p < q} + {q <= p}. intros p q. destruct (qx_dec p q) as [h1 | h1]. destruct h1 as [h1 | h1]. left; auto. right. right. apply sym_qx_eq. assumption. right. left. assumption. Qed. Lemma qx_is_q_dec : forall q:Qx, {qx_is_q q} + {~qx_is_q q}. intro q. destruct (qx_dec neg_inf q) as [h2 | h2]. destruct h2 as [h2 | h2]. destruct (qx_dec q pos_inf) as [h3 | h3]. destruct h3 as [h3 | h3]. left. destruct q. apply qx_lt_irrefl in h2; contradiction. apply qx_lt_irrefl in h3; contradiction. constructor. right. intro h4. destruct h4. inversion h3. right. inversion h3. right. intro h3. destruct h3. inversion h2. right. inversion h2. Qed. Lemma not_qx_lt_iff : forall (p q:Qx), ~ p < q <-> q <= p. intros p q. split. intro h1. pose proof (qx_dec p q) as h2. destruct h2 as [h2 | h2]. destruct h2 as [h2 | h2]. contradiction. right. apply sym_qx_eq. assumption. left; auto. intros h2 h3. destruct h2 as [h2 | h2]. pose proof (trans_qx_lt _ _ _ h2 h3) as h4. apply qx_lt_irrefl in h4; auto. pose proof (trans_qx_eq_lt _ _ _ h2 h3) as h4. apply qx_lt_irrefl in h4; auto. Qed. Lemma not_qx_le_iff : forall (p q:Qx), ~ p <= q <-> q < p. intros p q. split. intro h1. pose proof (qx_dec p q) as h2. destruct h2 as [h2 | h2]. destruct h2 as [h2 | h2]. contradict h1. left. assumption. contradict h1. right. assumption. assumption. intros h2 h3. pose proof (trans_qx_lt_le _ _ _ h2 h3) as h4. apply qx_lt_irrefl in h4; auto. Qed. Definition qx_min (p q:Qx) := if (qx_dec p q) then p else q. Definition qx_max (p q:Qx) := if (qx_dec p q) then q else p. Lemma qx_min_le : forall (p q:Qx), qx_min p q <= p /\ qx_min p q <= q. intros p q. unfold qx_min. destruct qx_dec as [h1 | h1]. split. right. apply qx_eq_refl. destruct h1 as [h1 | h1]. left. assumption. right. assumption. split. left. assumption. right. apply qx_eq_refl. Qed. Lemma le_qx_max : forall (p q:Qx), p <= qx_max p q /\ q <= qx_max p q. intros p q. unfold qx_max. destruct qx_dec as [h1 | h1]. destruct h1 as [h1 | h1]. split. left. assumption. right. apply qx_eq_refl. split. right. apply h1. right. apply qx_eq_refl. split. right. apply qx_eq_refl. left. assumption. Qed. Lemma qx_2lt_min : forall (p q q':Qx), p < q -> p < q' -> p < qx_min q q'. intros p q q' h1 h2. unfold qx_min. destruct qx_dec; auto. Qed. Lemma qx_2lt_min_min : forall (p p' q q':Qx), p < q -> p' < q' -> qx_min p p' < qx_min q q'. intros p p' q q' h1 h2. unfold qx_min. destruct qx_dec as [h3 | h3]; destruct qx_dec as [h4 | h4]; auto. destruct h3 as [h3 | h3]. eapply trans_qx_lt. apply h3. apply h2. eapply trans_qx_eq_lt. apply h3. assumption. destruct h4 as [h4 | h4]. eapply trans_qx_lt. apply h3. apply h1. apply sym_qx_eq in h4. eapply trans_qx_lt_eq. apply h2. assumption. Qed. Lemma qx_2lt_max : forall (p p' q:Qx), p < q -> p' < q -> qx_max p p' < q. intros p p' q h1 h2. unfold qx_max. destruct qx_dec; auto. Qed. Lemma qx_2le_max : forall (p p' q:Qx), p <= q -> p' <= q -> qx_max p p' <= q. intros p p' q h1 h2. unfold qx_max. destruct qx_dec; auto. Qed. Lemma not_pos_inf_lt : forall q:Qx, ~pos_inf < q. intros q h1. inversion h1. Qed. Lemma not_lt_neg_inf : forall q:Qx, ~q < neg_inf. intros q h1. inversion h1. Qed. Lemma lt_neg_inf : forall q:Q, neg_inf < (q_qx q). intro q. constructor. Qed. Lemma le_neg_inf : forall q:Qx, neg_inf <= q. intro q; destruct q. apply qx_le_refl. constructor. constructor. constructor. constructor. Qed. Lemma lt_pos_inf : forall q:Q, (q_qx q) < pos_inf. intro q. constructor. Qed. Lemma le_pos_inf : forall q:Qx, q <= pos_inf. intro q; destruct q. constructor. constructor. apply qx_le_refl. constructor. constructor. Qed. Lemma qx_lt_dec : forall (p q:Qx), p < q -> {exists p', p = q_qx p'} + {p = neg_inf}. intros p q h1. destruct p. right; auto. apply not_pos_inf_lt in h1; contradiction. left. exists q0. reflexivity. Qed. Lemma qx_lt_dec' : forall (p q:Qx), p < q -> {exists q', q = q_qx q'} + {q = pos_inf}. intros p q h1. destruct q. apply not_lt_neg_inf in h1. contradiction. right; auto. left. exists q. reflexivity. Qed. (*Half-open interval.*) Definition ho_invl (a b:Qx) : Ensemble Q := [x:Q | a <= q_qx x /\ q_qx x < b]. Definition ho_invl' (p:Qx*Qx) := ho_invl (fst p) (snd p). Definition hoq_invl (a b:Q) := ho_invl (q_qx a) (q_qx b). Definition hoq_invl' (p:Q*Q) := hoq_invl (fst p) (snd p). Lemma hoq_invl_compat' : forall (p:Q*Q), hoq_invl' p = ho_invl' (q_qx (fst p), q_qx (snd p)). auto. Qed. Lemma inh_ho_invl_iff : forall (a b:Qx), Inhabited (ho_invl a b) <-> a < b. intros a b. split. intro h1. destruct h1 as [x h1]. destruct h1 as [h1]. destruct h1 as [h1 h2]. eapply trans_qx_le_lt. apply h1. apply h2. intro h1. destruct a, b; try inversion h1. apply Inhabited_intro with 0. constructor. split; constructor. constructor. subst. apply Inhabited_intro with (q - 1). constructor. split; try constructor. constructor. apply Qminus_pos_Qlt. constructor. subst. apply Inhabited_intro with q. constructor. split; try right; constructor. constructor. subst. apply Inhabited_intro with q. constructor. split. right. constructor. constructor. assumption. Qed. Lemma inh_hoq_invl_iff : forall (a b:Q), Inhabited (hoq_invl a b) <-> Qlt a b. intros; rewrite <- q_qx_lt_iff. apply inh_ho_invl_iff. Qed. Lemma in_ho_invl_lr_iff : forall (q r s:Q), Inn (ho_invl' (q_qx q, q_qx r)) s <-> Qle q s /\ Qlt s r. intros q r s; split; intro h1. inversion h1. simpl in H. destruct H. inversion H. inversion H0. inversion H1; subst. split. apply Qlt_le_weak; auto. assumption. inversion H0; inversion H1; subst. split. rewrite H7. apply Qle_refl. assumption. destruct h1 as [h1 h2]. econstructor. simpl; split; try rewrite q_qx_le_iff; auto. constructor; auto. Qed. Lemma in_hoq_invl_lr_iff : forall (q r s:Q), Inn (hoq_invl' (q, r)) s <-> Qle q s /\ Qlt s r. intros; rewrite <- in_ho_invl_lr_iff; tauto. Qed. Lemma in_ho_invl_l_iff : forall (q s:Q), Inn (ho_invl' (q_qx q, pos_inf)) s <-> Qle q s. intros q s; split; intro h1. destruct h1 as [h1]. simpl in h1. destruct h1 as [h1 h2]. rewrite q_qx_le_iff in h1; auto. constructor. simpl. split. rewrite q_qx_le_iff; auto. constructor. Qed. Lemma in_ho_invl_r_iff : forall (r s:Q), Inn (ho_invl' (neg_inf, q_qx r)) s <-> Qlt s r. intros r s; split; intro h1. destruct h1 as [h1]. simpl in h1. destruct h1 as [h1 h2]. rewrite q_qx_lt_iff in h2; auto. constructor. simpl. split. apply le_neg_inf. rewrite q_qx_lt_iff; auto. Qed. Lemma in_ho_invl : forall (s:Q), Inn (ho_invl' (neg_inf, pos_inf)) s. intro s. constructor. simpl; split. apply le_neg_inf. apply lt_pos_inf. Qed. Lemma ho_invl_neg_inf_eq : forall x:Qx, ho_invl x neg_inf = Empty_set _. intro x. apply Extensionality_Ensembles; red; split; red; intros a h1; try contradiction. destruct h1 as [h1]. destruct h1 as [h1 h2]. inversion h2. Qed. Lemma ho_invl_pos_inf_eq : forall x:Qx, ho_invl pos_inf x = Empty_set _. intro x. apply Extensionality_Ensembles; red; split; red; intros a h1; try contradiction. destruct h1 as [h1]. destruct h1 as [h1 h2]. inversion h1 as [h3|h3]; inversion h3. Qed. Lemma ho_invl_neg_inf_eq_neg_inf : forall (a b c:Qx), Inhabited (ho_invl neg_inf b) -> ho_invl neg_inf b = ho_invl a c -> a = neg_inf. intros a b c h0 h1. destruct (qx_dec a neg_inf) as [[h2|h2]|h2]. inversion h2. inversion h2. reflexivity. destruct a; auto. rewrite ho_invl_pos_inf_eq in h1. destruct b. rewrite ho_invl_neg_inf_eq in h0. destruct h0; contradiction. assert (h3:Inn (ho_invl neg_inf pos_inf) 0). constructor. split; auto. left. constructor. constructor. rewrite h1 in h3. contradiction. assert (h4:Inn (ho_invl neg_inf (q_qx q)) (q-1)). constructor. split; auto. left. constructor. constructor. apply Qminus_pos_Qlt. constructor. rewrite h1 in h4. contradiction. pose proof h0 as h0'. rewrite h1 in h0'. rewrite inh_ho_invl_iff in h0'. assert (h3:Inn (ho_invl (q_qx q) c) q). constructor. split; auto; try right; try constructor; try constructor. rewrite <- h1 in h3. destruct h3 as [h3]. destruct h3 as [h3 h4]. assert (h5:Inn (ho_invl neg_inf b) (q - 1)). constructor. split. left. constructor. assert (h5:(q -1 < q)%Q). apply Qminus_pos_Qlt. constructor. assert (h6:q_qx (q - 1) < q_qx q). constructor. assumption. eapply trans_qx_lt. apply h6. apply h4. rewrite h1 in h5. destruct h5 as [h5]. destruct h5 as [h5 h6]. destruct h5 as [h5 | h5]. inversion h5. subst. assert (h7:(0 < 1)%Q). constructor. pose proof (Qminus_pos_Qlt 1 q h7) as h8. pose proof (Qlt_trans _ _ _ H1 h8) as h9. apply Qlt_irrefl in h9. contradiction. inversion h5. subst. assert (h7:(q + 0 == q + -(1))%Q). rewrite <- H1. rewrite Qplus_0_r. constructor. rewrite Qplus_inj_l in h7 at 1. inversion h7. Qed. Lemma ho_invl_neg_inf_pos_inf_eq : ho_invl neg_inf pos_inf = Full_set Q. apply Extensionality_Ensembles; red; split; red; intros x h1; try constructor. split. constructor. constructor. constructor. Qed. Lemma ho_invl_pos_inf_eq_pos_inf : forall (a b c:Qx), Inhabited (ho_invl a pos_inf) -> ho_invl a pos_inf = ho_invl b c -> c = pos_inf. intros a b c h0 h1. destruct (qx_dec a pos_inf) as [[h2|h2]|h2]. inversion h2. subst. rewrite ho_invl_neg_inf_pos_inf_eq in h1. destruct b, c; auto. rewrite ho_invl_neg_inf_eq in h1. assert (h3:Inn (Full_set Q) 0). constructor. rewrite h1 in h3. contradiction. assert (h3:Inn (Full_set Q) q). constructor. rewrite h1 in h3. destruct h3 as [h3]. destruct h3 as [h3 h4]. inversion h4. subst. apply Qlt_irrefl in H1. contradiction. rewrite ho_invl_pos_inf_eq in h1. assert (h3:Inn (Full_set Q) 0). constructor. rewrite h1 in h3. contradiction. rewrite ho_invl_pos_inf_eq in h1. assert (h3:Inn (Full_set Q) 0). constructor. rewrite h1 in h3. contradiction. rewrite ho_invl_neg_inf_eq in h1. assert (h3:Inn (Full_set Q) 0). constructor. rewrite h1 in h3. contradiction. assert (h3:Inn (Full_set Q) q0). constructor. rewrite h1 in h3. destruct h3 as [h3]. destruct h3 as [h3 h4]. inversion h4. subst. apply Qlt_irrefl in H1. contradiction. subst. pose proof h0 as h0'. rewrite h1 in h0'. rewrite inh_ho_invl_iff in h0'. destruct b, c; auto. apply qx_lt_irrefl in h0'. contradiction. assert (h3:Inn (ho_invl (q_qx a0) pos_inf) a0). constructor. split. right. constructor. constructor. constructor. rewrite h1 in h3. destruct h3 as [h3]. destruct h3 as [h3 h4]. assert (h5:Inn (ho_invl (q_qx a0) pos_inf) (q + 1)). constructor. split. constructor. constructor. inversion h4. subst. assert (h5:(q < q + 1)%Q). apply Qlt_plus_pos. constructor. eapply Qlt_trans. apply H1. apply h5. constructor. rewrite h1 in h5. destruct h5 as [h5]. destruct h5 as [h5]. inversion H. subst. assert (h6:(0 < 1)%Q). constructor. pose proof (Qlt_plus_pos 1 q h6) as h7. pose proof (Qlt_trans _ _ _ h7 H2) as h8. apply Qlt_irrefl in h8. contradiction. inversion h0'. inversion h0'. inversion h0'. assert (h3:Inn (ho_invl (q_qx q) (q_qx q0)) q). constructor; auto. split. right. constructor. constructor. assumption. rewrite <- h1 in h3. destruct h3 as [h3]. destruct h3 as [h3 h4]. pose proof (trans_qx_le_lt _ _ _ h3 h0') as h5. assert (h6:Inn (ho_invl (q_qx a0) pos_inf) q0). constructor. split. left. assumption. constructor. rewrite h1 in h6. destruct h6 as [h6]. destruct h6 as [h6 h7]. inversion h7. subst. apply Qlt_irrefl in H1. contradiction. inversion h2. subst. rewrite ho_invl_pos_inf_eq in h0. destruct h0; contradiction. inversion h2. Qed. Lemma ho_invl_neg_inf_eq2 : forall (q r:Qx), ho_invl neg_inf q = ho_invl neg_inf r -> q == r. intros q r h1. destruct q, r; try constructor; auto. rewrite ho_invl_neg_inf_eq in h1. rewrite ho_invl_neg_inf_pos_inf_eq in h1. assert (h2:Inn (Full_set Q) 0). constructor. rewrite <- h1 in h2. contradiction. rewrite ho_invl_neg_inf_eq in h1. assert (h2:Inn (ho_invl neg_inf (q_qx q)) (q-1)). constructor. split. left. constructor. constructor. apply Qminus_pos_Qlt. constructor. rewrite <- h1 in h2. contradiction. rewrite ho_invl_neg_inf_pos_inf_eq in h1. rewrite ho_invl_neg_inf_eq in h1. assert (h2:Inn (Full_set Q) 0). constructor. rewrite h1 in h2. contradiction. apply ho_invl_pos_inf_eq_pos_inf in h1. rewrite h1. constructor. rewrite ho_invl_neg_inf_pos_inf_eq. apply Inhabited_intro with 0. constructor. rewrite ho_invl_neg_inf_eq in h1. assert (h2:Inn (ho_invl neg_inf (q_qx q)) (q-1)). constructor. split. left. constructor. constructor. apply Qminus_pos_Qlt. constructor. rewrite h1 in h2. contradiction. symmetry in h1. apply ho_invl_pos_inf_eq_pos_inf in h1. discriminate h1. rewrite ho_invl_neg_inf_pos_inf_eq. apply Inhabited_intro with 0. constructor. destruct (Q_dec q q0) as [[h2|h2]|h2]. assert (h3:Inn (ho_invl neg_inf (q_qx q0)) q). constructor. split. left. constructor. constructor. assumption. rewrite <- h1 in h3. destruct h3 as [h3]. destruct h3 as [h3 h4]. inversion h4. subst. apply Qlt_irrefl in H1. contradiction. assert (h3:Inn (ho_invl neg_inf (q_qx q)) q0). constructor. split. left. constructor. constructor. assumption. rewrite h1 in h3. destruct h3 as [h3]. destruct h3 as [h3 h4]. inversion h4. subst. apply Qlt_irrefl in H1. contradiction. assumption. Qed. Lemma ho_invl_pos_inf_eq2 : forall (q r:Qx), ho_invl q pos_inf = ho_invl r pos_inf -> q == r. intros q r h1. destruct q, r; try constructor; auto. rewrite ho_invl_pos_inf_eq in h1. rewrite ho_invl_neg_inf_pos_inf_eq in h1. assert (h2:Inn (Full_set Q) 0). constructor. rewrite h1 in h2. contradiction. apply ho_invl_neg_inf_eq_neg_inf in h1. discriminate h1. rewrite ho_invl_neg_inf_pos_inf_eq. apply Inhabited_intro with 0. constructor. rewrite ho_invl_pos_inf_eq in h1. rewrite ho_invl_neg_inf_pos_inf_eq in h1. assert (h2:Inn (Full_set Q) 0). constructor. rewrite <- h1 in h2. contradiction. rewrite ho_invl_pos_inf_eq in h1. assert (h2:Inn (ho_invl (q_qx q) pos_inf) q). constructor. split. right. constructor. constructor. constructor. rewrite <- h1 in h2. contradiction. symmetry in h1. apply ho_invl_neg_inf_eq_neg_inf in h1. discriminate h1. rewrite ho_invl_neg_inf_pos_inf_eq. apply Inhabited_intro with 0. constructor. rewrite ho_invl_pos_inf_eq in h1. assert (h2:Inn (ho_invl (q_qx q) pos_inf) q). constructor. split. right. constructor. constructor. constructor. rewrite h1 in h2. contradiction. destruct (Q_dec q q0) as [[h2|h2]|h2]. assert (h3:~Inn (ho_invl (q_qx q0) pos_inf) q). intro h4. destruct h4 as [h4]. destruct h4 as [h4 h5]. rewrite q_qx_le_iff in h4. pose proof (Qle_lt_trans _ _ _ h4 h2) as h6. apply Qlt_irrefl in h6. contradiction. rewrite <- h1 in h3. contradict h3. constructor. split. right. constructor. constructor. constructor. assert (h3:~Inn (ho_invl (q_qx q) pos_inf) q0). intro h4. destruct h4 as [h4]. destruct h4 as [h4 h5]. rewrite q_qx_le_iff in h4. pose proof (Qle_lt_trans _ _ _ h4 h2) as h6. apply Qlt_irrefl in h6. contradiction. rewrite h1 in h3. contradict h3. constructor. split. right. constructor. constructor. constructor. assumption. Qed. Lemma ho_invl_qq_union : forall (p q r:Q), (p < q)%Q -> (q < r)%Q -> Union (ho_invl (q_qx p) (q_qx q)) (ho_invl (q_qx q) (q_qx r)) = ho_invl (q_qx p) (q_qx r). intros p q r hp hq. pose proof (Qlt_trans _ _ _ hp hq) as h0. apply Extensionality_Ensembles; red; split; red; intros x h1. destruct h1 as [x h1 | x h1]. destruct h1 as [h1]. destruct h1 as [h1 h2]. constructor. split; auto. inversion h2. subst. constructor. eapply Qlt_trans. apply H1. apply hq. destruct h1 as [h1]. destruct h1 as [h1 h2]. constructor. split. left. eapply trans_qx_lt_le. assert (h3:q_qx p < q_qx q). constructor; auto. apply h3. assumption. assumption. destruct h1 as [h1]. destruct h1 as [h1 h2]. destruct (qx_dec (q_qx x) (q_qx q)) as [[h3|h3]|h3]. left. constructor. split; auto. inversion h3. subst. right. constructor. split. right. constructor. rewrite H1. constructor. assumption. right. constructor. split. left. assumption. assumption. Qed. Lemma hoq_invl_qq_union : forall (p q r:Q), (p < q)%Q -> (q < r)%Q -> Union (hoq_invl p q) (hoq_invl q r) = hoq_invl p r. unfold hoq_invl; intros; apply ho_invl_qq_union; auto. Qed. Lemma ho_invl_qq_eq : forall q r q' r':Q, (q < r)%Q -> (q' < r')%Q -> ho_invl (q_qx q) (q_qx r) = ho_invl (q_qx q') (q_qx r') -> (q == q')%Q /\ (r == r')%Q. intros q r q' r' h1 h2 h3. pose proof (Q_dec r r') as h5. assert (h7:(q' <= q)%Q). assert (h8:Inn (ho_invl (q_qx q') (q_qx r')) q). rewrite <- h3. constructor. split. right. constructor. constructor. constructor. assumption. destruct h8 as [h8]. destruct h8 as [h8 h9]. rewrite q_qx_le_iff in h8. assumption. assert (h7':(q <= q')%Q). assert (h8:Inn (ho_invl (q_qx q) (q_qx r)) q'). rewrite h3. constructor. split. right. constructor. constructor. constructor. assumption. destruct h8 as [h8]. destruct h8 as [h8 h9]. rewrite q_qx_le_iff in h8. assumption. pose proof (Qle_antisym _ _ h7 h7') as h0. split. rewrite h0. constructor. destruct h5 as [[h5 | h5] | h5]. pose proof (Qlt_trans _ _ _ h1 h5) as h6. clear h7 h7'. assert (h7:Inn (ho_invl (q_qx q') (q_qx r')) r). constructor. split. rewrite q_qx_le_iff. rewrite h0. apply Qlt_impl_Qle. assumption. constructor. assumption. rewrite <- h3 in h7. destruct h7 as [h7]. destruct h7 as [h7 h8]. inversion h8. subst. apply Qlt_irrefl in H1. contradiction. assert (h9:Inn (ho_invl (q_qx q) (q_qx r)) r'). constructor. split. rewrite q_qx_le_iff. rewrite <- h0. apply Qlt_impl_Qle. assumption. constructor. assumption. rewrite h3 in h9. destruct h9 as [h9]. destruct h9 as [h9 h10]. inversion h10. subst. apply Qlt_irrefl in H1. contradiction. assumption. Qed. Lemma hoq_invl_qq_eq : forall q r q' r':Q, (q < r)%Q -> (q' < r')%Q -> hoq_invl q r = hoq_invl q' r' -> (q == q')%Q /\ (r == r')%Q. unfold hoq_invl; intros; apply ho_invl_qq_eq; auto. Qed. Lemma ho_invl_pair_inj_ : forall (p r:Qx*Qx), Inhabited (ho_invl' p) -> Inhabited (ho_invl' r) -> ho_invl' p = ho_invl' r -> fst p == fst r /\ snd p == snd r. intros p r hp hr h2. destruct p as [pa pb], r as [ra rb]. simpl. unfold ho_invl' in h2, hp, hr. simpl in h2, hp, hr. rewrite inh_ho_invl_iff in hp, hr. pose proof (qx_dec pa ra) as h3. destruct h3 as [[h3 |h3] | h3]. pose proof (qx_lt_dec _ _ h3) as h4. destruct h4 as [h4 | h5]. destruct h4 as [p h4]. assert (h5:~Inn (ho_invl ra rb) p). intro h6. destruct h6 as [h6]. destruct h6 as [h6 h7]. rewrite <- h4 in h6. pose proof (trans_qx_lt_le _ _ _ h3 h6) as h8. apply qx_lt_irrefl in h8; auto. rewrite <- h2 in h5. contradict h5. constructor. rewrite <- h4. split; auto. right. apply qx_eq_refl. subst. apply ho_invl_neg_inf_eq_neg_inf in h2. subst. split; auto. constructor. apply qx_lt_irrefl in h3. contradiction. destruct pb. apply qx_lt_irrefl in hp. contradiction. apply Inhabited_intro with 0. constructor. split. left. constructor. constructor. apply Inhabited_intro with (q - 1). constructor. split. left. constructor. constructor. apply Qminus_pos_Qlt. constructor. split; auto. inversion h3. subst. destruct pb, rb; try constructor; auto. inversion hp. inversion hp. inversion hr. apply ho_invl_pos_inf_eq_pos_inf in h2. discriminate h2. apply Inhabited_intro with a. constructor. split. right. constructor. constructor. constructor. rewrite ho_invl_neg_inf_eq in h2. rewrite <- inh_ho_invl_iff in hp. rewrite h2 in hp. destruct hp; contradiction. symmetry in h2. apply ho_invl_pos_inf_eq_pos_inf in h2. discriminate h2. apply Inhabited_intro with b. constructor. split. right. constructor. constructor. constructor. apply ho_invl_qq_eq in h2. destruct h2 as [h2a h2b]. assumption. inversion hp. subst. assumption. inversion hr. subst. assumption. subst. apply ho_invl_neg_inf_eq2 in h2. assumption. subst. inversion hp. pose proof (qx_lt_dec _ _ h3) as h4. destruct h4 as [h4 | h5]. destruct h4 as [p h4]. subst. pose proof (qx_lt_dec _ _ hr) as h5. destruct h5 as [h5 | h6]. destruct h5 as [p' h5]. injection h5. intro h6. subst. assert (h6:Inn (ho_invl (q_qx p') rb) p'). constructor. split. right. constructor. constructor. assumption. rewrite <- h2 in h6. destruct h6 as [h6]. destruct h6 as [h6 h7]. pose proof (trans_qx_lt_le _ _ _ h3 h6) as h8. apply qx_lt_irrefl in h8. contradiction. discriminate h6. subst. symmetry in h2. apply ho_invl_neg_inf_eq_neg_inf in h2. subst. split. constructor. inversion h3. destruct rb. inversion hr. rewrite ho_invl_neg_inf_pos_inf_eq. apply Inhabited_intro with 0. constructor. apply Inhabited_intro with (q - 1). constructor. split. left. constructor. constructor. apply Qminus_pos_Qlt. constructor. Qed. Lemma ho_invl_pair_inj : forall (q r s t:Qx), Inhabited (ho_invl' (q, r)) -> Inhabited (ho_invl' (s, t)) -> ho_invl' (q, r) = ho_invl' (s, t) -> q == s /\ r == t. intros q r s t h1 h2 h3. pose proof (ho_invl_pair_inj_ (q, r) (s, t) h1 h2 h3) as h4. apply h4. Qed. Lemma hoq_invl_pair_inj_ : forall (p q:Q*Q), Inhabited (hoq_invl' p) -> Inhabited (hoq_invl' q) -> hoq_invl' p = hoq_invl' q -> (fst p == fst q)%Q /\ (snd p == snd q)%Q. intros p r hp hr h1. do 2 rewrite <- q_qx_eq_iff. pose proof (ho_invl_pair_inj_ (q_qx (fst p), q_qx (snd p)) (q_qx (fst r), q_qx (snd r)) hp hr h1) as h2. simpl in h2; auto. Qed. Lemma eq_endpnts_not_disjoint : forall (p r:Qx*Qx), Inhabited (ho_invl' p) -> Inhabited (ho_invl' r) -> snd p == fst r -> Union (ho_invl' p) (ho_invl' r) = ho_invl' ((fst p), (snd r)). intros p r h1 h2 h3. destruct p as [pa pb], r as [ra rb]. simpl in h3. destruct pb, ra. destruct h1 as [x h1]. destruct h1 as [h1]. simpl in h1. destruct h1 as [h1 h4]. inversion h4. inversion h3. inversion h3. inversion h3. destruct h2 as [x h2]. destruct h2 as [h2]. simpl in h2. destruct h2 as [h2]. destruct h2 as [h2|h2]. inversion h2. inversion h2. inversion h3. inversion h3. inversion h3. inversion h3. subst. simpl. apply Extensionality_Ensembles; red; split; red; intros x h4; simpl. destruct h4 as [x h4|x h4]. simpl in h4. unfold ho_invl' in h1, h2. rewrite inh_ho_invl_iff in h1, h2. simpl in h1, h2. destruct h4 as [h4]. destruct h4 as [h4 h5]. constructor. split; auto. eapply trans_qx_lt. apply h5. eapply trans_qx_eq_lt. apply h3. apply h2. simpl in h4. destruct h4 as [h4]. destruct h4 as [h4 h5]. constructor. unfold ho_invl' in h1, h2. rewrite inh_ho_invl_iff in h1, h2. simpl in h1, h2. split. pose proof (trans_qx_eq_le _ _ _ h3 h4) as h6. left. eapply trans_qx_lt_le. apply h1. assumption. assumption. simpl in h4. destruct h4 as [h4]. destruct h4 as [h4 h5]. destruct (qx_dec (q_qx x) (q_qx q)) as [h6 | h6]. destruct h6 as [h6 | h6]. left. constructor. split; auto. pose proof (trans_qx_eq _ _ _ h6 h3) as h7. right. constructor. split. right. apply sym_qx_eq. assumption. assumption. right. constructor. split. apply sym_qx_eq in h3. left. eapply trans_qx_eq_lt. apply h3. assumption. assumption. Qed. Lemma disjoint_ho_invl_iff : forall (p r:Qx*Qx), Inhabited (ho_invl' p) -> Inhabited (ho_invl' r) -> (disjoint (ho_invl' p) (ho_invl' r) <-> snd p <= fst r \/ snd r <= fst p). intros p r hp hr. split. intro h1. apply NNPP. intro h2. apply not_or_and in h2. destruct h2 as [h2 h3]. rewrite not_qx_le_iff in h2, h3. unfold ho_invl' in hp, hr. rewrite inh_ho_invl_iff in hp, hr. pose proof (qx_2lt_max _ _ _ h3 hr) as h5. apply qx_lt_dec in h5. destruct h5 as [h5 | h5]. destruct h5 as [q h5]. pose proof (le_qx_max (fst p) (fst r)) as h6. destruct h6 as [h6 h7]. contradict h1. unfold disjoint. apply Inhabited_not_empty. apply Inhabited_intro with q. constructor. constructor. split. rewrite <- h5; auto. rewrite <- h5. unfold qx_max. destruct qx_dec; auto. constructor. split. rewrite <- h5; auto. rewrite <- h5. unfold qx_max. destruct qx_dec; auto. pose proof (le_qx_max (fst p) (fst r)) as h6. rewrite h5 in h6. destruct h6 as [h6 h7]. destruct h6 as [h6 | h6]. inversion h6. destruct h7 as [h7 | h7]. inversion h7. contradict h1. unfold disjoint. apply Inhabited_not_empty. pose proof (qx_min_le (snd p) (snd r)) as h8. destruct h8 as [h8 h9]. destruct h8 as [h8 | h8]. destruct h9 as [h9 | h9]. pose proof (qx_lt_dec (qx_min (snd p) (snd r)) _ h8) as h10. destruct h10 as [h10 | h10]. destruct h10 as [q h10]. apply Inhabited_intro with (q -1). constructor. constructor. split. assert (h11:neg_inf < q_qx (q - 1)). constructor. pose proof (trans_qx_eq_lt _ _ _ h6 h11) as h12. left. assert (h13:1 > 0). constructor. pose proof (Qminus_pos_Qlt 1 q h13) as h14. assumption. pose proof (qx_min_le (snd p) (snd r)) as h11. destruct h11 as [h11 h12]. assert (h13:1 > 0). constructor. assert (h15:q_qx (q - 1) < q_qx q). pose proof (Qminus_pos_Qlt 1 q h13) as h14. constructor; auto. rewrite <- h10 in h15. eapply trans_qx_lt_le. apply h15; auto. assumption. constructor. split. left. eapply trans_qx_eq_lt. apply h7. constructor. assert (h13:1 > 0). constructor. assert (h15:q_qx (q - 1) < q_qx q). pose proof (Qminus_pos_Qlt 1 q h13) as h14. constructor; auto. pose proof (qx_min_le (snd p) (snd r)) as h11. destruct h11 as [h11 h12]. eapply trans_qx_lt. apply h15. rewrite <- h10. assumption. unfold qx_min in h10. destruct qx_dec as [h11 | h11]. rewrite h10 in h2. apply sym_qx_eq in h7. pose proof (trans_qx_eq_lt _ _ _ h7 h2) as h12. inversion h12. rewrite h10 in h3. inversion h3. assert (h10:snd r < snd p). apply sym_qx_eq in h9. eapply trans_qx_eq_lt. apply h9. assumption. pose proof (qx_lt_dec _ _ h10) as h11. destruct h11 as [h11 | h11]. destruct h11 as [q h11]. apply Inhabited_intro with (q - 1). constructor. constructor. split. left. eapply trans_qx_eq_lt. apply h6. constructor. assert (h13:1 > 0). constructor. assert (h15:q_qx (q - 1) < q_qx q). pose proof (Qminus_pos_Qlt 1 q h13) as h14. constructor; auto. rewrite <- h11 in h15. eapply trans_qx_lt. apply h15. assumption. constructor. split. left. eapply trans_qx_eq_lt. apply h7. constructor. rewrite h11. constructor. apply Qminus_pos_Qlt. constructor. rewrite h11 in hr. apply sym_qx_eq in h7. pose proof (trans_qx_eq_lt _ _ _ h7 hr) as h12. inversion h12. destruct h9 as [h9 | h9]. assert (h10:snd p < snd r). apply sym_qx_eq in h8. eapply trans_qx_eq_lt. apply h8. assumption. pose proof (qx_lt_dec _ _ h10) as h11. destruct h11 as [h11 | h11]. destruct h11 as [q h11]. apply Inhabited_intro with (q - 1). constructor. constructor. split. left. eapply trans_qx_eq_lt. apply h6. constructor. assert (h13:1 > 0). constructor. assert (h15:q_qx (q - 1) < q_qx q). pose proof (Qminus_pos_Qlt 1 q h13) as h14. constructor; auto. rewrite <- h11 in h15. assumption. constructor. split. left. eapply trans_qx_eq_lt. apply h7. constructor. assert (h13:1 > 0). constructor. assert (h15:q_qx (q - 1) < q_qx q). pose proof (Qminus_pos_Qlt 1 q h13) as h14. constructor; auto. rewrite <- h11 in h15. eapply trans_qx_lt. apply h15. assumption. rewrite h11 in hp. inversion hp. assert (h10:snd p == snd r). apply sym_qx_eq in h8. pose proof (trans_qx_eq _ _ _ h8 h9). assumption. destruct p as [pa pb], r as [ra rb]. simpl in hp, hr, h2, h3, h5, h6, h7, h8, h9, h10. destruct pa, pb, ra, rb; try inversion h10; try inversion hp; try inversion h2; try inversion h7; try inversion h6; subst. apply Inhabited_intro with 0. constructor; constructor; simpl. split. left; constructor. constructor. split. left; constructor. constructor. apply Inhabited_intro with (q - 1). constructor. constructor; simpl; split. left; constructor. assert (h13:1 > 0). constructor. assert (h15:q_qx (q - 1) < q_qx q). pose proof (Qminus_pos_Qlt 1 q h13) as h14. constructor; auto. assumption. constructor; simpl; split. left; constructor. assert (h13:1 > 0). constructor. pose proof (Qminus_pos_Qlt 1 q h13) as h14. constructor. rewrite <- H1. assumption. intro h1. red. apply Extensionality_Ensembles; red; split; red; intros x h2; try contradiction. destruct h2 as [x h2 h3]. destruct h2 as [h2], h3 as [h3]. destruct h2 as [h2a h2b], h3 as [h3a h3b]. destruct h1 as [h1 | h1]. pose proof (trans_qx_le_lt _ _ _ h3a h2b) as h4. pose proof (trans_qx_le_lt _ _ _ h1 h4) as h5. apply qx_lt_irrefl in h5. contradiction. pose proof (trans_qx_le_lt _ _ _ h2a h3b) as h4. pose proof (trans_qx_lt_le _ _ _ h4 h1) as h5. apply qx_lt_irrefl in h5. contradiction. Qed. Lemma disjoint_fst_snd_compat_le : forall (p q:Qx*Qx), Inhabited (ho_invl' p) -> Inhabited (ho_invl' q) -> disjoint (ho_invl' p) (ho_invl' q) -> (fst p <= fst q <-> snd p <= snd q). intros p q h1 h2 h3. rewrite disjoint_ho_invl_iff in h3; auto. apply inh_ho_invl_iff in h1. apply inh_ho_invl_iff in h2. split. intro h4. destruct h3 as [h3 | h3]. left. eapply trans_qx_le_lt. apply h3. assumption. pose proof (trans_qx_le_lt _ _ _ h4 h2) as h5. pose proof (trans_qx_lt_le _ _ _ h5 h3) as h6. apply qx_lt_irrefl in h6. contradiction. intro h4. destruct h3 as [h3 | h3]. left. eapply trans_qx_lt_le. apply h1. assumption. pose proof (trans_qx_le _ _ _ h4 h3) as h5. pose proof (trans_qx_le_lt _ _ _ h5 h1) as h6. apply qx_lt_irrefl in h6. contradiction. Qed. Lemma disjoint_fst_snd_compat_lt : forall (p q:Qx*Qx), Inhabited (ho_invl' p) -> Inhabited (ho_invl' q) -> disjoint (ho_invl' p) (ho_invl' q) -> (fst p < fst q <-> snd p < snd q). intros p q h1 h2 h3. rewrite disjoint_ho_invl_iff in h3; auto. apply inh_ho_invl_iff in h1. apply inh_ho_invl_iff in h2. split. intro h4. destruct h3 as [h3 | h3]. eapply trans_qx_le_lt. apply h3. assumption. pose proof (trans_qx_lt_le _ _ _ h2 h3) as h5. pose proof (trans_qx_lt _ _ _ h5 h4) as h6. apply qx_lt_irrefl in h6. contradiction. intro h4. destruct h3 as [h3 | h3]. eapply trans_qx_lt_le. apply h1. assumption. pose proof (trans_qx_lt_le _ _ _ h4 h3) as h5. pose proof (trans_qx_lt _ _ _ h5 h1) as h6. eapply qx_lt_irrefl in h6. contradiction. Qed. Definition overlapping_ho_pair_join (p r:Qx*Qx) : Qx*Qx:= ((qx_min (fst p) (fst r)), (qx_max (snd p) (snd r))). Definition overlapping_ho_pair_meet (p r:Qx*Qx) : Qx *Qx := ((qx_max (fst p) (fst r)), (qx_min (snd p) (snd r))). Lemma ho_invl_union : forall (p r:Qx*Qx), Inhabited (ho_invl' p) -> Inhabited (ho_invl' r) -> ~disjoint (ho_invl' p) (ho_invl' r) -> Union (ho_invl' p) (ho_invl' r) = ho_invl' (overlapping_ho_pair_join p r). intros p r h1 h2 h3. simpl. apply Extensionality_Ensembles; red; split; red; intros x h4. destruct h4 as [x h4 | x h4]. destruct h4 as [h4]. constructor. pose proof (qx_min_le (fst p) (fst r)) as h5. destruct h5 as [h5 h6]. destruct h4 as [h4a h4b]. split. eapply trans_qx_le. apply h5. apply h4a. pose proof (le_qx_max (snd p) (snd r)) as h7. destruct h7 as [h7 h8]. eapply trans_qx_lt_le. apply h4b. assumption. destruct h4 as [h4]. destruct h4 as [h4 h5]. constructor. split. pose proof (qx_min_le (fst p) (fst r)) as h6. destruct h6 as [h6 h7]. eapply trans_qx_le. apply h7. assumption. pose proof (le_qx_max (snd p) (snd r)) as h6. destruct h6 as [h6 h7]. eapply trans_qx_lt_le. apply h5. assumption. rewrite disjoint_ho_invl_iff in h3; auto. apply not_or_and in h3. destruct h3 as [h3a h3b]. rewrite not_qx_le_iff in h3a. rewrite not_qx_le_iff in h3b. apply inh_ho_invl_iff in h1. apply inh_ho_invl_iff in h2. destruct h4 as [h4]. simpl in h4. destruct h4 as [h4 h5]. destruct (classic (Inn (ho_invl' (fst r, snd p)) x)) as [h6 | h6]. destruct h6 as [h6]. simpl in h6. destruct h6 as [h6 h7]. unfold qx_min in h4. unfold qx_max in h5. destruct (qx_dec (snd p) (snd r)) as [h8 | h8]. right. constructor; auto. destruct qx_dec as [h9 | h9]. apply sumbool_to_or in h9. left. constructor. split; auto. clear h7. clear h6. destruct (classic (Inn (ho_invl' r) x)) as [h6 | h6]. right; auto. unfold ho_invl', ho_invl in h6. apply sat_nin in h6. apply not_and_or in h6. rewrite not_qx_le_iff in h6. destruct h6 as [h6 | h6]. pose proof (trans_qx_le_lt _ _ _ h4 h6) as h7. apply qx_lt_irrefl in h7. contradiction. rewrite not_qx_lt_iff in h6. pose proof (trans_qx_lt_le _ _ _ h3b h6) as h10. left. constructor. split; auto. unfold ho_invl', ho_invl in h6. simpl in h6. apply sat_nin in h6. apply not_and_or in h6. rewrite not_qx_le_iff in h6. rewrite not_qx_lt_iff in h6. unfold qx_max in h5. destruct qx_dec as [h7 | h7]. apply sumbool_to_or in h7. unfold qx_min in h4. destruct qx_dec as [h8 | h8]. apply sumbool_to_or in h8. destruct h6 as [h6 | h6]. pose proof (trans_qx_lt _ _ _ h6 h3a) as h9. left. constructor; split; auto. right. constructor. split. left. eapply trans_qx_lt_le. apply h3a. apply h6. assumption. destruct h6 as [h6 | h6]. pose proof (trans_qx_lt_le _ _ _ h6 h4) as h9. apply qx_lt_irrefl in h9. contradiction. right. constructor; split; auto. unfold qx_min in h4. destruct qx_dec as [h8 | h8]. left; split; auto. destruct h6 as [h6 | h6]. pose proof (trans_qx_lt_le _ _ _ h6 h4) as h9. apply qx_lt_irrefl in h9. contradiction. pose proof (trans_qx_lt_le _ _ _ h5 h6) as h9. apply qx_lt_irrefl in h9. contradiction. Qed. Lemma strict_included_neg_inf_q : forall q:Q, Strict_Included (ho_invl' (neg_inf, q_qx q)) (ho_invl' (neg_inf, pos_inf)). intro q. red; split. red; intros x h1. destruct h1 as [h1]. destruct h1 as [h1 h2]. simpl in h1, h2. constructor. split; auto. simpl. constructor. intro h1. assert (h2:Inn (ho_invl' (neg_inf, pos_inf)) q). constructor. simpl. split. left; constructor. constructor. rewrite <- h1 in h2. destruct h2 as [h2]. simpl in h2. destruct h2 as [h2 h3]. apply qx_lt_irrefl in h3. assumption. Qed. Lemma strict_included_neg_inf_minus : forall q:Q, Strict_Included (ho_invl' (neg_inf, q_qx (q - 1))) (ho_invl' (neg_inf, q_qx q)). intro q. red; split. red; intros x h1. destruct h1 as [h1]. simpl in h1. destruct h1 as [h1 h2]. assert (h3:q_qx (q - 1) < q_qx q). constructor. apply Qminus_pos_Qlt. constructor. constructor. simpl. split; auto. eapply trans_qx_lt. apply h2. assumption. intro h1. assert (h2:Inn (ho_invl' (neg_inf, q_qx q)) (q - 1)). constructor. simpl; split. left. constructor. constructor. apply Qminus_pos_Qlt. constructor. rewrite <- h1 in h2. inversion h2 as [h3]. destruct h3 as [h3 h4]. simpl in h4. apply qx_lt_irrefl in h4. assumption. Qed. Lemma strict_included_Sq_pos_inf : forall (p:Q), Strict_Included (ho_invl' (q_qx p, q_qx (p+1))) (ho_invl' (q_qx p, pos_inf)). intro p. red; split. red; intros x h1. destruct h1 as [h1]. simpl in h1. destruct h1 as [h1 h2]. constructor. simpl. split; auto. constructor. intro h1. apply ho_invl_pair_inj_ in h1. simpl in h1. destruct h1 as [h1 h2]. inversion h2. apply Inhabited_intro with p. constructor. simpl. split. right; constructor. constructor. constructor. apply Qlt_plus_pos. constructor. apply Inhabited_intro with p. constructor. simpl. split. right. constructor; constructor. constructor. Qed. Lemma strict_included_midpoint : forall (p q:Q), (p < q)%Q -> Strict_Included (ho_invl' (q_qx p, q_qx (p + (q - p) / (2 # 1)))) (ho_invl' (q_qx p, q_qx q)). intros p q h1. pose proof (midpoint_Qlt p q h1) as h4. red. split. red. intros x h2. destruct h2 as [h2]. simpl in h2. destruct h2 as [h2 h3]. constructor. simpl. split; auto. eapply trans_qx_lt. apply h3. constructor. assumption. intro h2. apply ho_invl_pair_inj_ in h2. simpl in h2. destruct h2 as [h2 h3]. inversion h3. subst. rewrite H1 in h4. apply Qlt_irrefl in h4. assumption. unfold ho_invl'. rewrite inh_ho_invl_iff. simpl. constructor. apply Qlt_midpoint. assumption. unfold ho_invl'. rewrite inh_ho_invl_iff. simpl. constructor. assumption. Qed. Lemma strict_included_midpoint_q : forall (p q:Q), (p < q)%Q -> Strict_Included (hoq_invl' (p, (p + (q - p) / (2 # 1)))) (hoq_invl' (p, q)). intros; rewrite hoq_invl_compat'; apply strict_included_midpoint; auto. Qed. Lemma ho_invl_ex_strict_included : forall p:Qx*Qx, Inhabited (ho_invl' p) -> exists q:Qx*Qx, Strict_Included (ho_invl' q) (ho_invl' p) /\ Inhabited (ho_invl' q). intros p h1. destruct p as [pa pb]. unfold ho_invl' in h1. destruct pa, pb; simpl in h1; try rewrite ho_invl_neg_inf_eq in h1. apply Inhabited_not_empty in h1. contradict h1; auto. exists (neg_inf, q_qx 0). split. apply strict_included_neg_inf_q. unfold ho_invl'; rewrite inh_ho_invl_iff; simpl; constructor. exists (neg_inf, (q_qx (q - 1))). split. apply strict_included_neg_inf_minus. unfold ho_invl'; rewrite inh_ho_invl_iff; simpl; constructor. apply Inhabited_not_empty in h1. contradict h1. auto. rewrite ho_invl_pos_inf_eq in h1. apply Inhabited_not_empty in h1. contradict h1. auto. rewrite ho_invl_pos_inf_eq in h1. apply Inhabited_not_empty in h1. contradict h1. auto. apply Inhabited_not_empty in h1. contradict h1. auto. exists (q_qx q, q_qx (q+1)). split. apply strict_included_Sq_pos_inf. unfold ho_invl'; rewrite inh_ho_invl_iff; simpl; constructor. apply Qlt_plus_pos. constructor. apply inh_ho_invl_iff in h1. inversion h1. subst. exists (q_qx q, q_qx (Qplus q (Qdiv (Qminus q0 q) (Qmake 2 1)))). split. apply strict_included_midpoint; auto. apply inh_ho_invl_iff. simpl. constructor. apply Qlt_midpoint. assumption. Qed. Definition ho_pairs_to_ho_invls_ens : list (Qx*Qx) -> Ensemble Q := (fun l => (FamilyUnion (Im (list_to_set l) (fun p => ho_invl' p)))). Lemma ho_pairs_to_ho_invls_ens_nil : ho_pairs_to_ho_invls_ens nil = Empty_set _. unfold ho_pairs_to_ho_invls_ens; simpl. rewrite image_empty. rewrite empty_family_union. reflexivity. Qed. Lemma ho_pairs_to_ho_invls_ens_cons : forall (l:list (Qx*Qx)) (p:Qx*Qx), ho_pairs_to_ho_invls_ens (p::l) = Union (ho_invl' p) (ho_pairs_to_ho_invls_ens l). intros l p. unfold ho_pairs_to_ho_invls_ens. simpl. rewrite Im_add. rewrite family_union_add. unfold ho_invl. reflexivity. Qed. Lemma ho_pairs_to_ho_invls_ens_sing : forall (p:Qx*Qx), ho_pairs_to_ho_invls_ens (p::nil) = ho_invl' p. intro p. rewrite (ho_pairs_to_ho_invls_ens_cons nil p). rewrite ho_pairs_to_ho_invls_ens_nil. rewrite Union_commutative. rewrite empty_union. reflexivity. Qed. Lemma ho_pairs_to_ho_invls_ens_preserves_app : forall (lp lq:list (Qx*Qx)), ho_pairs_to_ho_invls_ens (lp ++ lq) = Union (ho_pairs_to_ho_invls_ens lp) (ho_pairs_to_ho_invls_ens lq). intros lp lq. unfold ho_pairs_to_ho_invls_ens. rewrite list_to_set_app. rewrite im_union. rewrite familyunion_union_eq. reflexivity. Qed. Lemma inh_ho_pairs_to_ho_invls_ens_impl : forall l:list (Qx*Qx), Inhabited (ho_pairs_to_ho_invls_ens l) -> exists p:Qx*Qx, In p l /\ Inhabited (ho_invl' p). intros l h1. destruct h1 as [q h1]. destruct h1 as [S q h3 h4]. destruct h3 as [p h3]. subst. exists p. split. rewrite <- list_to_set_in_iff in h3. assumption. apply Inhabited_intro with q; auto. Qed. Inductive union_of_ho_invls : Ensemble Q -> Prop := | union_of_ho_invls_intro : forall l:list (Qx*Qx), union_of_ho_invls (ho_pairs_to_ho_invls_ens l). Lemma ho_invl_incl_union_of_ho_invls : forall (l:list (Qx * Qx)) (p:Qx*Qx), In p l -> Included (ho_invl' p) (ho_pairs_to_ho_invls_ens l). intros l p h1. red; intros x h2. unfold ho_pairs_to_ho_invls_ens. eapply family_union_intro. apply Im_intro with p. rewrite <- list_to_set_in_iff; auto. apply eq_refl. assumption. Qed. Lemma union_of_ho_invls_ex_strict_included : forall S:Ensemble Q, Inhabited S -> union_of_ho_invls S -> exists S':Ensemble Q, Strict_Included S' S /\ union_of_ho_invls S' /\ Inhabited S'. intros S h1 h2. destruct h2 as [l]. destruct (classic (exists p:Qx*Qx, In p l /\ ho_invl' p <> ho_pairs_to_ho_invls_ens l /\ Inhabited (ho_invl' p))) as [h2 | h2]. destruct h2 as [p h2]. destruct h2 as [h2 [h3 h3']]. exists (ho_invl' p). split. split. red; intros x h4. unfold ho_pairs_to_ho_invls_ens. apply family_union_intro with (ho_invl' p). apply Im_intro with p. rewrite <- list_to_set_in_iff. assumption. reflexivity. assumption. assumption. split. rewrite <- ho_pairs_to_ho_invls_ens_sing. constructor. assumption. destruct (nil_dec l) as [h3 | h3]. subst. apply Inhabited_not_empty in h1. rewrite ho_pairs_to_ho_invls_ens_nil in h1. contradict h1; auto. pose proof (inh_ho_pairs_to_ho_invls_ens_impl _ h1) as ha. destruct ha as [p [ha hb]]. pose proof (not_ex_all_not _ _ h2 p) as h6. simpl in h6. apply not_and_or in h6. destruct h6 as [|h6]; try contradiction. apply not_and_or in h6. destruct h6 as [h6|h6]; try contradiction. apply NNPP in h6. rewrite <- h6. rewrite <- h6 in h1. pose proof (ho_invl_ex_strict_included _ h1) as h7. destruct h7 as [q [h7 h8]]. exists (ho_invl' q). split; auto. split. rewrite <- ho_pairs_to_ho_invls_ens_sing. constructor. assumption. Qed. Definition int_alg_t := {S:Ensemble Q | union_of_ho_invls S}. Lemma int_alg_t_plus_sig : forall (A B:int_alg_t), union_of_ho_invls (Union (proj1_sig A) (proj1_sig B)). intros A B. destruct A as [A h1], B as [B h2]. destruct h1 as [la], h2 as [lb]. simpl. rewrite <- ho_pairs_to_ho_invls_ens_preserves_app. constructor. Qed. Definition int_alg_t_plus (A B:int_alg_t) := exist _ _ (int_alg_t_plus_sig A B). Lemma intersection_ho_invls : forall (p q:Qx*Qx), let S := Intersection (ho_invl' p) (ho_invl' q) in S = Empty_set _ \/ S = ho_invl' (overlapping_ho_pair_meet p q). intros p q S. cut (S <> Empty_set Q -> S = ho_invl' (overlapping_ho_pair_meet p q)). tauto. intro h1. unfold S in h1. assert (h2:~disjoint (ho_invl' p) (ho_invl' q)). unfold disjoint. assumption. destruct (ens_eq_dec _ (ho_invl' p) (Empty_set _)) as [h3 | h3]. contradict h2. red. rewrite h3. rewrite empty_intersection. reflexivity. destruct (ens_eq_dec _ (ho_invl' q) (Empty_set _)) as [h4 | h4]. contradict h2. rewrite h4. red. rewrite Intersection_commutative. rewrite empty_intersection. reflexivity. apply not_empty_Inhabited in h3. apply not_empty_Inhabited in h4. rewrite disjoint_ho_invl_iff in h2; auto. apply not_or_and in h2. destruct h2 as [h2 h5]. rewrite not_qx_le_iff in h2, h5. apply Extensionality_Ensembles; red; split; red; intros x h6. constructor. unfold overlapping_ho_pair_meet. destruct h6 as [x h6 h7]. destruct h6 as [h6], h7 as [h7]. destruct h6 as [h6a h6b], h7 as [h7a h7b]. split; simpl. apply qx_2le_max; auto. apply qx_2lt_min; auto. destruct h6 as [h6]. destruct h6 as [h6 h7]. unfold overlapping_ho_pair_meet, qx_max, qx_min in h6, h7. simpl in h6, h7. destruct qx_dec as [h8 | h8]. destruct qx_dec as [h9 | h9]. constructor; constructor. split; auto. apply sumbool_to_or in h8. apply sumbool_to_or in h9. eapply trans_qx_le. apply h8. assumption. apply sumbool_to_or in h8. apply sumbool_to_or in h9. split; auto. eapply trans_qx_lt_le. apply h7. assumption. apply sumbool_to_or in h8. constructor; constructor. split; auto. eapply trans_qx_le. apply h8. assumption. eapply trans_qx_lt. apply h7. assumption. split; auto. destruct qx_dec as [h9 | h9]. apply sumbool_to_or in h9. constructor; constructor. split; auto. split; auto. left. eapply trans_qx_lt_le. apply h8. assumption. eapply trans_qx_lt_le. apply h7. assumption. constructor; constructor. split; auto. eapply trans_qx_lt. apply h7. assumption. split; auto. left. eapply trans_qx_lt_le. apply h8. assumption. Qed. Lemma union_of_ho_invls_union : forall (p:Qx*Qx) (S:Ensemble Q), union_of_ho_invls S -> union_of_ho_invls (Union (ho_invl' p) S). intros p S h1. destruct h1. rewrite <- ho_pairs_to_ho_invls_ens_cons. constructor. Qed. Lemma union_of_ho_invls_intersection_ho_invl_int_alg_t : forall (p:Qx*Qx) (S:int_alg_t), union_of_ho_invls (Intersection (ho_invl' p) (proj1_sig S)). intros p S. destruct S as [S h1]. simpl. destruct h1 as [l]. induction l as [|a l h1]; simpl. rewrite ho_pairs_to_ho_invls_ens_nil. rewrite Intersection_commutative. rewrite empty_intersection. rewrite <- ho_pairs_to_ho_invls_ens_nil. constructor. rewrite ho_pairs_to_ho_invls_ens_cons. rewrite Distributivity. pose proof (intersection_ho_invls p a) as h2. simpl in h2. destruct h2 as [h2 | h2]. rewrite h2. rewrite empty_union. assumption. rewrite h2. apply union_of_ho_invls_union; auto. Qed. Lemma union_of_ho_invls_family : forall (F:Family Q), Finite F -> (forall S, Inn F S -> union_of_ho_invls S) -> union_of_ho_invls (FamilyUnion F). intros F h1. induction h1 as [|F h2 h3 A h4]; simpl; auto. intros h2. rewrite empty_family_union. rewrite <- ho_pairs_to_ho_invls_ens_nil. constructor. intros h5. assert (h6:(forall S : Ensemble Q, Inn F S -> union_of_ho_invls S)). intros S h6. apply h5. left. assumption. specialize (h3 h6). rewrite family_union_add. assert (h7:Inn (Add F A) A). right. constructor; auto. specialize (h5 _ h7). inversion h3 as [l h8]. destruct h5 as [l']. rewrite <- ho_pairs_to_ho_invls_ens_preserves_app. constructor. Qed. Lemma int_alg_t_times_sig : forall (R S:int_alg_t), union_of_ho_invls (Intersection (proj1_sig R) (proj1_sig S)). intros R S. destruct R as [R h1]. simpl. destruct h1 as [l]. unfold ho_pairs_to_ho_invls_ens. rewrite Intersection_commutative. rewrite dist_union_infnt_1. unfold Int_Fam_Ens. rewrite im_im. assert (h1:forall A, Inn (Im (list_to_set l) (fun x : Qx * Qx => Intersection (proj1_sig S) (ho_invl' x))) A -> union_of_ho_invls A). intros A h2. destruct h2 as [A h2]. subst. rewrite Intersection_commutative. apply union_of_ho_invls_intersection_ho_invl_int_alg_t. apply union_of_ho_invls_family. apply finite_image. apply list_to_set_finite. assumption. Qed. Definition int_alg_t_times (R S:int_alg_t) := exist _ _ (int_alg_t_times_sig R S). Lemma union_of_ho_invls_full : union_of_ho_invls (Full_set Q). assert (h1:Full_set Q = ho_pairs_to_ho_invls_ens ((neg_inf, pos_inf) :: nil)). unfold ho_pairs_to_ho_invls_ens. apply Extensionality_Ensembles; red; split; red; intros x h1. eapply family_union_intro. simpl. rewrite Im_add. rewrite image_empty. rewrite add_empty_sing. unfold ho_invl'. simpl. rewrite ho_invl_neg_inf_pos_inf_eq. econstructor. assumption. constructor. rewrite h1; constructor. Qed. Lemma union_of_ho_invls_empty : union_of_ho_invls (Empty_set Q). assert (h1:Empty_set Q = ho_pairs_to_ho_invls_ens nil). rewrite ho_pairs_to_ho_invls_ens_nil. reflexivity. rewrite h1. constructor. Qed. Lemma union_of_ho_invls_ho_invl : forall (p q:Qx), union_of_ho_invls (ho_invl p q). intros p q. assert (h1:ho_invl p q = ho_pairs_to_ho_invls_ens ((p, q)::nil)). unfold ho_pairs_to_ho_invls_ens. simpl. rewrite Im_add. rewrite image_empty. rewrite add_empty_sing. rewrite family_union_sing. reflexivity. rewrite h1. constructor. Qed. Lemma union_of_ho_invls_union_ho_invl : forall (p p' q q':Qx), union_of_ho_invls (Union (ho_invl p q) (ho_invl p' q')). intros p p' q q'. assert (h1: Union (ho_invl p q) (ho_invl p' q') = ho_pairs_to_ho_invls_ens ((p, q)::(p', q')::nil)). unfold ho_pairs_to_ho_invls_ens. simpl. rewrite <- couple_add_add. rewrite im_couple. rewrite family_union_couple. rewrite Union_commutative. reflexivity. rewrite h1. constructor. Qed. Lemma comp_ho_invl_neg_inf_q_qx : forall q:Q, Comp (ho_invl neg_inf (q_qx q)) = ho_invl (q_qx q) pos_inf. intro q. apply Extensionality_Ensembles; red; split; red; intros x h1. red in h1. red in h1. unfold ho_invl in h1. apply sat_nin in h1. apply not_and_or in h1. rewrite not_qx_le_iff in h1. destruct h1 as [h1 | h1]. inversion h1. apply not_qx_lt_iff in h1. constructor; constructor; auto. constructor. inversion h1 as [h2]. destruct h2 as [h2 h3]. red. red. intro h4. inversion h4 as [h5]. destruct h5 as [h5 h6]. pose proof (trans_qx_lt_le _ _ _ h6 h2) as h7. apply qx_lt_irrefl in h7. contradiction. Qed. Lemma comp_ho_invl_q_qx_pos_inf : forall q:Q, Comp (ho_invl (q_qx q) pos_inf) = ho_invl neg_inf (q_qx q). intro q. rewrite <- Complement_Complement. f_equal. symmetry. apply comp_ho_invl_neg_inf_q_qx. Qed. Lemma comp_ho_invls_q_qx_q_qx : forall q q':Q, Comp (ho_invl (q_qx q) (q_qx q')) = Union (ho_invl neg_inf (q_qx q)) (ho_invl (q_qx q') pos_inf). intros q q'. apply Extensionality_Ensembles; red; split; red; intros x h1. red in h1. red in h1. unfold ho_invl in h1. apply sat_nin in h1. apply not_and_or in h1. rewrite not_qx_le_iff, not_qx_lt_iff in h1. destruct h1 as [h1 | h1]. left. constructor. split; auto. left. constructor. right. constructor. split; auto. constructor. red. red. intro h2. destruct h2 as [h2]. destruct h2 as [h2 h3]. destruct h1 as [x h1 | x h1]. destruct h1 as [h1]. destruct h1 as [h1 h1']. pose proof (trans_qx_lt_le _ _ _ h1' h2) as h4. apply qx_lt_irrefl in h4. assumption. destruct h1 as [h1]. destruct h1 as [h1 h1']. pose proof (trans_qx_le_lt _ _ _ h1 h3) as h4. apply qx_lt_irrefl in h4. assumption. Qed. Lemma union_of_ho_invls_comp : forall p:Qx*Qx, union_of_ho_invls (Comp (ho_invl' p)). intro p. unfold ho_invl'. destruct p as [pa pb]. destruct pa, pb; simpl; try rewrite ho_invl_neg_inf_eq; try rewrite ho_invl_pos_inf_eq; try rewrite ho_invl_neg_inf_pos_inf_eq; try rewrite complement_empty; try rewrite complement_full; try apply union_of_ho_invls_empty; try apply union_of_ho_invls_full; try rewrite comp_ho_invls_q_qx_q_qx; try rewrite comp_ho_invl_q_qx_pos_inf; try rewrite comp_ho_invl_neg_inf_q_qx; try apply union_of_ho_invls_ho_invl; try apply union_of_ho_invls_union_ho_invl. Qed. Lemma int_alg_t_comp_sig : forall S:int_alg_t, union_of_ho_invls (Comp (proj1_sig S)). intros S. destruct S as [S h1]. destruct h1. simpl. induction l as [|a l h1]. rewrite ho_pairs_to_ho_invls_ens_nil. rewrite complement_empty. apply union_of_ho_invls_full. rewrite ho_pairs_to_ho_invls_ens_cons. rewrite comp_union. pose proof (union_of_ho_invls_comp a) as h2. pose proof (int_alg_t_times_sig (exist _ _ h2) (exist _ _ h1)) as h3. simpl in h3. assumption. Qed. Definition int_alg_t_comp (R:int_alg_t) := exist _ _ (int_alg_t_comp_sig R). Definition int_alg_bc := Build_Bconst int_alg_t (Full_set int_alg_t) int_alg_t_plus int_alg_t_times (exist _ _ union_of_ho_invls_full) (exist _ _ union_of_ho_invls_empty) int_alg_t_comp. Lemma int_alg_und_set : BS int_alg_bc = Full_set (Btype int_alg_bc). auto. Qed. Lemma int_alg_assoc_sum : forall n m p : Btype int_alg_bc, (n + (m + p))%Bool_Alg = (n + m + p)%Bool_Alg. intros; simpl. unfold int_alg_t_plus. apply proj1_sig_injective. simpl. rewrite Union_associative. reflexivity. Qed. Lemma int_alg_assoc_prod : forall n m p : Btype int_alg_bc, (n * (m * p))%Bool_Alg = (n * m * p)%Bool_Alg. intros; simpl. unfold int_alg_t_times. apply proj1_sig_injective. simpl. rewrite Intersection_associative. reflexivity. Qed. Lemma int_alg_comm_sum : forall n m : Btype int_alg_bc, (n + m)%Bool_Alg = (m + n)%Bool_Alg. intros; simpl. unfold int_alg_t_plus. apply proj1_sig_injective. simpl. rewrite Union_commutative. reflexivity. Qed. Lemma int_alg_comm_prod : forall n m : Btype int_alg_bc, (n * m)%Bool_Alg = (m * n)%Bool_Alg. intros; simpl. unfold int_alg_t_times. apply proj1_sig_injective. simpl. rewrite Intersection_commutative. reflexivity. Qed. Lemma int_alg_abs_sum : forall n m : Btype int_alg_bc, (n + n * m)%Bool_Alg = n. intros; simpl. unfold int_alg_t_plus, int_alg_t_times. apply proj1_sig_injective. simpl. rewrite abs_sum_psa. reflexivity. Qed. Lemma int_alg_abs_prod : forall n m : Btype int_alg_bc, (n * (n + m))%Bool_Alg = n. intros; simpl. unfold int_alg_t_plus, int_alg_t_times. apply proj1_sig_injective. simpl. rewrite abs_prod_psa. reflexivity. Qed. Lemma int_alg_dist_sum : forall n m p : Btype int_alg_bc, (p * (n + m))%Bool_Alg = (p * n + p * m)%Bool_Alg. intros; simpl. unfold int_alg_t_times, int_alg_t_plus. apply proj1_sig_injective. simpl. rewrite Distributivity. reflexivity. Qed. Lemma int_alg_dist_prod : forall n m p : Btype int_alg_bc, (p + n * m)%Bool_Alg = ((p + n) * (p + m))%Bool_Alg. intros; simpl. unfold int_alg_t_times, int_alg_t_plus. apply proj1_sig_injective. simpl. rewrite dist_prod_psa. reflexivity. Qed. Lemma int_alg_comp_sum : forall n : Btype int_alg_bc, (n + - n)%Bool_Alg = 1%Bool_Alg. intros; simpl. unfold int_alg_t_plus. apply proj1_sig_injective; simpl. apply comp_sum_psa. Qed. Lemma int_alg_comp_prod : forall n : Btype int_alg_bc, (n * - n)%Bool_Alg = 0%Bool_Alg. intros; simpl. unfold int_alg_t_plus. apply proj1_sig_injective; simpl. apply comp_prod_psa. Qed. Definition int_alg_ba := Build_Bool_Alg int_alg_bc int_alg_und_set int_alg_assoc_sum int_alg_assoc_prod int_alg_comm_sum int_alg_comm_prod int_alg_abs_sum int_alg_abs_prod int_alg_dist_sum int_alg_dist_prod int_alg_comp_sum int_alg_comp_prod. Lemma int_alg_ba_le_iff : forall (x y:bt int_alg_ba), le x y <-> Included (proj1_sig x) (proj1_sig y). intros x y. destruct x as [R h1], y as [S h2]. simpl. split. intro h3. red in h3. simpl in h3. unfold int_alg_t_plus in h3. apply exist_injective in h3. simpl in h3. rewrite <- inclusion_iff_union in h3. assumption. intro h3. red. simpl. unfold int_alg_t_plus. apply proj1_sig_injective; simpl. rewrite <- inclusion_iff_union. assumption. Qed. Lemma int_alg_ba_lt_iff : forall (x y:bt int_alg_ba), lt x y <-> Strict_Included (proj1_sig x) (proj1_sig y). intros x y. split. intro h1. destruct h1 as [h1 h2]. rewrite int_alg_ba_le_iff in h1. split; auto. intro h3. apply proj1_sig_injective in h3. subst. contradict h2; auto. intro h1. destruct h1 as [h1 h2]. rewrite <- int_alg_ba_le_iff in h1. red; split; auto. intro h3. subst. contradict h2; auto. Qed. Section CountableIntAlg. Lemma countable_qx : CountableT Qx. pose proof Q_countable as h1. pose proof Q_infinite as h2. pose proof (inf_countable_index _ h1 h2) as h3. destruct h3 as [f h3]. pose proof h3 as h3'. apply bijective_impl_invertible in h3'. pose (proj1_sig (function_inverse _ h3')) as f'. pose proof (invertible_impl_inv_invertible _ h3') as h4. pose (fun q:Qx => match q with | neg_inf => 0%nat | pos_inf => 1%nat | q_qx q' => S (S (f' q')) end) as g. assert (h5:injective g). red; intros x y h5. unfold g in h5. destruct x, y; auto; try omega. do 2 apply S_inj in h5. unfold f' in h5. apply invertible_impl_bijective in h4. destruct h4 as [h4a h4b]. apply h4a in h5. subst. reflexivity. econstructor. apply h5. Qed. Lemma countable_qxp : CountableT (Qx*Qx). apply countable_product; apply countable_qx. Qed. Lemma countable_int_alg_t : CountableT int_alg_t. pose (fun x:all_tuples (Qx*Qx) => (exist _ _ (union_of_ho_invls_intro (nprod_to_list _ _ (projT2 x))))) as f. pose proof (countable_all_tuples_of_countable_type _ countable_qxp) as h1. apply (surj_countable f h1). red. intro y. destruct y as [S h2]. destruct h2 as [l]. unfold f. exists (existT _ (length l) (nprod_of_list _ l)). apply proj1_sig_injective. simpl. rewrite nprod_to_list_nprod_of_list_eq. reflexivity. Qed. Lemma countable_int_alg_ba : countable_ba int_alg_ba. apply countable_int_alg_t. Qed. End CountableIntAlg. End IntervalAlgebras. bool2-v0-3/Quotients.v0000644000175000017500000002142413575574055015527 0ustar dwyckoff76dwyckoff76(*This file is a complete copy and paste from Daniel Schepler's "Zorn's Lemma." If I add to it, then I will add a copyright.*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Import SetUtilities. Require Export Image. Require Import ImageImplicit. Require Export Relation_Definitions. Require Import Relation_Definitions_Implicit. Require Import Description. Require Import ProofIrrelevance. Require Import LogicUtilities. Set Implicit Arguments. Section Quotient. Variable A:Type. Variable R:relation A. Hypothesis equivR:equivalence R. Definition equiv_class (x:A) : Ensemble A := [ y:A | R x y ]. Lemma R_impl_equality_of_equiv_class: forall x y:A, R x y -> equiv_class x = equiv_class y. Proof. destruct equivR. intros. apply Extensionality_Ensembles; split; red; intros z ?. constructor. destruct H0. apply equiv_trans with x; trivial. apply equiv_sym; trivial. destruct H0. constructor. apply equiv_trans with y; trivial. Qed. Lemma equality_of_equiv_class_impl_R: forall x y:A, equiv_class x = equiv_class y -> R x y. Proof. destruct equivR. intros. assert (In (equiv_class x) x). constructor. apply equiv_refl. rewrite H in H0. destruct H0. apply equiv_sym. assumption. Qed. Definition equiv_classes : Ensemble (Ensemble A) := Im (Full_set _) equiv_class. Definition quotient : Type := { S:Ensemble A | In equiv_classes S }. Lemma equiv_class_in_quotient: forall x:A, In equiv_classes (equiv_class x). Proof. unfold equiv_classes. intro. apply Im_def. constructor. Qed. Definition quotient_projection (x:A) : quotient := exist _ (equiv_class x) (equiv_class_in_quotient x). Lemma quotient_projection_correct: forall x:A, proj1_sig (quotient_projection x) = equiv_class x. Proof. trivial. Qed. Lemma quotient_projection_surjective: forall xbar:quotient, exists x:A, quotient_projection x = xbar. Proof. intro. destruct xbar. destruct i. exists x. unfold quotient_projection. pose proof e. symmetry in H; destruct H. f_equal. apply proof_irrelevance. Qed. Lemma quotient_projection_collapses_R: forall x1 x2:A, R x1 x2 -> quotient_projection x1 = quotient_projection x2. Proof. intros. apply subset_eq_compatT. apply R_impl_equality_of_equiv_class. assumption. Qed. Lemma quotient_projection_minimally_collapses_R: forall x1 x2:A, quotient_projection x1 = quotient_projection x2 -> R x1 x2. Proof. intros. apply equality_of_equiv_class_impl_R. repeat rewrite <- quotient_projection_correct. rewrite H. reflexivity. Qed. End Quotient. Section InducedFunction. (* well-defined A->B induces a function A/R->B *) Variable A B:Type. Variable R:Relation A. Variable f:A->B. Hypothesis equiv:equivalence R. Hypothesis well_defined: forall x y:A, R x y -> f x = f y. Lemma description_of_fbar: forall xbar:quotient R, exists! y:B, exists x:A, quotient_projection R x = xbar /\ f x = y. Proof. intro. pose proof (quotient_projection_surjective xbar). destruct H. exists (f x). unfold unique. split. exists x. tauto. intros. destruct H0. destruct H0. rewrite <- H1. apply well_defined. apply equality_of_equiv_class_impl_R; trivial. transitivity (proj1_sig (quotient_projection R x)). trivial. rewrite H. rewrite <- H0. trivial. Qed. Definition induced_function (xbar:quotient R) : B := proj1_sig (constructive_definite_description _ (description_of_fbar xbar)). Lemma induced_function_correct: forall x:A, induced_function (quotient_projection R x) = f x. Proof. intro. unfold induced_function. destruct constructive_definite_description. simpl. destruct e. destruct H. rewrite <- H0. apply well_defined. apply quotient_projection_minimally_collapses_R; trivial. Qed. Lemma induced_function_unique: forall fbar:quotient R->B, (forall x:A, fbar (quotient_projection R x) = f x) -> (forall xbar:quotient R, fbar xbar = induced_function xbar). Proof. intros. destruct (quotient_projection_surjective xbar). rewrite <- H0. rewrite H. rewrite induced_function_correct. reflexivity. Qed. End InducedFunction. Section InducedFunction2. (* well-defined function A->B induces a function A/R->B/S *) Variable A B:Type. Variable R:relation A. Variable S:relation B. Variable f:A->B. Hypothesis equivR: equivalence R. Hypothesis equivS: equivalence S. Hypothesis well_defined2: forall a1 a2:A, R a1 a2 -> S (f a1) (f a2). Definition projf (a:A) : quotient S := quotient_projection S (f a). Lemma projf_well_defined: forall a1 a2:A, R a1 a2 -> projf a1 = projf a2. Proof. intros. unfold projf. apply quotient_projection_collapses_R; trivial. apply well_defined2. assumption. Qed. Definition induced_function2: quotient R -> quotient S := induced_function projf equivR projf_well_defined. Lemma induced_function2_correct: forall a:A, induced_function2 (quotient_projection R a) = quotient_projection S (f a). Proof. intros. unfold induced_function2. rewrite induced_function_correct. reflexivity. Qed. End InducedFunction2. Section InducedFunction2arg. (* well-defined function A x B -> C induces a function (A/R) x (B/S) -> C *) Variable A B C:Type. Variable R:relation A. Variable S:relation B. Variable f:A->B->C. Hypothesis equivR:equivalence R. Hypothesis equivS:equivalence S. Hypothesis well_defined_2arg: forall (a1 a2:A) (b1 b2:B), R a1 a2 -> S b1 b2 -> f a1 b1 = f a2 b2. Lemma slices_well_defined: forall (a:A) (b1 b2:B), S b1 b2 -> f a b1 = f a b2. Proof. intros. apply well_defined_2arg. apply (equiv_refl equivR). assumption. Qed. Definition induced1 (a:A) : quotient S -> C := induced_function (f a) equivS (slices_well_defined a). Definition eq_fn (f g:quotient S->C) := forall b:quotient S, f b = g b. Lemma eq_fn_equiv: equivalence eq_fn. Proof. constructor. unfold reflexive. unfold eq_fn. reflexivity. unfold transitive. unfold eq_fn. intros. transitivity (y b). apply H. apply H0. unfold symmetric. unfold eq_fn. symmetry. apply H. Qed. Lemma well_defined_induced1: forall a1 a2:A, R a1 a2 -> eq_fn (induced1 a1) (induced1 a2). Proof. unfold eq_fn. intros. pose proof (quotient_projection_surjective b). destruct H0. rewrite <- H0. unfold induced1. rewrite induced_function_correct. rewrite induced_function_correct. apply well_defined_2arg. assumption. apply (equiv_refl equivS). Qed. Definition induced2 := induced_function2 induced1 equivR eq_fn_equiv well_defined_induced1. Definition eval (b:quotient S) (f:quotient S->C): C := f b. Lemma well_defined_eval: forall (b:quotient S) (f g:quotient S->C), eq_fn f g -> eval b f = eval b g. Proof. intros. unfold eval. apply H. Qed. Definition induced_eval (b:quotient S) := induced_function (eval b) eq_fn_equiv (well_defined_eval b). Definition induced_function2arg (a:quotient R) (b:quotient S) : C := induced_eval b (induced2 a). Lemma induced_function2arg_correct: forall (a:A) (b:B), induced_function2arg (quotient_projection R a) (quotient_projection S b) = f a b. Proof. intros. unfold induced_function2arg. unfold induced_eval. unfold induced2. rewrite induced_function2_correct. rewrite induced_function_correct. unfold eval. unfold induced1. rewrite induced_function_correct. reflexivity. Qed. End InducedFunction2arg. Section InducedFunction3. (* well-defined function A x B -> C induces a function (A/R) x (B/S) -> C/T *) Variable A B C:Type. Variable R:relation A. Variable S:relation B. Variable T:relation C. Variable f:A->B->C. Hypothesis equivR:equivalence R. Hypothesis equivS:equivalence S. Hypothesis equivT:equivalence T. Hypothesis well_defined3: forall (a1 a2:A) (b1 b2:B), R a1 a2 -> S b1 b2 -> T (f a1 b1) (f a2 b2). Definition projf2 (a:A) (b:B) := quotient_projection T (f a b). Lemma projf2_well_defined: forall (a1 a2:A) (b1 b2:B), R a1 a2 -> S b1 b2 -> projf2 a1 b1 = projf2 a2 b2. Proof. intros. apply quotient_projection_collapses_R; trivial. apply well_defined3. assumption. assumption. Qed. Definition induced_function3 : quotient R -> quotient S -> quotient T := induced_function2arg projf2 equivR equivS projf2_well_defined. Lemma induced_function3_correct: forall (a:A) (b:B), induced_function3 (quotient_projection R a) (quotient_projection S b) = quotient_projection T (f a b). Proof. intros. unfold induced_function3. rewrite induced_function2arg_correct. reflexivity. Qed. End InducedFunction3. bool2-v0-3/ArithUtilities.v0000644000175000017500000050206013575574055016477 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-2019, Daniel Wyckoff*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- version 8.4 pl4 Coq.*) Require Export Arith. Require Import Omega. Require Import NPeano. Require Import ProofIrrelevance. Require Import LogicUtilities. Require Import Ensembles. Require Import Even. Require Import Equality. Require Import Description. Require Import FunctionProperties. Require Import EqDec. Require Import FunctionalExtensionality. Require Import DecidableDec. Notation le_lt := le_lt_n_Sm. Notation lt_le := lt_n_Sm_le. Section Trivialities. Lemma disc_nat : forall n, S n <> 0. auto. Qed. Lemma pred0 : pred 0 = 0. auto. Qed. Lemma lt_0_1 : 0 < 1. auto with arith. Qed. Lemma lt_0_2 : 0 < 2. auto with arith. Qed. Lemma lt_1_2 : 1 < 2. auto with arith. Qed. Lemma le0 : forall (m:nat), m <= 0 -> m = 0. intro m; destruct m; auto with arith. Qed. Lemma le0' : forall (m:nat), m = 0 -> m <= 0. intros; subst; auto with arith. Qed. Lemma le_1_S : forall n, 1 <= S n. auto with arith. Qed. Lemma ge_S_1 : forall n, S n >= 1. auto with arith. Qed. Lemma lt1 : forall (n:nat), n < 1 -> n = 0. intros; omega. Qed. Lemma lt1' : forall (n:nat), n = 0 -> n < 1. intros; subst; auto. Qed. Lemma lt_1_SSn : forall n, 1 < (S (S n)). intros; auto with arith. Qed. Lemma lt2_dec : forall (n:nat), n < 2 -> {n = 0} + {n = 1}. intros n h1. destruct n as [|n]. left; auto. right. apply lt_S_n in h1. apply lt1 in h1; subst; auto. Qed. Lemma Olt : forall m n, m < n -> 0 < n. intros; omega. Qed. Lemma gt0 : forall n:nat, 0 < n -> n = S (pred n). intros; omega. Qed. Lemma gt1 : forall n:nat, 1 < n -> n = S (S (pred (pred n))). intro n. induction n; omega. Qed. Lemma lt_pred' : forall m n, m < pred n -> m < n. intros; omega. Qed. Lemma lt_pred_S : forall m n, m < pred n -> S m < n. intros; omega. Qed. Lemma lt_le_pred : forall m n, m < n -> m <= pred n. intros; omega. Qed. Lemma le_S_pred : forall m n, m <= S n -> pred m <= n. intros; omega. Qed. Lemma le_pred_S : forall m n, pred m <= n -> m <= S n. intros; omega. Qed. Lemma ge_S_n : forall m n, S m >= S n -> m >= n. auto with arith. Qed. Lemma ge_n_S : forall m n, m >= n -> S m >= S n. auto with arith. Qed. Lemma ge_Sn_n : forall n, S n >= n. auto with arith. Qed. Lemma eq_le : forall (m n:nat), m = n -> m <= n. intros; subst; auto. Qed. Lemma eq_le' : forall (m n:nat), m = n -> n <= m. intros; subst; auto. Qed. Lemma eq_plus_l : forall m n, n = n + m -> m = 0. intros; omega. Qed. Lemma eq_plus_r : forall m n, n = m + n -> m = 0. intros; omega. Qed. Lemma lt_gt : forall m n, m < n -> n > m. auto. Qed. Lemma gt_lt : forall m n, m > n -> n < m. auto. Qed. Lemma le_ge : forall m n, m <= n -> n >= m. auto. Qed. Lemma ge_le : forall m n, m >= n -> n <= m. auto. Qed. Lemma nlt_ngt : forall m n, ~ m < n -> ~ n > m. auto. Qed. Lemma ngt_nlt : forall m n, ~ m > n -> ~ n < m. auto. Qed. Lemma nle_nge : forall m n, ~ m <= n -> ~ n >= m. auto. Qed. Lemma nge_nle : forall m n, ~ m >= n -> ~ n <= m. auto. Qed. Lemma lt_not_le : forall m n, m < n -> ~n <= m. intros; omega. Qed. Lemma not_le_lt : forall m n, ~n <= m -> m < n. intros; omega. Qed. Lemma gt_not_ge : forall m n, m > n -> ~n >= m. intros; omega. Qed. Lemma not_ge_gt : forall m n, ~n >= m -> m > n. intros; omega. Qed. Lemma lt_impl_le : forall m n, m < n -> m <= n. auto with arith. Qed. Lemma lt_neq : forall m n, m < n -> m <> n. intros; intro; subst; apply lt_irrefl in H; auto. Qed. Lemma gt_neq : forall m n, m > n -> m <> n. intros; intro; subst; apply lt_irrefl in H; auto. Qed. Lemma neq_lt0 : forall n, 0 <> n -> 0 < n. auto with arith. Qed. Lemma lt_neq_impl : forall m n, m < n -> m <> pred n -> m < pred n. intros; omega. Qed. Lemma neq_dec : forall m n, m <> n -> {m < n} + {n < m}. intros m n h1. destruct (lt_eq_lt_dec m n) as [[h2 | h2] | h2]; subst. left; auto. contradict h1; auto. right; auto. Qed. Lemma ge1_lt : forall (m w:nat), 1 <= m < 1 + w -> m < S w. intros; omega. Qed. Lemma neq0_iff : forall (n:nat), n <> 0 <-> 0 < n. intro n; destruct n; split; auto with arith. Qed. Lemma not_not0_iff : forall (n:nat), ~ n <> 0 <-> n = 0. intro n; destruct n; omega. Qed. Lemma not_lt0_iff : forall (n:nat), ~ 0 < n <-> n = 0. intro; omega. Qed. Lemma not_lt_le : forall (m n:nat), ~ n < m -> m <= n. intros; omega. Qed. Lemma not_not_nat : forall (m n:nat), ~ n <> m -> n = m. intros; omega. Qed. Lemma nat_eq_neq : forall (m n p:nat), m = n /\ n <> p -> m <> p. intros; omega. Qed. Lemma nat_eq_neq_iff : forall (m n p:nat), m = n /\ n <> p <-> m = n /\ m <> p. intros; split; omega. Qed. Lemma S_eq_iff : forall m n:nat, S m = n <-> m = pred n /\ n > 0. intros; split; omega. Qed. Lemma lt_eq_trans : forall (m n p:nat), m < n -> n = p -> m < p. intros; subst; auto with arith. Qed. Lemma le_eq_trans : forall (m n p:nat), m <= n -> n = p -> m <= p. intros; subst; auto with arith. Qed. Lemma eq_lt_trans : forall (m n p:nat), m = n -> n < p -> m < p. intros; subst; auto. Qed. Lemma eq_le_trans : forall (m n p:nat), m = n -> n <= p -> m <= p. intros; subst; auto. Qed. Lemma mult2 : forall n:nat, 2 * n = n + n. intro; omega. Qed. Lemma lt_mod : forall k n (pfn : 0 < n), k mod n < n. intros k n h1. pose proof (mod_bound_pos k n (le_0_n _) h1) as h2. destruct h2; auto with arith. Qed. Lemma lt_mod' : forall r k n (pfn : r < n), k mod n < n. intros; apply lt_mod; omega. Qed. Lemma le_mod : forall k n, k mod n <= k. intros k n. destruct (zerop n) as [|h1]; subst; simpl; auto with arith. apply neq0_iff in h1. pose proof (div_mod k n h1) as h2. pose proof (le_plus_l (k mod n) (n * (k / n))) as h3. rewrite plus_comm in h3. rewrite <- h2 in h3; auto. Qed. Lemma le_pred_sum : forall i n m, i < n -> i <= m + (pred n). intros; omega. Qed. Lemma le_pred_sum' : forall i n m, i < n -> i <= m + n - 1. intros; omega. Qed. Lemma le_pred_sum_minus : forall i m n, i < n -> m <= m + (pred n) - (pred i). intros; omega. Qed. (*This is the symmetric equality-flip of [minus_n_n].*) Lemma minus_same : forall n, n - n = 0. auto with arith. Qed. Lemma sum_pred_minus : forall i m n, 0 < i -> i < n -> m + pred n - pred i - m = n - i. intros; omega. Qed. Lemma S_pred_plus_minus : forall n i, 0 < i -> i < n -> S (pred i) + (n - i) = n. intros; omega. Qed. Lemma le_pred_sum_pos : forall m n, 0 < m -> 0 < n -> m <= pred (m + n). intros; omega. Qed. Lemma S_pred_sum_minus : forall m n, 0 < m -> 0 < n -> S (pred (m + n) - m) = n. intros; omega. Qed. Lemma O_mod : forall n, 0 mod n = 0. intro n. induction n; simpl; auto with arith. Qed. Lemma O_div : forall n, 0 / n = 0. intro n. induction n; simpl; auto with arith. Qed. Lemma plus_S_shift_l : forall (m n:nat), (S m + n) = S (m + n). intros; omega. Qed. Lemma S_compat : forall (m:nat), m + 1 = S m. intros m'. induction m' as [|m h1]; auto with arith. rewrite <- h1 at 2. rewrite h1. rewrite plus_comm. rewrite plus_Snm_nSm. auto with arith. Qed. Lemma plus_S_shift_r : forall (m n:nat), (m + S n) = S (m + n). intros m n. assert (h1:S n = 1 + n). auto with arith. rewrite h1. rewrite plus_permute. rewrite plus_comm. simpl. auto with sets. apply S_compat. Qed. Lemma plus_eq_iff : forall i j, i = i + j <-> j = 0. intros; omega. Qed. Lemma minus_eq_iff : forall i j, i = i - j <-> i = 0 \/ j = 0. intros; omega. Qed. Lemma le1_dec : forall n, n <= 1 -> {n = 0} + {n = 1}. intros n h1. destruct n as [|n]. left; auto. apply le_S_n in h1. apply le0 in h1; subst. right; auto. Qed. Lemma le2_dec : forall n, n <= 2 -> {n = 0} + {n = 1} + {n = 2}. intros n h1. destruct n as [|n]. left; auto. apply le_S_n in h1. apply le1_dec in h1. destruct h1; subst. left; right; auto. right; auto. Qed. Lemma S_plus_S : forall m n, S m + S n = S (S (m + n)). intros; omega. Qed. End Trivialities. Lemma mono_pred_lt : forall (m n:nat), m < n -> 0 < m -> pred m < pred n. intros m n h1 h2. apply lt_S_n. rewrite <- (S_pred _ _ h2); auto. assert (h3:0 < n). eapply lt_trans. apply h2. assumption. rewrite <- (S_pred _ _ h3). assumption. Qed. Section Minus. Lemma zero_minus : forall (m:nat), 0 - m = 0. intro; auto with arith. Qed. Lemma lt_minus_anti : forall i j n, i < j -> j <= n -> n - i > n - j. intros; omega. Qed. Lemma minus_gt_anti : forall i j n, j <= n -> n - i > n - j -> i < j. intros; omega. Qed. Lemma gt_minus_anti : forall i j n, i > j -> i <= n -> n - i < n - j. intros; omega. Qed. Lemma minus_lt_anti : forall i j n, i <= n -> n - i < n - j -> i > j. intros; omega. Qed. Lemma le_minus_anti : forall i j n, i <= j -> n - i >= n - j. intros; omega. Qed. Lemma minus_ge_anti : forall i j n, i <= n -> n - i >= n - j -> i <= j. intros; omega. Qed. Lemma ge_minus_anti : forall i j n, i >= j -> n - i <= n - j. intros; omega. Qed. Lemma minus_le_anti : forall i j n, j <= n -> n - i <= n - j -> i >= j. intros; omega. Qed. Lemma S_minus_S : forall (m n:nat), S m - S n = m - n. intro m. induction m as [|m ih]. intros; omega. intro n. specialize (ih n). omega. Qed. Lemma one_minus : forall m, 0 < m -> 1 - m = 0. intro m. destruct m as [|m]. intro h1. apply lt_irrefl in h1; contradiction. intro. rewrite S_minus_S. apply zero_minus. Qed. Lemma S_minus_shift : forall (m n:nat), n <= m -> (S m - n) = S (m - n). intros m n. revert m. induction n as [|n ih]. intro; omega. intros m h1. destruct m as [|m]. omega. apply le_S_n in h1. specialize (ih _ h1). do 2 rewrite S_minus_S; auto. Qed. Lemma le_S_minus_shift : forall (m n:nat), n <= m -> S (m - n) <= S m - n. intros; omega. Qed. Lemma S_minus_pred : forall m n, S n <= m -> m - S n = pred (m - n). intros; omega. Qed. Lemma le_minus_iff : forall (m n:nat), m <= n <-> m - n = 0. intuition. Qed. Lemma minus_plus_le_cancel : forall m n p, m <= n -> n = m + p -> n - m = p. intros; omega. Qed. Lemma minus_plus_lt_cancel : forall m n p, m < n -> n = m + p -> n - m = p. intros; omega. Qed. Lemma lt_iff : forall p q, p < q <-> q - p > 0. intuition. Qed. Lemma minus_lt_compat_r : forall i j n, n < j -> i < j -> i - n < j - n. intros; omega. Qed. Lemma plus_minus_assoc : forall (m n p:nat), p <= n -> m + (n - p) = m + n - p. intros m n p. induction m as [|m h1]; destruct n as [|n]; destruct p as [|p]; simpl; auto. intro h2. omega. intro h2. specialize (h1 h2). assert (h3:p <= n). omega. clear h2. rewrite S_minus_S in h1. rewrite h1. rewrite minus_Sn_m. rewrite S_minus_S. reflexivity. omega. Qed. Lemma minus_plus_assoc : forall (m n p:nat), n + p <= m -> m - (n + p) = m - n - p. intro m. induction m as [|m ih]. intros; do 2 rewrite zero_minus; auto. intros n p h1. apply le_lt_eq_dec in h1. destruct h1 as [h1 | h1]. rewrite S_minus_shift; auto with arith. specialize (ih n p). hfirst ih. auto with arith. specialize (ih hf). rewrite ih. rewrite S_minus_shift, <- (S_minus_shift (m - n) p); auto; omega. omega. Qed. Lemma minus_minus_assoc : forall (m n p:nat), n - p <= m -> p <= n -> n <= m -> m - (n - p) = m - n + p. intros; omega. Qed. Lemma plus_minus_same_r : forall m n:nat, m + n - m = n. intros; omega. Qed. Lemma plus_minus_same_l : forall m n:nat, n + m - m = n. intros; omega. Qed. Lemma plus_minus_same' : forall m n:nat, m <= n -> m + (n - m) = n. intros; omega. Qed. Lemma minus_plus_same : forall m n, m <= n -> n - m + m = n. intros; omega. Qed. Lemma minus_inj : forall (m n p:nat), n <= m -> p <= m -> m - n = m - p -> n = p. intro m. induction m as [|m ih]. intros n p h1 h2. apply le0 in h1. apply le0 in h2; subst; auto. intros n p h1 h2 h3. apply le_lt_eq_dec in h1. apply le_lt_eq_dec in h2. destruct h1 as [h1|], h2 as [h2|]; subst. red in h1, h2. apply le_S_n in h1. apply le_S_n in h2. specialize (ih _ _ h1 h2). do 2 rewrite S_minus_shift in h3; auto. omega. omega. reflexivity. Qed. Lemma le_minus_inj : forall (m n p:nat), p <= m -> p <= n -> m - p <= n - p -> m <= n. intros; omega. Qed. Lemma minus_inj_neq : forall (m n p:nat), m - n <> m - p -> n <> p. intros m n p h1 h2. subst. contradict h1; auto. Qed. Lemma le_minus_0 : forall m n, m <= n -> m - n = 0. intros; omega. Qed. Lemma minus_same_diff : forall i j n, i <= n -> j <= n -> n - i - (n - j) = j - i. intros; omega. Qed. Lemma le_plus_minus_same_iff : forall i j k, i + j - k <= i <-> j <= k. intros; omega. Qed. Lemma plus_minus_lt : forall i j k, j < k -> i + j - k = i - (k - j). intros; omega. Qed. Lemma plus_minus_le : forall i j k, j <= k -> i + j - k = i - (k - j). intros; omega. Qed. Lemma minus_minus_lt : forall i j k, j < k -> i - k - j = (i - j) - k. intros; omega. Qed. Lemma minus_minus_le : forall i j k, j <= k -> i - k - j = (i - j) - k. intros; omega. Qed. Lemma minus_minus_lt' : forall i j k, j < i -> i - k - j = (i - j) - k. intros; omega. Qed. Lemma minus_minus_le' : forall i j k, j <= i -> i - k - j = (i - j) - k. intros; omega. Qed. Lemma plus_minus_minus_lt : forall i j k m, m < i -> i + j - k - m = i - m + j - k. intros; omega. Qed. Lemma plus_minus_minus_le : forall i j k m, m <= i -> i + j - k - m = i - m + j - k. intros; omega. Qed. Lemma plus_minus_minus_same : forall i j k, i + j - i - k = j - k. intros; omega. Qed. Lemma plus_minus_minus_same': forall i j k, i + j - k - i = j - k. intros; omega. Qed. Lemma plus_minus_minus_same'': forall i j k, i + j - k - j = i - k. intros; omega. Qed. Lemma plus_minus_eq_iff : forall i j k, i = i + j - k <-> (j < k /\ i = 0) \/ j = k. intros; omega. Qed. Lemma plus_minus_plus_same_lt : forall i j k, k < j -> i + j - (i + k) = j - k. intros; omega. Qed. Lemma plus_minus_plus_same_le : forall i j k, k <= j -> i + j - (i + k) = j - k. intros; omega. Qed. Lemma plus_minus_plus_same_lt' : forall i j k, k < i -> i + j - (k + j) = i - k. intros; omega. Qed. Lemma plus_minus_plus_same_le' : forall i j k, k <=i -> i + j - (k + j) = i - k. intros; omega. Qed. Lemma minus_minus_plus_minus_same_le : forall i j k, j <= k -> k <= i + j -> i - (i + j - k) = k - j. intros; omega. Qed. Lemma plus_minus_anti : forall i j k, j <= i -> i - j + k + j - i = k. intros; omega. Qed. Lemma minus_plus_le : forall i j k, i - j + k <= i + k. intros; omega. Qed. Lemma plus_minus_minus_same_le : forall i j k, k <= i -> i - k <= j -> i + (j - (i - k)) = j + k. intros; omega. Qed. Lemma plus_plus_plus_minus_same : forall i j k, i + j + k - i = j + k. intros; omega. Qed. Lemma plus_minus_plus_diff_lt : forall i j k m, k < i -> m < j -> i + j - (k + m) = (i - k) + (j - m). intros; omega. Qed. Lemma plus_minus_plus_diff_le : forall i j k m, k <= i -> m <= j -> i + j - (k + m) = (i - k) + (j - m). intros; omega. Qed. Lemma plus_minus_plus_diff_lt_le : forall i j k m, k < i -> m <= j -> i + j - (k + m) = (i - k) + (j - m). intros; omega. Qed. Lemma plus_minus_plus_diff_le_lt : forall i j k m, k <= i -> m < j -> i + j - (k + m) = (i - k) + (j - m). intros; omega. Qed. Lemma plus_minus_plus_diff_lt' : forall i j k m, k < i -> j < m -> i + j - (k + m) = (i - k) - (m - j). intros; omega. Qed. Lemma plus_minus_plus_diff_le' : forall i j k m, k <= i -> j <= m -> i + j - (k + m) = (i - k) - (m - j). intros; omega. Qed. Lemma plus_minus_plus_diff_lt_le' : forall i j k m, k < i -> j <= m -> i + j - (k + m) = (i - k) - (m - j). intros; omega. Qed. Lemma plus_minus_plus_diff_le_lt' : forall i j k m, k <= i -> j < m -> i + j - (k + m) = (i - k) - (m - j). intros; omega. Qed. Lemma minus_lt_compat : forall i j k m, i < j -> k < m -> m < i -> i - m < j - k. intros; omega. Qed. Lemma minus_le_compat : forall i j k m, i <= j -> k <= m -> i - m <= j - k. intros; omega. Qed. Lemma minus_le_lt_compat : forall i j k m, i <= j -> k < m -> m <= i -> i - m < j - k. intros; omega. Qed. Lemma minus_lt_le_compat : forall i j k m, i < j -> k <= m -> m <= i -> i - m < j - k. intros; omega. Qed. Lemma minus_minus_minus_diff_le : forall i j k m, k <= i -> j <= m -> j <= i -> m <= k -> k - m <= i - j -> i - j - (k - m) = i - k + (m - j). intros; omega. Qed. (*I'm spelling out a lot of omegas for efficiency!*) Lemma minus_minus_eq_iff : forall m n p, m - n = m - p <-> (m <= n /\ m <= p \/ ((m > n /\ p <= m \/ m > p /\ n <= m) /\ n = p)). intros m n p. revert n m. induction p as [|p ih]; simpl; intros m n. rewrite <- minus_n_O, eq_sym_iff, minus_eq_iff. split; intro h1. destruct h1; subst. destruct (zerop m) as [|hn]; left; auto with arith. destruct (zerop n) as [|hn]; subst. left; auto with arith. right; auto with arith. destruct h1 as [h1|h1]. destruct h1 as [? h1]. apply le0 in h1; left; auto. destruct h1; right; auto. destruct m as [|m], n as [|n]; split; intro h1; auto with arith. rewrite <- minus_n_O, minus_eq_iff in h1. destruct h1; discriminate. destruct h1 as [h1|h1]. destruct h1 as [h1 h2 h1]. apply le_Sn_0 in h1; contradiction. destruct h1; discriminate. do 2 rewrite S_minus_S in h1. rewrite ih in h1. intuition. do 2 rewrite S_minus_S; rewrite ih; intuition. Qed. (*Am spelling out a slow [intros; omega]!*) Lemma minus_minus_minus_diff_le' : forall i j k m, k <= i -> m <= j -> j <= i -> m <= k -> k - m <= i - j -> i - j - (k - m) = i - k - (j - m). intros i j k m h1 h2 h3 h4 h5. pose proof (minus_minus_eq_iff i (j + (k - m)) (k + (j - m))) as h6. do 2 rewrite minus_plus_assoc in h6. rewrite h6. destruct (le_lt_dec i (j +(k - m))) as [h7 | h7]. left. split; auto. rewrite plus_minus_assoc; auto. rewrite plus_minus_assoc in h7; auto. rewrite plus_comm; auto. right. split. left. split; auto. rewrite plus_minus_assoc; auto. rewrite plus_minus_assoc in h7; auto. rewrite plus_comm; auto with arith. rewrite plus_minus_assoc; auto. rewrite (plus_minus_assoc k j m); auto. rewrite plus_comm; auto. pose proof (plus_le_compat_l _ _ j h5) as h7. rewrite plus_minus_assoc in h7; auto. rewrite plus_comm in h7. rewrite <- plus_minus_assoc in h7; auto. rewrite (plus_minus_assoc j i j) in h7; auto. rewrite <- (plus_comm i j) in h7. rewrite <- (plus_minus_assoc i j j) in h7; auto. rewrite minus_same, <- plus_n_O in h7; auto. pose proof (plus_le_compat_l _ _ j h5) as h7. rewrite plus_minus_assoc in h7; auto. rewrite <- plus_minus_assoc in h7; auto. rewrite (plus_minus_assoc j i j) in h7; auto. rewrite <- (plus_comm i j) in h7. rewrite <- (plus_minus_assoc i j j) in h7; auto. rewrite minus_same, <- plus_n_O in h7; auto. pose proof (plus_le_compat_l _ _ j h5) as h7. rewrite plus_minus_assoc in h7; auto. rewrite <- plus_minus_assoc in h7; auto. rewrite (plus_minus_assoc j i j) in h7; auto. rewrite <- (plus_comm i j) in h7. rewrite <- (plus_minus_assoc i j j) in h7; auto. rewrite minus_same, <- plus_n_O in h7; auto. Qed. Lemma minus_minus_plus_same_lt : forall i j k, k < j -> j - k < i -> i - (j - k) + j = i + k. intros; omega. Qed. Lemma minus_minus_plus_same_le : forall i j k, k <= j -> j - k <= i -> i - (j - k) + j = i + k. intros; omega. Qed. Lemma minus_plus_plus_same_lt : forall i j k, j < i -> i - j + k + j = i + k. intros; omega. Qed. Lemma minus_plus_plus_same_le : forall i j k, j <= i -> i - j + k + j = i + k. intros; omega. Qed. End Minus. Section S. Lemma S_minus : forall n, S n - n = 1. intros; omega. Qed. Lemma S_inj : forall (m n:nat), S m = S n -> m = n. intros ? ? h1; injection h1; auto. Qed. Lemma S_inj_neq : forall (m n:nat), S m <> S n -> m <> n. intros m n h1 h2. subst. contradict h1; auto. Qed. Lemma le_times_S : forall i j, i <= i * S j. intros; rewrite <- mult_1_r at 1; apply mult_le_compat; auto with arith. Qed. Lemma S_eq_plus : forall m n, S m = S m + n -> m = m + n. intros; omega. Qed. End S. Section Dec. Lemma nat_eq_dec : type_eq_dec nat. red; intros m n. destruct (lt_eq_lt_dec m n) as [h1 | h1]. destruct h1 as [h1 | h1]. right; intro; subst. apply lt_irrefl in h1; auto. subst. left; auto. right; intro; subst. apply lt_irrefl in h1; auto. Qed. Lemma neq_nat_dec : forall (m n:nat), m <> n -> {m < n} + {n < m}. intros m. induction m as [|m ih]. intros; left; auto with arith. intros n h1. destruct n as [|n]. right; auto with arith. apply S_inj_neq in h1. specialize (ih _ h1). destruct ih; [left; auto with arith | right; auto with arith]. Qed. Lemma lt_dec : forall m n, m < n -> {m < pred n} + {m = pred n}. intros m n h1. destruct n as [|n]. apply lt_n_O in h1; contradiction. simpl. red in h1. apply le_S_n in h1. apply le_lt_eq_dec in h1; auto. Qed. Lemma le_dec' : forall m n, m <= n -> {m <= pred n} + {m = n}. intros m n. revert m. induction n as [|n ih]. intros m h1; apply le0 in h1; subst. right; auto. simpl; intros m h1. destruct m as [|m]. left; auto with arith. apply le_S_n in h1. destruct (zerop n) as [|h3]; subst. right; auto with arith. specialize (ih _ h1). destruct ih as [h2|]; subst. left; omega. right; auto. Qed. Lemma pred0_dec : forall n:nat, pred n = O -> {n = 0} + {n = 1}. intro n. induction n; simpl; auto; omega. Qed. Lemma lt_le_dec : forall m n:nat, {m < n} + {n <= m}. intros m n. destruct (lt_eq_lt_dec m n) as [h1 | h1]. destruct h1 as [h1 | h1]. left; auto. subst. right; auto with arith. right; auto with arith. Qed. Lemma gt_ge_dec : forall m n:nat, {m > n} + {n >= m}. intros; apply lt_le_dec. Qed. Lemma lt_S_dec : forall m n:nat, m < S n -> {m < n} + {m = n}. intros m n h1. destruct (lt_le_dec m n) as [h2 | h2]. left; auto. right. omega. Qed. Lemma lt_ex_dec : forall (P:nat->Prop) (n:nat), pred_dec P -> {exists m:nat, m < n /\ P m} + {~exists m:nat, m < n /\ P m}. intros P n. induction n as [|n ih]; simpl. intros; right; intro h1. destruct h1 as [? [? ?]]; omega. intro h1. specialize (ih h1). destruct ih as [h2 | h2]. left. destruct h2 as [m [h2 h4]]. exists m. split; auto with arith. specialize (h1 n). destruct h1 as [h1 | h1]. left. exists n. split; auto with arith. right. intro h3. destruct h3 as [m [h3 h4]]. destruct (nat_eq_dec m n) as [|h5]; subst. contradiction. contradict h2. exists m. split; try omega; auto with arith. Qed. Lemma lt_ex_dep_nat_dec : forall {n:nat} (P:pred_dep_nat n), pred_dep_nat_dec P -> {exists (m:nat) (pfm:m < n), P m pfm} + {~exists (m:nat) (pfm:m < n), P m pfm}. intro n. induction n as [|n ih]; simpl. intros; right; intro h1. destruct h1 as [? [? ?]]; omega. intros P h1. specialize (ih (pred_dep_nat_pred P) (pred_dep_nat_dec_pred _ h1)). destruct ih as [h2 | h2]. left. destruct h2 as [m [h2 h3]]. exists m, (lt_trans _ _ _ h2 (lt_n_Sn _)). red in h3; auto. destruct (h1 n (lt_n_Sn _)) as [h3 | h3]. left. exists n, (lt_n_Sn _); auto. right. intro h4. destruct h4 as [m [h4 h5]]. red in h4. pose proof (le_S_n _ _ h4) as h6. destruct (le_lt_eq_dec _ _ h6) as [h7 |]; subst. contradict h2. exists m, h7; red. erewrite f_equal_dep. apply h5. reflexivity. contradict h3. erewrite f_equal_dep. apply h5. reflexivity. Qed. End Dec. Section Bound. (*This sets a weak bound [b] for the identity function on nat. So, [(nat_bound_weak b):nat->nat, i => i], for [i <= b]. So [nat_bound_weak 3 1 = 1], [nat_bound_weak 3 2 = 2], [nat_bound 3 3 = 3], [nat_bound 3 8 = 3].*) Fixpoint nat_bound_weak b : nat -> nat := match b with | 0 => fun _ => 0 | S b' => fun i => match i with | 0 => 0 | S i' => S (nat_bound_weak b' i') end end. (*This sets a strong bound [b] for the identity function on nat. So, [(nat_bound_strong b):nat->nat, i => i], for [i < b]. So [nat_bound 3 1 = 1], [nat_bound 3 2 = 2], [nat_bound 3 3 = 2], [nat_bound 3 8 = 2].*) Definition nat_bound b : nat -> nat := fun i => nat_bound_weak (pred b) i. Lemma nat_bound_weak_eq_same_iff : forall n i, nat_bound_weak n i = i <-> i <= n. intro n; induction n as [|n ih]; simpl. intuition. intro i; destruct i as [|i]; simpl. intuition. split; intro h1. apply S_inj in h1; apply le_n_S. rewrite ih in h1; auto. apply le_S_n in h1; f_equal. rewrite <- ih in h1; auto. Qed. Lemma nat_bound_eq_same_iff : forall n i, nat_bound n i = i <-> i < n \/ (n = 0 /\ i = 0). unfold nat_bound; intros; rewrite nat_bound_weak_eq_same_iff. omega. Qed. Lemma nat_bound_weak_eq_iff : forall n i m, nat_bound_weak n i = m <-> (i <= n /\ i = m \/ i > n /\ n = m). intro n; induction n as [|n ih]; simpl. intros i m; split; intro h1; subst. destruct i as [|i]. left; auto. right; auto with arith. destruct h1 as [h1|h1]. destruct h1; subst. apply le0 in H; auto. tauto. intro i. destruct i as [|i]. intro m. split; intro; subst. left; auto with arith. destruct H; omega. intro m. destruct m as [|m]. intuition. split; intro h1. apply S_inj in h1; subst. specialize (ih i n). destruct (le_lt_dec i n) as [h1|h1]. left. split; auto with arith. f_equal. rewrite <- nat_bound_weak_eq_same_iff in h1; auto. right. split; auto with arith. f_equal. symmetry. rewrite ih. tauto. f_equal. destruct h1 as [h1|h1]. destruct h1 as [h1 h2]. apply S_inj in h2; subst. apply le_S_n in h1. rewrite <- nat_bound_weak_eq_same_iff in h1; auto. destruct h1 as [h1 h2]. apply S_inj in h2; subst. rewrite ih. right; auto with arith. Qed. Lemma nat_bound_eq_iff : forall n i m, nat_bound n i = m <-> (i < n /\ i = m \/ i >= n /\ pred n = m). unfold nat_bound; intros n i m; rewrite nat_bound_weak_eq_iff; omega. Qed. Lemma le_bound : forall n i, nat_bound_weak n i <= n. unfold nat_bound; intro n; induction n as [|n ih]; simpl; auto with arith; intro i; destruct i; auto with arith. Qed. Lemma lt_bound : forall n i, nat_bound (S n) i < S n. unfold nat_bound; intros; red; apply le_n_S; simpl; apply le_bound. Qed. End Bound. Section Inj. Lemma pred_inj0: forall (m n:nat), 0 < m -> 0 < n -> pred m = pred n -> m = n. intros m n; destruct m, n; omega. Qed. Lemma pred_inj1: forall (m n:nat), m <> 1 -> n <> 1 -> pred m = pred n -> m = n. intros m n; destruct m, n; omega. Qed. Lemma pred_inj': forall (m n:nat), pred m > 0 -> pred m = pred n -> m = n. intros m n h0. induction m as [|m h1]; destruct n as [|n]. auto. simpl in h0. omega. simpl. intro h2. subst. simpl in h0. omega. intro h2. simpl in h2. f_equal. auto. Qed. Lemma pred_inj_lt : forall (m n:nat), pred m < pred n -> m < n. intros m n h1. destruct m as [|m]; destruct n as [|n]; simpl; auto. simpl in h1. apply lt_n_0 in h1. contradiction. simpl in h1. apply lt_n_S; auto. Qed. Lemma lt_S_pred : forall m n, S m < n -> m < pred n. intros; omega. Qed. Lemma lt_S_pred' : forall m n, 0 < m -> m < S n -> pred m < n. intros; omega. Qed. Lemma lt_S_neq : forall m n, m < S n -> m <> n -> m < n. intros; omega. Qed. Lemma pred_pred_lt : forall n:nat, 0 < n -> pred (pred n) < n. intros; omega. Qed. Lemma pred_minus_pred_pred_lt : forall n, n < 2 -> pred n - (pred (pred n)) = 0. intro; omega. Qed. Lemma pred_minus_pred_pred_ge : forall n, n >= 2 -> pred n - (pred (pred n)) = 1. intro; omega. Qed. Lemma pred_minus : forall n m, m <= n -> pred (n - m) = pred n - m. intros; omega. Qed. Lemma plus_inj_l : forall m n p:nat, m + n = m + p -> n = p. intros; omega. Qed. Lemma plus_inj_r : forall m n p:nat, m + p = n + p -> m = n. intros; omega. Qed. End Inj. Lemma le_prop_seg_one_t : forall (m n:nat), m <= n -> forall a:nat, 1 <= a <= m -> 1 <= a <= n. intros; omega. Qed. Lemma lt_prop_seg_one_t : forall (m n:nat), m < n -> forall a:nat, 1 <= a <= m -> 1 <= a <= n. intros; omega. Qed. Lemma lt_le_prop_seg_one_t : forall (m n:nat), m < n -> forall a:nat, m < a <= n -> 1 <= a <= n. intros; omega. Qed. Lemma lt_le_prop_seg_one_t' : forall (m n:nat), m <= n -> forall a:nat, m < a <= n -> 1 <= a <= n. intros; omega. Qed. Lemma inj_minus : forall (m n p:nat), p <= m -> p <= n -> m - p = n - p -> m = n. intros m n p. revert m n. induction p as [|p ih]; simpl. intros ? ? ? ? h3. do 2 rewrite <- minus_n_O in h3; auto. intros m n h1 h2 h3. do 2 rewrite S_minus_pred in h3; auto. apply pred_inj0 in h3; omega. Qed. (*This is a bit more handy than [NPeano]'s existential version [div], in that you avoid an extra variable and existential. Also, not too many properties were derived in that version, so I lose little.*) Inductive dv' (a:nat) : nat -> Prop := dv_intro : forall n, dv' a (a * n). Infix "dv" := dv' (at level 40, no associativity). Section Times. Lemma mult_S_lt_l : forall i j k, S i * j < S i * k -> j < k. intros i j k h1. destruct (lt_eq_lt_dec j k) as [[h2 | h2] | h2]; auto. subst. apply lt_irrefl in h1; contradiction. pose proof (mult_lt_compat_l _ _ (S i) h2 (lt_0_Sn _)) as h3. pose proof (lt_trans _ _ _ h1 h3) as h4. apply lt_irrefl in h4. contradiction. Qed. Lemma mult_S_lt_r : forall i j k, j * S i < k * S i -> j < k. intros i j k h1. destruct (lt_eq_lt_dec j k) as [[h2 | h2] | h2]; auto. subst. apply lt_irrefl in h1; contradiction. pose proof (mult_lt_compat_r _ _ (S i) h2 (lt_0_Sn _)) as h3. pose proof (lt_trans _ _ _ h1 h3) as h4. apply lt_irrefl in h4. contradiction. Qed. Lemma lt_mult_iff : forall i j, i * j < i <-> (j = 0 /\ 0 < i). intros i j; split; intro h1. destruct j as [|j]. rewrite mult_0_r in h1; auto. pose proof (le_times_S i j) as h2. pose proof (le_lt_trans _ _ _ h2 h1) as h3. apply lt_irrefl in h3; contradiction. destruct h1 as [? h1]; subst. rewrite mult_0_r; auto. Qed. Lemma S_times_l : forall x y, S x * y = x * y + y. intros x y. rewrite <- S_compat. rewrite mult_plus_distr_r. rewrite mult_1_l; auto. Qed. Lemma S_times_l' : forall x y, S x * y = y + x * y. intros; rewrite plus_comm; apply S_times_l. Qed. Lemma S_times_r : forall x y, y * S x = y * x + y. intros x y. rewrite <- S_compat. rewrite mult_plus_distr_l. rewrite mult_1_r; auto. Qed. Lemma S_times_r' : forall x y, y * S x = y + y * x. intros; rewrite plus_comm; apply S_times_r. Qed. Lemma S_times_S : forall x y, S x * S y = S (x * y + x + y). intros x y. rewrite <- S_compat. rewrite <- (S_compat y). rewrite mult_plus_distr_r. rewrite mult_1_l. rewrite mult_plus_distr_l. rewrite S_compat, mult_1_r. rewrite <- plus_S_shift_r; auto. Qed. Lemma minus_dist_r : forall i j m, i * m - j * m = (i - j) * m. intros i j m. revert i j. induction m as [|m ih]. intros. do 3 rewrite mult_0_r; auto with arith. intros i j. destruct (lt_eq_lt_dec i j) as [[h1 | h1] | h1]. pose proof (mult_S_lt_compat_l m i j h1) as h2. rewrite (mult_comm _ i), (mult_comm _ j) in h2. apply lt_le_weak in h2. rewrite le_minus_iff in h2. rewrite h2. apply lt_le_weak in h1. rewrite le_minus_iff in h1. rewrite h1. rewrite mult_0_l; auto. subst. do 2 rewrite <- minus_n_n. rewrite mult_0_l; auto. pose proof (mult_S_lt_compat_l m j i h1) as h2. rewrite mult_comm, (mult_comm j), (mult_comm (i - j)). do 3 rewrite S_times_l. specialize (ih i j). rewrite mult_comm, (mult_comm j), (mult_comm (i - j)) in ih. rewrite <- ih. rewrite plus_minus_assoc; auto with arith. destruct m as [|m]. do 2 rewrite mult_0_l. do 2 rewrite plus_0_l; auto with arith. pose proof (plus_minus_assoc i (S m*i) (S m*j)) as h3. hfirst h3. apply mult_le_compat_l; auto with arith. specialize (h3 hf). omega. Qed. Lemma minus_dist_l : forall i j m, m * i - m * j = m * (i - j). intros i j m. rewrite (mult_comm m i), (mult_comm m j), (mult_comm m (i - j)). apply minus_dist_r. Qed. Lemma times_S_eq_0_iff : forall i j, i * S j = 0 <-> i = 0. intro i. destruct i as [|i]. intro. rewrite mult_0_l; tauto. intro j. rewrite S_times_S. split; intro h1; discriminate h1. Qed. Lemma S_times_eq_iff : forall i j k, S i * j = S i * k <-> j = k. intros i j k; split; intro h1. destruct (lt_eq_lt_dec j k) as [[h2 | h2] | h2]. pose proof (mult_S_lt_compat_l i j k h2) as h3. rewrite h1 in h3. apply lt_irrefl in h3; contradiction. assumption. pose proof (mult_S_lt_compat_l i k j h2) as h3. rewrite <- h1 in h3. apply lt_irrefl in h3; contradiction. subst; auto. Qed. Lemma lt_times_S_plus_S : forall x y z, x < x * S y + S z. intro x. induction x as [|x ih]; simpl; auto with arith. intros y z. apply lt_n_S. rewrite (plus_comm y (x * S y)). rewrite <- plus_assoc. rewrite plus_S_shift_r. apply ih. Qed. Lemma lt_plus_S : forall x y, x < x + S y. intros; omega. Qed. Lemma times_le_iff : forall i j, i * j <= i <-> (i = 0 \/ j = 0 \/ j = 1). intros i j; split; intro h1. destruct i as [|i]. left; auto. right. destruct j as [|j]. left; auto. right. rewrite S_times_S in h1. apply le_S_n in h1. f_equal. destruct j as [|j]; auto. rewrite <- plus_assoc in h1. rewrite plus_S_shift_r in h1. pose proof (lt_times_S_plus_S i j (i + j)). omega. destruct h1 as [h1 | h1]; subst. rewrite mult_0_l; auto with arith. destruct h1; subst. rewrite mult_0_r; auto with arith. rewrite mult_1_r; auto with arith. Qed. Lemma eq_times_iff : forall x y, x = x * y <-> x = 0 \/ (x = 0 /\ y = 0) \/ y = 1. intros x y; split; intro h1. destruct x as [|x]. left; auto. right. destruct y as [|y]. left; auto. rewrite mult_0_r in h1. discriminate. right. f_equal. rewrite S_times_S in h1. apply S_inj in h1. destruct y as [|y]; auto. rewrite <- plus_assoc in h1. rewrite (plus_S_shift_r x y) in h1. pose proof (lt_times_S_plus_S x y (x + y)). omega. destruct h1 as [h1 | [[h0 h1] | h1]]; subst. rewrite mult_0_l; auto. rewrite mult_0_r; auto. rewrite mult_1_r; auto. Qed. Lemma eq_times_plus_iff : forall x y z, x = x * y + z <-> (x = 0 /\ z = 0) \/ (y = 1 /\ z = 0) \/ (y = 0 /\ x = z). intros x y z; split; intro h1. destruct x as [|x]. left. rewrite mult_0_l in h1. rewrite plus_0_l in h1; auto. right. destruct y as [|y]. right. rewrite mult_0_r in h1. rewrite plus_0_l in h1; auto. destruct z as [|z]. rewrite plus_0_r in h1. rewrite eq_times_iff in h1. destruct h1 as [h1 | h1]. discriminate. destruct h1 as [h1 | h1]. destruct h1; discriminate. left; auto. rewrite S_times_S in h1. rewrite plus_S_shift_l in h1. apply S_inj in h1. rewrite <- (mult_1_r x) in h1 at 3. rewrite <- mult_plus_distr_l in h1. rewrite S_compat in h1. rewrite <- plus_assoc in h1. rewrite plus_S_shift_r in h1. pose proof (lt_times_S_plus_S x y (y + z)). omega. destruct h1 as [[h0 h1] | [[h0 h1] | h1]]; subst. rewrite mult_0_l; auto with arith. rewrite mult_1_r; auto with arith. destruct h1; subst. rewrite mult_0_r; auto. Qed. Lemma one_eq_S_times_plus_iff : forall x y z, 1 = (S x) * y + z <-> (x = 0 /\ ((y = 0 /\ z = 1) \/ (y = 1 /\ z = 0))) \/ (y = 0 /\ z = 1). intros x y z; split; intro h1. destruct x as [|x]. left. rewrite mult_1_l in h1. split; auto. destruct y as [|y]. left. split; auto with arith. right. destruct z as [|z]. rewrite plus_0_r in h1. split; auto. omega. right. destruct y as [|y]. rewrite mult_0_r in h1. rewrite plus_0_l in h1; auto. rewrite S_times_S in h1. rewrite plus_S_shift_l in h1. apply S_inj in h1. rewrite <- plus_assoc in h1. rewrite <- plus_assoc in h1. rewrite plus_S_shift_l in h1. rewrite plus_comm in h1. rewrite plus_S_shift_l in h1. discriminate. destruct h1 as [h1 | h1]. destruct h1 as [? h1]; subst. destruct h1 as [[? ?] | [? ?]]; subst. rewrite mult_0_r; auto with arith. rewrite mult_1_r, plus_0_r; auto with arith. destruct h1; subst. rewrite mult_0_r; auto with arith. Qed. Lemma minus_eq_times_iff : forall i j k, i - j = i * k <-> (i = 0 \/ (j = 0 /\ (i = 0 \/ (i = 0 /\ k = 0) \/ k = 1)) \/ (k = 0 /\ i <= j)). intros i j k; split; intro h1. destruct i as [|i]. left; auto. right. destruct j as [|j]. rewrite <- minus_n_O in h1. rewrite eq_times_iff in h1. destruct h1 as [h1 | h1]. discriminate. destruct h1 as [h1 | h1]. destruct h1; discriminate. subst. left. split; auto. right. pose proof (le_minus (S i) (S j)) as h2. rewrite h1 in h2. rewrite times_le_iff in h2. destruct h2 as [h2 | h2]. discriminate. destruct h2; subst. rewrite mult_0_r in h1. rewrite <- le_minus_iff in h1; auto. rewrite mult_1_r in h1. symmetry in h1. rewrite minus_eq_iff in h1. destruct h1; discriminate. destruct h1 as [h1 | h1]; subst. rewrite zero_minus, mult_0_l; auto. destruct h1 as [h1 | h1]. destruct h1 as [? h1]; subst. destruct h1 as [|h1]; subst. rewrite <- minus_n_n, mult_0_l; auto. destruct h1 as [h1 | h1]. destruct h1; subst. rewrite <- minus_n_n, mult_0_l; auto. subst. rewrite minus_n_O, mult_1_r; auto. destruct h1 as [? h1]; subst. rewrite mult_0_r. rewrite <- le_minus_iff; auto. Qed. Lemma minus_eq_times_plus_iff : forall i j k m, i - j = i * k + m <-> (i = 0 /\ m = 0) \/ (j = 0 /\ ((i = 0 /\ m = 0) \/ (k = 1 /\ m = 0) \/ (k = 0 /\ i = m))) \/ (k = 0 /\ m = i - j) \/ (m = 0 /\ (i = 0 \/ (j = 0 /\ (i = 0 \/ (i = 0 /\ k = 0) \/ k = 1)) \/ (k = 0 /\ i <= j))). intros i j k m. destruct m as [|m]. rewrite plus_0_r. rewrite minus_eq_times_iff. split; intro h1. right. right. right. split; auto. destruct h1 as [h1 | h1]. destruct h1; subst. left; auto. destruct h1 as [h1 | h1]. destruct h1 as [? h1]; subst. destruct h1 as [h1 | h1]. destruct h1; subst. left; auto. destruct h1 as [h1 | h1]. destruct h1; subst. right. left. split; auto. destruct h1; subst. left; auto. destruct h1 as [h1 | h1]. destruct h1 as [? h1]; subst. right. right. symmetry in h1. rewrite <- le_minus_iff in h1; auto. destruct h1 as [h1 h2]. destruct h2 as [h2 | h2]; subst. left; auto. destruct h2 as [h2 | h2]. destruct h2 as [? h2]; subst. destruct h2 as [h2 | h2]; subst. left; auto. destruct h2 as [h2 | h2]. destruct h2; subst. left; auto. subst. right. left. split; auto. auto. split; intro h1. right. destruct j as [|j]. rewrite <- minus_n_O in h1. rewrite eq_times_plus_iff in h1. destruct h1 as [h1 | h1]. destruct h1; discriminate. destruct h1 as [h1 | h1]. destruct h1; discriminate. destruct h1; subst. left. auto. right. destruct k as [|k]. rewrite mult_0_r, plus_0_l in h1. left; auto. right. pose proof (lt_times_S_plus_S i k m) as h2. pose proof (le_minus i (S j)) as h3. pose proof (le_lt_trans _ _ _ h3 h2) as h4. rewrite h1 in h4. apply lt_irrefl in h4; contradiction. destruct h1 as [h1 | h1]. destruct h1; discriminate. destruct h1 as [h1 | h1]. destruct h1 as [? h1]; subst. destruct h1 as [h1 | h1]. destruct h1; discriminate. destruct h1 as [h1 | h1]. destruct h1; discriminate. destruct h1; subst. rewrite minus_n_O, mult_0_r, plus_0_l; auto. destruct h1 as [h1 | h1]. destruct h1 as [? h1]; subst. rewrite mult_0_r, plus_0_l; auto. destruct h1; discriminate. Qed. Lemma times_eq_times_iff : forall i j k, i * j = i * k <-> i = 0 \/ j = k. intros i j k. destruct i as [|i]. do 2 rewrite mult_0_l; tauto. split; intro h1. right. rewrite S_times_eq_iff in h1; auto. destruct h1 as [h1|]; try discriminate h1; subst; auto. Qed. Lemma mult_inj : forall i j n, 0 < n -> n * i = n * j -> i = j. intros i j n h1 h2. destruct (lt_eq_lt_dec i j) as [[h3 | h3] | h3]. pose proof (mult_lt_compat_l i j n h3 h1) as h4. rewrite h2 in h4. apply lt_irrefl in h4; contradiction. assumption. pose proof (mult_lt_compat_l j i n h3 h1) as h4. rewrite h2 in h4. apply lt_irrefl in h4; contradiction. Qed. Section Divides. Lemma dv0 : forall n, n dv 0. intro n. rewrite <- (mult_0_r n). constructor. Qed. Lemma one_dv : forall n, 1 dv n. intro n. rewrite <- (mult_1_l). constructor. Qed. Lemma dv_refl : forall a, a dv a. intro a. rewrite <- mult_1_r at 2. constructor. Qed. Lemma dv_trans : forall a b c, a dv b -> b dv c -> a dv c. intros a b c h1 h2. destruct h1, h2. rewrite <- mult_assoc. constructor. Qed. Lemma dv_antisym : forall a b, a dv b -> b dv a -> a = b. intros a b h1 h2. destruct h1. inversion h2 as [c h3]. rewrite h3. rewrite <- mult_assoc in h3. symmetry in h3. rewrite eq_times_iff in h3. destruct h3 as [h3 | h3]; subst. rewrite mult_0_l; auto. destruct h3 as [h3 | h3]. destruct h3; subst. rewrite mult_0_l; auto. apply mult_is_one in h3. destruct h3; subst. rewrite mult_1_r; auto. Qed. Lemma dv_le : forall a b, a dv b -> b = 0 \/ a <= b. intros a b h1. destruct h1. destruct n as [|n]. left. rewrite mult_0_r; auto. right. pose proof (mult_O_le a (S n)) as h1. destruct h1 as [h1 | h1]. discriminate. rewrite mult_comm; auto. Qed. Lemma times_eq_times_plus_div : forall i j k m, i * j = i * k + m -> i dv m. intros i j k m h1. destruct i as [|i]. do 2 rewrite mult_0_l in h1. rewrite plus_0_l in h1; subst. rewrite <- mult_1_r. constructor. destruct (lt_eq_lt_dec j k) as [[h2 | h2] | h2]; subst. pose proof (mult_S_lt_compat_l i j k h2) as h3. rewrite h1 in h3. destruct m as [|m]. rewrite plus_0_r in h3. apply lt_irrefl in h3; contradiction. omega. destruct m as [|m]. rewrite <- (mult_0_r (S i)). constructor. omega. pose proof (mult_S_lt_compat_l i k j h2) as h3. pose proof (minus_plus_lt_cancel _ _ _ h3 h1); subst. rewrite (mult_comm _ j), (mult_comm _ k). rewrite minus_dist_r. rewrite mult_comm. constructor. Qed. Lemma mult_dv_iff : forall i n, i * n dv n <-> n = 0 \/ i = 1. intros i n; split; intro h1. inversion h1. rewrite <- mult_assoc, mult_comm in H. symmetry in H. rewrite <- mult_assoc in H. rewrite eq_times_iff in H. destruct H as [h2 | h2]; subst. left. rewrite mult_0_r, mult_0_l; auto. destruct h2 as [h2 | h2]. destruct h2 as [? h2]; subst. left. rewrite mult_0_r, mult_0_l; auto. apply mult_is_one in h2. destruct h2; subst. right; auto. destruct h1; subst. rewrite mult_0_r. rewrite <- (mult_1_r 0) at 2. constructor. rewrite mult_1_l. rewrite <- (mult_1_r n) at 2. constructor. Qed. Lemma plus_mult_dv_iff : forall i j n, (i + j * n) dv n <-> i = 0 /\ (j = 1 \/ n = 0) \/ (j = 0 \/ n = 0) /\ i dv n. intros i j n. destruct i as [|i]. rewrite plus_0_l. split; intro h1. rewrite mult_dv_iff in h1. destruct h1; subst. left. split; auto. left. split; auto. destruct h1 as [h1 | h1]. destruct h1 as [? h1]. destruct h1; subst. rewrite mult_1_l. apply dv_refl. rewrite mult_0_r. apply dv_refl. destruct h1 as [? h1]; subst. destruct H; subst. rewrite mult_0_l; auto. rewrite mult_0_r; auto. split; intro h1. right. inversion h1. symmetry in H. rewrite mult_plus_distr_r, plus_comm in H. rewrite (mult_comm j), <- mult_assoc in H. rewrite eq_times_plus_iff in H. destruct H as [h2 | h2]. destruct h2; subst. apply mult_is_O in H0. destruct H0; subst. discriminate. split. right. rewrite mult_0_r; auto. rewrite mult_0_r; auto. apply dv0. destruct h2 as [h2 | h2]. destruct h2 as [h2 h3]. apply mult_is_one in h2. destruct h2; subst. rewrite mult_1_r in h3. discriminate. destruct h2 as [h2 h3]. apply mult_is_O in h2. destruct h2; subst. split. left; auto. rewrite mult_0_l, plus_0_r. constructor. split. right. rewrite mult_0_r; auto. rewrite mult_0_r; auto. apply dv0. destruct h1 as [h1 | h1]. destruct h1; discriminate. destruct h1 as [h1 h2]. destruct h1; subst. rewrite mult_0_l, plus_0_r; auto. apply dv0. Qed. Lemma dv_plus : forall i m n, i dv m -> i dv n -> i dv (m + n). intros i j n h1 h2. destruct h1, h2. rewrite <- mult_plus_distr_l. constructor. Qed. Lemma dv_minus : forall i m n, i dv m -> i dv n -> i dv (m - n). intros i j n h1 h2. destruct h1, h2. rewrite minus_dist_l. constructor. Qed. Lemma gt_not_div : forall i n, n > 0 -> i > n -> ~i dv n. intros i n h0 h1 h2. destruct h2. destruct n as [|n]. rewrite mult_0_r in h0. apply lt_irrefl in h0; auto. rewrite <- (mult_1_r i) in h1 at 1. destruct i as [|i]. rewrite mult_0_l in h0. apply lt_irrefl in h0; auto. apply mult_S_lt_l in h1. apply lt_S_n in h1. apply lt_n_0 in h1; auto. Qed. Lemma times_div : forall i j, 0 < i -> 0 < j -> i * j / i = j. intros i j h1 h2. rewrite (S_pred _ _ h1). simpl. pose proof (divmod_spec (j + pred i * j) (pred i ) 0 (pred i) (le_refl _)) as h3. destruct (divmod (j + pred i * j) (pred i ) 0 (pred i)) as [q u]. rewrite <- minus_n_n, mult_0_r, plus_0_r, plus_0_r, <- (S_pred _ _ h1) in h3. destruct h3 as [h3 h4]. simpl. rewrite <- (mult_1_l j) in h3 at 1. rewrite <- mult_plus_distr_r in h3. rewrite plus_comm in h3. rewrite S_compat in h3. rewrite <- (S_pred _ _ h1) in h3. pose proof h3 as h5. apply times_eq_times_plus_div in h5. pose proof (dv_le _ _ h5) as h6. destruct h6 as [h6 | h6]. rewrite h6, plus_0_r in h3. apply mult_inj in h3; auto. omega. Qed. Lemma times_mod_r : forall i n, (i * n) mod n = 0. intros i n. destruct n as [|n]; simpl; auto. pose proof (divmod_spec (i * S n) n 0 n (le_refl _)) as h1. destruct (divmod (i * S n) n 0 n) as [q u]. simpl. rewrite mult_0_r, <- minus_n_n, plus_0_r, plus_0_r in h1. destruct h1 as [h1 h2]. rewrite mult_comm in h1 at 1. pose proof h1 as h3. apply times_eq_times_plus_div in h3. pose proof (dv_le (S n) (n - u) h3) as h4. destruct h4 as [|h4]; auto with arith. omega. Qed. Lemma times_mod_l : forall i n, (n * i) mod n = 0. intros; rewrite mult_comm, times_mod_r; auto. Qed. Lemma dv_iff : forall m n, 0 < m -> (m dv n <-> (n mod m = 0)). intros m n h0; split; intro h1. destruct h1. rewrite mult_comm. apply times_mod_r. rewrite <- neq0_iff in h0. pose proof (div_mod n m h0) as h2. rewrite h1, plus_0_r in h2. rewrite h2. constructor. Qed. Lemma mod_eq0_iff : forall m n, m mod n = 0 <-> n = 0 \/ n dv m. intros m n. destruct (zerop n) as [|h1]; subst; simpl. tauto. rewrite <- dv_iff; auto. split; intro h2. right; auto. destruct h2; subst; auto. apply lt_irrefl in h1; contradiction. Qed. Lemma mod1 : forall n, n mod 1 = 0. intro n. rewrite mod_eq0_iff. right. apply one_dv. Qed. Lemma times_eq_iff : forall i j k, 0 < i -> 0 < j -> (i * j = k <-> (j = k / i /\ i dv k)). intros i j k h1; split; intro h2. subst. rewrite times_div; auto. split; auto. constructor. destruct h2; subst. destruct H1. rewrite times_div; auto. destruct n; simpl in H. rewrite mult_0_r, O_div in H. apply lt_irrefl in H. contradiction. auto with arith. Qed. Lemma div_eq_0_iff : forall i n, i / n = 0 <-> n = 0 \/ i < n. intros i n; split; intro h1. destruct n as [|n]. left; auto. simpl in h1. right. pose proof (divmod_spec i n 0 n (le_refl _)) as h2. destruct (divmod i n 0 n) as [q u]. simpl in h1; subst. rewrite <- minus_n_n, mult_0_r, plus_0_r, plus_0_r, plus_0_l in h2. destruct h2; subst. omega. destruct h1 as [|h0]; subst. simpl; auto. destruct n as [|n]. simpl; auto. simpl. pose proof (divmod_spec i n 0 n (le_refl _)) as h2. destruct (divmod i n 0 n) as [q u]. rewrite <- minus_n_n, mult_0_r, plus_0_r, plus_0_r in h2. destruct h2 as [h2 h3]. subst. simpl. pose proof (le_plus_trans (S n * q) (S n * q) (n - u) (le_refl _)) as h4. pose proof (le_lt_trans _ _ _ h4 h0) as h5. rewrite lt_mult_iff in h5. destruct h5; auto. Qed. Lemma lt_mod_eq : forall i n, i < n -> i mod n = i. intros i n h1. pose proof (div_eq_0_iff i n) as h3. hr h3. right; auto. rewrite <- h3 in hr. pose proof (div_mod i n) as h4. destruct (zerop n) as [|h5]; subst. apply lt_n_0 in h1. contradiction. rewrite <- neq0_iff in h5. specialize (h4 h5). rewrite hr, mult_0_r, plus_0_l in h4; auto. Qed. Lemma eq_div_iff : forall i j n, i = j / n <-> n = 0 /\ i = 0 \/ 0 < n /\ j - j mod n = i * n. intros i j n; split; intro h1; subst. destruct (zerop n) as [|h1]; subst. simpl. left; auto. right. split; auto. rewrite (S_pred _ _ h1) at 1 2. simpl. pose proof (divmod_spec j (pred n) 0 (pred n) (le_refl _)) as h2. destruct (divmod j (pred n) 0 (pred n)) as [q u]. simpl. rewrite <- (S_pred _ _ h1), <- minus_n_n, plus_0_r, mult_0_r, plus_0_r in h2. destruct h2 as [h0 h2]; subst. rewrite <- (plus_minus_assoc (n*q) (pred n - u) _ (le_refl _)). rewrite <- minus_n_n. rewrite plus_0_r. rewrite mult_comm; auto. destruct h1 as [h1 | h1]. destruct h1; subst; simpl; auto. destruct h1 as [h1 h2]. rewrite (S_pred _ _ h1) in h2 at 1. rewrite (S_pred _ _ h1). simpl in h2; simpl. pose proof (divmod_spec j (pred n) 0 (pred n) (le_refl _)) as h3. destruct (divmod j (pred n) 0 (pred n)) as [q u]. simpl. simpl in h2, h3. rewrite <- minus_n_n, plus_0_r, mult_0_r, plus_0_r in h3. destruct h3 as [h3 h4]. rewrite h3 in h2. rewrite <- (plus_minus_assoc (q + pred n*q) (pred n - u) _ (le_refl _)) in h2. rewrite <- minus_n_n, plus_0_r in h2. rewrite <- (mult_1_l q) in h2 at 1. rewrite <- mult_plus_distr_r in h2. rewrite plus_comm in h2. rewrite S_compat in h2. rewrite <- (S_pred _ _ h1) in h2. rewrite (mult_comm i) in h2. eapply mult_inj in h2; auto. Qed. Lemma times_div_eq_iff : forall i n, n * (i / n) = i <-> n dv i. intros i n; split; intro h1. unfold div in h1. destruct n as [|n]. rewrite mult_0_r in h1; subst. apply dv0. pose proof (divmod_spec i n 0 n (le_refl _)) as h2. destruct (divmod i n 0 n) as [q u]. rewrite <- h1. constructor. destruct h1. destruct n as [|n], n0 as [|n0]. simpl; auto. simpl; auto. rewrite mult_0_r, O_div, mult_0_r; auto. rewrite times_div; auto with arith. Qed. End Divides. Lemma le_div : forall i j n, i <= j -> i / n <= j / n. intros i j n h1. unfold div. destruct n as [|n]; auto with arith. pose proof (divmod_spec i n 0 n (le_refl _)) as h2. pose proof (divmod_spec j n 0 n (le_refl _)) as h3. destruct (divmod i n 0 n) as [q u], (divmod j n 0 n) as [q' u']. rewrite mult_0_r, plus_0_r, <- (minus_n_n), plus_0_r in h2, h3. destruct h2 as [h2 hu], h3 as [h3 hu']; subst. simpl. assert (h2:n - u <= n). omega. assert (h3:n - u' <= n). omega. pose proof (plus_le_compat_r _ _ (S n * q') h3) as h4. rewrite plus_comm in h4. pose proof (le_trans _ _ _ h1 h4) as h5. pose proof (minus_le_compat_r _ _ (n - u) h5) as h6. rewrite <- plus_minus_assoc, <- minus_n_n, plus_0_r in h6; auto. rewrite plus_comm in h6. rewrite minus_minus_assoc, <- plus_minus_assoc, <- minus_n_n, plus_0_r in h6; auto with arith. apply le_lt in hu. pose proof (plus_lt_compat_r _ _ (S n * q') hu) as h7. rewrite <- (mult_1_r (S n)) in h7 at 2. rewrite <- mult_plus_distr_l in h7. rewrite plus_comm in h6. pose proof (le_lt_trans _ _ _ h6 h7) as h8. apply mult_S_lt_l in h8; auto with arith. pose proof (le_minus n u) as h7. pose proof (le_plus_l n (S n * q')) as h8. rewrite plus_comm in h8; auto with arith. eapply le_trans. apply h7. assumption. Qed. Lemma times_eq_times_plus_iff : forall i j k m, i * j = i * k + m <-> ((i = 0 \/ j = k) /\ m = 0) \/ i > 0 /\ j > k /\ m = i * (j - k). intros i j k m. destruct (zerop i) as [|h1]; subst. do 2 rewrite mult_0_l. rewrite plus_0_l. split; intro h1; subst. left. split; auto. destruct h1 as [h1 | h1]; auto. destruct h1; auto. destruct h1 as [h1]. apply lt_irrefl in h1; contradiction. split; intro h2. destruct (lt_eq_lt_dec j k) as [[h3 | h3] | h3]. pose proof (mult_S_lt_compat_l (pred i) _ _ h3) as h4. rewrite <- (S_pred _ _ h1) in h4. right. omega. subst. left. split; auto. omega. right. split; auto. pose proof (mult_S_lt_compat_l (pred i) _ _ h3) as h4. rewrite <- (S_pred _ _ h1) in h4. pose proof (f_equal (fun c => minus c (i * k)) h2) as h5. simpl in h5. rewrite plus_comm in h5. rewrite <- plus_minus_assoc in h5; auto with arith. rewrite <- minus_n_n, plus_0_r in h5. rewrite mult_comm, (mult_comm i k), minus_dist_r in h5. rewrite times_eq_iff in h5. destruct h5 as [h5 h6]. inversion h6. subst. destruct (zerop n) as [|hn]; subst. rewrite mult_0_r in h1. rewrite O_div in h1. apply lt_irrefl in h1. contradiction. rewrite times_div in h1; auto. split; auto. rewrite times_div. apply mult_comm. omega. assumption. omega. omega. assumption. destruct h2 as [h2 | h2]. destruct h2 as [[? | ?] ?]; subst. apply lt_irrefl in h1; contradiction. auto with arith. destruct h2 as [h2 [h3 h4]]. apply (inj_minus _ _ (i * k)). apply mult_le_compat_l; auto with arith. auto with arith. rewrite (mult_comm i j), (mult_comm i k). rewrite minus_dist_r. rewrite (mult_comm (j - k)). rewrite <- h4. rewrite plus_comm. rewrite <- plus_minus_assoc; auto with arith. rewrite <- minus_n_n, plus_0_r. reflexivity. Qed. Lemma times_eq_times_minus_iff : forall i j k m, i * j = i * k - m <-> i = 0 \/ (j = 0 /\ i * k <= m) \/ ((i > 0 /\ j > 0) /\ (j = k /\ m = 0 \/ j < k /\ m = i * (k - j))). intros i j k m. destruct (zerop i) as [|h1]; subst. do 2 rewrite mult_0_l. simpl. tauto. destruct (zerop j) as [|hj]; subst. rewrite mult_0_r, <- minus_n_O. split; intro h2. symmetry in h2. rewrite <- le_minus_iff in h2. right. left; auto. destruct h2 as [h2 | h2]. subst. apply lt_irrefl in h1; contradiction. destruct h2 as [h2 | h2]. destruct h2 as [? h2]. rewrite le_minus_iff in h2; auto. destruct h2 as [hg [h2 | h2]]. destruct hg as [? hg]. apply lt_irrefl in hg; contradiction. destruct hg as [? hg]. apply lt_irrefl in hg. contradiction. split; intro h2. destruct (lt_eq_lt_dec j k) as [[h3 | h3] | h3]. pose proof (mult_S_lt_compat_l (pred i) _ _ h3) as h4. rewrite <- (S_pred _ _ h1) in h4. right. right. split; auto. right. split; auto. pose proof (f_equal (fun c => c + m) h2) as h5. simpl in h5. rewrite <- minus_minus_assoc, <- minus_n_n, <- minus_n_O in h5. rewrite plus_comm in h5. pose proof (f_equal (fun c => c - i * j) h5) as h6. simpl in h6. rewrite <- plus_minus_assoc, <- minus_n_n, plus_0_r in h6. rewrite minus_dist_l in h6; auto. auto. rewrite <- minus_n_n; auto with arith. auto. destruct (le_lt_dec m (i * k) ) as [h6 | h6]; auto. apply lt_le_weak in h6. rewrite le_minus_iff in h6. rewrite h6, plus_0_l in h5. symmetry in h5. rewrite plus_comm, plus_eq_iff in h5. apply mult_is_O in h5; destruct h5; subst. apply lt_irrefl in h1. contradiction. apply lt_irrefl in hj; contradiction. subst. rewrite minus_eq_iff in h2. destruct h2 as [h2 | h2]. apply mult_is_O in h2. destruct h2; subst. apply lt_irrefl in h1; contradiction. apply lt_irrefl in hj; contradiction. subst. right. right. split; auto. right. right. split; auto. right. pose proof (mult_lt_compat_l _ _ i h3 h1) as h4. pose proof (le_minus (i * k) m) as h5. rewrite <- h2 in h5. pose proof (lt_le_trans _ _ _ h4 h5) as h6. apply lt_irrefl in h6; contradiction. destruct h2 as [? | h2]; subst. apply lt_irrefl in h1; contradiction. destruct h2 as [h2 | h2]. destruct h2 as [? h2]; subst. apply lt_irrefl in hj; contradiction. destruct h2 as [hg [h2 | h2]]. destruct h2 as [? h2]; subst. rewrite <- minus_n_O; auto. destruct h2 as [? h2]; subst. rewrite <- minus_dist_l. rewrite minus_minus_assoc. rewrite <- minus_n_n, plus_0_l; auto. apply le_minus. apply mult_le_compat_l; auto with arith. auto. Qed. Lemma times_eq_times_plus_minus_iff : forall i j k m n, i * j = i * k + m - n <-> (i = 0 /\ m <= n) \/ (j = k /\ m = n) \/ (j = 0 /\ i * k + m <= n) \/ (i > 0 /\ j > 0 /\ ((j > k /\ (n < m /\ m = i * (j - k) + n)) \/ (j < k /\ (m <= n /\ n = i * (k - j) + m)))). intros i j k m n; split; intro h1. destruct (zerop i) as [|h2]; subst. do 2 rewrite mult_0_l in h1. rewrite plus_0_l in h1. symmetry in h1. rewrite <- le_minus_iff in h1. left. split; auto. destruct (zerop j) as [|h3]; subst. rewrite mult_0_r in h1. symmetry in h1. rewrite <- le_minus_iff in h1. right. right. left. split; auto. destruct (lt_eq_lt_dec j k) as [h4 | h4]. destruct h4 as [h4 | h4]; subst. destruct (le_lt_dec m n) as [h5 | h5]. pose proof h4 as h4'. rewrite lt_iff in h4; auto. pose proof (f_equal (fun c => c - (i * j)) h1) as h9. simpl in h9. rewrite <- minus_n_n in h9. rewrite plus_minus_le in h9; auto. rewrite minus_minus_lt' in h9. rewrite minus_dist_l in h9. symmetry in h9. rewrite <- le_minus_iff in h9. right. right. right. split; auto. split; auto. destruct (le_lt_dec (i *k + m) n) as [hl | hl]; auto. rewrite le_minus_iff in hl. rewrite hl in h1. apply mult_is_O in h1. destruct h1; subst. apply lt_irrefl in h2; contradiction. apply lt_irrefl in h3; contradiction. right. split; auto. split; auto. pose proof (f_equal (minus (i * k)) h1) as h10. rewrite minus_minus_plus_minus_same_le in h10; auto with arith. rewrite minus_dist_l in h10. pose proof (f_equal (fun c => c + m) h10) as he. simpl in he. rewrite minus_plus_same in he; auto with arith. rewrite (S_pred _ _ h2). apply mult_S_lt_compat_l; auto. pose proof (f_equal (fun c => c - (i * j)) h1) as h9. simpl in h9. rewrite <- minus_n_n in h9; auto. rewrite plus_minus_minus_lt in h9. rewrite minus_dist_l in h9. rewrite <- plus_minus_assoc in h9; auto with arith. symmetry in h9. apply plus_is_O in h9. destruct h9 as [h9 h10]. apply mult_is_O in h9. destruct h9 as [|h9]. subst. apply lt_irrefl in h2; contradiction. rewrite <- le_minus_iff in h9, h10. pose proof (lt_le_trans _ _ _ h4 h9) as h11. apply lt_irrefl in h11. contradiction. rewrite (S_pred _ _ h2). apply mult_S_lt_compat_l; auto. rewrite plus_minus_eq_iff in h1. destruct h1 as [h1 | h1]. destruct h1 as [h1 h4]. apply mult_is_O in h4. destruct h4; subst. apply lt_irrefl in h2; contradiction. apply lt_irrefl in h3; contradiction. subst. right. left; auto. right. right. right. split; auto. split; auto. left; auto. split; auto. destruct (le_lt_dec m n) as [h5 | h5]. rewrite plus_minus_le in h1; auto. pose proof (f_equal (fun c => c + (n - m)) h1) as h6. simpl in h6. rewrite <- minus_minus_assoc, minus_same, <- minus_n_O in h6; auto. pose proof (f_equal (minus (i * j)) h6) as h7. pose proof (le_plus_l (i * j) (n - m)) as h8. rewrite le_minus_iff in h8. rewrite h8 in h7. symmetry in h7. rewrite <- le_minus_iff in h7. rewrite (S_pred _ _ h2) in h7. apply mult_S_le_reg_l in h7. pose proof (lt_le_trans _ _ _ h4 h7) as h9. apply lt_irrefl in h9; contradiction. rewrite minus_same; auto with arith. destruct (le_lt_dec (n - m) (i * k)) as [h7 | h7]; auto. apply lt_le_weak in h7. rewrite le_minus_iff in h7. rewrite h7 in h1. apply mult_is_O in h1. destruct h1; subst. apply lt_irrefl in h2; contradiction. apply lt_irrefl in h3; contradiction. split; auto. rewrite <- plus_minus_assoc in h1; auto with arith. pose proof (f_equal (fun c => c - i * k) h1) as h6. simpl in h6. rewrite plus_comm in h6. rewrite <- plus_minus_assoc, minus_same, plus_0_r in h6; auto. rewrite minus_dist_l in h6. rewrite h6. rewrite minus_plus_same; auto with arith. destruct h1 as [h1 | h1]; subst. destruct h1 as [? h1]; subst. do 2 rewrite mult_0_l. rewrite plus_0_l. rewrite le_minus_iff in h1; auto. destruct h1 as [h1 | h1]. destruct h1; subst. rewrite plus_minus_le; auto. rewrite minus_same, <- minus_n_O; auto. destruct h1 as [h1 | h1]. destruct h1 as [? h1]; subst. rewrite mult_0_r. rewrite le_minus_iff in h1; auto. destruct h1 as [h1 [h2 h3]]. destruct h3 as [h3 | h3]. destruct h3 as [h3 [h4 h5]]. pose proof (f_equal (fun c => c - n) h5) as h6. simpl in h6. rewrite plus_minus_same_l in h6. rewrite <- plus_minus_assoc; auto with arith. rewrite h6. rewrite <- minus_dist_l. rewrite plus_minus_same'; auto. apply mult_le_compat_l; auto with arith. destruct h3 as [h3 [h4 h5]]. rewrite h5. rewrite <- minus_dist_l. rewrite minus_plus_assoc. rewrite plus_minus_minus_same''. rewrite minus_minus_assoc, <- minus_n_n; auto with arith. apply le_minus. apply minus_plus_le. Qed. Lemma times_plus_eq_times_plus_iff : forall i j k m n, i * j + m = i * k + n <-> ((i = 0 \/ j = k) /\ m = n) \/ (i > 0 /\ (j < k /\ m = i*(k - j) + n) \/ (k < j /\ n = i*(j - k) + m)). intros i j k m n. destruct (zerop i) as [|hi]; subst. do 2 rewrite mult_0_l, plus_0_l. split; intro h1; subst. left. split; auto. destruct h1 as [h1 | h1]. destruct h1; subst. reflexivity. destruct h1 as [h1 | h1]. destruct h1 as [h1 h2]. apply lt_irrefl in h1; contradiction. rewrite mult_0_l, plus_0_l in h1. destruct h1; auto. split; intro h1. destruct (lt_eq_lt_dec j k) as [[h2 | h2] | h2]. right. left. split; auto with arith. split; auto. apply (plus_inj_r _ _ (i*j)). rewrite plus_comm, h1. rewrite <- plus_assoc, (plus_comm n (i*j)), plus_assoc. rewrite <- mult_plus_distr_l. rewrite <- minus_minus_assoc, <- minus_n_n, <- minus_n_O; auto with arith. rewrite <- minus_n_n; auto with arith. subst. left. split. right; auto. omega. right. right. split; auto. rewrite mult_comm. rewrite <- minus_dist_r. rewrite plus_comm. rewrite plus_minus_assoc. rewrite mult_comm, plus_comm. rewrite h1. rewrite plus_comm, <- plus_minus_assoc. rewrite mult_comm, <- minus_n_n, plus_0_r; auto. rewrite mult_comm; auto with arith. apply mult_le_compat_r; auto with arith. destruct h1 as [h1 | h1]. destruct h1 as [[? | ?] ?]; subst. apply lt_irrefl in hi; contradiction. reflexivity. destruct h1 as [h1 | h1]. destruct h1 as [h1 [h2 h3]]; subst. rewrite (mult_comm _ (k - j)), <- minus_dist_r. rewrite plus_assoc. rewrite plus_minus_assoc. rewrite (plus_comm (i*j) (k*i)). rewrite <- plus_minus_assoc. rewrite (mult_comm i j), <- minus_n_n, plus_0_r. rewrite mult_comm; auto. rewrite mult_comm; auto. apply mult_le_compat_r; auto with arith. destruct h1 as [h1]; subst. rewrite (mult_comm _ (j - k)). rewrite <- minus_dist_r. rewrite plus_assoc, plus_minus_assoc. rewrite (plus_comm (i*k)). rewrite <- plus_minus_assoc. rewrite (mult_comm i k), <- minus_n_n, plus_0_r, mult_comm; auto. rewrite mult_comm; auto. apply mult_le_compat_r; auto with arith. Qed. End Times. Lemma dv_plus_rev_r : forall i m n, i dv m -> i dv (m + n) -> i dv n. intros i m n h1 h2. destruct h1. inversion h2. rewrite times_eq_times_plus_iff in H0. destruct H0 as [h3 | h3]. destruct h3 as [[? | ?] ?]; subst. apply dv0. apply dv0. destruct h3 as [? [? ?]]; subst. constructor. Qed. Lemma dv_plus_rev_l : forall i m n, i dv n -> i dv (m + n) -> i dv m. intros; eapply dv_plus_rev_r. apply H. rewrite plus_comm; auto. Qed. Lemma dv_minus_rev_l : forall i m n, i dv m -> i dv (m - n) -> i = 0 \/ m <= n \/ i dv n. intros i m n h1 h2. destruct h1. inversion h2. rewrite times_eq_times_minus_iff in H0. destruct H0 as [|h3]; subst. left ;auto. destruct h3 as [h3 | h3]. destruct h3 as [? h3]; subst. right. left; auto. destruct h3 as [h3 h4]. destruct h3 as [h3 h5]. destruct h4 as [h4 | h4]. destruct h4; subst. right. right. apply dv0. right. destruct h4 as [h4]; subst. right. constructor. Qed. Lemma dv_minus_rev_r : forall i m n, i dv n -> i dv (m - n) -> m <= n \/ i dv m. intros i m n h1 h2. destruct h1. inversion h2. pose proof (f_equal (fun c => c + (i * n)) H0) as h3. simpl in h3. destruct (le_lt_dec (i * n) m) as [h4 | h4]. rewrite <- minus_minus_assoc, <- minus_n_n, <- minus_n_O in h3; subst; auto with arith. rewrite <- mult_plus_distr_l. right. constructor. rewrite <- minus_n_n; auto with arith. apply lt_le_weak in h4. rewrite le_minus_iff in h4. symmetry in h3. rewrite h4, plus_0_l, plus_comm, plus_eq_iff in h3. apply mult_is_O in h3; destruct h3; subst. rewrite mult_0_l, <- minus_n_O in h4; subst. right. apply dv0. rewrite mult_0_r in H0. symmetry in H0. rewrite <- le_minus_iff in H0. left; auto. Qed. Lemma not_dv_plus_l : forall i m n, ~i dv m -> i dv n -> ~i dv (m + n). intros i m n h1 h2 h3. pose proof (dv_plus_rev_l _ _ _ h2 h3) as h4; contradiction. Qed. Lemma not_dv_plus_r : forall i m n, ~i dv n -> i dv m -> ~i dv (m + n). intros i m n h1 h2 h3. pose proof (dv_plus_rev_r _ _ _ h2 h3) as h4; contradiction. Qed. Lemma dv_not_plus_l : forall i m n, i dv m -> ~i dv (m + n) -> ~i dv n. intros i m n h1 h2 h3. pose proof (dv_plus _ _ _ h1 h3); contradiction. Qed. Lemma dv_not_plus_r : forall i m n, i dv n -> ~i dv (m + n) -> ~i dv m. intros i m n h1 h2 h3. pose proof (dv_plus _ _ _ h3 h1); contradiction. Qed. Lemma not_dv_minus_l : forall i m n, ~i dv m -> i dv n -> m <= n \/ ~i dv (m - n). intros i m n h1 h2. destruct (le_lt_dec m n) as [h3 | h3]. left; auto. right. intro h4. pose proof (dv_minus_rev_r _ _ _ h2 h4) as h5. destruct h5 as [h5 | h5]. pose proof (lt_le_trans _ _ _ h3 h5) as h6. apply lt_irrefl in h6; contradiction. contradiction. Qed. Lemma not_dv_minus_r : forall i m n, ~i dv n -> i dv m -> i = 0 \/ m <= n \/ ~i dv (m - n). intros i m n h1 h2. destruct (zerop i) as [|hz]. left; auto. right. destruct (le_lt_dec m n) as [h3 | h3]. left; auto. right. intro h4. pose proof (dv_minus_rev_l _ _ _ h2 h4) as h5. destruct h5 as [h5 | h5]. subst. apply lt_irrefl in hz; contradiction. destruct h5 as [h5 | h5]. pose proof (le_lt_trans _ _ _ h5 h3) as h6. apply lt_irrefl in h6; auto. contradiction. Qed. Lemma dv_not_minus_l : forall i m n, i dv m -> ~i dv (m - n) -> ~i dv n. intros i m n h1 h2 h3. pose proof (dv_minus _ _ _ h1 h3) as h4. contradiction. Qed. Lemma dv_not_minus_r : forall i m n, i dv n -> ~i dv (m - n) -> ~i dv m. intros i m n h1 h2 h3. pose proof (dv_minus _ _ _ h3 h1) as h4. contradiction. Qed. Lemma not_dv_plus_rev_l : forall i m n, ~i dv m -> i dv (m + n) -> ~i dv n. intros i m n h1 h2. intro h3. pose proof (dv_plus_rev_l _ _ _ h3 h2) as h4; contradiction. Qed. Lemma not_dv_plus_rev_r : forall i m n, ~i dv n -> i dv (m + n) -> ~i dv m. intros i m n h1 h2. intro h3. pose proof (dv_plus_rev_r _ _ _ h3 h2) as h4; contradiction. Qed. Lemma not_dv_minus_rev_l : forall i m n, ~i dv m -> i dv (m - n) -> m <= n \/ ~i dv n. intros i m n h1 h2. destruct (le_lt_dec m n) as [hl | hl]. left; auto. right. intro h3. pose proof (dv_minus_rev_r _ _ _ h3 h2) as h4. destruct h4 as [h4 | h4]. pose proof (lt_le_trans _ _ _ hl h4) as h5. apply lt_irrefl in h5; auto. contradiction. Qed. Lemma not_dv_minus_rev_r : forall i m n, ~i dv n -> i dv (m - n) -> i = 0 \/ m <= n \/ ~i dv m. intros i m n h1 h2. destruct (zerop i) as [hz | hz]. left; auto. right. destruct (le_lt_dec m n) as [hl | hl]. left; auto. right. intro h3. pose proof (dv_minus_rev_l _ _ _ h3 h2) as h4. destruct h4 as [h4 | h4]. subst. apply lt_irrefl in hz; contradiction. destruct h4 as [h4 | h4]. pose proof (lt_le_trans _ _ _ hl h4) as h5. apply lt_irrefl in h5; auto. contradiction. Qed. Lemma minus_div_le : forall i j n, (i - j) / n <= i / n - j / n. intros i j n. destruct (le_lt_dec i j) as [hlt | hlt]. rewrite le_minus_iff in hlt. rewrite hlt, O_div; auto with arith. unfold div. destruct n as [|n]; auto. pose proof (divmod_spec (i - j) n 0 n (le_refl _)) as h1. pose proof (divmod_spec i n 0 n (le_refl _)) as h2. pose proof (divmod_spec j n 0 n (le_refl _)) as h3. destruct (divmod (i - j) n 0 n) as [q u], (divmod i n 0 n) as [q' u'], (divmod j n 0 n) as [q0 u0]. rewrite <- minus_n_n, mult_0_r, plus_0_r, plus_0_r in h1, h2, h3. simpl. destruct h1 as [h1 h4], h2 as [h2 h5], h3 as [h3 h6]; subst. rewrite minus_plus_assoc in h1. pose proof (le_minus n u0) as h7. pose proof (le_minus n u') as h8. pose proof (le_minus n u) as h9. red in hlt. destruct (le_lt_dec u' u0) as [hl | hl]. pose proof (minus_le_compat_r _ _ (n - u0) hlt) as h10. rewrite <- plus_S_shift_r in h10. rewrite <- plus_minus_assoc in h10. rewrite S_minus in h10. rewrite S_compat in h10. apply le_lt in h10. apply lt_S_n in h10. rewrite <- plus_minus_assoc, minus_same_diff in h10; auto. pose proof (le_minus u0 u') as h11. pose proof (le_trans _ _ _ h11 h6) as h12. pose proof (plus_le_compat_r _ _ (S n * q') h12) as h13. apply le_lt in h13. rewrite <- plus_S_shift_l in h13. rewrite plus_comm in h13. pose proof (lt_trans _ _ _ h10 h13) as h14. rewrite <- (mult_1_r (S n)) in h14 at 2. rewrite <- mult_plus_distr_l in h14. rewrite plus_comm, S_compat in h14. apply mult_S_lt_l in h14. apply lt_le in h14. rewrite plus_comm in h1. rewrite <- plus_minus_assoc in h1. rewrite minus_dist_l in h1. rewrite plus_comm in h1. rewrite <- plus_minus_assoc, minus_same_diff in h1. rewrite times_plus_eq_times_plus_iff in h1. destruct h1 as [h1 | h1]. destruct h1 as [h1 h2]. destruct h1; subst; auto with arith. discriminate. destruct h1 as [h1 | h1]. destruct h1 as [h1 [h2 h3]]. rewrite h3 in h12. pose proof (le_plus_l (S n * (q - (q' - q0))) (n- u)) as h15. pose proof (le_trans _ _ _ h15 h12) as h16. rewrite lt_iff in h2. do 2 red in h2. pose proof (mult_le_compat_r _ _ (S n) h2) as h17. rewrite mult_1_l in h17. rewrite mult_comm in h17. pose proof (le_trans _ _ _ h17 h16) as h18. apply le_Sn_n in h18. contradiction. destruct h1; auto with arith. assumption. assumption. apply ge_minus_anti; auto. apply mult_le_compat_l; auto. apply ge_minus_anti; auto. auto with arith. change (S n * q0 + (n - u0) < S n * q' + (n - u')) in hlt. pose proof hl as h10. apply (minus_lt_compat_r _ _ (n - u0)) in hlt. rewrite <- plus_minus_assoc in hlt; auto. rewrite <- minus_n_n, plus_0_r in hlt. pose proof (le_minus (S n * q' + (n- u')) (n - u0)) as h11. pose proof (lt_le_trans _ _ _ hlt h11) as h12. pose proof (plus_le_compat_r _ _ (S n * q') h8) as h13. rewrite plus_comm in h13. pose proof (lt_le_trans _ _ _ h12 h13) as h14. pose proof (lt_n_Sn n) as h15. pose proof (plus_lt_compat_r _ _ (S n * q') h15) as h16. pose proof (lt_trans _ _ _ h14 h16) as h17. rewrite <- (mult_1_r (S n)) in h17 at 2. rewrite <- mult_plus_distr_l in h17. apply mult_S_lt_l in h17. rewrite plus_comm, S_compat in h17. apply lt_le in h17. rewrite plus_comm in h1. rewrite <- plus_minus_assoc in h1. rewrite minus_dist_l in h1. apply le_lt_eq_dec in h17. destruct h17 as [h17 | h17]; subst. pose proof h17 as h17'. rewrite lt_iff in h17'. rewrite plus_comm in h1. pose proof (f_equal (fun c => c + (n - u0)) h1) as h18. simpl in h18. change (S n * (q' - q0) + (n - u') - (n - u0) + (n - u0) = S n * q + (n - u) + (n - u0)) in h18. rewrite <- minus_minus_assoc in h18. rewrite <- minus_n_n, <- minus_n_O in h18. rewrite <- plus_assoc in h18. rewrite times_plus_eq_times_plus_iff in h18. destruct h18 as [h18 | h18]. destruct h18 as [h18 h19]. destruct h18; subst. discriminate. auto. destruct h18 as [h18 | h18]. destruct h18 as [h18 h19]. destruct h19 as [h19 h20]. pose proof (le_minus n u') as h21. pose proof (lt_n_Sn n) as h22. pose proof (le_lt_trans _ _ _ h21 h22) as h23. pose proof h19 as h24. rewrite lt_iff in h24. pose proof (mult_le_compat_l _ _ (S n) h24) as h25. rewrite mult_1_r in h25. pose proof (lt_le_trans _ _ _ h23 h25) as h26. pose proof (le_plus_l (S n * (q - (q' - q0))) (n - u + (n - u0))) as h27. pose proof (lt_le_trans _ _ _ h26 h27) as h28. rewrite h20 in h28. apply lt_irrefl in h28. contradiction. destruct h18; auto with arith. rewrite <- minus_n_n; auto with arith. auto. pose proof (mult_le_compat_l _ _ (S n) h17') as h19. rewrite mult_1_r in h19. pose proof (le_plus_l (S n * (q' - q0)) (n - u')) as h20. pose proof (le_trans _ _ _ h19 h20) as h21. pose proof (le_minus n u0) as h22. pose proof (lt_n_Sn n) as h23. pose proof (lt_le_trans _ _ _ h23 h21) as h24. pose proof (le_lt_trans _ _ _ h22 h24); auto with arith. rewrite <- minus_n_n, mult_0_r, plus_0_r in h1. apply (gt_minus_anti _ _ n) in h10. apply lt_le_weak in h10. rewrite le_minus_iff in h10. rewrite h10 in h1. simpl in h1. symmetry in h1. apply plus_is_O in h1. destruct h1 as [h1 h17]. apply plus_is_O in h1. destruct h1; subst. rewrite <- minus_n_n; auto. assumption. apply mult_le_compat_l; auto. destruct q' as [|q']; subst. rewrite mult_0_r, plus_0_l in h1. destruct q0 as [|q0]. rewrite mult_0_r, plus_0_l, plus_0_l in hlt. pose proof (gt_minus_anti _ _ n h10 h5) as h11. pose proof (lt_trans _ _ _ h11 hlt) as h12. apply lt_irrefl in h12; contradiction. rewrite mult_0_r, plus_0_l in hlt. pose proof (lt_n_Sn n) as h11. pose proof (le_lt_trans _ _ _ h8 h11) as h12. pose proof (mult_O_le (S n) (S q0)) as h13. rewrite mult_comm in h13. destruct h13 as [|h13]. discriminate. pose proof (lt_le_trans _ _ _ h12 h13) as h14. pose proof (le_plus_l (S n * S q0) (n - u0)) as h15. pose proof (lt_le_trans _ _ _ h14 h15) as h16. rewrite hlt in h16. apply lt_irrefl in h16. contradiction. pose proof (lt_n_Sn n) as h11. pose proof (le_lt_trans _ _ _ h7 h11) as h12. pose proof (mult_O_le (S n) (S q')) as h13. destruct h13 as [h13 | h13]. discriminate. pose proof (lt_le_trans _ _ _ h12 h13) as h14. pose proof (le_plus_l (S q' * S n) (n - u')) as h15. pose proof (lt_le_trans _ _ _ h14 h15) as h16. rewrite mult_comm in h16; auto. auto with arith. Qed. Definition seg_pred_dec_from_S {n:nat} {pf:seg_pred_type (S n)} (f:seg_pred_dec pf) : seg_pred_dec (fun i (pflt:i < n) => (pf _ (lt_trans _ _ _ pflt (lt_n_Sn _)))) := (fun i (pflt:i < n) => f i (lt_trans _ _ _ pflt (lt_n_Sn _))). Definition seg_fun_eq_dec_from_S {U n} {f:seg_fun (S n) U} (pf:seg_fun_eq_dec f) : seg_fun_eq_dec (seg_fun_from_S f) := (fun i j (pfi:i < n) (pfj:j < n) => pf i j (lt_S _ _ pfi) (lt_S _ _ pfj)). Lemma seg_pred_dec_from_S_compat : forall U i n (f:seg_fun (S n) U) (pfi:i < n) (pfi':i < S n) (pfs:seg_pred_dec (fun j (pfj:j < S n) => f i pfi' = f j pfj)), seg_pred_dec (fun (j : nat) (pfj : j < n) => seg_fun_from_S f i pfi = seg_fun_from_S f j pfj). intros U i n f hi hi' h3. red in h3; red. intros j hj. unfold seg_fun_from_S. specialize (h3 j (lt_S _ _ hj)); auto. destruct h3 as [h3 | h3]. left. erewrite f_equal_dep. rewrite h3 at 1; auto. reflexivity. right. contradict h3. erewrite f_equal_dep. rewrite h3 at 1; auto. reflexivity. Qed. Lemma seg_fun_eq_dec_from_S_compat : forall {U n} {f:seg_fun (S n) U}, seg_fun_eq_dec f -> seg_fun_eq_dec (seg_fun_from_S f). intros n f pf h1. red in h1; red. intros i j hi hj. unfold seg_pred_type_from_S. specialize (h1 i j (lt_S _ _ hi) (lt_S _ _ hj)); auto. Qed. Lemma seg_pred_dec_from_S_compat' : forall U i n (f:seg_fun (S n) U) (pfi:i < n), forall (pfi':i < S n) (pfs:seg_pred_dec (fun j (pfj:j < S n) => f i pfi' = f j pfj)), seg_pred_dec (fun (j : nat) (pflt : j < n) => f i pfi' = f j (lt_S _ _ pflt)). intros U i n f hi hi' h3. red in h3; red. intros j hj. unfold seg_fun_from_S. specialize (h3 j (lt_S _ _ hj)); auto . Qed. (*This section contains boolean analogues of many familiar arithmetical predicates.*) Section NatAnalogues. Fixpoint beq_nat (a b:nat) : bool := match a with | 0 => match b with | 0 => true | S b' => false end | S a' => match b with | 0 => false | S b' => beq_nat a' b' end end. Fixpoint lt_nat (a b:nat) : bool := match a with | 0 => match b with | 0 => false | S b' => true end | S a' => match b with | 0 => false | S b' => lt_nat a' b' end end. Fixpoint gt_nat (a b:nat) : bool := match a with | 0 => false | S a' => match b with | 0 => true | S b' => gt_nat a' b' end end. Fixpoint le_nat (a b:nat) : bool := match a with | 0 => true | S a' => match b with | 0 => false | S b' => le_nat a' b' end end. Fixpoint ge_nat (a b:nat) : bool := match a with | 0 => match b with | 0 => true | S b' => false end | S a' => match b with | 0 => true | S b' => ge_nat a' b' end end. Lemma beq_nat_sym : forall a b, beq_nat a b = beq_nat b a. intro a; induction a; intro b; destruct b; simpl; auto. Qed. Lemma lt_nat_gt : forall a b, lt_nat a b = gt_nat b a. intro a; induction a; simpl; auto; intro b; destruct b; simpl; auto. Qed. Lemma le_nat_ge : forall a b, le_nat a b = ge_nat b a. intro a; induction a; simpl; auto; intro b; destruct b; simpl; auto. Qed. Lemma not_lt_nat : forall a b, negb (lt_nat a b) = orb (beq_nat a b) (gt_nat a b). intro a; induction a; simpl; auto; intro b; destruct b; simpl; auto. Qed. Lemma not_lt_nat' : forall a b, negb (lt_nat a b) = ge_nat a b. intro a; induction a; simpl; auto; intro b; destruct b; simpl; auto. Qed. Lemma not_gt_nat : forall a b, negb (gt_nat a b) = orb (beq_nat a b) (lt_nat a b). intro a; induction a; simpl; auto; intro b; destruct b; simpl; auto. Qed. Lemma not_gt_nat' : forall a b, negb (gt_nat a b) = le_nat a b. intro a; induction a; simpl; auto; intro b; destruct b; simpl; auto. Qed. Lemma lt_eq_lt_nat : forall a b, (orb (orb (lt_nat a b) (beq_nat a b)) (lt_nat b a)) = true. intro a; induction a; simpl; auto; intro b; destruct b; simpl; auto. Qed. Lemma beq_nat_compat : forall a b, beq_nat a b = true <-> a = b. intro a; induction a; intro b; destruct b; simpl; auto; split; intro h1; try discriminate h1; auto. f_equal; auto. rewrite IHa in h1; auto. apply S_inj in h1; subst. rewrite IHa; auto. Qed. Lemma lt_nat_compat : forall a b, lt_nat a b = true <-> a < b. intro a; induction a as [|a]; intro b; destruct b as [|b]; simpl; split; intro h1; try discriminate; try contradict (lt_n_0 _ h1); auto with arith. apply lt_n_S; auto. apply IHa; auto. apply lt_S_n in h1. rewrite <- IHa in h1; auto. Qed. Lemma gt_nat_compat : forall a b, gt_nat a b = true <-> a > b. intros ? ?; split; intro h1. rewrite <- lt_nat_gt in h1; apply lt_gt; apply lt_nat_compat; auto. rewrite <- lt_nat_gt; apply lt_gt in h1; apply lt_nat_compat in h1; auto. Qed. Lemma le_nat_compat : forall a b, le_nat a b = true <-> a <= b. intros a; induction a as [|a ih]; simpl. intuition. intro b; destruct b. intuition. split; intro h1. rewrite ih in h1; auto with arith. apply le_S_n in h1. rewrite <- ih in h1; auto. Qed. Lemma ge_nat_compat : forall a b, ge_nat a b = true <-> a >= b. intros ? ?; split; intro h1. rewrite <- le_nat_ge in h1; apply le_ge; apply le_nat_compat; auto. rewrite <- le_nat_ge; apply le_ge in h1; apply le_nat_compat in h1; auto. Qed. Lemma beq_nat_same : forall r, beq_nat r r = true. intro r; induction r; auto. Qed. Lemma lt_nat_same : forall r, lt_nat r r = false. intro r; induction r; auto. Qed. Lemma gt_nat_same : forall r, gt_nat r r = false. intro r; induction r; auto. Qed. Lemma le_nat_same : forall r, le_nat r r = true. intro r; induction r; auto. Qed. Lemma ge_nat_same : forall r, ge_nat r r = true. intro r; induction r; auto. Qed. Lemma neq_nat_compat : forall a b, beq_nat a b = false <-> a <> b. intro a; induction a as [|a ih]; simpl; intro b; destruct b; split; intro h1; try discriminate; auto. contradict h1; auto. intro h2. apply S_inj in h2. rewrite ih in h1; contradiction. rewrite ih. auto with arith. Qed. End NatAnalogues. Lemma lt_decompose : forall (m n:nat), m < n -> exists p, n = m + S p. intros m n. revert m. induction n as [|a ih]. intros ? h1. apply lt_n_0 in h1; contradiction. intros m h1. red in h1. apply le_S_n in h1. apply le_lt_eq_dec in h1. destruct h1 as [h1 |]; subst. apply ih in h1. destruct h1 as [p]; subst. exists (S p); omega. exists 0; omega. Qed. Lemma minus_plus' : forall m n:nat, m < n -> m + (n - m) = n. intros; omega. Qed. Lemma mono_pred_le : forall (m n:nat), m <= n -> pred m <= pred n. intros; omega. Qed. Lemma mono_pred_ge : forall (m n:nat), m >= n -> pred m >= pred n. intros; omega. Qed. Lemma n_minus_pred_n : forall (n:nat), 0 < pred n -> n - pred n = 1. intro n. induction n as [|n h1]. simpl. intro h1. contradict h1. auto with arith. simpl. intro h2. destruct n. auto. rewrite <- S_compat. auto with arith. Qed. (*For use in [map_seq] in [ListUtilities].*) Definition seg_S_tl {n:nat} {U:Type} (f:seg_fun (S n) U) : seq_fun 1 n U := fun (m:nat) (pf:1 <= m < 1 + n) => f m (ge1_lt m n pf). Lemma lt_seq : forall (m n w:nat), m < w -> n <= n + m < n + w. auto with arith. Qed. Lemma seq_le : forall m n w, n <= n + m < n + w -> n <= n + m. intros ? ? ? h1; destruct h1; auto. Qed. Lemma seq_lt : forall m n w, n <= n + m < n + w -> n + m < n + w. intros ? ? ? h1; destruct h1; auto. Qed. Lemma seq_lt' : forall m n w, n <= n + m < n + w -> n < n + w. intros; omega. Qed. Lemma lt_seq_iff : forall (m n w:nat), m < w <-> n <= n + m < n + w. intros; split; omega. Qed. Lemma lt_seq' : forall (m n w:nat), n <= m < n + w -> (n <= m /\ m - n < w). intros; omega. Qed. Lemma lt_seq_iff' : forall (m n w:nat), n <= m < n + w <-> (n <= m /\ m - n < w). intros; split; omega. Qed. Lemma seq_S : forall (m n w:nat), n <= m < n + w -> n <= m < n + S w. intros; omega. Qed. Lemma seq_S_w : forall (n w:nat), n <= n + w < n + S w. auto with arith. Qed. Lemma le_lt_S : forall i n w, S n <= i < S n + pred w -> n <= i < n + w. intros; omega. Qed. Lemma one_le_lt : forall i n, 1 <= i < 1 + pred n -> i < n. intros; omega. Qed. Lemma zero_le_lt : forall i m, 0 <= i < 0 + m -> i < m. intros; omega. Qed. Lemma seg_fun_to_seq_compat : forall {m:nat} {U:Type} (f:seg_fun m U) (i:nat) (pf:0 <= i < 0 + m), seg_fun_to_seq f i pf = f i (zero_le_lt _ _ pf). intros m u f i h1; unfold seg_fun_to_seq; destruct h1; f_equal; apply proof_irrelevance. Qed. Definition seq_fun_from_S {U:Type} {n w:nat} (f:seq_fun n (S w) U) : seq_fun n w U := fun m pf => f m (seq_S _ _ _ pf). Definition seq_fun_to_seg {U:Type} {n w:nat} (f:seq_fun n w U) : seg_fun w U := fun m pf => f _ (lt_seq _ _ _ pf). Definition seg_fun_rest_seq {m n:nat} {U:Type} (f:seg_fun n U) (pfm:m < n) : seq_fun m (n - m) U := let pfe := minus_plus' m n pfm in (fun i pf => let pf' := proj2 pf in let pflt := lt_eq_trans _ _ _ pf' pfe in f i pflt). Definition seq_fun_tl {n w:nat} {U:Type} (f:seq_fun n w U) : seq_fun (S n) (pred w) U := fun i pf => f i (le_lt_S _ _ _ pf). Definition seg_fun_tl {n:nat} {U:Type} (f:seg_fun n U) : seq_fun 1 (pred n) U := seq_fun_tl (seg_fun_to_seq f). Lemma seq_fun_from_S_compat : forall {U:Type} {n w:nat} (f:seq_fun n (S w) U) i (pf:n <= i < n + w), seq_fun_from_S f i pf = f i (seq_S _ _ _ pf). unfold seq_fun_from_S; intros; reflexivity. Qed. Lemma seq_fun_tl_compat : forall {n w:nat} {U:Type} (f:seq_fun n w U) (i:nat) (pf:S n <= i < S n + pred w), seq_fun_tl f i pf = f i (le_lt_S _ _ _ pf). unfold seq_fun_tl; intros; f_equal. Qed. Lemma seg_fun_tl_compat : forall {n w:nat} {U:Type} (f:seg_fun n U) (i:nat) (pf:1 <= i < 1 + pred n), seg_fun_tl f i pf = f i (one_le_lt _ _ pf). unfold seg_fun_tl; intros; rewrite seq_fun_tl_compat, seg_fun_to_seq_compat; f_equal; apply proof_irrelevance. Qed. Lemma seq_fun_eq_seg : forall {U:Type} {n w:nat} (f:seq_fun n w U), f = fun m pf => let pf' := proj2 (lt_seq' _ _ _ pf) in (seq_fun_to_seg f) _ pf'. intros U n w. revert n. induction w as [|w]; simpl; auto. intros n f. apply functional_extensionality_dep. intro m. apply functional_extensionality. intros; omega. intros n f. apply functional_extensionality_dep. intro m. apply functional_extensionality. intros h1. unfold seq_fun_to_seg. apply f_equal_prop_dep. omega. Qed. (*[Not HalfOpen]*) Lemma no_ho : forall (m n:nat), ~ m <= n < m + 0. intros; omega. Qed. Lemma O_lt_pred_n : forall (n:nat), 0 < pred n -> n > 1. intros; induction n; omega. Qed. Lemma lt_pred_n : forall m n:nat, m < n -> pred n < n. intros; omega. Qed. Lemma O_lt_pred_n_lt_n : forall (n:nat), 0 < pred n -> pred n < n. intros n h1. induction n; omega. Qed. Lemma O_pred_n_eq_1 : forall n:nat, 0 < n -> 0 = pred n -> n = 1. intro n. induction n; simpl; auto; omega. Qed. Definition abs_diff (a b:nat) := max a b - min a b. Lemma lt_sum_iff : forall n a b, n < a + b <-> n - min a b < max a b. intros n a b. destruct (le_lt_dec a b) as [h2 | h2]. rewrite min_l, max_r; omega. rewrite min_r, max_l; omega. Qed. Lemma min_comm : forall (m n:nat), min m n = min n m. intros m n. destruct (le_lt_dec m n) as [h1 | h1]. rewrite (min_l m), (min_r n); auto. rewrite (min_r m), (min_l n); auto with arith. Qed. Lemma max_comm : forall (m n:nat), max m n = max n m. intros m n. destruct (le_lt_dec m n) as [h1 | h1]. rewrite (max_r m), (max_l n); auto. rewrite (max_l m), (max_r n); auto with arith. Qed. Lemma min_compat_l : forall (m n:nat), min m n <= m. intros; destruct (le_lt_dec m n). rewrite min_l; auto. rewrite min_r; auto with arith. Qed. Lemma min_compat_r : forall (m n:nat), min m n <= n. intros; destruct (le_lt_dec m n). rewrite min_l; auto. rewrite min_r; auto with arith. Qed. Lemma min_O : forall (m:nat), min m O = O. intro; apply min_r; apply le_0_n. Qed. Lemma max_O : forall (m:nat), max m O = m. intro; apply max_l; apply le_0_n. Qed. Lemma min_S : forall (m n:nat), min (S m) (S n) = S (min m n). intros; simpl; auto. Qed. Lemma le_min_iff : forall (m n a:nat), a <= min m n <-> (a <= m /\ a <= n). intros m n a. destruct (le_lt_dec m n). rewrite min_l; omega. rewrite min_r; omega. Qed. Lemma min_le_iff : forall (m n a:nat), min m n <= a <-> (m <= a \/ n <= a). intros m n a. destruct (le_lt_dec m n) as [h1 | h1]. rewrite min_l; omega. rewrite min_r; omega. Qed. Lemma le_max_iff : forall (m n a:nat), a <= max m n <-> (a <= m \/ a <= n). intros m n a. destruct (le_lt_dec m n) as [h1 | h1]. rewrite max_r; omega. rewrite max_l; omega. Qed. Lemma max_le_iff : forall (m n a:nat), max m n <= a <-> (m <= a /\ n <= a). intros m n a. destruct (le_lt_dec m n) as [h1 | h1]. rewrite max_r; omega. rewrite max_l; omega. Qed. Lemma min_max_compat : forall (a b:nat), min (max a b) a = a. intros. destruct (le_lt_dec a b) as [h2 | h3]. rewrite max_r, min_r; auto. rewrite max_l, min_l; auto with arith. Qed. Lemma O_lt_pow : forall (b n:nat), 0 < b -> 0 < b^n. intros b n. revert b. induction n as [|n h1]. simpl. auto with arith. simpl. simpl in h1. simpl. auto. intro b. simpl. specialize (h1 b). induction b as [|b h2]. omega. intro h3. specialize (h1 h3). pose proof mult_lt_compat_r. rewrite mult_comm. assert (h4:0 = 0 * S b). ring. rewrite h4. apply mult_lt_compat_r. assumption. assumption. Qed. Lemma pow_mono : forall (b m n:nat), 0 < b -> m <= n -> b^m <= b^n. intros b m n. revert b m. induction n as [|n h1]. simpl. intros b m h1 h2. assert (h3:m = 0). auto with arith. subst. simpl. auto. intros b m. revert b. rename h1 into hind. induction m as [|m h1]. intros b h1 h2. pose proof (O_lt_pow _ (S n) h1) as h3. simpl. simpl in h3. omega. intros b h2 h3. apply le_S_n in h3. specialize (hind _ _ h2 h3). simpl. pose proof (mult_le_compat_r (b ^ m) (b ^n) b hind) as h4. rewrite (mult_comm b (b ^ n)). rewrite mult_comm. assumption. Qed. Lemma k_nat_ind : forall (k:nat) (P:{n:nat | k <= n}->Prop), P (exist _ _ (le_refl _)) -> (forall n, P n -> P (exist _ _ (le_S _ _ (proj2_sig n)))) -> forall n, P n. intro k. induction k as [|k h1]. intros P h1 h2 n. destruct n as [n h3]. revert h3 h2. revert h1. revert P. induction n as [|n h4]. intros P h4 h5 h6. pose proof (proof_irrelevance _ h5 (le_refl 0)); subst; auto. intros P h1 h2 h3. specialize (h4 P h1 (le_0_n _) h3). specialize (h3 _ h4). simpl in h3. assert (h5:h2 = le_S 0 n (le_0_n n)). apply proof_irrelevance. subst. assumption. intros P h2 h3 n. destruct n as [n h4]. pose (fun x:{n:nat | k <= n} => P (exist _ _ (le_n_S _ _ (proj2_sig x)))) as P'. assert (h5: P' (exist (fun n : nat => k <= n) k (le_refl k))). unfold P'. simpl. assert (h6:exist _ _ (le_refl (S k)) = exist _ _ (le_n_S k k (le_refl k))). apply proj1_sig_injective; auto. rewrite <- h6 at 1. assumption. assert (h6: forall n : {n : nat | k <= n}, P' n -> P' (exist (fun n0 : nat => k <= n0) (S (proj1_sig n)) (le_S k (proj1_sig n) (proj2_sig n)))). intros x h7. specialize (h3 _ h7). simpl in h3. unfold P'. simpl. assert (h8:exist _ _ (le_S (S k) (S (proj1_sig x)) (le_n_S k (proj1_sig x) (proj2_sig x))) = exist _ _ (le_n_S k (S (proj1_sig x)) (le_S k (proj1_sig x) (proj2_sig x)))). apply proj1_sig_injective; auto. rewrite <- h8 at 1. assumption. pose proof (le_pred _ _ h4) as h4'. simpl in h4'. specialize (h1 P' h5 h6 (exist _ _ h4')). unfold P' in h1. simpl in h1. assert (h7:0 < n). omega. pose proof (S_pred _ _ h7) as h8. assert (h9:exist _ _ h4 = exist _ _ (le_n_S k (pred n) h4')). apply proj1_sig_injective. simpl. assumption. rewrite h9 at 1. assumption. Qed. Lemma divmod_x0S0 : forall x q, divmod x 0 (S q) 0 = (S (fst (divmod x 0 q 0)), 0). intro x; induction x; simpl; auto. Qed. Lemma divmod_x000 : forall x, divmod x 0 0 0 = (x, 0). intro x. induction x as [|x ih]; simpl; auto. rewrite divmod_x0S0. rewrite ih; simpl; auto. Qed. Lemma divmod_0y0u : forall y u, divmod 0 y 0 u = (0, u). intros y u. revert y. induction u as [|u]; simpl; auto. Qed. Lemma divmod_S00S : forall x u, divmod (S x) 0 0 (S u) = divmod x 0 0 u. auto. Qed. Lemma divmod_SyS0 : forall x y q, divmod (S x) y (S q) 0 = divmod x y (S (S q)) y. auto. Qed. Lemma divmod_Sy00 : forall x y, divmod (S x) y 0 0 = divmod x y 1 y. auto. Qed. Lemma divmod_Sy0S : forall x y u, divmod (S x) y 0 (S u) = divmod x y 0 u. auto. Qed. Lemma divmod_Sx0x : forall x, divmod (S x) x 0 x = (1, x). intro x. pose proof (divmod_spec (S x) x 0 x (le_refl _)) as h1. rewrite <- mult_n_O in h1. rewrite <- plus_n_O in h1. rewrite minus_same in h1. rewrite <- plus_n_O in h1. destruct (divmod (S x) x 0 x) as [f s]. destruct h1 as [h1 h2]. rewrite eq_times_plus_iff in h1. destruct h1 as [h1 | [h1 | h1]]. destruct h1; discriminate. destruct h1; subst. rewrite <- le_minus_iff in H0. apply injective_projections; auto with arith. destruct h1 as [h1 h3]; subst. omega. Qed. Lemma mod_functional : forall i j n, i = j -> i mod n = j mod n. intros i j n ?; subst; auto. Qed. Lemma mod_same : forall x, x mod x = 0. intro x. unfold modulo. destruct x as [|x]; auto. pose proof (divmod_spec (S x) x 0 x (le_refl _)) as h1. destruct (divmod (S x) x 0 x). rewrite minus_same in h1. rewrite mult_0_r in h1. do 2 rewrite plus_0_r in h1. destruct h1 as [h1 h2]. rewrite eq_times_plus_iff in h1. destruct h1 as [h1 | [h1 | h1]]. destruct h1; discriminate. destruct h1 as [? h1]; subst. simpl; auto. destruct h1 as [? h1]; subst; simpl. omega. Qed. Lemma one_mod : 1 mod 1 = 0. simpl; auto. Qed. Lemma one_mod_S : forall n, n > 1 -> 1 mod n = 1. intros n h1. destruct n as [|n]; simpl. apply lt_n_0 in h1; contradiction. destruct n as [|n]. apply lt_irrefl in h1; contradiction. simpl. destruct n as [|n]; auto. rewrite S_minus; auto. Qed. Lemma S_mod : forall x y, S x mod y = (x mod y + 1 mod y) mod y. intros x y. unfold modulo. destruct y as [|y]; auto. pose proof (divmod_spec (S x) y 0 y (le_refl _)) as h1. destruct (divmod (S x) y 0 y) as [q u]. rewrite minus_same, mult_0_r in h1. do 2 rewrite plus_0_r in h1. destruct h1 as [h1 h2]. pose proof (divmod_spec x y 0 y (le_refl _)) as h3. destruct (divmod x y 0 y) as [q' u']. rewrite minus_same, mult_0_r in h3. do 2 rewrite plus_0_r in h3. destruct h3 as [h3 h4]. pose proof (divmod_spec 1 y 0 y (le_refl _)) as h5. destruct (divmod 1 y 0 y) as [r v]. rewrite minus_same, mult_0_r in h5. do 2 rewrite plus_0_r in h5. destruct h5 as [h5 h6]. pose proof (divmod_spec (y - snd (q', u') + (y - snd (r, v))) y 0 y (le_refl _)) as h7. destruct (divmod (y - snd (q', u') + (y - snd (r, v))) y 0 y) as [r' v']. simpl. simpl in h7. rewrite minus_same, mult_0_r in h7. do 2 rewrite plus_0_r in h7. destruct h7 as [h7 h8]. f_equal. subst. rewrite one_eq_S_times_plus_iff in h5. destruct h5 as [h5 | h5]. destruct h5 as [? h5]; subst. apply le0 in h2; subst. apply le0 in h8; subst; auto. destruct h5 as [? h5]; subst. pose proof (f_equal (plus v) h5) as h9. rewrite <- le_plus_minus in h9; auto. subst. rewrite h5 in h7. rewrite <- (mult_1_l r') in h7 at 1. rewrite <- mult_plus_distr_r in h7. rewrite plus_comm in h7. rewrite (plus_minus_assoc 1 (v+1) u') in h7; auto. rewrite minus_eq_times_plus_iff in h7. destruct h7 as [h7 | h7]. destruct h7 as [h7 h9]. rewrite S_compat, plus_S_shift_r in h7. discriminate. destruct h7 as [h7 | h7]. destruct h7 as [? h7]; subst. destruct h7 as [h7 | h7]. destruct h7 as [h7]. rewrite S_compat, plus_S_shift_r in h7. discriminate. destruct h7 as [h7 | h7]. destruct h7 as [? h7]; subst. rewrite <- le_minus_iff in h7. pose proof (le_antisym _ _ h7 h8); subst. rewrite S_compat in h1. rewrite <- minus_n_O in h1. destruct u as [|u]. rewrite <- minus_n_O in h1. rewrite (plus_S_shift_r (S (S v) * q) v) in h1. apply S_inj in h1. rewrite times_plus_eq_times_plus_iff in h1. destruct h1 as [h1 | [h1 | h1]]. destruct h1; omega. destruct h1 as [h1 h9]. destruct h9 as [h9 h10]. pose proof (lt_n_Sn (S v)) as h11. pose proof (mult_O_le (S (S v)) (q - q')) as h12. destruct h12 as [? | h12]. omega. rewrite mult_comm in h12. pose proof (le_plus_l (S (S v) * (q - q')) v) as h13. rewrite <- h10 in h13. pose proof (lt_le_trans _ _ _ h11 h12) as h14. pose proof (lt_le_trans _ _ _ h14 h13); omega. destruct h1 as [h1 h9]. pose proof (lt_n_Sn v) as h11. pose proof (lt_n_Sn (S v)) as h11'. pose proof (mult_O_le (S (S v)) (q' - q)) as h12. destruct h12 as [? | h12]. omega. rewrite mult_comm in h12. pose proof (le_plus_l (S (S v) * (q' - q)) (S v)) as h13. rewrite <- h9 in h13. pose proof (lt_trans _ _ _ h11 h11') as h11a. pose proof (lt_le_trans _ _ _ h11a h12) as h14. pose proof (lt_le_trans _ _ _ h14 h13); omega. rewrite S_minus_S in h1. rewrite <- plus_S_shift_r in h1. rewrite <- (mult_1_r (S (S v))) in h1 at 2. rewrite <- mult_plus_distr_l in h1. rewrite times_eq_times_plus_iff in h1. destruct h1 as [h1 | h1]. destruct h1; omega. destruct h1 as [h1 [h10 he]]. pose proof (lt_n_Sn v) as h11. pose proof (lt_n_Sn (S v)) as h11'. pose proof (lt_trans _ _ _ h11 h11') as h12. pose proof (mult_O_le (S (S v)) (q' + 1 - q)) as h13. destruct h13 as [? | h13]. omega. rewrite mult_comm in h13. pose proof (lt_le_trans _ _ _ h12 h13) as h14. rewrite <- he in h14. omega. destruct h7 as [? h7]; subst. rewrite S_compat in h7. rewrite plus_comm, S_compat in h7. omega. destruct h7 as [h7 | h7]. destruct h7 as [? h7]; subst. apply le_lt_eq_dec in h2. destruct h2 as [h2 | h2]. rewrite S_compat in h1. pose proof (S_minus_shift v u) as h9. hfirst h9. red in h2. rewrite S_compat in h2. apply le_S_n in h2; auto. specialize (h9 hf). rewrite h9 in h1. rewrite plus_S_shift_r in h1. apply S_inj in h1. rewrite times_plus_eq_times_plus_iff in h1. destruct h1 as [h1 | h1]. destruct h1 as [h1 h10]. destruct h1; subst. discriminate. rewrite S_compat in h7. rewrite <- plus_minus_assoc in h7; auto with arith. rewrite h10 in h7. rewrite plus_comm, S_compat in h7. rewrite <- S_minus_shift in h7. apply minus_inj in h7; auto. rewrite S_compat in h8; auto. assumption. rewrite S_compat in h4; auto. rewrite S_compat in h7. rewrite <- plus_minus_assoc in h7. destruct h1 as [h1 | h1]. destruct h1 as [h1 h10]. destruct h10 as [h11 h12]. rewrite h12 in h7. rewrite plus_comm in h7. rewrite <- plus_assoc in h7. rewrite S_compat in h7. pose proof (mult_O_le (S (S v)) (q - q')) as h13. destruct h13 as [h13 | h13]. rewrite lt_iff in h11. rewrite h13 in h11. apply lt_irrefl in h11; contradiction. rewrite mult_comm in h13. pose proof (plus_le_compat_r _ _ (S (v - u)) h13) as h14. rewrite <- h7 in h14. pose proof (le_plus_l (S (S v)) (S (v - u))) as h15. pose proof (le_trans _ _ _ h15 h14) as h16. pose proof (le_minus (S v) v') as h17. pose proof (le_trans _ _ _ h16 h17) as h18. apply le_Sn_n in h18. contradiction. destruct h1 as [h1 h10]. rewrite plus_comm in h7. rewrite <- minus_minus_assoc in h7. apply minus_inj in h7; subst. pose proof (le_minus v u) as h11. pose proof (mult_O_le (S (S v)) (q' - q)) as h12. destruct h12 as [h12 | h12]. rewrite lt_iff in h1. rewrite h12 in h1. apply lt_irrefl in h1. contradiction. rewrite mult_comm in h12. pose proof (le_plus_r (S v - u') (S (S v) * (q' - q))) as h13. rewrite plus_comm in h13. rewrite <- h10 in h13. pose proof (lt_n_Sn v) as h14. pose proof (lt_n_Sn (S v)) as h15. pose proof (lt_trans _ _ _ h14 h15) as h16. pose proof (lt_le_trans _ _ _ h16 h12) as h17. pose proof (lt_le_trans _ _ _ h17 h13) as h18. pose proof (le_lt_trans _ _ _ h11 h18) as h19. apply lt_irrefl in h19. contradiction. rewrite S_compat in h8; auto. rewrite S_compat in h4. eapply le_trans. apply le_minus. assumption. rewrite S_compat in h4. eapply le_trans. apply le_minus. assumption. destruct (zerop u') as [|h11]; subst. rewrite <- minus_n_O in h7. pose proof (le_minus (S v) v') as h11. rewrite S_compat in h7. pose proof (lt_n_Sn (S v)) as h12. pose proof (le_lt_trans _ _ _ h11 h12) as h13. rewrite h7 in h13. apply lt_irrefl in h13. contradiction. red in h11; auto. rewrite S_compat in h4; auto. rewrite S_compat in h4; auto. subst. rewrite <- minus_n_n, plus_0_r in h1. destruct (zerop q) as [h9 | h9]; subst. rewrite mult_0_r in h1. discriminate. pose proof (le_times_S (S (v+1)) (pred q)) as h10. rewrite <- (S_pred _ _ h9) in h10. rewrite <- h1 in h10. destruct (zerop q') as [h11 | h11]; subst. rewrite mult_0_r in h1. rewrite plus_0_l in h1. rewrite <- S_minus_shift in h1. pose proof (le_minus (S (v+1)) u') as h11. rewrite h1 in h11. pose proof (le_times_S (S (v + 1)) (pred q)) as h12. rewrite <- (S_pred _ _ h9) in h12. pose proof (le_antisym _ _ h11 h12) as h13. rewrite <- mult_1_r in h13. apply mult_inj in h13; subst. rewrite mult_1_r in h1. symmetry in h1. rewrite minus_eq_iff in h1. destruct h1; subst. discriminate. rewrite <- minus_n_O in h7. rewrite S_compat in h7. rewrite plus_comm in h7. rewrite S_compat in h7. pose proof (le_minus (S v) v') as h13. rewrite h7 in h13. apply le_Sn_n in h13. contradiction. auto with arith. assumption. rewrite <- plus_S_shift_r in h1. symmetry in h1. rewrite times_eq_times_plus_iff in h1. destruct h1 as [h1 | h1]. destruct h1; discriminate. destruct h1 as [h1 h2]. destruct h2 as [h2 h3]. symmetry in h3. rewrite <- S_minus_shift in h3. pose proof (le_minus (S (v + 1)) u') as h12. rewrite <- h3 in h12. pose proof (mult_O_le (S (v + 1)) (q - q')) as h13. destruct h13 as [h13 | h13]. apply gt_lt in h2. rewrite lt_iff in h2. rewrite h13 in h2. apply lt_irrefl in h2; contradiction. rewrite mult_comm in h13. pose proof (le_antisym _ _ h12 h13) as h14. rewrite <- mult_1_r in h14. apply mult_inj in h14. pose proof (f_equal (plus q') h14) as h15. rewrite <- le_plus_minus in h15. subst. rewrite minus_plus, mult_1_r in h3. rewrite minus_eq_iff in h3. destruct h3; subst. discriminate. rewrite <- minus_n_O in h7. rewrite S_compat in h7. pose proof (le_minus (S v) v') as h15. rewrite h7 in h15. rewrite plus_comm, S_compat in h15. apply le_Sn_n in h15. contradiction. auto with arith. auto with arith. assumption. destruct h7 as [h7 h10]. pose proof (f_equal (plus v') h7) as h11. rewrite plus_minus_assoc in h11; auto. rewrite minus_plus, plus_0_r in h11. subst. destruct h10 as [h10 | h10]. rewrite plus_comm, S_compat in h10. discriminate. destruct h10 as [h10 | h10]. destruct h10 as [h10 h11]; subst. destruct h11 as [h11 | h11]. rewrite plus_comm, S_compat in h11. discriminate. destruct h11 as [h11 | h11]; subst. destruct h11 as [h11]. rewrite plus_comm, S_compat in h11. discriminate. rewrite <- minus_n_O, S_compat in h1. rewrite plus_minus_assoc in h1. rewrite (plus_S_shift_r (S (S v) * q) v) in h1. rewrite S_minus_shift in h1. apply S_inj in h1. pose proof (f_equal (plus u) h1) as h9. rewrite plus_comm in h9. rewrite <- plus_assoc in h9. rewrite plus_minus_assoc in h9. rewrite minus_plus in h9. rewrite times_plus_eq_times_plus_iff in h9. destruct h9 as [h9 | h9]. destruct h9 as [? h9]. pose proof (le_plus_l (S v) u) as h10. rewrite h9 in h10. apply le_Sn_n in h10. contradiction. destruct h9 as [h9 | h9]. destruct h9 as [h9 h10]. destruct h10 as [h10 h11]. destruct u as [|u]. rewrite plus_0_r in h11. pose proof (mult_O_le (S (S v)) (q - q')) as h12. destruct h12 as [h12 | h12]. rewrite lt_iff in h10. rewrite h12 in h10. apply lt_n_0 in h10. contradiction. rewrite mult_comm in h12. pose proof (le_plus_l (S (S v) * (q - q')) v) as h13. pose proof (le_trans _ _ _ h12 h13) as h14. rewrite <- h11 in h14. apply le_Sn_n in h14. contradiction. rewrite <- plus_Snm_nSm in h11. rewrite <- (mult_1_r (S (S v))) in h11 at 1. rewrite times_plus_eq_times_plus_iff in h11. destruct h11 as [h11 | h11]. destruct h11 as [h11]. destruct h11 as [|h11]. discriminate. subst. rewrite S_compat; auto. destruct h11 as [h11 | h11]. destruct h11 as [h11 h12]. destruct h12 as [h12 h13]. rewrite h13 in h1. rewrite S_compat in h2. apply le_S_n in h2. rewrite h13 in h2. pose proof h12 as h12'. rewrite lt_iff in h12'. assert (hlt : S (S v) * (q - q' -1) > 0). rewrite (S_pred _ _ h11), (S_pred _ _ h12'). rewrite S_times_S. auto with arith. pose proof (lt_plus_S v (pred (S (S v) * (q - q' - 1)))) as h14. rewrite <- (S_pred _ _ hlt) in h14. rewrite plus_comm in h14. pose proof (lt_le_trans _ _ _ h14 h2) as h15. apply lt_irrefl in h15. contradiction. destruct h11 as [h11 h12]. red in h11. apply le_S_n in h11. apply le0 in h11. rewrite <- le_minus_iff in h11. pose proof (lt_le_trans _ _ _ h10 h11) as h13. apply lt_irrefl in h13; contradiction. destruct h9 as [h9 h10]. pose proof (lt_n_Sn v) as h11. pose proof (le_plus_l (S v) u) as h12. pose proof (lt_le_trans _ _ _ h11 h12) as h13. pose proof (le_plus_l (S v + u) (S (S v) * (q' - q))) as h14. pose proof (lt_le_trans _ _ _ h13 h14) as h15. rewrite plus_comm in h15. rewrite <- h10 in h15. apply lt_irrefl in h15; contradiction. rewrite S_compat in h2. destruct q as [|q]. rewrite mult_0_r in h9, h1. rewrite plus_0_l in h1, h9. pose proof (le_minus v u) as h10. pose proof (lt_n_Sn v) as h11. pose proof (le_lt_trans _ _ _ h10 h11) as h12. pose proof (le_plus_l (S v) (S (S v) * q')) as h13. pose proof (lt_le_trans _ _ _ h12 h13) as h14. rewrite plus_comm in h14. rewrite h1 in h14. apply lt_irrefl in h14; contradiction. pose proof (mult_O_le (S (S v)) (S q)) as h12. pose proof (lt_n_Sn (S v)) as h10. pose proof (le_lt_trans _ _ _ h2 h10) as h11. destruct h12 as [|h12]; subst. discriminate. pose proof (lt_le_trans _ _ _ h11 h12) as h13. rewrite mult_comm in h13. pose proof (le_plus_l (S (S v) * S q) v) as h14. pose proof (lt_le_trans _ _ _ h13 h14); auto with arith. rewrite S_compat in h2. destruct q as [|q]. rewrite mult_0_r in h1. rewrite plus_0_l in h1. pose proof (le_minus (S v) u) as h10. pose proof (le_plus_l (S v) (S (S v) * q')) as h11. pose proof (le_trans _ _ _ h10 h11) as h12. rewrite plus_comm in h12. pose proof (lt_n_Sn (S (S v) * q' + S v)) as h13. pose proof (le_lt_trans _ _ _ h12 h13) as h14. rewrite h1 in h14. apply lt_irrefl in h14. contradiction. pose proof (lt_n_Sn (S v)) as h9. pose proof (le_lt_trans _ _ _ h2 h9) as h10. pose proof (mult_O_le (S (S v)) (S q)) as h11. destruct h11 as [|h11]. discriminate. pose proof (le_plus_l (S q * S (S v)) v) as h12. pose proof (le_trans _ _ _ h11 h12) as h13. rewrite mult_comm in h13. pose proof (lt_le_trans _ _ _ h10 h13); auto with arith. rewrite S_compat in h2; auto. destruct h10 as [? h10]; subst. rewrite plus_comm in h10. rewrite S_compat in h10. pose proof (le_trans _ _ _ h10 h4) as h11. apply le_Sn_n in h11. contradiction. Qed. Lemma gt_one_mod : forall n, n > 1 -> 1 mod n = 1. intros; rewrite S_mod, O_mod, plus_0_l, lt_mod_eq, lt_mod_eq; auto. apply lt_mod; auto with arith. Qed. Lemma plus_times_mod_r : forall i j n, (i + n * j) mod n = i mod n. intros i j n. unfold modulo. destruct n as [|n]; auto. pose proof (divmod_spec (i + S n * j) n 0 n (le_refl _)) as h1. pose proof (divmod_spec i n 0 n (le_refl _)) as h2. destruct (divmod (i + S n * j) n 0 n) as [q u], (divmod i n 0 n) as [q' u']. rewrite mult_0_r, minus_same, plus_0_r, plus_0_r in h1, h2. simpl. destruct h1 as [h1 h3], h2 as [h2 h4]. subst. rewrite <- plus_assoc in h1. rewrite (plus_comm (n - u')) in h1. rewrite plus_assoc in h1. rewrite <- mult_plus_distr_l in h1. rewrite times_plus_eq_times_plus_iff in h1. destruct h1 as [h1 | h1]. destruct h1; auto. destruct h1 as [h1 | h1]. destruct h1 as [h1 [h2 h5]]. pose proof (le_minus n u') as h6. rewrite h5 in h6. pose proof (mult_O_le (S n) (q - (q' + j))) as h8. destruct h8 as [h8 | h8]. rewrite <- le_minus_iff in h8. pose proof (le_lt_trans _ _ _ h8 h2) as h9. apply lt_irrefl in h9; contradiction. rewrite mult_comm in h8. pose proof (le_plus_l (S n * (q - (q' + j))) (n - u)) as h9. pose proof (le_trans _ _ _ h9 h6) as h10. pose proof (le_trans _ _ _ h8 h10) as h11. apply le_Sn_n in h11. contradiction. destruct h1 as [h1 h2]. pose proof (mult_O_le (S n) (q' + j - q)) as h5. destruct h5 as [h5 | h5]. rewrite <- le_minus_iff in h5. pose proof (lt_le_trans _ _ _ h1 h5) as h6. apply lt_irrefl in h6. contradiction. rewrite mult_comm in h5. pose proof (le_plus_l (S n * (q' + j - q)) (n - u')) as h6. pose proof (le_trans _ _ _ h5 h6) as h7. rewrite <- h2 in h7. pose proof (le_minus n u) as h8. pose proof (le_trans _ _ _ h7 h8) as h9. apply le_Sn_n in h9. contradiction. Qed. Lemma plus_times_mod_ : forall i j n, (n * j + i) mod n = i mod n. intros; rewrite plus_comm; apply plus_times_mod_r. Qed. Lemma minus_times_mod_r : forall i j n, i >= n * j -> (i - n * j) mod n = i mod n. intros i j n h1. unfold modulo. destruct n as [|n]; auto. pose proof (divmod_spec (i - S n * j) n 0 n (le_refl _)) as h2. pose proof (divmod_spec i n 0 n (le_refl _)) as h3. destruct (divmod (i - S n * j) n 0 n) as [u v], (divmod i n 0 n) as [u' v']. simpl. rewrite mult_0_r, minus_same, plus_0_r, plus_0_r in h2, h3. destruct h2 as [h2 h4], h3 as [h3 h5]; subst. pose proof (f_equal (fun c => c + (S n * j)) h2) as h3. simpl in h3. do 3 rewrite <- S_times_l' in h3. rewrite minus_plus_same, <- plus_assoc, (plus_comm _ (S n * j)), plus_assoc, <- mult_plus_distr_l in h3; auto. rewrite times_plus_eq_times_plus_iff in h3. destruct h3 as [h3 | h3]. destruct h3 as [h3 h6]. destruct h3; subst. discriminate. rewrite mult_plus_distr_l, <- plus_assoc, plus_comm, plus_plus_plus_minus_same, (plus_comm _ (n - v)) in h2. apply plus_inj_r in h2; auto. destruct h3 as [h3 | h3]. destruct h3 as [h3 h6]. destruct h6 as [h6 h7]. rewrite lt_iff in h6. pose proof (le_minus n v') as h8. pose proof (le_lt_trans _ _ _ h8 (lt_n_Sn _)) as h9. pose proof (mult_O_le (S n) (u + j - u')) as h10. destruct h10 as [h10 | h10]. rewrite h10 in h6. apply lt_irrefl in h6; contradiction. rewrite mult_comm in h10. pose proof (le_plus_l (S n * (u + j - u')) (n - v)) as h11. pose proof (le_trans _ _ _ h10 h11) as h12. rewrite <- h7 in h12. pose proof (le_lt_trans _ _ _ h12 h9) as h13. apply lt_irrefl in h13; contradiction. destruct h3 as [h3 h6]. rewrite plus_comm in h2. pose proof (mult_O_le (S n) (u' - (u + j))) as h7. destruct h7 as [h7 | h7]. rewrite <- le_minus_iff in h7. pose proof (le_lt_trans _ _ _ h7 h3) as h8. apply lt_irrefl in h8; contradiction. pose proof (le_plus_l (S n * (u' - (u + j))) (n - v')) as h8. rewrite <- h6, mult_comm in h8. pose proof (le_trans _ _ _ h7 h8) as h9. pose proof (le_minus n v) as h10. pose proof (le_trans _ _ _ h9 h10) as h11. apply le_Sn_n in h11. contradiction. Qed. Lemma mod_eq_iff : forall i j n, n > 0 -> i < j -> (i mod n = j mod n <-> (n dv (j - i))). intros i j n h1 hlt. pose proof h1 as h1'. apply gt_lt in h1'. rewrite <- neq0_iff in h1'. split; intro h2. pose proof (div_mod i n h1') as h3. pose proof (div_mod j n h1') as h4. rewrite h2 in h3. rewrite h3. rewrite h4 at 1. rewrite (plus_comm (n*(i/n)) (j mod n)). rewrite minus_plus_assoc. rewrite <- plus_minus_assoc. rewrite <- minus_n_n, plus_0_r. rewrite mult_comm, (mult_comm n). rewrite minus_dist_r. rewrite mult_comm. constructor. auto with arith. rewrite plus_comm. apply plus_le_compat_r. apply mult_le_compat_l. apply le_div; auto with arith. inversion h2. pose proof (div_mod i n h1') as h3. pose proof (div_mod j n h1') as h4. rewrite h3, h4 in H0. destruct (le_lt_dec (i mod n) (j mod n)) as [h5 | h5]. rewrite plus_comm in H0. rewrite minus_plus_assoc in H0. rewrite <- plus_minus_assoc in H0. rewrite minus_dist_l in H0. rewrite plus_comm in H0. rewrite times_eq_times_plus_minus_iff in H0. destruct H0 as [h6 | h6]. destruct h6 as [h6]. rewrite h6 in h1. apply lt_irrefl in h1; contradiction. destruct h6 as [h6 | h6]. destruct h6; auto. destruct h6 as [h6 | h6]. destruct h6 as [h6 h7]. pose proof (lt_mod i n h1) as h9. pose proof (le_lt_trans _ _ _ h7 h9) as h10. pose proof (le_plus_l (n * (j/n - i/n)) (j mod n)) as h11. pose proof (le_lt_trans _ _ _ h11 h10) as h12. rewrite <- (mult_1_r n) in h12 at 4. rewrite (S_pred _ _ h1) in h12. apply mult_S_lt_l in h12. red in h12. apply le_S_n in h12. apply le0 in h12. rewrite <- (S_pred _ _ h1) in h12. pose proof (minus_div_le j i n) as h13. rewrite h12 in h13. apply le0 in h13. rewrite div_eq_0_iff in h13. destruct h13 as [h13 | h13]. rewrite h13 in h1. apply lt_irrefl in h1; contradiction. pose proof (dv_le _ _ h2) as h14. destruct h14 as [h14 | h14]. rewrite <- le_minus_iff in h14. pose proof (lt_le_trans _ _ _ hlt h14) as h15. apply lt_irrefl in h15; contradiction. pose proof (lt_le_trans _ _ _ h13 h14) as h15. apply lt_irrefl in h15; contradiction. destruct h6 as [h6 [h7 h8]]. destruct h8 as [h8 | h8]. destruct h8 as [h8 [h9 h10]]. rewrite h10 in h4. rewrite plus_assoc in h4. rewrite <- mult_plus_distr_l in h4. rewrite plus_minus_minus_same_le in h4; auto with arith. rewrite mult_plus_distr_l in h4. rewrite <- plus_assoc in h4. rewrite <- h3 in h4. rewrite h4. rewrite plus_comm. rewrite plus_times_mod_r; auto. apply le_div; auto with arith. destruct h8 as [h8 h9]. destruct h9 as [h9 h10]. apply le_antisym; auto. apply mult_le_compat_l. apply le_div; auto with arith. rewrite plus_comm. apply plus_le_compat; auto. apply mult_le_compat_l. apply le_div; auto with arith. rewrite plus_minus_plus_diff_le_lt' in H0. rewrite minus_dist_l in H0. rewrite times_eq_times_minus_iff in H0. destruct H0 as [h6 | h6]. rewrite h6 in h1. apply lt_irrefl in h1; contradiction. destruct h6 as [h6 | h6]. destruct h6 as [h6 h7]. destruct (zerop (j / n - i / n)) as [h8 | h8]. pose proof (minus_div_le j i n) as h9. rewrite h8 in h9. apply le0 in h9. rewrite div_eq_0_iff in h9. destruct h9 as [h9 | h9]. rewrite h9 in h1. apply lt_irrefl in h1. contradiction. pose proof (dv_le _ _ h2) as h10. destruct h10 as [h10 | h10]. rewrite <- le_minus_iff in h10. pose proof (lt_le_trans _ _ _ hlt h10) as h11. apply lt_irrefl in h11; contradiction. pose proof (lt_le_trans _ _ _ h9 h10) as h11. apply lt_irrefl in h11; contradiction. pose proof (mult_O_le n (j / n - i / n)) as h9. destruct h9 as [h9 | h9]. rewrite h9 in h8. apply lt_irrefl in h8; contradiction. rewrite mult_comm in h9. pose proof (lt_mod i n h1) as h10. pose proof (le_minus (i mod n) (j mod n)) as h11. pose proof (lt_le_trans _ _ _ h10 h9) as h12. pose proof (le_lt_trans _ _ _ h11 h12) as h13. pose proof (lt_le_trans _ _ _ h13 h7) as h14. apply lt_irrefl in h14; contradiction. destruct h6 as [h6 h7]. destruct h6 as [? h6]. destruct h7 as [h7 | h7]. destruct h7 as [h7 h8]. rewrite <- le_minus_iff in h8. pose proof (lt_le_trans _ _ _ h5 h8) as h9. apply lt_irrefl in h9; contradiction. destruct h7 as [h7 h8]. pose proof (mult_O_le n (j / n - i / n - n0)) as h9. destruct h9 as [h9 | h9]. rewrite lt_iff in h7. rewrite h9 in h7. apply lt_irrefl in h7. contradiction. rewrite mult_comm in h9. rewrite <- h8 in h9. pose proof (lt_mod i n h1) as h10. pose proof (le_minus (i mod n) (j mod n)) as h11. pose proof (le_lt_trans _ _ _ h9 (le_lt_trans _ _ _ h11 h10)) as h12. apply lt_irrefl in h12; contradiction. apply mult_le_compat_l. apply le_div; auto with arith. assumption. Qed. Lemma plus_mod : forall i j n, (i + j) mod n = (i mod n + j mod n) mod n. intros i j n. destruct (zerop n) as [|h1]; subst. simpl; auto. pose proof h1 as h1'. rewrite <- neq0_iff in h1. pose proof (le_mod i n) as hli. pose proof (le_mod j n) as hlj. pose proof (plus_le_compat _ _ _ _ hli hlj) as hij. apply le_lt_eq_dec in hij. destruct hij as [hij | hij]. pose proof (div_mod i n h1) as hi. pose proof (div_mod j n h1) as hj. pose proof (f_equal (fun c => c - i mod n) hi) as h4. pose proof (f_equal (fun c => c - j mod n) hj) as h5. simpl in h4, h5. rewrite plus_minus_same_l in h4, h5. symmetry. rewrite mod_eq_iff. rewrite plus_minus_plus_diff_le, h4, h5. rewrite <- mult_plus_distr_l. constructor. apply le_mod. apply le_mod. assumption. assumption. rewrite hij; auto. Qed. Lemma minus_mod_le : forall i j n, j <= i -> j mod n <= i mod n -> (i - j) mod n = i mod n - j mod n. intros i j n hij hijn. destruct (zerop n) as [|h1]; subst; simpl; auto. pose proof h1 as h1' . rewrite <- neq0_iff in h1. pose proof (div_mod i n h1) as h2. pose proof (div_mod j n h1) as h3. symmetry. pose proof (f_equal (fun c => c - (n * (i / n))) h2) as h7. simpl in h7. rewrite plus_minus_same_r in h7. pose proof (f_equal (fun c => c - (n * (j / n))) h3) as h8. simpl in h8. rewrite plus_minus_same_r in h8. rewrite <- (lt_mod_eq (i mod n - j mod n) n). rewrite <- h7, <- h8, minus_minus_minus_diff_le'. rewrite minus_dist_l. apply minus_times_mod_r. rewrite h2, h3 at 1. rewrite plus_minus_plus_diff_le. rewrite minus_dist_l. apply le_plus_l. apply mult_le_compat_l. apply le_div; auto. assumption. assumption. apply mult_le_compat_l. apply le_div; auto. rewrite h2 at 2. apply le_plus_l. rewrite h3 at 2. apply le_plus_l. rewrite h8, h7; auto. eapply le_lt_trans. apply le_minus. apply lt_mod; auto. Qed. Lemma minus_mod_lt : forall i j n, j <= i -> i mod n < j mod n -> (i - j) mod n = n - (j mod n - i mod n). intros i j n h1 h2. destruct (zerop n) as [|h3]; subst; simpl; auto. pose proof h3 as h3'. rewrite <- neq0_iff in h3'. pose proof (div_mod (i - j) n h3') as h4. pose proof (div_mod i n h3') as h5. pose proof (div_mod j n h3') as h6. rewrite h5, h6 in h4 at 1. rewrite plus_minus_plus_diff_le' in h4; auto with arith. rewrite minus_dist_l in h4. symmetry in h4. pose proof (f_equal (fun c => c + (j mod n - i mod n)) h4) as h7. simpl in h7. pose proof (le_div j i n h1) as h8. apply le_lt_eq_dec in h8. destruct h8 as [h8 | h8]. rewrite minus_plus_same in h7. rewrite <- plus_assoc in h7. symmetry in h7. rewrite times_eq_times_plus_iff in h7. destruct h7 as [h7 | h7]. destruct h7 as [h7 h9]. destruct h7 as [h7 | h7]. rewrite h7 in h3'; contradict h3'; auto. apply plus_is_O in h9. destruct h9 as [h9 h10]. rewrite <- le_minus_iff in h10. pose proof (lt_le_trans _ _ _ h2 h10) as h11. apply lt_irrefl in h11; contradiction. destruct h7 as [h7 [h9 h10]]. pose proof (lt_mod (i - j) n h3) as h11. pose proof (le_minus (j mod n) (i mod n)) as h12. pose proof (lt_mod j n h3) as h13. pose proof (le_lt_trans _ _ _ h12 h13) as h14. pose proof (plus_lt_compat _ _ _ _ h11 h14) as h15. rewrite <- mult2 in h15. rewrite h10 in h15. rewrite (mult_comm 2) in h15. rewrite (S_pred _ _ h3) in h15 at 1 5. apply mult_S_lt_l in h15. apply gt_lt in h9. rewrite lt_iff in h9. rewrite (S_pred _ _ h9) in h15. apply lt_S_n in h15. apply lt1 in h15. apply pred0_dec in h15. destruct h15 as [h15 | h15]. rewrite h15 in h9. apply lt_irrefl in h9; contradiction. rewrite h15 in h10. rewrite mult_1_r in h10. rewrite <- h10 at 2. rewrite plus_minus_same_l; auto. rewrite lt_iff in h8. pose proof (lt_mod j n h3) as h9. pose proof (le_minus (j mod n) (i mod n)) as h10. pose proof (le_lt_trans _ _ _ h10 h9) as h11. pose proof (mult_O_le n (i / n - j / n)) as h12. destruct h12 as [h12 | h13]. rewrite h12 in h8. apply lt_irrefl in h8. contradiction. apply lt_le_weak. eapply lt_le_trans. apply h11. rewrite mult_comm; auto. rewrite h8 in h6. pose proof (plus_lt_compat_l _ _ (n * (i / n)) h2) as h9. rewrite <- h5, <- h6 in h9. pose proof (lt_le_trans _ _ _ h9 h1) as h10. apply lt_irrefl in h10. contradiction. apply mult_le_compat_l. apply le_div; auto. Qed. Lemma mod_plus_mod_l : forall i j n, ((i mod n) + j) mod n = (i + j) mod n. intros i j n. destruct (zerop n) as [|hn]; subst; simpl; auto. pose proof (le_mod i n) as hl. apply le_lt_eq_dec in hl. destruct hl as [hl|he]. pose proof hn as hn'. rewrite <- neq0_iff in hn. rewrite mod_eq_iff, plus_minus_plus_same_le'. pose proof (div_mod i n hn) as h1. rewrite h1 at 1. rewrite plus_minus_same_l. constructor. apply le_mod. assumption. apply plus_lt_compat_r; auto. rewrite he; auto. Qed. Lemma mod_plus_mod_r : forall i j n, (j + (i mod n)) mod n = (j + i) mod n. intros; rewrite plus_comm, mod_plus_mod_l, plus_comm; auto. Qed. Lemma mod_mod : forall i n, (i mod n) mod n = i mod n. intros i n. destruct (zerop n) as [|h1]; subst; simpl; auto. intros; apply lt_mod_eq. apply lt_mod; auto. Qed. Lemma plus_mod_inj_r : forall i j k n (pfn : 0 < n), (i + k) mod n = (j + k) mod n -> i mod n = j mod n. intros i j k n h1 h2. destruct (lt_eq_lt_dec i j) as [h3 | h3]. destruct h3 as [h3 | h3]. rewrite mod_eq_iff in h2; auto. rewrite plus_minus_plus_diff_lt_le' in h2; auto with arith. rewrite minus_same, <- minus_n_O in h2. rewrite mod_eq_iff; auto. omega. subst; auto. symmetry. symmetry in h2. rewrite mod_eq_iff in h2. rewrite plus_minus_plus_diff_lt_le' in h2; auto with arith. rewrite minus_same, <- minus_n_O in h2. rewrite mod_eq_iff; auto. assumption. omega. Qed. Lemma plus_mod_inj_l : forall i j k n (pfn : 0 < n), (k + i) mod n = (k + j) mod n -> i mod n = j mod n. intros i j k n h1 h2. rewrite plus_comm, (plus_comm k j) in h2. apply plus_mod_inj_r in h2; auto. Qed. Definition shift_seg_fun (n k:nat) : seg_fun n nat := fun i (pf:i < n) => (i + k) mod n. Lemma lt_shift_seg_fun : forall n k i (pf:i < n), shift_seg_fun n k i pf < n. unfold shift_seg_fun; intros; apply lt_mod; omega. Qed. Definition left_shift_seg_fun (n:nat) : seg_fun n nat := shift_seg_fun n (pred n). Definition right_shift_seg_fun (n:nat) : seg_fun n nat := shift_seg_fun n 1. Lemma lt_left_shift_seg_fun : forall n i (pf:i < n), left_shift_seg_fun n i pf < n. unfold left_shift_seg_fun; intros; apply lt_shift_seg_fun. Qed. Lemma lt_right_shift_seg_fun : forall n i (pf:i < n), right_shift_seg_fun n i pf < n. unfold left_shift_seg_fun; intros; apply lt_shift_seg_fun. Qed. Lemma inj_shift_seg_fun : forall n k, inj_seg (shift_seg_fun n k). intros n k. red. intros i j h1 h2 h3. unfold shift_seg_fun in h3. apply plus_mod_inj_r in h3. do 2 rewrite lt_mod_eq in h3; auto; subst. omega. Qed. Lemma inj_left_shift_seg_fun : forall n, inj_seg (left_shift_seg_fun n). unfold left_shift_seg_fun; intros; apply inj_shift_seg_fun. Qed. Lemma inj_right_shift_seg_fun : forall n, inj_seg (right_shift_seg_fun n). unfold left_shift_seg_fun; intros; apply inj_shift_seg_fun. Qed. Lemma shift_seg_fun0 : forall n, shift_seg_fun n 0 = id_seg n. unfold shift_seg_fun, id_seg. intro. apply functional_extensionality_dep. intro. apply functional_extensionality. intro. rewrite plus_0_r, lt_mod_eq; auto. Qed. (*This would go in FunctionProperties if not for dependencies.*) Definition shift_functor {U:Type} {n:nat} (f:seg_fun n U) (k:nat) (pf0 : 0 < n) : seg_fun n U := fun i (pf:i < n) => f ((i + k) mod n) (lt_mod _ _ pf0). Definition left_shift_functor {U:Type} {n:nat} (f:seg_fun n U) (pf0 : 0 < n) : seg_fun n U := shift_functor f (pred n) pf0. Definition right_shift_functor {U:Type} {n:nat} (f:seg_fun n U) (pf0 : 0 < n) : seg_fun n U := shift_functor f 1 pf0. Lemma inj_shift_functor : forall {U:Type} {n:nat} (f:seg_fun n U) (k:nat) (pf0 : 0 < n), inj_seg f -> inj_seg (shift_functor f k pf0). intros U n f k h1 h2. red in h2; red. intros i j hi hj. unfold shift_functor. intro h3. apply h2 in h3. apply plus_mod_inj_r in h3; auto. rewrite <- neq0_iff in h1. pose proof (div_eq_0_iff i n) as h4. pose proof (div_eq_0_iff j n) as h5. hr h4. right; auto. hr h5. right; auto. rewrite <- h4 in hr. rewrite <- h5 in hr0. pose proof (div_mod i n h1) as h6. pose proof (div_mod j n h1) as h7. rewrite hr in h6. rewrite hr0 in h7. rewrite mult_0_r, plus_0_l in h6, h7; auto. rewrite h6, h7; auto. Qed. Lemma inj_left_shift_functor : forall {U:Type} {n:nat} (f:seg_fun n U) (pf0 : 0 < n), inj_seg f -> inj_seg (left_shift_functor f pf0). unfold left_shift_functor; intros; apply inj_shift_functor; auto. Qed. Lemma inj_right_shift_functor : forall {U:Type} {n:nat} (f:seg_fun n U) (pf0 : 0 < n), inj_seg f -> inj_seg (right_shift_functor f pf0). unfold right_shift_functor; intros; apply inj_shift_functor; auto. Qed. Section EvenOdd. Scheme even_ind := Minimality for even Sort Prop with odd_ind := Minimality for odd Sort Prop. Lemma even_decompose : forall n:nat, even n -> exists k:nat, n = (2 * k)%nat. apply (even_ind (fun n => exists k, n = (2*k)%nat) (fun n => exists k, n = (2*k + 1)%nat)). exists 0%nat. omega. intros n h1 h2. destruct h2 as [k h2]. subst. exists (k + 1)%nat. omega. intros n h1 h2. destruct h2 as [k h2]. subst. exists k. omega. Qed. Lemma odd_decompose : forall n:nat, odd n -> exists k:nat, n = (2 * k + 1)%nat. pose proof (even_ind (fun n => exists k, S n = (2*k + 1)%nat) (fun n => exists k, S n = (2*k)%nat)) as h0. cut (forall n:nat, even n -> exists k:nat, S n = (2 * k + 1)%nat). intros h1 n h2. destruct h2 as [n h2]. apply h1; auto. apply h0. exists 0%nat. omega. intros n h1 h2. destruct h2 as [k h2]. rewrite h2. exists k. omega. intros n h1 h2. destruct h2 as [k h2]. rewrite h2. exists (k + 1)%nat. omega. Qed. Lemma odd_1 : odd 1. do 2 constructor. Qed. Lemma even_2 : even 2. do 3 constructor. Qed. Lemma odd_eq_1_dec : forall (n:nat), odd n -> {n = 1%nat} + {(n > 1)%nat}. intros n h1. destruct (lt_eq_lt_dec n 1) as [[h2 | h2] | h2]. assert (h3:n = 0%nat). omega. subst. contradict (not_even_and_odd _ even_O h1). left; auto. right; auto. Qed. Lemma odd_ge_1 : forall (n:nat), odd n -> 1 <= n. intros n h1. destruct (odd_eq_1_dec _ h1); subst; auto with arith. Qed. Lemma odd_Sn_iff : forall (n:nat), odd (S n) <-> even n. intro n. destruct n as [|n]. split. intro h1. destruct h1. constructor. intro; constructor; auto. split. intro h1. inversion h1. subst. assumption. intro h1. constructor. assumption. Qed. Lemma even_Sn_iff : forall (n:nat), even (S n) <-> odd n. intro n. destruct n as [|n]. split. intro h1. inversion h1. contradict (not_even_and_odd _ even_O H0). intro h1. contradict (not_even_and_odd _ even_O h1). split. intro h1. inversion h1. assumption. intro h1. constructor. assumption. Qed. Lemma odd_pred_iff : forall (n:nat), (0 < n)%nat -> (odd (pred n) <-> even n). intros n h0. destruct n as [|n]. omega. simpl. rewrite even_Sn_iff. tauto. Qed. Lemma even_pred_iff : forall (n:nat), (0 < n)%nat -> (even (pred n) <-> odd n). intros n h0. destruct n as [|n]. omega. simpl. rewrite odd_Sn_iff. tauto. Qed. Lemma le_even_odd : forall (m n:nat), even m -> odd n -> (m <= n)%nat -> (m <= pred n)%nat. intros m n h1 h2 h3. apply even_decompose in h1. apply odd_decompose in h2. destruct h1 as [k h1], h2 as [k' h2]. omega. Qed. Lemma le_odd_even : forall (m n:nat), odd m -> even n -> (m <= n)%nat -> (m <= pred n)%nat. intros m n h1 h2 h3. apply odd_decompose in h1. apply even_decompose in h2. destruct h1, h2; omega. Qed. Lemma odd_gt_0 : forall n:nat, odd n -> 0 < n. intros n h1. destruct (zerop n) as [h2 |]; auto. subst. inversion h1. Qed. End EvenOdd. (*This section defines the natural numbers from 1,2, ... (excluding 0) and is built so that inductive proofs involving it are more natural than using a lower-bounded sigma type.*) Section Nat1. Inductive nat1 : nat -> Set:= | nat1_1 : nat1 1 | nat1_S : forall {n:nat}, nat1 n -> nat1 (S n). Inductive nat1t : Set := nat1t_intro : forall {n':nat} (n:nat1 n'), nat1t. Definition nat1t_S (n:nat1t) : nat1t. destruct n as [n' n]. refine (nat1t_intro (nat1_S n)). Defined. Lemma not_nat1_0 : forall n:nat1 0, False. intro n. inversion n. Qed. Definition nat_of_1t (n:nat1t) : nat. destruct n as [n' n]. refine n'. Defined. Lemma nat1_functional : forall (n' : nat) (n : nat1 n') (m : nat1 n'), m = n. intros n' n m. dependent induction n; dependent induction m; auto. pose proof (not_nat1_0 m) as h1. contradiction. pose proof (not_nat1_0 n0) as h1. contradiction. f_equal. apply IHn. Qed. Lemma nat1t_functional : forall (m n:nat1t), (let (m'', _) := m in let (n'', _) := n in m'' = n'') -> m = n. intros m n. simpl. destruct m as [m' m], n as [n' n]. intro; subst. f_equal. apply nat1_functional. Qed. Lemma nat_of_1t_inj : Injective nat_of_1t. red. intros x y h1. unfold nat_of_1t in h1. destruct x as [x h2], y as [y h3]. subst. f_equal. apply nat1_functional. Qed. Lemma nat_of_1t_S : forall (n:nat1t), nat_of_1t (nat1t_S n) = S (nat_of_1t n). intro n. destruct n; auto. Qed. Definition plus1 {m' n':nat} (m:nat1 m') (n:nat1 n') : nat1 (m'+n'). induction m as [|m']. eapply nat1_S; auto. simpl. constructor; auto. Qed. Definition plus1t (m n:nat1t) : nat1t. destruct m as [m' m], n as [n' n]. pose (plus1 m n) as p. eapply nat1t_intro. refine p. Defined. Definition times1 {m' n':nat} (m:nat1 m') (n:nat1 n') : nat1 (m'*n'). induction m as [|m' s' s]. rewrite mult_1_l; auto. rewrite mult_comm. rewrite <- mult_n_Sm. pose (plus1 n s) as p. rewrite plus_comm in p. rewrite mult_comm in p. refine p. Defined. Definition times1t (m n:nat1t) : nat1t. destruct m as [m' m], n as [n' n]. pose (times1 m n) as p. eapply nat1t_intro. refine p. Defined. Definition n1_1 : nat1 1. constructor. Defined. Definition n1t_1 : nat1t := (nat1t_intro (n1_1)). Definition n1_2 : nat1 2. do 2 constructor. Defined. Definition n1t_2 : nat1t := (nat1t_intro (n1_2)). Definition even1 {n':nat} (n:nat1 n') : Prop := even n'. Definition odd1 {n':nat} (n:nat1 n') : Prop := odd n'. Definition even1t (n:nat1t) : Prop. destruct n as [n' n]. refine (even1 n). Defined. Definition odd1t (n:nat1t) : Prop. destruct n as [n' n]. refine (odd1 n). Defined. Lemma even_odd_dec1 : forall {n':nat} (n:nat1 n'), {even1 n} + {odd1 n}. intros n' n. destruct (even_odd_dec n') as [h1 | h1]. left. red. assumption. right. red. assumption. Qed. Lemma even_odd_dec_1t : forall (n:nat1t), {even1t n} + {odd1t n}. intro n. destruct n as [n' n]. apply (even_odd_dec1 n). Qed. Lemma not_even_and_odd_1t : forall n:nat1t, even1t n -> odd1t n -> False. intros n h1 h2. destruct n as [n' n]. red in h1, h2. apply (not_even_and_odd _ h1 h2). Qed. Lemma odd1_1 : odd1 (n1_1). red. apply odd_1. Qed. Lemma odd1t_1 : odd1t (n1t_1). unfold n1t_1, odd1t. apply odd1_1. Defined. Lemma even1_2 : even1 (n1_2). red. apply even_2. Qed. Lemma even1t_2 : even1t (n1t_2). unfold n1t_2, even1t. apply even1_2. Qed. Lemma odd_Sn_iff1 : forall {n':nat} (n:nat1 n'), odd1 (nat1_S n) <-> even1 n. intros n' n. split. intro h1. red in h1. red. rewrite <- odd_Sn_iff; auto. intro h1. red in h1. red. rewrite odd_Sn_iff; auto. Qed. Lemma odd_Sn_iff1t : forall (n:nat1t), odd1t (nat1t_S n) <-> even1t n. intro n. unfold nat1t_S. destruct n as [n' n]. simpl. apply odd_Sn_iff1. Qed. Lemma even_Sn_iff1 : forall {n':nat} (n:nat1 n'), even1 (nat1_S n) <-> odd1 n. intros n' n. split. intro h1. red in h1. red. rewrite <- even_Sn_iff; auto. intro h1. red in h1. red. rewrite even_Sn_iff; auto. Qed. Lemma even_Sn_iff1t : forall (n:nat1t), even1t (nat1t_S n) <-> odd1t n. intro n. unfold nat1t_S. destruct n as [n' n]. simpl. apply even_Sn_iff1. Qed. Definition pred1 {n':nat} (n:nat1 n') : nat1 (match n' with | 0 => 1 | 1 => 1 | k => pred k end). destruct n as [|n' n]. refine n1_1. destruct n as [|n' n]. simpl. refine n1_1. simpl. apply nat1_S; auto. Defined. Definition pred1t (n:nat1t) : nat1t. destruct n as [n' n]. refine (nat1t_intro (pred1 n)). Defined. Lemma gt0_to_1t_ex : forall {n:nat}, 0 < n -> exists! n1t:nat1t, nat_of_1t n1t = n. intros n h1. induction n as [|n h2]. contradict (lt_irrefl _ h1). destruct (zerop n) as [h3 | h3]. subst. exists n1t_1. red; simpl. split; auto. intros n h3. destruct n as [n' n]. simpl in h3. subst. pose proof (nat1t_functional n1t_1 (nat1t_intro n)) as h3. simpl in h3. apply h3; auto. specialize (h2 h3). destruct h2 as [n1t h2]. red in h2. destruct h2 as [h2 h4]. subst. exists (nat1t_S n1t). red; simpl; split; auto. unfold nat1t_S. destruct n1t. simpl. reflexivity. intros n h5. unfold nat1t_S. destruct n1t. simpl in h5. specialize (h4 (pred1t n)). simpl in h4. unfold nat_of_1t in h4. destruct n as [n h6]. simpl in h4. destruct n as [|n]. contradict (not_nat1_0 h6). destruct n as [|n]. simpl in h5. assert (h7:n' = 0). omega. subst. contradict (not_nat1_0 n0). simpl in h4. unfold nat_of_1t in h5. apply S_inj in h5. specialize (h4 h5). subst. injection h4. intro h7. unfold pred1 in h7. simpl in h7. dependent induction h6. dependent induction h6. f_equal. f_equal. apply inj_pair2 in h7. subst. simpl in h4. unfold n1_1. reflexivity. f_equal. f_equal. apply inj_pair2 in h7. assumption. Qed. Definition gt0_to_1t {n:nat} (pf:0 < n) := proj1_sig (constructive_definite_description _ (gt0_to_1t_ex pf)). Lemma gt0_to_1t_compat : forall {n:nat} (pf:0 < n), let n1t := gt0_to_1t pf in nat_of_1t n1t = n. unfold gt0_to_1t. intros. destruct constructive_definite_description. simpl. assumption. Qed. Lemma gt0_to_1t_S : forall (n:nat) (pf:0 < n) (pf':0 < S n), gt0_to_1t pf' = nat1t_S (gt0_to_1t pf). intros n h1 h2. pose proof (gt0_to_1t_compat h2) as h3. simpl in h3. pose proof (gt0_to_1t_compat h1) as h4. simpl in h4. apply nat_of_1t_inj. rewrite nat_of_1t_S. rewrite h3, h4. reflexivity. Qed. Lemma gt0_to_1t_inj : forall (m n:nat) (pfm:0 < m) (pfn:0 < n), gt0_to_1t pfm = gt0_to_1t pfn -> m = n. intros m n h1 h2 h3. pose proof (gt0_to_1t_compat h1) as h4. pose proof (gt0_to_1t_compat h2) as h5. simpl in h4, h5. rewrite <- h4, <- h5 at 1. f_equal. assumption. Qed. Lemma even_gt0_to_1t_iff : forall (n:nat) (pf:0 < n), even n <-> even1t (gt0_to_1t pf). intros n h1. pose proof (gt0_to_1t_compat h1) as h2. simpl in h2. rewrite <- h2 at 1. unfold even1t. unfold nat_of_1t. destruct (gt0_to_1t h1). subst. unfold even1. tauto. Qed. Definition odd_to_1t {n:nat} (pf:odd n) : nat1t. pose proof (odd_gt_0 _ pf) as h1. refine (gt0_to_1t h1). Defined. Lemma odd_to_1t_compat : forall {n:nat} (pf:odd n), odd1t (odd_to_1t pf). intros n h1. red. unfold odd_to_1t. unfold gt0_to_1t. destruct constructive_definite_description as [n' h2]. simpl. destruct n' as [n'' n']. red. simpl in h2. subst. assumption. Qed. Definition nat1t_lt (m n:nat1t) : Prop. destruct m as [m' m], n as [n' n]. refine (m' < n'). Defined. Definition nat1t_le (m n:nat1t) : Prop. destruct m as [m' m], n as [n' n]. refine (m' <= n'). Defined. Lemma nat1t_le_iff : forall (m n:nat1t), nat1t_le m n <-> let (m', _) := m in let (n', _) := n in m' <= n'. intros m n. destruct m as [m h1], n as [n h2]. unfold nat1t_le. tauto. Qed. Lemma nat1t_lt_iff : forall (m n:nat1t), nat1t_lt m n <-> let (m', _) := m in let (n', _) := n in m' < n'. intros m n. destruct m as [m h1], n as [n h2]. unfold nat1t_lt. tauto. Qed. Lemma le_nat_from_1t_iff : forall (m:nat) (n:nat1t) (pf:0 < m), m <= nat_of_1t n <-> nat1t_le (gt0_to_1t pf) n. intros m n h1. pose proof (gt0_to_1t_compat h1) as h2. simpl in h2. unfold nat1t_le. destruct (gt0_to_1t h1). destruct n as [a b]. subst. simpl. tauto. Qed. Lemma odd_pred_iff1 : forall {n':nat} (n:nat1 n'), 1 < n' -> ((odd1 (pred1 n) <-> even1 n)). intros n' n h0. split. intro h1. red in h1. red. destruct n' as [|n']. constructor. destruct n' as [|n']. omega. simpl in h1. rewrite <- odd_pred_iff. simpl. assumption. omega. intro h1. red in h1. red. destruct n' as [|n']. constructor. assumption. destruct n' as [|n']. omega. rewrite odd_pred_iff; auto. omega. Qed. Lemma odd_pred_iff1t : forall (n:nat1t), nat1t_lt n1t_1 n -> ((odd1t (pred1t n) <-> even1t n)). intros n h1. red in h1. destruct n as [n' n]. simpl. unfold n1t_1 in h1. apply odd_pred_iff1; auto. Qed. Lemma even_pred_iff1 : forall {n':nat} (n:nat1 n'), 1 < n' -> (even1 (pred1 n) <-> odd1 n). intros n' n h0. split. intro h1. red in h1. red. destruct n' as [|n']. omega. destruct n' as [|n']. omega. simpl in h1. rewrite <- even_pred_iff. simpl. assumption. omega. intro h1. red in h1. red. destruct n' as [|n']. constructor. assumption. destruct n' as [|n']. omega. rewrite even_pred_iff; auto. omega. Qed. Lemma even_pred_iff1t : forall (n:nat1t), nat1t_lt n1t_1 n -> (even1t (pred1t n) <-> odd1t n). intros n h1. unfold nat1t_lt, n1t_1 in h1. destruct n as [n' n]. simpl. apply even_pred_iff1; auto. Qed. Lemma gt1t_dec : forall (n:nat1t), {n = n1t_1} + {nat1t_lt n1t_1 n}. intro n. destruct n as [n' n]. destruct (lt_eq_lt_dec n' 1) as [h2 | h2]. destruct h2 as [h2 | h2]. apply lt1 in h2. subst. pose proof (not_nat1_0 n). contradiction. subst. left. unfold n1t_1, n1_1. f_equal. apply nat1_functional. right. red. simpl. assumption. Qed. Section EvenOddNat1. Inductive even_nat1t : Set := | even_nat1t_intro : forall (n:nat1t), even1t n -> even_nat1t. Inductive odd_nat1t : Set := | odd_nat1t_intro : forall (n:nat1t), odd1t n -> odd_nat1t. Definition odd_to_on1t {n:nat} (pf:odd n) : odd_nat1t := odd_nat1t_intro _ (odd_to_1t_compat pf). Lemma even_nat1t_functional : forall (m n:nat1t) (pfm:even1t m) (pfn:even1t n), m = n -> even_nat1t_intro m pfm = even_nat1t_intro n pfn. intros; subst; pose proof (proof_irrelevance _ pfn pfm); subst; auto. Qed. Lemma odd_nat1t_functional : forall (m n:nat1t) (pfm:odd1t m) (pfn:odd1t n), m = n -> odd_nat1t_intro m pfm = odd_nat1t_intro n pfn. intros; subst; pose proof (proof_irrelevance _ pfn pfm); subst; auto. Qed. Definition even_nat1ts := Full_set even_nat1t. Definition odd_nat1ts := Full_set odd_nat1t. Definition nat_of_en1t (n:even_nat1t) : nat. destruct n as [n n']. refine (nat_of_1t n). Defined. Definition nat_of_on1t (n:odd_nat1t) : nat. destruct n as [n n']. refine (nat_of_1t n). Defined. Definition en1t_lt : even_nat1t -> even_nat1t -> Prop. intros m n. destruct m as [m h1], n as [n h2]. refine (nat1t_lt m n). Defined. Definition on1t_lt : odd_nat1t -> odd_nat1t -> Prop. intros m n. destruct m as [m h1], n as [n h2]. refine (nat1t_lt m n). Defined. Definition en1t_le : even_nat1t -> even_nat1t -> Prop. intros m n. destruct m as [m h1], n as [n h2]. refine (nat1t_le m n). Defined. Definition on1t_le : odd_nat1t -> odd_nat1t -> Prop. intros m n. destruct m as [m h1], n as [n h2]. refine (nat1t_le m n). Defined. Lemma on1t_le_refl : forall n:odd_nat1t, on1t_le n n. intro n. red. destruct n as [n h1]. red. destruct n. auto with arith. Qed. Lemma en1t_le_refl : forall n:even_nat1t, en1t_le n n. intro n. red. destruct n as [n h1]. red. destruct n. auto with arith. Qed. Lemma lt_1_nat1t : forall n:nat1t, 1 <= nat_of_1t n. intro n. destruct n as [n' n]. simpl. destruct (lt_eq_lt_dec 1 n') as [h1 | h1]. destruct h1; auto with arith. subst. auto with arith. assert (n' = 0). omega. subst. inversion n. Qed. Lemma le_1_odd1t : forall (n:odd_nat1t), 1 <= nat_of_on1t n. intro n. destruct n as [n h1]. apply lt_1_nat1t; auto. Qed. Lemma onep_on1t : forall n:odd_nat1t, {1 = nat_of_on1t n} + {1 < nat_of_on1t n}. intro n. pose proof (le_1_odd1t n) as h1. apply le_lt_eq_dec in h1. destruct h1. right; auto. left; auto. Qed. Lemma en1t_ge_2 : forall m:even_nat1t, 2 <= nat_of_en1t m. intro m. destruct m as [m' h1]. destruct m' as [m'' m']. red in h1. simpl. red in h1. apply even_decompose in h1. destruct h1 as [k h1]. subst. destruct k as [|k]. simpl in m'. inversion m'. omega. Qed. Lemma en1t_gt_1 : forall m:even_nat1t, 1 < nat_of_en1t m. intro m. pose proof (en1t_ge_2 m) as h1. destruct m as [m' h2]. auto with arith. Qed. Lemma en1t_ge_1 : forall m:even_nat1t, 1 <= nat_of_en1t m. intro m. pose proof (en1t_gt_1 m) as h1. destruct m as [m' h2]. auto with arith. Qed. Lemma nat_of_en1t_le_compat : forall (m n:even_nat1t), nat_of_en1t m <= nat_of_en1t n -> en1t_le m n. intros m n h1. red. destruct m as [m h2], n as [n h3]. simpl in h1. red. destruct m as [m' m], n as [n' n]. simpl in h1. assumption. Qed. Lemma nat_of_on1t_le_compat : forall (m n:odd_nat1t), nat_of_on1t m <= nat_of_on1t n -> on1t_le m n. intros m n h1. red. destruct m as [m h2], n as [n h3]. simpl in h1. red. destruct m as [m' m], n as [n' n]. simpl in h1. assumption. Qed. Lemma even_nat_of_en1t : forall (n:even_nat1t), even (nat_of_en1t n). intro n. pose proof (en1t_ge_1 n) as h1. assert (h2:0 < nat_of_en1t n). auto with arith. rewrite (even_gt0_to_1t_iff _ h2). destruct n as [n h3]. unfold gt0_to_1t. destruct constructive_definite_description as [n' h4]. simpl. simpl in h4. apply nat_of_1t_inj in h4. subst. assumption. Qed. Lemma not_odd_nat_of_en1t : forall (n:even_nat1t), ~odd (nat_of_en1t n). intro n. pose proof (even_nat_of_en1t n) as h1. intro h2. apply (not_even_and_odd _ h1 h2). Qed. Lemma odd_nat_of_on1t : forall (n:odd_nat1t), odd (nat_of_on1t n). intro n. unfold nat_of_on1t. destruct n as [n' h1]. red in h1. destruct n' as [n'' n']. simpl. red in h1. assumption. Qed. Lemma not_even_nat_of_on1t : forall (n:odd_nat1t), ~even (nat_of_on1t n). intro n. pose proof (odd_nat_of_on1t n) as h1. intro h2. apply (not_even_and_odd _ h2 h1). Qed. Lemma match_gt0_to_1t_even_iff : forall (m:nat) (pf:0 < m), (match gt0_to_1t pf with | nat1t_intro n' n0 => even1 n0 end) <-> even (nat_of_1t (gt0_to_1t pf)). intros m h0. unfold gt0_to_1t. destruct constructive_definite_description as [n h1]. simpl. subst. destruct n as [n' n]. simpl. unfold even1. tauto. Qed. Lemma nat_of_1t_pred_compat : forall (n:nat1t), 1 < nat_of_1t n -> nat_of_1t (pred1t n) = pred (nat_of_1t n). intros n h0. unfold nat_of_1t, pred1t. destruct n as [n h1]. destruct n. inversion h1. dependent induction h1. simpl in h0. contradict (lt_irrefl _ h0). destruct n as [|n]. simpl in h0. contradict (lt_irrefl _ h0). reflexivity. Qed. Lemma on1t_pred_ex : forall n:odd_nat1t, 1 < nat_of_on1t n -> exists! m:even_nat1t, nat_of_en1t m = pred (nat_of_on1t n). intros n. destruct n as [n' h0]. intro h1. pose (pred1t n') as m. pose proof (even_pred_iff1t n'). simpl in h1. assert (h1':nat1t_lt n1t_1 n'). red. simpl. destruct n'. assumption. pose proof (even_pred_iff1t n' h1') as h2. pose proof (iff2 h2 h0) as h0'. exists (even_nat1t_intro _ h0'). red; split; auto. simpl. apply nat_of_1t_pred_compat; auto. intros e h3. destruct e as [e h4]. subst. f_equal. simpl in h3. rewrite <- nat_of_1t_pred_compat in h3; auto. pose proof nat_of_1t_inj as h5. red in h5. specialize (h5 _ _ h3). subst. f_equal. apply proof_irrelevance. Qed. Definition on1t_pred (n:odd_nat1t) (pf: 1 < nat_of_on1t n) : even_nat1t := proj1_sig (constructive_definite_description _ (on1t_pred_ex n pf)). Lemma on1t_pred_compat : forall (n:odd_nat1t) (pf:1 < nat_of_on1t n), let m := on1t_pred n pf in nat_of_en1t m = pred (nat_of_on1t n). unfold on1t_pred. intros. destruct constructive_definite_description. simpl. assumption. Qed. (*This defaults to "2" when passed an argument of "1"*) Definition on1t_pred' (n:odd_nat1t) : even_nat1t. destruct (onep_on1t n) as [h1 | h1]. refine (even_nat1t_intro _ even1t_2). refine (on1t_pred _ h1). Defined. Lemma en1t_pred_ex : forall n:even_nat1t, exists! m:odd_nat1t, nat_of_on1t m = pred (nat_of_en1t n). intros n. pose proof (en1t_gt_1 n) as h0. destruct n as [n' n]. assert (h1':nat1t_lt n1t_1 n'). red. simpl. destruct n' as [n'' n']. simpl in h0. assumption. pose proof (odd_pred_iff1t _ h1') as h2. pose proof (iff2 h2 n) as h0'. exists (odd_nat1t_intro _ h0'). red; split; auto. simpl. apply nat_of_1t_pred_compat; auto. intros e h3. destruct e as [e h4]. subst. f_equal. simpl in h3. rewrite <- nat_of_1t_pred_compat in h3; auto. pose proof nat_of_1t_inj as h5. red in h5. specialize (h5 _ _ h3). subst. f_equal. apply proof_irrelevance. Qed. Definition en1t_pred (n:even_nat1t) : odd_nat1t := proj1_sig (constructive_definite_description _ (en1t_pred_ex n)). Lemma en1t_pred_compat : forall (n:even_nat1t) (pf:1 < nat_of_en1t n), let m := en1t_pred n in nat_of_on1t m = pred (nat_of_en1t n). unfold en1t_pred. intros. destruct constructive_definite_description. simpl. assumption. Qed. Lemma nat_of_en1t_le_pred_compat : forall (m:even_nat1t) (n:odd_nat1t), nat_of_en1t m <= nat_of_on1t n -> en1t_le m (on1t_pred' n). intros m n h1. unfold on1t_pred'. destruct onep_on1t as [h2 | h2]. red. destruct m as [m' m]. rewrite <- h2 in h1. unfold nat_of_en1t in h1. red. destruct m' as [m'' m']. simpl. simpl in h1. auto with arith. red. destruct m as [m' m]. unfold on1t_pred. destruct constructive_definite_description as [n' h3]. simpl. destruct n' as [n'' n']. simpl in h2. simpl in h1. red. destruct m' as [m'' m']. destruct n'' as [n''' n'']. simpl in h1. simpl in h3. subst. red in m. red in m. assert (h3: m'' <> nat_of_on1t n). intro h3. subst. destruct n as [o h3]. simpl in m. red in h3. destruct o as [o' o]. red in h3. simpl in m. apply (not_even_and_odd _ m h3). omega. Qed. Lemma nat_of_on1t_le_pred_compat : forall (m:odd_nat1t) (n:even_nat1t), nat_of_on1t m <= nat_of_en1t n -> on1t_le m (en1t_pred n). intros m n h1. unfold en1t_pred. destruct m as [m' m]. unfold on1t_pred. destruct constructive_definite_description as [n' h3]. simpl. destruct n' as [n'' n']. simpl in h1. red. destruct m' as [m'' m']. destruct n'' as [n''' n'']. simpl in h1. simpl in h3. subst. red in m. red in m. assert (h3: m'' <> nat_of_en1t n). intro h3. subst. destruct n as [o h3]. simpl in m. red in h3. destruct o as [o' o]. red in h3. simpl in m. apply (not_even_and_odd _ h3 m). omega. Qed. Lemma en1t_le_iff : forall (m n:even_nat1t), en1t_le m n <-> let (m', _) := m in let (n', _) := n in nat1t_le m' n'. intros m n. destruct m as [m h1], n as [n h2]. split. intro h3. red in h3. assumption. intro h3. red. assumption. Qed. End EvenOddNat1. Lemma S_le_1 : forall (n:nat), 1 <= S n. intros; auto with arith. Qed. Lemma st01_pf : 0 < 1. auto with arith. Qed. Lemma st02_pf : 0 < 2. auto with arith. Qed. Lemma st12_pf : 1 < 2. auto with arith. Qed. (*For use in conjunction with [seg_one_t] defined in SetUtilities.*) Lemma sot11_pf : 1 <= 1 <= 1. auto with arith. Qed. Lemma sot12_pf : 1 <= 1 <= 2. auto with arith. Qed. Lemma sot22_pf : 1 <= 2 <= 2. auto with arith. Qed. Lemma sot_cond : forall n:nat1t, let (n', _) := n in 1 <= n' <= n'. intro n. pose proof (lt_1_nat1t n) as h1. destruct n. simpl in h1. auto with arith. Qed. Lemma lt_sot_iff : forall (m n:nat), m < n <-> 1 <= S m <= n. intros; omega. Qed. End Nat1. Section Integers. Require Import ZArith. Lemma Zpos_inj : forall p q, Zpos p = Zpos q -> p = q. intros ? ? h1. inversion h1; auto. Qed. Lemma Zneg_inj : forall p q, Zneg p = Zneg q -> p = q. intros ? ? h1. inversion h1; auto. Qed. Lemma Zlt_impl_Zle : forall (a b:Z), (a < b)%Z -> (a <= b)%Z. intros; omega. Qed. Lemma Z_decompose : forall (a:Z), (a = Z.sgn a * (Zpos (Z.to_pos (Z.abs a))))%Z. intro a; destruct a; simpl; auto. Qed. Lemma Zlt_neg : forall (p q:positive), (Z.pos p < Z.pos q)%Z -> (Z.neg p > Z.neg q)%Z. intros p q h1. inversion h1. red. simpl. rewrite H0. simpl; auto. Qed. Lemma Zlt_mult_mono : forall p q r, (Z.pos p < Z.pos q)%Z -> (Z.pos p * Z.pos r < Z.pos q * Z.pos r)%Z. intros p q r h1. red; red in h1. simpl; simpl in h1. rewrite Pos.mul_comm. rewrite (Pos.mul_comm q r). rewrite Pos.mul_compare_mono_l; auto. Qed. Lemma Zgt_mult_mono : forall p q r, (Z.pos p > Z.pos q)%Z -> (Z.pos p * Z.pos r > Z.pos q * Z.pos r)%Z. intros p q r h1. red; red in h1. simpl; simpl in h1. rewrite Pos.mul_comm. rewrite (Pos.mul_comm q r). rewrite Pos.mul_compare_mono_l; auto. Qed. Lemma Zlt_mult_neg : forall (a b:Z) (p:positive), (a < b)%Z -> (a * Z.neg p > b * Z.neg p)%Z. intros a b p h1. destruct a. simpl. destruct b. omega. simpl. constructor. inversion h1. simpl. destruct b; simpl. inversion h1. apply Zlt_neg. do 2 rewrite Pos2Z.inj_mul. apply Zlt_mult_mono; auto. inversion h1. destruct b; simpl. constructor. constructor. do 2 rewrite Pos2Z.inj_mul. apply Zgt_mult_mono; auto. red in h1; red. simpl in h1; simpl. rewrite <- Pos.compare_antisym in h1. pose proof (f_equal CompOpp h1) as h2. rewrite <- Pos.compare_antisym in h2. simpl in h2; auto. Qed. Lemma Zlt_mult : forall (a b c:Z), (c < 0)%Z -> (a < b)%Z -> (a * c > b * c)%Z. intros a b c. revert a b. induction c. intros; omega. intros a b h2 h1. inversion h2. intros a b h1 h2. apply Zlt_mult_neg; auto. Qed. Lemma Z_neg_prod : forall (a b:Z), (a < 0)%Z -> (b < 0)%Z -> (a * b > 0)%Z. intro a. destruct a. intro b. destruct b. intros; omega. intros; omega. intros; omega. intro b. destruct b. intros; omega. intros h1 h2. destruct (Z_dec (Z.pos p * Z.pos p0) 0) as [h3 | h3]. destruct h3 as [h3 | h3]. inversion h3. inversion h2. inversion h2. intro h1. inversion h1. intro b. destruct b. intros; omega. intros ? h1. inversion h1. intros. simpl. constructor. Qed. Lemma Z_mixed_prod : forall (a b:Z), (a < 0)%Z -> (b > 0)%Z -> (a * b < 0)%Z. intro a. destruct a. intro b. destruct b. intros; omega. intros; omega. intros; omega. intros b h1. inversion h1. intro b. destruct b. intros; omega. intros; simpl. constructor. intros ? h1. inversion h1. Qed. Lemma Z_pos_prod : forall (a b:Z), (a > 0)%Z -> (b > 0)%Z -> (a * b > 0)%Z. intro a. destruct a. intros; omega. intro b. destruct b. intros; omega. intros; simpl. constructor. intros ? h1. inversion h1. intros ? h1. inversion h1. Qed. Lemma no_zero_divisors : forall (a b:Z), (a * b)%Z = 0%Z -> {a = 0%Z} + {b = 0%Z}. intros a b h1. destruct (Z_dec a 0) as [h2 | h2]. destruct h2 as [h2 | h2]. right. destruct (Z_dec b 0) as [h3 | h3]. destruct h3 as [h3 | h3]. pose proof (Z_neg_prod _ _ h2 h3) as h4. omega. pose proof (Z_mixed_prod _ _ h2 h3) as h4. omega. assumption. right. destruct (Z_dec b 0) as [h3 | h3]. destruct h3 as [h3 | h3]. pose proof (Z_mixed_prod _ _ h3 h2) as h4. rewrite Zmult_comm in h4. omega. pose proof (Z_pos_prod _ _ h2 h3) as h4. omega. assumption. left; auto. Qed. Lemma Zpos1 : forall (p:positive), (Zpos (p~1) = 2 * Zpos p + 1)%Z. intro p; induction p; simpl; auto. Qed. Lemma Zpos0 : forall (p:positive), (Zpos (p~0) = 2 * Zpos p)%Z. intro p; induction p; simpl; auto. Qed. Lemma Zneg1 : forall (p:positive), (Zneg (p~1) = 2 * Zneg p - 1)%Z. intro p; induction p; simpl; auto. Qed. Lemma Zneg0 : forall (p:positive), (Zneg (p~0) = 2 * Zneg p)%Z. intro p; induction p; simpl; auto. Qed. Lemma pos_to_nat_compat : forall (p:positive), Pos.to_nat p = Z.to_nat (Zpos p). intro p; induction p; simpl; auto. Qed. Lemma mult_to_nat : forall (n:nat) (a:Z), n * Z.to_nat a = Z.to_nat ((Z.of_nat n) * a). intro n. induction n as [|n ih]; simpl; auto. intro a. destruct a; simpl. apply mult_0_r. specialize (ih (Zpos p)). simpl in ih. rewrite ih. rewrite Pos2Nat.inj_mul. rewrite SuccNat2Pos.id_succ. rewrite Z2Nat.inj_mul. rewrite Nat2Z.id. rewrite <- pos_to_nat_compat. rewrite <- S_compat. rewrite mult_plus_distr_r. rewrite mult_1_l. rewrite plus_comm; auto. omega. apply Zle_0_pos. omega. Qed. Lemma pos_to_nat1 : forall (p:positive), Pos.to_nat p~1 = 2 * Pos.to_nat p + 1. intro p. rewrite pos_to_nat_compat. rewrite Zpos1. rewrite pos_to_nat_compat. rewrite Z2Nat.inj_add. f_equal. simpl. rewrite plus_0_r. rewrite <- mult2. do 2 rewrite pos_to_nat_compat. rewrite mult_to_nat. f_equal. pose proof (Zle_0_pos p) as h1. pose proof (Z.mul_0_r 2) as h2. rewrite <- h2. apply Z.mul_le_mono_nonneg; omega. omega. Qed. Lemma pred_double_Z : forall p, Z.pos (Pos.pred_double p) = Z.pred_double (Zpos p). auto. Qed. Lemma Z_pred_inj : forall i j, Z.pred i = Z.pred j -> i = j. intro i. induction i; simpl. intros j h1. pose proof (f_equal Z.succ h1) as h2. simpl in h2. rewrite Z.succ_pred in h2; auto. destruct p. intros j h1. pose proof (f_equal Z.succ h1) as h2. simpl in h2. rewrite h2, Z.succ_pred; auto. intros j h1. pose proof (f_equal Z.succ h1) as h2. rewrite Zpos0. rewrite pred_double_Z in h1. rewrite Z.pred_double_spec in h1. omega. intros; omega. intros j h1. rewrite <- Pos2Z.add_neg_neg in h1. omega. Qed. End Integers. Section Rationals. Require Import QArith. Lemma Qminus_pos_Qlt : forall (b q:Q), (b > 0) -> (q - b < q). intros b q h1. rewrite Qlt_minus_iff. unfold Qminus. rewrite Qopp_plus at 1. rewrite Qopp_opp. rewrite Qplus_assoc. assert (h3:Qeq (q - q) 0). ring. rewrite h3. rewrite Qplus_0_l at 1. assumption. Qed. Lemma Qlt_plus_pos : forall (b q:Q), (b > 0) -> (q < q + b). intros b q h1. rewrite Qlt_minus_iff. rewrite <- Qplus_assoc. rewrite Qplus_comm. rewrite <- Qplus_assoc. rewrite (Qplus_comm (-q) q). assert (h2:q - q == 0). ring. rewrite h2. rewrite Qplus_0_r. assumption. Qed. Lemma Qlt_midpoint : forall (p q:Q), p < q -> p < (p + (q - p) / (2 # 1)). intros p q h1. pose proof (Qplus_lt_l p q (-p)) as h2. pose proof (iff2 h2 h1) as h3. rewrite Qplus_opp_r in h3. assert (h4: 0 < / (2 # 1)). constructor. pose proof (Qmult_lt_l 0 (q + -p) (Qinv (2#1)) h4) as h5. pose proof (iff2 h5 h3) as h6. rewrite Qmult_0_r in h6. rewrite Qmult_comm in h6. pose proof (Qplus_lt_l 0 ( (q + - p) * / (2 # 1)) p) as h7. rewrite Qplus_0_l in h7. rewrite Qplus_comm. rewrite h7 at 1. assumption. Qed. Lemma midpoint_Qlt : forall (p q:Q), p < q -> (p + (q - p) / (2 # 1)) < q. intros p q h1. assert (h7:p*(2 #1) + -p == p). simpl. rewrite <- (Qplus_inj_r (p * (2 # 1) + - p) p p). rewrite <- Qplus_assoc. rewrite (Qplus_comm (-p) p). rewrite Qplus_opp_r. rewrite Qplus_0_r. rewrite <- (Qmult_1_r p) at 2 3. rewrite <- Qmult_plus_distr_r. assert (h7:2#1 = 1 + 1). constructor. rewrite h7. reflexivity. assert (h7':q*(2 #1) + -q == q). simpl. rewrite <- (Qplus_inj_r (q * (2 # 1) + - q) q q). rewrite <- Qplus_assoc. rewrite (Qplus_comm (-q) q). rewrite Qplus_opp_r. rewrite Qplus_0_r. rewrite <- (Qmult_1_r q) at 2 3. rewrite <- Qmult_plus_distr_r. assert (h7':2#1 = 1 + 1). constructor. rewrite h7'. reflexivity. rewrite <- h7 in h1. rewrite <- h7' in h1. pose proof (Qplus_lt_l (p * (2 # 1) + - p) ( q * (2 # 1) + - q) q) as h8. pose proof (iff2 h8 h1) as h9. rewrite <- (Qplus_assoc (q * (2 # 1)) (- q) q) in h9. rewrite (Qplus_comm (-q) q) in h9. rewrite Qplus_opp_r in h9. rewrite Qplus_0_r in h9. rewrite <- Qplus_assoc in h9. rewrite (Qplus_comm (-p) q) in h9. assert (h10:0 < (Qinv (2#1))). constructor. pose proof (Qmult_lt_r (p * (2 # 1) +(q + - p)) (q * (2 # 1)) _ h10) as h11. pose proof (iff2 h11 h9) as h12. rewrite Qmult_plus_distr_l in h12. rewrite <- Qmult_assoc in h12. rewrite <- Qmult_assoc in h12. assert (h13: ~ 2 # 1 == 0). intro h13. inversion h13. pose proof (Qmult_inv_r (2 # 1) h13) as h14. rewrite h14 in h12. do 2 rewrite Qmult_1_r in h12. assumption. Qed. Lemma Qlt_impl_Qle : forall (a b:Q), a < b -> a <= b. intros a b h1. destruct a as [an ad], b as [bn bd]. apply Zlt_impl_Zle; auto. Qed. Definition p_to_Z (p:positive) : Z := Z.of_nat (Pos.to_nat p). Lemma qmake0 : forall (p q:positive), Qmake Z0 p == Qmake Z0 q. intro p. intros q. red; simpl; auto. Qed. Lemma Zpos_mul : forall (p q:positive), Zpos (Pos.mul p q) = (Zpos p * Zpos q)%Z. intros; simpl; auto. Qed. Lemma Q_inj_mult : forall q r s:Q, q == r -> s * q == s * r. intros q r s. revert q r. induction s as [n d]. intros q r h1. red. red in h1. simpl. rewrite <- Zmult_assoc. symmetry. rewrite <- Zmult_assoc. f_equal. rewrite Zpos_mul. rewrite Zmult_assoc. rewrite Zmult_comm. rewrite Zmult_assoc. rewrite (Zmult_comm _ (Qnum r)). rewrite <- h1. rewrite <- Zmult_assoc. f_equal. simpl. rewrite Zpos_mul. simpl. f_equal. apply Pos.mul_comm. Qed. Lemma Z_cancel_r : forall (a:Z) (p:positive), a * ' p # p == a # 1. intros a p; red; simpl. rewrite Z.mul_1_r; auto. Qed. Lemma Z_cancel_l : forall (a:Z) (p:positive), ' p * a # p == a # 1. intros a ?; red; simpl; destruct a; simpl; auto; rewrite Pos.mul_1_r, Pos.mul_comm; auto. Qed. Lemma Q_neq_dec : forall (q r:Q), ~q == r -> {q < r} + {q > r}. intros q r ?; destruct (Q_dec q r) as [| h2]; auto; contradiction. Qed. Lemma Qnum_recip : forall (n:Z) (d:positive), n <> 0%Z -> (QArith_base.Qnum (/(n # d)) = Zsgn n * Zpos d)%Z. intros n; induction n; intros; try omega; simpl; auto. Qed. Definition Qmin (a b:Q) := match (Q_dec a b) with | inleft P => if P then a else b | _ => b end. Definition Qmax (a b:Q) := match (Q_dec a b) with | inleft P => if P then b else a | _ => a end. Lemma Qmin_le : forall (a b:Q), Qmin a b <= a /\ Qmin a b <= b. intros a b. unfold Qmin. destruct Q_dec as [h1 | h1]. destruct h1 as [h1 | h1]. constructor. apply Qle_refl. apply Qlt_impl_Qle; auto. split. apply Qlt_impl_Qle; auto. apply Qle_refl. rewrite h1. split; apply Qle_refl. Qed. Lemma le_Qmax : forall (a b:Q), a <= Qmax a b /\ b <= Qmax a b. intros a b. unfold Qmax. destruct Q_dec as [h1 | h1]. destruct h1 as [h1 | h1]. constructor. apply Qlt_impl_Qle. assumption. apply Qle_refl. split. apply Qle_refl. apply Qlt_impl_Qle; auto. rewrite h1. split; apply Qle_refl. Qed. End Rationals. (*This section is used to develop lemmas that allow you to destruct a [nat] discriminant in a match clause, without losing information if that discriminant is a functional expression with non-primitive actual paramaters. Of course a different lemma will be needed for each form of the match clause that is currently in use. No attempt is made to be "comprehensive."*) Section MatchDestruct. Lemma match_destruct_nat_zerop : forall {T U V:Type} (P:nat->U->V->T) (u:U) (v:V)(n:nat) (a:T), match n with | O => a | S m => P m u v end = let f := (fun m:nat => if (zerop m) then a else (P (pred m) u v)) in f n. intros T U V P u v n a. destruct n. simpl. reflexivity. simpl. reflexivity. Qed. End MatchDestruct. bool2-v0-3/FunctionProperties3.v0000644000175000017500000001241713575574055017463 0ustar dwyckoff76dwyckoff76(*Copyright (C) 2019, Daniel Wyckoff*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import NaryFunctions. Require Import FunctionalExtensionality. Require Import LogicUtilities. Require Import SetUtilities. Require Import SetUtilities2. Require Import FunctionProperties. Require Import FunctionProperties2. Require Import TypeUtilities. (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Lemma in_seg_fun_depT_nth_to_tuple_prod : forall {T n} {U:T->Type} (A:forall t, Ensemble (U t)) (x:T^n) (f:seg_fun_depT_nth U x) (pfi:forall i (pfi:i < n), Inn (A (nth_prod i pfi x)) (f i pfi)), Inn (tuple_prod_ens_dep A x) (seg_fun_depT_nth_to_tuple_prod x f). intros T n. induction n as [|n ih]; simpl; intros; constructor; intros i hi; try omega. destruct i as [|i]; simpl. specialize (pfi 0 (lt_0_Sn _)); simpl in pfi; auto. specialize (ih U A (snd x) (seg_fun_depT_nth_to_S f)). hfirst ih. intros j hj. specialize (pfi (S j) (lt_n_S _ _ hj)); simpl in pfi; unfold seg_fun_depT_nth_to_S; simpl. unfold seg_fun_depT_to_S. term0 pfi. gterm0. simpl in c, c0. pose proof (f_equal_depT' (nth_prod j (lt_S_n _ _ (lt_n_S _ _ hj)) (snd x)) (nth_prod j hj (snd x)) c c0 (nth_prod_functional2 j (lt_S_n _ _ (lt_n_S _ _ hj)) _ _)) as h2. hfirst h2. unfold c, c0. symmetry. rewrite <- transfer_dep_r_eq_iff. fold c c0. eapply f_equal_depT'. unfold c. erewrite transfer_dep_transfer_dep_r_compat. reflexivity. specialize (h2 hf). fold c c0. symmetry in h2. gterm0. fold c1. assert (h3:c0 = c1). unfold c0, c1. repeat f_equal. rewrite <- h3. dependent rewrite -> h2; auto. specialize (ih hf). inversion ih; auto. Grab Existential Variables. apply nth_prod_functional2. Qed. Lemma sig_tuple_prod_ens_dep_const_to_seg_fun_depT_nth_to_tuple_prod : forall {T U n} (A:T->Ensemble U) (wt:T^n) (wu:U^n) (pfi:forall j (pfj:j < n), Inn (A (nth_prod j pfj wt)) (nth_prod j pfj wu)), proj1_sig (sig_tuple_prod_ens_dep_const_to A wt (exist (fun c:tuple_prod (fun _:T =>U) wt => Inn (tuple_prod_ens_dep A wt) c) (seg_fun_depT_nth_to_tuple_prod wt (fun j (pfj:j < n) => nth_prod j pfj wu) (U := (fun _:T => U))) (in_seg_fun_depT_nth_to_tuple_prod A wt (fun j (pfj:j < n) => nth_prod j pfj wu) pfi))) = wu. intros T U n. induction n as [|n ih]; simpl. intros; apply unit_eq. intros A wt wu h1. destruct wt as [wtf wts], wu as [wuf wus]; simpl. f_equal. specialize (ih A wts wus). hfirst ih. intros j hj. specialize (h1 (S j) (lt_n_S _ _ hj)); simpl in h1. pose proof (proof_irrelevance _ hj (lt_S_n _ _ (lt_n_S _ _ hj))) as h2. rewrite h2; auto. specialize (ih hf). simpl in ih. terml ih. redterm0 x. gterml. redterm0 x0. assert (h3:c = c0). unfold c, c0. f_equal. apply functional_extensionality_dep; intro k; apply functional_extensionality; intro hk. unfold seg_fun_depT_nth_to_S. symmetry. simpl. rewrite <- (transfer_dep_eq_iff (U := fun _:T => U)). pose proof (proof_irrelevance _ hk (lt_S_n _ _ (lt_n_S _ _ hk))) as h2. rewrite <- h2. f_equal. unfold seg_fun_depT_to_S. repeat f_equal; auto. fold c0; fold c in ih. rewrite <- h3; auto. Qed. Lemma nth_prod_sig_tuple_prod_ens_dep_const_to : forall {T U n} (A:T->Ensemble U) (wt:T^n) (wu:sig_set (tuple_prod_ens_dep A wt)) k (hk:k < n), nth_prod k hk (proj1_sig (sig_tuple_prod_ens_dep_const_to A wt wu)) = nth_tuple_prod (proj1_sig wu) k hk. intros T U n. induction n as [|n ih]; simpl. intros; omega. intros A wt wu k hk. destruct wt as [ft wt], wu as [u wu]; destruct u as [fu su]; simpl; auto. destruct k as [|k]; auto. simpl in su, wu; simpl. assert (h1:Inn (tuple_prod_ens_dep A wt) su). constructor. intros i hi. destruct wu as [wu]. specialize (wu (S i) (lt_n_S _ _ hi)). simpl in wu. pose proof (proof_irrelevance _ hi (lt_S_n _ _ (lt_n_S _ _ hi))) as h1. rewrite <- h1 in wu; auto. specialize (ih A wt (exist (fun c => Inn (tuple_prod_ens_dep A wt) c) _ h1) k (lt_S_n _ _ hk)). simpl in ih; auto. Qed. bool2-v0-3/ZornsLemma.v0000644000175000017500000002606213575574055015626 0ustar dwyckoff76dwyckoff76(*This file is a complete copy and paste from Daniel Schepler's "Zorn's Lemma." If I add to it, then I will add a copyright.*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Version 8.4 pl4 Coq.*) Require Export Classical. Require Import ClassicalChoice. Require Import Relations. Require Export Relation_Definitions. Require Import Relation_Definitions_Implicit. Require Import SetUtilities. Require Import LogicUtilities. Require Import SetUtilities. Section ZL'. Variable T:Type. Variable R:relation T. Hypothesis ord:order R. Definition chain (S: Ensemble T) : Prop := forall x y:T, In S x -> In S y -> (R x y \/ R y x). Definition maximal (x:T) : Prop := forall y:T, R x y -> x = y. Variable chain_sup: forall S: Ensemble T, chain S -> { x:T | (forall y:T, In S y -> R y x) /\ (forall z:T, (forall y:T, In S y -> R y z) -> R x z) }. Variable inflation: forall x:T, { y:T | R x y /\ x <> y /\ forall z:T, R x z -> R z y -> z = x \/ z = y }. Inductive tower : Ensemble T := | sup_intro: forall (S: Ensemble T), Included S tower -> forall c:chain S, In tower (proj1_sig (chain_sup S c)) | inflation_intro: forall x:T, In tower x -> In tower (proj1_sig (inflation x)). Lemma tower_is_chain: chain tower. Proof. unfold chain. intros. revert x H. induction H0. intros. case (classic (forall y:T, In S y -> R y x)). right. apply (proj2_sig (chain_sup S c)). assumption. intro. apply not_all_ex_not in H2. destruct H2. apply imply_to_and in H2. destruct H2. left. apply (ord_trans ord) with x0. pose proof (H0 x0 H2 x H1). tauto. apply (proj2_sig (chain_sup S c)). assumption. pose proof (proj2_sig (inflation x)). simpl in H0. destruct H0. destruct H1. remember (proj1_sig (inflation x)) as x'. assert (forall y:T, In tower y -> R y x \/ R x' y). intros. induction H3. case (classic (forall x0:T, In S x0 -> R x0 x)). left. apply (proj2_sig (chain_sup S c)). assumption. right. apply not_all_ex_not in H5. destruct H5. apply imply_to_and in H5. destruct H5. apply (ord_trans ord) with x0. pose proof (H4 x0). tauto. apply (proj2_sig (chain_sup S c)). assumption. assert (In tower x'). rewrite Heqx'. apply inflation_intro. assumption. destruct IHtower0. assert (In tower (proj1_sig (inflation x0))). apply inflation_intro. assumption. case (IHtower (proj1_sig (inflation x0)) H6). left. assumption. intro. pose proof (proj2_sig (inflation x0)). simpl in H8. assert (x = x0 \/ x = proj1_sig (inflation x0)). apply H8. assumption. assumption. case H9. right. rewrite Heqx'. rewrite H10. apply (ord_refl ord). left. rewrite H10. apply (ord_refl ord). right. apply (ord_trans ord) with x0. assumption. apply (proj2_sig (inflation x0)). intros. case (H3 x0 H4). left. apply (ord_trans ord) with x. assumption. assumption. right. assumption. Qed. (* can now show the given hypotheses are contradictory *) Lemma ZL': False. Proof. pose proof (proj2_sig (chain_sup tower tower_is_chain)). simpl in H. remember (proj1_sig (chain_sup tower tower_is_chain)) as x0. assert (In tower x0). rewrite Heqx0. constructor 1. auto with sets. pose (x' := proj1_sig (inflation x0)). assert (In tower x'). constructor 2. assumption. pose proof (proj2_sig (inflation x0)). simpl in H2. destruct H2. destruct H3. contradict H3. apply (ord_antisym ord). assumption. destruct H. apply H. assumption. Qed. End ZL'. Implicit Arguments chain [[T]]. Implicit Arguments maximal [[T]]. Section ZL. (* get rid of the need for a sup function and immediate successors *) Variable T:Type. Variable R:relation T. Hypothesis ord: order R. Hypothesis ub_of_chain: forall S:Ensemble T, chain R S -> exists x:T, forall y:T, In S y -> R y x. Definition chains := {S:Ensemble T | chain R S}. Definition chains_ord := (fun S1 S2:chains => Included (proj1_sig S1) (proj1_sig S2)). Lemma chains_order: order chains_ord. Proof. constructor. unfold reflexive. unfold chains_ord. auto with sets. unfold transitive. unfold chains_ord. auto with sets. unfold antisymmetric. unfold chains_ord. intros. apply proj1_sig_injective. auto with sets. Qed. Definition chains_sup_def : forall F: Ensemble chains, chain chains_ord F -> chains. refine (fun F H => exist _ [ x:T | exists S:chains, In F S /\ In (proj1_sig S) x ] _). red; intros. destruct H0. destruct H1. destruct H0. destruct H1. pose proof (H x0 x1). destruct x0. destruct x1. simpl in H0. destruct H0. simpl in H1. destruct H1. apply H2 in H0. unfold chains_ord in H0. simpl in H0. destruct H0. apply c0. apply H0. exact H3. exact H4. apply c. exact H3. apply H0. exact H4. assumption. Defined. Lemma chains_sup_correct: forall (F:Ensemble chains) (P:chain chains_ord F), let U := chains_sup_def F P in (forall S:chains, In F S -> chains_ord S U) /\ (forall T:chains, (forall S:chains, In F S -> chains_ord S T) -> chains_ord U T). Proof. intros. split. intros. unfold chains_ord. unfold Included. intros. unfold U. simpl. constructor. exists S. auto. intros. unfold chains_ord. unfold Included. intros. unfold U in H0. simpl in H0. destruct H0. destruct H0. destruct H0. apply H in H0. apply H0. assumption. Qed. Definition chains_sup (F:Ensemble chains) (P:chain chains_ord F) := let U := chains_sup_def F P in exist (fun U:chains => (forall S:chains, In F S -> chains_ord S U) /\ (forall T:chains, (forall S:chains, In F S -> chains_ord S T) -> chains_ord U T)) (chains_sup_def F P) (chains_sup_correct F P). Theorem ZornsLemma: exists x:T, maximal R x. Proof. pose proof (ZL' chains chains_ord chains_order chains_sup). apply NNPP. unfold maximal. intuition. assert (forall x:T, ~ forall y:T, R x y -> x = y). apply not_ex_all_not. exact H0. assert (forall x:T, exists y:T, ~ (R x y -> x = y)). intro. apply not_all_ex_not. apply H1. assert (forall x:T, exists y:T, R x y /\ x <> y). intro. pose proof (H2 x). destruct H3. apply imply_to_and in H3. exists x0. assumption. clear H0 H1 H2. assert (forall x:chains, exists y:chains, chains_ord x y /\ x <> y /\ forall z:chains, chains_ord x z -> chains_ord z y -> z = x \/ z = y). intro. destruct x. pose proof (ub_of_chain x c). destruct H0. pose proof (H3 x0). destruct H1. pose (x' := Add x x1). assert (chain R x'). unfold chain. intros. case H2. case H4. intros. apply c. assumption. assumption. intros. destruct H5. left. apply ord_trans with x0. apply H0. assumption. apply H1. intros. destruct H5. case H4. right. apply ord_trans with x0. apply H0. assumption. apply H1. intros. destruct H5. left. apply ord_refl. exists (exist _ x' H2). split. unfold chains_ord. simpl. unfold Included. intros. constructor. assumption. split. intuition. injection H4. intro. assert (In x x1). rewrite H1. constructor 2. auto with sets. contradict H6. apply ord_antisym. assumption. apply H0. assumption. intros. destruct z. unfold chains_ord in H4. simpl in H4. unfold chains_ord in H5. simpl in H5. case (classic (In x2 x1)). right. apply subset_eq_compatT. apply Extensionality_Ensembles. split. assumption. unfold Included. intros. case H7. exact H4. intros. destruct H8. assumption. left. apply subset_eq_compatT. apply Extensionality_Ensembles. split. unfold Included. intros. assert (In x' x3). apply H5. assumption. inversion H8. assumption. destruct H9. contradiction H6. assumption. apply choice in H0. destruct H0 as [f]. apply H. intro. apply exist with (f x). apply H0. Qed. End ZL. Implicit Arguments ZornsLemma [[T]]. Require Import Quotients. Section ZL_preorder. Variable T:Type. Variable R:relation T. Hypothesis Rpreord: preorder R. Hypothesis ub_of_chain: forall S:Ensemble T, chain R S -> exists x:T, forall y:T, In S y -> R y x. Definition premaximal (x:T) : Prop := forall y:T, R x y -> R y x. Lemma ZornsLemmaForPreorders: exists x:T, premaximal x. Proof. pose (Requiv (x y:T) := R x y /\ R y x). assert (equivalence Requiv). constructor. red; intros. split; apply preord_refl; trivial. red; intros. destruct H. destruct H0. split; apply preord_trans with y; trivial. red; intros. destruct H. split; trivial. pose (Rquo := quotient Requiv). let Hnew:=fresh in refine (let Hnew:=_ in let inducedR := induced_function2arg R H H Hnew in let inducedR_prop := induced_function2arg_correct R H H Hnew in _). intros. assert (True_rect (R a1 b1) = True_rect (R a2 b2)). apply Extensionality_Ensembles; split; red; intros. destruct x. red in H2. simpl in H2. red; simpl. apply preord_trans with a1; trivial. apply H0. apply preord_trans with b1; trivial. apply H1. destruct x. red in H2; simpl in H2. red; simpl. apply preord_trans with a2; trivial. apply H0. apply preord_trans with b2; trivial. apply H1. assert (True_rect (R a1 b1) I = True_rect (R a2 b2) I). rewrite H2; trivial. simpl in H3. assumption. clearbody inducedR_prop. fold inducedR in inducedR_prop. assert (exists x:Rquo, maximal inducedR x). apply ZornsLemma. constructor. red; intros xbar. destruct (quotient_projection_surjective xbar) as [x []]. rewrite inducedR_prop. apply preord_refl; trivial. red; intros xbar ybar zbar ? ?. destruct (quotient_projection_surjective xbar) as [x]. destruct H2. destruct (quotient_projection_surjective ybar) as [y]. destruct H2. destruct (quotient_projection_surjective zbar) as [z]. destruct H2. rewrite inducedR_prop in H0, H1. rewrite inducedR_prop. apply preord_trans with y; trivial. red; intros xbar ybar ? ?. destruct (quotient_projection_surjective xbar) as [x]. destruct H2. destruct (quotient_projection_surjective ybar) as [y]. destruct H2. rewrite inducedR_prop in H0, H1. unfold quotient_projection. apply subset_eq_compatT. apply R_impl_equality_of_equiv_class; trivial. split; trivial. intros Sbar ?. pose (S := inverse_image (quotient_projection _) Sbar). let H:=fresh in refine (let H:=ub_of_chain S _ in _). red; intros. pose proof (H0 (quotient_projection _ x) (quotient_projection _ y)). rewrite 2 inducedR_prop in H3. destruct H1. destruct H2. apply H3; trivial. destruct H1. exists (quotient_projection _ x). intros ybar ?. destruct (quotient_projection_surjective ybar) as [y]. destruct H3. rewrite inducedR_prop. apply H1. constructor; exact H2. destruct H0 as [xbar]. destruct (quotient_projection_surjective xbar) as [x]. destruct H1. exists x. red; intros. red in H0. let H:=fresh in refine (let H:=H0 (quotient_projection Requiv y) _ in _). rewrite inducedR_prop. assumption. unfold quotient_projection in H2. injection H2; intros. assert (In (equiv_class Requiv x) y). rewrite H3. constructor; apply equiv_refl; trivial. destruct H4. apply H4. Qed. End ZL_preorder. Arguments premaximal [T] _ _.bool2-v0-3/PartialFunctionAlgebras.v0000644000175000017500000002534013575574055020300 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2019, Daniel Wyckoff.*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Import NaryFunctions. Require Import FunctionAlgebrasBasics. Require Import PartialAlgebrasBasics. Require Import ListUtilities. Require Import FunctionProperties. Require Import ArithUtilities. Require Import SetUtilities. Require Import TypeUtilities. Require Import TypeUtilities2. Require Import SetUtilities1_5. Require Import FiniteMaps. Section PartialFunctionAlgebras. Definition function_algebra_sig : list nat := 1 :: 1 :: 1 :: 1 :: 2 :: nil. (*This is an interpretation of the Mal'Tsev symbols (integers 0-4) into operations. The specific numbering of the indexing of Mal'Tsev operations is arbitrary, yet conssitent with the text and previous code.*) Definition Jfun T : seg_fun 5 (arb_function (arb_function T)) := fun i => match i with | 0 => fun _ => unary_to_arb_function (@mzeta_to_arb T) | S i0 => match i0 with | 0 => fun _ => unary_to_arb_function (@mtau_to_arb T) | S i1 => match i1 with | 0 => fun _ => unary_to_arb_function (@mdelta_to_arb T) | S i2 => match i2 with | 0 => fun _ => unary_to_arb_function (@mnabla_to_arb T) | S i3 => match i3 with | 0 => fun _ => binary_to_arb_function (@mstar_to_arb T) | S i3 => fun pflt: S (S (S (S (S i3)))) < 5 => let pflt' := lt_S_n _ _ (lt_S_n _ _ (lt_S_n _ _ (lt_S_n _ _ (lt_S_n _ _ pflt)))) in False_rect _ (lt_n_O _ pflt') end end end end end. Lemma Jfun_ar_compat : forall T i (pfi:i < length function_algebra_sig), ar (Jfun T i pfi) = nth_lt _ i pfi. intros T i; do 5 (destruct i; simpl; auto). intro; omega. Qed. End PartialFunctionAlgebras. (*The purpose of this section is to explore the finitary aspects of function generation, so that rather than look at sublcasses (which are infinite and usually have a finite number of generators), we can use the Mal'Tsev Operations to create relations between elements, much like the Schreier vector in computational group theory (and its attendant graph).*) Section ExperimentalPointwiseGeneration. (*Example: [pointwise_generates {f g h i} (op h, op' i)*) Inductive pointwise_generates {T} (E:Ensemble (arb_function T)) : Ensemble (arb_function T) -> Prop := | pointwise_generates_empty : pointwise_generates E (Empty_set _) | mal_tsev_op1_generates : forall (F:Ensemble (arb_function T)), pointwise_generates E F -> forall m n (op:function T m -> function T n), mal_tsev_op1 op -> forall (f:function T m), Inn E (existT _ _ f) -> pointwise_generates E (Add F (existT _ _ (op f))) | mal_tsev_op2_generates : forall (F:Ensemble (arb_function T)), pointwise_generates E F -> forall m n (op:function T n -> function T m -> function T (pred (m + n))), mal_tsev_op2 op -> forall (f:function T n) (g:function T m), Inn E (existT _ _ f) -> Inn E (existT _ _ g) -> pointwise_generates E (Add F (existT _ _ (op f g))). (*Example: [pointwise_strictly_generates {h i} (op h, op' i)*) Inductive pointwise_strictly_generates {T} : Ensemble (arb_function T) -> Ensemble (arb_function T) -> Prop := | pointwise_strictly_generates_empty : pointwise_strictly_generates (Empty_set _) (Empty_set _) | mal_tsev_op1_strictly_generates : forall n m (E F:Ensemble (arb_function T)) (f:function T n) (op:function T n->function T m), mal_tsev_op1 op -> pointwise_strictly_generates E F -> pointwise_strictly_generates (Add E (existT _ _ f)) (Add E (existT _ _ (op f))) | mal_tsev_op2_strictly_generates : forall n m (E F:Ensemble (arb_function T)) (f:function T n) (g:function T m) (op:function T n->function T m->function T (pred (m+n))), mal_tsev_op2 op -> pointwise_strictly_generates E F -> pointwise_strictly_generates (Add (Add E (existT _ _ f)) (existT _ _ g)) (Add E (existT _ _ (op f g))). Lemma pointwise_generates_included1 : forall {T} (E E' F:Ensemble (arb_function T)), Included E E' -> pointwise_generates E F -> pointwise_generates E' F. intros T E E' F h1 h2. revert h1. revert E'. induction h2. intros; constructor. intros E' h1. specialize (IHh2 _ h1). econstructor. assumption. apply H. apply h1; auto. intros E' h1. specialize (IHh2 _ h1). eapply mal_tsev_op2_generates; auto. Qed. Lemma pointwise_generates_empty1 : forall {T:Type} (E:Ensemble (arb_function T)), pointwise_generates (Empty_set _) E -> E = Empty_set _. intros T E h1. inversion h1; subst; auto; try contradiction. Qed. Lemma pointwise_generates_finite : forall {T:Type} (E F:Ensemble (arb_function T)), pointwise_generates E F -> Finite F. intros ? ? ? h1; induction h1; try apply Empty_is_finite; try apply Add_preserves_Finite; auto. Qed. Lemma pointwise_generates_add1 : forall {T n} (E F:Ensemble (arb_function T)) (f:function T n), Inn E (existT _ _ f) -> forall (op:function T n -> function T n), mal_tsev_op1 op -> pointwise_generates E F -> pointwise_generates E (Add F (existT _ _ (op f))). intros; constructor; auto. Qed. Lemma pointwise_generates_add2 : forall {T n m} (E F:Ensemble (arb_function T)) (f:function T n) (g:function T m), Inn E (existT _ _ f) -> Inn E (existT _ _ g) -> forall (op:function T n -> function T m -> function T (pred (m+n))), mal_tsev_op2 op -> pointwise_generates E F -> pointwise_generates E (Add F (existT _ _ (op f g))). intros; apply mal_tsev_op2_generates; auto. Qed. Lemma pointwise_generates_in_impl : forall {T}, FiniteT T -> forall (E F:Ensemble (arb_function T)) f, F <> Empty_set _ -> pointwise_generates E F -> Inn F f -> exists e, Inn E e /\ ((exists (op:function T (ar e) -> function T (ar f)), mal_tsev_op1 op /\ projT2 f = op (projT2 e)) \/ (exists g (op:function T (ar e) -> function T (ar g) -> function T (pred (ar g + ar e))) (pfe:ar f = pred (ar g + ar e)), mal_tsev_op2 op /\ function_transfer pfe (projT2 f) = op (projT2 e) (projT2 g))). intros T hfin E F f he h1. revert f he. induction h1. intros; contradiction. intros f0 h0 h2. pose proof (pointwise_generates_finite _ _ h1) as h3. destruct (eq_empty_dec (arb_function_eq_dec hfin) _ h3) as [h4|h4]; subst. rewrite add_empty_sing in h2. destruct h2; subst. exists (existT _ _ f). split; auto. left. simpl. exists op; auto. inversion h2 as [c h5|c h5]; subst. specialize (IHh1 f0 h4 h5); auto. destruct h5; subst. exists (existT _ _ f). split; auto. left. simpl. exists op; auto. intros f0 h2 h3. pose proof (pointwise_generates_finite _ _ h1) as h4. destruct (eq_empty_dec (arb_function_eq_dec hfin) _ h4) as [h5|h5]; subst. rewrite add_empty_sing in h3. destruct h3; subst. exists (existT _ _ f). split; auto. right. simpl. exists (existT _ _ g), op, (eq_refl _); split; auto. inversion h3 as [c h6|c h6]; subst. specialize (IHh1 f0 h5 h6); auto. destruct h6; subst. exists (existT _ _ f). split; auto. right; simpl. exists (existT _ _ g), op, (eq_refl _); split; auto. Qed. End ExperimentalPointwiseGeneration. (*This is merely an experiment in finite homomoprhisms on function algebras.*) Section FiniteHomomorphisms. Inductive Fin_hom {T} : Ensemble (arb_function T) -> Ensemble (arb_function T) -> arb_function T -> Type := fin_hom_intro1 : forall (E F:Ensemble (arb_function T)) (def:arb_function T), Finite E -> Finite F -> forall E' F', pointwise_generates E' E -> pointwise_generates F' F -> forall (h:Fin_map E F def), (forall m n (e:function T n) (pfe:n = ar (h |-> existT _ _ e)), Inn E' (existT _ _ e) -> forall (op:function T n -> function T m) (pfe':m = ar (h |-> existT _ _ (op e))), Inn E (existT _ _ (op e)) -> mal_tsev_op1 op -> Inn F' (h |-> (existT _ _ e)) /\ Inn F (h |-> (existT _ _ (op e))) /\ h |-> (existT _ _ (op e)) = existT _ _ (op (function_transfer (eq_sym pfe) (projT2 (h |-> existT _ _ e))))) -> Fin_hom E F def. End FiniteHomomorphisms. bool2-v0-3/FiniteOperations.v0000644000175000017500000112710213575574055017017 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014-2015 Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Import SetUtilities. Require Import LogicUtilities. Require Export BoolAlgBasics. Require Import FieldsOfSets. Require Import InfiniteOperations. Require Import TypeUtilities. Require Import ArithUtilities. Require Import FunctionProperties. Require Import Equality. Require Import FunctionalExtensionality. Require Import FiniteMaps. Require Import DecidableDec. Require Import Equality. Require Import ListUtilities. Require Import ListUtilities2. Require Import Description. Require Import EqDec. Section ListOperations. Variable B:Bool_Alg. Let Bt := bt B. Fixpoint times_list {B0:Bool_Alg} (l:list (bt B0)) : (bt B0) := match l with | nil => 1 | cons a l' => a * (times_list l') end. Fixpoint plus_list {B0:Bool_Alg} (l:list (bt B0)) : (bt B0) := match l with | nil => 0 | cons a l' => a + (plus_list l') end. Lemma plus_list_app : forall (l1 l2:list Bt), plus_list (l1 ++ l2) = plus_list l1 + plus_list l2. intros l1 l2. induction l1 as [|a l1' h1]. simpl. rewrite comm_sum. rewrite zero_sum. reflexivity. (* a::l1'*) simpl. rewrite h1. apply assoc_sum. Qed. Lemma times_list_app : forall (l1 l2:list Bt), times_list (l1 ++ l2) = times_list l1 * times_list l2. intros l1 l2. induction l1 as [|a l1' h1]. simpl. rewrite comm_prod. rewrite one_prod. reflexivity. (* a::l1'*) simpl. rewrite h1. apply assoc_prod. Qed. (* maybe turn Bt into a general Boolean algebraic type *) Lemma dist_list_sing_plus : forall (l:list Bt) (x:Bt), x * (plus_list l) = plus_list (map (fun y:Bt => (x*y)) l). intros l x. induction l as [|a l h1]. (* nil *) simpl. apply zero_prod. simpl. rewrite dist_sum. rewrite h1. reflexivity. Qed. Lemma dist_list_sing_times : forall (l:list Bt) (x:Bt), x + (times_list l) = times_list (map (fun y:Bt => (x+y)) l). intros l x. induction l as [|a l h1]. (* nil *) simpl. apply one_sum. simpl. rewrite dist_prod. rewrite h1. reflexivity. Qed. Lemma dist_list_2_plus : forall (l1 l2:list Bt), (plus_list l1) * (plus_list l2) = plus_list (map (fun p:Bt*Bt => (fst p * snd p)) (list_prod l1 l2)). intros l1 l2. induction l1 as [|a l1' h1]; destruct l2 as [|b l2']; simpl. (* nil nil *) apply zero_prod. (* nil b::l2' *) rewrite comm_prod. apply zero_prod. (* a::l1' nil *) simpl in h1. rewrite <- h1. do 2 rewrite zero_prod. reflexivity. (* a::l1' b::l2'*) simpl in h1. rewrite comm_prod at 1. rewrite dist_sum. rewrite comm_prod in h1. rewrite h1. rewrite comm_prod. rewrite dist_sum. rewrite map_app. assert (h3:a*plus_list l2' = plus_list (map (fun p:Bt * Bt => fst p * snd p) (map (fun y:Bt => (a, y)) l2'))). rewrite dist_list_sing_plus. assert (h4:map (fun y:Bt => a * y) l2' = map (fun p : Bt * Bt => fst p * snd p) (map (fun y:Bt => (a, y)) l2')). rewrite map_map. simpl. reflexivity. rewrite h4. reflexivity. rewrite h3. rewrite plus_list_app. rewrite <- assoc_sum. reflexivity. Qed. Lemma dist_list_2_times : forall (l1 l2:list Bt), (times_list l1) + (times_list l2) = times_list (map (fun p:Bt*Bt => (fst p + snd p)) (list_prod l1 l2)). intros l1 l2. induction l1 as [|a l1' h1]; destruct l2 as [|b l2']; simpl. (* nil nil *) apply one_sum. (* nil b::l2' *) rewrite comm_sum. apply one_sum. (* a::l1' nil *) simpl in h1. rewrite <- h1. do 2 rewrite one_sum. reflexivity. (* a::l1' b::l2'*) simpl in h1. rewrite comm_sum at 1. rewrite dist_prod. rewrite comm_sum in h1. rewrite h1. rewrite comm_sum. rewrite dist_prod. rewrite map_app. assert (h3:a+times_list l2' = times_list (map (fun p:Bt * Bt => fst p + snd p) (map (fun y:Bt => (a, y)) l2'))). rewrite dist_list_sing_times. assert (h4:map (fun y:Bt => a + y) l2' = map (fun p : Bt * Bt => fst p + snd p) (map (fun y:Bt => (a, y)) l2')). rewrite map_map. simpl. reflexivity. rewrite h4. reflexivity. rewrite h3. rewrite times_list_app. rewrite <- assoc_prod. reflexivity. Qed. Definition plus_times_list_of_lists (l:faml Bt) : Bt := plus_list (map (fun l':(list Bt) => times_list l') l). Definition times_plus_list_of_lists (l:faml Bt) : Bt := times_list (map (fun l':(list Bt) => plus_list l') l). Definition times_fun {T:Type} (p:T->Bt) (l:list T) := times_list (map p l). Definition plus_fun {T:Type} (p:T->Bt) (l:list T) := plus_list (map p l). Definition times_fun_in_dep {T:Type} {l:list T} (p:fun_in_dep l Bt) := times_list (map_in_dep p). Definition plus_fun_in_dep {T:Type} {l:list T} (p:fun_in_dep l Bt) := plus_list (map_in_dep p). Definition plus_times_fun1 {T U:Type} (li:list T) (lj:list U) (p:T->U->Bt) := plus_fun (fun i:T => times_fun (p i) lj) li. Definition plus_times_fun2 {T U:Type} (li:list T) (lj:list U) (p:T->U->Bt) := plus_fun (fun j:U => times_fun (fun i:T => (p i j)) li) lj. Definition times_plus_fun1 {T U:Type} (li:list T) (lj:list U) (p:T->U->Bt) := times_fun (fun i:T => plus_fun (p i) lj) li. Definition times_plus_fun2 {T U:Type} (li:list T) (lj:list U) (p:T->U->Bt) := times_fun (fun j:U => plus_fun (fun i:T => (p i j)) li) lj. Definition plus_times_all_funs1 {T U:Type} (pfdt:type_eq_dec T) (pfdu: type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Bt): Bt. pose (list_power li lj) as lp. pose (fun (l:list (T*U)) (pfinlp:In l lp) => (fpl_to_fun_in_dep (in_list_power_fpl pfdt pfdu _ _ l pfi (list_power_no_dup _ _ _ pfi pfj pfinlp) pfinlp))) as f. pose (fun (l:list (T*U)) (pfinlp:In l lp) => times_fun_in_dep (fun (i:T) (pfini:In i li) => p i (f l pfinlp i pfini))) as g. refine (plus_fun_in_dep g). Defined. Definition plus_times_all_funs2 {T U:Type} (pfdt:type_eq_dec T) (pfdu: type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Bt) : Bt. pose (list_power lj li) as lp. pose (fun (l:list (U*T)) (pfinlp:In l lp) => (fpl_to_fun_in_dep (in_list_power_fpl pfdu pfdt _ _ l pfj (list_power_no_dup _ _ _ pfj pfi pfinlp) pfinlp))) as f. pose (fun (l:list (U*T)) (pfinlp:In l lp) => times_fun_in_dep (fun (j:U) (pfinj:In j lj) => p (f l pfinlp j pfinj) j)) as g. refine (plus_fun_in_dep g). Defined. Definition times_plus_all_funs1 {T U:Type} (pfdt:type_eq_dec T) (pfdu: type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Bt): Bt. pose (list_power li lj) as lp. pose (fun (l:list (T*U)) (pfinlp:In l lp) => (fpl_to_fun_in_dep (in_list_power_fpl pfdt pfdu _ _ l pfi (list_power_no_dup _ _ _ pfi pfj pfinlp) pfinlp))) as f. pose (fun (l:list (T*U)) (pfinlp:In l lp) => plus_fun_in_dep (fun (i:T) (pfini:In i li) => p i (f l pfinlp i pfini))) as g. refine (times_fun_in_dep g). Defined. Definition times_plus_all_funs2 {T U:Type} (pfdt:type_eq_dec T) (pfdu: type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Bt) : Bt. pose (list_power lj li) as lp. pose (fun (l:list (U*T)) (pfinlp:In l lp) => (fpl_to_fun_in_dep (in_list_power_fpl pfdu pfdt _ _ l pfj (list_power_no_dup _ _ _ pfj pfi pfinlp) pfinlp))) as f. pose (fun (l:list (U*T)) (pfinlp:In l lp) => plus_fun_in_dep (fun (j:U) (pfinj:In j lj) => p (f l pfinlp j pfinj) j)) as g. refine (times_fun_in_dep g). Defined. (* This is the complete distributive rule for finite sets, as quoted in (8.3), but mine is more general since it allows you to take the meet of the joins of lists of arbitrary length (not fixed as in the book).*) Lemma complete_dist_list_times : forall (l:faml Bt), times_plus_list_of_lists l = plus_times_list_of_lists (list_of_lists_seqs l). intro l. induction l as [|al l' h1]. (* nil *) unfold plus_times_list_of_lists. simpl. rewrite zero_sum. reflexivity. (* al::l'*) simpl. unfold times_plus_list_of_lists. simpl. unfold times_plus_list_of_lists in h1. rewrite h1. unfold plus_times_list_of_lists. rewrite dist_list_2_plus. rewrite map_map. simpl. induction al as [|a t h2]. (*nil*) simpl. reflexivity. (*a::t*) simpl. rewrite map_map. rewrite map_app. rewrite plus_list_app. rewrite h2. rewrite map_app. rewrite plus_list_app. do 2 rewrite map_map. simpl. reflexivity. Qed. Lemma complete_dist_list_plus : forall (l:faml Bt), plus_times_list_of_lists l = times_plus_list_of_lists (list_of_lists_seqs l). intro l. induction l as [|al l' h1]. (* nil *) unfold times_plus_list_of_lists. simpl. unfold plus_times_list_of_lists. simpl. rewrite comm_prod. rewrite zero_prod. reflexivity. (* al::l'*) simpl. unfold plus_times_list_of_lists. simpl. unfold plus_times_list_of_lists in h1. rewrite h1. unfold times_plus_list_of_lists. rewrite dist_list_2_times. rewrite map_map. simpl. induction al as [|a t h2]. (*nil*) simpl. reflexivity. (*a::t*) simpl. rewrite map_map. rewrite map_app. rewrite times_list_app. rewrite h2. rewrite map_app. rewrite times_list_app. do 2 rewrite map_map. simpl. reflexivity. Qed. Lemma inf_union : forall (U V:Ensemble Bt) (u v:Bt), inf U u -> inf V v -> (inf (Union U V) (u * v)). intros U V u v h1 h2. red in h1. destruct h1 as [h1l h1r]. red in h2. destruct h2 as [h2l h2r]. red. split. (* lower bound *) red. red in h1l. red in h2l. intros s h3. pose proof (Union_inv _ _ _ _ h3) as h4. destruct h4 as [h4l | h4r]. (* h4l *) assert (h5:le (u * v) u). red. rewrite eq_ord. rewrite comm_prod. rewrite assoc_prod. rewrite idem_prod. reflexivity. specialize (h1l _ h4l). apply (trans_le _ _ _ h5 h1l). assert (h5:le (u * v) v). red. rewrite eq_ord. rewrite <- assoc_prod. rewrite idem_prod. reflexivity. specialize (h2l _ h4r). apply (trans_le _ _ _ h5 h2l). (* greatest lower bound *) intros b h3. red in h3. assert (h4:lb U b). red. intros s h5. assert (Inn (Union U V) s). auto with sets. apply h3; assumption. assert (h5:lb V b). red. intros s h6. assert (Inn (Union U V) s). auto with sets. apply h3; assumption. specialize (h1r b h4). specialize (h2r b h5). pose proof (mono_prod _ _ _ _ h1r h2r) as h6. rewrite idem_prod in h6. assumption. Qed. Lemma sup_union : forall (U V:Ensemble Bt) (u v:Bt), sup U u -> sup V v -> (sup (Union U V) (u + v)). intros U V u v h1 h2. red in h1. destruct h1 as [h1l h1r]. red in h2. destruct h2 as [h2l h2r]. red. split. (* upper bound *) red. red in h1l. red in h2l. intros s h3. pose proof (Union_inv _ _ _ _ h3) as h4. destruct h4 as [h4l | h4r]. (* h4l *) assert (h5:le u (u + v)). red. rewrite assoc_sum. rewrite idem_sum. reflexivity. specialize (h1l _ h4l). apply (trans_le _ _ _ h1l h5). assert (h5:le v (u + v)). red. rewrite comm_sum. rewrite <- assoc_sum. rewrite idem_sum. reflexivity. specialize (h2l _ h4r). apply (trans_le _ _ _ h2l h5). (* least upper bound *) intros b h3. red in h3. assert (h4:ub U b). red. intros s h5. assert (Inn (Union U V) s). auto with sets. apply h3; assumption. assert (h5:ub V b). red. intros s h6. assert (Inn (Union U V) s). auto with sets. apply h3; assumption. specialize (h1r b h4). specialize (h2r b h5). pose proof (mono_sum _ _ _ _ h1r h2r) as h6. rewrite idem_sum in h6. assumption. Qed. Lemma decompose_inf : forall (S X:Ensemble Bt) (s x x':Bt), Included X S -> inf S s -> inf (Setminus S X) x' -> inf X x -> inf (Couple x x') s. intros S X s x x' h1 h2 h3 h4. red. red in h2. destruct h2 as [h2l h2r]. red in h3. destruct h3 as [h3l h3r]. red in h4. destruct h4 as [h4l h4r]. split. (* lower bound *) red. intros b h5. pose proof (Couple_inv _ _ _ _ h5) as h6. destruct h6 as [h6l | h6r]. (* b = x *) subst. red in h2l. assert (h6: forall y:Bt, Inn X y -> Inn S y). auto with sets. assert (h7: forall y:Bt, Inn X y -> le s y). intros y h8. specialize (h6 _ h8). apply h2l; assumption. assert (h8: lb X s). red. assumption. apply h4r; assumption. (* b = x'*) subst. assert (h6:Included (Setminus S X) S). red. intros b h7. unfold Setminus in h7. red in h7. tauto. red in h3l. assert (h7: forall y:Bt, Inn (Setminus S X) y -> Inn S y). auto with sets. assert (h8: forall y:Bt, Inn (Setminus S X) y -> le s y). intros y h9. specialize (h7 _ h9). apply h2l; assumption. assert (h9: lb (Setminus S X) s). red. assumption. apply h3r; assumption. (* greatest lower bound *) intros a h5. assert (h6: lb S a). red in h5. pose proof (Couple_l _ x x') as h6. pose proof (Couple_r _ x x') as h7. pose proof (h5 _ h6) as h8. pose proof (h5 _ h7) as h9. red in h3l. red in h4l. red. intros c h10. case (classic (Inn X c)) as [h11 | h12]. (* In X c *) pose proof (h4l _ h11) as h13. apply trans_le with x; assumption. (* ~ In X c *) assert (h13: Inn (Setminus S X) c). auto with sets. pose proof (h3l _ h13) as h14. apply trans_le with x'; assumption. apply h2r; assumption. Qed. Lemma decompose_inf': forall (S X:Ensemble Bt) (x x':Bt), Included X S -> inf X x -> inf (Setminus S X) x' -> inf S (x * x'). intros S X x x' h1 h2 h3. pose proof decompose_inf. case (classic (exists s, inf S s)) as [h4 | h5]. (* h4 *) destruct h4 as [s h4]. assert (h5:s = x * x'). pose proof (decompose_inf _ _ _ _ _ h1 h4 h3 h2) as h5. pose proof (lat_prod x x') as h6. pose proof (inf_unq _ _ _ h5 h6). assumption. rewrite h5 in h4. assumption. (* h5 *) contradict h5. exists (x * x'). assert (h6:S = Union (Setminus S X) X). apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h7. case (classic (Inn X y)) as [h8 | h9]. (* h8 *) apply Union_intror; assumption. (* h9 *) apply Union_introl. unfold Setminus. red. split; assumption. (* >= *) red. intros y h7. pose proof (Union_inv _ _ _ _ h7) as h8. destruct h8 as [h8l | h8r]. unfold Setminus in h8l. red in h8l. tauto. auto with sets. rewrite h6. pose proof (inf_union _ _ _ _ h3 h2) as h7. rewrite comm_prod in h7. assumption. Qed. Lemma inf_times_cons : forall (E:Ensemble Bt) (a b:Bt), Inn E a -> inf (Subtract E a) b -> inf E (a * b). intros E a b h1 h2. pose proof decompose_inf'. pose proof (inf_singleton _ a) as h3. unfold Subtract in h2. assert (h4:Included (Singleton a) E). red. intros y h5. destruct h5; assumption. apply (decompose_inf' _ _ _ _ h4 h3 h2). Qed. Lemma inf_times_finite : forall (l:list Bt), inf (list_to_set l) (times_list l). intros l. induction l. (* nil *) simpl. red. split. red. intros. contradiction. intros; apply one_max. (* a::l *) simpl. unfold Add. rewrite Union_commutative. apply inf_union. apply inf_singleton. assumption. Qed. Lemma sup_plus_finite : forall (l:list Bt), sup (list_to_set l) (plus_list l). intros l. induction l. (* nil *) simpl. red. split. red. intros. contradiction. intros; apply zero_min. (* a::l *) simpl. unfold Add. rewrite Union_commutative. apply sup_union. apply sup_singleton. assumption. Qed. Lemma neg_times_list : forall (l:list Bt), - times_list l = plus_list (map (@Bcomp (Bc B)) l). intro l. induction l as [|a l ih]; simpl. rewrite one_comp; auto. rewrite de_mor_prod. f_equal; auto. Qed. Lemma neg_plus_list : forall (l:list Bt), - plus_list l = times_list (map (@Bcomp (Bc B)) l). intro l. induction l as [|a l ih]; simpl. rewrite zero_comp; auto. rewrite de_mor_sum. f_equal; auto. Qed. Lemma le_times_coarse_list : forall (l l':list Bt), coarse_list l l' -> le (times_list l') (times_list l). intros l l' h1. induction h1; simpl. apply one_max. eapply trans_le. rewrite comm_prod. apply times_le; auto. assumption. apply mono_prod. apply refl_le. assumption. Qed. Lemma times_coarse_list_zero : forall (l l':list Bt), coarse_list l l' -> times_list l = 0 -> times_list l' = 0. intros l l' h1 h2. apply le_times_coarse_list in h1. rewrite h2 in h1. apply le_x_0 in h1; auto. Qed. Definition zero_coarse_lists (l:list Bt) : Ensemble (list Bt) := [l':list Bt | l' <> l /\ coarse_list l' l /\ times_list l' = 0]. Lemma zero_coarse_lists_nil : zero_coarse_lists (@nil Bt) = Empty_set _. apply not_inhabited_empty. intro h1. inversion h1 as [l h2]. destruct h2 as [h2]. destruct h2 as [h2 [h3 h4]]. rewrite coarse_list_nil2_iff in h3. contradiction. Qed. Lemma incl_union_zero_coarse_lists_cons : forall (l:list Bt) (p:Bt), Included (Union (zero_coarse_lists l) (Im (zero_coarse_lists l) (fun l0 : list Bt => p :: l0))) (zero_coarse_lists (p :: l)). intros l p; red; intros lx h1. destruct h1 as [lx h1 | lx h1]. destruct h1 as [h1]. destruct h1 as [h1 [h2 h3]]. constructor. split. intro; subst. apply length_coarse_list in h2. simpl in h2; omega. split. eapply coarse_list_trans. red; intro; apply ba_eq_dec. apply h2. apply coarse_list_cons1. apply coarse_list_refl. assumption. destruct h1 as [lx h1]; subst. destruct h1 as [h1]. destruct h1 as [h1 [h2 h3]]. constructor. split. intro h4. inversion h4. contradiction. split. apply coarse_list_cons2; auto. simpl. rewrite h3. apply zero_prod. Qed. Lemma zero_coarse_lists_cons : forall (l:list Bt) (p:Bt), let Z := zero_coarse_lists (p::l) in let S := Union (zero_coarse_lists l) (Im (zero_coarse_lists l) (fun l => p::l)) in Z = Union S (Setminus Z S). intros; simpl; apply decompose_setminus_inc, incl_union_zero_coarse_lists_cons. Qed. Lemma incl_zero_coarse_lists : forall (l:list Bt), Included (zero_coarse_lists l) (coarse_lists l). intro l; red; intros lx h1. destruct h1 as [h1]. destruct h1 as [h1 [h2 h3]]. constructor; auto. Qed. Lemma finite_zero_coarse_lists : forall (l:list Bt), Finite (zero_coarse_lists l). intro l. eapply Finite_downward_closed. apply finite_coarse_lists. apply incl_zero_coarse_lists. Qed. Lemma incl_union_zero_coarse_lists : forall (l:list Bt), Included (FamilyUnion (ens_list_to_fam (zero_coarse_lists l))) (list_to_set l). intros l x h1. destruct h1 as [S x h1 h2]. destruct h1 as [lx h1]; subst. destruct h1 as [h1]. rewrite <- list_to_set_in_iff in h2. destruct h1 as [h1 [h3 h4]]. rewrite <- list_to_set_in_iff. eapply coarse_list_incl. apply h3. assumption. Qed. Lemma times_list_times_in : forall (l:list Bt) (x:Bt), In x l -> x * times_list l = times_list l. intro l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h1. destruct h1 as [|h1]; subst. rewrite assoc_prod, idem_prod; auto. rewrite assoc_prod, (comm_prod _ x a), <- assoc_prod. f_equal; auto. Qed. Lemma plus_list_plus_in : forall (l:list Bt) (x:Bt), In x l -> x + plus_list l = plus_list l. intro l. induction l as [|a l ih]; simpl. intros; contradiction. intros x h1. destruct h1 as [|h1]; subst. rewrite assoc_sum, idem_sum; auto. rewrite assoc_sum, (comm_sum _ x a), <- assoc_sum. f_equal; auto. Qed. Lemma prod_list_dup_eq : forall (l:list Bt) (x:Bt), In x l -> times_list l = times_list (x::l). intros l. induction l as [|a l h1]. intros x h2; try contradiction. intros x h2. destruct h2 as [h2l | h2r]. subst. simpl. rewrite assoc_prod. rewrite idem_prod. reflexivity. simpl. rewrite assoc_prod. rewrite (comm_prod _ x a). rewrite <- assoc_prod. simpl in h1. rewrite <- (h1 x h2r). auto. Qed. Lemma sum_list_dup_eq : forall (l:list Bt) (x:Bt), In x l -> plus_list l = plus_list (x::l). intro l. induction l as [|a l h1]. intros x h2; try contradiction. intros x h2. destruct h2 as [h2l | h2r]. subst. simpl. rewrite assoc_sum. rewrite idem_sum. reflexivity. simpl. rewrite assoc_sum. rewrite (comm_sum _ x a). rewrite <- assoc_sum. simpl in h1. rewrite <- (h1 x h2r). auto. Qed. Lemma prod_preserves_list_singularize : forall (l:list Bt), times_list l = times_list (list_singularize ba_eq_dec l). intro l. induction l as [|a l h1]. simpl; auto. simpl. rewrite h1. destruct (in_dec ba_eq_dec a l) as [h2 | h3]. rewrite <- h1. pose proof (prod_list_dup_eq _ _ h2) as h3. simpl in h3. rewrite <- h3. auto. rewrite list_singularize_in_cons; auto. rewrite list_singularize_nin; auto. Qed. Lemma sum_preserves_list_singularize : forall (l:list Bt), plus_list l = plus_list (list_singularize ba_eq_dec l). intro l. induction l as [|a l h1]. simpl; auto. simpl. rewrite h1. destruct (in_dec ba_eq_dec a l) as [h2 | h3]. rewrite <- h1. pose proof (sum_list_dup_eq _ _ h2) as h3. simpl in h3. rewrite <- h3. auto. rewrite list_singularize_in_cons; auto. rewrite list_singularize_nin; auto. Qed. Lemma times_sing_preserves_new_head : forall (l:list Bt) (x:Bt), In x l -> times_list (list_singularize ba_eq_dec l) = times_list (new_head ba_eq_dec (list_singularize ba_eq_dec l) x). intro l. induction l as [|a l h1]. intros; contradiction. intros x h2. simpl. simpl in h1. destruct (in_dec ba_eq_dec a l) as [h3 | h4]. destruct h2 as [h2a | h2b]. subst. rewrite list_singularize_in_cons; auto. rewrite list_singularize_in_cons; auto. destruct h2 as [|h2]; subst; auto. rewrite list_singularize_nin; auto; simpl. f_equal. lr_if_goal'. rewrite <- list_singularize_new_head_not_in_compat; auto. contradict hr; auto. rewrite list_singularize_nin; auto; simpl. lr_if_goal'; subst. contradiction. simpl. rewrite assoc_prod. rewrite (comm_prod _ x a). rewrite <- assoc_prod. f_equal; auto. Qed. Lemma plus_sing_preserves_new_head : forall (l:list Bt) (x:Bt), In x l -> plus_list (list_singularize ba_eq_dec l) = plus_list (new_head ba_eq_dec (list_singularize ba_eq_dec l) x). intro l. induction l as [|a l h1]. intros; contradiction. intros x h2. simpl; simpl in h1. destruct (in_dec ba_eq_dec a l) as [h3 | h4]. destruct h2 as [h2a | h2b]. subst. simpl. rewrite list_singularize_in_cons; auto. rewrite list_singularize_in_cons; auto. destruct h2 as [|h2]; subst; auto. rewrite list_singularize_nin; auto; simpl. lr_if_goal'. f_equal. rewrite <- list_singularize_new_head_not_in_compat; auto. contradict hr; auto. rewrite list_singularize_nin; auto. destruct (ba_eq_dec a x) as [h6 | h7]. subst. contradiction. simpl. lr_if_goal'. subst. contradict h7; auto. simpl. rewrite assoc_sum. rewrite (comm_sum _ x a). rewrite <- assoc_sum. f_equal; auto. Qed. Lemma times_list_sing_cons : forall (l:list Bt) (x:Bt), In x (list_singularize ba_eq_dec l) -> times_list (list_singularize ba_eq_dec l) = x * times_list (list_singularize ba_eq_dec (remove ba_eq_dec x l)). intros l. induction l as [|a l h1]. intros x h1. simpl in h1. contradiction. intros x h2. simpl in h2. unfold Bt in h1. unfold Bt. destruct (in_dec ba_eq_dec a l) as [h3 | h4]; destruct (ba_eq_dec a x) as [h5 | h6]; subst. rewrite list_singularize_in_cons; auto. rewrite remove_cons_eq; auto. apply h1; auto. rewrite <- list_singularize_in_iff; auto. rewrite list_singularize_in_cons; auto. simpl. lr_if_goal'; subst. contradict h6; auto. rewrite list_singularize_in_cons; auto. apply h1; auto. rewrite list_singularize_in_cons in h2; auto. rewrite list_singularize_in_cons in h2; auto. rewrite in_remove_iff. split; auto. rewrite list_singularize_nin in h2; auto. rewrite remove_cons_eq. rewrite list_singularize_nin; auto. simpl. f_equal. rewrite nin_remove_eq; auto. rewrite list_singularize_nin; auto. rewrite list_singularize_nin in h2; auto. destruct h2 as [|h2]; subst. contradict h6; auto. simpl. lr_if_goal'; subst. contradict h6; auto. rewrite list_singularize_nin; auto. simpl. rewrite assoc_prod, (comm_prod _ x a), <- assoc_prod. f_equal; auto. contradict h4. rewrite in_remove_iff in h4. destruct h4; auto. Qed. Lemma plus_list_sing_cons : forall (l:list Bt) (x:Bt), In x (list_singularize ba_eq_dec l) -> plus_list (list_singularize ba_eq_dec l) = x + plus_list (list_singularize ba_eq_dec (remove ba_eq_dec x l)). intros l. induction l as [|a l h1]. intros x h1. simpl in h1. contradiction. intros x h2. simpl in h2. unfold Bt in h1; unfold Bt; simpl. destruct (in_dec ba_eq_dec a l) as [h3 | h4]. lr_if_goal'; subst. rewrite list_singularize_in_cons; auto. apply h1; auto. rewrite <- list_singularize_in_iff; auto. rewrite list_singularize_in_cons; auto. simpl. rewrite list_singularize_in_cons; auto. apply h1; auto. rewrite list_singularize_in_cons in h2; auto. rewrite in_remove_iff. split; auto. rewrite list_singularize_nin in h2; auto. destruct h2 as [|h2]; subst. lr_if_goal'; subst. rewrite list_singularize_nin; simpl; f_equal. rewrite nin_remove_eq; auto. auto. contradict hr; auto. specialize (h1 _ h2). rewrite list_singularize_nin; auto; simpl. rewrite h1. lr_if_goal'; subst. rewrite <- list_singularize_in_iff in h2; contradiction. rewrite list_singularize_nin; simpl. rewrite assoc_sum, <- (comm_sum _ x a),<- assoc_sum; auto. contradict h4. rewrite in_remove_iff in h4. destruct h4; auto. Qed. End ListOperations. Arguments plus_list_app [B] _ _. Arguments times_list_app [B] _ _. Arguments dist_list_sing_plus [B] _ _. Arguments dist_list_sing_times [B] _ _. Arguments dist_list_2_plus [B] _ _. Arguments dist_list_2_times [B] _ _. Arguments plus_times_list_of_lists [B] _. Arguments times_plus_list_of_lists [B] _. Arguments times_fun [B] [T] _ _. Arguments plus_fun [B] [T] _ _. Arguments plus_times_fun1 [B] [T] [U] _ _ _. Arguments plus_times_fun2 [B] [T] [U] _ _ _. Arguments times_plus_fun1 [B] [T] [U] _ _ _. Arguments times_plus_fun2 [B] [T] [U] _ _ _. Arguments plus_times_all_funs1 [B] [T] [U] _ _ _ _ _ _ _. Arguments plus_times_all_funs2 [B] [T] [U] _ _ _ _ _ _ _. Arguments times_plus_all_funs1 [B] [T] [U] _ _ _ _ _ _ _. Arguments times_plus_all_funs2 [B] [T] [U] _ _ _ _ _ _ _. Arguments complete_dist_list_times [B] _. Arguments complete_dist_list_plus [B] _. Arguments inf_union [B] _ _ _ _ _ _. Arguments sup_union [B] _ _ _ _ _ _. Arguments decompose_inf [B] _ _ _ _ _ _ _ _ _. Arguments decompose_inf' [B] _ _ _ _ _ _ _. Arguments inf_times_cons [B] _ _ _ _ _. Arguments inf_times_finite [B] _. Arguments sup_plus_finite [B] _. Arguments plus_fun_in_dep [B] [T] [l] _. Arguments times_fun_in_dep [B] [T] [l] _. Section OpProp. Variable B:Bool_Alg. Let Bt := bt B. Lemma le_times_finite_member : forall (l:list Bt) (x:Bt), (In x l) -> le (times_list l) x. intros l x h1. induction l as [|a l h2]. (* nil *) simpl in h1. contradiction. (* a::l *) apply in_inv in h1. destruct h1 as [h1l | h1r]. (* a = x *) subst. simpl. apply times_le. (* In x l *) specialize (h2 h1r). simpl. pose proof (times_le (times_list l) a) as h3. rewrite comm_prod in h3. apply trans_le with (times_list l); assumption. Qed. Lemma le_member_plus_finite : forall (l:list Bt) (x:Bt), (In x l) -> le x (plus_list l). intros l x h1. induction l as [|a l h2]. (* nil *) simpl in h1. contradiction. (* a:: l*) apply in_inv in h1. destruct h1 as [h1l | h1r]. (* a = x *) subst. simpl. apply le_plus. (* In x l *) specialize (h2 h1r). simpl. pose proof (le_plus (plus_list l) a) as h3. rewrite comm_sum in h3. apply trans_le with (plus_list l); assumption. Qed. End OpProp. Arguments le_times_finite_member [B] _ _ _. Arguments le_member_plus_finite [B] _ _ _. Section SetOperations. Variable B:Bool_Alg. Let Bt := bt B. Lemma list_to_sets_eq_times_sing_eq : forall (E:Ensemble Bt), forall (l l':list Bt), E = list_to_set (list_singularize ba_eq_dec l) -> E = list_to_set (list_singularize ba_eq_dec l') -> times_list (list_singularize ba_eq_dec l) = times_list (list_singularize ba_eq_dec l'). intros E l l' h1. pose proof (list_to_set_finite (list_singularize ba_eq_dec l)) as h2. rewrite <- h1 in h2. revert h1. revert l l'. induction h2 as [|A h0 h4 x h5]. intros l l' h1 h3. symmetry in h1. apply empty_set_nil in h1. symmetry in h3. apply empty_set_nil in h3. rewrite h1. rewrite h3. reflexivity. intros l l' h6 h7. pose proof (add_list_to_set_to_remove ba_eq_dec _ _ _ h5 h6) as h8. pose proof (add_list_to_set_to_remove ba_eq_dec _ _ _ h5 h7) as h9. rewrite remove_list_singularize_compat in h8. rewrite remove_list_singularize_compat in h9. specialize (h4 _ _ h8 h9). pose proof (Add_intro2 _ A x) as h10. pose proof h10 as h11. rewrite h6 in h10. rewrite h7 in h11. rewrite <- list_to_set_in_iff in h10. rewrite <- list_to_set_in_iff in h11. rewrite (times_list_sing_cons _ _ x h10). rewrite (times_list_sing_cons _ _ x h11). rewrite h4. reflexivity. Qed. Lemma list_to_sets_eq_plus_sing_eq : forall (E:Ensemble Bt), forall (l l':list Bt), E = list_to_set (list_singularize ba_eq_dec l) -> E = list_to_set (list_singularize ba_eq_dec l') -> plus_list (list_singularize ba_eq_dec l) = plus_list (list_singularize ba_eq_dec l'). intros E l l' h1. pose proof (list_to_set_finite (list_singularize ba_eq_dec l)) as h2. rewrite <- h1 in h2. revert h1. revert l l'. induction h2 as [|A h0 h4 x h5]. intros l l' h1 h3. symmetry in h1. apply empty_set_nil in h1. symmetry in h3. apply empty_set_nil in h3. rewrite h1. rewrite h3. reflexivity. intros l l' h6 h7. pose proof (add_list_to_set_to_remove ba_eq_dec _ _ _ h5 h6) as h8. pose proof (add_list_to_set_to_remove ba_eq_dec _ _ _ h5 h7) as h9. rewrite remove_list_singularize_compat in h8. rewrite remove_list_singularize_compat in h9. specialize (h4 _ _ h8 h9). pose proof (Add_intro2 _ A x) as h10. pose proof h10 as h11. rewrite h6 in h10. rewrite h7 in h11. rewrite <- list_to_set_in_iff in h10. rewrite <- list_to_set_in_iff in h11. rewrite (plus_list_sing_cons _ _ x h10). rewrite (plus_list_sing_cons _ _ x h11). rewrite h4. reflexivity. Qed. Lemma list_to_sets_eq_times_eq : forall (E:Ensemble Bt), forall (l l':list Bt), E = list_to_set l -> E = list_to_set l' -> times_list l = times_list l'. intros E l l' h1 h2. pose proof (list_to_set_singularize_compat ba_eq_dec l) as h3. unfold Bt in h3, h1. rewrite h3 in h1 at 1. pose proof (list_to_set_singularize_compat ba_eq_dec l') as h4. unfold Bt in h4, h2. rewrite h4 in h2 at 1. rewrite prod_preserves_list_singularize. rewrite (prod_preserves_list_singularize _ l'). apply list_to_sets_eq_times_sing_eq with E; assumption. Qed. Lemma list_to_sets_eq_plus_eq : forall (E:Ensemble Bt), forall (l l':list Bt), E = list_to_set l -> E = list_to_set l' -> plus_list l = plus_list l'. intros E l l' h1 h2. pose proof (list_to_set_singularize_compat ba_eq_dec l) as h3. unfold Bt in h3, h1. rewrite h3 in h1 at 1. pose proof (list_to_set_singularize_compat ba_eq_dec l') as h4. unfold Bt in h4, h2. rewrite h4 in h2 at 1. rewrite sum_preserves_list_singularize. rewrite (sum_preserves_list_singularize _ l'). apply list_to_sets_eq_plus_sing_eq with E; assumption. Qed. Lemma times_list_functional : forall (l l':list Bt), list_to_set l = list_to_set l' -> times_list l = times_list l'. intros l l' h1. apply (list_to_sets_eq_times_eq _ l l' (eq_refl _) h1). Qed. Lemma plus_list_functional : forall (l l':list Bt), list_to_set l = list_to_set l' -> plus_list l = plus_list l'. intros l l' h1. apply (list_to_sets_eq_plus_eq _ l l' (eq_refl _) h1). Qed. Lemma times_list_unq : forall (E:Ensemble Bt) (pf:Finite E), exists! x, forall (l:list Bt), E = list_to_set l -> x = times_list l. intros E h1. pose proof (finite_set_list _ h1) as h2. destruct h2 as [l h3]. exists (times_list l). red. split. intros l' h4. apply list_to_sets_eq_times_eq with E; assumption. intros x h4. rewrite (h4 _ h3). reflexivity. Qed. Lemma plus_list_unq : forall (E:Ensemble Bt) (pf:Finite E), exists! x, forall (l:list Bt), E = list_to_set l -> x = plus_list l. intros E h1. pose proof (finite_set_list _ h1) as h2. destruct h2 as [l h3]. exists (plus_list l). red. split. intros l' h4. apply list_to_sets_eq_plus_eq with E; assumption. intros x h4. rewrite (h4 _ h3). reflexivity. Qed. Lemma times_times_list_remove : forall (l:list Bt), NoDup l -> forall (a:Bt), In a l -> times_list l = a * times_list (remove eq_dec a l). intro l. induction l as [|b l h1]; auto; intros h2 x h3. simpl in h3. contradiction. simpl. pose proof (no_dup_cons_nin _ _ h2) as h4. apply no_dup_cons in h2. destruct h3 as [h3a | h3b]. subst. f_equal. destruct (eq_dec x x). rewrite <- remove_not_in'; auto. contradict n. reflexivity. destruct (eq_dec x b) as [h5 | h6]. subst. contradiction. simpl. rewrite assoc_prod. rewrite (comm_prod _ x b). specialize (h1 h2 _ h3b). rewrite h1 at 1. rewrite assoc_prod. reflexivity. Qed. Lemma plus_plus_list_remove : forall (l:list Bt), NoDup l -> forall (a:Bt), In a l -> plus_list l = a + plus_list (remove eq_dec a l). intro l. induction l as [|b l h1]; auto; intros h2 x h3. simpl in h3. contradiction. simpl. pose proof (no_dup_cons_nin _ _ h2) as h4. apply no_dup_cons in h2. destruct h3 as [h3a | h3b]. subst. f_equal. destruct (eq_dec x x). rewrite <- remove_not_in'; auto. contradict n. reflexivity. destruct (eq_dec x b) as [h5 | h6]. subst. contradiction. simpl. rewrite assoc_sum. rewrite (comm_sum _ x b). specialize (h1 h2 _ h3b). rewrite h1 at 1. rewrite assoc_sum. reflexivity. Qed. (*This is the fruit of all the "singularize" code above, a function that returns the product (sum) of a finite set, without the user having to deal with lists at all.*) Definition times_set (E:Ensemble Bt) (pf:Finite E) : Bt. refine (proj1_sig (constructive_definite_description _ (times_list_unq E pf))). Defined. Definition plus_set (E:Ensemble Bt) (pf:Finite E) : Bt. refine (proj1_sig (constructive_definite_description _ (plus_list_unq E pf))). Defined. Lemma times_set_compat : forall (E:Ensemble Bt) (pf:Finite E), exists l:list Bt, times_set E pf = times_list l. intros E h1. unfold times_set. destruct constructive_definite_description as [x h2]. simpl. pose proof (finite_set_list _ h1) as h3. destruct h3 as [l h4]. exists l. apply h2. assumption. Qed. Lemma times_set_compat' : forall (E:Ensemble Bt) (l:list Bt) (pf:Finite E), E = list_to_set l -> times_set E pf = times_list l. intros E l pf h1. unfold times_set. destruct constructive_definite_description as [x h2]. simpl. apply h2; assumption. Qed. Lemma times_list_compat : forall (l:list Bt), times_list l = times_set (list_to_set l) (list_to_set_finite l). intro; erewrite times_set_compat'; auto. Qed. Lemma times_set_list_to_set : forall (l:list Bt) (pf:Finite (list_to_set l)), times_set (list_to_set l) pf = times_list l. intros; apply times_set_compat'; auto. Qed. Lemma plus_set_compat : forall (E:Ensemble Bt) (pf:Finite E), exists l:list Bt, plus_set E pf = plus_list l. intros E h1. unfold plus_set. destruct constructive_definite_description as [x h2]. simpl. pose proof (finite_set_list _ h1) as h3. destruct h3 as [l h4]. exists l. apply h2. assumption. Qed. Lemma plus_set_compat' : forall (E:Ensemble Bt) (l:list Bt) (pf:Finite E), E = list_to_set l -> plus_set E pf = plus_list l. intros E l pf h1. unfold plus_set. destruct constructive_definite_description as [x h2]. simpl. apply h2; assumption. Qed. Lemma plus_list_compat : forall (l:list Bt), plus_list l = plus_set (list_to_set l) (list_to_set_finite l). intro; erewrite plus_set_compat'; auto. Qed. Lemma plus_set_list_to_set : forall (l:list Bt) (pf:Finite (list_to_set l)), plus_set (list_to_set l) pf = plus_list l. intros; apply plus_set_compat'; auto. Qed. Lemma times_set_functional : forall (A C:Ensemble Bt), A = C -> forall (pfa:Finite A) (pfc:Finite C), times_set A pfa = times_set C pfc. intros A C h3 h1 h2. pose proof (subsetT_eq_compat _ _ _ _ h1 h2 h3) as h4. dependent rewrite -> h4. reflexivity. Qed. Lemma plus_set_functional : forall (A C:Ensemble Bt), A = C -> forall (pfa:Finite A) (pfc:Finite C), plus_set A pfa = plus_set C pfc. intros A C h3 h1 h2. pose proof (subsetT_eq_compat _ _ _ _ h1 h2 h3) as h4. dependent rewrite -> h4. reflexivity. Qed. Lemma times_set_empty : times_set (Empty_set Bt) (Empty_is_finite Bt) = 1. unfold times_set. destruct constructive_definite_description as [x h1]. simpl. specialize (h1 nil). simpl in h1. apply h1. reflexivity. Qed. Lemma times_set_empty' : forall (pf:Finite (Empty_set Bt)), times_set (Empty_set Bt) pf = 1. intro h1. pose (Empty_is_finite Bt) as h2. pose proof (proof_irrelevance _ h1 h2); subst. apply times_set_empty. Qed. Lemma plus_set_empty : plus_set (Empty_set Bt) (Empty_is_finite Bt) = 0. unfold plus_set. destruct constructive_definite_description as [x h1]. simpl. specialize (h1 nil). simpl in h1. apply h1. reflexivity. Qed. Lemma plus_set_empty' : forall (pf:Finite (Empty_set Bt)), plus_set (Empty_set Bt) pf = 0. intro h1. pose (Empty_is_finite Bt) as h2. pose proof (proof_irrelevance _ h1 h2); subst. apply plus_set_empty. Qed. Lemma times_set_add : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt), times_set (Add E x) (Add_preserves_Finite _ _ _ pf) = x * (times_set E pf). intros E pf x. unfold times_set. destruct constructive_definite_description as [y h1]. simpl. destruct constructive_definite_description as [z h2]. simpl. pose proof (finite_set_list _ pf) as h3. destruct h3 as [l h3]. pose proof (h2 l h3) as h5. pose proof (finite_set_list _ (Add_preserves_Finite _ _ x pf)) as h4. destruct h4 as [l' h4]. pose proof (h1 l' h4) as h6. destruct (in_dec eq_dec x l) as [h9 | h10]. (* h9 *) pose proof h9 as h9'. rewrite (list_to_set_in_iff l x) in h9'. rewrite <- h3 in h9'. pose proof (in_add_eq E x h9') as h10. rewrite h10 in h4. specialize (h2 _ h4). pose proof (prod_list_dup_eq _ _ _ h9) as h11. simpl in h11. subst. rewrite h5. assumption. (* h10 *) pose proof h10 as h10'. rewrite (list_to_set_in_iff l x) in h10'. rewrite <- h3 in h10'. assert (h7:Subtract (Add E x) x = Subtract (list_to_set l') x). f_equal. auto. pose proof (subtract_new_head_compat ba_eq_dec l' x) as h8. unfold Bt in h7, h8. rewrite h8 in h7. pose proof (sub_add_same_nin _ _ h10') as hs. unfold Bt in hs. rewrite hs in h7. pose proof (h2 _ h7) as h11. subst. rewrite h11. pose proof (Add_intro2 _ (list_to_set l) x) as h12. rewrite h4 in h12. rewrite <- list_to_set_in_iff in h12. rewrite prod_preserves_list_singularize. rewrite (prod_preserves_list_singularize _ (remove ba_eq_dec x l')). apply times_list_sing_cons. pose proof (list_to_set_singularize_compat ba_eq_dec l') as h13. rewrite list_to_set_in_iff in h12. unfold Bt in h12, h13. rewrite h13 in h12. rewrite <- list_to_set_in_iff in h12. assumption. Qed. Lemma times_set_add' : forall (E:Ensemble Bt) (pf1:Finite E) (x:Bt) (pf2:Finite (Add E x)), times_set (Add E x) pf2 = x * (times_set E pf1). intros E h1 x h2. pose (Add_preserves_Finite _ _ x h1) as h3. pose proof (proof_irrelevance _ h2 h3); subst. unfold h3. rewrite times_set_add. reflexivity. Qed. Lemma plus_set_add : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt), plus_set (Add E x) (Add_preserves_Finite _ _ _ pf) = x + (plus_set E pf). intros E pf x. unfold plus_set. destruct constructive_definite_description as [y h1]. simpl. destruct constructive_definite_description as [z h2]. simpl. pose proof (finite_set_list _ pf) as h3. destruct h3 as [l h3]. pose proof (h2 l h3) as h5. pose proof (finite_set_list _ (Add_preserves_Finite _ _ x pf)) as h4. destruct h4 as [l' h4]. pose proof (h1 l' h4) as h6. destruct (in_dec eq_dec x l) as [h9 | h10]. (* h9 *) pose proof h9 as h9'. rewrite (list_to_set_in_iff l x) in h9'. rewrite <- h3 in h9'. pose proof (in_add_eq E x h9') as h10. rewrite h10 in h4. specialize (h2 _ h4). pose proof (sum_list_dup_eq _ _ _ h9) as h11. simpl in h11. subst. rewrite h5. assumption. (* h10 *) pose proof h10 as h10'. rewrite (list_to_set_in_iff l x) in h10'. rewrite <- h3 in h10'. assert (h7:Subtract (Add E x) x = Subtract (list_to_set l') x). f_equal. auto. pose proof (subtract_new_head_compat ba_eq_dec l' x) as h8. unfold Bt in h7, h8. rewrite h8 in h7. pose proof (sub_add_same_nin _ _ h10') as hs. unfold Bt in hs. rewrite hs in h7. pose proof (h2 _ h7) as h11. subst. rewrite h11. pose proof (Add_intro2 _ (list_to_set l) x) as h12. rewrite h4 in h12. rewrite <- list_to_set_in_iff in h12. rewrite sum_preserves_list_singularize. rewrite (sum_preserves_list_singularize _ (remove ba_eq_dec x l')). apply plus_list_sing_cons. pose proof (list_to_set_singularize_compat ba_eq_dec l') as h13. rewrite list_to_set_in_iff in h12. unfold Bt in h12, h13. rewrite h13 in h12. rewrite <- list_to_set_in_iff in h12. assumption. Qed. Lemma plus_set_add' : forall (E:Ensemble Bt) (pf1:Finite E) (x:Bt) (pf2:Finite (Add E x)), plus_set (Add E x) pf2 = x + (plus_set E pf1). intros E h1 x h2. pose (Add_preserves_Finite _ _ x h1) as h3. pose proof (proof_irrelevance _ h2 h3) as h4. rewrite h4. unfold h3. rewrite plus_set_add. reflexivity. Qed. Lemma times_set_sing : forall x:Bt, times_set (Singleton x) (Singleton_is_finite _ x) = x. intro x. pose proof (Empty_is_finite Bt) as h1. pose proof (Add_preserves_Finite _ _ x h1) as h2. pose proof (add_empty_sing x) as h3. symmetry in h3. pose proof (subsetT_eq_compat _ _ _ _ (Singleton_is_finite _ x) h2 h3) as h4. dependent rewrite -> h4. rewrite (times_set_add' _ h1 x h2). rewrite times_set_empty'. apply one_prod. Qed. Lemma times_set_sing' : forall (x:Bt) (pf:Finite (Singleton x)), times_set _ pf = x. intros x h1. pose proof (times_set_sing x) as h2. pose proof (proof_irrelevance _ h1 (Singleton_is_finite Bt x)) as h3. rewrite <- h3 in h2. assumption. Qed. Lemma times_list_sing : forall (x:Bt), times_list (x::nil) = x. intros; simpl. rewrite one_prod; auto. Qed. Lemma times_list_all_sing : forall (l:list Bt) (x:Bt), l <> nil -> all_sing l x -> times_list l = x. intros l x h0 h1. rewrite prod_preserves_list_singularize. erewrite all_sing_list_singularize. rewrite times_list_sing. reflexivity. assumption. assumption. Qed. Lemma times_set_one_or : forall (E:Ensemble Bt) (pf:Finite E), (E = Empty_set _) \/ (E = Singleton 1) -> 1 = times_set E pf. intros E h1 h2. destruct h2. subst. rewrite times_set_empty'. reflexivity. subst. rewrite times_set_sing'. reflexivity. Qed. Lemma plus_set_sing : forall x:Bt, plus_set (Singleton x) (Singleton_is_finite _ x) = x. intro x. pose proof (Empty_is_finite Bt) as h1. pose proof (Add_preserves_Finite _ _ x h1) as h2. pose proof (add_empty_sing x) as h3. symmetry in h3. pose proof (subsetT_eq_compat _ _ _ _ (Singleton_is_finite _ x) h2 h3) as h4. dependent rewrite -> h4. rewrite (plus_set_add' _ h1 x h2). rewrite plus_set_empty'. apply zero_sum. Qed. Lemma plus_set_sing' : forall (x:Bt) (pf:Finite (Singleton x)), plus_set _ pf = x. intros x h1. pose proof (plus_set_sing x) as h2. pose proof (proof_irrelevance _ h1 (Singleton_is_finite Bt x)) as h3. rewrite <- h3 in h2. assumption. Qed. Lemma plus_list_sing : forall (x:Bt), plus_list (x::nil) = x. intros; simpl. rewrite zero_sum; auto. Qed. Lemma plus_list_all_sing : forall (l:list Bt) (x:Bt), l <> nil -> all_sing l x -> plus_list l = x. intros l x h0 h1. rewrite sum_preserves_list_singularize. erewrite all_sing_list_singularize. rewrite plus_list_sing. reflexivity. assumption. assumption. Qed. Lemma plus_set_zero_or : forall (E:Ensemble Bt) (pf:Finite E), (E = Empty_set _) \/ (E = Singleton 0) -> 0 = plus_set E pf. intros E h1 h2. destruct h2. subst. rewrite plus_set_empty'. reflexivity. subst. rewrite plus_set_sing'. reflexivity. Qed. Lemma times_set_couple : forall (x y:Bt), times_set (Couple x y) (finite_couple x y) = x * y. intros x y. pose proof (couple_add_sing x y) as h1. pose proof (Add_preserves_Finite _ _ y (Singleton_is_finite _ x)) as h2. pose proof (subsetT_eq_compat _ _ _ _ (finite_couple x y) h2 h1) as h3. dependent rewrite -> h3. rewrite (times_set_add' _ (Singleton_is_finite _ x) _ h2). rewrite times_set_sing. rewrite comm_prod. reflexivity. Qed. Lemma times_set_couple' : forall (x y:Bt) (pf:Finite (Couple x y)), times_set (Couple x y) pf = x * y. intros x y h1. rewrite <- times_set_couple. apply times_set_functional. reflexivity. Qed. Lemma plus_set_couple : forall (x y:Bt), plus_set (Couple x y) (finite_couple x y) = x + y. intros x y. pose proof (couple_add_sing x y) as h1. pose proof (Add_preserves_Finite _ _ y (Singleton_is_finite _ x)) as h2. pose proof (subsetT_eq_compat _ _ _ _ (finite_couple x y) h2 h1) as h3. dependent rewrite -> h3. rewrite (plus_set_add' _ (Singleton_is_finite _ x) _ h2). rewrite plus_set_sing. rewrite comm_sum. reflexivity. Qed. Lemma plus_set_couple' : forall (x y:Bt) (pf:Finite (Couple x y)), plus_set (Couple x y) pf = x + y. intros x y h1. rewrite <- plus_set_couple. apply plus_set_functional. reflexivity. Qed. Lemma le_times_set : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt), (Inn E x) -> le (times_set E pf) x. intros E h1 x h2. pose proof h1 as h3. induction h3 as [|E h4 h5 a]. contradiction. destruct h2 as [x h2a | x h2b]. assert (h6:Included E (Add E a)). auto with sets. pose proof (Finite_downward_closed _ _ h1 _ h6) as h7. rewrite (times_set_add' _ h7 _ h1). pose proof (times_le (times_set E h7) a) as h3. specialize (h5 h7 h2a). apply trans_le with (times_set E h7). rewrite comm_prod. assumption. assumption. destruct h2b. rewrite (times_set_add' _ h4 _ h1). apply times_le. Qed. Definition times_set_def (def:Bt) (E:Ensemble Bt) := match (classic_dec (Finite E)) with | left pf => times_set _ pf | _ => def end. Definition plus_set_def (def:Bt) (E:Ensemble Bt) := match (classic_dec (Finite E)) with | left pf => plus_set _ pf | _ => def end. Lemma le_plus_set : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt), (Inn E x) -> le x (plus_set E pf). intros E h1 x h2. pose proof h1 as h3. induction h3 as [|E h4 h5 a]. contradiction. destruct h2 as [x h2a | x h2b]. assert (h6:Included E (Add E a)). auto with sets. pose proof (Finite_downward_closed _ _ h1 _ h6) as h7. rewrite (plus_set_add' _ h7 _ h1). pose proof (le_plus (plus_set E h7) a) as h3. specialize (h5 h7 h2a). rewrite comm_sum. apply trans_le with (plus_set E h7). assumption. assumption. destruct h2b. rewrite (plus_set_add' _ h4 _ h1). apply le_plus. Qed. Lemma inf_times_set : forall (E:Ensemble Bt) (pf:Finite E), inf E (times_set E pf). intros E h1. pose proof h1 as h2. induction h2 as [|E h3 h4 e h5]. (*Empty*) rewrite times_set_empty'. red. split. red. intros; contradiction. intros; apply one_max. (*Add*) assert (h6:Included E (Add E e)). auto with sets. specialize (h4 h3). rewrite (times_set_add' _ h3). unfold Add. rewrite Union_commutative. apply inf_union. apply inf_singleton. assumption. Qed. Lemma sup_plus_set : forall (E:Ensemble Bt) (pf:Finite E), sup E (plus_set E pf). intros E h1. pose proof h1 as h2. induction h2 as [|E h3 h4 e h5]. (*Empty*) rewrite plus_set_empty'. red. split. red. intros; contradiction. intros; apply zero_min. (*Add*) assert (h6:Included E (Add E e)). auto with sets. specialize (h4 h3). rewrite (plus_set_add' _ h3). unfold Add. rewrite Union_commutative. apply sup_union. apply sup_singleton. assumption. Qed. Lemma plus_set_preserves_le : forall (E:Ensemble Bt) (pf:Finite E) (y:Bt), (forall x:Bt, Inn E x -> le x y) -> le (plus_set E pf) y. intros E h1 y h2. pose proof (finite_set_list_no_dup _ h1) as h3. destruct h3 as [l h3]. destruct h3 as [h3 h4]. subst. revert y h4 h1 h2. induction l as [|a l h1]; simpl; auto. intros y h1 h2 h3. rewrite plus_set_empty' at 1. apply zero_min. intros y h2 h3 h4. pose proof (no_dup_cons _ _ h2) as h5. pose proof (no_dup_cons_nin _ _ h2) as h6. assert (h7:Finite (list_to_set l)). eapply Finite_downward_closed. apply h3. auto with sets. rewrite (plus_set_add' _ h7). assert (h8:forall x : Bt, Inn (list_to_set l) x -> le x y). intros x h9. assert (h10:Inn (Add (list_to_set l) a) x). left; auto. apply h4; auto. specialize (h1 y h5 h7 h8). assert (h9:Inn (Add (list_to_set l) a) a). right; auto. constructor. apply plus_preserves_le; auto. Qed. Lemma times_set_preserves_le : forall (E:Ensemble Bt) (pf:Finite E) (y:Bt), (forall x:Bt, Inn E x -> le y x) -> le y (times_set E pf). intros E h1 y h2. pose proof (finite_set_list_no_dup _ h1) as h3. destruct h3 as [l h3]. destruct h3 as [h3 h4]. subst. revert y h4 h1 h2. induction l as [|a l h1]; simpl; auto. intros y h1 h2 h3. rewrite times_set_empty' at 1. apply one_max. intros y h2 h3 h4. pose proof (no_dup_cons _ _ h2) as h5. pose proof (no_dup_cons_nin _ _ h2) as h6. assert (h7:Finite (list_to_set l)). eapply Finite_downward_closed. apply h3. auto with sets. rewrite (times_set_add' _ h7). assert (h8:forall x : Bt, Inn (list_to_set l) x -> le y x). intros x h9. assert (h10:Inn (Add (list_to_set l) a) x). left; auto. apply h4; auto. specialize (h1 y h5 h7 h8). assert (h9:Inn (Add (list_to_set l) a) a). right; auto. constructor. apply times_preserves_le; auto. Qed. Lemma times_set_times_in : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt), Inn E x -> x * (times_set E pf) = times_set E pf. intros E h1 x h2. pose proof (add_sub_compat_in _ _ h2) as h3. assert (h4:Finite (Add (Subtract E x) x)). rewrite h3; auto. assert (h5:times_set E h1 = times_set _ h4). apply times_set_functional; auto. rewrite h5 at 1. assert (h6:Finite (Subtract E x)). eapply Finite_downward_closed. apply h4. auto with sets. rewrite (times_set_add' _ h6 _ h4). rewrite assoc_prod. rewrite idem_prod. rewrite <- times_set_add. pose proof (proof_irrelevance _ h4 (Add_preserves_Finite Bt (Subtract E x) x h6)) as h7. rewrite <- h7. rewrite h5. reflexivity. Qed. Lemma times_list_0 : forall (l:list Bt), In 0 l -> times_list l = 0. intros l; induction l as [|a l ih]; simpl; auto. intro; contradiction. intro h1. destruct h1 as [|h1]; subst. rewrite comm_prod. apply zero_prod; auto. rewrite ih; auto. rewrite zero_prod; auto. Qed. Lemma times_list_opps : forall (l:list Bt) (x y:bt B), In x l -> In y l -> x = - y -> times_list l = 0. intros l x y h1 h2 h3; subst. revert h1 h2. revert y. induction l as [|a l ih]; simpl. intros; contradiction. intros y h1 h2. destruct h1 as [h1|h1], h2 as [h2|h2]; subst. symmetry in h2. rewrite eq_comp_iff in h2. apply degen_eq; auto. rewrite <- (times_list_times_in _ _ _ h2). rewrite assoc_prod. rewrite (comm_prod _ (-y) y). rewrite comp_prod. rewrite comm_prod. rewrite zero_prod; auto. rewrite <- (times_list_times_in _ _ _ h1). rewrite assoc_prod. rewrite comp_prod. rewrite comm_prod. rewrite zero_prod; auto. rewrite (ih _ h1 h2); auto. rewrite zero_prod; auto. Qed. Lemma times_list_neq_zero_ex : forall (l:list Bt), l <> nil -> times_list l <> 0 -> exists x:Bt, In x l /\ x <> 0. intro l. induction l as [|a l ih]; simpl. intro h1; contradict h1; auto. intros h1 h2. destruct l as [|b l ]; simpl. simpl in h2. rewrite one_prod in h2. exists a. split; try left; auto. specialize (ih (cons_neq_nil _ _)). simpl in ih. simpl in h2. hfirst ih. contradict h2; rewrite h2, zero_prod; auto. specialize (ih hf). destruct ih as [x [h3 h4]]. destruct h3 as [|h3]; subst. exists x. split; auto. exists x; split; auto. Qed. Lemma times_set_in_subtract : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt), let pfs := subtract_preserves_finite E x pf in Inn E x -> times_set E pf = (times_set (Subtract E x) pfs) * x. intros E h1 x h2 h3. pose proof (add_sub_compat_in _ _ h3) as h4. pose proof (times_set_functional _ _ h4 (Add_preserves_Finite _ _ x h2) h1) as h5. rewrite times_set_add in h5. rewrite comm_prod. rewrite h5; auto. Qed. (*kinda weak (assumes injectivity) -- it's usually better to use the one after this.*) Lemma times_list_map_in_dep_remove : forall {T:Type} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l Bt) (x:T) (pf:In x l), inj_dep f -> let f' := fun_in_dep_remove pfdt f x in times_list (map_in_dep f) = (f x pf) * times_list (map_in_dep f'). intros T l h1 f x h2 hi. simpl. do 2 rewrite times_list_compat. assert (h3:Inn (list_to_set (map_in_dep f)) (f x h2)). rewrite <- list_to_set_in_iff. apply in_map_in_dep. rewrite (times_set_in_subtract _ (list_to_set_finite _) _ h3). rewrite comm_prod. f_equal. apply times_set_functional. erewrite list_to_set_map_in_dep_remove_inj. reflexivity. assumption. Qed. Lemma times_list_map_in_dep_remove_in : forall {T:Type} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l Bt) (x:T) (pf:In x l), let f' := fun_in_dep_remove_in pfdt ba_eq_dec f x pf in times_list (map_in_dep f) = (f x pf) * times_list (map_in_dep f'). intros T l h1 f x h2. simpl. do 2 rewrite times_list_compat. assert (h3:Inn (list_to_set (map_in_dep f)) (f x h2)). rewrite <- list_to_set_in_iff. apply in_map_in_dep. rewrite (times_set_in_subtract _ (list_to_set_finite _) _ h3). rewrite comm_prod. f_equal. apply times_set_functional. rewrite list_to_set_map_in_dep_remove_in; auto. Qed. Inductive ba_disjointl (l:list Bt) : Prop := ba_disjointl_intro : forall (x y:bt B), x <> 0 -> y <> 0 -> In x l -> In y l -> x <> y -> x * y = 0 -> ba_disjointl l. Lemma ba_disjointl_impl_zero : forall (l:list Bt), ba_disjointl l -> times_list l = 0. intros l h1. destruct h1. generalize dependent l. intro l. induction l as [|a l ih]; simpl. intro; contradiction. intros h1 h2. destruct h1 as [|h1], h2 as [|h2]; subst. subst. rewrite idem_prod in H4. subst. rewrite zero_prod_l; auto. rewrite <- (times_list_times_in _ _ _ h2); auto. rewrite assoc_prod. rewrite H4. apply zero_prod_l. rewrite <- (times_list_times_in _ _ _ h1); auto. rewrite assoc_prod. rewrite comm_prod in H4. rewrite H4. apply zero_prod_l. rewrite ih, zero_prod; auto. Qed. Lemma not_ba_disjoint_nil : ~ba_disjointl nil. intro h1; inversion h1; contradiction. Qed. Lemma ba_disjointl_cons : forall (l:list Bt) (a:Bt), ba_disjointl l -> ba_disjointl (a :: l). intros l a h1. destruct h1. econstructor. apply H. apply H0. right. assumption. right. assumption. assumption. assumption. Qed. Lemma ba_disjointl_cons_impl : forall (l:list Bt) (a:Bt), ba_disjointl (a :: l) -> ba_disjointl l \/ a * times_list l = 0. intros l a h1. inversion h1. destruct H1 as [|h2], H2 as [|h3]; subst. contradiction. rewrite <- (times_list_times_in _ _ _ h3). right. rewrite assoc_prod, H4, zero_prod_l; auto. rewrite <- (times_list_times_in _ _ _ h2); auto. right. rewrite assoc_prod, (comm_prod _ y x), H4, zero_prod_l; auto. left. econstructor. apply H. apply H0. apply h2. apply h3. assumption. assumption. Qed. Lemma ba_disjointl_0_cons : forall (l:list Bt), ba_disjointl (0::l) -> ba_disjointl l. intros l h1. inversion h1. destruct H1 as [|h2], H2 as [|h3]; subst. contradict H0; auto. contradict H; auto. contradict H0; auto. econstructor. apply H. apply H0. assumption. assumption. assumption. assumption. Qed. Lemma ba_disjointl_comm_cons : forall (l:list Bt) (a b:Bt), ba_disjointl (a::b::l) -> ba_disjointl (b::a::l). intros l a b h1. inversion h1. econstructor. apply H. apply H0. apply in_list2_comm; auto. apply in_list2_comm; auto. assumption. assumption. Qed. Lemma times_list_zero_impl : forall (l:list Bt), times_list l = 0 -> {In 0 l} + {degen B} + {exists (x:Bt) (l':list Bt), In x l /\ ~In x l' /\ x <> 0 /\ times_list l' <> 0 /\ coarse_list l' l /\ x * times_list l' = 0}. intro l. destruct (in_dec ba_eq_dec 0 l) as [hi | hi]. intro; left; auto. revert hi. induction l as [|a l ih]; simpl. intros; left; right; red; auto. intro h1. destruct (in_dec ba_eq_dec 0 l) as [h3 | h3]. intro; left; auto. destruct (ba_eq_dec (times_list l) 0) as [h4 | h4]. apply ih in h4. destruct h4 as [[h4 | h4] | h4]; try contradiction. intro. left; right; auto. intro hz. right. destruct h4 as [a' [l' [h5 [h6 [h7 [h8 [h9 h10]]]]]]]. pose proof (le_times_coarse_list _ _ _ h9) as h11. apply (mono_prod a' a') in h11. rewrite h10 in h11. apply le_x_0 in h11. exists a', l'. split. right; auto. repeat split; auto. apply coarse_list_cons1; auto. apply refl_le. assumption. assert (h5:a <> 0). intro; subst; contradict h1; left; auto. destruct (in_dec ba_eq_dec a l) as [h6 | h6]. intro h2. rewrite (times_list_times_in _ _ _ h6) in h2. contradiction. right. exists a, l. split. left; auto. repeat esplit; auto. apply coarse_list_cons1. apply coarse_list_refl. Qed. Lemma impl_times_list_zero : forall (l:list Bt) (x:Bt) (l':list Bt), In x l -> coarse_list l' l -> x * times_list l' = 0 -> times_list l = 0. intros l x l' h1 h2 h3. apply le_times_coarse_list in h2. apply (mono_prod x x) in h2. rewrite h3 in h2. apply le_x_0 in h2. rewrite <- (times_list_times_in _ _ _ h1); auto. apply refl_le. Qed. Lemma ba_disjointl_dec : forall (l:list Bt), {ba_disjointl l} + {~ ba_disjointl l}. intro l. induction l as [|a l ih]. right; apply not_ba_disjoint_nil. destruct ih as [ih | ih]. left. apply ba_disjointl_cons; auto. destruct (ba_eq_dec a 0) as [h1 | h1]. right. subst. contradict ih. apply ba_disjointl_0_cons; auto. pose proof (l_ex_dec ba_eq_dec (fun c => a <> c /\ c <> 0 /\ a * c = 0) l) as h2. hfirst h2. red. intro x. destruct (eq_dec a x) as [|h3]; subst. right. intro h3. destruct h3 as [h3]; contradict h3; auto. destruct (ba_eq_dec x 0) as [|h4]; subst. right. intro h4. destruct h4 as [? [h4 ?]]. contradict h4; auto. destruct (ba_eq_dec (a*x) 0) as [h5 | h5]. left. repeat split; auto. right. intro h6. destruct h6 as [? [? ?]]. contradiction. specialize (h2 hf). destruct h2 as [h2 | h2]. left. destruct h2 as [x [h2 [h3 [h4 h5]]]]. econstructor. apply h1. apply h4. left; auto. right; auto. assumption. assumption. right. contradict h2. inversion h2. destruct H1 as [|h5], H2 as [|h6]; subst. subst. contradict H3; auto. exists y. repeat split; auto. exists x. repeat split; auto. rewrite comm_prod; auto. contradict ih. econstructor. apply H. apply H0. assumption. assumption. assumption. assumption. Qed. Lemma finite_le_times_set_bounded : forall (x:Bt), let E := [y : Bt | lt y x] in forall (pf:Finite E), Inhabited E -> lt (times_set E pf) x. intros x E h1 h2. destruct h2 as [a h2]. red. split. pose proof (le_times_set _ h1 _ h2) as h3. destruct h2 as [h2]. red in h2. destruct h2 as [h2 h4]. eapply trans_le. apply h3. apply h2. intro h3. pose proof (f_equal (Btimes (Bc B) a) h3) as h4. rewrite times_set_times_in in h4. destruct h2 as [h2]. destruct h2 as [h2a h2b]. red in h2a. rewrite eq_ord in h2a. rewrite h2a in h4. rewrite h3 in h4. clear h3. symmetry in h4. contradiction. assumption. Qed. Lemma times_set_union : forall (C D:Ensemble Bt) (pfc:Finite C) (pfd: Finite D), times_set (Union C D) (Union_preserves_Finite _ _ _ pfc pfd) = (times_set C pfc) * (times_set D pfd). intros C D h1. pose proof h1 as h2. induction h2 as [|C h3 h4]. intros h3. rewrite times_set_empty'. pose proof (empty_union D) as h4. pose proof (subsetT_eq_compat _ _ _ _ (Union_preserves_Finite Bt (Empty_set Bt) D h1 h3) h3 h4) as h5. dependent rewrite -> h5. rewrite comm_prod. rewrite one_prod. reflexivity. intro h5. pose proof (union_add_comm C D x) as h6. pose proof (Union_preserves_Finite _ _ _ h3 h5) as h7. pose proof (Add_preserves_Finite _ _ x h7) as h8. pose proof (subsetT_eq_compat _ _ _ _ (Union_preserves_Finite Bt (Add C x) D h1 h5) h8 h6) as h9. dependent rewrite -> h9. rewrite (times_set_add' C h3 x h1). rewrite (times_set_add' (Union C D) h7 x h8). specialize (h4 h3 h5). pose proof (proof_irrelevance _ (Union_preserves_Finite Bt C D h3 h5) h7) as h10. pose proof (subsetT_eq_compat _ _ _ h7 (Union_preserves_Finite Bt C D h3 h5) h7 h10) as h11. dependent rewrite <- h11. rewrite h4. apply assoc_prod. Qed. Lemma times_set_union' : forall (C D:Ensemble Bt) (pfc:Finite C) (pfd: Finite D) (pfu:Finite (Union C D)), times_set (Union C D) pfu = (times_set C pfc) * (times_set D pfd). intros C D h1 h2 h3. pose proof (times_set_union _ _ h1 h2) as h4. rewrite <- h4. apply times_set_functional. reflexivity. Qed. Lemma times_set_union'' : forall (C D:Ensemble Bt) (pf:Finite (Union C D)), let pfc := Finite_downward_closed _ _ pf _ (Union_increases_l _ _ _) in let pfd := Finite_downward_closed _ _ pf _ (Union_increases_r _ _ _) in times_set (Union C D) pf = (times_set C pfc) * (times_set D pfd). intros; apply times_set_union'. Qed. Lemma times_set_inc_le : forall (C D:Ensemble Bt) (pfc:Finite C) (pfd:Finite D) (pfi:Included C D), le (times_set D pfd) (times_set C pfc). intros C D h1 h2 h3. pose proof h2 as h2'. pose proof (decompose_setminus_inc D _ h3) as h4. generalize dependent h2. rewrite h4. intro h2. pose proof (setminus_inc D C) as h5. pose proof (Finite_downward_closed _ _ h2' _ h5) as h6. rewrite (times_set_union' _ _ h1 h6). apply times_le. Qed. Lemma plus_set_union : forall (C D:Ensemble Bt) (pfc:Finite C) (pfd: Finite D), plus_set (Union C D) (Union_preserves_Finite _ _ _ pfc pfd) = (plus_set C pfc) + (plus_set D pfd). intros C D h1. pose proof h1 as h2. induction h2 as [|C h3 h4]. intros h3. rewrite plus_set_empty'. pose proof (empty_union D) as h4. pose proof (subsetT_eq_compat _ _ _ _ (Union_preserves_Finite Bt (Empty_set Bt) D h1 h3) h3 h4) as h5. dependent rewrite -> h5. rewrite comm_sum. rewrite zero_sum. reflexivity. intro h5. pose proof (union_add_comm C D x) as h6. pose proof (Union_preserves_Finite _ _ _ h3 h5) as h7. pose proof (Add_preserves_Finite _ _ x h7) as h8. pose proof (subsetT_eq_compat _ _ _ _ (Union_preserves_Finite Bt (Add C x) D h1 h5) h8 h6) as h9. dependent rewrite -> h9. rewrite (plus_set_add' C h3 x h1). rewrite (plus_set_add' (Union C D) h7 x h8). specialize (h4 h3 h5). pose proof (proof_irrelevance _ (Union_preserves_Finite Bt C D h3 h5) h7) as h10. pose proof (subsetT_eq_compat _ _ _ h7 (Union_preserves_Finite Bt C D h3 h5) h7 h10) as h11. dependent rewrite <- h11. rewrite h4. apply assoc_sum. Qed. Lemma plus_set_union' : forall (C D:Ensemble Bt) (pfc:Finite C) (pfd: Finite D) (pfu:Finite (Union C D)), plus_set (Union C D) pfu = (plus_set C pfc) + (plus_set D pfd). intros C D h1 h2 h3. pose proof (plus_set_union _ _ h1 h2) as h4. rewrite <- h4. apply plus_set_functional. reflexivity. Qed. Lemma plus_set_union'' : forall (C D:Ensemble Bt) (pf:Finite (Union C D)), let pfc := Finite_downward_closed _ _ pf _ (Union_increases_l _ _ _) in let pfd := Finite_downward_closed _ _ pf _ (Union_increases_r _ _ _) in plus_set (Union C D) pf = (plus_set C pfc) + (plus_set D pfd). intros; apply plus_set_union'. Qed. Lemma plus_set_plus_in : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt), Inn E x -> x + (plus_set E pf) = plus_set E pf. intros E h1 x h2. pose proof (add_sub_compat_in _ _ h2) as h3. assert (h4:Finite (Add (Subtract E x) x)). rewrite h3; auto. assert (h5:plus_set E h1 = plus_set _ h4). apply plus_set_functional; auto. rewrite h5 at 1. assert (h6:Finite (Subtract E x)). eapply Finite_downward_closed. apply h4. auto with sets. rewrite (plus_set_add' _ h6 _ h4). rewrite assoc_sum. rewrite idem_sum. rewrite <- plus_set_add. pose proof (proof_irrelevance _ h4 (Add_preserves_Finite Bt (Subtract E x) x h6)) as h7. rewrite <- h7. rewrite h5. reflexivity. Qed. Lemma plus_list_1 : forall (l:list Bt), In 1 l -> plus_list l = 1. intros l; induction l as [|a l ih]; simpl; auto. intro; contradiction. intro h1. destruct h1 as [|h1]; subst. rewrite comm_sum. apply one_sum; auto. rewrite ih; auto. rewrite one_sum; auto. Qed. Lemma plus_list_opps : forall (l:list Bt) (x y:bt B), In x l -> In y l -> x = - y -> plus_list l = 1. intros l x y h1 h2 h3; subst. revert h1 h2. revert y. induction l as [|a l ih]; simpl. intros; contradiction. intros y h1 h2. destruct h1 as [h1|h1], h2 as [h2|h2]; subst. symmetry in h2. rewrite eq_comp_iff in h2. apply degen_eq; auto. rewrite <- (plus_list_plus_in _ _ _ h2). rewrite assoc_sum. rewrite (comm_sum _ (-y) y). rewrite comp_sum. rewrite comm_sum. rewrite one_sum; auto. rewrite <- (plus_list_plus_in _ _ _ h1). rewrite assoc_sum. rewrite comp_sum. rewrite comm_sum. rewrite one_sum; auto. rewrite (ih _ h1 h2); auto. rewrite one_sum; auto. Qed. Lemma plus_set_in_subtract : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt), let pfs := subtract_preserves_finite E x pf in Inn E x -> plus_set E pf = (plus_set (Subtract E x) pfs) + x. intros E h1 x h2 h3. pose proof (add_sub_compat_in _ _ h3) as h4. pose proof (plus_set_functional _ _ h4 (Add_preserves_Finite _ _ x h2) h1) as h5. rewrite plus_set_add in h5. rewrite comm_sum. rewrite h5; auto. Qed. (*kinda weak -- assumes injectivity -- maybe use the next lemma*) Lemma plus_list_map_in_dep_remove : forall {T:Type} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l Bt) (x:T) (pf:In x l), inj_dep f -> let f' := fun_in_dep_remove pfdt f x in plus_list (map_in_dep f) = (f x pf) + plus_list (map_in_dep f'). intros T l h1 f x h2 hi. simpl. do 2 rewrite plus_list_compat. assert (h3:Inn (list_to_set (map_in_dep f)) (f x h2)). rewrite <- list_to_set_in_iff. apply in_map_in_dep. rewrite (plus_set_in_subtract _ (list_to_set_finite _) _ h3). rewrite comm_sum. f_equal. apply plus_set_functional. erewrite list_to_set_map_in_dep_remove_inj. reflexivity. assumption. Qed. Lemma plus_list_map_in_dep_remove_in : forall {T:Type} {l:list T} (pfdt:type_eq_dec T) (f:fun_in_dep l Bt) (x:T) (pf:In x l), let f' := fun_in_dep_remove_in pfdt ba_eq_dec f x pf in plus_list (map_in_dep f) = (f x pf) + plus_list (map_in_dep f'). intros T l h1 f x h2. simpl. do 2 rewrite plus_list_compat. assert (h3:Inn (list_to_set (map_in_dep f)) (f x h2)). rewrite <- list_to_set_in_iff. apply in_map_in_dep. rewrite (plus_set_in_subtract _ (list_to_set_finite _) _ h3). rewrite comm_sum. f_equal. apply plus_set_functional. rewrite list_to_set_map_in_dep_remove_in; auto. Qed. Lemma plus_set_inc_le : forall (C D:Ensemble Bt) (pfc:Finite C) (pfd:Finite D) (pfi:Included C D), le (plus_set C pfc) (plus_set D pfd). intros C D h1 h2 h3. pose proof h2 as h2'. pose proof (decompose_setminus_inc D _ h3) as h4. generalize dependent h2. rewrite h4. intro h2. pose proof (setminus_inc D C) as h5. pose proof (Finite_downward_closed _ _ h2' _ h5) as h6. rewrite (plus_set_union' _ _ h1 h6). apply le_plus. Qed. Lemma plus_set_im_add : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt) (f:Bt->Bt), plus_set (Im (Add E x) f) (finite_image _ _ _ f (Add_preserves_Finite _ _ x pf)) = f x + plus_set (Im E f) (finite_image _ _ _ f pf). intros E h1 x f. pose proof (Im_add _ _ E x f) as h2. pose proof (finite_image _ _ _ f h1) as h3. pose proof (Add_preserves_Finite _ _ (f x) h3) as h4. pose proof (subsetT_eq_compat _ _ _ _ (finite_image Bt Bt (Add E x) f (Add_preserves_Finite Bt E x h1)) h4 h2) as h7. dependent rewrite -> h7. rewrite (plus_set_add' _ (finite_image Bt Bt E f h1) _ h4). reflexivity. Qed. Lemma plus_set_im_add' : forall (E:Ensemble Bt) (x:Bt) (f:Bt->Bt) (pf0: Finite E) (pf1:Finite (Im (Add E x) f)) (pf2:Finite (Im E f)), plus_set (Im (Add E x) f) pf1 = f x + (plus_set (Im E f) pf2). intros E x f h1 h2 h3. pose proof (plus_set_im_add _ h1 x f) as h4. pose proof (proof_irrelevance _ h2 (finite_image Bt Bt (Add E x) f (Add_preserves_Finite Bt E x h1))) as h5. pose proof (proof_irrelevance _ h3 (finite_image Bt Bt E f h1)) as h6. rewrite <- h5 in h4. rewrite <- h6 in h4. assumption. Qed. Lemma plus_set_im_im : forall {T U:Type} (C:Ensemble T) (f:T->U) (g:U->Bt) (pfc : Finite C) (pfiic:Finite (Im (Im C f) g)), let k := (fun x:T => g (f x)) in plus_set (Im (Im C f) g) pfiic = plus_set (Im C k) (finite_image _ _ _ k pfc). intros T U C f g h1 h3 k. pose proof (im_im C f g) as h4. apply (plus_set_functional _ _ h4 h3 (finite_image _ _ _ k h1)). Qed. Lemma times_set_im_add : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt) (f:Bt->Bt), times_set (Im (Add E x) f) (finite_image _ _ _ f (Add_preserves_Finite _ _ x pf)) = f x * times_set (Im E f) (finite_image _ _ _ f pf). intros E h1 x f. pose proof (Im_add _ _ E x f) as h2. pose proof (finite_image _ _ _ f h1) as h3. pose proof (Add_preserves_Finite _ _ (f x) h3) as h4. pose proof (subsetT_eq_compat _ _ _ _ (finite_image Bt Bt (Add E x) f (Add_preserves_Finite Bt E x h1)) h4 h2) as h7. dependent rewrite -> h7. rewrite (times_set_add' _ (finite_image Bt Bt E f h1) _ h4). reflexivity. Qed. Lemma times_set_im_add' : forall (E:Ensemble Bt) (x:Bt) (f:Bt->Bt) (pf0: Finite E) (pf1:Finite (Im (Add E x) f)) (pf2:Finite (Im E f)), times_set (Im (Add E x) f) pf1 = f x * (times_set (Im E f) pf2). intros E x f h1 h2 h3. pose proof (times_set_im_add _ h1 x f) as h4. pose proof (proof_irrelevance _ h2 (finite_image Bt Bt (Add E x) f (Add_preserves_Finite Bt E x h1))). pose proof (proof_irrelevance _ h3 (finite_image Bt Bt E f h1)). subst; auto. Qed. Lemma times_set_im_im : forall {T U:Type} (C:Ensemble T) (f:T->U) (g:U->Bt) (pfc : Finite C) (pfiic:Finite (Im (Im C f) g)), let k := (fun x:T => g (f x)) in times_set (Im (Im C f) g) pfiic = times_set (Im C k) (finite_image _ _ _ k pfc). intros T U C f g h1 h3 k. pose proof (im_im C f g) as h4. apply (times_set_functional _ _ h4 h3 (finite_image _ _ _ k h1)). Qed. Lemma times_set_zero_incl : forall (E: Ensemble (bt B)) (pf: Finite E) (E':Ensemble (bt B)) (pf':Finite E'), times_set E pf = 0 -> Included E E' -> times_set E' pf' = 0. intros E h1 E' h2 h3 h4. pose proof (times_set_inc_le _ _ h1 h2 h4) as h5. apply le_x_0. eapply trans_le. apply h5. rewrite <- h3. apply refl_le. Qed. Lemma dist_set_plus1 : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt), let f := (fun y:Bt => x *y) in x * (plus_set E pf) = plus_set (Im E f) (finite_image _ _ _ f pf). intros E h1. pose proof h1 as h2. induction h2 as [|E h3 h4 a h5]. (* Empty *) intros x f. rewrite plus_set_empty'. rewrite zero_prod. pose proof (image_empty Bt Bt f) as h3. pose proof (subsetT_eq_compat _ _ (Im (Empty_set Bt) f) (Empty_set Bt) (finite_image _ _ (Empty_set Bt) f h1) h1 h3) as h4. dependent rewrite -> h4. rewrite plus_set_empty'. reflexivity. (* Add *) intros x f. pose proof (plus_set_im_add' E a f h3 (finite_image Bt (Btype (Bc B)) (Add E a) f h1) (finite_image Bt (Btype (Bc B)) E f h3)) as h6. replace ( plus_set (Im (Add E a) f) (finite_image Bt (Btype (Bc B)) (Add E a) f h1)) with ( f a + plus_set (Im E f) (finite_image Bt (Btype (Bc B)) E f h3)). rewrite (plus_set_add' _ h3 _ h1). rewrite dist_sum. specialize (h4 h3 x). rewrite h4. unfold f. reflexivity. Qed. Lemma dist_set_plus1' : forall (E:Ensemble Bt) (pf1:Finite E) (x:Bt), let f := (fun y:Bt => x *y) in forall (pf2:Finite (Im E f)), x * (plus_set E pf1) = plus_set (Im E f) pf2. intros E h1 x f h2. pose proof (dist_set_plus1 _ h1 x) as h3. pose proof (proof_irrelevance _ h2 (finite_image Bt (Btype (Bc B)) E f h1)); subst; auto. Qed. Lemma dist_set_times1 : forall (E:Ensemble Bt) (pf:Finite E) (x:Bt), let f := (fun y:Bt => x + y) in x + (times_set E pf) = times_set (Im E f) (finite_image _ _ _ f pf). intros E h1. pose proof h1 as h2. induction h2 as [|E h3 h4 a h5]. (* Empty *) intros x f. rewrite times_set_empty'. rewrite one_sum. pose proof (image_empty Bt Bt f) as h3. pose proof (subsetT_eq_compat _ _ (Im (Empty_set Bt) f) (Empty_set Bt) (finite_image _ _ (Empty_set Bt) f h1) h1 h3) as h4. dependent rewrite -> h4. rewrite times_set_empty'. reflexivity. (* Add *) intros x f. pose proof (times_set_im_add' E a f h3 (finite_image Bt (Btype (Bc B)) (Add E a) f h1) (finite_image Bt (Btype (Bc B)) E f h3)) as h6. replace ( times_set (Im (Add E a) f) (finite_image Bt (Btype (Bc B)) (Add E a) f h1)) with ( f a * times_set (Im E f) (finite_image Bt (Btype (Bc B)) E f h3)). rewrite (times_set_add' _ h3 _ h1). rewrite dist_prod. specialize (h4 h3 x). rewrite h4. unfold f. reflexivity. Qed. Lemma dist_set_times1' : forall (E:Ensemble Bt) (pf1:Finite E) (x:Bt), let f := (fun y:Bt => x + y) in forall (pf2:Finite (Im E f)), x + (times_set E pf1) = times_set (Im E f) pf2. intros E h1 x f h2. pose proof (dist_set_times1 E h1 x) as h3. pose proof (proof_irrelevance _ h2 (finite_image Bt (Btype (Bc B)) E f h1)); subst; auto. Qed. Lemma dist_set_plus2 : forall (D E:Ensemble Bt) (pfd: Finite D) (pfe: Finite E), let f := (fun p:Bt*Bt => (fst p) * (snd p)) in (plus_set D pfd) * (plus_set E pfe) = plus_set (Im (cart_prod D E) f) (finite_image _ _ _ f (cart_prod_fin _ _ pfd pfe)). intros D E h1. pose proof h1 as h2. induction h2 as [|D h3 h4 d h5]. (*Empty D*) pose proof (@cart_prod_empty Bt _ E) as h3. intros h4 f. pose proof (image_empty _ _ f) as h5. rewrite <- h3 in h5. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Bt * Bt) (Bt ) (cart_prod (Empty_set Bt) E) f (cart_prod_fin (Empty_set Bt) E h1 h4)) (Empty_is_finite Bt) h5) as h6. unfold Bt, bt in h6. unfold Bt, bt. dependent rewrite -> h6. pose proof (plus_set_empty) as he. unfold Bt in he. rewrite he at 1. rewrite plus_set_empty'. rewrite comm_prod. rewrite zero_prod. reflexivity. (*Add D*) intros h6 f. pose proof h6 as h7. rewrite (plus_set_add' _ h3). rewrite dist_sum_r. specialize (h4 h3 h6). rewrite h4. pose (fun y:Bt => d * y) as g. pose proof (finite_image _ _ _ g h7) as h8. rewrite (dist_set_plus1' _ h6 _ h8). pose proof (cart_prod_eq' (Add D d) E) as h9. pose proof (feq_im _ _ f h9) as h10. pose proof (cart_prod_fin (Add D d) E h1 h6) as h11. pose proof (finite_image _ _ _ f h11) as h12. rewrite h10 in h12. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Bt * Bt) (Btype (Bc B)) (cart_prod (Add D d) E) f (cart_prod_fin (Add D d) E h1 h6)) h12 h10) as h13. dependent rewrite -> h13. pose proof (add_ex_family D d (fun y:Bt => cart_prod (Singleton y) E)) as h14. pose proof (f_equal (@FamilyUnion _) h14) as h15. pose proof (feq_im _ _ f h15) as h16. simpl in h16. pose proof (cart_prod_fin D E h3 h6) as h17. pose proof h17 as h17'. rewrite cart_prod_eq in h17. pose proof (family_union_add ( [S : Ensemble (Bt * Bt) | exists u : Bt, Inn D u /\ S = cart_prod (Singleton u) E]) (cart_prod (Singleton d) E)) as h18. rewrite h18 in h16. rewrite im_union in h16. pose proof (cart_prod_sing_fin' E d h7) as h19. pose proof (cart_prod_fin D E h3 h6) as h20. pose proof (finite_image _ _ _ f h19) as h21. rewrite cart_prod_eq' in h20. pose proof (finite_image _ _ _ f h20) as h22. pose proof (plus_set_union _ _ h21 h22) as h23. pose proof (subsetT_eq_compat _ _ _ _ h12 (Union_preserves_Finite Bt (Im (cart_prod (Singleton d) E) f) (Im (FamilyUnion [S : Ensemble (Bt * Bt) | exists x : Bt, Inn D x /\ S = cart_prod (Singleton x) E]) f) h21 h22) h16) as h24. dependent rewrite -> h24. rewrite plus_set_union. pose proof (cart_prod_eq' D E) as h25. pose proof (feq_im _ _ f h25) as h26. pose proof (finite_image _ _ _ f h20) as h27. pose proof (finite_image _ _ _ f h17') as h28. pose proof (subsetT_eq_compat _ _ _ _ h28 h27 h26) as h29. pose proof (proof_irrelevance _ h27 h22) as h30. rewrite h30 in h29. dependent rewrite <- h29. assert (h31:plus_set (Im E (fun y : Bt => d * y)) h8 = plus_set (Im (cart_prod (Singleton d) E) f) h21). assert (h32:(Im E (fun y : Bt => d * y)) = (Im (cart_prod (Singleton d) E) f)). apply Extensionality_Ensembles. red. split. (* <= *) red. intros z h31. destruct h31 as [z h31]. subst. apply Im_intro with (d, z). constructor. simpl. split. constructor. assumption. unfold f. simpl. reflexivity. (* >= *) red. intros pr h31. destruct h31 as [pr h31]. subst. destruct h31 as [h31]. apply Im_intro with (snd pr). tauto. unfold f. destruct h31 as [h31l h31r]. destruct h31l. reflexivity. pose proof (subsetT_eq_compat _ _ _ _ h8 h21 h32) as h33. dependent rewrite -> h33. reflexivity. rewrite h31. pose proof (proof_irrelevance _ h28 (finite_image (Bt * Bt) (Btype (Bc B)) (cart_prod D E) (fun p : Bt * Bt => fst p * snd p) (cart_prod_fin D E h3 h6))); subst; auto. Qed. Lemma dist_set_plus2' : forall (D E:Ensemble Bt) (pfd: Finite D) (pfe: Finite E), let f := (fun p:Bt*Bt => (fst p) * (snd p)) in forall (pfc: Finite (Im (cart_prod D E) f)), (plus_set D pfd) * (plus_set E pfe) = plus_set (Im (cart_prod D E) f) pfc. intros D E h1 h2 f h3. pose proof (dist_set_plus2 D E h1 h2) as h4. pose proof (proof_irrelevance _ h3 (finite_image (Bt * Bt) (Btype (Bc B)) (cart_prod D E) f (cart_prod_fin D E h1 h2))); subst; auto. Qed. Lemma dist_set_times2 : forall (D E:Ensemble Bt) (pfd: Finite D) (pfe: Finite E), let f := (fun p:Bt*Bt => (fst p) + (snd p)) in (times_set D pfd) + (times_set E pfe) = times_set (Im (cart_prod D E) f) (finite_image _ _ _ f (cart_prod_fin _ _ pfd pfe)). intros D E h1. pose proof h1 as h2. induction h2 as [|D h3 h4 d h5]. (*Empty D*) pose proof (@cart_prod_empty Bt _ E) as h3. intros h4 f. pose proof (image_empty _ _ f) as h5. rewrite <- h3 in h5. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Bt * Bt) (Btype (Bc B)) (cart_prod (Empty_set Bt) E) f (cart_prod_fin (Empty_set Bt) E h1 h4)) (Empty_is_finite Bt) h5) as h6. dependent rewrite -> h6. fold Bt. pose proof times_set_empty as he. unfold Bt in he. unfold Bt. rewrite he at 1. rewrite times_set_empty'. rewrite comm_sum. rewrite one_sum. reflexivity. (*Add D*) intros h6 f. pose proof h6 as h7. rewrite (times_set_add' _ h3). rewrite dist_prod_r. specialize (h4 h3 h6). rewrite h4. pose (fun y:Bt => d + y) as g. pose proof (finite_image _ _ _ g h7) as h8. rewrite (dist_set_times1' _ h6 _ h8). pose proof (cart_prod_eq' (Add D d) E) as h9. pose proof (feq_im _ _ f h9) as h10. pose proof (cart_prod_fin (Add D d) E h1 h6) as h11. pose proof (finite_image _ _ _ f h11) as h12. rewrite h10 in h12. pose proof (subsetT_eq_compat _ _ _ _ (finite_image (Bt * Bt) (Btype (Bc B)) (cart_prod (Add D d) E) f (cart_prod_fin (Add D d) E h1 h6)) h12 h10) as h13. dependent rewrite -> h13. pose proof (add_ex_family D d (fun y:Bt => cart_prod (Singleton y) E)) as h14. pose proof (f_equal (@FamilyUnion _) h14) as h15. pose proof (feq_im _ _ f h15) as h16. simpl in h16. pose proof (cart_prod_fin D E h3 h6) as h17. pose proof h17 as h17'. rewrite cart_prod_eq in h17. pose proof (family_union_add ( [S : Ensemble (Bt * Bt) | exists u : Bt, Inn D u /\ S = cart_prod (Singleton u) E]) (cart_prod (Singleton d) E)) as h18. rewrite h18 in h16. rewrite im_union in h16. pose proof (cart_prod_sing_fin' E d h7) as h19. pose proof (cart_prod_fin D E h3 h6) as h20. pose proof (finite_image _ _ _ f h19) as h21. rewrite cart_prod_eq' in h20. pose proof (finite_image _ _ _ f h20) as h22. pose proof (times_set_union _ _ h21 h22) as h23. pose proof (subsetT_eq_compat _ _ _ _ h12 (Union_preserves_Finite Bt (Im (cart_prod (Singleton d) E) f) (Im (FamilyUnion [S : Ensemble (Bt * Bt) | exists x : Bt, Inn D x /\ S = cart_prod (Singleton x) E]) f) h21 h22) h16) as h24. dependent rewrite -> h24. rewrite times_set_union. pose proof (cart_prod_eq' D E) as h25. pose proof (feq_im _ _ f h25) as h26. pose proof (finite_image _ _ _ f h20) as h27. pose proof (finite_image _ _ _ f h17') as h28. pose proof (subsetT_eq_compat _ _ _ _ h28 h27 h26) as h29. pose proof (proof_irrelevance _ h27 h22) as h30. rewrite h30 in h29. dependent rewrite <- h29. assert (h31:times_set (Im E (fun y : Bt => d + y)) h8 = times_set (Im (cart_prod (Singleton d) E) f) h21). assert (h32:(Im E (fun y : Bt => d + y)) = (Im (cart_prod (Singleton d) E) f)). apply Extensionality_Ensembles. red. split. (* <= *) red. intros z h31. destruct h31 as [z h31]. subst. apply Im_intro with (d, z). constructor. simpl. split. constructor. assumption. unfold f. simpl. reflexivity. (* >= *) red. intros pr h31. destruct h31 as [pr h31]. subst. destruct h31 as [h31]. apply Im_intro with (snd pr). tauto. unfold f. destruct h31 as [h31l h31r]. destruct h31l. reflexivity. pose proof (subsetT_eq_compat _ _ _ _ h8 h21 h32) as h33. dependent rewrite -> h33. reflexivity. rewrite h31. pose proof (proof_irrelevance _ h28 (finite_image (Bt * Bt) (Btype (Bc B)) (cart_prod D E) (fun p : Bt * Bt => fst p + snd p) (cart_prod_fin D E h3 h6))); subst; auto. Qed. Lemma dist_set_times2' : forall (D E:Ensemble Bt) (pfd: Finite D) (pfe: Finite E), let f := (fun p:Bt*Bt => (fst p) + (snd p)) in forall (pfc: Finite (Im (cart_prod D E) f)), (times_set D pfd) + (times_set E pfe) = times_set (Im (cart_prod D E) f) pfc. intros D E h1 h2 f h3. pose proof (dist_set_times2 D E h1 h2) as h4. pose proof (proof_irrelevance _ h3 (finite_image (Bt * Bt) (Btype (Bc B)) (cart_prod D E) f (cart_prod_fin D E h1 h2))); subst; auto. Qed. Definition plus_fin_pair_map1 {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E 0) := fun_to_fin_map pfdt C 0 pfc (fun x:T => plus_set (im1 F x) (finite_im1 F x)). Definition plus_fin_pair_map2 {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E 0) := fun_to_fin_map pfdu D 0 pfd (fun y:U => plus_set (im2 F y) (finite_im2 F y)). Definition times_fin_pair_map1 {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E 0) := fun_to_fin_map pfdt C 0 pfc (fun x:T => times_set (im1 F x) (finite_im1 F x)). Definition times_fin_pair_map2 {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E 0) := fun_to_fin_map pfdu D 0 pfd (fun y:U => times_set (im2 F y) (finite_im2 F y)). Lemma plus_fin_pair_map1_list_compat : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E 0) (nml:nice_map_lists2 F), let f := fun (x:T) (pf:Inn C x) => plus_list (im1l (fpl2 F nml) x (impl_in_n_la2 F nml pfc _ pf)) in fin_map_eq (plus_fin_pair_map1 pfdt pfc F) (fun_in_ens_to_fin_map pfdt f pfc 0). intros T U C D E hdt h1 F nml. simpl. red. ex_goal. red. intros x h3. destruct h3 as [x h3]. subst. econstructor. apply plus_set_compat'. apply im1_im1l_compat. assumption. exists hex. apply fin_map_ext. intros x h3. rewrite <- fin_map_new_ran_compat. erewrite fun_in_ens_to_fin_map_compat. unfold plus_fin_pair_map1. rewrite fun_to_fin_map_compat; auto. apply plus_set_compat'; auto. apply im1_im1l_compat; auto. assumption. Grab Existential Variables. assumption. assumption. Qed. Lemma plus_fin_pair_map2_list_compat : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E 0) (nml:nice_map_lists2 F), let f := fun (y:U) (pf:Inn D y) => plus_list (im2l (fpl2 F nml) y (impl_in_n_lb2 F nml pfd _ pf)) in fin_map_eq (plus_fin_pair_map2 pfdu pfd F) (fun_in_ens_to_fin_map pfdu f pfd 0). intros T U C D E hdu h1 F nml. simpl. red. ex_goal. red. intros x h3. destruct h3 as [x h3]. subst. econstructor. apply plus_set_compat'. apply im2_im2l_compat. assumption. exists hex. apply fin_map_ext. intros y h3. rewrite <- fin_map_new_ran_compat. erewrite fun_in_ens_to_fin_map_compat. unfold plus_fin_pair_map2. rewrite fun_to_fin_map_compat; auto. apply plus_set_compat'; auto. apply im2_im2l_compat; auto. assumption. Grab Existential Variables. assumption. assumption. Qed. Lemma times_fin_pair_map1_list_compat : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E 0) (nml:nice_map_lists2 F), let f := fun (x:T) (pf:Inn C x) => times_list (im1l (fpl2 F nml) x (impl_in_n_la2 F nml pfc _ pf)) in fin_map_eq (times_fin_pair_map1 pfdt pfc F) (fun_in_ens_to_fin_map pfdt f pfc 0). intros T U C D hdt E h1 F nml. simpl. red. ex_goal. red. intros x h3. destruct h3 as [x h3]. subst. econstructor. apply times_set_compat'. apply im1_im1l_compat. assumption. exists hex. apply fin_map_ext. intros x h3. rewrite <- fin_map_new_ran_compat. erewrite fun_in_ens_to_fin_map_compat. unfold times_fin_pair_map1. rewrite fun_to_fin_map_compat; auto. apply times_set_compat'; auto. apply im1_im1l_compat; auto. assumption. Grab Existential Variables. assumption. assumption. Qed. Lemma times_fin_pair_map2_list_compat : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E 0) (nml:nice_map_lists2 F), let f := fun (y:U) (pf:Inn D y) => times_list (im2l (fpl2 F nml) y (impl_in_n_lb2 F nml pfd _ pf)) in fin_map_eq (times_fin_pair_map2 pfdu pfd F) (fun_in_ens_to_fin_map pfdu f pfd 0). intros T U C D E hdu h1 F nml. simpl. red. ex_goal. red. intros x h3. destruct h3 as [x h3]. subst. econstructor. apply times_set_compat'. apply im2_im2l_compat. assumption. exists hex. apply fin_map_ext. intros y h3. rewrite <- fin_map_new_ran_compat. erewrite fun_in_ens_to_fin_map_compat. unfold times_fin_pair_map2. rewrite fun_to_fin_map_compat; auto. apply times_set_compat'; auto. apply im2_im2l_compat; auto. assumption. Grab Existential Variables. assumption. assumption. Qed. Lemma plus_fin_pair_map2_functional : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (pf:Finite D), forall (F G:Fin_map (cart_prod C D) E 0), F = G -> fin_map_eq (plus_fin_pair_map2 pfdu pf F) (plus_fin_pair_map2 pfdu pf G). intros T U C D E hdu h1 F G h2. subst. red. ex_goal. auto with sets. exists hex. apply fin_map_ext. intros x h4. symmetry. apply fin_map_new_ran_compat; auto. Qed. Definition fin_map_times {T:Type} {A:Ensemble T} {C:Ensemble Bt} (f:Fin_map A C 0) : Bt. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ A (fin_map_app f) h1) as h2. refine (times_set _ h2). Defined. Lemma fin_map_times_list_compat : forall {T:Type} {A:Ensemble T} {C:Ensemble Bt} (F:Fin_map A C 0) (nml:nice_map_lists F), fin_map_times F = times_list (n_im F nml). intros T A C F nml. unfold fin_map_times. pose proof (n_im_im_fin_map_compat F nml) as h1. apply times_set_compat'. rewrite h1. unfold im_fin_map. reflexivity. Qed. Lemma fin_map_times_empty1 : forall {T:Type} {C:Ensemble Bt} (F:Fin_map (Empty_set T) C 0), fin_map_times F = 1. intros T C F. pose proof (fin_map_empty1 F) as h1. unfold fin_map_times. pose proof (image_empty T Bt (fin_map_app F)) as h2. pose proof (Empty_is_finite Bt) as h3. pose proof (subsetT_eq_compat _ _ _ _ (finite_image T Bt (Empty_set T) (fin_map_app F) (fin_map_fin_dom F)) h3 h2) as h4. dependent rewrite -> h4. apply times_set_empty'. Qed. Lemma fin_map_eq_times : forall {T:Type} (A:Ensemble T) (C E:Ensemble Bt) (F:Fin_map A C 0) (G:Fin_map A E 0), fin_map_eq F G -> fin_map_times F = fin_map_times G. intros T A C E F G h1. destruct h1 as [h1 h2]. pose proof (fin_map_new_ran_compat F (fin_map_fin_ran G) h1) as h3. rewrite h2 in h3. unfold fin_map_times. apply times_set_functional. apply im_ext_in. intros x h4. apply h3; auto. Qed. Lemma plus_fin_pair_map2_cart_empty_eq1 : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (C:Ensemble Bt) (D:Ensemble U) (pfd:Finite D) (F:Fin_map (cart_prod (Empty_set T) D) C 0), fin_map_eq (plus_fin_pair_map2 pfdu pfd F) (fun_to_fin_map pfdu D 0 pfd (fun x => 0)). intros T U hdt hdu C D h1 F. pose proof (cart_empty_map11_compat hdt hdu F h1) as hc. pose proof (fin_map_fin_ran F) as hf. pose proof (proof_irrelevance _ hf (fin_map_fin_ran F)) as hpi. rewrite <- hpi in hc. clear hpi. subst. apply fps_eq_fin_map_eq. pose proof h1 as h3. induction h3. do 2 rewrite image_empty. auto with sets. red. intros a h4. destruct h4 as [a h5 b h6]. apply Im_intro with a. assumption. symmetry in h6. rewrite <- h6. gen0. rewrite im2_empty1. intro. rewrite plus_set_empty'; auto. unfold plus_fin_pair_map2. apply fin_map_to_fps_fun_to_fin_map_functional; auto. apply functional_extensionality. intro y. gen0. rewrite im2_empty1. intro. apply plus_set_empty'. Qed. Lemma plus_fin_pair_map2_cart_empty : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (C:Ensemble Bt) (pfe:Finite (Empty_set _)) (F:Fin_map (cart_prod A (Empty_set U)) C 0) (y:U), Finite A -> plus_fin_pair_map2 pfdu pfe F |-> y = 0. intros T U hdt A C h1 F y h2. unfold plus_fin_pair_map2. rewrite fun_to_fin_map_empty_set1. apply empty_map1_def. Qed. Lemma plus_fin_pair_map2_cart_empty_eq2 : forall {T U:Type} (A:Ensemble T) (C:Ensemble Bt) (pfdu:type_eq_dec U) (pfa:Finite A) (pfe:Finite (Empty_set U)) (F:Fin_map (cart_prod A (Empty_set U)) C 0), fin_map_eq (plus_fin_pair_map2 pfdu pfe F) (empty_map1 U Bt pfdu 0 _ (fin_map_fin_ran F)). intros T U A C hdu h1 h3 F. apply fps_eq_fin_map_eq. rewrite image_empty. auto with sets. apply fin_map_to_fps_ext. intro x. rewrite plus_fin_pair_map2_cart_empty. rewrite empty_map1_def. reflexivity. assumption. Qed. Lemma fin_map_times_sing : forall {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (pf:Finite A) (val:Bt), A <> Empty_set _ -> fin_map_times (fin_map_sing pfdt A pf 0 val) = val. intros T hdt A h1 val h2. pose proof (im_fin_map_sing hdt A 0 val h1 h2) as h3. rewrite <- times_set_sing. unfold fin_map_times. apply times_set_functional. assumption. Qed. Lemma fin_map_times_cart_empty11 : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (C:Ensemble Bt) (D:Ensemble U) (pfd:Finite D) (pfde:D <> Empty_set _) (F:Fin_map (cart_prod (Empty_set T) D) C 0), fin_map_times (plus_fin_pair_map2 pfdu pfd F) = 0. intros T U hdt hdu C D h1 h0 F. pose proof (fin_map_cart_empty11 F) as h3. pose proof (plus_fin_pair_map2_cart_empty_eq1 hdt hdu C D h1 F) as h4. rewrite (fin_map_eq_times _ _ _ _ _ h4). pose proof (fun_to_fin_map_sing_im hdu D 0 0 h1 (U:=Bt)) as h5. pose proof (fin_map_eq_times _ _ _ _ _ h5) as h6. replace ( fin_map_times (fun_to_fin_map hdu D 0 h1 (fun _ : U => 0))) with (fin_map_times (fin_map_sing hdu D h1 0 0)). apply fin_map_times_sing. assumption. Qed. Lemma fin_map_times_cart_empty21 : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (C:Ensemble Bt) (pfa:Finite A) (pfe:Finite (Empty_set U)) (F:Fin_map (cart_prod A (Empty_set U)) C 0), fin_map_times (plus_fin_pair_map2 pfdu pfe F) = 1. intros T U hdu A C h1 h2 F. pose proof (plus_fin_pair_map2_cart_empty_eq2 _ _ hdu h1 h2 F) as h3. rewrite (fin_map_eq_times _ _ _ _ _ h3). rewrite fin_map_times_empty1. reflexivity. Qed. Lemma fin_map_times_empty2 : forall {T:Type} {A:Ensemble T} (F:Fin_map A (Empty_set Bt) 0), fin_map_times F = 1. intros T A F. inversion F as [hdt h1 h2 S h3]. induction h1 as [|A h4 h5 x h6]. apply fin_map_times_empty1. pose proof (add_not_empty A x) as h7. pose proof (Add_preserves_Finite _ _ x h4) as h8. pose proof (empty_map_ex2_fin _ 0 h8 h7 F). contradiction. Qed. Definition fin_map_plus {T:Type} {A:Ensemble T} {C:Ensemble Bt} (f:Fin_map A C 0) : Bt. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ A (fin_map_app f) h1) as h2. refine (plus_set _ h2). Defined. Lemma fin_map_plus_list_compat : forall {T:Type} {A:Ensemble T} {C:Ensemble Bt} (F:Fin_map A C 0) (nml:nice_map_lists F), fin_map_plus F = plus_list (n_im F nml). intros T A C F nml. unfold fin_map_times. pose proof (n_im_im_fin_map_compat F nml) as h1. apply plus_set_compat'. rewrite h1. unfold im_fin_map. reflexivity. Qed. Lemma fin_map_plus_empty1 : forall {T:Type} {C:Ensemble Bt} (F:Fin_map (Empty_set T) C 0), fin_map_plus F = 0. intros T C F. pose proof (fin_map_empty1 F) as h1. unfold fin_map_plus. pose proof (image_empty T Bt (fin_map_app F)) as h2. pose proof (Empty_is_finite Bt) as h3. pose proof (subsetT_eq_compat _ _ _ _ (finite_image T Bt (Empty_set T) (fin_map_app F) (fin_map_fin_dom F)) h3 h2) as h4. dependent rewrite -> h4. apply plus_set_empty'. Qed. Lemma fin_map_plus_empty2 : forall {T:Type} {A:Ensemble T} (F:Fin_map A (Empty_set Bt) 0), fin_map_plus F = 0. intros T A F. inversion F as [hdt h1 h2 S h3]. induction h1 as [|A h4 h5 x h6]. apply fin_map_plus_empty1. pose proof (add_not_empty A x) as h7. pose proof (Add_preserves_Finite _ _ x h4) as h8. pose proof (empty_map_ex2_fin _ 0 h8 h7 F). contradiction. Qed. Lemma fin_map_eq_plus : forall {T:Type} (A:Ensemble T) (C E:Ensemble Bt) (F:Fin_map A C 0) (G:Fin_map A E 0), fin_map_eq F G -> fin_map_plus F = fin_map_plus G. intros T A C E F G h1. destruct h1 as [h1 h2]. pose proof (fin_map_new_ran_compat F (fin_map_fin_ran G) h1) as h3. rewrite h2 in h3. unfold fin_map_plus. apply plus_set_functional. apply im_ext_in. apply h3; auto. Qed. Definition times_plus_fin_pair_map1 {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E 0) := fin_map_times (plus_fin_pair_map1 pfdt pfc F). Definition times_plus_fin_pair_map2 {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E 0) := fin_map_times (plus_fin_pair_map2 pfdu pfd F). Definition plus_times_fin_pair_map1 {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E 0) := fin_map_plus (times_fin_pair_map1 pfdt pfc F). Definition plus_times_fin_pair_map2 {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E 0) := fin_map_plus (times_fin_pair_map2 pfdu pfd F). Lemma times_plus_fin_pair_map1_list_compat : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E 0) (nml:nice_map_lists (plus_fin_pair_map1 pfdt pfc F)), times_plus_fin_pair_map1 pfdt pfc F = times_list (n_im _ nml). intros T U C D E hdt pfc F nml. apply fin_map_times_list_compat. Qed. Lemma plus_fun_fin_map_to_fun_comm : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (F:Fin_map (cart_prod C D) E 0) (pfc:Finite C) (x:T) (lb:list U), list_to_set lb = D -> Inn C x -> plus_fun (fun y:U => fin_map_to_fun F (x, y)) lb = fun_to_fin_map pfdt C 0 pfc (fun x:T => plus_set (im1 F x) (finite_im1 F x)) |-> x. intros T U C D E hdt F h1 x lb h0 h2. unfold plus_fun. rewrite fun_to_fin_map_compat. pose proof (plus_set_compat'). rewrite <- (plus_set_compat' _ _ (list_to_set_finite _) (eq_refl _)). apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros z h5. rewrite <- list_to_set_in_iff in h5. unfold Bt in h5. rewrite in_map_iff in h5. destruct h5 as [b h5]. destruct h5 as [h5l h5r]. econstructor. assumption. rewrite list_to_set_in_iff in h5r. unfold Bt in h0. rewrite h0 in h5r. apply h5r. subst. symmetry. apply fin_map_to_fun_compat. red. intros z h5. destruct h5 as [h b h5 z h6]. rewrite <- list_to_set_in_iff. unfold Bt. rewrite in_map_iff. exists b. split. subst. symmetry. apply fin_map_to_fun_compat. unfold Bt in h0. rewrite <- h0 in h5. unfold Bt in h5. rewrite <- list_to_set_in_iff in h5. assumption. assumption. Qed. Lemma plus_fun_fin_map_to_fun_comm' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (F:Fin_map (cart_prod C D) E 0) (pfd:Finite D) (y:U) (la:list T), list_to_set la = C -> Inn D y -> plus_fun (fun x:T => fin_map_to_fun F (x, y)) la = fun_to_fin_map pfdu D 0 pfd (fun x:U => plus_set (im2 F x) (finite_im2 F x)) |-> y. intros T U C D E hdu F h1 y la h0 h2. unfold plus_fun. rewrite fun_to_fin_map_compat. pose proof (plus_set_compat'). rewrite <- (plus_set_compat' _ _ (list_to_set_finite _) (eq_refl _)). apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. intros z h5. rewrite <- list_to_set_in_iff in h5. unfold Bt in h5. rewrite in_map_iff in h5. destruct h5 as [a h5]. destruct h5 as [h5l h5r]. econstructor. assumption. rewrite list_to_set_in_iff in h5r. unfold Bt in h0. rewrite h0 in h5r. apply h5r. subst. symmetry. apply fin_map_to_fun_compat. red. intros z h5. destruct h5 as [h a h5 z h6]. rewrite <- list_to_set_in_iff. unfold Bt. rewrite in_map_iff. exists a. split. subst. symmetry. apply fin_map_to_fun_compat. unfold Bt in h0. rewrite <- h0 in h5. unfold Bt in h5. rewrite <- list_to_set_in_iff in h5. assumption. assumption. Qed. Lemma times_plus_fin_pair_map1_list_compat' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E 0) (nml2:nice_map_lists2 F) (nml:nice_map_lists (plus_fin_pair_map1 pfdt pfc F)), (n_la2 _ nml2) = (n_la _ nml) -> times_list (n_im _ nml) = times_plus_fun1 (n_la2 _ nml2) (n_lb2 _ nml2) (f_no_pr (fin_map_to_fun F)). intros T U C D E hdt pfc F nml2 nml h0. unfold times_plus_fun1. unfold times_fun. do 2 rewrite <- (times_set_compat' _ _ (list_to_set_finite _) (eq_refl _)). apply times_set_functional. apply Extensionality_Ensembles. red. split. red. (* <= *) intros y h1. rewrite <- list_to_set_in_iff in h1. rewrite <- list_to_set_in_iff. unfold Bt. rewrite in_map_iff. unfold n_im in h1. unfold Bt in h1. simpl in h1. pose proof (in_map_iff (snd (B:=Btype (Bc B))) (n_lp (plus_fin_pair_map1 hdt pfc F) nml) y) as hi. unfold bt, Bt in hi, h1. simpl in hi, h1. rewrite hi in h1. destruct h1 as [pr h1]. destruct h1 as [h1l h1r]. exists (fst pr). split. unfold f_no_pr. rewrite list_to_set_in_iff in h1r. pose proof (lp_compat _ nml) as h2. unfold Bt in h2. unfold Bt in h1r. rewrite h2 in h1r. pose proof (fin_map_to_fps_fin_map_app_compat _ _ h1r) as h3. subst. unfold plus_fin_pair_map1 in h3. pose proof (fin_map_to_fps_compat (plus_fin_pair_map1 hdt pfc F)) as h4. destruct h4 as [h4 h5]. destruct h4 as [h4a h4b]. pose proof (h4b _ h1r) as h6. destruct h6 as [h6l h6r]. rewrite <- h3. apply plus_fun_fin_map_to_fun_comm. destruct h6r as [y h6r x h8]. pose proof (lab_empty2 _ nml2) as h11. destruct (nil_dec' (n_lab2 _ nml2)) as [h9 | h10]. destruct h11 as [h11l | h11r]. destruct h11l as [h11a h11b]. rewrite h11b in h6r. contradiction. destruct h11r as [h11a h11b]. unfold Bt. unfold Bt in h11a. rewrite h11a. rewrite h11b. simpl. reflexivity. destruct h11 as [h11l h11r]. assumption. assumption. pose proof (fin_map_to_fps_compat (plus_fin_pair_map1 hdt pfc F)) as h4. destruct h4 as [h4l h4r]. destruct h4l as [h4a h4b]. pose proof (lp_compat _ nml) as h2. pose proof h1r as h1'. rewrite list_to_set_in_iff in h1r. unfold Bt in h2. unfold Bt in h1r. rewrite h2 in h1r. pose proof (h4b _ h1r) as h5. destruct h5 as [h5l h5r]. unfold Bt. unfold Bt in h0. rewrite h0. rewrite list_to_set_in_iff. pose proof (la_compat _ nml) as h6. unfold Bt in h6. rewrite h6. assumption. (* >= *) red. intros y h1. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h1. unfold Bt in h1. rewrite in_map_iff in h1. destruct h1 as [x h1]. destruct h1 as [h1l h1r]. unfold n_im. unfold Bt, bt. rewrite in_map_iff at 1. exists (x, y). simpl. split; auto. rewrite list_to_set_in_iff. pose proof (lp_compat _ nml) as h2. unfold Bt, bt in h2. rewrite h2. unfold Bt. unfold Bt in h0. rewrite h0 in h1r. pose proof (la_compat _ nml) as h3. unfold Bt in h3. rewrite list_to_set_in_iff in h1r. rewrite h3 in h1r. apply fin_map_app_fin_map_to_fps_compat. simpl; auto. simpl. rewrite <- h1l. unfold f_no_pr. unfold plus_fin_pair_map1. symmetry. apply plus_fun_fin_map_to_fun_comm. pose proof (lab_empty2 _ nml2) as h5. destruct (nil_dec' (n_lab2 F nml2)) as [h6 | h7]. destruct h5 as [h5l | h5r]. destruct h5l as [h5a h5b]. rewrite h5b in h1r. contradiction. destruct h5r as [h5a h5b]. unfold Bt in h5a. rewrite h5a. rewrite h5b. simpl. reflexivity. destruct h5. assumption. assumption. Qed. Lemma times_plus_fin_pair_map2_list_compat : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E 0) (nml:nice_map_lists (plus_fin_pair_map2 pfdu pfd F)), times_plus_fin_pair_map2 pfdu pfd F = times_list (n_im _ nml). intros T U C D E pfc F nml. apply fin_map_times_list_compat. Qed. Lemma times_plus_fin_pair_map2_list_compat' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E 0) (nml2:nice_map_lists2 F) (nml:nice_map_lists (plus_fin_pair_map2 pfdu pfd F)), (n_lb2 _ nml2) = (n_la _ nml) -> times_list (n_im _ nml) = times_plus_fun2 (n_la2 _ nml2) (n_lb2 _ nml2) (f_no_pr (fin_map_to_fun F)). intros T U C D E hdu pfd F nml2 nml h0. unfold times_plus_fun2. unfold times_fun. do 2 rewrite <- (times_set_compat' _ _ (list_to_set_finite _) (eq_refl _)). apply times_set_functional. apply Extensionality_Ensembles. red. split. red. (* <= *) intros y h1. rewrite <- list_to_set_in_iff in h1. rewrite <- list_to_set_in_iff. unfold Bt. rewrite in_map_iff. unfold n_im in h1. unfold Bt, bt in h1. rewrite in_map_iff in h1. destruct h1 as [pr h1]. destruct h1 as [h1l h1r]. exists (fst pr). split. unfold f_no_pr. rewrite list_to_set_in_iff in h1r. pose proof (lp_compat _ nml) as h2. unfold Bt, bt in h2. unfold Bt, bt in h1r. rewrite h2 in h1r. pose proof (fin_map_to_fps_fin_map_app_compat _ _ h1r) as h3. subst. unfold plus_fin_pair_map2 in h3. pose proof (fin_map_to_fps_compat (plus_fin_pair_map2 hdu pfd F)) as h4. destruct h4 as [h4 h5]. destruct h4 as [h4a h4b]. pose proof (h4b _ h1r) as h6. destruct h6 as [h6l h6r]. rewrite <- h3. apply plus_fun_fin_map_to_fun_comm'. destruct h6r as [y h6r x h8]. pose proof (lab_empty2 _ nml2) as h11. destruct (nil_dec' (n_lab2 _ nml2)) as [h9 | h10]. destruct h11 as [h11l | h11r]. Focus 2. destruct h11r as [h11a h11b]. rewrite h11b in h6r. contradiction. destruct h11l as [h11a h11b]. unfold Bt. unfold Bt in h11a. rewrite h11a. rewrite h11b. simpl. reflexivity. destruct h11 as [h11l h11r]. assumption. assumption. pose proof (fin_map_to_fps_compat (plus_fin_pair_map2 hdu pfd F)) as h4. destruct h4 as [h4l h4r]. destruct h4l as [h4a h4b]. pose proof (lp_compat _ nml) as h2. pose proof h1r as h1'. rewrite list_to_set_in_iff in h1r. unfold Bt, bt in h2. unfold Bt, bt in h1r. rewrite h2 in h1r. pose proof (h4b _ h1r) as h5. destruct h5 as [h5l h5r]. unfold Bt. unfold Bt in h0. rewrite h0. rewrite list_to_set_in_iff. pose proof (la_compat _ nml) as h6. unfold Bt in h6. rewrite h6. assumption. (* >= *) red. intros y h1. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h1. unfold Bt in h1. rewrite in_map_iff in h1. destruct h1 as [x h1]. destruct h1 as [h1l h1r]. unfold n_im. unfold Bt, bt. rewrite in_map_iff. exists (x, y). simpl. split; auto. rewrite list_to_set_in_iff. pose proof (lp_compat _ nml) as h2. unfold Bt, bt in h2. rewrite h2. unfold Bt. unfold Bt in h0. rewrite h0 in h1r. pose proof (la_compat _ nml) as h3. unfold Bt in h3. rewrite list_to_set_in_iff in h1r. rewrite h3 in h1r. apply fin_map_app_fin_map_to_fps_compat. simpl; auto. simpl. rewrite <- h1l. unfold f_no_pr. unfold plus_fin_pair_map2. symmetry. apply plus_fun_fin_map_to_fun_comm'. pose proof (lab_empty2 _ nml2) as h5. destruct (nil_dec' (n_lab2 F nml2)) as [h6 | h7]. destruct h5 as [h5l | h5r]. Focus 2. destruct h5r as [h5a h5b]. rewrite h5b in h1r. contradiction. destruct h5l as [h5a h5b]. unfold Bt in h5a. rewrite h5a. rewrite h5b. simpl. reflexivity. destruct h5. assumption. assumption. Qed. Lemma plus_times_fin_pair_map1_list_compat : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E 0) (nml:nice_map_lists (times_fin_pair_map1 pfdt pfc F)), plus_times_fin_pair_map1 pfdt pfc F = plus_list (n_im _ nml). intros T U C D E pfc F nml. apply fin_map_plus_list_compat. Qed. Lemma times_fun_fin_map_to_fun_comm : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (F:Fin_map (cart_prod C D) E 0) (pfc:Finite C) (x:T) (lb:list U), list_to_set lb = D -> Inn C x -> times_fun (fun y:U => fin_map_to_fun F (x, y)) lb = fun_to_fin_map pfdt C 0 pfc (fun x:T => times_set (im1 F x) (finite_im1 F x)) |-> x. intros T U C D E hdt F h1 x lb h0 h2. unfold times_fun. rewrite fun_to_fin_map_compat. pose proof (times_set_compat'). rewrite <- (times_set_compat' _ _ (list_to_set_finite _) (eq_refl _)). apply times_set_functional. apply Extensionality_Ensembles. red. split. red. intros z h5. rewrite <- list_to_set_in_iff in h5. unfold Bt in h5. rewrite in_map_iff in h5. destruct h5 as [b h5]. destruct h5 as [h5l h5r]. econstructor. assumption. subst. rewrite list_to_set_in_iff in h5r. apply h5r. subst. symmetry. apply fin_map_to_fun_compat. red. intros z h5. destruct h5 as [h b h5 z h6]. rewrite <- list_to_set_in_iff. unfold Bt. rewrite in_map_iff. exists b. split. subst. symmetry. apply fin_map_to_fun_compat. unfold Bt in h0. rewrite <- h0 in h5. unfold Bt in h5. rewrite <- list_to_set_in_iff in h5. assumption. assumption. Qed. Lemma plus_times_fin_pair_map1_list_compat' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E 0) (nml2:nice_map_lists2 F) (nml:nice_map_lists (times_fin_pair_map1 pfdt pfc F)), (n_la2 _ nml2) = (n_la _ nml) -> plus_list (n_im _ nml) = plus_times_fun1 (n_la2 _ nml2) (n_lb2 _ nml2) (f_no_pr (fin_map_to_fun F)). intros T U C D E hdt pfc F nml2 nml h0. unfold plus_times_fun1. unfold plus_fun. do 2 rewrite <- (plus_set_compat' _ _ (list_to_set_finite _) (eq_refl _)). apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. (* <= *) intros y h1. rewrite <- list_to_set_in_iff in h1. rewrite <- list_to_set_in_iff. unfold Bt. rewrite in_map_iff. unfold n_im in h1. unfold Bt, bt in h1. rewrite in_map_iff in h1. destruct h1 as [pr h1]. destruct h1 as [h1l h1r]. exists (fst pr). split. unfold f_no_pr. rewrite list_to_set_in_iff in h1r. pose proof (lp_compat _ nml) as h2. unfold Bt, bt in h2. unfold Bt, bt in h1r. rewrite h2 in h1r. pose proof (fin_map_to_fps_fin_map_app_compat _ _ h1r) as h3. subst. unfold times_fin_pair_map1 in h3. pose proof (fin_map_to_fps_compat (times_fin_pair_map1 hdt pfc F)) as h4. destruct h4 as [h4 h5]. destruct h4 as [h4a h4b]. pose proof (h4b _ h1r) as h6. destruct h6 as [h6l h6r]. rewrite <- h3. apply times_fun_fin_map_to_fun_comm. destruct h6r as [y h6r x h8]. pose proof (lab_empty2 _ nml2) as h11. destruct (nil_dec'(n_lab2 _ nml2)) as [h9 | h10]. destruct h11 as [h11l | h11r]. destruct h11l as [h11a h11b]. rewrite h11b in h6r. contradiction. destruct h11r as [h11a h11b]. unfold Bt. unfold Bt in h11a. rewrite h11a. rewrite h11b. simpl. reflexivity. destruct h11 as [h11l h11r]. assumption. assumption. pose proof (fin_map_to_fps_compat (times_fin_pair_map1 hdt pfc F)) as h4. destruct h4 as [h4l h4r]. destruct h4l as [h4a h4b]. pose proof (lp_compat _ nml) as h2. pose proof h1r as h1'. rewrite list_to_set_in_iff in h1r. unfold Bt, bt in h2. unfold Bt, bt in h1r. rewrite h2 in h1r. pose proof (h4b _ h1r) as h5. destruct h5 as [h5l h5r]. unfold Bt. unfold Bt in h0. rewrite h0. rewrite list_to_set_in_iff. pose proof (la_compat _ nml) as h6. unfold Bt in h6. rewrite h6. assumption. (* >= *) red. intros y h1. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h1. unfold Bt in h1. rewrite in_map_iff in h1. destruct h1 as [x h1]. destruct h1 as [h1l h1r]. unfold n_im. unfold Bt, bt. rewrite in_map_iff. exists (x, y). simpl. split; auto. rewrite list_to_set_in_iff. pose proof (lp_compat _ nml) as h2. unfold Bt, bt in h2. rewrite h2. unfold Bt. unfold Bt in h0. rewrite h0 in h1r. pose proof (la_compat _ nml) as h3. unfold Bt in h3. rewrite list_to_set_in_iff in h1r. rewrite h3 in h1r. apply fin_map_app_fin_map_to_fps_compat. simpl; auto. simpl. rewrite <- h1l. unfold f_no_pr. unfold times_fin_pair_map1. symmetry. apply times_fun_fin_map_to_fun_comm. pose proof (lab_empty2 _ nml2) as h5. destruct (nil_dec' (n_lab2 F nml2)) as [h6 | h7]. destruct h5 as [h5l | h5r]. destruct h5l as [h5a h5b]. rewrite h5b in h1r. contradiction. destruct h5r as [h5a h5b]. unfold Bt in h5a. rewrite h5a. rewrite h5b. simpl. reflexivity. destruct h5. assumption. assumption. Qed. Lemma plus_times_fin_pair_map2_list_compat : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E 0) (nml:nice_map_lists (times_fin_pair_map2 pfdu pfd F)), plus_times_fin_pair_map2 pfdu pfd F = plus_list (n_im _ nml). intros. apply fin_map_plus_list_compat. Qed. Lemma times_fun_fin_map_to_fun_comm' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (F:Fin_map (cart_prod C D) E 0) (pfd:Finite D) (y:U) (la:list T), list_to_set la = C -> Inn D y -> times_fun (fun x:T => fin_map_to_fun F (x, y)) la = fun_to_fin_map pfdu D 0 pfd (fun x:U => times_set (im2 F x) (finite_im2 F x)) |-> y. intros T U C D E hdu F h1 y la h0 h2. unfold times_fun. rewrite fun_to_fin_map_compat. pose proof (times_set_compat'). rewrite <- (times_set_compat' _ _ (list_to_set_finite _) (eq_refl _)). apply times_set_functional. apply Extensionality_Ensembles. red. split. red. intros z h5. rewrite <- list_to_set_in_iff in h5. unfold Bt in h5. rewrite in_map_iff in h5. destruct h5 as [a h5]. destruct h5 as [h5l h5r]. econstructor. assumption. subst. rewrite list_to_set_in_iff in h5r. apply h5r. unfold Bt in h0. subst. rewrite fin_map_to_fun_compat; auto. red. intros z h5. destruct h5 as [h a h5 z h6]. rewrite <- list_to_set_in_iff. unfold Bt. rewrite in_map_iff. exists a. split. subst. symmetry. apply fin_map_to_fun_compat. unfold Bt in h0. rewrite <- h0 in h5. unfold Bt in h5. rewrite <- list_to_set_in_iff in h5. assumption. assumption. Qed. Lemma plus_times_fin_pair_map2_list_compat' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Bt} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E 0) (nml2:nice_map_lists2 F) (nml:nice_map_lists (times_fin_pair_map2 pfdu pfd F)), (n_lb2 _ nml2) = (n_la _ nml) -> plus_list (n_im _ nml) = plus_times_fun2 (n_la2 _ nml2) (n_lb2 _ nml2) (f_no_pr (fin_map_to_fun F)). intros T U C D E hdu pfd F nml2 nml h0. unfold plus_times_fun2. unfold plus_fun. do 2 rewrite <- (plus_set_compat' _ _ (list_to_set_finite _) (eq_refl _)). apply plus_set_functional. apply Extensionality_Ensembles. red. split. red. (* <= *) intros y h1. rewrite <- list_to_set_in_iff in h1. rewrite <- list_to_set_in_iff. unfold Bt. rewrite in_map_iff. unfold n_im in h1. unfold Bt, bt in h1. rewrite in_map_iff in h1. destruct h1 as [pr h1]. destruct h1 as [h1l h1r]. exists (fst pr). split. unfold f_no_pr. rewrite list_to_set_in_iff in h1r. pose proof (lp_compat _ nml) as h2. unfold Bt, bt in h2. unfold Bt, bt in h1r. rewrite h2 in h1r. pose proof (fin_map_to_fps_fin_map_app_compat _ _ h1r) as h3. subst. unfold times_fin_pair_map2 in h3. pose proof (fin_map_to_fps_compat (times_fin_pair_map2 hdu pfd F)) as h4. destruct h4 as [h4 h5]. destruct h4 as [h4a h4b]. pose proof (h4b _ h1r) as h6. destruct h6 as [h6l h6r]. rewrite <- h3. apply times_fun_fin_map_to_fun_comm'. destruct h6r as [y h6r x h8]. pose proof (lab_empty2 _ nml2) as h11. destruct (nil_dec' (n_lab2 _ nml2)) as [h9 | h10]. destruct h11 as [h11l | h11r]. Focus 2. destruct h11r as [h11a h11b]. rewrite h11b in h6r. contradiction. destruct h11l as [h11a h11b]. unfold Bt. unfold Bt in h11a. rewrite h11a. rewrite h11b. simpl. reflexivity. destruct h11 as [h11l h11r]. assumption. assumption. pose proof (fin_map_to_fps_compat (times_fin_pair_map2 hdu pfd F)) as h4. destruct h4 as [h4l h4r]. destruct h4l as [h4a h4b]. pose proof (lp_compat _ nml) as h2. pose proof h1r as h1'. rewrite list_to_set_in_iff in h1r. unfold Bt, bt in h2. unfold Bt, bt in h1r. rewrite h2 in h1r. pose proof (h4b _ h1r) as h5. destruct h5 as [h5l h5r]. unfold Bt. unfold Bt in h0. rewrite h0. rewrite list_to_set_in_iff. pose proof (la_compat _ nml) as h6. unfold Bt in h6. rewrite h6. assumption. (* >= *) red. intros y h1. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h1. unfold Bt in h1. rewrite in_map_iff in h1. destruct h1 as [x h1]. destruct h1 as [h1l h1r]. unfold n_im. unfold Bt, bt. rewrite in_map_iff. exists (x, y). simpl. split; auto. rewrite list_to_set_in_iff. pose proof (lp_compat _ nml) as h2. unfold Bt, bt in h2. rewrite h2. unfold Bt. unfold Bt in h0. rewrite h0 in h1r. pose proof (la_compat _ nml) as h3. unfold Bt in h3. rewrite list_to_set_in_iff in h1r. rewrite h3 in h1r. apply fin_map_app_fin_map_to_fps_compat. simpl; auto. simpl. rewrite <- h1l. unfold f_no_pr. unfold times_fin_pair_map2. symmetry. apply times_fun_fin_map_to_fun_comm'. pose proof (lab_empty2 _ nml2) as h5. destruct (nil_dec' (n_lab2 F nml2)) as [h6 | h7]. destruct h5 as [h5l | h5r]. Focus 2. destruct h5r as [h5a h5b]. rewrite h5b in h1r. contradiction. destruct h5l as [h5a h5b]. unfold Bt in h5a. rewrite h5a. rewrite h5b. simpl. reflexivity. destruct h5. assumption. assumption. Qed. (*The meet of (p (i, a i)) as i ranges over I*) Definition times_fun_fin_map1 {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Bt) (a:Fin_map I J def) : Bt. pose proof (fin_map_fin_dom a) as h1. pose proof (fin_map_fin_ran a) as h2. pose proof (cart_prod_fin _ _ h1 h2) as h4. pose (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu) _ 0 h4 p) as P. refine (fin_map_times (fun_to_fin_map (fin_map_type_eq_dec a) I 0 h1 (fun i:T => (P |-> (i, (a |-> i)))))). Defined. (*The meet of (p (a j, j)) as j ranges over J*) Definition times_fun_fin_map2 {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Bt) (a:Fin_map J I def) : Bt. pose proof (fin_map_fin_dom a) as h1. pose proof (fin_map_fin_ran a) as h2. pose proof (cart_prod_fin _ _ h2 h1) as h4. pose (fun_to_fin_map (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 h4 p) as P. refine (fin_map_times (fun_to_fin_map (fin_map_type_eq_dec a) J 0 h1 (fun j:U => (P |-> (a |-> j, j))))). Defined. Definition times_fun_fin_map1_l {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pf:type_eq_dec U) (p:(T*U)->Bt) (a:Fin_map I J def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pf) _ 0 (cart_prod_fin _ _ (fin_map_fin_dom a) (fin_map_fin_ran a)) p)) := times_list (map (fun i:T => (p (i, a |-> i))) (n_la2 _ nml2)). Definition times_fun_fin_map2_l {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pf:type_eq_dec T) (p:(T*U)->Bt) (a:Fin_map J I def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod pf (fin_map_type_eq_dec a)) _ 0 (cart_prod_fin _ _ (fin_map_fin_ran a) (fin_map_fin_dom a)) p)) := times_list (map (fun j:U => (p (a |-> j, j))) (n_lb2 _ nml2)). Lemma im_sing_times_fun_fin_map1 : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pf:type_eq_dec U) (p:T*U->Bt) (a:Fin_map I J def), Im (Singleton a) (times_fun_fin_map1 pf p) = Singleton ((times_fun_fin_map1 pf p) a). intros T U I J def hdt p a. rewrite <- add_empty_sing. rewrite Im_add. rewrite image_empty. rewrite add_empty_sing. reflexivity. Qed. Lemma times_fun_fin_map1_empty : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:U) (p:(T*U)->Bt), times_fun_fin_map1 pfdu p (empty_map T U pfdt def) = 1. intros T U hdt hdu def p. unfold times_fun_fin_map1. destruct (empty_map Bt Bt ba_eq_dec 0) as [h1 h2 S h3]. apply fin_map_times_empty1. Qed. Lemma times_fun_fin_map2_empty : forall {T U:Type} (def:T) (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (p:(T*U)->Bt), times_fun_fin_map2 pfdt p (empty_map U T pfdu def) = 1. intros T U def hdt hdu p. unfold times_fun_fin_map2. destruct (empty_map Bt Bt ba_eq_dec 0) as [h1 h2 S h3]. apply fin_map_times_empty1. Qed. Lemma times_fun_fin_map1_empty1 : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:U) (p:(T*U)->Bt) (J:Ensemble U) (pf:Finite J), times_fun_fin_map1 pfdu p (empty_map1 T U pfdt def J pf) = 1. intros T U hdt hdu def p J h1. unfold times_fun_fin_map1. destruct (empty_map1 T U hdt def J h1) as [h2 h3 S h4]. apply fin_map_times_empty1. Qed. Lemma times_fun_fin_map2_empty1 : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:T) (p:(T*U)->Bt) (I:Ensemble T) (pf:Finite I), times_fun_fin_map2 pfdt p (empty_map1 U T pfdu def I pf) = 1. intros T U hdt hdu def p I h1. unfold times_fun_fin_map2. destruct (empty_map1 U T hdu def I h1) as [h2 h3 S h4]. apply fin_map_times_empty1. Qed. Lemma times_fun_fin_map1_list_compat : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Bt) (a:Fin_map I J def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod(fin_map_type_eq_dec a) pfdu) _ 0 (cart_prod_fin _ _ (fin_map_fin_dom a) (fin_map_fin_ran a)) p)), times_fun_fin_map1 pfdu p a = times_fun_fin_map1_l pfdu p a nml2. intros T U I J def hdu p a nml2. unfold times_fun_fin_map1_l. pose proof (lab_empty2 _ nml2) as h1. destruct (nil_dec' (n_lab2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) hdu) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_dom a) (fin_map_fin_ran a)) p) nml2)) as [h2 | h3]. destruct h1 as [h4 | h5]. (* h4 *) destruct h4 as [h4l h4r]. rewrite h4l. simpl. generalize dependent a. rewrite h4r. intros. pose proof (fin_map_fin_ran a) as h5. pose proof (empty_map1_compat a h5) as h6. rewrite h6. apply times_fun_fin_map1_empty1. (* h5 *) destruct h5 as [h5l h5r]. generalize dependent a. rewrite h5r. intros. pose proof (fin_map_fin_dom a) as h6. induction h6 as [|I h7 h8 x h9]. pose proof (empty_map_compat a) as h6. rewrite h6 at 1. rewrite times_fun_fin_map1_empty. pose proof (lab_empty2' _ nml2) as h7. destruct (nil_dec' (n_lab2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) hdu) (cart_prod (Empty_set T) (Empty_set U)) 0 (cart_prod_fin (Empty_set T) (Empty_set U) (fin_map_fin_dom a) (fin_map_fin_ran a)) p) nml2)) as [h8 | h9]. destruct h7 as [h7l h7r]. specialize (h7l (eq_refl _)). rewrite h7l. simpl. reflexivity. contradiction. pose proof (fin_map_fin_dom a) as h10. pose proof (Add_intro2 _ I x) as h11. pose proof (Inhabited_intro _ _ _ h11) as h12. pose proof (Inhabited_not_empty _ _ h12) as h13. pose proof (empty_map_ex2_fin _ _ h10 h13 a). contradiction. unfold times_fun_fin_map1. unfold fin_map_times. pose proof (list_to_set_finite (map (fun i : T => p (i, a |-> i)) (n_la2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) hdu) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_dom a) (fin_map_fin_ran a)) p) nml2))) as h2. pose proof (times_set_compat' _ (map (fun i : T => p (i, a |-> i)) (n_la2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) hdu) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_dom a) (fin_map_fin_ran a)) p) nml2)) h2 (eq_refl _)) as h4. rewrite <- h4. apply times_set_functional. apply Extensionality_Ensembles. red. split. (* <=*) red. intros y h5. rewrite <- list_to_set_in_iff. rewrite in_map_iff. destruct h5 as [x h5 y h6]. rewrite fun_to_fin_map_compat in h6. rewrite fun_to_fin_map_compat in h6. exists x. split. rewrite h6. reflexivity. destruct h1 as [h1a h1b]. rewrite list_to_set_in_iff. rewrite h1a. assumption. constructor. simpl. split. assumption. apply fin_map_app_in. assumption. assumption. (* >= *) red. intros y h5. rewrite <- list_to_set_in_iff in h5. rewrite in_map_iff in h5. destruct h5 as [x h5]. destruct h5 as [h5l h5r]. destruct h1 as [h1l h1r]. rewrite list_to_set_in_iff in h5r. rewrite h1l in h5r. apply Im_intro with x. assumption. rewrite fun_to_fin_map_compat. rewrite fun_to_fin_map_compat. rewrite h5l. reflexivity. constructor. simpl. split. assumption. apply fin_map_app_in. assumption. assumption. Qed. Lemma times_fun_fin_map2_list_compat : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Bt) (a:Fin_map J I def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 (cart_prod_fin _ _ (fin_map_fin_ran a) (fin_map_fin_dom a)) p)), times_fun_fin_map2 pfdt p a = times_fun_fin_map2_l pfdt p a nml2. intros T U I J def hdt p a nml2. unfold times_fun_fin_map2_l. pose proof (lab_empty2 _ nml2) as h1. destruct (nil_dec' (n_lab2 (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec a)) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_ran a) (fin_map_fin_dom a) ) p) nml2)) as [h2 | h3]. destruct h1 as [h4 | h5]. (* h5 *) Focus 2. destruct h5 as [h5l h5r]. rewrite h5l. simpl. generalize dependent a. rewrite h5r. intros. pose proof (fin_map_fin_ran a) as h5. pose proof (empty_map1_compat a h5) as h6. rewrite h6. apply times_fun_fin_map2_empty1. (* h4 *) destruct h4 as [h4l h4r]. generalize dependent a. rewrite h4r. intros. pose proof (fin_map_fin_dom a) as h6. induction h6 as [|J h7 h8 y h9]. pose proof (empty_map_compat a) as h6. rewrite h6 at 1. rewrite times_fun_fin_map2_empty. pose proof (lab_empty2' _ nml2) as h7. destruct (nil_dec' (n_lab2 (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec a)) (cart_prod (Empty_set T) (Empty_set U)) 0 (cart_prod_fin (Empty_set T) (Empty_set U) (fin_map_fin_ran a) (fin_map_fin_dom a)) p) nml2)) as [h8 | h9]. destruct h7 as [h7l h7r]. specialize (h7r (eq_refl _)). rewrite h7r. simpl. reflexivity. contradiction. pose proof (fin_map_fin_dom a) as h10. pose proof (Add_intro2 _ J y) as h11. pose proof (Inhabited_intro _ _ _ h11) as h12. pose proof (Inhabited_not_empty _ _ h12) as h13. pose proof (empty_map_ex2_fin _ _ h10 h13 a). contradiction. unfold times_fun_fin_map2. unfold fin_map_times. pose proof (list_to_set_finite (map (fun j : U => p (a |-> j, j)) (n_lb2 (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec a)) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_ran a) (fin_map_fin_dom a)) p) nml2))) as h2. pose proof (times_set_compat' _ (map (fun j : U => p (a |-> j, j)) (n_lb2 (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec a)) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_ran a) (fin_map_fin_dom a)) p) nml2)) h2 (eq_refl _)) as h4. rewrite <- h4. apply times_set_functional. apply Extensionality_Ensembles. red. split. (* <=*) red. intros x h5. rewrite <- list_to_set_in_iff. rewrite in_map_iff. destruct h5 as [y h5 x h6]. rewrite fun_to_fin_map_compat in h6. rewrite fun_to_fin_map_compat in h6. exists y. split. rewrite h6. reflexivity. destruct h1 as [h1a h1b]. rewrite list_to_set_in_iff. rewrite h1b. assumption. constructor. simpl. split. Focus 2. assumption. apply fin_map_app_in. assumption. assumption. (* >= *) red. intros z h5. rewrite <- list_to_set_in_iff in h5. rewrite in_map_iff in h5. destruct h5 as [y h5]. destruct h5 as [h5l h5r]. destruct h1 as [h1l h1r]. rewrite list_to_set_in_iff in h5r. rewrite h1r in h5r. apply Im_intro with y. assumption. rewrite fun_to_fin_map_compat. rewrite fun_to_fin_map_compat. rewrite h5l. reflexivity. constructor. simpl. split. Focus 2. assumption. apply fin_map_app_in. assumption. assumption. Qed. Definition plus_fun_fin_map1 {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Bt) (a:Fin_map I J def) : Bt. pose proof (fin_map_fin_dom a) as h1. pose proof (fin_map_fin_ran a) as h2. pose proof (cart_prod_fin _ _ h1 h2) as h4. pose (@fun_to_fin_map (T*U) Bt (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu) _ 0 h4 p) as P. refine (fin_map_plus (fun_to_fin_map (fin_map_type_eq_dec a) I 0 h1 (fun i:T => (P |-> (i, (a |-> i)))))). Defined. (*The join of (p (a j, j)) as j ranges over J*) Definition plus_fun_fin_map2 {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Bt) (a:Fin_map J I def) : Bt. pose proof (fin_map_fin_dom a) as h1. pose proof (fin_map_fin_ran a) as h2. pose proof (cart_prod_fin _ _ h2 h1) as h4. pose (@fun_to_fin_map (T*U) Bt (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 h4 p) as P. refine (fin_map_plus (fun_to_fin_map (fin_map_type_eq_dec a) J 0 h1 (fun j:U => (P |-> (a |-> j, j))))). Defined. Lemma plus_fun_fin_map1_empty : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:U) (p:(T*U)->Bt), plus_fun_fin_map1 pfdu p (empty_map T U pfdt def) = 0. intros T U hdt hdu def p. unfold plus_fun_fin_map1. destruct (empty_map Bt Bt ba_eq_dec 0) as [h1 h2 S h3]. apply fin_map_plus_empty1. Qed. Lemma plus_fun_fin_map1_empty1 : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:U) (p:(T*U)->Bt) (J:Ensemble U) (pf:Finite J), plus_fun_fin_map1 pfdu p (empty_map1 T U pfdt def J pf) = 0. intros T U hdt hdu def p J h1. unfold plus_fun_fin_map1. destruct (empty_map1 T U hdt def J h1) as [h2 h3 S h4]. apply fin_map_plus_empty1. Qed. Definition plus_fun_fin_map1_l {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Bt) (a:Fin_map I J def) (nml2:nice_map_lists2 (@fun_to_fin_map (T*U) Bt (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu) _ 0 (cart_prod_fin _ _ (fin_map_fin_dom a) (fin_map_fin_ran a)) p)) := plus_list (map (fun i:T => (p (i, a |-> i))) (n_la2 _ nml2)). Definition plus_fun_fin_map2_l {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Bt) (a:Fin_map J I def) (nml2:nice_map_lists2 (@fun_to_fin_map (T*U) Bt (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 (cart_prod_fin _ _ (fin_map_fin_ran a) (fin_map_fin_dom a)) p)) := plus_list (map (fun j:U => (p (a |-> j, j))) (n_lb2 _ nml2)). Lemma plus_fun_fin_map2_empty : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:T) (p:(T*U)->Bt), plus_fun_fin_map2 pfdt p (empty_map U T pfdu def) = 0. intros T U hdt hdu def p. unfold plus_fun_fin_map2. destruct (empty_map Bt Bt ba_eq_dec 0) as [h1 h2 S h3]. apply fin_map_plus_empty1. Qed. Lemma plus_fun_fin_map2_empty1 : forall {T U:Type} (def:T) (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (p:(T*U)->Bt) (I:Ensemble T) (pf:Finite I), plus_fun_fin_map2 pfdt p (empty_map1 U T pfdu def I pf) = 0. intros T U def hdt hdu p I h1. unfold plus_fun_fin_map2. destruct (empty_map1 U T hdu def I h1) as [h2 h3 S h4]. apply fin_map_plus_empty1. Qed. Lemma plus_fun_fin_map1_list_compat : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Bt) (a:Fin_map I J def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu )_ 0 (cart_prod_fin _ _ (fin_map_fin_dom a) (fin_map_fin_ran a)) p)), plus_fun_fin_map1 pfdu p a = plus_fun_fin_map1_l pfdu p a nml2. intros T U I J def hdu p a nml2. unfold plus_fun_fin_map1_l. pose proof (lab_empty2 _ nml2) as h1. destruct (nil_dec' (n_lab2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) hdu) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_dom a) (fin_map_fin_ran a)) p) nml2)) as [h2 | h3]. destruct h1 as [h4 | h5]. (* h4 *) destruct h4 as [h4l h4r]. unfold Bt. rewrite h4l at 1. simpl. generalize dependent a. rewrite h4r. intros. pose proof (fin_map_fin_ran a) as h5. pose proof (empty_map1_compat a h5) as h6. rewrite h6. apply plus_fun_fin_map1_empty1. (* h5 *) destruct h5 as [h5l h5r]. generalize dependent a. rewrite h5r. intros. pose proof (fin_map_fin_dom a) as h6. induction h6 as [|I h7 h8 x h9]. pose proof (empty_map_compat a) as h6. rewrite h6 at 1. rewrite plus_fun_fin_map1_empty. pose proof (lab_empty2' _ nml2) as h7. destruct (nil_dec' (n_lab2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) hdu) (cart_prod (Empty_set T) (Empty_set U)) 0 (cart_prod_fin (Empty_set T) (Empty_set U) (fin_map_fin_dom a) (fin_map_fin_ran a)) p) nml2)) as [h8 | h9]. destruct h7 as [h7l h7r]. specialize (h7l (eq_refl _)). unfold Bt. rewrite h7l at 1. simpl. reflexivity. contradiction. pose proof (fin_map_fin_dom a) as h10. pose proof (Add_intro2 _ I x) as h11. pose proof (Inhabited_intro _ _ _ h11) as h12. pose proof (Inhabited_not_empty _ _ h12) as h13. pose proof (empty_map_ex2_fin _ _ h10 h13 a). contradiction. unfold plus_fun_fin_map1. unfold fin_map_plus. pose proof (list_to_set_finite (map (fun i : T => p (i, a |-> i)) (n_la2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) hdu) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_dom a) (fin_map_fin_ran a)) p) nml2))) as h2. pose proof (plus_set_compat' _ (map (fun i : T => p (i, a |-> i)) (n_la2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) hdu) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_dom a) (fin_map_fin_ran a)) p) nml2)) h2 (eq_refl _)) as h4. unfold Bt in h4. unfold Bt. rewrite <- h4 at 1. apply plus_set_functional. apply Extensionality_Ensembles. red. split. (* <=*) red. intros y h5. unfold Bt. rewrite <- list_to_set_in_iff. rewrite in_map_iff. destruct h5 as [x h5 y h6]. rewrite fun_to_fin_map_compat in h6. rewrite fun_to_fin_map_compat in h6. exists x. split. rewrite h6. reflexivity. destruct h1 as [h1a h1b]. rewrite list_to_set_in_iff. rewrite h1a. assumption. constructor. simpl. split. assumption. apply fin_map_app_in. assumption. assumption. (* >= *) red. intros y h5. unfold Bt in h5. unfold Bt. rewrite <- list_to_set_in_iff in h5. rewrite in_map_iff in h5. destruct h5 as [x h5]. destruct h5 as [h5l h5r]. destruct h1 as [h1l h1r]. rewrite list_to_set_in_iff in h5r. rewrite h1l in h5r. apply Im_intro with x. assumption. rewrite fun_to_fin_map_compat. rewrite fun_to_fin_map_compat. rewrite h5l. reflexivity. constructor. simpl. split. assumption. apply fin_map_app_in. assumption. assumption. Qed. Lemma plus_fun_fin_map2_list_compat : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Bt) (a:Fin_map J I def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 (cart_prod_fin _ _ (fin_map_fin_ran a) (fin_map_fin_dom a)) p)), plus_fun_fin_map2 pfdt p a = plus_fun_fin_map2_l pfdt p a nml2. intros T U I J def hdt p a nml2. unfold plus_fun_fin_map2_l. pose proof (lab_empty2 _ nml2) as h1. destruct (nil_dec' (n_lab2 (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec a)) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_ran a) (fin_map_fin_dom a) ) p) nml2)) as [h2 | h3]. destruct h1 as [h4 | h5]. (* h5 *) Focus 2. destruct h5 as [h5l h5r]. unfold Bt in h5l. unfold Bt. rewrite h5l at 1. simpl. generalize dependent a. rewrite h5r. intros. pose proof (fin_map_fin_ran a) as h5. pose proof (empty_map1_compat a h5) as h6. rewrite h6. apply plus_fun_fin_map2_empty1. (* h4 *) destruct h4 as [h4l h4r]. generalize dependent a. rewrite h4r. intros. pose proof (fin_map_fin_dom a) as h6. induction h6 as [|J h7 h8 y h9]. pose proof (empty_map_compat a) as h6. rewrite h6 at 1. rewrite plus_fun_fin_map2_empty. pose proof (lab_empty2' _ nml2) as h7. destruct (nil_dec' (n_lab2 (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec a)) (cart_prod (Empty_set T) (Empty_set U)) 0 (cart_prod_fin (Empty_set T) (Empty_set U) (fin_map_fin_ran a) (fin_map_fin_dom a)) p) nml2)) as [h8 | h9]. destruct h7 as [h7l h7r]. specialize (h7r (eq_refl _)). unfold Bt in h7r. unfold Bt. rewrite h7r at 1. simpl. reflexivity. contradiction. pose proof (fin_map_fin_dom a) as h10. pose proof (Add_intro2 _ J y) as h11. pose proof (Inhabited_intro _ _ _ h11) as h12. pose proof (Inhabited_not_empty _ _ h12) as h13. pose proof (empty_map_ex2_fin _ _ h10 h13 a). contradiction. unfold times_fun_fin_map2. unfold fin_map_times. pose proof (list_to_set_finite (map (fun j : U => p (a |-> j, j)) (n_lb2 (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec a)) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_ran a) (fin_map_fin_dom a)) p) nml2))) as h2. pose proof (plus_set_compat' _ (map (fun j : U => p (a |-> j, j)) (n_lb2 (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec a)) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_ran a) (fin_map_fin_dom a)) p) nml2)) h2 (eq_refl _)) as h4. unfold Bt in h4. unfold Bt. rewrite <- h4 at 1. apply plus_set_functional. apply Extensionality_Ensembles. red. split. (* <=*) red. intros x h5. unfold Bt in h5. unfold Bt. rewrite <- list_to_set_in_iff. rewrite in_map_iff. destruct h5 as [y h5 x h6]. rewrite fun_to_fin_map_compat in h6. rewrite fun_to_fin_map_compat in h6. exists y. split. rewrite h6. reflexivity. destruct h1 as [h1a h1b]. rewrite list_to_set_in_iff. rewrite h1b. assumption. constructor. simpl. split. Focus 2. assumption. apply fin_map_app_in. assumption. assumption. (* >= *) red. intros z h5. unfold Bt in h5. unfold Bt. rewrite <- list_to_set_in_iff in h5. rewrite in_map_iff in h5. destruct h5 as [y h5]. destruct h5 as [h5l h5r]. destruct h1 as [h1l h1r]. rewrite list_to_set_in_iff in h5r. rewrite h1r in h5r. apply Im_intro with y. assumption. rewrite fun_to_fin_map_compat. rewrite fun_to_fin_map_compat. rewrite h5l. reflexivity. constructor. simpl. split. Focus 2. assumption. apply fin_map_app_in. assumption. assumption. Qed. Definition plus_times_fun_all_maps1 {T U:Type} (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:U) (pfi:Finite I) (pfj:Finite J) (p:(T*U)->Bt): Bt := let f := (@times_fun_fin_map1 T U I J def pfdu p) in let S := Full_set (Fin_map I J def) in plus_set (Im S f) (finite_image _ _ S f (finite_fin_maps _ _ def pfi pfj)). Definition plus_times_fun_all_maps2 {T U:Type} (pfdt:type_eq_dec T) (I:Ensemble T) (J:Ensemble U) (def:T) (pfi:Finite I) (pfj:Finite J) (p:(T*U)->Bt): Bt := let f:= (@times_fun_fin_map2 T U I J def pfdt p) in let S := Full_set (Fin_map J I def) in plus_set (Im S f) (finite_image _ _ S f (finite_fin_maps _ _ def pfj pfi)). Definition times_plus_fun_all_maps1 {T U:Type} (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:U) (pfi:Finite I) (pfj:Finite J) (p:T*U->Bt) : Bt := let f:= (@plus_fun_fin_map1 T U I J def pfdu p) in let S := Full_set (Fin_map I J def) in times_set (Im S f) (finite_image _ _ S f (finite_fin_maps _ _ def pfi pfj)). Definition times_plus_fun_all_maps2 {T U:Type} (pfdt:type_eq_dec T) (I:Ensemble T) (J:Ensemble U) (def:T) (pfi:Finite I) (pfj:Finite J) (p:T*U->Bt) : Bt := let f:= (@plus_fun_fin_map2 T U I J def pfdt p) in let S := Full_set (Fin_map J I def) in times_set (Im S f) (finite_image _ _ S f (finite_fin_maps _ _ def pfj pfi)). Lemma complete_dist_list_times1' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Bt), times_plus_fun1 li lj p = plus_times_all_funs1 pfdt pfdu li lj pfi pfj p. intros T U hdt hdu li lj h1 h2 p. pose proof (complete_dist_list_times) as h3. unfold times_plus_fun1. unfold times_fun. unfold plus_fun. unfold times_plus_list_of_lists in h3. specialize (h3 _ (map (fun i : T => (map (p i) lj)) li)). rewrite map_map in h3. unfold Bt. unfold Bt in h3. rewrite h3. clear h3. unfold plus_times_list_of_lists. unfold plus_times_all_funs1. unfold plus_fun. pose proof (list_to_set_finite (map (fun l' : list (Btype (Bc B)) => times_list l') (list_of_lists_seqs (map (fun i : T => map (p i) lj) li)))) as h3. pose proof (list_to_set_finite (map_in_dep (fun (l : list (T * U)) (pfinlp : In l (list_power li lj)) => times_fun_in_dep (fun (i : T) (pfini : In i li) => p i (fpl_to_fun_in_dep (in_list_power_fpl hdt hdu li lj l h1 (list_power_no_dup li lj l h1 h2 pfinlp) pfinlp) i pfini))))) as h4. pose proof (plus_set_compat' _ _ h3 (eq_refl _)) as h5. pose proof (plus_set_compat' _ _ h4 (eq_refl _)) as h6. unfold Bt, bt. unfold Bt, bt in h5. unfold Bt, bt in h6. rewrite <- h5. unfold plus_fun_in_dep. rewrite <- h6 at 1. clear h5 h6. apply plus_set_functional. apply Extensionality_Ensembles. red. split. (* <= *) red. intros b h5. unfold Bt, bt in h5. rewrite <- list_to_set_in_iff in h5. rewrite in_map_iff in h5. destruct h5 as [l h5]. destruct h5 as [h5l h5r]. unfold Bt, bt. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff at 1. unfold times_fun. pose proof (in_list_of_lists_seqs_map_map hdt hdu p _ _ h1 h2 _ h5r) as h6. destruct h6 as [lp h7]. destruct h7 as [h7 h8]. exists lp, h7. simpl. rewrite <- h5l. unfold times_fun_in_dep. simpl in h8. rewrite h8. reflexivity. (* >= *) red. intros x h5. unfold Bt, bt. unfold Bt, bt in h5. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h5. rewrite in_map_in_dep_iff in h5. rewrite in_map_iff. destruct h5 as [ls h6]. destruct h6 as [h6l h6r]. unfold times_fun in h6l. pose proof (in_map_list_of_lists_seqs_map_in_dep hdt hdu p _ _ ls h1 h2 h6l) as h7. destruct h7 as [l [h7a h7b]]. exists l. split. subst. unfold times_fun_in_dep. reflexivity. assumption. Qed. Lemma complete_dist_list_times2' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Bt) (def:T), times_plus_fun2 li lj p = plus_times_all_funs2 pfdt pfdu li lj pfi pfj p. intros T U hdt hdu li lj h1 h2 p def. pose proof (complete_dist_list_times) as h3. unfold times_plus_fun2. unfold times_fun. unfold plus_fun. unfold times_plus_list_of_lists in h3. specialize (h3 _ (map (fun j : U => map (fun i : T => p i j) li) lj)). rewrite map_map in h3. unfold Bt. unfold Bt in h3. rewrite h3. clear h3. unfold plus_times_list_of_lists. unfold plus_times_all_funs2. unfold plus_fun. pose proof (list_to_set_finite (map (fun l' : list (Btype (Bc B)) => times_list l') (list_of_lists_seqs (map (fun j : U => map (fun i : T => p i j) li) lj)))) as h3. match goal with |- _ = plus_fun_in_dep ?f => pose proof (list_to_set_finite (map_in_dep f)) as h4 end. pose proof (plus_set_compat' _ _ h3 (eq_refl _)) as h5. pose proof (plus_set_compat' _ _ h4 (eq_refl _)) as h6. unfold Bt, bt. unfold Bt, bt in h5. unfold Bt, bt in h6. rewrite <- h5. unfold plus_fun_in_dep. rewrite <- h6 at 1. clear h5 h6. apply plus_set_functional. apply Extensionality_Ensembles. red. split. (* <= *) red. intros b h5. unfold Bt, bt in h5. rewrite <- list_to_set_in_iff in h5. rewrite in_map_iff in h5. destruct h5 as [l h5]. destruct h5 as [h5l h5r]. unfold Bt, bt. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. unfold times_fun. pose proof (in_list_of_lists_seqs_map_map' hdt hdu _ _ _ h1 h2 _ def h5r) as h6. destruct h6 as [lp h7]. destruct h7 as [h7 h8]. exists lp, h7. simpl. rewrite <- h5l. simpl in h8. rewrite h8. unfold times_fun_in_dep; auto. (* >= *) red. intros x h5. unfold Bt, bt. unfold Bt, bt in h5. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h5. rewrite in_map_in_dep_iff in h5. rewrite in_map_iff. destruct h5 as [ls h6]. destruct h6 as [h6l h6r]. unfold times_fun in h6l. pose proof (in_map_in_dep_list_of_lists_seqs_map' hdt hdu p li lj ls h1 h2 h6l) as h7. destruct h7 as [l [h7a h7b]]. exists l. split. rewrite <- h6r. rewrite h7b at 1. unfold times_fun_in_dep. reflexivity. assumption. Qed. Lemma complete_dist_list_plus1' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Bt), plus_times_fun1 li lj p = times_plus_all_funs1 pfdt pfdu li lj pfi pfj p. intros T U hdt hdu li lj h1 h2 p. pose proof (complete_dist_list_plus) as h3. unfold plus_times_fun1. unfold plus_fun. unfold times_fun. unfold plus_times_list_of_lists in h3. specialize (h3 _ (map (fun i : T => (map (p i) lj)) li)). rewrite map_map in h3. unfold Bt. unfold Bt in h3. rewrite h3. clear h3. unfold times_plus_list_of_lists, times_plus_all_funs1. unfold times_fun_in_dep. pose proof (list_to_set_finite (map (fun l' : list (Btype (Bc B)) => plus_list l') (list_of_lists_seqs (map (fun i : T => map (p i) lj) li)))) as h3. pose proof (list_to_set_finite (map_in_dep (fun (l : list (T * U)) (pfinlp : In l (list_power li lj)) => plus_fun_in_dep (fun (i : T) (pfini : In i li) => p i (fpl_to_fun_in_dep (in_list_power_fpl hdt hdu li lj l h1 (list_power_no_dup li lj l h1 h2 pfinlp) pfinlp) i pfini))))) as h4. pose proof (times_set_compat' _ _ h3 (eq_refl _)) as h5. pose proof (times_set_compat' _ _ h4 (eq_refl _)) as h6. unfold Bt, bt. unfold Bt, bt in h5. unfold Bt, bt in h6. rewrite <- h5 at 1. unfold times_fun_in_dep. rewrite <- h6 at 1. clear h5 h6. apply times_set_functional. apply Extensionality_Ensembles. red. split. (* <= *) red. intros b h5. unfold Bt, bt in h5. rewrite <- list_to_set_in_iff in h5. rewrite in_map_iff in h5. destruct h5 as [l h5]. destruct h5 as [h5l h5r]. unfold Bt, bt. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff at 1. unfold times_fun. pose proof (in_list_of_lists_seqs_map_map hdt hdu p _ _ h1 h2 _ h5r) as h6. destruct h6 as [lp h7]. destruct h7 as [h7 h8]. exists lp, h7. simpl. rewrite <- h5l. unfold times_fun_in_dep. simpl in h8. rewrite h8. reflexivity. (* >= *) red. intros x h5. unfold Bt, bt. unfold Bt, bt in h5. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h5. rewrite in_map_in_dep_iff in h5. rewrite in_map_iff. destruct h5 as [ls h6]. destruct h6 as [h6l h6r]. unfold times_fun in h6l. pose proof (in_map_list_of_lists_seqs_map_in_dep hdt hdu p _ _ ls h1 h2 h6l) as h7. destruct h7 as [l [h7a h7b]]. exists l. split. subst. unfold times_fun_in_dep. reflexivity. assumption. Qed. Lemma complete_dist_list_plus2' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Bt) (def:T), plus_times_fun2 li lj p = times_plus_all_funs2 pfdt pfdu li lj pfi pfj p. intros T U hdt hdu li lj h1 h2 p def. pose proof (complete_dist_list_plus) as h3. unfold plus_times_fun2. unfold plus_fun. unfold times_fun. unfold plus_times_list_of_lists in h3. specialize (h3 _ (map (fun j : U => map (fun i : T => p i j) li) lj)). rewrite map_map in h3. unfold Bt. unfold Bt in h3. rewrite h3. clear h3. unfold times_plus_list_of_lists. unfold times_plus_all_funs2. unfold plus_fun. pose proof (list_to_set_finite (map (fun l' : list (Btype (Bc B)) => plus_list l') (list_of_lists_seqs (map (fun j : U => map (fun i : T => p i j) li) lj)))) as h3. match goal with |- _ = times_fun_in_dep ?f => pose proof (list_to_set_finite (map_in_dep f)) as h4 end. pose proof (times_set_compat' _ _ h3 (eq_refl _)) as h5. pose proof (times_set_compat' _ _ h4 (eq_refl _)) as h6. unfold Bt, bt. unfold Bt, bt in h5. unfold Bt, bt in h6. rewrite <- h5. unfold plus_fun_in_dep. rewrite <- h6 at 1. clear h5 h6. apply times_set_functional. apply Extensionality_Ensembles. red. split. (* <= *) red. intros b h5. unfold Bt, bt in h5. rewrite <- list_to_set_in_iff in h5. rewrite in_map_iff in h5. destruct h5 as [l h5]. destruct h5 as [h5l h5r]. unfold Bt, bt. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. unfold times_fun. pose proof (in_list_of_lists_seqs_map_map' hdt hdu _ _ _ h1 h2 _ def h5r) as h6. destruct h6 as [lp h7]. destruct h7 as [h7 h8]. exists lp, h7. simpl. rewrite <- h5l. simpl in h8. rewrite h8. unfold times_fun_in_dep; auto. (* >= *) red. intros x h5. unfold Bt, bt. unfold Bt, bt in h5. rewrite <- list_to_set_in_iff. rewrite <- list_to_set_in_iff in h5. rewrite in_map_in_dep_iff in h5. rewrite in_map_iff. destruct h5 as [ls h6]. destruct h6 as [h6l h6r]. unfold times_fun in h6l. pose proof (in_map_in_dep_list_of_lists_seqs_map' hdt hdu p li lj ls h1 h2 h6l) as h7. destruct h7 as [l [h7a h7b]]. exists l. split. rewrite <- h6r. rewrite h7b at 1. unfold times_fun_in_dep. reflexivity. assumption. Qed. Lemma plus_times_all_maps1_funs_compat : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:U) (pfi:Finite I) (pfj: Finite J) (li:list T) (lj:list U) (pfndpi: NoDup li) (pfndpj: NoDup lj) (p:T->U->Bt), list_to_set li = I -> list_to_set lj = J -> plus_times_fun_all_maps1 pfdu _ _ def pfi pfj (f_pr p) = plus_times_all_funs1 pfdt pfdu _ _ pfndpi pfndpj p. intros T U I J hdt hdu def h1 h2 li lj h3 h4 p h5 h6. unfold plus_times_fun_all_maps1. unfold plus_times_all_funs1. unfold plus_fun. match goal with |- _ = plus_fun_in_dep ?f => pose proof (list_to_set_finite (map_in_dep f)) as h7 end. pose proof (plus_set_compat' _ _ h7 (eq_refl _)) as h8. unfold Bt. unfold Bt in h8. rewrite <- h8 at 1. apply plus_set_functional. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h9. destruct h9 as [F h9 x]. subst. unfold Bt. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. pose proof (fin_map_ex_nice_map_lists_list_to_set (fin_map_type_eq_dec F) hdu F h3 h4) as h10. destruct h10 as [nml h11]. destruct h11 as [h11l h11r]. pose proof (in_lp _ nml) as h12. rewrite <- h11l in h12. rewrite <- h11r in h12. exists (n_lp F nml), h12. simpl. pose proof (fin_map_ex_nice_map_lists2_list_to_set_dom (fin_map_type_eq_dec F) hdu ba_eq_dec (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec F) hdu) (cart_prod (list_to_set li) (list_to_set lj)) 0 (cart_prod_fin (list_to_set li) (list_to_set lj) (fin_map_fin_dom F) (fin_map_fin_ran F)) (f_pr p)) h3 h4) as h13. destruct h13 as [nml2 h13]. destruct h13 as [h13l h13r]. pose proof (times_fun_fin_map1_list_compat hdu _ _ nml2) as h14. unfold Bt. unfold Bt in h14. rewrite h14. unfold times_fun_fin_map1_l. unfold times_fun_in_dep. f_equal. unfold Bt. unfold Bt in h13l. rewrite <- h13l at 1. symmetry. apply map_map_in_dep_ext. intros x h15. unfold f_pr. simpl. f_equal. pose proof (type_eq_dec_eq hdt (fin_map_type_eq_dec F)) as h16. rewrite h16. symmetry. apply fpl_f_compat_list_to_set. (* >= *) red. intros x h9. unfold Bt in h9. rewrite <- list_to_set_in_iff in h9. rewrite in_map_in_dep_iff in h9. destruct h9 as [lp h9]. destruct h9 as [h10 h9r]. simpl in h9r. pose proof (in_list_power_synced _ _ _ h10) as hs. pose proof (list_power_no_dup _ _ _ h3 h4 h10) as h11. pose proof (in_list_power_fpl hdt hdu _ _ _ h3 h11 h10) as h12. pose proof (fp_fpl_compat _ _ (list_to_set lp) _ _ lp h5 h6 (eq_refl _)) as h13. unfold Bt in h12. rewrite <- h13 in h12. apply Im_intro with (fin_map_intro _ _ def hdt h1 h2 _ h12). constructor. pose (f_pr p). pose proof (fin_map_ex_nice_map_lists2_dom hdt hdu ba_eq_dec (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec (fin_map_intro I J def hdt h1 h2 (list_to_set lp) h12)) hdu) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_dom (fin_map_intro I J def hdt h1 h2 (list_to_set lp) h12)) (fin_map_fin_ran (fin_map_intro I J def hdt h1 h2 (list_to_set lp) h12))) (f_pr p)) _ _ h5 h6 h3 h4) as h14. destruct h14 as [nml2 h14]. destruct h14 as [h14l h14r]. pose proof (times_fun_fin_map1_list_compat hdu (f_pr p) _ nml2) as h14. unfold Bt. unfold Bt in h14. rewrite h14. rewrite <- h9r. unfold times_fun_fin_map1_l. unfold times_fun_in_dep. f_equal. unfold Bt. unfold Bt in h14l. rewrite <- h14l at 1. symmetry. apply map_map_in_dep_ext. intros i h15. unfold f_pr. simpl. f_equal. pose proof (fin_map_ex_nice_map_lists_intro hdt hdu _ _ li lj lp def h1 h2 h10 h12 h5 h6 h3 h4) as h16. destruct h16 as [nml h16]. destruct h16 as [h16l [h16r h16f]]. pose proof (in_lp _ nml) as h17. rewrite <- h16l in h17. rewrite <- h16r in h17. rewrite <- h16f in h17. pose proof (lp_compat _ nml) as h18. rewrite <- h16f in h18. pose proof (fpl_f_compat_pseudo_list_to_set hdu _ li lj lp h10 h3 h4 i h15 h5 h6 h18) as h19. simpl in h19. rewrite <- h19. f_equal. apply proof_irrelevance. Qed. Lemma plus_times_all_maps2_funs_compat : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:T) (pfi:Finite I) (pfj: Finite J) (li:list T) (lj:list U) (pfndpi: NoDup li) (pfndpj: NoDup lj) (p:T->U->Bt), list_to_set li = I -> list_to_set lj = J -> plus_times_fun_all_maps2 pfdt _ _ def pfi pfj (f_pr p) = plus_times_all_funs2 pfdt pfdu _ _ pfndpi pfndpj p. intros T U I J hdt hdu def h1 h2 li lj h3 h4 p h5 h6. unfold plus_times_fun_all_maps2. unfold plus_times_all_funs2. unfold plus_fun. match goal with |- _ = plus_fun_in_dep ?f => pose proof (list_to_set_finite (map_in_dep f)) as h7 end. pose proof (plus_set_compat' _ _ h7 (eq_refl _)) as h8. unfold Bt. unfold Bt in h8. rewrite <- h8 at 1. apply plus_set_functional. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h9. destruct h9 as [F h9 x]. subst. unfold Bt. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. pose proof (fin_map_ex_nice_map_lists_list_to_set (fin_map_type_eq_dec F) hdt F h4 h3) as h10. destruct h10 as [nml h11]. destruct h11 as [h11l h11r]. pose proof (in_lp _ nml) as h12. rewrite <- h11l in h12. rewrite <- h11r in h12. exists (n_lp F nml), h12. simpl. pose proof (fin_map_ex_nice_map_lists2_list_to_set_dom hdt (fin_map_type_eq_dec F) ba_eq_dec (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec F)) (cart_prod (list_to_set li) (list_to_set lj)) 0 (cart_prod_fin (list_to_set li) (list_to_set lj) (fin_map_fin_ran F) (fin_map_fin_dom F)) (f_pr p)) h3 h4) as h13. destruct h13 as [nml2 h13]. destruct h13 as [h13l h13r]. pose proof (times_fun_fin_map2_list_compat hdt (f_pr p) _ nml2) as h14. unfold Bt. unfold Bt in h14. rewrite h14. unfold times_fun_fin_map2_l. unfold times_fun_in_dep. f_equal. unfold Bt. unfold Bt in h13r. rewrite <- h13r at 1. symmetry. apply map_map_in_dep_ext. intros x h15. unfold f_pr. simpl. f_equal. pose proof (type_eq_dec_eq hdu (fin_map_type_eq_dec F)) as h16. rewrite h16. symmetry. apply fpl_f_compat_list_to_set. (* >= *) red. intros x h9. unfold Bt in h9. rewrite <- list_to_set_in_iff in h9. rewrite in_map_in_dep_iff in h9. destruct h9 as [lp h9]. destruct h9 as [h10 h9r]. simpl in h9r. pose proof (in_list_power_synced _ _ _ h10) as hs. pose proof (list_power_no_dup _ _ _ h4 h3 h10) as h11. pose proof (in_list_power_fpl hdu hdt _ _ _ h4 h11 h10) as h12. pose proof (fp_fpl_compat _ _ (list_to_set lp) _ _ lp h6 h5 (eq_refl _)) as h13. unfold Bt in h12. rewrite <- h13 in h12. apply Im_intro with (fin_map_intro _ _ def hdu h2 h1 _ h12). constructor. pose (f_pr p). pose proof (fin_map_ex_nice_map_lists2_dom hdt hdu ba_eq_dec (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec (fin_map_intro J I def hdu h2 h1 (list_to_set lp) h12))) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_ran (fin_map_intro J I def hdu h2 h1 (list_to_set lp) h12)) (fin_map_fin_dom (fin_map_intro J I def hdu h2 h1 (list_to_set lp) h12))) (f_pr p)) _ _ h5 h6 h3 h4) as h14. destruct h14 as [nml2 h14]. destruct h14 as [h14l h14r]. pose proof (times_fun_fin_map2_list_compat hdt (f_pr p) _ nml2) as h14. unfold Bt. unfold Bt in h14. rewrite h14. rewrite <- h9r. unfold times_fun_fin_map2_l. unfold times_fun_in_dep. f_equal. unfold Bt. unfold Bt in h14r. rewrite <- h14r at 1. symmetry. apply map_map_in_dep_ext. intros i h15. unfold f_pr. simpl. f_equal. pose proof (fin_map_ex_nice_map_lists_intro hdu hdt _ _ lj li lp def h2 h1 h10 h12 h6 h5 h4 h3) as h16. destruct h16 as [nml h16]. destruct h16 as [h16l [h16r h16f]]. pose proof (in_lp _ nml) as h17. rewrite <- h16l in h17. rewrite <- h16r in h17. rewrite <- h16f in h17. pose proof (lp_compat _ nml) as h18. rewrite <- h16f in h18. pose proof (fpl_f_compat_pseudo_list_to_set hdt _ lj li lp h10 h4 h3 i h15 h6 h5 h18) as h19. simpl in h19. rewrite <- h19. f_equal. apply proof_irrelevance. Qed. Lemma times_plus_all_maps1_funs_compat : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:U) (pfi:Finite I) (pfj: Finite J) (li:list T) (lj:list U) (pfndpi: NoDup li) (pfndpj: NoDup lj) (p:T->U->Bt), list_to_set li = I -> list_to_set lj = J -> times_plus_fun_all_maps1 pfdu _ _ def pfi pfj (f_pr p) = times_plus_all_funs1 pfdt pfdu _ _ pfndpi pfndpj p. intros T U I J hdt hdu def h1 h2 li lj h3 h4 p h5 h6. unfold times_plus_fun_all_maps1. unfold times_plus_all_funs1. unfold times_fun. match goal with |- _ = times_fun_in_dep ?f => pose proof (list_to_set_finite (map_in_dep f)) as h7 end. pose proof (times_set_compat' _ _ h7 (eq_refl _)) as h8. unfold Bt. unfold Bt in h8. rewrite <- h8 at 1. apply times_set_functional. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h9. destruct h9 as [F h9 x]. subst. unfold Bt. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. pose proof (fin_map_ex_nice_map_lists_list_to_set (fin_map_type_eq_dec F) hdu F h3 h4) as h10. destruct h10 as [nml h11]. destruct h11 as [h11l h11r]. pose proof (in_lp _ nml) as h12. rewrite <- h11l in h12. rewrite <- h11r in h12. exists (n_lp F nml), h12. simpl. pose proof (fin_map_ex_nice_map_lists2_list_to_set_dom (fin_map_type_eq_dec F) hdu ba_eq_dec (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec F) hdu) (cart_prod (list_to_set li) (list_to_set lj)) 0 (cart_prod_fin (list_to_set li) (list_to_set lj) (fin_map_fin_dom F) (fin_map_fin_ran F)) (f_pr p)) h3 h4) as h13. destruct h13 as [nml2 h13]. destruct h13 as [h13l h13r]. pose proof (plus_fun_fin_map1_list_compat hdu _ _ nml2) as h14. unfold Bt. unfold Bt in h14. rewrite h14. unfold plus_fun_fin_map1_l. unfold plus_fun_in_dep. f_equal. unfold Bt. unfold Bt in h13l. rewrite <- h13l at 1. symmetry. apply map_map_in_dep_ext. intros x h15. unfold f_pr. simpl. f_equal. pose proof (type_eq_dec_eq hdt (fin_map_type_eq_dec F)) as h16. rewrite h16. symmetry. apply fpl_f_compat_list_to_set. (* >= *) red. intros x h9. unfold Bt in h9. rewrite <- list_to_set_in_iff in h9. rewrite in_map_in_dep_iff in h9. destruct h9 as [lp h9]. destruct h9 as [h10 h9r]. simpl in h9r. pose proof (in_list_power_synced _ _ _ h10) as hs. pose proof (list_power_no_dup _ _ _ h3 h4 h10) as h11. pose proof (in_list_power_fpl hdt hdu _ _ _ h3 h11 h10) as h12. pose proof (fp_fpl_compat _ _ (list_to_set lp) _ _ lp h5 h6 (eq_refl _)) as h13. unfold Bt in h12. rewrite <- h13 in h12. apply Im_intro with (fin_map_intro _ _ def hdt h1 h2 _ h12). constructor. pose (f_pr p). pose proof (fin_map_ex_nice_map_lists2_dom hdt hdu ba_eq_dec (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec (fin_map_intro I J def hdt h1 h2 (list_to_set lp) h12)) hdu) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_dom (fin_map_intro I J def hdt h1 h2 (list_to_set lp) h12)) (fin_map_fin_ran (fin_map_intro I J def hdt h1 h2 (list_to_set lp) h12))) (f_pr p)) _ _ h5 h6 h3 h4) as h14. destruct h14 as [nml2 h14]. destruct h14 as [h14l h14r]. pose proof (plus_fun_fin_map1_list_compat hdu (f_pr p) _ nml2) as h14. unfold Bt. unfold Bt in h14. rewrite h14. rewrite <- h9r. unfold plus_fun_fin_map1_l. unfold plus_fun_in_dep. f_equal. unfold Bt. unfold Bt in h14l. rewrite <- h14l at 1. symmetry. apply map_map_in_dep_ext. intros i h15. unfold f_pr. simpl. f_equal. pose proof (fin_map_ex_nice_map_lists_intro hdt hdu _ _ li lj lp def h1 h2 h10 h12 h5 h6 h3 h4) as h16. destruct h16 as [nml h16]. destruct h16 as [h16l [h16r h16f]]. pose proof (in_lp _ nml) as h17. rewrite <- h16l in h17. rewrite <- h16r in h17. rewrite <- h16f in h17. pose proof (lp_compat _ nml) as h18. rewrite <- h16f in h18. pose proof (fpl_f_compat_pseudo_list_to_set hdu _ li lj lp h10 h3 h4 i h15 h5 h6 h18) as h19. simpl in h19. rewrite <- h19. f_equal. apply proof_irrelevance. Qed. Lemma times_plus_all_maps2_funs_compat : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:T) (pfi:Finite I) (pfj: Finite J) (li:list T) (lj:list U) (pfndpi: NoDup li) (pfndpj: NoDup lj) (p:T->U->Bt), list_to_set li = I -> list_to_set lj = J -> times_plus_fun_all_maps2 pfdt _ _ def pfi pfj (f_pr p) = times_plus_all_funs2 pfdt pfdu _ _ pfndpi pfndpj p. intros T U I J hdt hdu def h1 h2 li lj h3 h4 p h5 h6. unfold times_plus_fun_all_maps2. unfold times_plus_all_funs2. unfold times_fun. match goal with |- _ = times_fun_in_dep ?f => pose proof (list_to_set_finite (map_in_dep f)) as h7 end. pose proof (times_set_compat' _ _ h7 (eq_refl _)) as h8. unfold Bt. unfold Bt in h8. rewrite <- h8 at 1. apply times_set_functional. apply Extensionality_Ensembles. red. split. (* <= *) red. intros y h9. destruct h9 as [F h9 x]. subst. unfold Bt. rewrite <- list_to_set_in_iff. rewrite in_map_in_dep_iff. pose proof (fin_map_ex_nice_map_lists_list_to_set (fin_map_type_eq_dec F) hdt F h4 h3) as h10. destruct h10 as [nml h11]. destruct h11 as [h11l h11r]. pose proof (in_lp _ nml) as h12. rewrite <- h11l in h12. rewrite <- h11r in h12. exists (n_lp F nml), h12. simpl. pose proof (fin_map_ex_nice_map_lists2_list_to_set_dom hdt (fin_map_type_eq_dec F) ba_eq_dec (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec F)) (cart_prod (list_to_set li) (list_to_set lj)) 0 (cart_prod_fin (list_to_set li) (list_to_set lj) (fin_map_fin_ran F) (fin_map_fin_dom F)) (f_pr p)) h3 h4) as h13. destruct h13 as [nml2 h13]. destruct h13 as [h13l h13r]. pose proof (plus_fun_fin_map2_list_compat hdt (f_pr p) _ nml2) as h14. unfold Bt. unfold Bt in h14. rewrite h14. unfold plus_fun_fin_map2_l. unfold plus_fun_in_dep. f_equal. unfold Bt. unfold Bt in h13r. rewrite <- h13r at 1. symmetry. apply map_map_in_dep_ext. intros x h15. unfold f_pr. simpl. f_equal. pose proof (type_eq_dec_eq hdu (fin_map_type_eq_dec F)) as h16. rewrite h16. symmetry. apply fpl_f_compat_list_to_set. (* >= *) red. intros x h9. unfold Bt in h9. rewrite <- list_to_set_in_iff in h9. rewrite in_map_in_dep_iff in h9. destruct h9 as [lp h9]. destruct h9 as [h10 h9r]. simpl in h9r. pose proof (in_list_power_synced _ _ _ h10) as hs. pose proof (list_power_no_dup _ _ _ h4 h3 h10) as h11. pose proof (in_list_power_fpl hdu hdt _ _ _ h4 h11 h10) as h12. pose proof (fp_fpl_compat _ _ (list_to_set lp) _ _ lp h6 h5 (eq_refl _)) as h13. unfold Bt in h12. rewrite <- h13 in h12. apply Im_intro with (fin_map_intro _ _ def hdu h2 h1 _ h12). constructor. pose (f_pr p). pose proof (fin_map_ex_nice_map_lists2_dom hdt hdu ba_eq_dec (fun_to_fin_map (impl_type_eq_dec_prod hdt (fin_map_type_eq_dec (fin_map_intro J I def hdu h2 h1 (list_to_set lp) h12))) (cart_prod I J) 0 (cart_prod_fin I J (fin_map_fin_ran (fin_map_intro J I def hdu h2 h1 (list_to_set lp) h12)) (fin_map_fin_dom (fin_map_intro J I def hdu h2 h1 (list_to_set lp) h12))) (f_pr p)) _ _ h5 h6 h3 h4) as h14. destruct h14 as [nml2 h14]. destruct h14 as [h14l h14r]. pose proof (plus_fun_fin_map2_list_compat hdt (f_pr p) _ nml2) as h14. unfold Bt. unfold Bt in h14. rewrite h14. rewrite <- h9r. unfold plus_fun_fin_map2_l. unfold plus_fun_in_dep. f_equal. unfold Bt. unfold Bt in h14r. rewrite <- h14r at 1. symmetry. apply map_map_in_dep_ext. intros i h15. unfold f_pr. simpl. f_equal. pose proof (fin_map_ex_nice_map_lists_intro hdu hdt _ _ lj li lp def h2 h1 h10 h12 h6 h5 h4 h3) as h16. destruct h16 as [nml h16]. destruct h16 as [h16l [h16r h16f]]. pose proof (in_lp _ nml) as h17. rewrite <- h16l in h17. rewrite <- h16r in h17. rewrite <- h16f in h17. pose proof (lp_compat _ nml) as h18. rewrite <- h16f in h18. pose proof (fpl_f_compat_pseudo_list_to_set hdt _ lj li lp h10 h4 h3 i h15 h6 h5 h18) as h19. simpl in h19. rewrite <- h19. f_equal. apply proof_irrelevance. Qed. Lemma times_plus_fun1_fin_map_to_fun_undoes_fun_to_fin_map : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (li:list T) (lj:list U) (pf:Finite (cart_prod I J)) (p:T*U->Bt), list_to_set li = I -> list_to_set lj = J -> let pfdtu := impl_type_eq_dec_prod pfdt pfdu in times_plus_fun1 li lj (f_no_pr (fin_map_to_fun (fun_to_fin_map pfdtu (cart_prod I J) 0 pf p))) = times_plus_fun1 li lj (f_no_pr p). intros T U hdt hdu I J li lj pf p h1 h2. simpl. unfold times_plus_fun1. unfold times_fun. f_equal. apply map_ext_in. intros x h3. unfold f_no_pr. simpl. unfold plus_fun. f_equal. apply map_ext_in. intros j h4. apply fin_map_to_fun_undoes_fun_to_fin_map. constructor. simpl. split. rewrite <- h1. rewrite <- list_to_set_in_iff. assumption. rewrite <- h2. rewrite <- list_to_set_in_iff. assumption. Qed. Lemma times_plus_fun2_fin_map_to_fun_undoes_fun_to_fin_map : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (li:list T) (lj:list U) (pf:Finite (cart_prod I J)) (p:T*U->Bt), list_to_set li = I -> list_to_set lj = J -> let pfdtu := impl_type_eq_dec_prod pfdt pfdu in times_plus_fun2 li lj (f_no_pr (fin_map_to_fun (fun_to_fin_map pfdtu (cart_prod I J) 0 pf p))) = times_plus_fun2 li lj (f_no_pr p). intros T U hdt hdu I J li lj pf p h1 h2. unfold times_plus_fun2. unfold times_fun. f_equal. apply map_ext_in. intros x h3. unfold f_no_pr. simpl. unfold plus_fun. f_equal. apply map_ext_in. intros j h4. apply fin_map_to_fun_undoes_fun_to_fin_map. constructor. simpl. split. rewrite <- h1. rewrite <- list_to_set_in_iff. assumption. rewrite <- h2. rewrite <- list_to_set_in_iff. assumption. Qed. Lemma plus_times_fun1_fin_map_to_fun_undoes_fun_to_fin_map : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (li:list T) (lj:list U) (pf:Finite (cart_prod I J)) (p:T*U->Bt), list_to_set li = I -> list_to_set lj = J -> let pfdtu := impl_type_eq_dec_prod pfdt pfdu in plus_times_fun1 li lj (f_no_pr (fin_map_to_fun (fun_to_fin_map pfdtu (cart_prod I J) 0 pf p))) = plus_times_fun1 li lj (f_no_pr p). intros T U hdt hdu I J li lj pf p h1 h2. unfold plus_times_fun1. unfold plus_fun. f_equal. apply map_ext_in. intros x h3. unfold f_no_pr. simpl. unfold times_fun. f_equal. apply map_ext_in. intros j h4. apply fin_map_to_fun_undoes_fun_to_fin_map. constructor. simpl. split. rewrite <- h1. rewrite <- list_to_set_in_iff. assumption. rewrite <- h2. rewrite <- list_to_set_in_iff. assumption. Qed. Lemma plus_times_fun2_fin_map_to_fun_undoes_fun_to_fin_map : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (li:list T) (lj:list U) (pf:Finite (cart_prod I J)) (p:T*U->Bt), list_to_set li = I -> list_to_set lj = J -> let pfdtu := impl_type_eq_dec_prod pfdt pfdu in plus_times_fun2 li lj (f_no_pr (fin_map_to_fun (fun_to_fin_map pfdtu (cart_prod I J) 0 pf p))) = plus_times_fun2 li lj (f_no_pr p). intros T U hdt hdu I J li lj pf p h1 h2. unfold plus_times_fun2. unfold plus_fun. f_equal. apply map_ext_in. intros x h3. unfold f_no_pr. simpl. unfold times_fun. f_equal. apply map_ext_in. intros j h4. apply fin_map_to_fun_undoes_fun_to_fin_map. constructor. simpl. split. rewrite <- h1. rewrite <- list_to_set_in_iff. assumption. rewrite <- h2. rewrite <- list_to_set_in_iff. assumption. Qed. Lemma complete_dist_times_plus1 : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:U) (pfi:Finite I) (pfj:Finite J) (p:T*U->Bt), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in times_plus_fin_pair_map1 pfdt pfi (fun_to_fin_map pfdtu _ 0 (cart_prod_fin _ _ pfi pfj) p) = plus_times_fun_all_maps1 pfdu _ _ def pfi pfj p. intros T U hdt hdu I J def h1 h2 p hdtu. pose proof fin_map_ex_nice_map_lists2 hdt hdu ba_eq_dec (fun_to_fin_map hdtu (cart_prod I J) 0 (cart_prod_fin I J h1 h2) p) as h3. destruct h3 as [nml2 h4]. clear h4. pose proof (la_lb_compat2 _ nml2) as h3. specialize (h3 h1 h2). destruct h3 as [h3l h3r]. pose proof (nda2 _ nml2) as h4. pose proof (ndb2 _ nml2) as h5. rewrite (f_pr_f_no_pr p). rewrite (plus_times_all_maps1_funs_compat hdt hdu def h1 h2 _ _ h4 h5 (f_no_pr p) h3l h3r). rewrite <- complete_dist_list_times1'. pose proof (fin_map_ex_nice_map_lists_list_to_set_dom hdt ba_eq_dec (plus_fin_pair_map1 hdt h1 (fun_to_fin_map hdtu (cart_prod I J) 0 (cart_prod_fin I J h1 h2) p)) _ h3l h4) as h6. destruct h6 as [nml h6]. pose proof (times_plus_fin_pair_map1_list_compat' hdt _ _ nml2 nml h6) as h7. pose proof (times_plus_fun1_fin_map_to_fun_undoes_fun_to_fin_map hdt hdu _ _ _ _ (cart_prod_fin I J h1 h2) p h3l h3r) as ht. simpl in ht. unfold Bt, hdtu in ht, h7. rewrite ht in h7 at 1. unfold Bt, bt, hdtu. unfold Bt, bt in h7. rewrite <- h7 at 1. unfold Bt. unfold Bt in nml. unfold Bt in h2. pose proof (times_plus_fin_pair_map1_list_compat hdt _ _ nml) as h8. unfold Bt, bt, hdtu in h8. rewrite <- h8 at 1. unfold Bt, bt. pose proof (f_pr_f_no_pr p) as h9. unfold Bt, bt in h9. rewrite <- h9. reflexivity. Qed. Lemma complete_dist_times_plus2 : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:T) (pfi:Finite I) (pfj:Finite J) (p:T*U->Bt), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in times_plus_fin_pair_map2 pfdu pfj (fun_to_fin_map pfdtu _ 0 (cart_prod_fin _ _ pfi pfj) p) = plus_times_fun_all_maps2 pfdt _ _ def pfi pfj p. intros T U hdt hdu I J def h1 h2 p hdtu. pose proof fin_map_ex_nice_map_lists2 hdt hdu ba_eq_dec (fun_to_fin_map hdtu (cart_prod I J) 0 (cart_prod_fin I J h1 h2) p) as h3. destruct h3 as [nml2 h4]. clear h4. pose proof (la_lb_compat2 _ nml2) as h3. specialize (h3 h1 h2). destruct h3 as [h3l h3r]. pose proof (nda2 _ nml2) as h4. pose proof (ndb2 _ nml2) as h5. rewrite (f_pr_f_no_pr p). rewrite (plus_times_all_maps2_funs_compat hdt hdu def h1 h2 _ _ h4 h5 (f_no_pr p) h3l h3r). rewrite <- complete_dist_list_times2'. pose proof (fin_map_ex_nice_map_lists_list_to_set_dom hdu ba_eq_dec (plus_fin_pair_map2 hdu h2 (fun_to_fin_map hdtu (cart_prod I J) 0 (cart_prod_fin I J h1 h2) p)) _ h3r h5) as h6. destruct h6 as [nml h6]. pose proof (times_plus_fin_pair_map2_list_compat' hdu _ _ nml2 nml h6) as h7. pose proof (times_plus_fun2_fin_map_to_fun_undoes_fun_to_fin_map hdt hdu _ _ _ _ (cart_prod_fin I J h1 h2) p h3l h3r) as ht. simpl in ht. unfold Bt, hdtu in ht, h7. rewrite ht in h7 at 1. unfold Bt, bt, hdtu. unfold Bt, bt in h7. rewrite <- h7 at 1. unfold Bt. unfold Bt in nml. unfold Bt in h2. pose proof (times_plus_fin_pair_map2_list_compat hdu _ _ nml) as h8. unfold Bt, bt, hdtu in h8. rewrite <- h8 at 1. unfold Bt, bt. pose proof (f_pr_f_no_pr p) as h9. unfold Bt, bt in h9. rewrite <- h9. reflexivity. refine def. Qed. Lemma complete_dist_plus_times1 : forall {T U:Type} (I:Ensemble T) (J:Ensemble U) (def:U) (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (pfi:Finite I) (pfj:Finite J) (p:T*U->Bt), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in plus_times_fin_pair_map1 pfdt pfi (fun_to_fin_map pfdtu _ 0 (cart_prod_fin _ _ pfi pfj) p) = times_plus_fun_all_maps1 pfdu _ _ def pfi pfj p. intros T U I J def hdt hdu h1 h2 p hdtu. pose proof (fin_map_ex_nice_map_lists2 hdt hdu ba_eq_dec (fun_to_fin_map hdtu (cart_prod I J) 0 (cart_prod_fin I J h1 h2) p)) as h3. destruct h3 as [nml2 h4]. clear h4. pose proof (la_lb_compat2 _ nml2) as h3. specialize (h3 h1 h2). destruct h3 as [h3l h3r]. pose proof (nda2 _ nml2) as h4. pose proof (ndb2 _ nml2) as h5. rewrite (f_pr_f_no_pr p). rewrite (times_plus_all_maps1_funs_compat hdt hdu def h1 h2 _ _ h4 h5 (f_no_pr p) h3l h3r). rewrite <- complete_dist_list_plus1'. pose proof (fin_map_ex_nice_map_lists_list_to_set_dom hdt ba_eq_dec (times_fin_pair_map1 hdt h1 (fun_to_fin_map hdtu (cart_prod I J) 0 (cart_prod_fin I J h1 h2) p)) _ h3l h4) as h6. destruct h6 as [nml h6]. pose proof (plus_times_fin_pair_map1_list_compat' hdt _ _ nml2 nml h6) as h7. pose proof (plus_times_fun1_fin_map_to_fun_undoes_fun_to_fin_map hdt hdu _ _ _ _ (cart_prod_fin I J h1 h2) p h3l h3r) as ht. simpl in ht. unfold Bt, hdtu in ht, h7. rewrite ht in h7 at 1. unfold Bt, bt, hdtu. unfold Bt, bt in h7. rewrite <- h7 at 1. unfold Bt. unfold Bt in nml. unfold Bt in h2. pose proof (plus_times_fin_pair_map1_list_compat hdt _ _ nml) as h8. unfold Bt, bt, hdtu in h8. rewrite <- h8 at 1. unfold Bt, bt. pose proof (f_pr_f_no_pr p) as h9. unfold Bt, bt in h9. rewrite <- h9. reflexivity. Qed. Lemma complete_dist_plus_times2 : forall {T U:Type} (I:Ensemble T) (J:Ensemble U) (def:T) (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (pfi:Finite I) (pfj:Finite J) (p:T*U->Bt), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in plus_times_fin_pair_map2 pfdu pfj (fun_to_fin_map pfdtu _ 0 (cart_prod_fin _ _ pfi pfj) p) = times_plus_fun_all_maps2 pfdt _ _ def pfi pfj p. intros T U I J def hdt hdu h1 h2 p hdtu. pose proof fin_map_ex_nice_map_lists2 hdt hdu ba_eq_dec (fun_to_fin_map hdtu (cart_prod I J) 0 (cart_prod_fin I J h1 h2) p) as h3. destruct h3 as [nml2 h4]. clear h4. pose proof (la_lb_compat2 _ nml2) as h3. specialize (h3 h1 h2). destruct h3 as [h3l h3r]. pose proof (nda2 _ nml2) as h4. pose proof (ndb2 _ nml2) as h5. rewrite (f_pr_f_no_pr p). rewrite (times_plus_all_maps2_funs_compat hdt hdu def h1 h2 _ _ h4 h5 (f_no_pr p) h3l h3r). rewrite <- complete_dist_list_plus2'. pose proof (fin_map_ex_nice_map_lists_list_to_set_dom hdu ba_eq_dec (times_fin_pair_map2 hdu h2 (fun_to_fin_map hdtu (cart_prod I J) 0 (cart_prod_fin I J h1 h2) p)) _ h3r h5) as h6. destruct h6 as [nml h6]. pose proof (plus_times_fin_pair_map2_list_compat' hdu _ _ nml2 nml h6) as h7. pose proof (plus_times_fun2_fin_map_to_fun_undoes_fun_to_fin_map hdt hdu _ _ _ _ (cart_prod_fin I J h1 h2) p h3l h3r) as ht. simpl in ht. unfold Bt, hdtu in ht, h7. rewrite ht in h7 at 1. unfold Bt, bt, hdtu. unfold Bt, bt in h7. rewrite <- h7 at 1. unfold Bt. unfold Bt in nml. unfold Bt in h2. pose proof (plus_times_fin_pair_map2_list_compat hdu _ _ nml) as h8. unfold Bt, bt, hdtu in h8. rewrite <- h8 at 1. unfold Bt, bt. pose proof (f_pr_f_no_pr p) as h9. unfold Bt, bt in h9. rewrite <- h9. reflexivity. refine def. Qed. End SetOperations. Arguments prod_list_dup_eq [B] _ _ _. Arguments sum_list_dup_eq [B] _ _ _. Arguments prod_preserves_list_singularize [B] _. Arguments sum_preserves_list_singularize [B] _. Arguments times_sing_preserves_new_head [B] _ _ _. Arguments plus_sing_preserves_new_head [B] _ _ _. Arguments times_list_sing_cons [B] _ _ _. Arguments plus_list_sing_cons [B] _ _ _. Arguments list_to_sets_eq_times_sing_eq [B] _ _ _ _ _. Arguments list_to_sets_eq_plus_sing_eq [B] _ _ _ _ _. Arguments list_to_sets_eq_times_eq [B] _ _ _ _ _. Arguments list_to_sets_eq_plus_eq [B] _ _ _ _ _. Arguments times_list_unq [B] _ _. Arguments plus_list_unq [B] _ _. Arguments times_set [B] _ _. Arguments plus_set [B] _ _. Arguments times_set_compat [B] _ _. Arguments times_set_compat' [B] _ _ _ _. Arguments plus_set_compat [B] _ _. Arguments plus_set_compat' [B] _ _ _ _. Arguments times_set_functional [B] _ _ _ _ _. Arguments plus_set_functional [B] _ _ _ _ _. Arguments times_set_empty [B]. Arguments times_set_empty' [B] _. Arguments plus_set_empty [B]. Arguments plus_set_empty' [B] _. Arguments times_set_add [B] _ _ _. Arguments times_set_add' [B] _ _ _ _. Arguments plus_set_add [B] _ _ _. Arguments plus_set_add' [B] _ _ _ _. Arguments times_set_sing [B] _. Arguments times_set_sing' [B] _ _. Arguments times_set_one_or [B] _ _ _. Arguments plus_set_sing [B] _. Arguments plus_set_sing' [B] _ _. Arguments plus_set_zero_or [B] _ _ _. Arguments times_set_couple [B] _ _. Arguments times_set_couple' [B] _ _ _. Arguments plus_set_couple [B] _ _. Arguments plus_set_couple' [B] _ _ _. Arguments le_times_set [B] _ _ _ _. Arguments le_plus_set [B] _ _ _ _. Arguments inf_times_set [B] _ _. Arguments sup_plus_set [B] _ _. Arguments times_set_union [B] _ _ _ _. Arguments times_set_union' [B] _ _ _ _ _. Arguments times_set_inc_le [B] _ _ _ _ _. Arguments plus_set_union [B] _ _ _ _. Arguments plus_set_union' [B] _ _ _ _ _. Arguments plus_set_inc_le [B] _ _ _ _ _. Arguments plus_set_im_add [B] _ _ _ _. Arguments plus_set_im_add' [B] _ _ _ _ _ _. Arguments times_set_im_add [B] _ _ _ _. Arguments times_set_im_add' [B] _ _ _ _ _ _. Arguments dist_set_plus1 [B] _ _ _. Arguments dist_set_plus1' [B] _ _ _ _. Arguments dist_set_times1 [B] _ _ _. Arguments dist_set_times1' [B] _ _ _ _. Arguments dist_set_plus2 [B] _ _ _ _. Arguments dist_set_plus2' [B] _ _ _ _ _. Arguments dist_set_times2 [B] _ _ _ _. Arguments dist_set_times2' [B] _ _ _ _ _. Arguments plus_fin_pair_map1 [B] [T] [U] [C] [D] [E] _ _ _. Arguments plus_fin_pair_map2 [B] [T] [U] [C] [D] [E] _ _ _. Arguments times_fin_pair_map1 [B] [T] [U] [C] [D] [E] _ _ _. Arguments times_fin_pair_map2 [B] [T] [U] [C] [D] [E] _ _ _. Arguments plus_fin_pair_map1_list_compat [B] [T] [U] [C] [D] [E] _ _ _ _. Arguments plus_fin_pair_map2_list_compat [B] [T] [U] [C] [D] [E] _ _ _ _. Arguments times_fin_pair_map1_list_compat [B] [T] [U] [C] [D] [E] _ _ _ _. Arguments times_fin_pair_map2_list_compat [B] [T] [U] [C] [D] [E] _ _ _ _. Arguments plus_fin_pair_map2_functional [B] [T] [U] [C] [D] [E] _ _ _ _ _. Arguments fin_map_times [B] [T] [A] [C] _. Arguments fin_map_times_list_compat [B] [T] [A] [C] _ _. Arguments fin_map_times_empty1 [B] [T] [C] _. Arguments fin_map_eq_times [B] [T] _ _ _ _ _ _. Arguments plus_fin_pair_map2_cart_empty_eq1 [B] [T] [U] _ _ _ _ _ _. Arguments plus_fin_pair_map2_cart_empty [B] [T] [U] _ _ _ _ _ _ _. Arguments plus_fin_pair_map2_cart_empty_eq2 [B] [T] [U] _ _ _ _ _ _. Arguments fin_map_times_sing [B] [T] _ _ _ _ _. Arguments fin_map_times_cart_empty11 [B] [T] [U] _ _ _ _ _ _ _. Arguments fin_map_times_cart_empty21 [B] [T] [U] _ _ _ _ _ _. Arguments fin_map_times_empty2 [B] [T] _ _. Arguments fin_map_plus [B] [T] [A] [C] _. Arguments fin_map_plus_list_compat [B] [T] [A] [C] _ _. Arguments fin_map_plus_empty1 [B] [T] [C] _. Arguments fin_map_plus_empty2 [B] [T] [A] _. Arguments fin_map_eq_plus [B] [T] _ _ _ _ _ _. Arguments times_plus_fin_pair_map1 [B] [T] [U] [C] [D] [E] _ _ _. Arguments times_plus_fin_pair_map2 [B] [T] [U] [C] [D] [E] _ _ _. Arguments times_plus_fin_pair_map2 [B] [T] [U] [C] [D] [E] _ _ _. Arguments plus_times_fin_pair_map1 [B] [T] [U] [C] [D] [E] _ _ _. Arguments plus_times_fin_pair_map2 [B] [T] [U] [C] [D] [E] _ _ _. Arguments times_plus_fin_pair_map1_list_compat [B] [T] [U] [C] [D] [E] _ _ _ _. Arguments plus_fun_fin_map_to_fun_comm [B] [T] [U] [C] [D] [E] _ _ _ _ _ _ _. Arguments plus_fun_fin_map_to_fun_comm' [B] [T] [U] [C] [D] [E] _ _ _ _ _ _ _. Arguments times_plus_fin_pair_map1_list_compat' [B] [T] [U] [C] [D] [E] _ _ _ _ _ _. Arguments times_plus_fin_pair_map2_list_compat [B] [T] [U] [C] [D] [E] _ _ _ _. Arguments times_plus_fin_pair_map2_list_compat' [B] [T] [U] [C] [D] [E] _ _ _ _ _ _. Arguments plus_times_fin_pair_map1_list_compat [B] [T] [U] [C] [D] [E] _ _ _ _. Arguments times_fun_fin_map_to_fun_comm [B] [T] [U] [C] [D] [E] _ _ _ _ _ _ _. Arguments plus_times_fin_pair_map1_list_compat' [B] [T] [U] [C] [D] [E] _ _ _ _ _ _. Arguments plus_times_fin_pair_map2_list_compat [B] [T] [U] [C] [D] [E] _ _ _ _. Arguments times_fun_fin_map_to_fun_comm' [B] [T] [U] [C] [D] [E] _ _ _ _ _ _ _. Arguments plus_times_fin_pair_map2_list_compat' [B] [T] [U] [C] [D] [E] _ _ _ _ _ _. Arguments times_fun_fin_map1 [B] [T] [U] [I] [J] [def] _ _ _. Arguments times_fun_fin_map2 [B] [T] [U] [I] [J] [def] _ _ _. Arguments times_fun_fin_map1_l [B] [T] [U] [I] [J] [def] _ _ _ _. Arguments times_fun_fin_map2_l [B] [T] [U] [I] [J] [def] _ _ _ _. Arguments im_sing_times_fun_fin_map1 [B] [T] [U] [I] [J] [def] _ _ _. Arguments times_fun_fin_map1_empty [B] [T] [U] _ _ _ _. Arguments times_fun_fin_map2_empty [B] [T] [U] _ _ _ _. Arguments times_fun_fin_map1_empty1 [B] [T] [U] _ _ _ _ _ _. Arguments times_fun_fin_map2_empty1 [B] [T] [U] _ _ _ _ _ _. Arguments times_fun_fin_map1_list_compat [B] [T] [U] [I] [J] [def] _ _ _ _. Arguments times_fun_fin_map2_list_compat [B] [T] [U] [I] [J] [def] _ _ _ _. Arguments plus_fun_fin_map1 [B] [T] [U] [I] [J] [def] _ _ _. Arguments plus_fun_fin_map2 [B] [T] [U] [I] [J] [def] _ _ _. Arguments plus_fun_fin_map1_empty [B] [T] [U] _ _ _ _. Arguments plus_fun_fin_map1_empty1 [B] [T] [U] _ _ _ _ _ _. Arguments plus_fun_fin_map1_l [B] [T] [U] [I] [J] [def] _ _ _ _. Arguments plus_fun_fin_map2_l [B] [T] [U] [I] [J] [def] _ _ _ _. Arguments plus_fun_fin_map2_empty [B] [T] [U] _ _ _ _. Arguments plus_fun_fin_map2_empty1 [B] [T] [U] _ _ _ _ _ _. Arguments plus_fun_fin_map1_list_compat [B] [T] [U] [I] [J] [def] _ _ _ _. Arguments plus_fun_fin_map2_list_compat [B] [T] [U] [I] [J] [def] _ _ _ _. Arguments plus_times_fun_all_maps1 [B] [T] [U] _ _ _ _ _ _ _. Arguments plus_times_fun_all_maps2 [B] [T] [U] _ _ _ _ _ _ _. Arguments times_plus_fun_all_maps1 [B] [T] [U] _ _ _ _ _ _ _. Arguments times_plus_fun_all_maps2 [B] [T] [U] _ _ _ _ _ _ _. Arguments complete_dist_list_times1' [B] [T] [U] _ _ _ _ _ _ _. Arguments complete_dist_list_times2' [B] [T] [U] _ _ _ _ _ _ _ _. Arguments complete_dist_list_plus1' [B] [T] [U] _ _ _ _ _ _ _. Arguments complete_dist_list_plus2' [B] [T] [U] _ _ _ _ _ _ _ _. Arguments plus_times_all_maps1_funs_compat [B] [T] [U] [I] [J] _ _ _ _ _ _ _ _ _ _ _ _. Arguments plus_times_all_maps2_funs_compat [B] [T] [U] _ _ _ _ _ _ _ _ _ _ _ _ _ _. Arguments times_plus_all_maps1_funs_compat [B] [T] [U] _ _ _ _ _ _ _ _ _ _ _ _ _ _. Arguments times_plus_all_maps2_funs_compat [B] [T] [U] _ _ _ _ _ _ _ _ _ _ _ _ _ _. Arguments times_plus_fun1_fin_map_to_fun_undoes_fun_to_fin_map [B] [T] [U] _ _ _ _ _ _ _ _ _ _. Arguments times_plus_fun2_fin_map_to_fun_undoes_fun_to_fin_map [B] [T] [U] _ _ _ _ _ _ _ _ _ _. Arguments plus_times_fun1_fin_map_to_fun_undoes_fun_to_fin_map [B] [T] [U] _ _ _ _ _ _ _ _ _ _. Arguments plus_times_fun2_fin_map_to_fun_undoes_fun_to_fin_map [B] [T] [U] _ _ _ _ _ _ _ _ _ _. Arguments complete_dist_times_plus1 [B] [T] [U] _ _ _ _ _ _ _ _. Arguments complete_dist_times_plus2 [B] [T] [U] _ _ _ _ _ _ _ _. Arguments complete_dist_plus_times1 [B] [T] [U] _ _ _ _ _ _ _ _. Arguments complete_dist_plus_times2 [B] [T] [U] _ _ _ _ _ _ _ _. (*Versions of all of the above theorems and definitions with [Bool_Alg_p]s instead of [Bool_Alg]s.*) Section ParametricAnalogues. Variable T':Type. Variable Bp:Bool_Alg_p T'. Let Btp := btp Bp. Definition ba_conv_fun_in_dep {T:Type} {l:list T} (f:fun_in_dep l Btp) : fun_in_dep l (bt (ba_conv Bp)). apply f. Defined. Lemma ba_conv_fun_in_dep_eq : forall {T:Type} (l:list T) (f:fun_in_dep l Btp), ba_conv_fun_in_dep f = (fun (x:T) (pf:In x l) => ba_conv_elt (f x pf)). intros T l f. unfold ba_conv_fun_in_dep. apply fun_in_dep_ext. intros; auto. Qed. Fixpoint times_list_p {T0:Type} {Bp0:Bool_Alg_p T0} (l:list (btp Bp0)) : Btype_p T0 (Bc_p T0 Bp0) := match l with | nil => %1 | cons a l' => a %* (times_list_p l') end. Fixpoint plus_list_p {T0:Type} {Bp0:Bool_Alg_p T0} (l:list (btp Bp0)) : Btype_p T0 (Bc_p T0 Bp0) := match l with | nil => %0 | cons a l' => a %+ (plus_list_p l') end. Lemma times_list_p_eq : forall {T0:Type} {Bp0:Bool_Alg_p T0} (l:list (btp Bp0)), times_list_p l = times_list (ba_conv_list l). intros T0 Bp0 l. induction l as [|a l h1]. simpl. reflexivity. simpl. rewrite h1 at 1. reflexivity. Qed. Lemma plus_list_p_eq : forall {T0:Type} {Bp0:Bool_Alg_p T0} (l:list (btp Bp0)), plus_list_p l = plus_list (ba_conv_list l). intros T0 Bp0 l. induction l as [|a l h1]. simpl. reflexivity. simpl. rewrite h1 at 1. reflexivity. Qed. Lemma plus_list_app_p : forall (l1 l2:list Btp), plus_list_p (l1 ++ l2) = plus_list_p l1 %+ plus_list_p l2. intros l1 l2. do 3 rewrite plus_list_p_eq. apply (@plus_list_app (ba_conv Bp)). Qed. Lemma times_list_app_p : forall (l1 l2:list Btp), times_list_p (l1 ++ l2) = times_list_p l1 %* times_list_p l2. intros l1 l2. do 3 rewrite times_list_p_eq. apply (@times_list_app (ba_conv Bp)). Qed. Lemma dist_list_sing_plus_p : forall (l:list Btp) (x:Btp), x %* (plus_list_p l) = plus_list_p (map (fun y:Btp => (x%*y)) l). intros l x. do 2 rewrite plus_list_p_eq. apply (@dist_list_sing_plus (ba_conv Bp)). Qed. Lemma dist_list_sing_times_p : forall (l:list Btp) (x:Btp), x %+ (times_list_p l) = times_list_p (map (fun y:Btp => (x%+y)) l). intros l x. do 2 rewrite times_list_p_eq. apply (@dist_list_sing_times (ba_conv Bp)). Qed. Lemma dist_list_2_plus_p : forall (l1 l2:list Btp), (plus_list_p l1) %* (plus_list_p l2) = plus_list_p (map (fun p:Btp*Btp => (fst p %* snd p)) (list_prod l1 l2)). intros l1 l2. do 3 rewrite plus_list_p_eq. apply (@dist_list_2_plus (ba_conv Bp)). Qed. Lemma dist_list_2_times_p : forall (l1 l2:list Btp), (times_list_p l1) %+ (times_list_p l2) = times_list_p (map (fun p:Btp*Btp => (fst p %+ snd p)) (list_prod l1 l2)). intros l1 l2. do 3 rewrite times_list_p_eq. apply (@dist_list_2_times (ba_conv Bp)). Qed. Definition plus_times_list_of_lists_p (l:faml Btp) : Btp := plus_list_p (map (fun l':(list Btp) => times_list_p l') l). Lemma plus_times_list_of_lists_p_eq : forall (l:faml Btp), plus_times_list_of_lists_p l = plus_times_list_of_lists (ba_conv_faml l). intro l. induction l as [|a l h1]; simpl. reflexivity. unfold plus_times_list_of_lists_p, plus_times_list_of_lists. rewrite plus_list_p_eq. assert (h2:(fun l' : list Btp => times_list_p l') = (fun l' : list (Btype (Bc (ba_conv Bp))) => times_list l')). apply functional_extensionality. intro x. rewrite times_list_p_eq. reflexivity. rewrite h2. reflexivity. Qed. Definition times_plus_list_of_lists_p (l:faml Btp) : Btp := times_list_p (map (fun l':(list Btp) => plus_list_p l') l). Lemma times_plus_list_of_lists_p_eq : forall (l:faml Btp), times_plus_list_of_lists_p l = times_plus_list_of_lists (ba_conv_faml l). intro l. induction l as [|a l h1]; simpl. reflexivity. unfold times_plus_list_of_lists_p, plus_times_list_of_lists. rewrite times_list_p_eq. assert (h2:(fun l' : list Btp => plus_list_p l') = (fun l' : list (Btype (Bc (ba_conv Bp))) => plus_list l')). apply functional_extensionality. intro x. rewrite plus_list_p_eq. reflexivity. rewrite h2. reflexivity. Qed. Definition times_fun_p {T:Type} (p:T->Btp) (l:list T) := times_list_p (map p l). Lemma times_fun_p_eq : forall {T:Type} (p:T->Btp) (l:list T), times_fun_p p l = times_fun (ba_conv_ind p) l. intros T p l. unfold times_fun_p, times_fun. rewrite times_list_p_eq. reflexivity. Qed. Definition plus_fun_p {T:Type} (p:T->Btp) (l:list T) := plus_list_p (map p l). Lemma plus_fun_p_eq : forall {T:Type} (p:T->Btp) (l:list T), plus_fun_p p l = plus_fun (ba_conv_ind p) l. intros T p l. unfold plus_fun_p, times_fun. rewrite plus_list_p_eq. reflexivity. Qed. Definition times_fun_in_dep_p {T:Type} {l:list T} (p:fun_in_dep l Btp) := times_list_p (map_in_dep p). Definition plus_fun_in_dep_p {T:Type} {l:list T} (p:fun_in_dep l Btp) := plus_list_p (map_in_dep p). Lemma times_fun_in_dep_p_eq : forall {T:Type} {l:list T} (p:fun_in_dep l Btp), times_fun_in_dep_p p = times_fun_in_dep (ba_conv_fun_in_dep p). intros T l p. unfold ba_conv_fun_in_dep, times_fun_in_dep_p, times_fun_in_dep. rewrite times_list_p_eq. reflexivity. Qed. Lemma plus_fun_in_dep_p_eq : forall {T:Type} {l:list T} (p:fun_in_dep l Btp), plus_fun_in_dep_p p = plus_fun_in_dep (ba_conv_fun_in_dep p). intros T l p. unfold ba_conv_fun_in_dep, plus_fun_in_dep_p, plus_fun_in_dep. rewrite plus_list_p_eq. reflexivity. Qed. Definition ba_conv_ind_ind {T U:Type} (p:T->U->Btp) := transfer_dep (U:=fun V=>T->U->V) (ba_conv_type Bp) p. Definition plus_times_fun1_p {T U:Type} (li:list T) (lj:list U) (p:T->U->Btp) := plus_fun_p (fun i:T => times_fun_p (p i) lj) li. Lemma plus_times_fun1_p_eq : forall {T U:Type} (li:list T) (lj:list U) (p:T->U->Btp), plus_times_fun1_p li lj p = plus_times_fun1 li lj (ba_conv_ind_ind p). intros T U li lj p. unfold plus_times_fun1_p, plus_times_fun1. rewrite plus_fun_p_eq. assert (h1: (fun i : T => times_fun_p (p i) lj) = (fun i : T => times_fun (ba_conv_ind_ind p i) lj)). apply functional_extensionality. intro x. rewrite times_fun_p_eq. reflexivity. rewrite h1. reflexivity. Qed. Definition plus_times_fun2_p {T U:Type} (li:list T) (lj:list U) (p:T->U->Btp) := plus_fun_p (fun j:U => times_fun_p (fun i:T => (p i j)) li) lj. Lemma plus_times_fun2_p_eq : forall {T U:Type} (li:list T) (lj:list U) (p:T->U->Btp), plus_times_fun2_p li lj p = plus_times_fun2 li lj (ba_conv_ind_ind p). intros T U li lj p. unfold plus_times_fun2_p, plus_times_fun2. rewrite plus_fun_p_eq. assert (h1: (fun j : U => times_fun_p (fun i : T => p i j) li) = (fun j : U => times_fun (fun i : T => ba_conv_ind_ind p i j) li)). apply functional_extensionality. intro x. rewrite times_fun_p_eq. reflexivity. rewrite h1. reflexivity. Qed. Definition times_plus_fun1_p {T U:Type} (li:list T) (lj:list U) (p:T->U->Btp) := times_fun_p (fun i:T => plus_fun_p (p i) lj) li. Lemma times_plus_fun1_p_eq : forall {T U:Type} (li:list T) (lj:list U) (p:T->U->Btp), times_plus_fun1_p li lj p = times_plus_fun1 li lj (ba_conv_ind_ind p). intros T U li lj p. unfold times_plus_fun1_p, times_plus_fun1. rewrite times_fun_p_eq. assert (h1: (fun i : T => plus_fun_p (p i) lj) = (fun i : T => plus_fun (ba_conv_ind_ind p i) lj)). apply functional_extensionality. intro x. rewrite plus_fun_p_eq. reflexivity. rewrite h1. reflexivity. Qed. Definition times_plus_fun2_p {T U:Type} (li:list T) (lj:list U) (p:T->U->Btp) := times_fun_p (fun j:U => plus_fun_p (fun i:T => (p i j)) li) lj. Lemma times_plus_fun2_p_eq : forall {T U:Type} (li:list T) (lj:list U) (p:T->U->Btp), times_plus_fun2_p li lj p = times_plus_fun2 li lj (ba_conv_ind_ind p). intros T U li lj p. unfold times_plus_fun2_p, times_plus_fun2. rewrite times_fun_p_eq. assert (h1: (fun j : U => plus_fun_p (fun i : T => p i j) li) = (fun j : U => plus_fun (fun i : T => ba_conv_ind_ind p i j) li)). apply functional_extensionality. intro x. rewrite plus_fun_p_eq. reflexivity. rewrite h1. reflexivity. Qed. Definition plus_times_all_funs1_p {T U:Type} (pfdt:type_eq_dec T) (pfdu: type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp): Btp. pose (list_power li lj) as lp. pose (fun (l:list (T*U)) (pfinlp:In l lp) => (fpl_to_fun_in_dep (in_list_power_fpl pfdt pfdu _ _ l pfi (list_power_no_dup _ _ _ pfi pfj pfinlp) pfinlp))) as f. pose (fun (l:list (T*U)) (pfinlp:In l lp) => times_fun_in_dep_p (fun (i:T) (pfini:In i li) => p i (f l pfinlp i pfini))) as g. refine (plus_fun_in_dep_p g). Defined. Lemma plus_times_all_funs1_p_eq : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp), plus_times_all_funs1_p pfdt pfdu li lj pfi pfj p = plus_times_all_funs1 pfdt pfdu li lj pfi pfj (ba_conv_ind_ind p). intros T U hdt hdu li lj h1 h2 p. unfold plus_times_all_funs1_p. unfold plus_times_all_funs1. rewrite plus_fun_in_dep_p_eq. f_equal. rewrite ba_conv_fun_in_dep_eq. apply fun_in_dep_ext; intros. unfold ba_conv_elt, transfer, eq_rect_r, eq_rect. simpl. rewrite times_fun_in_dep_p_eq; auto. Qed. Definition plus_times_all_funs2_p {T U:Type} (pfdt:type_eq_dec T) (pfdu: type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp): Btp. pose (list_power lj li) as lp. pose (fun (l:list (U*T)) (pfinlp:In l lp) => (fpl_to_fun_in_dep (in_list_power_fpl pfdu pfdt _ _ l pfj (list_power_no_dup _ _ _ pfj pfi pfinlp) pfinlp))) as f. pose (fun (l:list (U*T)) (pfinlp:In l lp) => times_fun_in_dep_p (fun (j:U) (pfinj:In j lj) => p (f l pfinlp j pfinj) j)) as g. refine (plus_fun_in_dep_p g). Defined. Lemma plus_times_all_funs2_p_eq : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp), plus_times_all_funs2_p pfdt pfdu li lj pfi pfj p = plus_times_all_funs2 pfdt pfdu li lj pfi pfj (ba_conv_ind_ind p). intros T U hdt hdu li lj h1 h2 p. unfold plus_times_all_funs2_p. unfold plus_times_all_funs2. rewrite plus_fun_in_dep_p_eq. f_equal. rewrite ba_conv_fun_in_dep_eq. apply fun_in_dep_ext; intros. unfold ba_conv_elt, transfer, eq_rect_r, eq_rect. simpl. rewrite times_fun_in_dep_p_eq; auto. Qed. Definition times_plus_all_funs1_p {T U:Type} (pfdt:type_eq_dec T) (pfdu: type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp): Btp. pose (list_power li lj) as lp. pose (fun (l:list (T*U)) (pfinlp:In l lp) => (fpl_to_fun_in_dep (in_list_power_fpl pfdt pfdu _ _ l pfi (list_power_no_dup _ _ _ pfi pfj pfinlp) pfinlp))) as f. pose (fun (l:list (T*U)) (pfinlp:In l lp) => plus_fun_in_dep_p (fun (i:T) (pfini:In i li) => p i (f l pfinlp i pfini))) as g. refine (times_fun_in_dep_p g). Defined. Lemma times_plus_all_funs1_p_eq : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp), times_plus_all_funs1_p pfdt pfdu li lj pfi pfj p = times_plus_all_funs1 pfdt pfdu li lj pfi pfj (ba_conv_ind_ind p). intros T U hdt hdu li lj h1 h2 p. unfold times_plus_all_funs1_p. unfold times_plus_all_funs1. rewrite times_fun_in_dep_p_eq. f_equal. rewrite ba_conv_fun_in_dep_eq. apply fun_in_dep_ext; intros. unfold ba_conv_elt, transfer, eq_rect_r, eq_rect. simpl. rewrite plus_fun_in_dep_p_eq; auto. Qed. Definition times_plus_all_funs2_p {T U:Type} (pfdt:type_eq_dec T) (pfdu: type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp): Btp. pose (list_power lj li) as lp. pose (fun (l:list (U*T)) (pfinlp:In l lp) => (fpl_to_fun_in_dep (in_list_power_fpl pfdu pfdt _ _ l pfj (list_power_no_dup _ _ _ pfj pfi pfinlp) pfinlp))) as f. pose (fun (l:list (U*T)) (pfinlp:In l lp) => plus_fun_in_dep_p (fun (j:U) (pfinj:In j lj) => p (f l pfinlp j pfinj) j)) as g. refine (times_fun_in_dep_p g). Defined. Lemma times_plus_all_funs2_p_eq : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp), times_plus_all_funs2_p pfdt pfdu li lj pfi pfj p = times_plus_all_funs2 pfdt pfdu li lj pfi pfj (ba_conv_ind_ind p). intros T U hdt hdu li lj h1 h2 p. unfold times_plus_all_funs2_p. unfold times_plus_all_funs2. rewrite times_fun_in_dep_p_eq. f_equal. rewrite ba_conv_fun_in_dep_eq. apply fun_in_dep_ext; intros. unfold ba_conv_elt, transfer, eq_rect_r, eq_rect. simpl. rewrite plus_fun_in_dep_p_eq; auto. Qed. Lemma complete_dist_list_times_p : forall (l:faml Btp), times_plus_list_of_lists_p l = plus_times_list_of_lists_p (list_of_lists_seqs l). intro l. rewrite times_plus_list_of_lists_p_eq. rewrite plus_times_list_of_lists_p_eq. apply (@complete_dist_list_times (ba_conv Bp)). Qed. Lemma complete_dist_list_plus_p : forall (l:faml Btp), plus_times_list_of_lists_p l = times_plus_list_of_lists_p (list_of_lists_seqs l). intro l. rewrite plus_times_list_of_lists_p_eq. rewrite times_plus_list_of_lists_p_eq. apply (@complete_dist_list_plus (ba_conv Bp)). Qed. Lemma inf_union_p : forall (U V:Ensemble Btp) (u v:Btp), inf_p U u -> inf_p V v -> (inf_p (Union U V) (u %* v)). intros U V u v. do 3 rewrite inf_p_iff. apply (@inf_union (ba_conv Bp)). Qed. Lemma sup_union_p : forall (U V:Ensemble Btp) (u v:Btp), sup_p U u -> sup_p V v -> (sup_p (Union U V) (u %+ v)). intros U V u v. do 3 rewrite sup_p_iff. apply (@sup_union (ba_conv Bp)). Qed. Lemma decompose_inf_p : forall (S X:Ensemble Btp) (s x x':Btp), Included X S -> inf_p S s -> inf_p (Setminus S X) x' -> inf_p X x -> inf_p (Couple x x') s. intros S X s x x'. do 4 rewrite inf_p_iff. apply (@decompose_inf (ba_conv Bp)). Qed. Lemma decompose_inf_p': forall (S X:Ensemble Btp) (x x':Btp), Included X S -> inf_p X x -> inf_p (Setminus S X) x' -> inf_p S (x %* x'). intros S X x x'. do 3 rewrite inf_p_iff. apply (@decompose_inf' (ba_conv Bp)). Qed. Lemma inf_times_cons_p : forall (E:Ensemble Btp) (a b:Btp), Inn E a -> inf_p (Subtract E a) b -> inf_p E (a %* b). intros E a b. do 2 rewrite inf_p_iff. apply (@inf_times_cons (ba_conv Bp)). Qed. Lemma inf_times_finite_p : forall (l:list Btp), inf_p (list_to_set l) (times_list_p l). intro l. rewrite inf_p_iff. rewrite times_list_p_eq. apply (@inf_times_finite (ba_conv Bp)). Qed. Lemma sup_plus_finite_p : forall (l:list Btp), sup_p (list_to_set l) (plus_list_p l). intro l. rewrite sup_p_iff. rewrite plus_list_p_eq. apply (@sup_plus_finite (ba_conv Bp)). Qed. Lemma le_times_finite_member_p : forall (l:list Btp) (x:Btp), (In x l) -> le_p (times_list_p l) x. intros l x. rewrite le_p_iff. rewrite times_list_p_eq. apply (@le_times_finite_member (ba_conv Bp)). Qed. Lemma le_member_plus_finite_p : forall (l:list Btp) (x:Btp), (In x l) -> le_p x (plus_list_p l). intros l x. rewrite le_p_iff. rewrite plus_list_p_eq. apply (@le_member_plus_finite (ba_conv Bp)). Qed. Lemma prod_list_dup_eq_p : forall (l:list Btp) (x:Btp), In x l -> times_list_p l = times_list_p (x::l). intros l x. do 2 rewrite times_list_p_eq. apply (@prod_list_dup_eq (ba_conv Bp)). Qed. Lemma sum_list_dup_eq_p : forall (l:list Btp) (x:Btp), In x l -> plus_list_p l = plus_list_p (x::l). intros l x. rewrite plus_list_p_eq. rewrite plus_list_p_eq. apply (@sum_list_dup_eq (ba_conv Bp)). Qed. Lemma prod_preserves_list_singularize_p : forall (l:list Btp), times_list_p l = times_list_p (list_singularize ba_eq_dec_p l). intro l. do 2 rewrite times_list_p_eq. rewrite (prod_preserves_list_singularize (ba_conv_list l)). f_equal. unfold ba_conv_list, transfer_dep, ba_conv_type, eq_rect_r, eq_rect. simpl. f_equal. apply type_eq_dec_eq. Qed. Lemma sum_preserves_list_singularize_p : forall (l:list Btp), plus_list_p l = plus_list_p (list_singularize ba_eq_dec_p l). intro l. do 2 rewrite plus_list_p_eq. rewrite (sum_preserves_list_singularize (ba_conv_list l)). f_equal. unfold ba_conv_list, transfer_dep, ba_conv_type, eq_rect_r, eq_rect. simpl. f_equal. apply type_eq_dec_eq. Qed. Lemma times_sing_preserves_new_head_p : forall (l:list Btp) (x:Btp), In x l -> let pf := ba_eq_dec_p in times_list_p (list_singularize pf l) = times_list_p (new_head pf (list_singularize pf l) x). intros l x. simpl. do 2 rewrite times_list_p_eq. intro h1. pose proof (@times_sing_preserves_new_head (ba_conv Bp) _ _ h1) as h2. unfold ba_conv_list, transfer_dep, ba_conv_type, eq_rect_r; simpl. simpl in h2. pose proof (type_eq_dec_eq (@ba_eq_dec (ba_conv Bp)) ba_eq_dec_p) as h3. rewrite <- h3 at 1. rewrite h2 at 1. repeat f_equal; auto. Qed. Lemma plus_sing_preserves_new_head_p : forall (l:list Btp) (x:Btp), In x l -> plus_list_p (list_singularize ba_eq_dec_p l) = plus_list_p (new_head ba_eq_dec_p (list_singularize ba_eq_dec_p l) x). intros l x. simpl. do 2 rewrite plus_list_p_eq. intro h1. pose proof (@plus_sing_preserves_new_head (ba_conv Bp) _ _ h1) as h2. unfold ba_conv_list, transfer_dep, ba_conv_type, eq_rect_r; simpl. simpl in h2. pose proof (type_eq_dec_eq (@ba_eq_dec (ba_conv Bp)) ba_eq_dec_p) as h3. rewrite <- h3 at 1. rewrite h2 at 1. repeat f_equal; auto. Qed. Lemma times_list_sing_cons_p : forall (l:list Btp) (x:Btp), In x (list_singularize ba_eq_dec_p l) -> times_list_p (list_singularize ba_eq_dec_p l) = x%*times_list_p (list_singularize ba_eq_dec_p (remove ba_eq_dec_p x l)). intros l x h1. do 2 rewrite times_list_p_eq. pose proof (type_eq_dec_eq (@ba_eq_dec (ba_conv Bp))) ba_eq_dec_p as h3. rewrite <- h3. apply (@times_list_sing_cons (ba_conv Bp)). rewrite h3; auto. Qed. Lemma plus_list_sing_cons_p : forall (l:list Btp) (x:Btp), In x (list_singularize ba_eq_dec_p l) -> plus_list_p (list_singularize ba_eq_dec_p l) = x%+plus_list_p (list_singularize ba_eq_dec_p (remove ba_eq_dec_p x l)). intros l x h1. do 2 rewrite plus_list_p_eq. pose proof (type_eq_dec_eq (@ba_eq_dec (ba_conv Bp))) ba_eq_dec_p as h3. rewrite <- h3. apply (@plus_list_sing_cons (ba_conv Bp)). rewrite h3; auto. Qed. Lemma list_to_sets_eq_times_sing_eq_p : forall (E:Ensemble Btp), forall (l l':list Btp), E = list_to_set (list_singularize ba_eq_dec_p l) -> E = list_to_set (list_singularize ba_eq_dec_p l') -> times_list_p (list_singularize ba_eq_dec_p l) = times_list_p (list_singularize ba_eq_dec_p l'). intros E l l'. do 2 rewrite times_list_p_eq. pose proof (type_eq_dec_eq (@ba_eq_dec (ba_conv Bp)) ba_eq_dec_p) as h3. rewrite <- h3. apply (@list_to_sets_eq_times_sing_eq (ba_conv Bp)). Qed. Lemma list_to_sets_eq_plus_sing_eq_p : forall (E:Ensemble Btp), forall (l l':list Btp), E = list_to_set (list_singularize ba_eq_dec_p l) -> E = list_to_set (list_singularize ba_eq_dec_p l') -> plus_list_p (list_singularize ba_eq_dec_p l) = plus_list_p (list_singularize ba_eq_dec_p l'). intros E l l'. do 2 rewrite plus_list_p_eq. pose proof (type_eq_dec_eq (@ba_eq_dec (ba_conv Bp)) ba_eq_dec_p) as h3. rewrite <- h3. apply (@list_to_sets_eq_plus_sing_eq (ba_conv Bp)). Qed. Lemma list_to_sets_eq_times_eq_p : forall (E:Ensemble Btp), forall (l l':list Btp), E = list_to_set l -> E = list_to_set l' -> times_list_p l = times_list_p l'. intros E l l'. do 2 rewrite times_list_p_eq. apply (@list_to_sets_eq_times_eq (ba_conv Bp)). Qed. Lemma list_to_sets_eq_plus_eq_p : forall (E:Ensemble Btp), forall (l l':list Btp), E = list_to_set l -> E = list_to_set l' -> plus_list_p l = plus_list_p l'. intros E l l'. do 2 rewrite plus_list_p_eq. apply (@list_to_sets_eq_plus_eq (ba_conv Bp)). Qed. Lemma times_list_unq_p : forall (E:Ensemble Btp) (pf:Finite E), exists! x, forall (l:list Btp), E = list_to_set l -> x = times_list_p l. intros E h1. pose proof (@times_list_unq (ba_conv Bp) E h1) as h2. destruct h2 as [x h2]. exists x. red in h2. destruct h2 as [h2l h2r]. red. split. intros l h3. rewrite times_list_p_eq. apply h2l; auto. intros x' h3. assert (h4:forall l:list (Btype (Bc (ba_conv Bp))), E = list_to_set l -> x' = times_list l). intros l h4. specialize (h3 _ h4). rewrite times_list_p_eq in h3. assumption. apply h2r; auto. Qed. Lemma plus_list_unq_p : forall (E:Ensemble Btp) (pf:Finite E), exists! x, forall (l:list Btp), E = list_to_set l -> x = plus_list_p l. intros E h1. pose proof (@plus_list_unq (ba_conv Bp) E h1) as h2. destruct h2 as [x h2]. exists x. red in h2. destruct h2 as [h2l h2r]. red. split. intros l h3. rewrite plus_list_p_eq. apply h2l; auto. intros x' h3. assert (h4:forall l:list (Btype (Bc (ba_conv Bp))), E = list_to_set l -> x' = plus_list l). intros l h4. specialize (h3 _ h4). rewrite plus_list_p_eq in h3. assumption. apply h2r; auto. Qed. Definition times_set_p (E:Ensemble Btp) (pf:Finite E) : Btp. refine (proj1_sig (constructive_definite_description _ (times_list_unq_p E pf))). Defined. Lemma times_set_p_eq : forall (E:Ensemble Btp) (pf:Finite E), times_set_p E pf = ba_conv_elt (times_set (ba_conv_set E) pf). intros E h1. unfold times_set_p. unfold times_set. destruct constructive_definite_description as [x h2]. destruct constructive_definite_description as [y h3]. simpl. pose proof (finite_set_list _ h1) as h4. destruct h4 as [l h4]. specialize (h2 _ h4). specialize (h3 _ h4). subst. rewrite times_list_p_eq. reflexivity. Qed. Definition plus_set_p (E:Ensemble Btp) (pf:Finite E) : Btp. refine (proj1_sig (constructive_definite_description _ (plus_list_unq_p E pf))). Defined. Lemma plus_set_p_eq : forall (E:Ensemble Btp) (pf:Finite E), plus_set_p E pf = ba_conv_elt (plus_set (ba_conv_set E) pf). intros E h1. unfold plus_set_p. unfold plus_set. destruct constructive_definite_description as [x h2]. destruct constructive_definite_description as [y h3]. simpl. pose proof (finite_set_list _ h1) as h4. destruct h4 as [l h4]. specialize (h2 _ h4). specialize (h3 _ h4). subst. rewrite plus_list_p_eq. reflexivity. Qed. Lemma times_set_compat_p : forall (E:Ensemble Btp) (pf:Finite E), exists l:list Btp, times_set_p E pf = times_list_p l. intros E h1. pose proof (@times_set_compat (ba_conv Bp) _ h1) as h2. destruct h2 as [l h2]. exists l. rewrite times_set_p_eq. rewrite times_list_p_eq. assumption. Qed. Lemma times_set_compat_p' : forall (E:Ensemble Btp) (l:list Btp) (pf:Finite E), E = list_to_set l -> times_set_p E pf = times_list_p l. intros E l h1. rewrite times_set_p_eq. rewrite times_list_p_eq. apply (@times_set_compat' (ba_conv Bp)). Qed. Lemma times_list_compat_p : forall (l:list Btp), times_list_p l = times_set_p (list_to_set l) (list_to_set_finite l). intro; erewrite times_set_compat_p'; auto. Qed. Lemma times_set_list_to_set_p : forall (l:list Btp) (pf:Finite (list_to_set l)), times_set_p (list_to_set l) pf = times_list_p l. intros; apply times_set_compat_p'; auto. Qed. Lemma plus_set_compat_p : forall (E:Ensemble Btp) (pf:Finite E), exists l:list Btp, plus_set_p E pf = plus_list_p l. intros E h1. pose proof (@plus_set_compat (ba_conv Bp) _ h1) as h2. destruct h2 as [l h2]. exists l. rewrite plus_list_p_eq. rewrite plus_set_p_eq. assumption. Qed. Lemma plus_set_compat_p' : forall (E:Ensemble Btp) (l:list Btp) (pf:Finite E), E = list_to_set l -> plus_set_p E pf = plus_list_p l. intros E l h1. rewrite plus_set_p_eq. rewrite plus_list_p_eq. apply (@plus_set_compat' (ba_conv Bp)). Qed. Lemma plus_list_compat_p : forall (l:list Btp), plus_list_p l = plus_set_p (list_to_set l) (list_to_set_finite l). intro; erewrite plus_set_compat_p'; auto. Qed. Lemma plus_set_list_to_set_p : forall (l:list Btp) (pf:Finite (list_to_set l)), plus_set_p (list_to_set l) pf = plus_list_p l. intros; apply plus_set_compat_p'; auto. Qed. Lemma times_set_functional_p : forall (A C:Ensemble Btp), A = C -> forall (pfa:Finite A) (pfc:Finite C), times_set_p A pfa = times_set_p C pfc. intros A C ? h1 h2. subst. do 2 rewrite times_set_p_eq. apply (@times_set_functional (ba_conv Bp)). pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. Lemma plus_set_functional_p : forall (A C:Ensemble Btp), A = C -> forall (pfa:Finite A) (pfc:Finite C), plus_set_p A pfa = plus_set_p C pfc. intros A C ? h1 h2. subst. do 2 rewrite plus_set_p_eq. apply (@plus_set_functional (ba_conv Bp)). pose proof (proof_irrelevance _ h1 h2); subst; auto. Qed. Lemma times_set_empty_p : times_set_p (Empty_set Btp) (Empty_is_finite Btp) = %1. rewrite times_set_p_eq. apply (@times_set_empty (ba_conv Bp)). Qed. Lemma times_set_empty_p' : forall (pf:Finite (Empty_set Btp)), times_set_p (Empty_set Btp) pf = %1. intro h1. rewrite times_set_p_eq. apply (@times_set_empty' (ba_conv Bp)). Qed. Lemma plus_set_empty_p : plus_set_p (Empty_set Btp) (Empty_is_finite Btp) = %0. rewrite plus_set_p_eq. apply (@plus_set_empty (ba_conv Bp)). Qed. Lemma plus_set_empty_p' : forall (pf:Finite (Empty_set Btp)), plus_set_p (Empty_set Btp) pf = %0. intro h1. rewrite plus_set_p_eq. apply (@plus_set_empty' (ba_conv Bp)). Qed. Lemma times_set_add_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp), times_set_p (Add E x) (Add_preserves_Finite _ _ _ pf) = x %* (times_set_p E pf). intros E h1 x. do 2 rewrite times_set_p_eq. apply (@times_set_add (ba_conv Bp)). Qed. Lemma times_set_add_p' : forall (E:Ensemble Btp) (pf1:Finite E) (x:Btp) (pf2:Finite (Add E x)), times_set_p (Add E x) pf2 = x %* (times_set_p E pf1). intros E h1 x h2. do 2 rewrite times_set_p_eq. apply (@times_set_add' (ba_conv Bp)). Qed. Lemma plus_set_add_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp), plus_set_p (Add E x) (Add_preserves_Finite _ _ _ pf) = x %+ (plus_set_p E pf). intros E h1 x. do 2 rewrite plus_set_p_eq. apply (@plus_set_add (ba_conv Bp)). Qed. Lemma plus_set_add_p' : forall (E:Ensemble Btp) (pf1:Finite E) (x:Btp) (pf2:Finite (Add E x)), plus_set_p (Add E x) pf2 = x %+ (plus_set_p E pf1). intros E h1 x h2. do 2 rewrite plus_set_p_eq. apply (@plus_set_add' (ba_conv Bp)). Qed. Lemma times_set_sing_p : forall x:Btp, times_set_p (Singleton x) (Singleton_is_finite _ x) = x. intro x. rewrite times_set_p_eq. apply (@times_set_sing (ba_conv Bp)). Qed. Lemma times_set_sing_p' : forall (x:Btp) (pf:Finite (Singleton x)), times_set_p _ pf = x. intros x h1. rewrite times_set_p_eq. apply (@times_set_sing' (ba_conv Bp)). Qed. Lemma times_list_sing_p : forall (x:Btp), times_list_p (x::nil) = x. intros; simpl. rewrite one_prod_p; auto. Qed. Lemma times_list_all_sing_p : forall (l:list Btp) (x:Btp), l <> nil -> all_sing l x -> times_list_p l = x. intros; rewrite times_list_p_eq. apply (@times_list_all_sing (ba_conv Bp)); auto. Qed. Lemma times_set_one_or_p : forall (E:Ensemble Btp) (pf:Finite E), (E = Empty_set _) \/ (E = Singleton %1) -> %1 = times_set_p E pf. intros E h1. rewrite times_set_p_eq. apply (@times_set_one_or (ba_conv Bp)). Qed. Lemma plus_set_sing_p : forall x:Btp, plus_set_p (Singleton x) (Singleton_is_finite _ x) = x. intro x. rewrite plus_set_p_eq. apply (@plus_set_sing (ba_conv Bp)). Qed. Lemma plus_set_sing_p' : forall (x:Btp) (pf:Finite (Singleton x)), plus_set_p _ pf = x. intros x h1. rewrite plus_set_p_eq. apply (@plus_set_sing' (ba_conv Bp)). Qed. Lemma plus_list_sing_p : forall (x:Btp), plus_list_p (x::nil) = x. intros; simpl. rewrite zero_sum_p; auto. Qed. Lemma plus_list_all_sing_p : forall (l:list Btp) (x:Btp), l <> nil -> all_sing l x -> plus_list_p l = x. intros; rewrite plus_list_p_eq. apply (@plus_list_all_sing (ba_conv Bp)); auto. Qed. Lemma plus_set_zero_or_p : forall (E:Ensemble Btp) (pf:Finite E), (E = Empty_set _) \/ (E = Singleton %0) -> %0 = plus_set_p E pf. intros E h1. rewrite plus_set_p_eq. apply (@plus_set_zero_or (ba_conv Bp)). Qed. Lemma times_set_couple_p : forall (x y:Btp), times_set_p (Couple x y) (finite_couple x y) = x %* y. intros x y. rewrite times_set_p_eq. apply (@times_set_couple (ba_conv Bp)). Qed. Lemma times_set_couple_p' : forall (x y:Btp) (pf:Finite (Couple x y)), times_set_p (Couple x y) pf = x %* y. intros x y h1. rewrite times_set_p_eq. apply (@times_set_couple' (ba_conv Bp)). Qed. Lemma plus_set_couple_p : forall (x y:Btp), plus_set_p (Couple x y) (finite_couple x y) = x %+ y. intros x y. rewrite plus_set_p_eq. apply (@plus_set_couple (ba_conv Bp)). Qed. Lemma plus_set_couple_p' : forall (x y:Btp) (pf:Finite (Couple x y)), plus_set_p (Couple x y) pf = x %+ y. intros x y h1. rewrite plus_set_p_eq. apply (@plus_set_couple' (ba_conv Bp)). Qed. Lemma le_times_set_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp), (Inn E x) -> le_p (times_set_p E pf) x. intros E h1 x. rewrite le_p_iff. rewrite times_set_p_eq. apply (@le_times_set (ba_conv Bp)). Qed. Lemma le_plus_set_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp), (Inn E x) -> le_p x (plus_set_p E pf). intros E h1 x h2. rewrite le_p_iff. rewrite plus_set_p_eq. apply (@le_plus_set (ba_conv Bp)); auto. Qed. Lemma inf_times_set_p : forall (E:Ensemble Btp) (pf:Finite E), inf_p E (times_set_p E pf). intros E h1. rewrite inf_p_iff. rewrite times_set_p_eq. apply (@inf_times_set (ba_conv Bp)). Qed. Lemma sup_plus_set_p : forall (E:Ensemble Btp) (pf:Finite E), sup_p E (plus_set_p E pf). intros E h1. rewrite sup_p_iff. rewrite plus_set_p_eq. apply (@sup_plus_set (ba_conv Bp)). Qed. Lemma times_set_union_p : forall (C D:Ensemble Btp) (pfc:Finite C) (pfd: Finite D), times_set_p (Union C D) (Union_preserves_Finite _ _ _ pfc pfd) = (times_set_p C pfc) %* (times_set_p D pfd). intros C D h1 h2. rewrite times_set_p_eq. do 2 rewrite times_set_p_eq. apply (@times_set_union (ba_conv Bp)). Qed. Lemma times_set_union_p' : forall (C D:Ensemble Btp) (pfc:Finite C) (pfd: Finite D) (pfu:Finite (Union C D)), times_set_p (Union C D) pfu = (times_set_p C pfc) %* (times_set_p D pfd). intros C D h1 h2 h3. do 3 rewrite times_set_p_eq. apply (@times_set_union' (ba_conv Bp)). Qed. Lemma times_set_union_p'' : forall (C D:Ensemble Btp) (pf:Finite (Union C D)), let pfc := Finite_downward_closed _ _ pf _ (Union_increases_l _ _ _) in let pfd := Finite_downward_closed _ _ pf _ (Union_increases_r _ _ _) in times_set_p (Union C D) pf = (times_set_p C pfc) %* (times_set_p D pfd). intros; apply times_set_union_p'. Qed. Lemma times_set_inc_le_p : forall (C D:Ensemble Btp) (pfc:Finite C) (pfd:Finite D) (pfi:Included C D), le_p (times_set_p D pfd) (times_set_p C pfc). intros C D h1 h2 h3. rewrite le_p_iff. do 2 rewrite times_set_p_eq. apply (@times_set_inc_le (ba_conv Bp)); auto. Qed. Lemma plus_set_union_p : forall (C D:Ensemble Btp) (pfc:Finite C) (pfd: Finite D), plus_set_p (Union C D) (Union_preserves_Finite _ _ _ pfc pfd) = (plus_set_p C pfc) %+ (plus_set_p D pfd). intros C D h1 h2. do 3 rewrite plus_set_p_eq. apply (@plus_set_union (ba_conv Bp)). Qed. Lemma plus_set_union_p' : forall (C D:Ensemble Btp) (pfc:Finite C) (pfd: Finite D) (pfu:Finite (Union C D)), plus_set_p (Union C D) pfu = (plus_set_p C pfc) %+ (plus_set_p D pfd). intros C D h1 h2 h3. do 3 rewrite plus_set_p_eq. apply (@plus_set_union' (ba_conv Bp)). Qed. Lemma plus_set_union_p'' : forall (C D:Ensemble Btp) (pf:Finite (Union C D)), let pfc := Finite_downward_closed _ _ pf _ (Union_increases_l _ _ _) in let pfd := Finite_downward_closed _ _ pf _ (Union_increases_r _ _ _) in plus_set_p (Union C D) pf = (plus_set_p C pfc) %+ (plus_set_p D pfd). intros; apply plus_set_union_p'. Qed. Lemma plus_set_inc_le_p : forall (C D:Ensemble Btp) (pfc:Finite C) (pfd:Finite D) (pfi:Included C D), le_p (plus_set_p C pfc) (plus_set_p D pfd). intros C D h1 h2 h3. rewrite le_p_iff. do 2 rewrite plus_set_p_eq. apply (@plus_set_inc_le (ba_conv Bp)); auto. Qed. Lemma plus_set_im_add_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp) (f:Btp->Btp), plus_set_p (Im (Add E x) f) (finite_image _ _ _ f (Add_preserves_Finite _ _ x pf)) = f x %+ plus_set_p (Im E f) (finite_image _ _ _ f pf). intros E h1 x f. do 2 rewrite plus_set_p_eq. apply (@plus_set_im_add (ba_conv Bp)). Qed. Lemma plus_set_im_add_p' : forall (E:Ensemble Btp) (x:Btp) (f:Btp->Btp) (pf0: Finite E) (pf1:Finite (Im (Add E x) f)) (pf2:Finite (Im E f)), plus_set_p (Im (Add E x) f) pf1 = f x %+ (plus_set_p (Im E f) pf2). intros E x f h1 h2 h3. do 2 rewrite plus_set_p_eq. apply (@plus_set_im_add' (ba_conv Bp)). assumption. Qed. Lemma plus_set_im_im_p : forall {T U:Type} (C:Ensemble T) (f:T->U) (g:U->Btp) (pfc : Finite C) (pfiic:Finite (Im (Im C f) g)), let k := (fun x:T => g (f x)) in plus_set_p (Im (Im C f) g) pfiic = plus_set_p (Im C k) (finite_image _ _ _ k pfc). intros T U C f g h1 h3 k. do 2 rewrite plus_set_p_eq. apply (@plus_set_im_im (ba_conv Bp)); auto. Qed. Lemma times_set_im_add_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp) (f:Btp->Btp), times_set_p (Im (Add E x) f) (finite_image _ _ _ f (Add_preserves_Finite _ _ x pf)) = f x %* times_set_p (Im E f) (finite_image _ _ _ f pf). intros E h1 x f. do 2 rewrite times_set_p_eq. apply (@times_set_im_add (ba_conv Bp)). Qed. Lemma times_set_im_add_p' : forall (E:Ensemble Btp) (x:Btp) (f:Btp->Btp) (pf0: Finite E) (pf1:Finite (Im (Add E x) f)) (pf2:Finite (Im E f)), times_set_p (Im (Add E x) f) pf1 = f x %* (times_set_p (Im E f) pf2). intros E x f h1 h2 h3. do 2 rewrite times_set_p_eq. apply (@times_set_im_add' (ba_conv Bp)). assumption. Qed. Lemma times_set_im_im_p : forall {T U:Type} (C:Ensemble T) (f:T->U) (g:U->Btp) (pfc : Finite C) (pfiic:Finite (Im (Im C f) g)), let k := (fun x:T => g (f x)) in times_set_p (Im (Im C f) g) pfiic = times_set_p (Im C k) (finite_image _ _ _ k pfc). intros T U C f g h1 h3 k. do 2 rewrite times_set_p_eq. apply (@times_set_im_im (ba_conv Bp)); auto. Qed. Lemma dist_set_plus1_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp), let f := (fun y:Btp => x %* y) in x %* (plus_set_p E pf) = plus_set_p (Im E f) (finite_image _ _ _ f pf). intros E h1 x f. do 2 rewrite plus_set_p_eq. apply (@dist_set_plus1 (ba_conv Bp)); auto. Qed. Lemma times_list_to_set_map_in_dep_decompose_sumbool_p : forall {U:Type} (l:list U) (Q:U->Prop) (pfs:forall x:U, In x l -> {Q x} + {~Q x}) (fq:fun_in_dep_prop l Q Btp) (fr:fun_in_dep_prop l (fun x => ~ Q x) Btp), let f := fun (x:U) (pf:In x l) => match (pfs x pf) with | left pfq => fq x pf pfq | right pfr => fr x pf pfr end in times_list_p (map_in_dep f) = times_list_p (map_in_dep_prop_l pfs fq) %* times_list_p (map_in_dep_prop_r pfs fr). intros U l Q h1 fq fr. simpl. rewrite times_list_compat_p. gen0. rewrite list_to_set_map_in_dep_decompose_sumbool. intro h2. rewrite times_set_union_p''. do 2 rewrite times_set_list_to_set_p. reflexivity. Qed. Lemma dist_set_plus1_p' : forall (E:Ensemble Btp) (pf1:Finite E) (x:Btp), let f := (fun y:Btp => x %* y) in forall (pf2:Finite (Im E f)), x %* (plus_set_p E pf1) = plus_set_p (Im E f) pf2. intros E h1 x f h2. do 2 rewrite plus_set_p_eq. apply (@dist_set_plus1' (ba_conv Bp)). Qed. Lemma dist_set_times1_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp), let f := (fun y:Btp => x %+ y) in x %+ (times_set_p E pf) = times_set_p (Im E f) (finite_image _ _ _ f pf). intros E h1 x f. do 2 rewrite times_set_p_eq. apply (@dist_set_times1 (ba_conv Bp)). Qed. Lemma dist_set_times1_p' : forall (E:Ensemble Btp) (pf1:Finite E) (x:Btp), let f := (fun y:Btp => x %+ y) in forall (pf2:Finite (Im E f)), x %+ (times_set_p E pf1) = times_set_p (Im E f) pf2. intros E h1 x f h2. do 2 rewrite times_set_p_eq. apply (@dist_set_times1' (ba_conv Bp)). Qed. Lemma dist_set_plus2_p : forall (D E:Ensemble Btp) (pfd: Finite D) (pfe: Finite E), let f := (fun p:Btp*Btp => (fst p) %* (snd p)) in (plus_set_p D pfd) %* (plus_set_p E pfe) = plus_set_p (Im (cart_prod D E) f) (finite_image _ _ _ f (cart_prod_fin _ _ pfd pfe)). intros D E h1 h2 f. do 3 rewrite plus_set_p_eq. apply (@dist_set_plus2 (ba_conv Bp)). Qed. Lemma dist_set_plus2_p' : forall (D E:Ensemble Btp) (pfd: Finite D) (pfe: Finite E), let f := (fun p:Btp*Btp => (fst p) %* (snd p)) in forall (pfc: Finite (Im (cart_prod D E) f)), (plus_set_p D pfd) %* (plus_set_p E pfe) = plus_set_p (Im (cart_prod D E) f) pfc. intros D E h1 h2 f h3. do 3 rewrite plus_set_p_eq. apply (@dist_set_plus2' (ba_conv Bp)). Qed. Lemma dist_set_times2_p : forall (D E:Ensemble Btp) (pfd: Finite D) (pfe: Finite E), let f := (fun p:Btp*Btp => (fst p) %+ (snd p)) in (times_set_p D pfd) %+ (times_set_p E pfe) = times_set_p (Im (cart_prod D E) f) (finite_image _ _ _ f (cart_prod_fin _ _ pfd pfe)). intros D E h1 h2 f. do 3 rewrite times_set_p_eq. apply (@dist_set_times2 (ba_conv Bp)). Qed. Lemma dist_set_times2_p' : forall (D E:Ensemble Btp) (pfd: Finite D) (pfe: Finite E), let f := (fun p:Btp*Btp => (fst p) %+ (snd p)) in forall (pfc: Finite (Im (cart_prod D E) f)), (times_set_p D pfd) %+ (times_set_p E pfe) = times_set_p (Im (cart_prod D E) f) pfc. intros D E h1 h2 f h3. do 3 rewrite times_set_p_eq. apply (@dist_set_times2' (ba_conv Bp)); auto. Qed. Definition ba_conv_fin_map_dom {U:Type} {C:Ensemble Btp} {D:Ensemble U} {def:U} (F:Fin_map C D def) : Fin_map (ba_conv_set C) D def. unfold ba_conv_set, ba_conv_type, transfer_dep_eq_refl. refine F. Defined. Definition ba_conv_fin_map_ran {T:Type} {C:Ensemble T} {D:Ensemble Btp} (F:Fin_map C D %0) : Fin_map C (ba_conv_set D) 0. unfold ba_conv_set, ba_conv_type, transfer_dep_eq_refl. refine F. Defined. Definition ba_conv_ens_fin_map_dom {U:Type} {C:Ensemble Btp} {D:Ensemble U} {def:U} (A:Ensemble (Fin_map C D def)) : Ensemble (Fin_map (ba_conv_set C) D def). unfold ba_conv_set, ba_conv_type, transfer_dep_eq_refl. refine A. Defined. Definition ba_conv_ens_fin_map_ran {T:Type} {C:Ensemble T} {D:Ensemble Btp} (A:Ensemble (Fin_map C D %0)) : Ensemble (Fin_map C (ba_conv_set D) 0). unfold ba_conv_set, ba_conv_type, transfer_dep_eq_refl. refine A. Defined. Ltac match1_goal := let pf := fresh "her" in match goal with |- context [match ?h' with | _ => _ end] => pose h' as pf end. Ltac match1 h := let P := type of h in let pf := fresh "her" in match P with context [match ?h' with | _ => _ end] => pose h' as pf end. Lemma incl_ba_conv_ens_fin_map_dom_impl : forall {U:Type} {C:Ensemble Btp} {D:Ensemble U} {def:U} (A B:Ensemble (Fin_map C D def)), Included A B -> Included (ba_conv_ens_fin_map_dom A) (ba_conv_ens_fin_map_dom B). intros U C D def A B h1. red. intros x h2. unfold ba_conv_ens_fin_map_dom, eq_rect_r, eq_rect, transfer_dep_eq_refl in h2; simpl in h2. apply h1 in h2. unfold ba_conv_ens_fin_map_dom, eq_rect_r, eq_rect, transfer_dep_eq_refl; simpl; auto. Qed. Lemma incl_ba_conv_ens_fin_map_dom_iff : forall {U:Type} {C:Ensemble Btp} {D:Ensemble U} {def:U} (A B:Ensemble (Fin_map C D def)), Included A B <-> Included (ba_conv_ens_fin_map_dom A) (ba_conv_ens_fin_map_dom B). intros; red; split. apply incl_ba_conv_ens_fin_map_dom_impl. intros h1. unfold ba_conv_ens_fin_map_dom, eq_rect_r, eq_rect, transfer_dep_eq_refl in h1; simpl in h1; auto. Qed. Definition plus_fin_pair_map1_p {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0) : Fin_map C (Im C (fun x : T => plus_set (im1 (ba_conv_fin_map_ran F) x) (finite_im1 (ba_conv_fin_map_ran F) x))) %0 := (fun_to_fin_map pfdt C %0 pfc (fun x:T => plus_set (im1 (ba_conv_fin_map_ran F) x) (finite_im1 (ba_conv_fin_map_ran F) x))). Lemma plus_fin_pair_map1_p_eq : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0), plus_fin_pair_map1_p pfdt pfc F = plus_fin_pair_map1 pfdt pfc (ba_conv_fin_map_ran F). auto. Qed. Definition plus_fin_pair_map2_p {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0) : Fin_map D (Im D (fun y : U => plus_set (im2 (ba_conv_fin_map_ran F) y) (finite_im2 (ba_conv_fin_map_ran F) y))) %0 := (fun_to_fin_map pfdu _ %0 pfd (fun y:U => plus_set (im2 (ba_conv_fin_map_ran F) y) (finite_im2 (ba_conv_fin_map_ran F) y))). Lemma plus_fin_pair_map2_p_eq : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0), plus_fin_pair_map2_p pfdu pfd F = plus_fin_pair_map2 pfdu pfd (ba_conv_fin_map_ran F). auto. Qed. Definition times_fin_pair_map1_p {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0) : Fin_map C (Im C (fun x : T => times_set (im1 (ba_conv_fin_map_ran F) x) (finite_im1 (ba_conv_fin_map_ran F) x))) %0 := (fun_to_fin_map pfdt C %0 pfc (fun x:T => times_set (im1 (ba_conv_fin_map_ran F) x) (finite_im1 (ba_conv_fin_map_ran F) x))). Lemma times_fin_pair_map1_p_eq : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} (pfdt:type_eq_dec T) {E:Ensemble Btp} (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0), times_fin_pair_map1_p pfdt pfc F = times_fin_pair_map1 pfdt pfc (ba_conv_fin_map_ran F). auto. Qed. Definition times_fin_pair_map2_p {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0) : Fin_map D (Im D (fun y : U => times_set (im2 (ba_conv_fin_map_ran F) y) (finite_im2 (ba_conv_fin_map_ran F) y))) %0 := (fun_to_fin_map pfdu D %0 pfd (fun y:U => times_set (im2 (ba_conv_fin_map_ran F) y) (finite_im2 (ba_conv_fin_map_ran F) y))). Lemma times_fin_pair_map2_p_eq : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0), times_fin_pair_map2_p pfdu pfd F = times_fin_pair_map2 pfdu pfd (ba_conv_fin_map_ran F). auto. Qed. Lemma plus_fin_pair_map1_list_compat_p : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfted:type_eq_dec T) (pfued:type_eq_dec U) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0) (nml:nice_map_lists2 F), let f := fun (x:T) (pf:Inn C x) => plus_list_p (im1l (fpl2 F nml) x (impl_in_n_la2 F nml pfc _ pf)) in fin_map_eq (plus_fin_pair_map1_p pfted pfc F) (fun_in_ens_to_fin_map pfted f pfc %0). intros T U C D E hted hued h1 F nml. rewrite plus_fin_pair_map1_p_eq. pose proof (@plus_fin_pair_map1_list_compat (ba_conv Bp) _ _ _ _ _ hted h1 (ba_conv_fin_map_ran F) nml) as h2. red in h2. destruct h2 as [h2 h3]. red. ex_goal. red. intros x h4. apply h2 in h4. destruct h4 as [x h4 ]. subst. pose (impl_in_n_la2 _ nml h1 _ h4) as h6. pose proof (plus_list_p_eq (im1l (fpl2 F nml) x h6)) as h5. unfold ba_conv_list in h5. unfold ba_conv_type in h5. rewrite transfer_dep_eq_refl in h5. simpl. unfold ba_conv_fin_map_ran. unfold eq_rect_r. simpl. unfold Btp, btp in h5. unfold Btype_p in h5. unfold ba_conv_set. unfold transfer_dep. unfold eq_rect_r. unfold h6 in h5. simpl in h5. unfold Btp in h5. unfold Btp. rewrite <- h5 at 1. econstructor. reflexivity. exists hex. apply fin_map_ext. intros x h5. terml h3. gterml. assert (he:x1 = x0 |-> x). unfold x1, x0. rewrite <- fin_map_new_ran_compat. rewrite <- fin_map_new_ran_compat; auto. assumption. fold x1. rewrite he. termr h3. gtermr. assert (h6:y0 = y |-> x). unfold y0, y. erewrite fun_in_ens_to_fin_map_compat. erewrite fun_in_ens_to_fin_map_compat. rewrite plus_list_p_eq. f_equal. unfold ba_conv_list, transfer_dep, ba_conv_type, eq_rect_r. simpl. f_equal. fold y0. rewrite h6. unfold x0, y. rewrite h3; auto. Grab Existential Variables. assumption. Qed. Lemma plus_fin_pair_map2_list_compat_p : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfted:type_eq_dec T) (pfued:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0) (nml:nice_map_lists2 F), let f := fun (y:U) (pf:Inn D y) => plus_list_p (im2l (fpl2 F nml) y (impl_in_n_lb2 F nml pfd _ pf)) in fin_map_eq (plus_fin_pair_map2_p pfued pfd F) (fun_in_ens_to_fin_map pfued f pfd %0). intros T U C D E hted hued h1 F nml. rewrite plus_fin_pair_map2_p_eq. pose proof (@plus_fin_pair_map2_list_compat (ba_conv Bp) _ _ _ _ _ hued h1 (ba_conv_fin_map_ran F) nml) as h2. red in h2. destruct h2 as [h2 h3]. red. ex_goal. red. intros x h4. apply h2 in h4. destruct h4 as [x h4 ]. subst. pose (impl_in_n_lb2 _ nml h1 _ h4) as h6. pose proof (plus_list_p_eq (im2l (fpl2 F nml) x h6)) as h5. unfold ba_conv_list in h5. unfold ba_conv_type in h5. rewrite transfer_dep_eq_refl in h5. simpl. unfold ba_conv_fin_map_ran. unfold eq_rect_r. simpl. unfold Btp, btp in h5. unfold Btype_p in h5. unfold ba_conv_set. unfold transfer_dep. unfold eq_rect_r. unfold h6 in h5. simpl in h5. unfold Btp in h5. unfold Btp. rewrite <- h5 at 1. econstructor. reflexivity. exists hex. apply fin_map_ext. intros x h5. terml h3. gterml. assert (he:x1 = x0 |-> x). unfold x1, x0. rewrite <- fin_map_new_ran_compat. rewrite <- fin_map_new_ran_compat; auto. assumption. fold x1. rewrite he. termr h3. gtermr. assert (h6:y0 = y |-> x). unfold y0, y. erewrite fun_in_ens_to_fin_map_compat. erewrite fun_in_ens_to_fin_map_compat. rewrite plus_list_p_eq. f_equal. unfold ba_conv_list, transfer_dep, ba_conv_type, eq_rect_r. simpl. f_equal. fold y0. rewrite h6. unfold x0, y. rewrite h3; auto. Grab Existential Variables. assumption. Qed. Lemma times_fin_pair_map1_list_compat_p : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfted:type_eq_dec T) (pfued:type_eq_dec U) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0) (nml:nice_map_lists2 F), let f := fun (x:T) (pf:Inn C x) => times_list_p (im1l (fpl2 F nml) x (impl_in_n_la2 F nml pfc _ pf)) in fin_map_eq (times_fin_pair_map1_p pfted pfc F) (fun_in_ens_to_fin_map pfted f pfc %0). intros T U C D E hted hued h1 F nml. rewrite times_fin_pair_map1_p_eq. pose proof (@times_fin_pair_map1_list_compat (ba_conv Bp) _ _ _ _ _ hted h1 (ba_conv_fin_map_ran F) nml) as h2. red in h2. destruct h2 as [h2 h3]. red. ex_goal. red. intros x h4. apply h2 in h4. destruct h4 as [x h4 ]. subst. pose (impl_in_n_la2 _ nml h1 _ h4) as h6. pose proof (times_list_p_eq (im1l (fpl2 F nml) x h6)) as h5. unfold ba_conv_list in h5. unfold ba_conv_type in h5. rewrite transfer_dep_eq_refl in h5. simpl. unfold ba_conv_fin_map_ran. unfold eq_rect_r. simpl. unfold Btp, btp in h5. unfold Btype_p in h5. unfold ba_conv_set. unfold transfer_dep. unfold eq_rect_r. unfold h6 in h5. simpl in h5. unfold Btp in h5. unfold Btp. rewrite <- h5 at 1. econstructor. reflexivity. exists hex. apply fin_map_ext. intros x h5. terml h3. gterml. assert (he:x1 = x0 |-> x). unfold x1, x0. rewrite <- fin_map_new_ran_compat. rewrite <- fin_map_new_ran_compat; auto. assumption. fold x1. rewrite he. termr h3. gtermr. assert (h6:y0 = y |-> x). unfold y0, y. erewrite fun_in_ens_to_fin_map_compat. erewrite fun_in_ens_to_fin_map_compat. rewrite times_list_p_eq. f_equal. unfold ba_conv_list, transfer_dep, ba_conv_type, eq_rect_r. simpl. f_equal. fold y0. rewrite h6. unfold x0, y. rewrite h3; auto. Grab Existential Variables. assumption. Qed. Lemma times_fin_pair_map2_list_compat_p : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfted:type_eq_dec T) (pfued:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0) (nml:nice_map_lists2 F), let f := fun (y:U) (pf:Inn D y) => times_list_p (im2l (fpl2 F nml) y (impl_in_n_lb2 F nml pfd _ pf)) in fin_map_eq (times_fin_pair_map2_p pfued pfd F) (fun_in_ens_to_fin_map pfued f pfd %0). intros T U C D E hted hued h1 F nml. rewrite times_fin_pair_map2_p_eq. pose proof (@times_fin_pair_map2_list_compat (ba_conv Bp) _ _ _ _ _ hued h1 (ba_conv_fin_map_ran F) nml) as h2. red in h2. destruct h2 as [h2 h3]. red. ex_goal. red. intros x h4. apply h2 in h4. destruct h4 as [x h4 ]. subst. pose (impl_in_n_lb2 _ nml h1 _ h4) as h6. pose proof (times_list_p_eq (im2l (fpl2 F nml) x h6)) as h5. unfold ba_conv_list in h5. unfold ba_conv_type in h5. rewrite transfer_dep_eq_refl in h5. simpl. unfold ba_conv_fin_map_ran. unfold eq_rect_r. simpl. unfold Btp, btp in h5. unfold Btype_p in h5. unfold ba_conv_set. unfold transfer_dep. unfold eq_rect_r. unfold h6 in h5. simpl in h5. unfold Btp in h5. unfold Btp. rewrite <- h5 at 1. econstructor. reflexivity. exists hex. apply fin_map_ext. intros x h5. terml h3. gterml. assert (he:x1 = x0 |-> x). unfold x1, x0. rewrite <- fin_map_new_ran_compat. rewrite <- fin_map_new_ran_compat; auto. assumption. fold x1. rewrite he. termr h3. gtermr. assert (h6:y0 = y |-> x). unfold y0, y. erewrite fun_in_ens_to_fin_map_compat. erewrite fun_in_ens_to_fin_map_compat. rewrite times_list_p_eq. f_equal. unfold ba_conv_list, transfer_dep, ba_conv_type, eq_rect_r. simpl. f_equal. fold y0. rewrite h6. unfold x0, y. rewrite h3; auto. Grab Existential Variables. assumption. Qed. Lemma plus_fin_pair_map2_functional_p : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pf:Finite D), forall (F G:Fin_map (cart_prod C D) E %0), F = G -> fin_map_eq (plus_fin_pair_map2_p pfdu pf F) (plus_fin_pair_map2_p pfdu pf G). intros; subst. red. exists (inclusion_reflexive _). apply fin_map_ext. intros x ha. symmetry. apply fin_map_new_ran_compat; auto. Qed. Definition fin_map_times_p {T:Type} {A:Ensemble T} {C:Ensemble Btp} (f:Fin_map A C %0) : Btp. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ A (fin_map_app f) h1) as h2. refine (times_set_p _ h2). Defined. Definition fin_map_times_p_eq : forall {T:Type} {A:Ensemble T} {C:Ensemble Btp} (f:Fin_map A C %0), fin_map_times_p f = fin_map_times (ba_conv_fin_map_ran f). intros. unfold fin_map_times_p. unfold fin_map_times. rewrite times_set_p_eq. reflexivity. Qed. Lemma fin_map_times_list_compat_p : forall {T:Type} {A:Ensemble T} {C:Ensemble Btp} (F:Fin_map A C %0) (nml:nice_map_lists F), fin_map_times_p F = times_list_p (n_im F nml). intros T A C F nml. unfold fin_map_times_p. pose proof (n_im_im_fin_map_compat F nml) as h1. apply times_set_compat_p'. rewrite h1. unfold im_fin_map. reflexivity. Qed. Lemma fin_map_times_empty1_p : forall {T:Type} {C:Ensemble Btp} (F:Fin_map (Empty_set T) C %0), fin_map_times_p F = %1. intros. rewrite fin_map_times_p_eq. apply (@fin_map_times_empty1 (ba_conv Bp)). Qed. Lemma fin_map_eq_times_p : forall {T:Type} (A:Ensemble T) (C E:Ensemble Btp) (F:Fin_map A C %0) (G:Fin_map A E %0), fin_map_eq F G -> fin_map_times_p F = fin_map_times_p G. intros. do 2 rewrite fin_map_times_p_eq. apply (@fin_map_eq_times (ba_conv Bp)). assumption. Qed. Lemma im2_empty1_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (D:Ensemble U) (C:Ensemble Btp) (pfd:Finite D) (pfc:Finite C) (y:U), im2 (cart_empty_map11 T U Btp pfdt pfdu %0 D C pfd pfc) y = Empty_set _. intros. apply (@im2_empty1 _ _ _ pfdt pfdu). Qed. Lemma im2_empty2_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (A:Ensemble T) (C:Ensemble Btp) (pfa:Finite A) (pfc:Finite C) (y:U), im2 (cart_empty_map21 T U Btp pfdt pfdu %0 A C pfa pfc) y = Empty_set _. intros. apply (@im2_empty2 _ _ _ pfdt pfdu). Qed. Lemma plus_fin_pair_map2_cart_empty_eq1_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (C:Ensemble Btp) (D:Ensemble U) (pfd:Finite D) (F:Fin_map (cart_prod (Empty_set T) D) C %0), fin_map_eq (plus_fin_pair_map2_p pfdu pfd F) (fun_to_fin_map pfdu D %0 pfd (fun x => %0)). intros. rewrite plus_fin_pair_map2_p_eq. apply (@plus_fin_pair_map2_cart_empty_eq1 (ba_conv Bp)); auto. Qed. Lemma plus_fin_pair_map2_cart_empty_p : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (C:Ensemble Btp) (pfe:Finite (Empty_set _)) (F:Fin_map (cart_prod A (Empty_set U)) C %0) (y:U), Finite A -> plus_fin_pair_map2_p pfdu pfe F |-> y = 0. intros. rewrite plus_fin_pair_map2_p_eq. apply (@plus_fin_pair_map2_cart_empty (ba_conv Bp)); auto. Qed. Lemma plus_fin_pair_map2_cart_empty_eq2_p : forall {T U:Type} (A:Ensemble T) (C:Ensemble Btp) (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (pfa:Finite A) (pfe:Finite (Empty_set U)) (F:Fin_map (cart_prod A (Empty_set U)) C %0), fin_map_eq (plus_fin_pair_map2_p pfdu pfe F) (empty_map1 U Btp pfdu %0 _ (fin_map_fin_ran F)). intros. rewrite plus_fin_pair_map2_p_eq. apply (@plus_fin_pair_map2_cart_empty_eq2 (ba_conv Bp)); auto. Qed. Lemma fin_map_times_sing_p : forall {T:Type} (pfdt:type_eq_dec T) (A:Ensemble T) (pf:Finite A) (val:Btp), A <> Empty_set _ -> fin_map_times_p (fin_map_sing pfdt A pf %0 val) = val. intros. rewrite fin_map_times_p_eq. apply (@fin_map_times_sing (ba_conv Bp)); auto. Qed. Lemma fin_map_times_cart_empty11_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (C:Ensemble Btp) (D:Ensemble U) (pfd:Finite D) (pfde:D <> Empty_set _) (F:Fin_map (cart_prod (Empty_set T) D) C %0), fin_map_times_p (plus_fin_pair_map2_p pfdu pfd F) = %0. intros. rewrite fin_map_times_p_eq. apply (@fin_map_times_cart_empty11 (ba_conv Bp)); auto. Qed. Lemma fin_map_times_cart_empty21_p : forall {T U:Type} (pfdu:type_eq_dec U) (A:Ensemble T) (C:Ensemble Btp) (pfa:Finite A) (pfe:Finite (Empty_set U)) (F:Fin_map (cart_prod A (Empty_set U)) C %0), fin_map_times_p (plus_fin_pair_map2_p pfdu pfe F) = %1. intros. rewrite fin_map_times_p_eq. apply (@fin_map_times_cart_empty21 (ba_conv Bp)); auto. Qed. Lemma fin_map_times_empty2_p : forall {T:Type} {A:Ensemble T} (F:Fin_map A (Empty_set Btp) %0), fin_map_times_p F = %1. intros. rewrite fin_map_times_p_eq. apply (@fin_map_times_empty2 (ba_conv Bp)); auto. Qed. Definition fin_map_plus_p {T:Type} {A:Ensemble T} {C:Ensemble Btp} (f:Fin_map A C %0) : Btp. pose proof (fin_map_fin_dom f) as h1. pose proof (finite_image _ _ A (fin_map_app f) h1) as h2. refine (plus_set_p _ h2). Defined. Definition fin_map_plus_p_eq : forall {T:Type} {A:Ensemble T} {C:Ensemble Btp} (f:Fin_map A C %0), fin_map_plus_p f = fin_map_plus (ba_conv_fin_map_ran f). intros. unfold fin_map_plus_p. unfold fin_map_plus. rewrite plus_set_p_eq. reflexivity. Qed. Lemma fin_map_plus_list_compat_p : forall {T:Type} {A:Ensemble T} {C:Ensemble Btp} (F:Fin_map A C %0) (nml:nice_map_lists F), fin_map_plus_p F = plus_list_p (n_im F nml). intros. rewrite fin_map_plus_p_eq. rewrite plus_list_p_eq. apply (@fin_map_plus_list_compat (ba_conv Bp)). Qed. Lemma fin_map_plus_empty1_p : forall {T:Type} {C:Ensemble Btp} (F:Fin_map (Empty_set T) C %0), fin_map_plus_p F = %0. intros. rewrite fin_map_plus_p_eq. apply (@fin_map_plus_empty1 (ba_conv Bp)). Qed. Lemma fin_map_plus_empty2_p : forall {T:Type} {A:Ensemble T} (F:Fin_map A (Empty_set Btp) %0), fin_map_plus_p F = %0. intros. rewrite fin_map_plus_p_eq. apply (@fin_map_plus_empty2 (ba_conv Bp)). Qed. Lemma fin_map_eq_plus_p : forall {T:Type} (A:Ensemble T) (C E:Ensemble Btp) (F:Fin_map A C %0) (G:Fin_map A E %0), fin_map_eq F G -> fin_map_plus_p F = fin_map_plus_p G. intros. do 2 rewrite fin_map_plus_p_eq. apply (@fin_map_eq_plus (ba_conv Bp)); auto. Qed. Definition times_plus_fin_pair_map1_p {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0) := fin_map_times_p (plus_fin_pair_map1_p pfdt pfc F). Lemma times_plus_fin_pair_map1_p_eq : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0), times_plus_fin_pair_map1_p pfdt pfc F = times_plus_fin_pair_map1 pfdt pfc (ba_conv_fin_map_ran F). intros. unfold times_plus_fin_pair_map1_p, times_plus_fin_pair_map1. rewrite fin_map_times_p_eq. reflexivity. Qed. Definition times_plus_fin_pair_map2_p {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0) := fin_map_times_p (plus_fin_pair_map2_p pfdu pfd F). Definition times_plus_fin_pair_map2_p_eq : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0), times_plus_fin_pair_map2_p pfdu pfd F = times_plus_fin_pair_map2 pfdu pfd (ba_conv_fin_map_ran F). intros. unfold times_plus_fin_pair_map2_p, times_plus_fin_pair_map2. rewrite fin_map_times_p_eq. reflexivity. Qed. Definition plus_times_fin_pair_map1_p {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfc:Finite C) (pfdt:type_eq_dec T) (F:Fin_map (cart_prod C D) E %0) := fin_map_plus_p (times_fin_pair_map1_p pfdt pfc F). Lemma plus_times_fin_pair_map1_p_eq : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0), plus_times_fin_pair_map1_p pfc pfdt F = plus_times_fin_pair_map1 pfdt pfc (ba_conv_fin_map_ran F). intros. unfold plus_times_fin_pair_map1_p, plus_times_fin_pair_map1. rewrite fin_map_plus_p_eq. reflexivity. Qed. Definition plus_times_fin_pair_map2_p {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0) := fin_map_plus_p (times_fin_pair_map2_p pfdu pfd F). Lemma plus_times_fin_pair_map2_p_eq : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} (pfdu:type_eq_dec U) {E:Ensemble Btp} (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0), plus_times_fin_pair_map2_p pfdu pfd F = plus_times_fin_pair_map2 pfdu pfd (ba_conv_fin_map_ran F). intros. unfold plus_times_fin_pair_map2_p, plus_times_fin_pair_map2. rewrite fin_map_plus_p_eq. reflexivity. Qed. Lemma times_plus_fin_pair_map1_list_compat_p : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0) (nml:nice_map_lists (plus_fin_pair_map1_p pfdt pfc F)), times_plus_fin_pair_map1_p pfdt pfc F = times_list_p (n_im _ nml). intros. rewrite times_plus_fin_pair_map1_p_eq. rewrite times_list_p_eq. apply (@times_plus_fin_pair_map1_list_compat (ba_conv Bp)). Qed. Lemma plus_fun_fin_map_to_fun_comm_p : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (F:Fin_map (cart_prod C D) E %0) (pfc:Finite C) (x:T) (lb:list U), list_to_set lb = D -> Inn C x -> plus_fun_p (fun y:U => fin_map_to_fun F (x, y)) lb = fun_to_fin_map pfdt C %0 pfc (fun x:T => plus_set_p (im1 F x) (finite_im1 F x)) |-> x. intros T U C D E hdt F h1 x lb h2 h3. pose proof (plus_fun_fin_map_to_fun_comm hdt (ba_conv_fin_map_ran F) h1 x lb h2 h3) as h4. rewrite plus_fun_p_eq. unfold ba_conv_ind. unfold ba_conv_type. rewrite transfer_dep_eq_refl. unfold ba_conv_fin_map_ran in h4. unfold eq_rect_r in h4. simpl in h4. unfold Btp, btp. unfold Btype_p. unfold ba_conv_set in h4. unfold transfer_dep in h4. unfold eq_rect_r in h4. simpl in h4. rewrite h4 at 1. assert (h5:(fun x0 : T => @plus_set (ba_conv Bp) (im1 F x0) (finite_im1 F x0)) = (fun x0 : T => plus_set_p (im1 F x0) (finite_im1 F x0))). apply functional_extensionality. intro. rewrite plus_set_p_eq. reflexivity. unfold Btp, btp, Btype_p in h5. unfold bt. simpl. simpl in h5. rewrite h5. reflexivity. Qed. Lemma plus_fun_fin_map_to_fun_comm_p' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (F:Fin_map (cart_prod C D) E %0) (pfd:Finite D) (y:U) (la:list T), list_to_set la = C -> Inn D y -> plus_fun_p (fun x:T => fin_map_to_fun F (x, y)) la = fun_to_fin_map pfdu D %0 pfd (fun x:U => plus_set_p (im2 F x) (finite_im2 F x)) |-> y. intros T U C D E hdu F h1 y la h2 h3. pose proof (plus_fun_fin_map_to_fun_comm' hdu (ba_conv_fin_map_ran F) h1 y la h2 h3) as h4. rewrite plus_fun_p_eq. unfold ba_conv_ind. unfold ba_conv_type. rewrite transfer_dep_eq_refl. unfold ba_conv_fin_map_ran in h4. unfold eq_rect_r in h4. simpl in h4. unfold Btp, btp. unfold Btype_p. unfold ba_conv_set in h4. unfold transfer_dep in h4. unfold eq_rect_r in h4. simpl in h4. rewrite h4 at 1. assert (h5:(fun x : U => @plus_set (ba_conv Bp) (im2 F x) (finite_im2 F x)) = (fun x : U => plus_set_p (im2 F x) (finite_im2 F x))). apply functional_extensionality. intro x. rewrite plus_set_p_eq. reflexivity. unfold Btp, btp in h5. unfold Btype_p in h5. unfold bt. simpl. simpl in h5. rewrite h5. reflexivity. Qed. Lemma times_plus_fin_pair_map1_list_compat_p' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0) (nml2:nice_map_lists2 F) (nml:nice_map_lists (plus_fin_pair_map1_p pfdt pfc F)), (n_la2 _ nml2) = (n_la _ nml) -> times_list_p (n_im _ nml) = times_plus_fun1_p (n_la2 _ nml2) (n_lb2 _ nml2) (f_no_pr (fin_map_to_fun F)). intros. rewrite times_list_p_eq, times_plus_fun1_p_eq. apply (@times_plus_fin_pair_map1_list_compat' (ba_conv Bp)); auto. Qed. Lemma times_plus_fin_pair_map2_list_compat_p : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0) (nml:nice_map_lists (plus_fin_pair_map2_p pfdu pfd F)), times_plus_fin_pair_map2_p pfdu pfd F = times_list_p (n_im _ nml). intros. rewrite times_plus_fin_pair_map2_p_eq, times_list_p_eq. apply (@times_plus_fin_pair_map2_list_compat (ba_conv Bp)). Qed. Lemma times_plus_fin_pair_map2_list_compat_p' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0) (nml2:nice_map_lists2 F) (nml:nice_map_lists (plus_fin_pair_map2_p pfdu pfd F)), (n_lb2 _ nml2) = (n_la _ nml) -> times_list_p (n_im _ nml) = times_plus_fun2_p (n_la2 _ nml2) (n_lb2 _ nml2) (f_no_pr (fin_map_to_fun F)). intros. rewrite times_list_p_eq, times_plus_fun2_p_eq. apply (@times_plus_fin_pair_map2_list_compat' (ba_conv Bp)); auto. Qed. Lemma plus_times_fin_pair_map1_list_compat_p : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0) (nml:nice_map_lists (times_fin_pair_map1_p pfdt pfc F)), plus_times_fin_pair_map1_p pfc pfdt F = plus_list_p (n_im _ nml). intros T U C D E pfc F nml. apply fin_map_plus_list_compat_p. Qed. Lemma times_fun_fin_map_to_fun_comm_p : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (F:Fin_map (cart_prod C D) E %0) (pfc:Finite C) (x:T) (lb:list U), list_to_set lb = D -> Inn C x -> times_fun_p (fun y:U => fin_map_to_fun F (x, y)) lb = fun_to_fin_map pfdt C %0 pfc (fun x:T => times_set_p (im1 F x) (finite_im1 F x)) |-> x. intros T U C D E hdt F h1 x lb h2 h3. pose proof (times_fun_fin_map_to_fun_comm hdt (ba_conv_fin_map_ran F) h1 x lb h2 h3) as h4. rewrite times_fun_p_eq. unfold ba_conv_ind. unfold ba_conv_type. rewrite transfer_dep_eq_refl. unfold ba_conv_fin_map_ran in h4. unfold eq_rect_r in h4. simpl in h4. unfold Btp, btp. unfold Btype_p. unfold ba_conv_set in h4. unfold transfer_dep in h4. unfold eq_rect_r in h4. simpl in h4. rewrite h4 at 1. assert (h5:(fun x0 : T => @times_set (ba_conv Bp) (im1 F x0) (finite_im1 F x0)) = (fun x0 : T => times_set_p (im1 F x0) (finite_im1 F x0))). apply functional_extensionality. intro. rewrite times_set_p_eq. reflexivity. unfold Btp, btp in h5. unfold Btype_p in h5. unfold bt. simpl in h5. simpl. rewrite h5. reflexivity. Qed. Lemma plus_times_fin_pair_map1_list_compat_p' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdt:type_eq_dec T) (pfc:Finite C) (F:Fin_map (cart_prod C D) E %0) (nml2:nice_map_lists2 F) (nml:nice_map_lists (times_fin_pair_map1_p pfdt pfc F)), (n_la2 _ nml2) = (n_la _ nml) -> plus_list_p (n_im _ nml) = plus_times_fun1_p (n_la2 _ nml2) (n_lb2 _ nml2) (f_no_pr (fin_map_to_fun F)). intros. rewrite plus_list_p_eq, plus_times_fun1_p_eq. apply (@plus_times_fin_pair_map1_list_compat' (ba_conv Bp)); auto. Qed. Lemma plus_times_fin_pair_map2_list_compat_p : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0) (nml:nice_map_lists (times_fin_pair_map2_p pfdu pfd F)), plus_times_fin_pair_map2_p pfdu pfd F = plus_list_p (n_im _ nml). intros T U C D E pfc F nml. apply fin_map_plus_list_compat_p. Qed. Lemma times_fun_fin_map_to_fun_comm_p' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (F:Fin_map (cart_prod C D) E %0) (pfd:Finite D) (y:U) (la:list T), list_to_set la = C -> Inn D y -> times_fun_p (fun x:T => fin_map_to_fun F (x, y)) la = fun_to_fin_map pfdu D %0 pfd (fun x:U => times_set_p (im2 F x) (finite_im2 F x)) |-> y. intros T U C D E hdu F h1 y la h2 h3. rewrite times_fun_p_eq. pose proof (times_fun_fin_map_to_fun_comm' hdu (ba_conv_fin_map_ran F) h1 y la h2 h3) as h4. unfold ba_conv_ind. unfold ba_conv_type. rewrite transfer_dep_eq_refl. unfold ba_conv_fin_map_ran in h4. unfold eq_rect_r in h4. simpl in h4. unfold Btp, bt. unfold Btype_p. unfold ba_conv_set in h4. unfold transfer_dep in h4. unfold eq_rect_r in h4. simpl in h4. rewrite h4 at 1. assert (h5: (fun x : U => @times_set (ba_conv Bp) (im2 F x) (finite_im2 F x)) = (fun x : U => times_set_p (im2 F x) (finite_im2 F x))). apply functional_extensionality. intro. rewrite times_set_p_eq. reflexivity. unfold Btp, btp in h5. unfold Btype_p in h5. unfold bt. simpl in h5. simpl. rewrite h5. reflexivity. Qed. Lemma plus_times_fin_pair_map2_list_compat_p' : forall {T U:Type} {C:Ensemble T} {D:Ensemble U} {E:Ensemble Btp} (pfdu:type_eq_dec U) (pfd:Finite D) (F:Fin_map (cart_prod C D) E %0) (nml2:nice_map_lists2 F) (nml:nice_map_lists (times_fin_pair_map2_p pfdu pfd F)), (n_lb2 _ nml2) = (n_la _ nml) -> plus_list_p (n_im _ nml) = plus_times_fun2_p (n_la2 _ nml2) (n_lb2 _ nml2) (f_no_pr (fin_map_to_fun F)). intros. rewrite plus_list_p_eq, plus_times_fun2_p_eq. apply (@plus_times_fin_pair_map2_list_compat' (ba_conv Bp)); auto. Qed. Definition times_fun_fin_map1_p {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Btp) (a:Fin_map I J def) : Btp. pose proof (fin_map_fin_dom a) as h1. pose proof (fin_map_fin_ran a) as h2. pose proof (cart_prod_fin _ _ h1 h2) as h4. pose (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu) _ 0 h4 (ba_conv_ind p)) as P. refine (fin_map_times_p (fun_to_fin_map (fin_map_type_eq_dec a) I 0 h1 (fun i:T => (P |-> (i, (a |-> i)))))). Defined. Lemma times_fun_fin_map1_p_eq : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Btp) (a:Fin_map I J def), times_fun_fin_map1_p pfdu p a = times_fun_fin_map1 pfdu (ba_conv_ind p) a. intros. unfold times_fun_fin_map1_p, times_fun_fin_map1. rewrite fin_map_times_p_eq. reflexivity. Qed. Definition times_fun_fin_map2_p {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Btp) (a:Fin_map J I def) : Btp. pose proof (fin_map_fin_dom a) as h1. pose proof (fin_map_fin_ran a) as h2. pose proof (cart_prod_fin _ _ h2 h1) as h4. pose (fun_to_fin_map (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 h4 (ba_conv_ind p)) as P. refine (fin_map_times_p (fun_to_fin_map (fin_map_type_eq_dec a) J 0 h1 (fun j:U => (P |-> (a |-> j, j))))). Defined. Lemma times_fun_fin_map2_p_eq : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Btp) (a:Fin_map J I def), times_fun_fin_map2_p pfdt p a = times_fun_fin_map2 pfdt (ba_conv_ind p) a. intros. unfold times_fun_fin_map2_p, times_fun_fin_map2. rewrite fin_map_times_p_eq. reflexivity. Qed. Definition times_fun_fin_map1_l_p {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Btp) (a:Fin_map I J def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu) _ 0 (cart_prod_fin _ _ (fin_map_fin_dom a) (fin_map_fin_ran a)) (ba_conv_ind p))) := times_list_p (map (fun i:T => (p (i, a |-> i))) (n_la2 _ nml2)). Lemma times_fun_fin_map1_l_p_eq : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Btp) (a:Fin_map I J def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu) _ 0 (cart_prod_fin _ _ (fin_map_fin_dom a) (fin_map_fin_ran a)) (ba_conv_ind p))), times_fun_fin_map1_l_p pfdu p a nml2 = times_fun_fin_map1_l pfdu (ba_conv_ind p) a nml2. intros. unfold times_fun_fin_map1_l_p. unfold times_fun_fin_map1_l. rewrite times_list_p_eq. reflexivity. Qed. Definition times_fun_fin_map2_l_p {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Btp) (a:Fin_map J I def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 (cart_prod_fin _ _ (fin_map_fin_ran a) (fin_map_fin_dom a)) (ba_conv_ind p))) := times_list_p (map (fun j:U => (p (a |-> j, j))) (n_lb2 _ nml2)). Lemma times_fun_fin_map2_l_p_eq : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Btp) (a:Fin_map J I def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 (cart_prod_fin _ _ (fin_map_fin_ran a) (fin_map_fin_dom a)) (ba_conv_ind p))), times_fun_fin_map2_l_p pfdt p a nml2 = times_fun_fin_map2_l pfdt (ba_conv_ind p) a nml2. intros. unfold times_fun_fin_map2_l_p, times_fun_fin_map2_l. rewrite times_list_p_eq. reflexivity. Qed. Lemma im_sing_times_fun_fin_map1_p : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:T*U->Btp) (a:Fin_map I J def), Im (Singleton a) (times_fun_fin_map1_p pfdu p) = Singleton ((times_fun_fin_map1_p pfdu p) a). intros T U I J def hdu p a. rewrite times_fun_fin_map1_p_eq. pose proof (im_sing_times_fun_fin_map1 hdu (ba_conv_ind p) a) as h1. simpl in h1. unfold Btp, btp. unfold Btype_p. rewrite <- h1 at 1. f_equal. apply functional_extensionality. intro F. rewrite times_fun_fin_map1_p_eq. reflexivity. Qed. Lemma times_fun_fin_map1_empty_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:U) (p:(T*U)->Btp), times_fun_fin_map1_p pfdu p (empty_map T U pfdt def) = %1. intros. rewrite times_fun_fin_map1_p_eq. apply (@times_fun_fin_map1_empty (ba_conv Bp)). Qed. Lemma times_fun_fin_map2_empty_p : forall {T U:Type} (def:T) (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (p:(T*U)->Btp), times_fun_fin_map2_p pfdt p (empty_map U T pfdu def) = %1. intros. rewrite times_fun_fin_map2_p_eq. apply (@times_fun_fin_map2_empty (ba_conv Bp)). Qed. Lemma times_fun_fin_map1_empty1_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:U) (p:(T*U)->Btp) (J:Ensemble U) (pf:Finite J), times_fun_fin_map1_p pfdu p (empty_map1 T U pfdt def J pf) = %1. intros. rewrite times_fun_fin_map1_p_eq. apply (@times_fun_fin_map1_empty1 (ba_conv Bp)). Qed. Lemma times_fun_fin_map2_empty1_p : forall {T U:Type} (def:T) (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (p:(T*U)->Btp) (I:Ensemble T) (pf:Finite I), times_fun_fin_map2_p pfdt p (empty_map1 U T pfdu def I pf) = %1. intros. rewrite times_fun_fin_map2_p_eq. apply (@times_fun_fin_map2_empty1 (ba_conv Bp)). Qed. Lemma times_fun_fin_map1_list_compat_p : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Btp) (a:Fin_map I J def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu) _ 0 (cart_prod_fin _ _ (fin_map_fin_dom a) (fin_map_fin_ran a)) (ba_conv_ind p))), times_fun_fin_map1_p pfdu p a = times_fun_fin_map1_l_p pfdu p a nml2. intros. rewrite times_fun_fin_map1_p_eq, times_fun_fin_map1_l_p_eq. apply (@times_fun_fin_map1_list_compat (ba_conv Bp)). Qed. Lemma times_fun_fin_map2_list_compat_p : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Btp) (a:Fin_map J I def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 (cart_prod_fin _ _ (fin_map_fin_ran a) (fin_map_fin_dom a)) (ba_conv_ind p))), times_fun_fin_map2_p pfdt p a = times_fun_fin_map2_l_p pfdt p a nml2. intros. rewrite times_fun_fin_map2_p_eq, times_fun_fin_map2_l_p_eq. apply (@times_fun_fin_map2_list_compat (ba_conv Bp)). Qed. Definition plus_fun_fin_map1_p {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Btp) (a:Fin_map I J def) : Btp. pose proof (fin_map_fin_dom a) as h1. pose proof (fin_map_fin_ran a) as h2. pose proof (cart_prod_fin _ _ h1 h2) as h4. pose (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu) _ 0 h4 (ba_conv_ind p)) as P. refine (fin_map_plus_p (fun_to_fin_map (fin_map_type_eq_dec a) I 0 h1 (fun i:T => (P |-> (i, (a |-> i)))))). Defined. Lemma plus_fun_fin_map1_p_eq : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Btp) (a:Fin_map I J def), plus_fun_fin_map1_p pfdu p a = plus_fun_fin_map1 pfdu (ba_conv_ind p) a. intros. unfold plus_fun_fin_map1_p, plus_fun_fin_map1. rewrite fin_map_plus_p_eq. reflexivity. Qed. Definition plus_fun_fin_map2_p {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pf:type_eq_dec T) (p:(T*U)->Btp) (a:Fin_map J I def) : Btp. pose proof (fin_map_fin_dom a) as h1. pose proof (fin_map_fin_ran a) as h2. pose proof (cart_prod_fin _ _ h2 h1) as h4. pose (fun_to_fin_map (impl_type_eq_dec_prod pf (fin_map_type_eq_dec a)) _ 0 h4 (ba_conv_ind p)) as P. refine (fin_map_plus_p (fun_to_fin_map (fin_map_type_eq_dec a) J 0 h1 (fun j:U => (P |-> (a |-> j, j))))). Defined. Lemma plus_fun_fin_map2_p_eq : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Btp) (a:Fin_map J I def), plus_fun_fin_map2_p pfdt p a = plus_fun_fin_map2 pfdt (ba_conv_ind p) a. intros. unfold plus_fun_fin_map2_p, plus_fun_fin_map2. rewrite fin_map_plus_p_eq. reflexivity. Qed. Lemma plus_fun_fin_map1_empty_p : forall {T U:Type} (def:U) (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (p:(T*U)->Btp), plus_fun_fin_map1_p pfdu p (empty_map _ U pfdt def) = %0. intros. rewrite plus_fun_fin_map1_p_eq. apply (@plus_fun_fin_map1_empty (ba_conv Bp)). Qed. Lemma plus_fun_fin_map1_empty1_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:U) (p:(T*U)->Btp) (J:Ensemble U) (pf:Finite J), plus_fun_fin_map1_p pfdu p (empty_map1 T U pfdt def J pf) = %0. intros. rewrite plus_fun_fin_map1_p_eq. apply (@plus_fun_fin_map1_empty1 (ba_conv Bp)). Qed. Definition plus_fun_fin_map1_l_p {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Btp) (a:Fin_map I J def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu) _ 0 (cart_prod_fin _ _ (fin_map_fin_dom a) (fin_map_fin_ran a)) (ba_conv_ind p))) := plus_list_p (map (fun i:T => (p (i, a |-> i))) (n_la2 _ nml2)). Lemma plus_fun_fin_map1_l_p_eq : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Btp) (a:Fin_map I J def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu) _ 0 (cart_prod_fin _ _ (fin_map_fin_dom a) (fin_map_fin_ran a)) (ba_conv_ind p))), plus_fun_fin_map1_l_p pfdu p a nml2 = plus_fun_fin_map1_l pfdu (ba_conv_ind p) a nml2. intros. unfold plus_fun_fin_map1_l_p, plus_fun_fin_map1_l. rewrite plus_list_p_eq. reflexivity. Qed. Definition plus_fun_fin_map2_l_p {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Btp) (a:Fin_map J I def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 (cart_prod_fin _ _ (fin_map_fin_ran a) (fin_map_fin_dom a)) (ba_conv_ind p))) := plus_list_p (map (fun j:U => (p (a |-> j, j))) (n_lb2 _ nml2)). Lemma plus_fun_fin_map2_l_p_eq : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Btp) (a:Fin_map J I def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 (cart_prod_fin _ _ (fin_map_fin_ran a) (fin_map_fin_dom a)) (ba_conv_ind p))), plus_fun_fin_map2_l_p pfdt p a nml2 = plus_fun_fin_map2_l pfdt (ba_conv_ind p) a nml2. intros. unfold plus_fun_fin_map2_l_p, plus_fun_fin_map2_l. rewrite plus_list_p_eq. reflexivity. Qed. Lemma plus_fun_fin_map2_empty_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:T) (p:(T*U)->Btp), plus_fun_fin_map2_p pfdt p (empty_map U T pfdu def) = %0. intros. rewrite plus_fun_fin_map2_p_eq. apply (@plus_fun_fin_map2_empty (ba_conv Bp)). Qed. Lemma plus_fun_fin_map2_empty1_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:T) (p:(T*U)->Btp) (I:Ensemble T) (pf:Finite I), plus_fun_fin_map2_p pfdt p (empty_map1 U T pfdu def I pf) = %0. intros. rewrite plus_fun_fin_map2_p_eq. apply (@plus_fun_fin_map2_empty1 (ba_conv Bp)). Qed. Lemma plus_fun_fin_map1_list_compat_p : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:U} (pfdu:type_eq_dec U) (p:(T*U)->Btp) (a:Fin_map I J def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod (fin_map_type_eq_dec a) pfdu) _ 0 (cart_prod_fin _ _ (fin_map_fin_dom a) (fin_map_fin_ran a)) (ba_conv_ind p))), plus_fun_fin_map1_p pfdu p a = plus_fun_fin_map1_l_p pfdu p a nml2. intros. rewrite plus_fun_fin_map1_p_eq, plus_fun_fin_map1_l_p_eq. apply (@plus_fun_fin_map1_list_compat (ba_conv Bp)). Qed. Lemma plus_fun_fin_map2_list_compat_p : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} {def:T} (pfdt:type_eq_dec T) (p:(T*U)->Btp) (a:Fin_map J I def) (nml2:nice_map_lists2 (fun_to_fin_map (impl_type_eq_dec_prod pfdt (fin_map_type_eq_dec a)) _ 0 (cart_prod_fin _ _ (fin_map_fin_ran a) (fin_map_fin_dom a)) (ba_conv_ind p))), plus_fun_fin_map2_p pfdt p a = plus_fun_fin_map2_l_p pfdt p a nml2. intros. rewrite plus_fun_fin_map2_p_eq, plus_fun_fin_map2_l_p_eq. apply (@plus_fun_fin_map2_list_compat (ba_conv Bp)). Qed. Definition plus_times_fun_all_maps1_p {T U:Type} (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:U) (pfi:Finite I) (pfj:Finite J) (p:(T*U)->Btp): Btp := let f:= (@times_fun_fin_map1_p _ _ I J def pfdu p) in let S := Full_set (Fin_map I J def) in plus_set_p (Im S f) (finite_image _ _ S f (finite_fin_maps _ _ def pfi pfj)). Lemma plus_times_fun_all_maps1_p_eq : forall {T U:Type} (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:U) (pfi:Finite I) (pfj:Finite J) (p:(T*U)->Btp), plus_times_fun_all_maps1_p pfdu I J def pfi pfj p = plus_times_fun_all_maps1 pfdu I J def pfi pfj (ba_conv_ind p). intros. unfold plus_times_fun_all_maps1_p, plus_times_fun_all_maps1. rewrite plus_set_p_eq. assert (h1:(ba_conv_set (Im (Full_set (Fin_map I J def)) (times_fun_fin_map1_p pfdu p))) = (Im (Full_set (Fin_map I J def)) (times_fun_fin_map1 pfdu (ba_conv_ind p)))). apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. subst. apply Im_intro with x; auto. rewrite times_fun_fin_map1_p_eq. reflexivity. red. intros x h1. destruct h1 as [x h1]. subst. apply Im_intro with x; auto. rewrite times_fun_fin_map1_p_eq. reflexivity. apply (plus_set_functional _ _ h1). Qed. Definition plus_times_fun_all_maps2_p {T U:Type} (pfdt:type_eq_dec T) (I:Ensemble T) (J:Ensemble U) (def:T) (pfi:Finite I) (pfj:Finite J) (p:(T*U)->Btp): Btp := let f:= (@times_fun_fin_map2_p T U I J def pfdt p) in let S := Full_set (Fin_map J I def) in plus_set_p (Im S f) (finite_image _ _ S f (finite_fin_maps _ _ def pfj pfi)). Lemma plus_times_fun_all_maps2_p_eq : forall {T U:Type} (pfdt:type_eq_dec T) (I:Ensemble T) (J:Ensemble U) (def:T) (pfi:Finite I) (pfj:Finite J) (p:(T*U)->Btp), plus_times_fun_all_maps2_p pfdt I J def pfi pfj p = plus_times_fun_all_maps2 pfdt I J def pfi pfj (ba_conv_ind p). intros. unfold plus_times_fun_all_maps2_p, plus_times_fun_all_maps1. rewrite plus_set_p_eq. assert (h1:(ba_conv_set (Im (Full_set (Fin_map J I def)) (times_fun_fin_map2_p pfdt p))) = (Im (Full_set (Fin_map J I def)) (times_fun_fin_map2 pfdt (ba_conv_ind p)))). apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. subst. apply Im_intro with x; auto. rewrite times_fun_fin_map2_p_eq. reflexivity. red. intros x h1. destruct h1 as [x h1]. subst. apply Im_intro with x; auto. rewrite times_fun_fin_map2_p_eq. reflexivity. apply (plus_set_functional _ _ h1 _ _). Qed. Definition times_plus_fun_all_maps1_p {T U:Type} (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:U) (pfi:Finite I) (pfj:Finite J) (p:T*U->Btp) : Btp := let f:= (@plus_fun_fin_map1_p T U I J def pfdu p) in let S := Full_set (Fin_map I J def) in times_set_p (Im S f) (finite_image _ _ S f (finite_fin_maps _ _ def pfi pfj)). Lemma times_plus_fun_all_maps1_p_eq : forall {T U:Type} (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:U) (pfi:Finite I) (pfj:Finite J) (p:T*U->Btp), times_plus_fun_all_maps1_p pfdu I J def pfi pfj p = times_plus_fun_all_maps1 pfdu I J def pfi pfj (ba_conv_ind p). intros. unfold times_plus_fun_all_maps1_p, times_plus_fun_all_maps1. rewrite times_set_p_eq. assert (h1:(ba_conv_set (Im (Full_set (Fin_map I J def)) (plus_fun_fin_map1_p pfdu p))) = (Im (Full_set (Fin_map I J def)) (plus_fun_fin_map1 pfdu (ba_conv_ind p)))). apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. subst. apply Im_intro with x; auto. rewrite plus_fun_fin_map1_p_eq. reflexivity. red. intros x h1. destruct h1 as [x h1]. subst. apply Im_intro with x; auto. rewrite plus_fun_fin_map1_p_eq. reflexivity. apply (times_set_functional _ _ h1). Qed. Definition times_plus_fun_all_maps2_p {T U:Type} (pfdt:type_eq_dec T) (I:Ensemble T) (J:Ensemble U) (def:T) (pfi:Finite I) (pfj:Finite J) (p:T*U->Btp) : Btp := let f:= (@plus_fun_fin_map2_p T U I J def pfdt p) in let S := Full_set (Fin_map J I def) in times_set_p (Im S f) (finite_image _ _ S f (finite_fin_maps _ _ def pfj pfi)). Lemma times_plus_fun_all_maps2_p_eq : forall {T U:Type} (pfdt:type_eq_dec T) (I:Ensemble T) (J:Ensemble U) (def:T) (pfi:Finite I) (pfj:Finite J) (p:T*U->Btp), times_plus_fun_all_maps2_p pfdt I J def pfi pfj p = times_plus_fun_all_maps2 pfdt I J def pfi pfj (ba_conv_ind p). intros. unfold times_plus_fun_all_maps2_p, times_plus_fun_all_maps2. rewrite times_set_p_eq. assert (h1:(ba_conv_set (Im (Full_set (Fin_map J I def)) (plus_fun_fin_map2_p pfdt p))) = (Im (Full_set (Fin_map J I def)) (plus_fun_fin_map2 pfdt (ba_conv_ind p)))). apply Extensionality_Ensembles. red. split. red. intros x h1. destruct h1 as [x h1]. subst. apply Im_intro with x; auto. rewrite plus_fun_fin_map2_p_eq. reflexivity. red. intros x h1. destruct h1 as [x h1]. subst. apply Im_intro with x; auto. rewrite plus_fun_fin_map2_p_eq. reflexivity. apply (times_set_functional _ _ h1). Qed. Lemma complete_dist_list_times1_p' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp) (def:U), times_plus_fun1_p li lj p = plus_times_all_funs1_p pfdt pfdu li lj pfi pfj p. intros. rewrite times_plus_fun1_p_eq, plus_times_all_funs1_p_eq. apply (@complete_dist_list_times1' (ba_conv Bp)). Qed. Lemma complete_dist_list_times2_p' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp) (def:T), times_plus_fun2_p li lj p = plus_times_all_funs2_p pfdt pfdu li lj pfi pfj p. intros. rewrite times_plus_fun2_p_eq, plus_times_all_funs2_p_eq. apply (@complete_dist_list_times2' (ba_conv Bp)). refine def. Qed. Lemma complete_dist_list_plus1_p' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp) (def:U), plus_times_fun1_p li lj p = times_plus_all_funs1_p pfdt pfdu li lj pfi pfj p. intros. rewrite plus_times_fun1_p_eq, times_plus_all_funs1_p_eq. apply (@complete_dist_list_plus1' (ba_conv Bp)). Qed. Lemma complete_dist_list_plus2_p' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pfi:NoDup li) (pfj:NoDup lj) (p:T->U->Btp) (def:T), plus_times_fun2_p li lj p = times_plus_all_funs2_p pfdt pfdu li lj pfi pfj p. intros. rewrite plus_times_fun2_p_eq, times_plus_all_funs2_p_eq. apply (@complete_dist_list_plus2' (ba_conv Bp)). refine def. Qed. Lemma plus_times_all_maps1_funs_compat_p : forall {T U:Type} {I:Ensemble T} {J:Ensemble U} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (def:U) (pfi:Finite I) (pfj: Finite J) (li:list T) (lj:list U) (pfndpi: NoDup li) (pfndpj: NoDup lj) (p:T->U->Btp), list_to_set li = I -> list_to_set lj = J -> plus_times_fun_all_maps1_p pfdu _ _ def pfi pfj (f_pr p) = plus_times_all_funs1_p pfdt pfdu _ _ pfndpi pfndpj p. intros. rewrite plus_times_fun_all_maps1_p_eq, plus_times_all_funs1_p_eq. apply (@plus_times_all_maps1_funs_compat (ba_conv Bp)); auto. Qed. Lemma plus_times_all_maps2_funs_compat_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:T) (pfi:Finite I) (pfj: Finite J) (li:list T) (lj:list U) (pfndpi: NoDup li) (pfndpj: NoDup lj) (p:T->U->Btp), list_to_set li = I -> list_to_set lj = J -> plus_times_fun_all_maps2_p pfdt _ _ def pfi pfj (f_pr p) = plus_times_all_funs2_p pfdt pfdu _ _ pfndpi pfndpj p. intros. rewrite plus_times_fun_all_maps2_p_eq, plus_times_all_funs2_p_eq. apply (@plus_times_all_maps2_funs_compat (ba_conv Bp)); auto. Qed. Lemma times_plus_all_maps1_funs_compat_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:U) (pfi:Finite I) (pfj: Finite J) (li:list T) (lj:list U) (pfndpi: NoDup li) (pfndpj: NoDup lj) (p:T->U->Btp), list_to_set li = I -> list_to_set lj = J -> times_plus_fun_all_maps1_p pfdu _ _ def pfi pfj (f_pr p) = times_plus_all_funs1_p pfdt pfdu _ _ pfndpi pfndpj p. intros. rewrite times_plus_fun_all_maps1_p_eq, times_plus_all_funs1_p_eq. apply (@times_plus_all_maps1_funs_compat (ba_conv Bp)); auto. Qed. Lemma times_plus_all_maps2_funs_compat_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:T) (pfi:Finite I) (pfj: Finite J) (li:list T) (lj:list U) (pfndpi: NoDup li) (pfndpj: NoDup lj) (p:T->U->Btp), list_to_set li = I -> list_to_set lj = J -> times_plus_fun_all_maps2_p pfdt _ _ def pfi pfj (f_pr p) = times_plus_all_funs2_p pfdt pfdu _ _ pfndpi pfndpj p. intros. rewrite times_plus_fun_all_maps2_p_eq, times_plus_all_funs2_p_eq. apply (@times_plus_all_maps2_funs_compat (ba_conv Bp)); auto. Qed. Lemma times_plus_fun1_fin_map_to_fun_undoes_fun_to_fin_map_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (li:list T) (lj:list U) (pf:Finite (cart_prod I J)) (p:T*U->Btp), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in list_to_set li = I -> list_to_set lj = J -> times_plus_fun1_p li lj (f_no_pr (fin_map_to_fun (fun_to_fin_map pfdtu (cart_prod I J) 0 pf (ba_conv_ind p)))) = times_plus_fun1_p li lj (f_no_pr p). intros. do 2 rewrite times_plus_fun1_p_eq. apply (@times_plus_fun1_fin_map_to_fun_undoes_fun_to_fin_map (ba_conv Bp)); auto. Qed. Lemma times_plus_fun2_fin_map_to_fun_undoes_fun_to_fin_map_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (li:list T) (lj:list U) (pf:Finite (cart_prod I J)) (p:T*U->Btp), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in list_to_set li = I -> list_to_set lj = J -> times_plus_fun2_p li lj (f_no_pr (fin_map_to_fun (fun_to_fin_map pfdtu (cart_prod I J) 0 pf (ba_conv_ind p)))) = times_plus_fun2_p li lj (f_no_pr p). intros. do 2 rewrite times_plus_fun2_p_eq. apply (@times_plus_fun2_fin_map_to_fun_undoes_fun_to_fin_map (ba_conv Bp)); auto. Qed. Lemma plus_times_fun1_fin_map_to_fun_undoes_fun_to_fin_map_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (li:list T) (lj:list U) (pf:Finite (cart_prod I J)) (p:T*U->Btp), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in list_to_set li = I -> list_to_set lj = J -> plus_times_fun1_p li lj (f_no_pr (fin_map_to_fun (fun_to_fin_map pfdtu (cart_prod I J) 0 pf (ba_conv_ind p)))) = plus_times_fun1_p li lj (f_no_pr p). intros. do 2 rewrite plus_times_fun1_p_eq. apply (@plus_times_fun1_fin_map_to_fun_undoes_fun_to_fin_map (ba_conv Bp)); auto. Qed. Lemma plus_times_fun2_fin_map_to_fun_undoes_fun_to_fin_map_p : forall {T U:Type} (I:Ensemble T) (J:Ensemble U) (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (li:list T) (lj:list U) (pf:Finite (cart_prod I J)) (p:T*U->Btp), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in list_to_set li = I -> list_to_set lj = J -> plus_times_fun2_p li lj (f_no_pr (fin_map_to_fun (fun_to_fin_map pfdtu (cart_prod I J) 0 pf (ba_conv_ind p)))) = plus_times_fun2_p li lj (f_no_pr p). intros. do 2 rewrite plus_times_fun2_p_eq. apply (@plus_times_fun2_fin_map_to_fun_undoes_fun_to_fin_map (ba_conv Bp)); auto. Qed. Lemma complete_dist_times_plus1_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:U) (pfi:Finite I) (pfj:Finite J) (p:T*U->Btp), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in times_plus_fin_pair_map1_p pfdt pfi (fun_to_fin_map pfdtu _ 0 (cart_prod_fin _ _ pfi pfj) (ba_conv_ind p)) = plus_times_fun_all_maps1_p pfdu _ _ def pfi pfj p. intros. rewrite times_plus_fin_pair_map1_p_eq, plus_times_fun_all_maps1_p_eq. apply (@complete_dist_times_plus1 (ba_conv Bp)). Qed. Lemma complete_dist_times_plus2_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:T) (pfi:Finite I) (pfj:Finite J) (p:T*U->Btp), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in times_plus_fin_pair_map2_p pfdu pfj (fun_to_fin_map pfdtu _ 0 (cart_prod_fin _ _ pfi pfj) (ba_conv_ind p)) = plus_times_fun_all_maps2_p pfdt _ _ def pfi pfj p. intros. rewrite times_plus_fin_pair_map2_p_eq, plus_times_fun_all_maps2_p_eq. apply (@complete_dist_times_plus2 (ba_conv Bp)). Qed. Lemma complete_dist_plus_times1_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:U) (pfi:Finite I) (pfj:Finite J) (p:T*U->Btp), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in plus_times_fin_pair_map1_p pfi pfdt (fun_to_fin_map pfdtu _ 0 (cart_prod_fin _ _ pfi pfj) (ba_conv_ind p)) = times_plus_fun_all_maps1_p pfdu _ _ def pfi pfj p. intros. rewrite plus_times_fin_pair_map1_p_eq, times_plus_fun_all_maps1_p_eq. apply (@complete_dist_plus_times1 (ba_conv Bp)). Qed. Lemma complete_dist_plus_times2_p : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (I:Ensemble T) (J:Ensemble U) (def:T) (pfi:Finite I) (pfj:Finite J) (p:T*U->Btp), let pfdtu := impl_type_eq_dec_prod pfdt pfdu in plus_times_fin_pair_map2_p pfdu pfj (fun_to_fin_map pfdtu _ 0 (cart_prod_fin _ _ pfi pfj) (ba_conv_ind p)) = times_plus_fun_all_maps2_p pfdt _ _ def pfi pfj p. intros. rewrite plus_times_fin_pair_map2_p_eq, times_plus_fun_all_maps2_p_eq. apply (@complete_dist_plus_times2 (ba_conv Bp)). Qed. Lemma ba_p_subst_times_set : forall (Ap:Bool_Alg_p T'), Ap = Bp -> forall (D:Ensemble Btp) (pfd:Finite D) (pfa:Inn (ba_p_ens Ap) (proj1_sig (times_set_p D pfd))) (pfb:Inn (ba_p_ens Bp) (proj1_sig (times_set_p D pfd))), proj1_sig (exist _ _ pfa) = proj1_sig (exist _ _ pfb). intros Ap h1 D h2 h3 h4. subst. reflexivity. Qed. Lemma ba_p_subst_plus_set : forall (Ap:Bool_Alg_p T'), Ap = Bp -> forall (D:Ensemble Btp) (pfd:Finite D) (pfa:Inn (ba_p_ens Ap) (proj1_sig (plus_set_p D pfd))) (pfb:Inn (ba_p_ens Bp) (proj1_sig (plus_set_p D pfd))), proj1_sig (exist _ _ pfa) = proj1_sig (exist _ _ pfb). intros Ap h1 D h2 h3 h4. subst. reflexivity. Qed. Lemma times_set_times_in_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp), Inn E x -> x %* (times_set_p E pf) = times_set_p E pf. intros E h1 x h2. rewrite times_set_p_eq. rewrite ba_conv_times. unfold ba_conv_elt, ba_conv_set, ba_conv_type. do 3 rewrite transfer_eq_refl. rewrite times_set_times_in; auto. Qed. Lemma plus_set_plus_in_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp), Inn E x -> x %+ (plus_set_p E pf) = plus_set_p E pf. intros E h1 x h2. rewrite plus_set_p_eq. rewrite ba_conv_plus. unfold ba_conv_elt, ba_conv_set, ba_conv_type. do 3 rewrite transfer_eq_refl. rewrite plus_set_plus_in; auto. Qed. Lemma times_list_times_in_p : forall (l:list Btp) (x:Btp), In x l -> x %* (times_list_p l) = times_list_p l. intros l x h1. rewrite times_list_compat_p. apply times_set_times_in_p. rewrite <- list_to_set_in_iff; auto. Qed. Lemma plus_list_plus_in_p : forall (l:list Btp) (x:Btp), In x l -> x %+ (plus_list_p l) = plus_list_p l. intros l x h1. rewrite plus_list_compat_p. apply plus_set_plus_in_p. rewrite <- list_to_set_in_iff; auto. Qed. Lemma times_list_0_p : forall (l:list Btp), In %0 l -> times_list_p l = %0. intros l h1; rewrite times_list_p_eq, ba_conv_zero. unfold ba_conv_list, ba_conv_type. rewrite transfer_dep_eq_refl. apply (@times_list_0 (ba_conv Bp) l h1). Qed. Lemma times_list_opps_p : forall (l:list Btp) (x y:btp Bp), In x l -> In y l -> x = %- y -> times_list_p l = %0. intros l x y h1 h2 h3. rewrite times_list_p_eq. rewrite ba_conv_zero. apply (@times_list_opps (ba_conv Bp) (ba_conv_list l) _ _ h1 h2 h3); auto. Qed. Lemma plus_list_1_p : forall (l:list Btp), In %1 l -> plus_list_p l = %1. intros l h1; rewrite plus_list_p_eq, ba_conv_one. unfold ba_conv_list, ba_conv_type. rewrite transfer_dep_eq_refl. apply (@plus_list_1 (ba_conv Bp) l h1). Qed. Lemma plus_list_opps_p : forall (l:list Btp) (x y:btp Bp), In x l -> In y l -> x = %- y -> plus_list_p l = %1. intros l x y h1 h2 h3. rewrite plus_list_p_eq. rewrite ba_conv_one. apply (@plus_list_opps (ba_conv Bp) (ba_conv_list l) _ _ h1 h2 h3); auto. Qed. Lemma times_set_in_subtract_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp), let pfs := subtract_preserves_finite E x pf in Inn E x -> times_set_p E pf = (times_set_p (Subtract E x) pfs) %* x. intros E h1 x h2 h3. do 2 rewrite times_set_p_eq. rewrite ba_conv_times. unfold ba_conv_elt, ba_conv_set, ba_conv_type. do 3 rewrite transfer_eq_refl. erewrite times_set_in_subtract. match goal with |- ?x * _ = ?x' * _ => assert (h4:x = x') end. apply times_set_functional. do 2 rewrite transfer_dep_eq_refl. reflexivity. rewrite h4. f_equal. apply h3. Qed. Lemma times_list_map_in_dep_remove_p : forall {U:Type} {l:list U} (pfdt:type_eq_dec U) (f:fun_in_dep l Btp) (x:U) (pf:In x l), inj_dep f -> let f' := fun_in_dep_remove pfdt f x in times_list_p (map_in_dep f) = (f x pf) %* times_list_p (map_in_dep f'). intros U l h1 f x h2 h3. simpl. pose proof (@times_list_map_in_dep_remove (ba_conv Bp) U l h1 f x h2 h3) as h4. simpl in h4. do 2 rewrite times_list_p_eq. assumption. Qed. Lemma times_list_map_in_dep_remove_in_p : forall {U:Type} {l:list U} (pfdt:type_eq_dec U) (f:fun_in_dep l Btp) (x:U) (pf:In x l), let f' := fun_in_dep_remove_in pfdt ba_eq_dec_p f x pf in times_list_p (map_in_dep f) = (f x pf) %* times_list_p (map_in_dep f'). intros U l h1 f x h2. simpl. pose proof (@times_list_map_in_dep_remove_in (ba_conv Bp) U l h1 f x h2) as h4. simpl in h4. do 2 rewrite times_list_p_eq. unfold Btp in h4. unfold Btp. rewrite h4 at 1. f_equal. f_equal. assert (h5:@ba_eq_dec (ba_conv Bp) = ba_eq_dec_p). apply type_eq_dec_eq. rewrite h5. reflexivity. Qed. Lemma plus_list_map_in_dep_remove_p : forall {U:Type} {l:list U} (pfdt:type_eq_dec U) (f:fun_in_dep l Btp) (x:U) (pf:In x l), inj_dep f -> let f' := fun_in_dep_remove pfdt f x in plus_list_p (map_in_dep f) = (f x pf) %+ plus_list_p (map_in_dep f'). intros U l h1 f x h2 h3. simpl. pose proof (@plus_list_map_in_dep_remove (ba_conv Bp) U l h1 f x h2 h3) as h4. simpl in h4. do 2 rewrite plus_list_p_eq. assumption. Qed. Lemma plus_list_map_in_dep_remove_in_p : forall {U:Type} {l:list U} (pfdt:type_eq_dec U) (f:fun_in_dep l Btp) (x:U) (pf:In x l), let f' := fun_in_dep_remove_in pfdt ba_eq_dec_p f x pf in plus_list_p (map_in_dep f) = (f x pf) %+ plus_list_p (map_in_dep f'). intros U l h1 f x h2. simpl. pose proof (@plus_list_map_in_dep_remove_in (ba_conv Bp) U l h1 f x h2) as h4. simpl in h4. do 2 rewrite plus_list_p_eq. unfold Btp in h4. unfold Btp. rewrite h4 at 1. f_equal. f_equal. assert (h5:@ba_eq_dec (ba_conv Bp) = ba_eq_dec_p). apply type_eq_dec_eq. rewrite h5. reflexivity. Qed. Lemma plus_set_in_subtract_p : forall (E:Ensemble Btp) (pf:Finite E) (x:Btp), let pfs := subtract_preserves_finite E x pf in Inn E x -> plus_set_p E pf = (plus_set_p (Subtract E x) pfs) %+ x. intros E h1 x h2 h3. do 2 rewrite plus_set_p_eq. rewrite ba_conv_plus. unfold ba_conv_elt, ba_conv_set, ba_conv_type. do 3 rewrite transfer_eq_refl. erewrite plus_set_in_subtract. match goal with |- ?x + _ = ?x' + _ => assert (h4:x = x') end. apply plus_set_functional. do 2 rewrite transfer_dep_eq_refl. reflexivity. rewrite h4. f_equal. apply h3. Qed. End ParametricAnalogues. Arguments plus_list_app_p [T'] [Bp] _ _. Arguments times_list_app_p [T'] [Bp] _ _. Arguments dist_list_sing_plus_p [T'] [Bp] _ _. Arguments dist_list_sing_times_p [T'] [Bp] _ _. Arguments dist_list_2_plus_p [T'] [Bp] _ _. Arguments dist_list_2_times_p [T'] [Bp] _ _. Arguments plus_times_list_of_lists_p [T'] [Bp] _. Arguments times_plus_list_of_lists_p [T'] [Bp] _. Arguments times_fun_p [T'] [Bp] [T] _ _. Arguments plus_fun_p [T'] [Bp] [T] _ _. Arguments plus_times_fun1_p [T'] [Bp] [T] [U] _ _ _. Arguments plus_times_fun2_p [T'] [Bp] [T] [U] _ _ _. Arguments times_plus_fun1_p [T'] [Bp] [T] [U] _ _ _. Arguments times_plus_fun2_p [T'] [Bp] [T] [U] _ _ _. Arguments plus_times_all_funs1_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments plus_times_all_funs2_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments times_plus_all_funs1_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments times_plus_all_funs2_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments complete_dist_list_times_p [T'] [Bp] _. Arguments complete_dist_list_plus_p [T'] [Bp] _. Arguments inf_union_p [T'] [Bp] _ _ _ _ _ _. Arguments sup_union_p [T'] [Bp] _ _ _ _ _ _. Arguments decompose_inf_p [T'] [Bp] _ _ _ _ _ _ _ _ _. Arguments decompose_inf_p' [T'] [Bp] _ _ _ _ _ _ _. Arguments inf_times_cons_p [T'] [Bp] _ _ _ _ _. Arguments inf_times_finite_p [T'] [Bp] _. Arguments sup_plus_finite_p [T'] [Bp] _. Arguments le_times_finite_member_p [T'] [Bp] _ _ _. Arguments le_member_plus_finite_p [T'] [Bp] _ _ _. Arguments prod_list_dup_eq_p [T'] [Bp] _ _ _. Arguments sum_list_dup_eq_p [T'] [Bp] _ _ _. Arguments prod_preserves_list_singularize_p [T'] [Bp] _. Arguments sum_preserves_list_singularize_p [T'] [Bp] _. Arguments times_sing_preserves_new_head_p [T'] [Bp] _ _ _. Arguments plus_sing_preserves_new_head_p [T'] [Bp] _ _ _. Arguments times_list_sing_cons_p [T'] [Bp] _ _ _. Arguments plus_list_sing_cons_p [T'] [Bp] _ _ _. Arguments list_to_sets_eq_times_sing_eq_p [T'] [Bp] _ _ _ _ _. Arguments list_to_sets_eq_plus_sing_eq_p [T'] [Bp] _ _ _ _ _. Arguments list_to_sets_eq_times_eq_p [T'] [Bp] _ _ _ _ _. Arguments list_to_sets_eq_plus_eq_p [T'] [Bp] _ _ _ _ _. Arguments times_list_unq_p [T'] [Bp] _ _. Arguments plus_list_unq_p [T'] [Bp] _ _. Arguments times_set_p [T'] [Bp] _ _. Arguments plus_set_p [T'] [Bp] _ _. Arguments times_set_compat_p [T'] [Bp] _ _. Arguments times_set_compat_p' [T'] [Bp] _ _ _ _. Arguments plus_set_compat_p [T'] [Bp] _ _. Arguments plus_set_compat_p' [T'] [Bp] _ _ _ _. Arguments times_set_functional_p [T'] [Bp] _ _ _ _ _. Arguments plus_set_functional_p [T'] [Bp] _ _ _ _ _. Arguments times_set_empty_p [T'] [Bp]. Arguments times_set_empty_p' [T'] [Bp] _. Arguments plus_set_empty_p [T'] [Bp]. Arguments plus_set_empty_p' [T'] [Bp] _. Arguments times_set_add_p [T'] [Bp] _ _ _. Arguments times_set_add_p' [T'] [Bp] _ _ _ _. Arguments plus_set_add_p [T'] [Bp] _ _ _. Arguments plus_set_add_p' [T'] [Bp] _ _ _ _. Arguments times_set_sing_p [T'] [Bp] _. Arguments times_set_sing_p' [T'] [Bp] _ _. Arguments times_set_one_or_p [T'] [Bp] _ _ _. Arguments plus_set_sing_p [T'] [Bp] _. Arguments plus_set_sing_p' [T'] [Bp] _ _. Arguments plus_set_zero_or_p [T'] [Bp] _ _ _. Arguments times_set_couple_p [T'] [Bp] _ _. Arguments times_set_couple_p' [T'] [Bp] _ _ _. Arguments plus_set_couple_p [T'] [Bp] _ _. Arguments plus_set_couple_p' [T'] [Bp] _ _ _. Arguments le_times_set_p [T'] [Bp] _ _ _ _. Arguments le_plus_set_p [T'] [Bp] _ _ _ _. Arguments inf_times_set_p [T'] [Bp] _ _. Arguments sup_plus_set_p [T'] [Bp] _ _. Arguments times_set_union_p [T'] [Bp] _ _ _ _. Arguments times_set_union_p' [T'] [Bp] _ _ _ _ _. Arguments times_set_inc_le_p [T'] [Bp] _ _ _ _ _. Arguments plus_set_union_p [T'] [Bp] _ _ _ _. Arguments plus_set_union_p' [T'] [Bp] _ _ _ _ _. Arguments plus_set_inc_le_p [T'] [Bp] _ _ _ _ _. Arguments plus_set_im_add_p [T'] [Bp] _ _ _ _. Arguments plus_set_im_add_p' [T'] [Bp] _ _ _ _ _ _. Arguments times_set_im_add_p [T'] [Bp] _ _ _ _. Arguments times_set_im_add_p' [T'] [Bp] _ _ _ _ _ _. Arguments dist_set_plus1_p [T'] [Bp] _ _ _. Arguments dist_set_plus1_p' [T'] [Bp] _ _ _ _. Arguments dist_set_times1_p [T'] [Bp] _ _ _. Arguments dist_set_times1_p' [T'] [Bp] _ _ _ _. Arguments dist_set_plus2_p [T'] [Bp] _ _ _ _. Arguments dist_set_plus2_p' [T'] [Bp] _ _ _ _ _. Arguments dist_set_times2_p [T'] [Bp] _ _ _ _. Arguments dist_set_times2_p' [T'] [Bp] _ _ _ _ _. Arguments plus_fin_pair_map1_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _. Arguments plus_fin_pair_map2_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _. Arguments times_fin_pair_map1_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _. Arguments times_fin_pair_map2_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _. Arguments plus_fin_pair_map1_list_compat_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _ _. Arguments plus_fin_pair_map2_list_compat_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _ _. Arguments times_fin_pair_map1_list_compat_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _ _. Arguments times_fin_pair_map2_list_compat_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _ _. Arguments plus_fin_pair_map2_functional_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _ _ . Arguments fin_map_times_p [T'] [Bp] [T] [A] [C] _. Arguments fin_map_times_list_compat_p [T'] [Bp] [T] [A] [C] _ _. Arguments fin_map_times_empty1_p [T'] [Bp] [T] [C] _. Arguments fin_map_eq_times_p [T'] [Bp] [T] _ _ _ _ _ _. Arguments im2_empty1_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments im2_empty2_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments plus_fin_pair_map2_cart_empty_eq1_p [T'] [Bp] [T] [U] _ _ _ _ _ _. Arguments plus_fin_pair_map2_cart_empty_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments plus_fin_pair_map2_cart_empty_eq2_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments fin_map_times_sing_p [T'] [Bp] [T] _ _ _ _ _. Arguments fin_map_times_cart_empty11_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments fin_map_times_cart_empty21_p [T'] [Bp] [T] [U] _ _ _ _ _ _. Arguments fin_map_times_empty2_p [T'] [Bp] [T] _ _. Arguments fin_map_plus_p [T'] [Bp] [T] [A] [C] _. Arguments fin_map_plus_list_compat_p [T'] [Bp] [T] [A] [C] _ _. Arguments fin_map_plus_empty1_p [T'] [Bp] [T] [C] _. Arguments fin_map_plus_empty2_p [T'] [Bp] [T] [A] _. Arguments fin_map_eq_plus_p [T'] [Bp] [T] _ _ _ _ _ _. Arguments times_plus_fin_pair_map1_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _. Arguments times_plus_fin_pair_map2_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _. Arguments times_plus_fin_pair_map2_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _. Arguments plus_times_fin_pair_map1_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _. Arguments plus_times_fin_pair_map2_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _. Arguments times_plus_fin_pair_map1_list_compat_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _. Arguments plus_fun_fin_map_to_fun_comm_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _ _ _ _. Arguments plus_fun_fin_map_to_fun_comm_p' [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _ _ _ _. Arguments times_plus_fin_pair_map1_list_compat_p' [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _ _ _. Arguments times_plus_fin_pair_map2_list_compat_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _. Arguments times_plus_fin_pair_map2_list_compat_p' [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _ _ _. Arguments plus_times_fin_pair_map1_list_compat_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _. Arguments times_fun_fin_map_to_fun_comm_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _ _ _ _. Arguments plus_times_fin_pair_map1_list_compat_p' [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _ _ _. Arguments plus_times_fin_pair_map2_list_compat_p [T'] [Bp] [T] [U] [C] [D] [E] _ _ _ _. Arguments times_fun_fin_map1_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _ . Arguments times_fun_fin_map2_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _. Arguments times_fun_fin_map1_l_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _ _. Arguments times_fun_fin_map2_l_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _ _. Arguments im_sing_times_fun_fin_map1_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _. Arguments times_fun_fin_map1_empty_p [T'] [Bp] [T] [U] _ _ _ _. Arguments times_fun_fin_map2_empty_p [T'] [Bp] [T] [U] _ _ _ _. Arguments times_fun_fin_map1_empty1_p [T'] [Bp] [T] [U] _ _ _ _ _ _. Arguments times_fun_fin_map2_empty1_p [T'] [Bp] [T] [U] _ _ _ _ _ _. Arguments times_fun_fin_map1_list_compat_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _ _. Arguments times_fun_fin_map2_list_compat_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _ _. Arguments plus_fun_fin_map1_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _. Arguments plus_fun_fin_map2_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _. Arguments plus_fun_fin_map1_empty_p [T'] [Bp] [T] [U] _ _ _ _. Arguments plus_fun_fin_map1_empty1_p [T'] [Bp] [T] [U] _ _ _ _ _ _. Arguments plus_fun_fin_map1_l_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _ _. Arguments plus_fun_fin_map2_l_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _ _. Arguments plus_fun_fin_map2_empty_p [T'] [Bp] [T] [U] _ _ _ _. Arguments plus_fun_fin_map2_empty1_p [T'] [Bp] [T] [U] _ _ _ _ _ _. Arguments plus_fun_fin_map1_list_compat_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _ _. Arguments plus_fun_fin_map2_list_compat_p [T'] [Bp] [T] [U] [I] [J] [def] _ _ _ _. Arguments plus_times_fun_all_maps1_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments plus_times_fun_all_maps2_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments times_plus_fun_all_maps1_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments times_plus_fun_all_maps2_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _. Arguments complete_dist_list_times1_p' [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _. Arguments complete_dist_list_times2_p' [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _. Arguments complete_dist_list_plus1_p' [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _. Arguments complete_dist_list_plus2_p' [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _. Arguments plus_times_all_maps1_funs_compat_p [T'] [Bp] [T] [U] [I] [J] _ _ _ _ _ _ _ _ _ _ _ _. Arguments plus_times_all_maps2_funs_compat_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _ _ _ _ _ _ _. Arguments times_plus_all_maps1_funs_compat_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _ _ _ _ _ _ _. Arguments times_plus_all_maps2_funs_compat_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _ _ _ _ _ _ _. Arguments times_plus_fun1_fin_map_to_fun_undoes_fun_to_fin_map_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _ _ _. Arguments times_plus_fun2_fin_map_to_fun_undoes_fun_to_fin_map_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _ _ _. Arguments plus_times_fun1_fin_map_to_fun_undoes_fun_to_fin_map_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _ _ _. Arguments plus_times_fun2_fin_map_to_fun_undoes_fun_to_fin_map_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _ _ _. Arguments complete_dist_times_plus1_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _. Arguments complete_dist_times_plus2_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _. Arguments complete_dist_plus_times1_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _. Arguments complete_dist_plus_times2_p [T'] [Bp] [T] [U] _ _ _ _ _ _ _ _. Arguments ba_conv_fin_map_dom [T'] [Bp] [U] [C] [D] [def] _. Arguments ba_conv_fin_map_ran [T'] [Bp] [T] [C] [D] _. Arguments incl_ba_conv_ens_fin_map_dom_iff [T'] [Bp] [U] [C] [D] [def] _ _. Arguments ba_conv_ens_fin_map_dom [T'] [Bp] [U] [C] [D] [def] _ _. Arguments ba_conv_ens_fin_map_ran [T'] [Bp] [T] [C] [D] _ _. bool2-v0-3/ListUtilities2.v0000644000175000017500000002563613575574055016436 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2016 Daniel Wyckoff *) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Import SetUtilities. Require Import SetUtilities1_5. Require Import SetUtilities2. Require Import ListUtilities. Require Import LogicUtilities. Require Import ArithUtilities. Require Import FiniteBags. Require Import DecidableDec. Require Import Description. Require Import FunctionalExtensionality. Require Import FunctionProperties. Require Import FunctionProperties2. Require Import TypeUtilities. Require Import EqDec. Require Import FiniteMaps. Require Import NPeano. (*Version 0.3 Alpha bool2 -- certified -- Version 8.4 pl4 Coq*) (*[apply proof_irrelevance]-certified.*) (*"Conve"rt 'f'irstn*) Definition conv_fn {T:Type} {l:list T} {m:(length l) @} (r:(S (proj1_sig m)) @) : (length (firstn (S (proj1_sig m)) l)) @ := segT_transfer (eq_sym (length_firstn_S _ _ (proj2_sig m))) r. Lemma subtract_remove_ens_compat : forall {T:Type} (pfd:type_eq_dec T) (l:list (Ensemble T)) (B:Ensemble T) (pfb:Finite B) (pfp:p_mem_l (@Finite _) l), Subtract (list_to_set l) B = list_to_set (remove_dep (@Finite _) (fin_ens_eq_dec T pfd) l B pfb pfp). intros T hdt l. induction l as [|A l ih]; simpl. intros; apply sub_empty. intros B h1 h2. destruct fin_ens_eq_dec as [h3 | h3]; subst. rewrite sub_add_eq. apply ih. simpl. rewrite sub_add_comm. specialize (ih B h1). hfirst ih. apply (p_mem_l_cons _ _ _ h2); auto. specialize (ih hf). rewrite ih. f_equal. f_equal. f_equal. apply proof_irrelevance. apply neq_sym; auto. Qed. Lemma plus_bag_nat_im_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), plus_bag_nat (im_set_bag nat_eq_dec (list_to_set (linds pfdt l)) (list_to_set_finite _) (@length _)) = length l. intros T h1 l. induction l as [|a l ih]; simpl. unfold linds. simpl. pose proof (im_set_bag_empty nat_eq_dec (@length T)) as h2. erewrite im_set_bag_functional2. rewrite im_set_bag_empty. apply plus_bag_nat_empty. reflexivity. reflexivity. pose proof (list_to_set_sort_list _ lt_list_nat_sort_dec (linds h1 (a::l))) as h2. simpl in h2. rewrite (im_set_bag_functional1 nat_eq_dec _ _ (list_to_set_finite _) (list_to_set_finite _) (@length _) h2). fold (sorted_linds h1 (a::l)). destruct (in_dec h1 a l) as [h3 | h3]. rewrite sorted_linds_cons_in. simpl. pose proof (im_set_bag_add nat_eq_dec (list_to_set (famlS (remove list_nat_eq_dec (linds_occ h1 l a) (sorted_linds h1 l)))) (list_to_set_finite _) (0 :: lS (linds_occ h1 l a))) as h4. hfirst h4. intro h5. rewrite <- list_to_set_in_iff in h5. rewrite in_famlS_iff in h5. destruct h5 as [h5 h6]. destruct h6 as [h6 h7]. pose proof (in_eq 0 (lS (linds_occ h1 l a))) as h8. rewrite h7 in h8. rewrite in_lS_iff in h8. destruct h8 as [? [? ?]]; discriminate. specialize (h4 hf (@length nat)). erewrite im_set_bag_functional2. rewrite h4 at 1. rewrite plus_bag_nat_add. simpl. rewrite length_lS. pose proof (linds_decompose_remove h1 _ _ h3) as h5. pose proof (list_to_set_sort_list _ lt_list_nat_sort_dec (linds h1 l)) as h6. simpl in h6. unfold sorted_linds in h5. rewrite h5 in h6. rewrite lt_list_nat_tso_dec_eq in h6. rewrite <- list_to_set_sort_list in h6. simpl in h6. rewrite (im_set_bag_functional1 _ _ _ _ (Add_preserves_Finite _ _ _ (list_to_set_finite _)) (@length _) h6) in ih. rewrite im_set_bag_add, plus_bag_nat_add in ih. unfold linds_remove in ih. rewrite plus_S_shift_r. f_equal. rewrite <- ih. f_equal. symmetry. apply plus_bag_nat_im_set_bag_eq_f'. exists lS. split. apply inj_lS. split. unfold famlS. rewrite map_im_compat. f_equal. unfold sorted_linds. rewrite <- subtract_remove_compat. rewrite <- subtract_remove_compat. f_equal. rewrite <- list_to_set_sort_list; auto. intros l' h7. rewrite length_lS; auto. rewrite <- list_to_set_in_iff. apply nin_linds_remove. assumption. reflexivity. reflexivity. assumption. rewrite sorted_linds_cons_nin; auto. simpl. rewrite im_set_bag_add'. rewrite <- ih. simpl. rewrite plus_bag_nat_add. rewrite S_compat. f_equal. symmetry. apply plus_bag_nat_im_set_bag_eq_f'. exists lS. split. apply inj_lS. split. unfold famlS. rewrite map_im_compat. f_equal. unfold sorted_linds. rewrite <- list_to_set_sort_list; auto. intros x h4. rewrite length_lS; auto. rewrite <- list_to_set_in_iff. intro h4. rewrite in_famlS_iff in h4. destruct h4 as [ln [h4 h5]]. pose proof (in_eq 0 nil) as h6. rewrite h5 in h6. rewrite in_lS_iff in h6. destruct h6 as [? [? ?]]; discriminate. Qed. Lemma inv_dep_nil : forall (T:Type), inv_dep (fun_in_dep_nil T False). intros; apply bij_impl_inv_dep; apply bij_dep_nil. Qed. Lemma inv_dep_nil_nil : forall T U (f:fun_in_dep (@nil T) U), inv_dep f -> equi U False. intros T U f h1. pose proof (inv_dep_sig f h1) as h2. pose proof (invertible_impl_inv_invertible _ h2) as h3. eapply equi_trans. econstructor. apply h3. apply equi_sym. apply equi_sig_false. Qed. Lemma inv_from_cons_in : forall {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U), In a l -> inv_dep f -> inv_dep (fun_from_cons f). intros ? ? ? ? ? ? h1; apply inv_impl_bij_dep in h1; apply bij_impl_inv_dep; apply bij_from_cons_in; auto. Qed. Lemma inv_from_cons_ran : forall {T U:Type} {l:list T} {a:T} (f:fun_in_dep (a::l) U) (pfnin : ~In a l) (pfinv:inv_dep f), let pfinj := inv_impl_inj_dep _ pfinv in inv_dep (fun_from_cons_ran f pfinj pfnin). intros; apply bij_impl_inv_dep. apply inj_surj_bij_from_cons_ran. apply inv_impl_surj_dep; auto. Qed. Lemma inv_tl_ran : forall {T U} {l:list T} (f:fun_in_dep l U) (pfn:l <> nil) (pfinj:inj_dep f) (pfnin : ~In (hd' l pfn) (tl l)), inv_dep f -> inv_dep (tl_in_dep_ran f pfn pfinj pfnin). intros; apply bij_impl_inv_dep. apply bij_tl_ran. apply inv_impl_bij_dep; auto. Qed. Lemma card_fun1_family_union_linds : forall {T:Type} (pfdt:type_eq_dec T) (l:list T), card_fun1 (FamilyUnion (faml_to_fam (linds pfdt l))) = length l. intros T hdt l. pose proof (partition_linds hdt l) as h1. do 2 red in h1. destruct h1 as [h1 h2]. pose proof (card_pairwise_disjoint (faml_to_fam (linds hdt l)) (finite_faml_to_fam _)) as h3. hfirst h3. intros S h4. apply in_faml_to_fam_impl in h4. destruct h4 as [l' [h4 h5]]; subst. apply list_to_set_finite. specialize (h3 hf h1). rewrite h3. pose proof (plus_bag_nat_im_linds hdt l) as h4. rewrite <- h4. symmetry. apply plus_bag_nat_im_set_bag_eq_f_hetero_dep. exists (fun l pf => list_to_set l). split. red. intros lm ln h5 h6 h7. rewrite <- list_to_set_in_iff in h5, h6. eapply in_linds_list_to_set_injable. apply h5. assumption. assumption. split. apply Extensionality_Ensembles; red; split; red; intros S h5. apply in_faml_to_fam_impl in h5. destruct h5 as [ln [h5 h6]]; subst. econstructor. rewrite <- list_to_set_in_iff. apply h5. reflexivity. destruct h5 as [S h5]; subst. rewrite <- list_to_set_in_iff in h5. apply in_faml_impl; auto. intros ln h5. erewrite card_fun1_length_eq. rewrite <- list_to_set_in_iff in h5. apply no_dup_faml_linds in h5. rewrite (list_singularize_no_dup nat_eq_dec) in h5. rewrite h5 at 1; auto. Qed. Lemma sorted_linds_eq_length : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), sorted_linds pfdt l = sorted_linds pfdu l' -> length l = length l'. intros T U h1 h2 l l' h3. pose proof (partition_linds h1 l) as h4. do 2 red in h4. destruct h4 as [h4 h5]. pose proof h5 as h5'. rewrite <- faml_to_fam_sorted_linds in h5'. rewrite h3 in h5'. pose proof (f_equal card_fun1 h5') as h6. rewrite list_to_set_seg_list_eq, faml_to_fam_sorted_linds, card_fun1_seg_eq, card_fun1_family_union_linds in h6; auto. Qed. Lemma sorted_linds_eq_iff' : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), (sorted_linds pfdt l = sorted_linds pfdu l') <-> (exists pfe:length l = length l', forall (m:nat) (pfm:m < length l), let pfm' := lt_eq_trans _ _ _ pfm pfe in lec pfdt pfm = lec pfdu pfm'). intros T U h1 h2 l l'; split; intro h3. unfold sorted_linds in h3. pose proof (sorted_linds_eq_length _ _ _ _ h3) as he. exists he. intros m h4; simpl. pose proof (partition_linds h1 l) as h5. pose proof (partition_linds h2 l') as h6. pose proof (in_linds_list_to_set_injable h1 l) as h9. red in h9. apply h9. apply in_lec. rewrite <- (in_sort_list_iff _ lt_list_nat_sort_dec (linds h1 l)). rewrite h3, in_sort_list_iff. apply in_lec. eapply partition_faml_eq. apply h5. apply in_lec. rewrite <- (in_sort_list_iff _ lt_list_nat_sort_dec (linds h1 l)). rewrite h3, in_sort_list_iff. apply in_lec. apply rep_in_lec. rewrite in_lec_iff. reflexivity. destruct h3 as [h3 h4]. pose proof (list_to_set_injable_increasingR _ lt_list_nat_sort_dec) as h5. red in h5. apply h5. unfold sorted_linds. apply weak_incrR_no_dup_impl. apply weak_incrR_sort_list. apply no_dup_sort_list. apply no_dup_linds. apply weak_incrR_no_dup_impl. apply weak_incrR_sort_list. apply no_dup_sort_list. apply no_dup_linds. apply Extensionality_Ensembles; red; split; red; intros ln h6; rewrite <- list_to_set_in_iff in h6; rewrite <- list_to_set_in_iff; unfold sorted_linds in h6; unfold sorted_linds; rewrite in_sort_list_iff in h6; rewrite in_sort_list_iff; rewrite in_linds_lec_iff in h6; rewrite in_linds_lec_iff; destruct h6 as [i h7]; destruct h7 as [h7]; subst; exists i. exists (lt_eq_trans _ _ _ h7 h3). apply h4. exists (lt_eq_trans _ _ _ h7 (eq_sym h3)). rewrite h4. f_equal. apply proof_irrelevance. Qed. Lemma linds_eq_iff : forall {T U:Type} (pfdt:type_eq_dec T) (pfdu:type_eq_dec U) (l:list T) (l':list U), (linds pfdt l = linds pfdu l') <-> (exists pfe:length l = length l', forall (m:nat) (pfm:m < length l), let pfm' := lt_eq_trans _ _ _ pfm pfe in lec pfdt pfm = lec pfdu pfm'). intros; rewrite <- sorted_linds_eq_iff; apply sorted_linds_eq_iff'. Qed. bool2-v0-3/DecidableDec.v0000644000175000017500000001154513575574055016007 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014, Daniel Wyckoff, except for the portions so labeleled which I got from Daniel Schepler*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*This file is copied and pasted from Daniel Schepler's "Zorn's Lemma" except for lemmas labeled "Wyckoff" *) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Import Description. Require Import EqDec. Require Import ProofIrrelevance. Require Import Arith. (*Wyckoff*) Definition pred_dec {T:Type} (P:T->Prop) : Type := forall x:T, {P x} + {~P x}. (*Wyckoff*) Definition pred_dep_nat (n:nat) : Type := forall m:nat, m < n->Prop. (*Wyckoff*) Definition pred_dep_nat_pred {n:nat} (P:pred_dep_nat (S n)) : pred_dep_nat n := fun m pf => P m (lt_trans _ _ _ pf (lt_n_Sn _)). (*Wyckoff*) Definition pred_dep_nat_dec {n:nat} (P:pred_dep_nat n) : Type := forall (m:nat) (pfm:m < n), {P m pfm} + {~P m pfm}. (*Wyckoff*) Lemma pred_dep_nat_dec_pred : forall {n:nat} (P:pred_dep_nat (S n)), pred_dep_nat_dec P -> pred_dep_nat_dec (pred_dep_nat_pred P). intros n P h1; red; intros m h2. pose proof (h1 _ (lt_trans _ _ _ h2 (lt_n_Sn _))) as h3. destruct h3 as [h3 | h3]. left. red; intros; auto. right; intros; auto. Qed. (*Schepler*) Lemma exclusive_dec: forall P Q:Prop, ~(P /\ Q) -> (P \/ Q) -> {P} + {Q}. Proof. intros. assert ({x:bool | if x then P else Q}). apply constructive_definite_description. case H0. exists true. red; split. assumption. destruct x'. reflexivity. tauto. exists false. red; split. assumption. destruct x'. tauto. reflexivity. destruct H1. destruct x. left. assumption. right. assumption. Qed. (*Schepler*) Lemma decidable_dec: forall P:Prop, P \/ ~P -> {P} + {~P}. Proof. intros. apply exclusive_dec. tauto. assumption. Qed. (*Wyckoff*) Lemma classic_eq : forall {P:Prop} (pf pf': {P} + {~P}), pf = pf'. intros P h1 h2. destruct h1, h2; f_equal; try apply proof_irrelevance; try contradiction. Qed. (*Wyckoff*) Lemma eq_dec_same : forall {T:Type} (pfdt:type_eq_dec T) (a:T) (pf: {a = a} + {a <> a}), pf = left (a <> a) eq_refl. intros T h1 z h2. destruct h2; subst; auto; f_equal; try apply proof_irrelevance. contradict n; auto. Qed. (*Wyckoff*) Lemma eq_dec_neq_l : forall {T:Type} (pfdt:type_eq_dec T) (a a':T) (pfneq:a <> a') (pf: {a = a'} + {a <> a'}), pf = right pfneq. intros T h1 a a' h2 h3. destruct h3 as [|h3]; subst. contradict h2; auto. f_equal. apply proof_irrelevance. Qed. (*See [LogicUtilities] for [eq_dec_neq_r]*) Require Import Classical. (*Schepler*) Lemma classic_dec: forall P:Prop, {P} + {~P}. Proof. intro P. apply decidable_dec. apply classic. Qed. (*Wyckoff*) Lemma or_to_sumbool : forall P Q:Prop, P \/ Q -> {P} + {Q}. intros P Q h1. case (classic_dec (P /\ Q)) as [h0l | h0r]. destruct h0l; left; assumption. apply exclusive_dec; assumption. Qed. (*Wyckoff*) Lemma sumbool_to_or : forall P Q:Prop, {P} + {Q} -> P \/ Q. intros ? ? h1. destruct h1; [left | right]; auto. Qed. (*Wyckoff*) Lemma fun_abstract_destruct_l : forall {T V:Type} {U:Type->Type} (l:U T) (a:T) (PL PR : U T-> T -> Prop) (pf:{PL l a}+ {PR l a}) (fl :PL l a -> V) (fr : PR l a -> V) (pfl:PL l a), ~ (PL l a /\ PR l a) -> match pf with | left pfl => fl pfl | right pfr => fr pfr end = fl pfl. intros T V U l a PL PR h2 fl fr h3 hn. destruct h2 as [h2 | h2]. pose proof (proof_irrelevance _ h2 h3); subst. reflexivity. tauto. Qed. (*Wyckoff*) Lemma fun_abstract_destruct_r : forall {T V:Type} {U:Type->Type} (l:U T) (a:T) (PL PR : U T-> T -> Prop) (pf:{PL l a}+ {PR l a}) (fl : PL l a -> V) (fr : PR l a -> V) (pfr:PR l a), ~ (PL l a /\ PR l a) -> match pf with | left pfl => fl pfl | right pfr => fr pfr end = fr pfr. intros T V U l a PL PR h2 fl fr h3 hn. destruct h2 as [h2 | h2]. tauto. pose proof (proof_irrelevance _ h2 h3); subst; auto. Qed. (*Wyckoff*) Lemma bool_eq_dec : type_eq_dec bool. red; intros x y; destruct x, y. left; auto. right; intro; discriminate. right; intro; discriminate. left; auto. Qed. (*Wyckoff*) Lemma bool_dec : forall b, {b = true} + {b = false}. intro b; destruct b; auto. Qed. bool2-v0-3/EnsemblesImplicit.v0000644000175000017500000000253313575574055017144 0ustar dwyckoff76dwyckoff76(* This simple "meta file" is taken directly from Schepler's "Zorn's Lemma", with one trivial addition*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) Require Export Ensembles. Arguments In [U] _ _. Arguments Included [U] _ _. Arguments Singleton [U] _ _. Arguments Union [U] _ _ _. Arguments Add [U] _ _ _. Arguments Intersection [U] _ _ _. Arguments Couple [U] _ _ _. Arguments Triple [U] _ _ _ _. Arguments Complement [U] _ _. Arguments Setminus [U] _ _ _. Arguments Subtract [U] _ _ _. Arguments Disjoint [U] _ _. Arguments Inhabited [U] _. Arguments Strict_Included [U] _ _. Arguments Same_set [U] _ _. Arguments Extensionality_Ensembles [U] _ _ _. Require Export Finite_sets. Arguments Finite [U] _. bool2-v0-3/Makefile0000644000175000017500000001520013575600573014773 0ustar dwyckoff76dwyckoff76############################################################################# ## v # The Coq Proof Assistant ## ## "$@" || ( RV=$$?; rm -f "$@"; exit $${RV} ) %.v.beautified: $(COQC) $(COQDEBUG) $(COQFLAGS) -beautify $* # WARNING # # This Makefile has been automagically generated # Edit at your own risks ! # # END OF WARNING bool2-v0-3/RegularOpenSets.v0000644000175000017500000004302113575574055016613 0ustar dwyckoff76dwyckoff76(* Copyright (C) 2014, Daniel Wyckoff*) (*This file is part of BooleanAlgebrasIntro2. BooleanAlgebrasIntro2 is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. BooleanAlgebrasIntro2 is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with BooleanAlgebrasIntro2. If not, see .*) (*Version 0.3 bool2 -- certified -- Coq 8.4 pl4.*) Require Import ProofIrrelevance. Require Export BoolAlgBasics. Require Import Topology. Require Import SetUtilities. Require Import LogicUtilities. Require Import InfiniteOperations. Lemma empty_closed_reg : forall (X:TopologicalSpace), regular _ (Empty_set (point_set X)). intro X. red. rewrite closure_empty. rewrite interior_empty. reflexivity. Qed. Lemma full_closed_reg : forall (X:TopologicalSpace), regular _ (Full_set (point_set X)). intro X. red. rewrite closure_full. rewrite interior_full. reflexivity. Qed. Section RegularOpenAlgebra. Variable T:Type. Variable X:TopologicalSpace. (*"Lemma 1. If P <= Q, then Q - ' <= P - ' ."*) Lemma comp_closure_inclusion_quasi_sym : forall (P Q:(Ensemble (point_set X))), Included P Q -> Included (Q - ') (P - '). intros P Q h1. apply complement_inclusion. apply closure_increasing. assumption. Qed. (*"Lemma 2. If P is open, then P <= P - ' - ' ."*) Lemma open_included_reg : forall (P:(Ensemble (point_set X))), open P -> Included P (P - ' - '). intros P h0. pose proof (closure_inflationary P) as h1. pose proof (complement_inclusion _ _ h1) as h2. pose proof (Complement_Complement _ P) as h3. rewrite <- h3 in h0. change (open (P ' ')) with (closed (P ')) in h0. pose proof (closure_minimal _ _ h0 h2) as h4. pose proof (complement_inclusion _ _ h4) as h5. rewrite Complement_Complement in h5. assumption. Qed. (*" Lemma 3. If P is open, then P - ' = P - ' - ' - '"*) Lemma open_regularization_comp_clo : forall (P:Ensemble (point_set X)), open P -> (P - ') = P - ' - ' - '. intros P h1. apply Extensionality_Ensembles. red. split. (* >= *) Focus 2. apply comp_closure_inclusion_quasi_sym. apply open_included_reg. assumption. (* <= *) assert (h2:open (P - ')). pose proof (closure_closed P) as h2. red in h2. assumption. apply open_included_reg. assumption. Qed. Corollary open_regular_comp_clo : forall (P:Ensemble (point_set X)), open P -> regular _ (P - '). intros P h1. red. rewrite interior_eq. apply open_regularization_comp_clo. assumption. Qed. Lemma regularization_increasing : forall (P Q:Ensemble (point_set X)), Included P Q -> Included (P - ' - ') (Q - ' - '). intros P Q h1. pose proof (comp_closure_inclusion_quasi_sym _ _ h1) as h2. apply comp_closure_inclusion_quasi_sym; assumption. Qed. (*"Lemma 4. If P and Q are open, then (P /\ Q)-'-' = P-'-' /\ Q-'-'"*) Lemma reg_preserves_int : forall (P Q:Ensemble (point_set X)), open P -> open Q -> (Intersection P Q) - ' - ' = Intersection (P - ' - ') (Q - ' - '). intros P Q h1 h2. apply Extensionality_Ensembles. red. split. (* <= *) assert (h3:Included (Intersection P Q) P). auto with sets. assert (h4:Included (Intersection P Q) Q). auto with sets. pose proof (regularization_increasing _ _ h3). pose proof (regularization_increasing _ _ h4). constructor; auto with sets. (* >= *) assert (h0:forall (P0 Q0:Ensemble (point_set X)), open P0 -> Included (Intersection P0 (Q0 - ' - ')) (Intersection P0 Q0 - ' - ')). intros P0 Q0 hp0. (*hq0.*) assert (h3:Included (Intersection P0 (Q0 -)) (Intersection P0 Q0 -)). red. intros x h3. apply meets_every_open_neighborhood_impl_closure. intros U h4 h5. destruct h3 as [x h6 h7]. pose proof (closure_impl_meets_every_open_neighborhood _ _ _ h7) as h8. pose proof (open_intersection2 _ _ _ h4 hp0) as h9. assert (h10:In (Intersection U P0) x). auto with sets. pose proof (h8 _ h9 h10) as h11. clear h8. inversion h11 as [a h12]. apply Inhabited_intro with a. destruct h12 as [? ? h12]. destruct h12. repeat split; assumption. pose proof (complement_inclusion _ _ h3) as h4. rewrite comp_int in h4. pose proof (closure_increasing _ _ h4) as h5. pose proof (complement_inclusion _ _ h5) as h6. pose proof (Complement_Complement _ P0) as h7. rewrite <- h7 in hp0. pose proof (closure_union (P0 ') (Q0 - ')) as h8. rewrite h8 in h6. rewrite comp_union in h6. assert (h9: P0 ' - ' = P0). change (open (P ' ')) with (closed (P ')) in h1. pose proof (closure_fixes_closed _ hp0) as h9. pose proof (f_equal (closure (X:=X)) h9). congruence. rewrite h9 in h6. assumption. pose proof (h0 P Q h1) as h3. assert (h4:open (P - ' - ')). change (open (P - ' - ')) with (closed (P - ' -)). apply closure_closed. pose proof (h0 (P - ' - ') Q h4) as h5. pose proof (h0 Q P h2) as h6. rewrite Intersection_commutative in h6. pose proof (regularization_increasing _ _ h6) as h7. assert (h8:Intersection Q P = Intersection P Q). apply Intersection_commutative. rewrite h8 in h7. rewrite open_regularization_comp_clo with (Intersection P Q - '). auto with sets. apply closure_closed. Qed. Corollary intersection_regular : forall (P Q:Ensemble (point_set X)), regular _ P -> regular _ Q -> regular _ (Intersection P Q). intros P Q h1 h2. rewrite regular_iff. rewrite reg_preserves_int; try (apply regular_open; assumption). rewrite regular_iff in h1. rewrite regular_iff in h2. rewrite <- h1. rewrite <- h2. reflexivity. Qed. Definition Btype_ro := {S:(Ensemble (point_set X)) | regular _ S}. Definition BS_ro := Full_set Btype_ro. Definition Bplus_ro (P Q:Btype_ro) : Btype_ro. destruct P as [P h1]. destruct Q as [Q h2]. refine (exist _ ((Union P Q) - ' - ') _). apply open_regular_comp_clo. apply closure_closed. Defined. Definition Btimes_ro (P Q:Btype_ro) : Btype_ro. destruct P as [P h1]. destruct Q as [Q h2]. refine (exist _ (Intersection P Q) _). apply intersection_regular; assumption. Defined. Definition Bone_ro : Btype_ro := (exist _ (Full_set (point_set X)) (full_closed_reg X)). Definition Bzero_ro : Btype_ro := (exist _ (Empty_set (point_set X)) (empty_closed_reg X)). Definition Bcomp_ro (S:Btype_ro) : Btype_ro. destruct S as [S h1]. refine (exist _ (S - ') _). apply open_regular_comp_clo. apply regular_open. assumption. Defined. Definition Bc_ro := Build_Bconst Btype_ro BS_ro Bplus_ro Btimes_ro Bone_ro Bzero_ro Bcomp_ro. Infix "ro+" := (Bplus_ro) (at level 50, left associativity). Infix "ro*" := (Btimes_ro) (at level 40, left associativity). Definition ro0 := Bzero_ro. Definition ro1 := Bone_ro. Notation "'ro-' x" := (Bcomp_ro x) (at level 30). Lemma complement_closure_union : forall (P Q:Ensemble (point_set X)), Union P Q - ' = Intersection (P - ') (Q - '). intros P Q. pose proof (closure_union P Q) as h1. rewrite h1. apply comp_union. Qed. Lemma regular_comp_clo_inj : forall (P Q:Ensemble (point_set X)), regular _ P -> regular _ Q -> P - ' = Q - ' -> P = Q. intros P Q h1 h2 h3. pose proof (f_equal (@closure X) h3) as h4. pose proof (f_equal (@Comp (point_set X)) h4) as h5. red in h1. rewrite interior_eq in h1. red in h2. rewrite interior_eq in h2. congruence. Qed. Lemma assoc_sum_ro : forall P Q R : Btype_ro, P ro+ (Q ro+ R) = P ro+ Q ro+ R. intros P Q R. pose proof (proj2_sig (P ro+ (Q ro+ R))) as h1. pose proof (proj2_sig (P ro+ Q ro+ R)) as h2. unfold Btimes_ro, Bplus_ro. destruct P as [P h3]. destruct Q as [Q h4]. destruct R as [R h5]. simpl in h1. simpl in h2. apply existTexist. apply subsetT_eq_compat. apply regular_comp_clo_inj; try assumption. assert (h6:open (Union P (Union Q R - ' - '))). apply open_union2; [apply regular_open; assumption | apply closure_closed]. pose proof (open_regularization_comp_clo _ h6) as h7. rewrite complement_closure_union in h7 at 1. assert (h8: open (Union Q R)). apply open_union2; apply regular_open; assumption. pose proof (open_regularization_comp_clo _ h8) as h9. rewrite <- h9 in h7. rewrite complement_closure_union in h7 at 1. rewrite <- h7. assert (h10:open (Union (Union P Q - ' - ') R)). apply open_union2; [apply closure_closed | apply regular_open; assumption]. pose proof (open_regularization_comp_clo _ h10) as h11. rewrite complement_closure_union in h11 at 1. assert (h12: open (Union P Q)). apply open_union2; apply regular_open; assumption. pose proof (open_regularization_comp_clo _ h12) as h13. rewrite <- h13 in h11. rewrite complement_closure_union in h11 at 1. rewrite <- h11. rewrite <- Intersection_associative; auto. Qed. Lemma assoc_prod_ro : forall P Q R : Btype_ro, P ro* (Q ro* R) = P ro* Q ro* R. intros P Q R. destruct P. destruct Q. destruct R. unfold Btimes_ro, Bplus_ro. apply existTexist. apply subsetT_eq_compat. rewrite <- Intersection_associative; auto. Qed. Lemma comm_sum_ro : forall P Q: Btype_ro, P ro+ Q = Q ro+ P. intros P Q. destruct P. destruct Q. unfold Btimes_ro, Bplus_ro. apply existTexist. apply subsetT_eq_compat. do 4 f_equal. auto with sets. Qed. Lemma comm_prod_ro : forall P Q : Btype_ro, P ro* Q = Q ro* P. intros P Q. destruct P. destruct Q. unfold Btimes_ro, Bplus_ro. apply existTexist. apply subsetT_eq_compat. apply Intersection_commutative. Qed. Lemma abs_sum_ro : forall P Q:Btype_ro, P ro+ P ro* Q = P. intros P Q. destruct P as [P h1]. destruct Q as [Q h2]. unfold Btimes_ro, Bplus_ro. apply existTexist. apply subsetT_eq_compat. rewrite dist_prod_psa. assert (h5:Union P P = P). apply Extensionality_Ensembles; split; red; auto with sets. intros ? h6. destruct h6; assumption. rewrite h5. assert (h6:Included P (Union P Q)). auto with sets. rewrite inclusion_iff_intersection_eq in h6. rewrite h6. red in h1. rewrite interior_eq in h1. symmetry. assumption. Qed. Lemma abs_prod_ro : forall P Q:Btype_ro, P ro* (P ro+ Q) = P. intros P Q. destruct P as [P h1]. destruct Q as [Q h2]. unfold Btimes_ro, Bplus_ro. apply existTexist. apply subsetT_eq_compat. pose proof (regular_open _ _ h1) as h3. pose proof (regular_open _ _ h2) as h4. pose proof (open_union2 _ _ _ h3 h4) as h5. red in h1. rewrite interior_eq in h1. rewrite h1 at 1. rewrite <- reg_preserves_int; try assumption. assert (h6: Included P (Union P Q)). auto with sets. rewrite inclusion_iff_intersection_eq in h6. rewrite h6. rewrite h1 at 2. reflexivity. Qed. Lemma dist_sum_ro : forall (Q R P:Btype_ro), P ro* (Q ro+ R) = (P ro* Q) ro+ (P ro* R). unfold Btimes_ro, Bplus_ro. intros Q R P. destruct Q as [Q h1]. destruct R as [R h2]. destruct P as [P h3]. apply existTexist. apply subsetT_eq_compat. red in h3. pose proof (interior_eq _ (P -)) as h4. rewrite <- h3 in h4. rewrite h4 at 1. rewrite <- reg_preserves_int. rewrite Distributivity with _ Q R P. reflexivity. apply regular_open. assumption. apply regular_open in h1. apply regular_open in h2. apply open_union2; assumption. Qed. Lemma dist_prod_ro : forall (Q R P:Btype_ro), P ro+ Q ro* R = (P ro+ Q) ro* (P ro+ R). unfold Btimes_ro, Bplus_ro. intros Q R P. destruct Q as [Q h1]. destruct R as [R h2]. destruct P as [P h3]. apply existTexist. apply subsetT_eq_compat. rewrite dist_prod_psa. apply reg_preserves_int; apply open_union2; apply regular_open; assumption. Qed. Lemma comp_prod_ro : forall (P:Btype_ro), P ro* (ro- P) = ro0. intro P. destruct P as [P h1]. unfold ro0. unfold Bzero_ro. unfold Btimes_ro. simpl. apply existTexist. apply subsetT_eq_compat. pose proof (excl_middle_empty (P -)) as h2. pose proof (closure_inflationary P) as h3. pose proof (intersection_preserves_inclusion _ _ (P - ') h3) as h4. rewrite Intersection_commutative in h4. rewrite Intersection_commutative in h2. rewrite h2 in h4. apply Extensionality_Ensembles; red; split; auto with sets. Qed. (*"Lemma 5. The boundary of an open set is a nowhere dense closed set."*) Lemma boundary_open_set : forall (P:Ensemble (point_set X)), open P -> let B := boundary P in closed B /\ now_dense B. Proof. intros P h1 B. pose proof (open_boundary_difference _ _ h1) as h2. rewrite setminus_int_complement in h2. assert (h3:closed B). unfold B. rewrite h2. apply closed_intersection2; [apply closure_closed | red; rewrite Complement_Complement; assumption]. split. (*closed*) assumption. (*now_dense*) rewrite now_dense_iff2. pose proof (closure_fixes_closed _ h3) as h4. rewrite h4. intros U h5 h6. apply NNPP. intro h7. apply not_empty_Inhabited in h7. destruct h7 as [x h7]. assert (h8:In (P -) x). unfold B in h6. rewrite h2 in h6. red in h6. specialize (h6 x h7). inversion h6. assumption. pose proof (closure_impl_meets_every_open_neighborhood _ P x h8 _ h5 h7) as h9. assert (h10:Intersection P U = Empty_set _). assert (h11:Included U (P ')). unfold B in h6. rewrite h2 in h6. red. intros y h11. red in h6. specialize (h6 _ h11). auto with sets. destruct h6; assumption. pose proof (included_empty_complement_int U (P ')) as h12. rewrite h12 in h11. rewrite Complement_Complement in h11. rewrite Intersection_commutative in h11. assumption. apply Inhabited_not_empty in h9. contradiction. Qed. Lemma comp_sum_ro : forall P:Btype_ro, P ro+ (ro- P) = ro1. intros P. unfold Btimes_ro, Bplus_ro, ro1. unfold Bone_ro. destruct P as [P h1]. simpl. apply existTexist. apply subsetT_eq_compat. apply regular_open in h1. pose proof (boundary_open_set _ h1) as h2. destruct h2 as [h2l h2r]. rewrite now_dense_iff4 in h2r. pose proof (closure_fixes_closed _ h2l) as h3. rewrite h3 in h2r. pose proof (open_boundary_difference _ _ h1) as h4. rewrite setminus_int_complement in h4. rewrite h4 in h2r. rewrite comp_int in h2r. rewrite Complement_Complement in h2r. red in h2r. rewrite Union_commutative in h2r. rewrite h2r. rewrite complement_full. rewrite closure_empty. apply complement_empty. Qed. Definition RO := Build_Bool_Alg Bc_ro (eq_refl BS_ro) assoc_sum_ro assoc_prod_ro comm_sum_ro comm_prod_ro abs_sum_ro abs_prod_ro dist_sum_ro dist_prod_ro comp_sum_ro comp_prod_ro. Variable It:Type. Lemma sup_regular : forall P:It->(Ensemble (point_set X)), regular _ ((IndexedUnion P) - ' - '). intro. rewrite regular_iff. rewrite <- open_regularization_comp_clo; [reflexivity | apply closure_closed]. Qed. Lemma inf_regular : forall P:It->(Ensemble (point_set X)), regular _ (IndexedIntersection P - ' - '). intro. rewrite regular_iff. rewrite <- open_regularization_comp_clo; [reflexivity | apply closure_closed]. Qed. (*End RegularOpenAlgebra.*) Let Bt := (Btype (Bc RO)). (*". . . the Boolean order relation for regular open sets is the same as ordinary set-theoretic inclusion."*) Lemma le_iff_inclusion_ro : forall (P Q:Bt), le P Q <-> Included (proj1_sig P) (proj1_sig Q). intros P Q. unfold le. rewrite eq_ord. simpl. unfold Btimes_ro. destruct P as [P]. destruct Q as [Q]. simpl. split. (* -> *) intro h1. apply exist_injective in h1. apply inclusion_iff_intersection_eq; assumption. (* <- *) intro h1. apply existTexist. apply subsetT_eq_compat. apply inclusion_iff_intersection_eq. assumption. Qed. Definition proj1_sig_ind_ro {It:Type} (P:It->Bt) := fun (i:It) => proj1_sig (P i). (*"Lemma 6. The supremum of a family {P_i} of regular open sets is (\/P_i) - ' - '"*) Lemma sup_ro : forall (P:It->Bt), Sup P (exist _ (IndexedUnion (proj1_sig_ind_ro P) - ' - ') (sup_regular (proj1_sig_ind_ro P))). intros P. do 3 red. split. (* Upper bound *) red. intros Pi h1. apply le_iff_inclusion_ro. simpl. assert (h2:Included (proj1_sig Pi) (IndexedUnion (proj1_sig_ind_ro P))). red. intros x h3. destruct h1 as [i ? ? h4]. apply indexed_union_intro with i. unfold proj1_sig_ind_ro. rewrite h4 in h3. assumption. assert (h3:open (IndexedUnion (proj1_sig_ind_ro P))). apply open_indexed_union. intro i. apply regular_open. unfold proj1_sig_ind_ro. apply proj2_sig. pose proof (open_included_reg _ h3) as h4. auto with sets. (* Least Upper Bound *) intros Q h1. destruct Q as [Q h2]. assert (h3: Included (IndexedUnion (proj1_sig_ind_ro P)) Q). red. intros x h4. destruct h4 as [i x h5]. red in h1. assert (h6:In (Im (Full_set It) P) (P i)). apply Im_intro with i. apply Full_intro. reflexivity. specialize (h1 _ h6). rewrite le_iff_inclusion_ro in h1. simpl in h1. auto with sets. rewrite le_iff_inclusion_ro. simpl. red in h2. clear h1. rewrite interior_eq in h2. rewrite h2. do 2 (apply comp_closure_inclusion_quasi_sym). assumption. Qed. Lemma inf_ro : forall (P:It->Bt), Inf P (exist _ (IndexedIntersection (proj1_sig_ind_ro P) - ' - ') (inf_regular (proj1_sig_ind_ro P))). Proof. intros P. do 3 red. split. (* Lower bound *) red. intros Pi h1. apply le_iff_inclusion_ro. simpl. assert (h2:Included (IndexedIntersection (proj1_sig_ind_ro P)) (proj1_sig Pi)). red. intros x h3. destruct h1 as [i ? ? h4]. rewrite h4. unfold proj1_sig_ind_ro in h3. inversion h3 as [? h5]. apply h5. destruct Pi as [Pi h3]. simpl. simpl in h2. red in h3. clear h1. rewrite interior_eq in h3. rewrite h3. do 2 (apply comp_closure_inclusion_quasi_sym). assumption. (* Greatest Lower Bound *) intros Q h1. destruct Q as [Q h2]. assert (h3:Included Q (IndexedIntersection (proj1_sig_ind_ro P))). red. intros x h4. unfold proj1_sig_ind_ro. red in h1. constructor. intro i. assert (h5:In (Im (Full_set It) P )(P i)). apply Im_intro with i. apply Full_intro. reflexivity. specialize (h1 _ h5). rewrite le_iff_inclusion_ro in h1. simpl in h1. auto with sets. rewrite le_iff_inclusion_ro. simpl. red in h2. clear h1. rewrite interior_eq in h2. rewrite h2. do 2 (apply comp_closure_inclusion_quasi_sym). assumption. Qed. End RegularOpenAlgebra.