Require Export set11. Set Implicit Arguments. (* $Id: sete2.v,v 1.52 2010/07/09 16:24:50 grimm Exp $ *) Module Exercice2. Definition example_worder := union2(doubleton (J TPa TPa) (J TPb TPb)) (singleton (J TPa TPb)). Lemma example_worder_related : forall x y, related example_worder x y = ( (x= TPa & y = TPa) \/ (x= TPb & y = TPb) \/ (x= TPa & y = TPb)). Proof. ir. uf related. uf example_worder. app iff_eq. ir. nin (union2_or H). nin (doubleton_or H0). left. ee. inter2tac. inter2tac. right. left. ee. inter2tac. inter2tac. cp (singleton_eq H0). right. right. ee. inter2tac. inter2tac. ir. nin H. ee. rw H; rw H0. fprops. app union2_first. fprops. nin H. app union2_first. ee. rw H; rw H0. fprops. app union2_second. ee. rw H; rw H0. fprops. Qed. Lemma substrate_example_worder : substrate example_worder = two_points. Proof. assert (Hb:is_graph example_worder). red. uf canonical_doubleton_order. ir. nin (union2_or H). nin (doubleton_or H0). rw H1. fprops. rw H1. fprops. db_tac. set_extens. rwi (inc_substrate_rw x Hb) H. nin H. nin H. cp (example_worder_related x x0). ufi related H0. rwi H0 H. nin H. ee. rw H. uf TPa. app R_inc. nin H. ee. rw H. uf TPb; ap R_inc. ee; rw H. uf TPa. ap R_inc. nin H. cp (example_worder_related x0 x). ufi related H0. rwi H0 H. nin H. ee. rw H1. uf TPa; ap R_inc. nin H. ee. rw H1. uf TPb; ap R_inc. ee; rw H1. uf TPb; ap R_inc. rwi two_points_pr H. nin H; assert (related example_worder x x); try rw example_worder_related. left. eee. substr_tac. right. left. rw H. eee. substr_tac. Qed. Lemma example_is_worder: worder example_worder. Proof. assert (Hb:is_graph example_worder). red. uf canonical_doubleton_order. ir. nin (union2_or H). nin (doubleton_or H0). rw H1. fprops. rw H1. fprops. db_tac. assert (order example_worder). red; ee; red; split; tv; ir ; rewrite example_worder_related in *. ir. rwi substrate_example_worder H. rwi two_points_pr H. nin H. left. eee. right. left. eee. ir. nin H. nin H. wri H1 H. rww H. nin H. nin H. wri H1 H. rww H. nin H0. left. intuition. right. right. nin H0. intuition. intuition. ir. nin H. nin H. rww H1. nin H. nin H. rww H1. nin H0. nin H0. rww H0. nin H0. nin H0. rww H0. nin H0. rw H0. nin H; am. red. ee. am. ir. rwi substrate_example_worder H0. nin (inc_or_not TPa x). ir. exists TPa. red. aw. ee. am. ir. aw. rw example_worder_related. cp (H0 _ H3). rwi two_points_pr H4. nin H4. rw H4. intuition. rw H4. intuition. rw substrate_example_worder. am. nin H1. exists y. red. aw. ee. am. ir. aw. rw example_worder_related. cp (H0 _ H3). rwi two_points_pr H4. nin H4. wri H4 H2. elim H2. am. rw H4. cp (H0 _ H1). rwi two_points_pr H5. nin H5. intuition. intuition. rww substrate_example_worder. Qed. Lemma inductive_example1: forall A F, sub A (powerset F) -> (forall S, (forall x y, inc x S -> inc y S -> sub x y \/ sub y x) -> inc (union S) A) -> inductive_set (inclusion_suborder A). Proof. ir. red. ir. exists (union X). red. assert (inc (union X) A). app H0. ir. red in H2. ee. awi H5. cp (H5 _ _ H3 H4). awi H6. nin H6. ee. left. am. red in H6. awi H6. ee; right. am. am. am. am. am. fprops. am. split. aw. ir. aw. awi H1. ee. app H1. am. app union_sub. Qed. Lemma inductive_graphs: forall a b, inductive_set (opposite_order (extension_order a b)). Proof. ir. cp (extension_is_order a b). red. ir. nin H1. ee. awii H2. awii H0. assert (Hd: forall i j, inc i X -> inc j X -> agrees_on (intersection2 (source i) (source j)) i j). ir. cp (H0 _ H3). cp (H0 _ H4). bwi H5; bwi H6. ee. cp (H2 _ _ H3 H4). ufi gge H11. red. ee. app intersection2sub_first. app intersection2sub_second. awii H11. nin H11; ee. ir. app W_extends. inter2tac. ir. sy. app W_extends. inter2tac. assert (He:forall i, inc i X -> function_prop i (source i) b). ir. red. cp (H0 _ H3). bwi H4. eee. cp (extension_covering _ _ He Hd). nin H3. clear H4. nin H3. ee. red in H3. ee. assert (sub (source x) a). rw H5. red. ir. nin (unionf_exists H7). nin H8. cp (H0 _ H8). awi H10. ee. bwi H10. ee. app H11. assert (inc x (set_of_sub_functions a b)). bw. eee. exists x. red. ee. aw. ir. aw. eee. red. cp (H0 _ H9). bwi H10. ee. am. am. cp (H4 _ H9). red in H13. ee. red. ir. cp (in_graph_W H10 H16). rwi H17 H16. cp (inc_pr1graph_source H10 H16). rw H17. wr (H15 _ H18). app W_pr3. app H13. rw H6. rw H12. fprops. app opposite_is_order. Qed. Lemma Exercise1_1: forall r, let E:= substrate r in let s := fun x y => (inc x E & inc y E & glt r x y) in order r -> (exists x, exists y, x <> y & related r x y) -> (transitive_r s & antisymmetric_r s & ~(reflexive_rr s)). Proof. ir. split. red. uf s. uf glt. ir. eee. order_tac. red. ir. wri H9 H4. elim H8. order_tac. split. red. uf s. uf glt. ir. ee. order_tac. red. ir. red in H1. nin H0. nin H0. nin H0. assert (s x x0). uf s. split. uf E. substr_tac. split. uf E. substr_tac. red. intuition. cp (H1 _ _ H3). nin H4. ufi s H4. ufi glt H4. intuition. Qed. Definition quotient_order_r r s X Y := inc X (quotient s) & inc Y (quotient s) & forall x, inc x X -> exists y, inc y Y & gle r x y. Definition quotient_order r s := graph_on (quotient_order_r r s) (quotient s). Definition increasing_pre f r r':= is_function f & preorder r & preorder r' & substrate r = source f & substrate r' = target f & forall x y, gle r x y -> gle r' (W x f) (W y f). Lemma Exercise1_2a: forall r s, is_equivalence s -> preorder r -> substrate s = substrate r -> preorder_r (quotient_order_r r s). Proof. ir. red. split. uf quotient_order_r . red. ir. eee. ir. cp (H7 _ H8). destruct H9 as [ a [ H9 H10]]. cp (H5 _ H9). nin H11. nin H11. ex_tac. red in H0; ee. nin H13. app (H14 _ _ _ H10 H12). red. uf quotient_order_r. red in H0. ee. nin H0. ir. eee. ir. ex_tac. app H3. wr H1. app (inc_in_quotient_substrate H H7 H4). ir. ex_tac. app H3. wr H1. app (inc_in_quotient_substrate H H7 H5). Qed. Lemma quotient_order_pr: forall r s x y, related (quotient_order r s) x y = quotient_order_r r s x y. Proof. ir. uf quotient_order. uf graph_on. uf related. ap iff_eq. ir. Ztac. awi H1. am. ir. Ztac. red in H. eee. aw. Qed. Lemma quotient_is_preorder: forall r s, is_equivalence s -> preorder r -> substrate s = substrate r -> preorder (quotient_order r s). Proof. ir. uf quotient_order. app preorder_from_rel. app Exercise1_2a. Qed. Lemma substrate_quotient_order: forall r s, is_equivalence s -> preorder r -> substrate s = substrate r -> substrate (quotient_order r s) = quotient s. Proof. ir. cp (quotient_is_preorder H H0 H1). set_extens. rwi preorder_reflexivity H3. rwi quotient_order_pr H3. red in H3. nin H3. am. am. rw preorder_reflexivity. rw quotient_order_pr. red. split. am. split. am. ir. ex_tac. wr preorder_reflexivity. wr H1. app (inc_in_quotient_substrate H H4 H3). am. am. Qed. Lemma Exercise1_2b1: forall r s g r', is_equivalence s -> preorder r -> substrate s = substrate r -> is_function g -> preorder r' -> quotient s = source g -> increasing_pre (compose g (canon_proj s)) r r' -> increasing_pre g (quotient_order r s) r'. Proof. ir. red. eee. app quotient_is_preorder. rww substrate_quotient_order. red in H5. ee. awi H9. am. assert (composable g (canon_proj s)). red. ir. eee. sy; aw. red. ir. rwi quotient_order_pr H7. red in H5. red in H7. ee. cp (inc_rep_itself H H7). nin (H9 _ H15). nin H16. cp (H14 _ _ H17). awi H12. assert (inc x0 (substrate s)). wr H12. order_tac. assert (inc (rep x) (substrate s)). wr H12. order_tac. rwii compose_W H18. rwii compose_W H18. awii H18. assert (class s (rep x) = x). app class_rep. rwii H21 H18. assert (y = class s x0). assert (class s (rep y) = y). app class_rep. wr H22. cp (related_rep_in_class H H8 H16). rwi related_rw H23. eee. am. ue. aw. aw. Qed. Definition weak_order_compatibility r s:= preorder r & is_equivalence s & substrate s = substrate r & forall x y x', gle r x y -> related s x x' -> exists y', (related s y y' & gle r x' y'). Lemma strong_order_compatibility: forall r s, preorder r -> is_equivalence s -> substrate s = substrate r -> (forall x x' y, gle r x y -> related s x x' -> gle r x' y) -> weak_order_compatibility r s. Proof. ir. red. eee. ir. exists y. split. app reflexivity_e. rw H1. red in H3. substr_tac. ap (H2 _ _ _ H3 H4). Qed. Lemma compatibility_proj_increasing: forall r s, is_equivalence s -> preorder r -> substrate s = substrate r -> weak_order_compatibility r s = increasing_pre (canon_proj s) r (quotient_order r s). Proof. ir. cp (quotient_is_preorder H H0 H1). uf increasing_pre. aw. wr H1. rww substrate_quotient_order. assert (Ha: is_graph s). fprops. ap iff_eq; ir; eee. ir. rw quotient_order_pr. uf quotient_order_r. assert (inc x (substrate s)). rw H1. order_tac. assert (inc y (substrate s)). rw H1. order_tac. aw. ee. gprops. gprops. ir. red in H3. bwi H7. ee. nin (H10 _ _ _ H4 H7). exists x1. bw. am. red. eee. ir. cp (H8 _ _ H9). rwi quotient_order_pr H11. ufi quotient_order_r H11. ee. wri inc_class H10. awii H13. nin (H13 _ H10). nin H14. bwi H14. exists x0. au. am. rw H1. order_tac. rw H1. order_tac. am. Qed. Lemma exercise1_2c1: forall r1 r2, preorder r1 -> preorder r2 -> weak_order_compatibility (product2_order r1 r2) (first_proj_eq (substrate r1) (substrate r2)). Proof. ir. red. ee. app product2_order_preorder. ap first_proj_equivalence. rw first_proj_substrate. rww product2_order_preorder_substrate. ir. rwi product2_order_pr H1. ee. rwi first_proj_eq_related H2. ee. exists (J (P y) (Q x')). rw product2_order_pr. rw first_proj_eq_related. awi H2. awi H6. aw. awi H3. eee. wrr preorder_reflexivity. Qed. Lemma exercise1_2c2: forall r1 r2, let compatibility r s := (forall x x' y, gle r x y -> related s x x' -> gle r x' y) in preorder r1 -> preorder r2 -> nonempty (substrate r1) -> compatibility (product2_order r1 r2) (first_proj_eq (substrate r1) (substrate r2)) -> r2 = coarse (substrate r2). Proof. ir. assert (preorder (product2_order r1 r2)). app product2_order_preorder. set_extens. uf coarse. assert (is_pair x). nin H0. nin H0. app H0. app product_inc. substr_tac. substr_tac. ufi coarse H4. awi H4. ee. nin H1. set (x1:= J y (Q x)). set (x2:= J y (P x)). assert (inc x1 (product (substrate r1) (substrate r2))). uf x1. fprops. assert (inc x2 (product (substrate r1) (substrate r2))). uf x2. fprops. assert (gle (product2_order r1 r2) x1 x1). wrr preorder_reflexivity. rww product2_order_preorder_substrate. set (s:=first_proj_eq (substrate r1) (substrate r2)). assert (related s x1 x2). uf s. rw first_proj_eq_related. split. am. split. am. uf x1. uf x2. aw. fold s in H2. cp (H2 _ _ _ H9 H10). rwi product2_order_pr H11. eee. ufi x1 H14; ufi x2 H14. red in H14. awi H14. red in H14. awi H14. am. am. Qed. Lemma exercise1_2c3: forall r1 r2, let compatibility r s := (forall x y y', gle r x y -> related s y y' -> gle r x y') in preorder r1 -> preorder r2 -> nonempty (substrate r1) -> compatibility (product2_order r1 r2) (first_proj_eq (substrate r1) (substrate r2)) -> r2 = coarse (substrate r2). Proof. ir. assert (preorder (product2_order r1 r2)). app product2_order_preorder. set_extens. uf coarse. assert (is_pair x). nin H0; nin H0. app H0. app product_inc. substr_tac. substr_tac. ufi coarse H4. awi H4. ee. nin H1. set (x1:= J y (P x)). set (x2:= J y (Q x)). assert (inc x1 (product (substrate r1) (substrate r2))). uf x1. fprops. assert (inc x2 (product (substrate r1) (substrate r2))). uf x2. fprops. assert (gle (product2_order r1 r2) x1 x1). wrr preorder_reflexivity. rww product2_order_preorder_substrate. set (s:=first_proj_eq (substrate r1) (substrate r2)). assert (related s x1 x2). uf s. rw first_proj_eq_related. split. am. split. am. uf x1. uf x2. aw. fold s in H2. cp (H2 _ _ _ H9 H10). rwi product2_order_pr H11. ee. ee. ufi x1 H14; ufi x2 H14. red in H14. red in H14. awi H14. am. am. Qed. Definition preorder_isomorphism f r r' := (preorder r) & (preorder r') & (bijective f) & (substrate r = source f) & (substrate r' = target f) & (forall x y, inc x (source f) -> inc y (source f) -> gle r x y = gle r' (W x f) (W y f)). Lemma exercise1_2c4: forall r1 r2 f, let s := (first_proj_eq (substrate r1) (substrate r2)) in let r:= product2_order r1 r2 in is_function f -> source f = quotient s -> target f = (substrate r1) -> preorder r1 -> preorder r2 -> nonempty (substrate r2) -> compose f (canon_proj s)=(first_proj (product (substrate r1) (substrate r2))) -> preorder_isomorphism f (quotient_order r s) r1. Proof. ir. set (E1:= substrate r1) in *. set (E2:= substrate r2) in *. assert (Ha:substrate r = product E1 E2). uf r. rww product2_order_preorder_substrate. assert(Hb:substrate s = product E1 E2). uf s. rww first_proj_substrate. assert (Hc:is_equivalence s). uf s; app first_proj_equivalence. assert (composable f (canon_proj s)). red. eee;aw; app function_canon_proj. assert (Hd:preorder r). uf r. app product2_order_preorder. assert (source (canon_proj s)= product E1 E2). aw. red. ee. app quotient_is_preorder. ue. am. red. ee. red. ee. am. ir. assert (inc x (quotient s)). wrr H0. cp (canon_proj_show_surjective Hc H11). assert (inc y (quotient s)). wrr H0. cp (canon_proj_show_surjective Hc H13). wr (related_rep_rep Hc H11 H13). uf s. rw first_proj_eq_related. eee. wri H12 H10. wri H14 H10. assert (inc (rep x) (source (canon_proj s))). rww H7. ue. wri (compose_W H6 H15) H10. assert (inc (rep y) (source (canon_proj s))). rww H7. ue. wri (compose_W H6 H16) H10. rwi H5 H10. rwii first_proj_W H10. rwii first_proj_W H10. fprops. ue. fprops. ue. app surjective_pr6. rw H1. ir. nin H4. exists (W (J y y0) (canon_proj s)). assert (inc (J y y0) (source (canon_proj s))). ue. ee. assert (source f = target (canon_proj s)). aw. rw H10. app inc_W_target. app canon_proj_function. wr (compose_W H6 H9). rw H5. rw first_proj_W. aw. fprops. fprops. sy; rww substrate_quotient_order. rww Ha. sy. am. ir. set (u:= W x f). set (v:= W y f). assert (inc x (quotient s)). wrr H0. cp (canon_proj_show_surjective Hc H10). assert (inc y (quotient s)). wrr H0. cp (canon_proj_show_surjective Hc H12). assert (u = W x f). tv. wri H11 H14. assert (v = W y f). tv. wri H13 H15. assert (inc (rep x) (source (canon_proj s))). rww H7. ue. wri (compose_W H6 H16) H14. assert (inc (rep y) (source (canon_proj s))). rww H7. ue. wri (compose_W H6 H17) H15. rwi H5 H14. rwi H5 H15. rwi first_proj_W H14. rwi first_proj_W H15. rw H14; rw H15. rw quotient_order_pr. uf quotient_order_r. app iff_eq. ir. ee. assert (inc (rep x) x). app (inc_rep_itself Hc H10). nin (H20 _ H21). nin H22. ufi r H23. rwi product2_order_pr H23. ee. cp (is_class_pr Hc H22 H19). assert (inc x0 (substrate s)). app (inc_in_quotient_substrate Hc H22 H19). cp (related_rep_class Hc H28). wri H27 H29. ufi s H29. rwi first_proj_eq_related H29. eee. ir. eee. ir. assert(inc (J v (Q x0)) (product (substrate r1) (substrate r2))). aw. ee. fprops. rwi H7 H17. awi H17. eee. cp (inc_in_quotient_substrate Hc H19 H10). rwi Hb H20; awi H20. eee. assert (inc x0 (product (substrate r1) (substrate r2))). cp (inc_in_quotient_substrate Hc H19 H10). rwi Hb H21. am. exists (J v (Q x0)). ee. wr (class_rep Hc H12). bw. uf s. rw first_proj_eq_related. ee. wr H7. am. am. sy. aw. uf r. rw product2_order_pr. eee. aw. rw H15. assert (inc (rep x) x). app (inc_rep_itself Hc H10). cp (is_class_pr Hc H19 H10). assert (inc x0 (substrate s)). app (inc_in_quotient_substrate Hc H19 H10). cp (related_rep_class Hc H24). wri H23 H25. ufi s H25. rwi first_proj_eq_related H25. eee. aw. red in H3. ee. red in H22. ee. nin H3. app H24. awi H21; eee. fprops. ue. fprops. ue. Qed. Definition quotient_order_axiom r s:= forall x y z, gle r x y -> gle r y z -> related s x z -> related s x y. Lemma order_is_preorder: forall r, order r -> preorder r. Proof. ir. assert (is_graph r). fprops. red in H. red. intuition. Qed. Lemma Exercise1_2d: forall r s, is_equivalence s -> order r -> substrate s = substrate r -> quotient_order_axiom r s -> order (quotient_order r s). Proof. ir. cp (order_is_preorder H0). cp (quotient_is_preorder H H3 H1). red in H4. ee. red. intuition. red. ee. am. ir. rwi quotient_order_pr H7. rwi quotient_order_pr H8. red in H7. red in H8. ee. set_extens. nin (H12 _ H13). ee. nin (H10 _ H14). ee. assert (related s x0 x2). rw in_class_related. exists x. eee. wrr inc_quotient. am. cp (H2 _ _ _ H15 H17 H18). cp (is_class_pr H H14 H11). rw H20. bw. app symmetricity. nin (H10 _ H13). ee. nin (H12 _ H14). ee. assert (related s x0 x2). rw in_class_related. exists y. eee. wrr inc_quotient. am. cp (H2 _ _ _ H15 H17 H18). cp (is_class_pr H H14 H7). rw H20. bw. app symmetricity. Qed. Lemma Exercise1_2f1: forall r r' f, increasing_fun f r r' -> quotient_order_axiom r (equivalence_associated f). Proof. ir. red. ir. red in H. ee. cp (H7 _ _ H0). cp (H7 _ _ H1). rww related_ea. rwii related_ea H2. ee. am. wr H5. order_tac. wri H11 H9. order_tac. Qed. Lemma Exercise1_2f2: forall r r' f, increasing_fun f r r' -> weak_order_compatibility r (equivalence_associated f) = (forall x y x', gle r x y -> inc x' (source f) -> W x f = W x' f -> exists y', inc y' (source f) & gle r x' y' & W y f = W y' f). Proof. ir. uf weak_order_compatibility. red in H. ee. ap iff_eq. ir. ee. assert (related (equivalence_associated f) x x'). rw related_ea. ee. wr H2. order_tac. am. am. am. nin (H11 _ _ _ H6 H12). rwi related_ea H13. ee. exists x0. intuition. am. ir. ee. app order_is_preorder. app equivalence_graph_ea. rww substrate_graph_ea. sy; am. ir. rwi related_ea H7. ee. nin (H5 _ _ _ H6 H8 H9). exists x0. rw related_ea. intuition. wr H2. order_tac. am. am. Qed. Lemma Exercise1_2f3: forall r r' f g, let CC:= forall x y x', gle r x y -> inc x' (source f) -> W x f = W x' f -> exists y', inc y' (source f) & gle r x' y' & W y f = W y' f in let DD:= forall x y, inc x (source f) -> inc y (source f) -> gle r' (W x f) (W y f) -> exists x', exists y', W x f = W x' f & W y f = W y' f & gle r x' y' in increasing_fun f r r' -> composable g (canon_proj (equivalence_associated f)) -> f = compose g (canon_proj (equivalence_associated f)) -> order_morphism g (quotient_order r (equivalence_associated f)) r' = (CC & DD). Proof. ir. cp (Exercise1_2f1 H). rename H2 into Hc. red in H. ee. set (s:= equivalence_associated f) in *. assert (Ha:is_equivalence s). uf s. app equivalence_graph_ea. assert (Hb: substrate s = substrate r). uf s. sy; rww substrate_graph_ea. assert (Hd:preorder r). app order_is_preorder. assert (He: forall x, inc x (quotient s) ->( inc (rep x) (source f) & W (rep x) f = W x g)). ir. cp (inc_rep_substrate Ha H7). split. wr H4. wrr Hb. cp (W_canon_proj Ha H8). rwii class_rep H9. assert (W (W (rep x) (canon_proj s)) g = W x g). rww H9. wri compose_W H10. wri H1 H10. am. am. aw. assert (Hf:forall x a, inc x (quotient s) -> inc a x = (inc (rep x)(source f) &inc a(source f) & W (rep x) f = W a f)). ir. ap iff_eq. ir. cp (related_rep_in_class Ha H7 H8). ufi s H9. rwi related_ea H9. am. am. ir. wr (class_rep Ha H7). bw. uf s. rw related_ea. am. am. app iff_eq. ir. assert (forall x y x', inc x (source f) -> inc y (source f) -> gle r' (W x f) (W y f) -> inc x' (class s x) -> exists y', inc y' (source f) & gle r x' y' & W y f = W y' f). ir. assert(inc x (substrate s)). ue. ue. assert (inc x (class s x)). app inc_itself_class. assert (inc (class s x) (quotient s)). gprops. rwi (Hf _ x H14) H13. cp (He _ H14). ee. wri H18 H10. rwi H16 H10. assert(inc y (substrate s)). ue. ue. assert (inc y (class s y)). app inc_itself_class. assert (inc (class s y) (quotient s)). gprops. rwi (Hf _ y H21) H20. cp (He _ H21). ee. wri H25 H10. rwi H23 H10. red in H7. ee. wri H30 H10. rwi quotient_order_pr H10. red in H10. ee. nin (H32 _ H11). exists x0. nin H33. rwi (Hf _ x0 H21) H33. ee. am. am. assert (inc y (class s y)). app inc_itself_class. rwi (Hf _ y H21) H37. ee. wrr H39. red in H7. ee. wr H28. rww substrate_quotient_order. red in H7. ee. wr H28. rww substrate_quotient_order. ee. red. ir. cp (H6 _ _ H9). rwi H11 H12. assert (inc x' (class s x')). app inc_itself_class. ue. ue. assert (inc y (source f)). wr H4. order_tac. nin (H8 _ _ _ H10 H14 H12 H13). exists x0. am. red. ir. assert (inc x (class s x)). app inc_itself_class. ue. ue. cp (H8 _ _ _ H9 H10 H11 H12). nin H13. exists x. exists x0. eee. red. ir. assert (order (quotient_order r s)). app Exercise1_2d. ufi CC H7; ufi DD H7. nin H7. assert (source g = quotient s). red in H0. ee. awi H11. am. assert(substrate (quotient_order r s) = source g). rw H10. rww substrate_quotient_order. assert (forall x y, inc x (source g) -> inc y (source g) -> gle (quotient_order r s) x y = gle r' (W x g) (W y g)). ir. rwi H10 H12. rwi H10 H13. nin (He _ H12). nin (He _ H13). wr H15. wr H17. ap iff_eq. ir. rwi quotient_order_pr H18. red in H18. ee. nin (H20 _ (inc_rep_itself Ha H18)). nin H21. rwi (Hf y x0 H13) H21. ee. rw H24. app H6. ir. rw quotient_order_pr. red. eee. cp (H9 _ _ H14 H16 H18). nin H19; nin H19. ee. ir. rwi (Hf x x2 H12) H22. nin H22. nin H23. rwi H19 H24. nin (H7 _ _ _ H21 H23 H24). exists x3. ee. rw Hf. eee. am. am. red in H0. eee. red. ee. am. ir. assert (gle r' (W y g) (W y g)). order_tac. rw H5. rw H1. aw. app inc_W_target. assert (gle r' (W x g) (W y g)). rww H17. assert (gle r' (W y g) (W x g)). rww H17. wri H12 H19. wri H12 H20. order_tac. am. am. am. am. rww H1. aw. Qed. Section Exercise1_3a. Variables r f g: Set. Definition E13_lam := second_proj (disjoint_union f). Definition E13_S:= equivalence_associated (second_proj (disjoint_union f)). Definition E13_F:= ordinal_sum r f g. Definition E13_sF:= disjoint_union f. Definition E13_H1:= ordinal_sum_axioms r f g. Definition E13_H2:= is_graph f & (forall i, inc i (domain f) -> nonempty (V i f)). Lemma Exercise1_3a1: is_graph E13_sF. Proof. ir. uf E13_sF. uf disjoint_union. uf disjoint_union_fam. red. ir. srwi H. nin H. nin H. bwi H0. awi H0. nin H0; am. bwi H. am. Qed. Lemma Exercise1_3a0: is_function E13_lam. Proof. uf E13_lam. app function_second_proj. ap Exercise1_3a1. Qed. Lemma Exercise1_3a2: surjective E13_lam. Proof. ap surjective_pr6. ap Exercise1_3a0. cp Exercise1_3a1. uf E13_lam. ir. ufi second_proj H0. awi H0. nin H0. exists (J x y). ee. uf second_proj. aw. rw W_second_proj. aw. am. am. am. Qed. Lemma Exercise1_3a3: E13_H2 -> domain f = target E13_lam. Proof. ir. nin H. uf E13_lam. uf second_proj. aw. set_extens. nin (H0 _ H1). aw. exists y. uf disjoint_union. uf disjoint_union_fam. srw. bw. ex_tac. bw. fprops. ap Exercise1_3a1. awi H1. nin H1. ufi disjoint_union H1. ufi disjoint_union_fam H1. srwi H1. nin H1. nin H1. bwi H1. bwi H2. awi H2. ee. ue. am. ap Exercise1_3a1. Qed. Lemma Exercise1_3a4: E13_H1 -> E13_H2 -> increasing_fun E13_lam E13_F r. Proof. ir. red in H. cp H; red in H. red. ee. ap Exercise1_3a0. uf E13_F. fprops. am. uf E13_F; uf E13_lam; uf second_proj. aw. rw H2. app Exercise1_3a3. uf E13_lam. ir. cp Exercise1_3a1. ufi E13_F H8. awi H8. ee. rww W_second_proj. rww W_second_proj. nin H11. nin H11; am. nin H11. rw H11. order_tac. cp (du_index_pr H10); eee. am. Qed. Definition disjoint_union_function f := corresp (domain f)(range (disjoint_union_fam f))(disjoint_union_fam f). Lemma disjoint_union_function_pr: is_function (disjoint_union_function f) & graph (disjoint_union_function f) = (disjoint_union_fam f). Proof. ir. uf disjoint_union_function. split. app is_function_pr. uf disjoint_union_fam. gprops. fprops. uf disjoint_union_fam. bw. aw. Qed. Lemma Exercise1_3a5: forall x y, fgraph f -> related (partition_relation (disjoint_union_function f) (disjoint_union f)) x y = (inc x E13_sF & inc y E13_sF & Q x = Q y). Proof. ir. uf E13_sF. rename H into Ha. cp disjoint_union_function_pr. ee. rw partition_relation_pr. uf in_same_coset. uf W. uf disjoint_union. uf disjoint_union_function. aw. uf disjoint_union_fam. ap iff_eq. ir. nin H1. ee. srw. bw. ex_tac. union_tac. bw. bwi H2. bwi H3. awi H2; awi H3. eee. am. am. ir. ee. srwi H1. srwi H2. nin H1;nin H2. ee. bwi H1. bwi H2. assert (x0 = x1). bwi H5; bwi H4. awi H5; awi H4. ee. wr H9. ue. am. am. am. ex_tac. split. ue. am. am. rw H0. app partion_union_disjoint. Qed. Lemma Exercise1_3a6: forall x y, related E13_S x y = (inc x E13_sF & inc y E13_sF & Q x = Q y). Proof. ir. cp Exercise1_3a1. assert (is_function E13_lam). ap Exercise1_3a0. uf E13_S. rw related_ea. assert (source (second_proj (disjoint_union f)) = E13_sF). uf second_proj. aw. rw H1. app iff_eq. ir. eee. rwii W_second_proj H4. rwii W_second_proj H4. ir. nin H2. nin H3. eee. rww W_second_proj. rww W_second_proj. am. Qed. Lemma Exercise1_3a7: is_equivalence E13_S. Proof. uf E13_S. app equivalence_graph_ea. ap Exercise1_3a0. Qed. Lemma Exercise1_3a8: forall a, E13_H2 -> inc a (quotient E13_S) = exists i, inc i (domain f) & a = product (V i f) (singleton i). Proof. ir. cp Exercise1_3a7. rww inc_quotient. rww is_class_rw. assert (substrate E13_S = disjoint_union f). uf E13_S. rww substrate_graph_ea. uf second_proj. aw. app function_second_proj. ap Exercise1_3a1. ap iff_eq. ir. ee. nin H2. assert (inc y (disjoint_union f)). wr H1. app H3. cp (du_index_pr H5). ee. ex_tac. set_extens. rwi (H4 y x H2) H9. ufi E13_S H9. rwi Exercise1_3a6 H9. ee. cp (du_index_pr H10). ee. app product_inc;ue. rw (H4 _ x H2). uf E13_S. awi H9. ee. cp H11. rw Exercise1_3a6. ee. am. uf E13_sF. uf disjoint_union. srw. uf disjoint_union_fam. bw. ex_tac. bw. aw. au. sy; am. ir. nin H2. nin H2. assert (sub a (disjoint_union f)). red. ir. uf disjoint_union. srw. uf disjoint_union_fam. bw. ex_tac. bw. ue. ee. nin H. ee. cp (H5 _ H2). nin H6. exists (J y x). ue. ue. ir. uf E13_S. rw Exercise1_3a6. ap iff_eq. ir. ee. app H4. app H4. rwi H3 H5. awi H5. rwi H3 H6. awi H6. ee. ue. ir. ee. rw H3. cp (du_index_pr H7). rwi H3 H5. awi H5. nin H5. nin H10. rwi H11 H8. aw. eee. Qed. Lemma Exercise1_3a9: is_function(fun_on_quotient E13_S E13_lam). Proof. ir. cp Exercise1_3a0. app function_foqc. ap Exercise1_3a7. uf E13_S. rww substrate_graph_ea. Qed. Lemma Exercise1_3a10: composable (fun_on_quotient E13_S E13_lam) (canon_proj E13_S). Proof. red. ee. ap Exercise1_3a9. ap function_canon_proj. ap Exercise1_3a7. uf fun_on_quotient. uf section_canon_proj. aw. Qed. Lemma Exercise1_3a11: E13_lam = compose (fun_on_quotient E13_S E13_lam) (canon_proj E13_S). Proof. ap (canonical_decomposition_surj2 Exercise1_3a2). Qed. Lemma Exercise1_3a12: forall x, E13_H2 -> inc x (quotient E13_S) -> exists i, inc i (domain f) & x = product (V i f) (singleton i) & W x (fun_on_quotient E13_S E13_lam) = i. Proof. ir. cp Exercise1_3a0. cp Exercise1_3a7. assert (source E13_lam = substrate E13_S). uf E13_S. sy; app substrate_graph_ea. assert (Q (rep x) = W (rep x) E13_lam). uf E13_lam. rww W_second_proj. ap Exercise1_3a1. ufi E13_lam H3. ufi second_proj H3. awi H3. rw H3. app inc_rep_substrate. rww W_foqc. wr H4. rwi (Exercise1_3a8 x H) H0. nin H0. nin H0. ex_tac. assert (inc (rep x) x). app nonempty_rep. nin H. ee. nin (H6 _ H0). exists (J y x0). rw H5. aw. ee. fprops. am. fprops. set (t:=rep x) in *. rwi H5 H6. awi H6. eee. Qed. Lemma Exercise1_3a13: E13_H2 -> bijective (fun_on_quotient E13_S E13_lam). Proof. ir. assert (Ha: source (fun_on_quotient E13_S E13_lam) = quotient E13_S). uf fun_on_quotient. uf section_canon_proj. aw. cp Exercise1_3a9. red. split. red. rw Ha. split. am. ir. ir. nin (Exercise1_3a12 H H1). nin (Exercise1_3a12 H H2). ee. rw H8; rw H6; wr H9; rw H3;rw H7. tv. app surjective_pr6. rw Ha. cp (Exercise1_3a3 H). ir. ufi fun_on_quotient H2. awi H2. ir. set (a:=product (V y f) (singleton y)). assert (inc a (quotient E13_S)). rww Exercise1_3a8. exists y. rw H1. au. ex_tac. nin (Exercise1_3a12 H H3). ee. rw H6. cp Exercise1_3a7. cp (inc_rep_itself H7 H3). set (t:= rep a) in H8. cp H8. ufi a H8. awi H8. rwi H5 H9. awi H9. eee. Qed. Lemma Exercise1_3a14: E13_H1 -> quotient_order_axiom E13_F E13_S. Proof. ir. red. uf E13_F. uf E13_S. ir. assert (order (ordinal_sum r f g)). red in H; fprops. rw Exercise1_3a6. rwi Exercise1_3a6 H2. awi H0. awi H1. eee. nin H. nin H9. nin H7. assert (glt r (Q x) (Q z)). order_tac. nin H11. elim H12. am. nin H7. ue. nin H9. ue. am. am. Qed. Lemma Exercise1_3a15: E13_H1 -> E13_H2 -> order_isomorphism (fun_on_quotient E13_S E13_lam) (quotient_order E13_F E13_S) r. Proof. ir. cp Exercise1_3a1. assert (order (ordinal_sum r f g)). fprops. assert (order_morphism (fun_on_quotient E13_S E13_lam) (quotient_order E13_F (equivalence_associated E13_lam)) r). rw Exercise1_2f3. assert (source E13_lam = disjoint_union f). uf E13_lam. uf second_proj. aw. rw H3. uf E13_F. split. ir. awi H4. ee. assert (W x E13_lam = Q x). uf E13_lam. rww W_second_proj. assert (W x' E13_lam = Q x'). uf E13_lam. rww W_second_proj. assert (W y E13_lam = Q y). uf E13_lam. rww W_second_proj. nin H8. exists y. ee. am. aw. eee. left. wr H10. wr H6. rww H9. tv. exists x'. ee. am. order_tac. rww substrate_ordinal_sum. wr H6. rw H11. rw H9. sy. am. am. ir. awi H6. ee. exists x. nin (equal_or_not (W x E13_lam) (W y E13_lam)). exists x. split. tv. split. sy. am. order_tac. rww substrate_ordinal_sum. exists y. eee. aw. eee. left. assert (W x E13_lam = Q x). uf E13_lam. rww W_second_proj. assert (W y E13_lam = Q y). uf E13_lam. rww W_second_proj. wr H8; wr H9; red; split;am. app Exercise1_3a4. app Exercise1_3a10. app Exercise1_3a11. cp Exercise1_3a13. red in H3. red. intuition. Qed. End Exercise1_3a. Lemma ordinal_sum2_order_greatest: forall r r' x, order r -> order r' -> nonempty (substrate r') -> greatest_element (ordinal_sum2 r r') x -> greatest_element r' (P x). Proof. ir. nin H1. red in H2. nin H2. assert (inc (J y TPb) (substrate (ordinal_sum2 r r'))). rw substrate_ordinal_sum2. rw canonical_du2_pr. split. fprops. right. aw. au. am. am. cp (H3 _ H4). rwi substrate_ordinal_sum2 H2. rwi canonical_du2_pr H2. nin H2. nin H6. nin H6. cp (related_ordinal_sum2_order_spec H H0 H6 H1). assert (x = J (P x) TPa). app pair_extensionality. fprops. aw. aw. wri H9 H8. elim (not_le_gt (order_ordinal_sum2 H H0) H5 H8). ee. red. ee. am. ir. assert (inc (J x0 TPb) (substrate (ordinal_sum2 r r'))). rw substrate_ordinal_sum2. rw canonical_du2_pr. split. fprops. right. aw. au. am. am. cp (H3 _ H9). rwi related_ordinal_sum2_order H10. ee. nin H12. nin H12. awi H12. elim two_points_distinctb. am. nin H12. nin H12. nin H13. awi H14. am. nin H12. awi H12. elim two_points_distinctb. am. am. am. am. am. Qed. Lemma directed_ordinal_sum_order: forall r f g, ordinal_sum_axioms1 r f g -> right_directed (ordinal_sum r f g) = (right_directed r & forall i, maximal_element r i -> right_directed (V i g)). Proof. ir. red in H. ee. cp H. red in H1. ee. assert (Ha: order (ordinal_sum r f g)). fprops. ap iff_eq. ir. rwi right_directed_pr H8. awi H8. ee. rw right_directed_pr. ee. am. ir. rwi H2 H10; rwi H2 H11. nin (H0 _ H10). nin (H0 _ H11). cp (inc_disjoint_union H10 H12). cp (inc_disjoint_union H11 H13). nin (H9 _ _ H14 H15). ee. exists (Q x0). cp (du_index_pr H16). ee. rw H2. am. cp (related_ordinal_sum_order_id H H17). awi H22. am. cp (related_ordinal_sum_order_id H H18). awi H22. am. ee. ir. rw right_directed_pr. red in H10. ee. app H6. wrr H2. ir. rwi H2 H10. rwi (H7 _ H10) H12. cp (inc_disjoint_union H10 H12). rwi (H7 _ H10) H13. cp (inc_disjoint_union H10 H13). nin (H9 _ _ H14 H15). ee. cp (related_ordinal_sum_order_id H H17). awi H19. cp (du_index_pr H16). ee. cp (H11 _ H19). awi H17; awi H18. ee. exists (P x0). ee. rww H7. wrr H23. nin H27. nin H27. elim H28. sy; am. ee; am. nin H25. nin H25. elim H28. sy;am. eee. am. am. am. am. ir. ee. rwi right_directed_pr H8. ee. rw right_directed_pr. ee. am. aw. ir. cp (du_index_pr H11). cp (du_index_pr H12). ee. rwi H2 H10. nin (H10 _ _ H13 H14). ee. nin (equal_or_not (Q x) x0). nin (equal_or_not (Q y) x0). nin (p_or_not_p (maximal_element r x0)). cp (H9 _ H24). rwi right_directed_pr H25. ee. rwi H7 H26. rwi H22 H17; rwi H23 H15. nin (H26 _ _ H17 H15). ee. exists (J x1 x0). ee. app inc_disjoint_union. aw. ee. am. app inc_disjoint_union. right. eee. aw. ee. am. app inc_disjoint_union. right. eee. am. ufi maximal_element H24. assert (exists x1, inc x1 (substrate r) & glt r x0 x1). app exists_proof. dneg. ee. rww H2. ir. nin (equal_or_not x0 x1). sy; am. elim (H25 x1). ee. order_tac. split; am. nin H25. ee. rwi H2 H25. nin (H0 _ H25). cp (inc_disjoint_union H25 H27). ex_tac. aw. eee. exists x. ee. am. order_tac. aw. aw. eee. left. rw H22. red. eee. nin (equal_or_not (Q y) x0). exists y. ee. am. aw. eee. left. rw H23. order_tac. order_tac. aw. nin (H0 _ H19). cp (inc_disjoint_union H19 H24). ex_tac. ee. aw. eee. left. order_tac. aw. eee. left. order_tac. Qed. Lemma ordinal_sum_pr1: forall r f g, ordinal_sum_axioms1 r f g -> forall i, inc i (domain f) -> exists y, inc y (V i f) & inc (J y i) (substrate (ordinal_sum r f g)). Proof. ir. red in H. ee. nin (H1 _ H0). ex_tac. set(inc_disjoint_union H0 H2). aw. Qed. Lemma ordinal_sum_lattice1: forall r f g, ordinal_sum_axioms1 r f g -> lattice (ordinal_sum r f g) -> (lattice r). Proof. ir. set (F:= ordinal_sum r f g) in H0. assert (Ha:forall i, inc i (domain f) -> exists y, inc y (V i f) & inc (J y i) (substrate F)). ir. uf F. app ordinal_sum_pr1. nin H. assert (order F). uf F. fprops. cp H; red in H;ee. assert (Hb: substrate F = disjoint_union f). uf F. aw. red. ee. am. ir. rwi H4 H10; rwi H4 H11. nin (Ha _ H10). nin (Ha _ H11). ee. cp (lattice_sup_pr H0 H15 H14). fold F in H16. ee. red. nin (p_or_not_p (gle r x y)). exists y. app sup_comparable. ir. nin (p_or_not_p (gle r y x)). exists x. rw doubleton_symm. app sup_comparable. ir. set (z:= (sup F (J x0 x) (J x1 y))) in *. cp (inc_arg2_substrate H16). rwi Hb H21. cp (du_index_pr H21). ee. exists (Q z). app least_upper_bound_doubleton. ufi F H16. cp (related_ordinal_sum_order_id H3 H16). awi H25. am. ufi F H1. cp (related_ordinal_sum_order_id H3 H17). awi H25. am. ir. cp (inc_arg2_substrate H25). rwi H4 H27. nin (Ha _ H27). nin H28. assert (gle F (J x0 x) (J x2 t)). uf F. aw. eee. left. order_tac. wri H30 H26. contradiction. assert (gle F (J x1 y) (J x2 t)). uf F. aw. eee. left. order_tac. wri H31 H25. contradiction. cp (H18 _ H30 H31). ufi F H32. cp (related_ordinal_sum_order_id H3 H32). awi H33. am. cp (lattice_inf_pr H0 H15 H14). fold F in H16. ee. red. nin (p_or_not_p (gle r x y)). exists x. app inf_comparable. ir. nin (p_or_not_p (gle r y x)). exists y. rw doubleton_symm. app inf_comparable. ir. set (z:= (inf F (J x0 x) (J x1 y))) in *. cp (inc_arg2_substrate H16). rwi Hb H21. cp (du_index_pr H21). ee. exists (Q z). app greatest_lower_bound_doubleton. ufi F H16. cp (related_ordinal_sum_order_id H3 H16). awi H25. am. ufi F H1. cp (related_ordinal_sum_order_id H3 H17). awi H25. am. ir. cp (inc_arg1_substrate H25). rwi H4 H27. nin (Ha _ H27). nin H28. assert (gle F (J x2 t) (J x0 x)). uf F. aw. eee. left. order_tac. rwi H30 H26. contradiction. assert (gle F (J x2 t) (J x1 y)). uf F. aw. eee. left. order_tac. rwi H31 H25. elim H20. am. cp (H18 _ H30 H31). ufi F H32. cp (related_ordinal_sum_order_id H3 H32). awi H33. am. Qed. Lemma ordinal_sum_lattice2: forall r f g, ordinal_sum_axioms1 r f g -> lattice (ordinal_sum r f g) -> (forall i j, inc i (domain f) -> inc j (domain f) -> (gle r i j \/ gle r j i \/ (exists u, exists v, least_element (V (sup r i j) g) u & greatest_element (V (inf r i j) g) v))). Proof. ir. cp (ordinal_sum_lattice1 H H0). set (F:= ordinal_sum r f g) in H0. assert (Ha:forall i, inc i (domain f) -> exists y, inc y (V i f) & inc (J y i) (substrate F)). ir. uf F. app ordinal_sum_pr1. nin H. assert (order F). uf F. fprops. cp H; red in H;ee. assert (Hb: substrate F = disjoint_union f). uf F. aw. nin (p_or_not_p (gle r i j)). au. nin (p_or_not_p (gle r j i)). au. right. right. nin (Ha _ H1). nin (Ha _ H2). ee. cp (lattice_sup_pr H0 H18 H17). cp (lattice_inf_pr H0 H18 H17). set (A:= inf F (J x i) (J x0 j)) in H20. set (a:= inf r i j). set (B:= sup F (J x i) (J x0 j)) in H19. set (b:= sup r i j). ee. wri H7 H1. wri H7 H2. cp (lattice_inf_pr H3 H1 H2). fold a in H25. ee. cp (lattice_sup_pr H3 H1 H2). fold b in H28. ee. assert (Hc:Q A = a). assert (gle r (Q A) a). ufi F H20. cp (related_ordinal_sum_order_id H6 H20). ufi F H21. cp (related_ordinal_sum_order_id H6 H21). awi H31. awi H32. app H27. cp (inc_arg1_substrate H25). rwi H7 H32. nin (Ha _ H32). ee. assert (gle F (J x1 a) (J x i)). uf F. aw. eee. left. order_tac. rwi H35 H26. contradiction. assert (gle F (J x1 a) (J x0 j)). uf F. aw. eee. left. order_tac. rwi H36 H25. contradiction. cp (H22 _ H35 H36). uf F. cp (related_ordinal_sum_order_id H6 H37). awi H38. order_tac. assert (Hd:Q B = b). assert (gle r b (Q B)). ufi F H19. cp (related_ordinal_sum_order_id H6 H19). ufi F H23. cp (related_ordinal_sum_order_id H6 H23). awi H31. awi H32. app H30. cp (inc_arg2_substrate H28). rwi H7 H32. nin (Ha _ H32). ee. assert (gle F (J x i) (J x1 b)). uf F. aw. eee. left. order_tac. wri H35 H29. contradiction. assert (gle F (J x0 j) (J x1 b)). uf F. aw. eee. left. order_tac. wri H36 H28. contradiction. cp (H24 _ H35 H36). uf F. cp (related_ordinal_sum_order_id H6 H37). awi H38. order_tac. cp (inc_arg2_substrate H23). rwi Hb H31. cp (du_index_pr H31). ee. wri (H12 _ H32) H33. rwi Hd H33. cp (inc_arg1_substrate H21). rwi Hb H35. cp (du_index_pr H35). ee. wri (H12 _ H36) H37. rwi Hc H36. exists (P B). exists (P A). ee. red. ee. am. ir. rwii H12 H39. assert (inc (J x1 b) (disjoint_union f)). app inc_disjoint_union. wrr Hd. assert (gle F (J x i) (J x1 b)). uf F. aw. eee. left. order_tac. wri H41 H29. contradiction. assert (gle F (J x0 j) (J x1 b)). uf F. aw. eee. left. order_tac. wri H42 H28. contradiction. cp (H24 _ H41 H42). ufi F H43. awi H43. ee. nin H45. red in H45. ee. contradiction. ee. wrr H45. am. ue. red. ee. ue. ir. rwii H12 H39. assert (inc (J x1 a) (disjoint_union f)). app inc_disjoint_union. assert (gle F (J x1 a) (J x i) ). uf F. aw. eee. left. order_tac. rwi H41 H26. contradiction. assert (gle F (J x1 a) (J x0 j)). uf F. aw. eee. left. order_tac. rwi H42 H25. contradiction. cp (H22 _ H41 H42). ufi F H43. awi H43. ee. nin H45. red in H45. ee. elim H46. sy. am. ee. am. am. Qed. Lemma ordinal_sum_lattice3: forall r f g i x y t, ordinal_sum_axioms1 r f g -> lattice (ordinal_sum r f g) -> inc i (domain f) -> gle (V i g) x t -> gle (V i g) y t -> has_supremum (V i g) (doubleton x y). Proof. ir. set (F:= ordinal_sum r f g) in H0. assert (Ha:forall i, inc i (domain f) -> exists y, inc y (V i f) & inc (J y i) (substrate F)). ir. uf F. app ordinal_sum_pr1. nin H. assert (order F). uf F. fprops. cp H; red in H;ee. assert (Hb: substrate F = disjoint_union f). uf F. aw. cp (inc_arg1_substrate H2). cp (inc_arg2_substrate H2). cp (inc_arg1_substrate H3). assert (inc (J x i) (substrate F)). uf F. aw. app inc_disjoint_union. wrr H12. assert (inc (J y i) (substrate F)). uf F. aw. app inc_disjoint_union. wrr H12. assert (inc (J t i) (substrate F)). uf F. aw. app inc_disjoint_union. wrr H12. cp (lattice_sup_pr H0 H17 H16). ee. assert (gle F (J y i) (J t i)). uf F. aw. eee. assert (gle F (J x i) (J t i)). uf F. aw. eee. cp (H21 _ H22 H23). set (z:= sup F (J y i) (J x i)) in *. cp (related_ordinal_sum_order_id H6 H19). awi H25. cp (related_ordinal_sum_order_id H6 H24). awi H26. cp (order_antisymmetry H H25 H26). exists (P z). ap least_upper_bound_doubleton. app H11. ufi F H20. awi H20. ee. nin H29. red in H29. nin H29. contradiction. eee. am. ufi F H19. awi H19. ee. nin H29. red in H29. nin H29. contradiction. eee. am. ir. assert (inc (J t0 i) (substrate F)). uf F. aw. app inc_disjoint_union. wrr H12. cp (inc_arg2_substrate H28). am. assert (gle F (J y i) (J t0 i)). uf F. aw. eee. assert (gle F (J x i) (J t0 i)). uf F. aw. eee. cp (H21 _ H31 H32). ufi F H33. awi H33. ee. nin H35. red in H35. ee. elim H36. sy; am. ee. rwi H35 H36. am. am. Qed. Lemma ordinal_sum_lattice4: forall r f g i x y t, ordinal_sum_axioms1 r f g -> lattice (ordinal_sum r f g) -> inc i (domain f) -> gle (V i g) t x -> gle (V i g) t y -> has_infimum (V i g) (doubleton x y). Proof. ir. set (F:= ordinal_sum r f g) in H0. assert (Ha:forall i, inc i (domain f) -> exists y, inc y (V i f) & inc (J y i) (substrate F)). ir. uf F. app ordinal_sum_pr1. nin H. assert (order F). uf F. fprops. cp H; red in H;ee. assert (Hb: substrate F = disjoint_union f). uf F. aw. cp (inc_arg2_substrate H2). cp (inc_arg1_substrate H2). cp (inc_arg2_substrate H3). assert (inc (J x i) (substrate F)). uf F. aw. app inc_disjoint_union. wrr H12. assert (inc (J y i) (substrate F)). uf F. aw. app inc_disjoint_union. wrr H12. assert (inc (J t i) (substrate F)). uf F. aw. app inc_disjoint_union. wrr H12. cp (lattice_inf_pr H0 H17 H16). ee. assert (gle F (J t i) (J y i)). uf F. aw. eee. assert (gle F (J t i) (J x i)). uf F. aw. eee. cp (H21 _ H22 H23). set (z:= inf F (J y i) (J x i)) in *. cp (related_ordinal_sum_order_id H6 H19). awi H25. cp (related_ordinal_sum_order_id H6 H24). awi H26. cp (order_antisymmetry H H25 H26). exists (P z). ap greatest_lower_bound_doubleton. app H11. ufi F H20. awi H20. ee. nin H29. red in H29. nin H29. contradiction. wr H27. eee. am. ufi F H19. awi H19. ee. nin H29. red in H29. nin H29. contradiction. wr H27. eee. am. ir. assert (inc (J t0 i) (substrate F)). uf F. aw. app inc_disjoint_union. wrr H12. cp (inc_arg1_substrate H28). am. assert (gle F (J t0 i) (J y i)). uf F. aw. eee. assert (gle F (J t0 i) (J x i)). uf F. aw. eee. cp (H21 _ H31 H32). ufi F H33. awi H33. ee. nin H35. red in H35. ee. elim H36. sy; am. ee. rwi H35 H36. wrr H27. am. Qed. Lemma ordinal_sum_lattice5: forall r f g i x y, ordinal_sum_axioms1 r f g -> lattice (ordinal_sum r f g) -> inc i (domain f) -> inc x (V i f) -> inc y (V i f) -> (forall t, inc t (V i f) -> ~ (gle (V i g) x t & gle (V i g) y t)) -> exists j, inc j (domain f) & least_element (induced_order r (Zo (domain f) (fun k=> glt r i k))) j & exists z, least_element (V j g) z. Proof. ir. set (F:= ordinal_sum r f g) in H0. assert (Ha:forall i, inc i (domain f) -> exists y, inc y (V i f) & inc (J y i) (substrate F)). ir. uf F. app ordinal_sum_pr1. nin H. assert (order F). uf F. fprops. cp H; red in H;ee. assert (Hb: substrate F = disjoint_union f). uf F. aw. assert (inc (J x i) (substrate F)). uf F. aw. app inc_disjoint_union. assert (inc (J y i) (substrate F)). uf F. aw. app inc_disjoint_union. cp (lattice_sup_pr H0 H14 H15). ee. set (Z:=sup F (J x i) (J y i)) in *. ufi F H16. ufi F H17. awi H16; awi H17. ee. assert (inc (Q Z) (domain f)). cp (du_index_pr H21). eee. assert (glt r i (Q Z)). nin H22. am. nin H20. am. ee. cp (du_index_pr H21). ee. wri H20 H27. elim (H4 _ H27). eee. assert (sub (Zo (domain f) (fun k => glt r i k)) (substrate r)). red. ir. Ztac. red in H27. ee. order_tac. exists (Q Z). ee. am. red. ee. aw. Ztac. ir. awi H26. Ztac. clear H26. aw. nin (Ha _ H27). ee. assert (gle F (J x i) (J x1 x0)). uf F. aw. eee. assert (gle F (J y i) (J x1 x0)). uf F. aw. eee. cp (H18 _ H30 H31). cp (related_ordinal_sum_order_id H7 H32). awi H33. am. Ztac. Ztac. am. am. exists (P Z). red. ee. rww H13. cp (du_index_pr H19). ee. am. ir. rwi H13 H26. assert (inc (J x0 (Q Z)) (disjoint_union f)). app inc_disjoint_union. assert (gle F (J x i) (J x0 (Q Z))). uf F. aw. eee. assert (gle F (J y i) (J x0 (Q Z))). uf F. aw. eee. cp (H18 _ H28 H29). ufi F H30. awi H30. ee. nin H32. red in H32. ee. elim H33. tv. ee. am. am. am. am. am. am. Qed. Lemma ordinal_sum_lattice6: forall r f g i x y, ordinal_sum_axioms1 r f g -> lattice (ordinal_sum r f g) -> inc i (domain f) -> inc x (V i f) -> inc y (V i f) -> (forall t, inc t (V i f) -> ~ (gle (V i g) t x & gle (V i g) t y)) -> exists j, inc j (domain f) & greatest_element (induced_order r (Zo (domain f) (fun k=> glt r k i))) j & exists z, greatest_element (V j g) z. Proof. ir. set (F:= ordinal_sum r f g) in H0. assert (Ha:forall i, inc i (domain f) -> exists y, inc y (V i f) & inc (J y i) (substrate F)). ir. uf F. app ordinal_sum_pr1. nin H. assert (order F). uf F. fprops. cp H; red in H;ee. assert (Hb: substrate F = disjoint_union f). uf F. aw. assert (inc (J x i) (substrate F)). uf F. aw. app inc_disjoint_union. assert (inc (J y i) (substrate F)). uf F. aw. app inc_disjoint_union. cp (lattice_inf_pr H0 H14 H15). ee. set (Z:=inf F (J x i) (J y i)) in *. ufi F H16. ufi F H17. awi H16; awi H17. ee. assert (inc (Q Z) (domain f)). cp (du_index_pr H16). ee. am. assert (glt r (Q Z) i). nin H22. am. nin H20. am. ee. cp (du_index_pr H16). ee. rwi H20 H27. elim (H4 _ H27). wr H22. eee. assert (sub (Zo (domain f) (fun k => glt r k i)) (substrate r)). red. ir. Ztac. red in H27. ee. order_tac. exists (Q Z). ee. am. red. ee. aw. Ztac. ir. awi H26. Ztac. clear H26. aw. nin (Ha _ H27). ee. assert (gle F (J x1 x0) (J x i)). uf F. aw. eee. assert (gle F (J x1 x0) (J y i)). uf F. aw. eee. cp (H18 _ H30 H31). cp (related_ordinal_sum_order_id H7 H32). awi H33. am. Ztac. Ztac. am. am. exists (P Z). red. ee. rww H13. cp (du_index_pr H17). ee. am. ir. rwi H13 H26. assert (inc (J x0 (Q Z)) (disjoint_union f)). app inc_disjoint_union. assert (gle F (J x0 (Q Z))(J x i)). uf F. aw. eee. assert (gle F (J x0 (Q Z)) (J y i)). uf F. aw. eee. cp (H18 _ H28 H29). ufi F H30. awi H30. ee. nin H32. red in H32. ee. elim H33. tv. ee. am. am. am. am. am. am. Qed. Lemma ordinal_sum_lattice: forall r f g, ordinal_sum_axioms1 r f g -> lattice (ordinal_sum r f g) = ((lattice r) & (forall i j, inc i (domain f) -> inc j (domain f) -> (gle r i j \/ gle r j i \/ (exists u, exists v, least_element (V (sup r i j) g) u & greatest_element (V (inf r i j) g) v))) & (forall i x y t, inc i (domain f) -> gle (V i g) x t -> gle (V i g) y t -> has_supremum (V i g) (doubleton x y)) & (forall i x y t, inc i (domain f) -> gle (V i g) t x -> gle (V i g) t y -> has_infimum (V i g) (doubleton x y)) & (forall i x y, inc i (domain f) -> inc x (V i f) -> inc y (V i f) -> (forall t, inc t (V i f) -> ~ (gle (V i g) x t & gle (V i g) y t)) -> exists j, inc j (domain f) & least_element (induced_order r (Zo (domain f) (fun k=> glt r i k))) j & exists z, least_element (V j g) z) & (forall i x y, inc i (domain f) -> inc x (V i f) -> inc y (V i f) -> (forall t, inc t (V i f) -> ~ (gle (V i g) t x & gle (V i g) t y)) -> exists j, inc j (domain f) & greatest_element (induced_order r (Zo (domain f) (fun k=> glt r k i))) j & exists z, greatest_element (V j g) z)). Proof. ir. assert (Ht:=H). nin H. set (F:= ordinal_sum r f g). assert (order F). uf F. fprops. cp H; red in H;ee. assert (Hb: substrate F = disjoint_union f). uf F. aw. assert (Ha:forall i, inc i (domain f) -> exists y, inc y (V i f) & inc (J y i) (substrate F)). ir. nin (H0 _ H9). exists y. ee. am. cp (inc_disjoint_union H9 H10). ue. app iff_eq. ir. cp (ordinal_sum_lattice1 Ht H9). ee. am. ir. app (ordinal_sum_lattice2 Ht). ir. app (ordinal_sum_lattice3 Ht H9 H11 H12 H13). ir. app (ordinal_sum_lattice4 Ht H9 H11 H12 H13). ir. app (ordinal_sum_lattice5 Ht H9 H11 H12 H13 H14). ir. app (ordinal_sum_lattice6 Ht H9 H11 H12 H13 H14). ir. ee. red. ee. ir. am. ir. ee. red. rwi Hb H15. rwi Hb H16. cp (du_index_pr H15). cp (du_index_pr H16). ee. wri H3 H17. wri H3 H18. cp (lattice_sup_pr H9 H17 H18). ee. set (a:= sup r (Q x) (Q y)) in *. assert (forall z, gle F x z -> gle F y z -> gle r a (Q z)). ir. ufi F H26. cp (related_ordinal_sum_order_id H2 H26). ufi F H26. cp (related_ordinal_sum_order_id H2 H27). cp (H25 _ H28 H29). am. rwi H3 H17. rwi H3 H18. nin (equal_or_not (Q x) (Q y)). wri H27 H19. nin (p_or_not_p (exists t, gle (V (Q x) g) (P x) t &gle (V (Q x) g) (P y) t)). nin H28. nin H28. cp (H11 _ _ _ _ H17 H28 H29). wri H8 H21;wri H8 H19. cp (sup_pr (H7 _ H17) H21 H19 H30). ee. set (z:= sup (V (Q x) g) (P x) (P y)) in *. assert (inc (J z (Q x)) (substrate F)). rw Hb. app inc_disjoint_union. wr H8. app (inc_arg2_substrate H31). am. exists (J z (Q x)). ap least_upper_bound_doubleton. am. uf F. aw. eee. uf F. aw. eee. ir. ufi F H35; ufi F H36; awi H35; awi H36. ee. uf F. aw. eee. nin H40. au. nin H38. left. ue. right. ee. am. app H33. ue. am. am. am. am. am. am. assert (forall t, inc t (V (Q x) f) -> ~ (gle (V (Q x) g) (P x) t & gle (V (Q x) g) (P y) t)). ir. red. ir. elim H28. exists t. am. nin (H13 _ _ _ H17 H21 H19 H29). ee. nin H32. nin H32. ee. assert (inc (J x1 x0) (disjoint_union f)). app inc_disjoint_union. wrr H8. assert(Hc:sub (Zo (domain f) (fun k => glt r (Q x) k)) (substrate r)). rw H3. app Z_sub. red in H31. ee. awi H35. awi H31. Ztac. clear H31. exists (J x1 x0). ap least_upper_bound_doubleton. am. uf F. aw. eee. uf F. aw. eee. ir. ufi F H31; ufi F H38; awi H31; awi H38. ee. uf F. aw. eee. nin (equal_or_not x0 (Q t)). right. ee. am. app H33. cp (du_index_pr H39). ee. rw H43. rww H8. left. nin (p_or_not_p (glt r (Q x) (Q t))). ir. assert (inc (Q t) (Zo (domain f) (fun k => glt r (Q x) k))). Ztac. wr H3. order_tac. cp (H35 _ H45). awi H46. red. ee. am. am. cp (inc_arg1_substrate H46). awi H47. am. am. am. am. ir. nin H42. elim H44. am. nin H40. elim H44. rww H27. ee. cp (du_index_pr H41). ee. wri H42 H48. elim (H29 _ H48). ee. am. rww H27. am. am. am. am. am. am. am. cp (H10 _ _ H17 H18). nin (equal_or_not (Q y) a). assert (gle F x y). uf F. aw. eee. left. red. ee. rww H29. am. exists y. app sup_comparable. nin (equal_or_not (Q x) a). assert (gle F y x). uf F. aw. eee. left. red. ee. rww H30. intuition. exists x. rw doubleton_symm. app sup_comparable. nin H28. assert (gle r (Q y) (Q y)). order_tac. ue. cp (order_antisymmetry H H24 (H25 _ H28 H31)). elim H29; am. nin H28. assert (gle r (Q x) (Q x)). order_tac. ue. cp (order_antisymmetry H H23 (H25 _ H31 H28)). elim H30. am. nin H28; nin H28; nin H28. fold a in H28. red in H28. ee. assert (inc (J x0 a) (disjoint_union f)). app inc_disjoint_union. wr H3. order_tac. wrr H8. wr H3. order_tac. exists (J x0 a). ap least_upper_bound_doubleton. am. uf F. aw. eee. left. red. au. uf F. aw. eee. left. red. au. ir. cp (H26 _ H34 H35). ufi F H34. awi H34. ufi F H35. awi H35. uf F. aw. eee. nin (equal_or_not a (Q t)). right. ee. am. app H32. rw H8. rw H41. cp (du_index_pr H37). ee; am. wr H3. order_tac. left. order_tac. am. am. ee. red. rwi Hb H15. rwi Hb H16. cp (du_index_pr H15). cp (du_index_pr H16). ee. wri H3 H17. wri H3 H18. cp (lattice_inf_pr H9 H17 H18). ee. set (a:= inf r (Q x) (Q y)) in *. assert (forall z, gle F z x-> gle F z y -> gle r (Q z) a). ir. ufi F H26. cp (related_ordinal_sum_order_id H2 H26). ufi F H26. cp (related_ordinal_sum_order_id H2 H27). cp (H25 _ H28 H29). am. rwi H3 H17. rwi H3 H18. nin (equal_or_not (Q x) (Q y)). wri H27 H19. nin (p_or_not_p (exists t, gle (V (Q x) g) t(P x) &gle (V (Q x) g) t (P y))). nin H28. nin H28. cp (H12 _ _ _ _ H17 H28 H29). wri H8 H21;wri H8 H19. cp (inf_pr (H7 _ H17) H21 H19 H30). ee. set (z:= inf (V (Q x) g) (P x) (P y)) in *. assert (inc (J z (Q x)) (substrate F)). rw Hb. app inc_disjoint_union. wr H8. order_tac. ue. exists (J z (Q x)). ap greatest_lower_bound_doubleton. am. uf F. aw. eee. uf F. aw. eee. ir. ufi F H35; ufi F H36; awi H35; awi H36. ee. uf F. aw. eee. nin H40. left. am. nin H38. left. rww H27. right. ee. am. rw H40. app H33. rww H27. wr H38. am. wr H40. am. am. am. am. am. am. am. assert (forall t, inc t (V (Q x) f) -> ~ (gle (V (Q x) g) t (P x) & gle (V (Q x) g) t (P y))). ir. red. ir. elim H28. exists t. am. nin (H14 _ _ _ H17 H21 H19 H29). ee. nin H32. nin H32. ee. assert (inc (J x1 x0) (disjoint_union f)). app inc_disjoint_union. wrr H8. assert(Hc:sub (Zo (domain f) (fun k => glt r k (Q x))) (substrate r)). rw H3. app Z_sub. red in H31. ee. awi H35. awi H31. Ztac. clear H31. exists (J x1 x0). ap greatest_lower_bound_doubleton. am. uf F. aw. eee. uf F. aw. eee. ir. ufi F H31; ufi F H38; awi H31; awi H38. ee. uf F. aw. eee. nin (equal_or_not x0 (Q t)). right. ee. sy;am. wr H43. app H33. cp (du_index_pr H38). ee. rw H43. rww H8. left. nin (p_or_not_p (glt r (Q t) (Q x))). ir. assert (inc (Q t) (Zo (domain f) (fun k => glt r k (Q x)))). Ztac. wr H3. order_tac. cp (H35 _ H45). awi H46. red. ee. am. intuition. am. cp (inc_arg2_substrate H46). awi H47. am. am. am. ir. nin H42. elim H44. am. nin H40. elim H44. rww H27. ee. cp (du_index_pr H31). ee. rwi H42 H48. elim (H29 _ H48). ee. wrr H42. wrr H42. am. am. am. am. am. am. am. cp (H10 _ _ H17 H18). nin (equal_or_not (Q y) a). assert (gle F y x). uf F. aw. eee. left. red. ee. rww H29. intuition. exists y. rw doubleton_symm. app inf_comparable. nin (equal_or_not (Q x) a). assert (gle F x y). uf F. aw. eee. left. red. ee. rww H30. am. exists x. app inf_comparable. nin H28. assert (gle r (Q x) (Q x)). order_tac. rww H3. cp (order_antisymmetry H H23 (H25 _ H31 H28)). elim H30; sy;am. nin H28. assert (gle r (Q y) (Q y)). order_tac. rww H3. cp (order_antisymmetry H H24 (H25 _ H28 H31)). elim H29. sy;am. nin H28; nin H28; nin H28. fold a in H31. red in H31. ee. assert (inc (J x1 a) (disjoint_union f)). app inc_disjoint_union. wr H3. order_tac. wrr H8. wr H3. order_tac. exists (J x1 a). ap greatest_lower_bound_doubleton. am. uf F. aw. eee. left. red. ee. am. intuition. uf F. aw. ee. am. am. left. red. au. ir. cp (H26 _ H34 H35). ufi F H34. awi H34. ufi F H35. awi H35. uf F. aw. eee. nin (equal_or_not a (Q t)). right. ee. sy;am. wr H41. app H32. rw H8. rw H41. cp (du_index_pr H35). ee; am. wr H3. order_tac. left. order_tac. symmetry in H42. contradiction. am. am. Qed. Require Export sete1. Definition not_comp_rel r := fun x y => inc x (substrate r) & inc y (substrate r) & x = y \/ ~ ((gle r x y) \/ (gle r y x)). Definition ncr_equiv r := Exercice1.Sgraph (not_comp_rel r) (substrate r). Definition ncr_component r := Exercice1.connected_comp (not_comp_rel r) (substrate r). Lemma ncr_properties: forall r, order r -> (is_equivalence (ncr_equiv r) & substrate (ncr_equiv r) = substrate r & (forall x, inc x (substrate r) -> class (ncr_equiv r) x = ncr_component r x)& (forall x y, not_comp_rel r x y -> related (ncr_equiv r) x y)). Proof. ir. uf ncr_equiv. uf ncr_component. assert (reflexive_r (not_comp_rel r) (substrate r)). red. ir. uf not_comp_rel. app iff_eq. ir. eee. ir. ee. am. assert (symmetric_r (not_comp_rel r)). red. uf not_comp_rel. ir. eee. nin H3. left. sy; am. right. dneg. intuition. uf not_comp_rel. assert (forall x y, not_comp_rel r x y -> inc x (substrate r)). ir. red in H2. ee; am. ee. app Exercice1.equivalence_Sgraph. app Exercice1.substrate_Sgraph. ir. app Exercice1.connected_comp_class. ir. red. uf Exercice1.Sgraph. uf graph_on. ee. Ztac. fprops. red. aw. exists (Exercice1.chain_pair x y). eee. simpl. intuition. Qed. Lemma Exercice1_4a1: forall r x y, order r -> inc x (substrate r) -> inc y (substrate r) -> gle r x y \/ gle r y x \/ class (ncr_equiv r) x = class (ncr_equiv r) y. Proof. ir. nin (p_or_not_p ((gle r x y) \/ (gle r y x))). nin H2. au. au. right. right. assert (not_comp_rel r x y). red. eee. cp (ncr_properties H). ee. wrr related_class_eq. app H7. app reflexivity_e. rww H5. Qed. Lemma Exercice1_4a2: forall r y, order r -> inc y (substrate r) -> forall a b, union2 a b = class (ncr_equiv r) y -> (forall u v, inc u a -> inc v b -> gle r u v \/ gle r v u) -> a = emptyset \/ b = emptyset \/ nonempty (intersection2 a b). Proof. ir. nin (emptyset_dichot a). left. am. right. nin (emptyset_dichot b). au. right. nin H3. nin H4. cp (ncr_properties H). ee. assert (Ha:is_class (ncr_equiv r) (union2 a b)). rw H1. app is_class_class. ue. assert (related (ncr_equiv r) y0 y1). rw in_class_related. exists (union2 a b). ee. am. app union2_first. app union2_second. am. red in H9. ufi ncr_equiv H9. ufi Exercice1.Sgraph H9. ufi graph_on H9. Ztac. awi H11. red in H11. nin H11. ee. assert (inc (Exercice1.chain_head x) a). ue. clear H12. assert (inc (Exercice1.chain_tail x) b). ue. clear H13. assert (exists u, exists v, inc u a & inc v b & (not_comp_rel r) u v). nin x. simpl in H11. simpl in H12. simpl in H14. exists P. exists P0. intuition. simpl in H11. simpl in H14. simpl in H12. ee. nin (inc_or_not (Exercice1.chain_head x) a). ir. app IHx. ir. set (v:= (Exercice1.chain_head x)) in *. cp (H8 _ _ H11). rwi is_class_rw Ha. ee. wri H19 H16. assert (inc v b). nin (union2_or H16). contradiction. am. exists P. exists v. intuition. inter2tac. am. nin H13. nin H13. ee. red in H16. ee. nin H18. exists x0. app intersection2_inc. rww H18. elim H18. app H2. Qed. Lemma Exercice1_4a3: forall r x y y', order r -> inc x (substrate r) -> related (ncr_equiv r) y y' -> gle r x y -> related (ncr_equiv r) x y \/ gle r x y'. Proof. ir. cp (ncr_properties H). ee. set (C:= class (ncr_equiv r) y). set (A:=Zo C (fun z => gle r x z)). set (B:=Zo C (fun z => gle r z x)). assert (forall u v, inc u A -> inc v B -> gle r u v \/ gle r v u). ir. ufi A H7. Ztac. clear H7. ufi B H8. Ztac. right. order_tac. nin (p_or_not_p (related (ncr_equiv r) x y)). left. am. right. assert(inc y (substrate r)). app (inc_arg2_substrate H2). assert (union2 A B = class (ncr_equiv r) y). set_extens. nin (union2_or H10). ufi A H11. Ztac. am. ufi B H11. Ztac. am. assert (inc x0 (substrate r)). wr H4. app (sub_class_substrate H3 H10). nin (Exercice1_4a1 H H0 H11). app union2_first. uf A. Ztac. nin H12. app union2_second. uf B. Ztac. elim H8. apply transitivity_e with x0. am. rww related_class_eq. app reflexivity_e. rww H4. app symmetricity. wr inc_class. am. am. cp (Exercice1_4a2 H H9 H10 H7). nin H11. elim (emptyset_pr (x:=y)). wr H11. uf A. Ztac. uf C. bw. app reflexivity_e. rww H4. nin H11. assert (inc y' C). uf C. bw. ufi C H12. wri H10 H12. nin (union2_or H12). ufi A H13. Ztac. am. rwi H11 H13. elim (emptyset_pr H13). nin H11. cp (intersection2_first H11). ufi A H12. Ztac. clear H12. cp (intersection2_second H11). ufi B H12. Ztac. elim H8. rw (order_antisymmetry H H14 H16). app symmetricity. wr inc_class. am. am. Qed. Lemma Exercice1_4b1: forall r x y x' y', order r -> related (ncr_equiv r) x x' -> related (ncr_equiv r) y y' -> class (ncr_equiv r) x <> class (ncr_equiv r) y -> gle r x y -> gle r x' y'. Proof. ir. cp (ncr_properties H). ee. assert (inc x' (substrate r)). wr H5. app (inc_arg2_substrate H0). assert (inc y' (substrate r)). wr H5. app (inc_arg2_substrate H1). assert (inc x (substrate r)). wr H5. app (inc_arg1_substrate H0). assert (inc y (substrate r)). wr H5. app (inc_arg1_substrate H1). nin (Exercice1_4a1 H H8 H9). am. nin H12. assert (Ha:related (ncr_equiv r) x x). app reflexivity_e. rww H5. nin (Exercice1_4a3 H H10 H1 H3). elim H2. wrr related_class_eq. assert (related (ncr_equiv r) x' x). app symmetricity. nin (Exercice1_4a3 H H9 H14 H12). elim H2. wrr related_class_eq. apply transitivity_e with x'. am. am. app symmetricity. apply transitivity_e with y'. am . am. am. wri (order_antisymmetry H H13 H15) H1. elim H2. wrr related_class_eq. app symmetricity. elim H2. wrr related_class_eq. apply transitivity_e with x'. am. am. apply transitivity_e with y'. am. rww related_class_eq. app reflexivity_e. rww H5. app symmetricity. app reflexivity_e. rww H5. Qed. Lemma Exercice1_4b: forall r, order r -> total_order(quotient_order r (ncr_equiv r)). Proof. ir. cp (ncr_properties H). ee. assert (substrate (quotient_order r (ncr_equiv r)) = quotient (ncr_equiv r)). rww substrate_quotient_order. app order_is_preorder. assert (order (quotient_order r (ncr_equiv r))). app Exercise1_2d. red. ir. nin (p_or_not_p (related (ncr_equiv r) x y)). am. assert (related (ncr_equiv r) y y). app reflexivity_e. rww H1. order_tac. assert (class (ncr_equiv r) x <> class (ncr_equiv r) y). dneg. app symmetricity. rww related_class_eq. sy; am. cp (Exercice1_4b1 H H7 H9 H10 H5). rww (order_antisymmetry H H6 H11). red. ee. am. rw H4. ir. uf gge. nin (equal_or_not x y). left. rw H8. order_tac. ue. rw quotient_order_pr. rw quotient_order_pr. uf quotient_order_r. cp (inc_rep_itself H0 H6). cp (inc_rep_itself H0 H7). assert (inc (rep x) (substrate r)). wr H1. app inc_rep_substrate. assert (inc (rep y) (substrate r)). wr H1. app inc_rep_substrate. assert (Ha:class (ncr_equiv r) (rep x) <> class (ncr_equiv r) (rep y)). red. ir. rwi (inc_quotient x H0) H6. rwi (inc_quotient y H0) H7. elim H8. red in H6; red in H7. ee. rw H17; rw H15. am. cp (Exercice1_4a1 H H11 H12). nin H13. left. eee. ir. exists (rep y). ee. am. assert (related (ncr_equiv r) (rep x) x0). app related_rep_in_class. assert (related (ncr_equiv r) (rep y)(rep y)). app related_rep_in_class. ap (Exercice1_4b1 H H15 H16 Ha H13). nin H13. right. eee. ir. exists (rep x). ee. am. assert (related (ncr_equiv r) (rep y) x0). app related_rep_in_class. assert (related (ncr_equiv r) (rep x)(rep x)). app related_rep_in_class. assert (class (ncr_equiv r) (rep y) <> class (ncr_equiv r) (rep x)). intuition. ap (Exercice1_4b1 H H15 H16 H17 H13). contradiction. Qed. (* Exercise 5 *) Definition free_subset r X := forall x y, inc x X -> inc y X -> gle r x y -> x = y. Definition set_of_free_subsets r:= Zo (powerset (substrate r)) (fun X=> free_subset r X). Definition free_subset_compare r X Y:= inc X (set_of_free_subsets r) & inc Y (set_of_free_subsets r) & forall x, inc x X -> exists y, inc y Y & gle r x y. Definition free_subset_order r:= graph_on (free_subset_compare r) (set_of_free_subsets r). Lemma Exercise1_5a: forall r, order r -> order_r (free_subset_compare r). Proof. ir. red. uf free_subset_compare. split. red. ir. eee. ir. nin (H5 _ H6). nin H7. nin (H3 _ H7). nin H9. ex_tac. order_tac. split. red. ir. ee. set_extens; [nin (H5 _ H6) | nin (H3 _ H6)]; nin H7; [nin (H3 _ H7) | nin (H5 _ H7) ]; nin H9; cp (order_transitivity H H8 H10); [ ufi set_of_free_subsets H0 |ufi set_of_free_subsets H4]; Ztac; wri (H13 _ _ H6 H9 H11) H10; rww (order_antisymmetry H H8 H10). red. ir. eee. ir. ex_tac. order_tac. ufi set_of_free_subsets H0. Ztac. rwi powerset_inc_rw H4. app H4. ir. ex_tac. order_tac. ufi set_of_free_subsets H1. Ztac. rwi powerset_inc_rw H4. app H4. Qed. Lemma fs_order_pr: forall r x y, related (free_subset_order r) x y = free_subset_compare r x y. Proof. ir. uf free_subset_order. uf graph_on. uf related. ap iff_eq. ir. Ztac. awi H1. am. ir. Ztac. red in H. ee. fprops. aw. Qed. Lemma fs_is_order: forall r, order r -> order (free_subset_order r). Proof. ir. uf free_subset_order. app order_from_rel. app Exercise1_5a. Qed. Lemma substrate_fs_order: forall r, order r -> substrate (free_subset_order r) = set_of_free_subsets r. Proof. ir. cp (fs_is_order H). set_extens. rwi order_reflexivity H1. rwi fs_order_pr H1. red in H1. nin H1. am. am. rw order_reflexivity. rw fs_order_pr. red. split. am. split. am. ir. exists x0. split. am. wr order_reflexivity. ufi set_of_free_subsets H1. Ztac. rwi powerset_inc_rw H3. app H3. am. am. Qed. Lemma Exercise1_5b: forall r x, order r -> inc x (substrate r) -> inc (singleton x) (set_of_free_subsets r). Proof. ir. uf set_of_free_subsets. Ztac. app powerset_inc. red. ir. rw (singleton_eq H1). am. red. ir. rw (singleton_eq H1). rww (singleton_eq H2). Qed. Lemma Exercise1_5c: forall r x y, order r -> inc x (substrate r) ->inc y (substrate r) -> gle r x y = gle (free_subset_order r) (singleton x) (singleton y). Proof. ir. rw fs_order_pr. app iff_eq. ir. red. ee. app Exercise1_5b. app Exercise1_5b. ir. rw (singleton_eq H3). exists y. split. fprops. am. ir. red in H2. ee. assert (inc x (singleton x)). fprops. cp (H4 _ H5). nin H6. nin H6. wrr (singleton_eq H6). Qed. Lemma Exercise1_5d: forall r, order r -> order_morphism (BL singleton (substrate r) (set_of_free_subsets r)) r (free_subset_order r). Proof. ir. assert (transf_axioms singleton (substrate r) (set_of_free_subsets r)). red. ir. app Exercise1_5b. assert (is_function (BL singleton (substrate r) (set_of_free_subsets r))). app bl_function. red. aw. ee. am. app fs_is_order. app injective_bl_function. ir. app singleton_inj. tv. rww substrate_fs_order. ir. rww W_bl_function. rww W_bl_function. app Exercise1_5c. Qed. Lemma Exercise1_5e: forall r X, total_order r -> inc X (set_of_free_subsets r)-> small_set X. Proof. ir. red. ir. red in H. ee. ufi set_of_free_subsets H0. Ztac. rwi powerset_inc_rw H4. red in H5. nin (H3 _ _ (H4 _ H1) (H4 _ H2)). app H5. red in H6. sy. app H5. Qed. Lemma Exercise1_5f: forall r X Y, order r -> inc X (set_of_free_subsets r) -> inc Y (set_of_free_subsets r) -> sub X Y -> gle (free_subset_order r) X Y. Proof. ir. rw fs_order_pr. red. ee. am. am. ir. exists x. split. app H2. wr order_reflexivity. ufi set_of_free_subsets H0. Ztac. rwi powerset_inc_rw H4. app H4. am. Qed. Lemma Exercise1_5g: forall r, total_order r -> total_order (free_subset_order r). Proof. ir. cp H. nin H. red. split. app fs_is_order. rww substrate_fs_order. assert (inc emptyset (set_of_free_subsets r)). uf set_of_free_subsets. Ztac. app powerset_inc. app sub_emptyset_any. red. ir. elim (emptyset_pr H2). ir. cp (Exercise1_5e H0 H3). cp (Exercise1_5e H0 H4). nin (emptyset_dichot x). rw H7. left. app Exercise1_5f. app sub_emptyset_any. nin (emptyset_dichot y). rw H8. right. red. app Exercise1_5f. app sub_emptyset_any. nin H7. nin H8. ufi set_of_free_subsets H3; Ztac. clear H3. rwi powerset_inc_rw H9. ufi set_of_free_subsets H4; Ztac. rwi powerset_inc_rw H3. cp (H9 _ H7). cp (H3 _ H8). assert (x = singleton y0). set_extens. red in H5. rw (H5 _ _ H7 H14). fprops. rw (singleton_eq H14). am. assert (y = singleton y1). set_extens. red in H6. rw (H6 _ _ H8 H15). fprops. rw (singleton_eq H15). am. rw H14; rw H15. nin (H1 _ _ H12 H13). left. wrr Exercise1_5c. right. red. wrr Exercise1_5c. Qed. Lemma Exercise1_5h: forall r, order r -> total_order (free_subset_order r) -> total_order r. Proof. ir. red. split. am. nin H0. ir. rwi substrate_fs_order H1. nin (H1 _ _ (Exercise1_5b H H2) (Exercise1_5b H H3)). left. rww Exercise1_5c. right. red. red in H4. rww Exercise1_5c. am. Qed. (* Exercise 6 *) Definition set_of_increasing_mappings r r' := Zo (set_of_functions (substrate r) (substrate r')) (fun z=> increasing_fun z r r'). Definition increasing_mappings_order r r' := induced_order (function_order (substrate r) (substrate r') r') (set_of_increasing_mappings r r'). Lemma set_of_increasing_mappings_pr: forall r r' f, order r -> order r' -> inc f (set_of_increasing_mappings r r') = (is_function f & source f = (substrate r) & target f = substrate r' & increasing_fun f r r'). Proof. ir. uf set_of_increasing_mappings. app iff_eq. ir. Ztac. awi H2. eee. ir. ee. Ztac. aw. eee. Qed. Hint Rewrite set_of_increasing_mappings_pr : aw. Lemma imo_order: forall r r', order r -> order r' -> order (increasing_mappings_order r r'). Proof. ir. uf increasing_mappings_order. app order_induced_order. rw substrate_function_order. uf set_of_increasing_mappings. app Z_sub. am. tv. app order_function_order. Qed. Lemma imo_substrate: forall r r', order r -> order r' -> substrate (increasing_mappings_order r r') = set_of_increasing_mappings r r'. Proof. ir. uf increasing_mappings_order. rw substrate_induced_order. tv. app order_function_order. uf set_of_increasing_mappings. rw substrate_function_order. app Z_sub. am. tv. Qed. Lemma imo_pr: forall r r' f g, order r -> order r' -> gle (increasing_mappings_order r r') f g = (inc f (set_of_increasing_mappings r r') & inc g (set_of_increasing_mappings r r') & function_order_r (substrate r) (substrate r') r' f g). Proof. ir. app iff_eq. ir. cp (imo_order H H0). cp (inc_arg1_substrate H1). cp (inc_arg2_substrate H1). rwi imo_substrate H3. rwi imo_substrate H4. ufi increasing_mappings_order H1. awi H1. ufi function_order H1. ufi graph_on H1. eee. red in H1. red in H1. Ztac. clear H1. awi H5. ee. awi H6. am. am. am. am. am. am. am. ir. ee. uf increasing_mappings_order. aw. ufi set_of_increasing_mappings H1. Ztac. clear H1. ufi set_of_increasing_mappings H2. Ztac. clear H2. uf function_order. red. red. uf graph_on. Ztac. aw. awi H4. awi H1. eee. aw. Qed. Definition first_projection f a b:= BL (fun z=> P (W z f)) a b. Definition secnd_projection f a c:= BL (fun z=> Q (W z f)) a c. Definition two_projections a b c := BL (fun z => (J (first_projection z a b) (secnd_projection z a c))) (set_of_functions a (product b c)) (product (set_of_functions a b) (set_of_functions a c)). Lemma Exercice1_6a: forall f a b c, is_function f -> source f =a -> target f = product b c -> (transf_axioms (fun z=> P (W z f)) a b & transf_axioms (fun z=> Q (W z f)) a c & is_function (first_projection f a b) & is_function (secnd_projection f a c) & (forall x, inc x a -> W x (first_projection f a b) = P (W x f)) & (forall x, inc x a -> W x (secnd_projection f a c) = Q (W x f))). Proof. ir. assert (transf_axioms (fun z=> P (W z f)) a b). red. ir. wri H0 H2. cp (inc_W_target H H2). rwi H1 H3. awi H3; eee. assert (transf_axioms (fun z=> Q (W z f)) a c). red. ir. wri H0 H3. cp (inc_W_target H H3). rwi H1 H4. awi H4; eee. uf first_projection. uf secnd_projection. ee. am. am. app bl_function. app bl_function. ir. aw. ir. aw. Qed. Lemma Exercice1_6b: forall a b c, transf_axioms (fun z => (J (first_projection z a b) (secnd_projection z a c))) (set_of_functions a (product b c)) (product (set_of_functions a b) (set_of_functions a c)). Proof. ir. red. ir. awi H. ee. cp (Exercice1_6a H H0 H1). ee. aw. split. fprops. eee; uf first_projection; uf secnd_projection; aw. Qed. Lemma Exercice1_6c: forall a b c, bijective (two_projections a b c). Proof. ir. cp (Exercice1_6b (a:=a) (b:=b)(c:=c)). uf two_projections. app bijective_bl_function. (* injectivity *) ir. awi H0. awi H1. cp (pr1_injective H2). cp (pr2_injective H2). ee. app funct_extensionality. ue. ue. rw H7. ir. cp (Exercice1_6a H0 H7 H8). cp (Exercice1_6a H1 H5 H6). ee. app pair_extensionality. assert (inc (W x u) (product b c)). wr H8. app inc_W_target. ue. awi H22; eee. assert (inc (W x v) (product b c)). wr H6. app inc_W_target. ue. awi H22; eee. wr H20. wr H15. rww H3. am. am. wr H16. wr H21. rww H4. am. am. (* surjectivity *) ir. awi H0. ee. set (f:= BL (fun z=> J (W z (P y)) (W z (Q y))) a (product b c)). assert (transf_axioms (fun z=> J (W z (P y)) (W z (Q y))) a (product b c)). red. ir. aw. ee. fprops. wr H6. app inc_W_target. ue. wr H4. app inc_W_target. ue. assert (is_function f). uf f. app bl_function. assert (source f = a). uf f. aw. assert (target f = product b c). uf f. aw. cp (Exercice1_6a H8 H9 H10). ee. exists f. ee. aw. eee. sy. app pair_extensionality. fprops. aw. app funct_extensionality. uf first_projection. aw. uf first_projection. aw. rw H5. ir. rww H15. uf f. aw. aw. app funct_extensionality. uf secnd_projection. aw. uf secnd_projection. aw. rw H3. ir. rww H16. uf f. aw. Qed. Hint Rewrite substrate_order_product2_order: bw. Lemma Exercice1_6d: forall f r r' r'', order r -> order r' -> order r'' -> increasing_fun f r (product2_order r' r'') -> (increasing_fun (first_projection f (substrate r) (substrate r')) r r' & increasing_fun (secnd_projection f (substrate r) (substrate r'')) r r''). Proof. ir. red in H2. nin H2; nin H3; nin H4; nin H5; nin H6. symmetry in H5; symmetry in H6. bwi H6. cp (Exercice1_6a H2 H5 H6). ee. red. eee. uf first_projection. aw. uf first_projection. aw. ir. rw H12. rw H12. cp (H7 _ _ H14). rwi product2_order_pr H15. red in H15. ee. am. order_tac. order_tac. red. eee. ir. uf secnd_projection. aw. uf secnd_projection. aw. ir. rw H13. rw H13. cp (H7 _ _ H14). rwi product2_order_pr H15. red in H15. ee. am. order_tac. order_tac. am. am. Qed. Definition two_projections_increasing r r' r'' := restriction2 (two_projections (substrate r) (substrate r')(substrate r'')) (set_of_increasing_mappings r (product2_order r' r'')) (product (set_of_increasing_mappings r r') (set_of_increasing_mappings r r'')). Lemma Exercice1_6e: forall r r' r'', order r -> order r' -> order r'' -> (restriction2_axioms (two_projections (substrate r) (substrate r')(substrate r'')) (set_of_increasing_mappings r (product2_order r' r'')) (product (set_of_increasing_mappings r r') (set_of_increasing_mappings r r''))). Proof. ir. cp (Exercice1_6c (substrate r) (substrate r') (substrate r'')). cp (Exercice1_6b (a:=substrate r) (b:=substrate r') (c:=substrate r'')). red. ee. fct_tac. uf two_projections. aw. uf set_of_increasing_mappings. bw. app Z_sub. uf two_projections. aw. app product_monotone. uf set_of_increasing_mappings. app Z_sub. uf set_of_increasing_mappings. app Z_sub. uf image_by_fun. red. ir. cp (bij_is_function H2). awi H4. nin H4. nin H4. wr (W_pr H5 H6). uf two_projections. awi H4. ee. assert (inc x0 (set_of_functions (substrate r) (product (substrate r') (substrate r'')))). aw. eee. bw. rww W_bl_function. nin (Exercice1_6d H H0 H1 H9). bwi H8. cp (Exercice1_6a H4 H7 H8). ee. aw. eee. uf first_projection; aw. uf first_projection; aw. uf secnd_projection; aw. uf secnd_projection; aw. am. am. am. app order_product2_order. Qed. Lemma Exercice1_6f: forall r r' r'', order r -> order r' -> order r'' -> bijective (two_projections_increasing r r' r''). Proof. ir. cp (Exercice1_6e H H0 H1). cp (Exercice1_6c (substrate r) (substrate r') (substrate r'')). nin H3. assert (is_function (two_projections_increasing r r' r'')). uf two_projections_increasing. app function_restriction2. red. split. uf two_projections_increasing. app injective_restriction2. uf two_projections_increasing. app surjective_pr6. ir. assert (inc y (target (two_projections (substrate r) (substrate r') (substrate r'')))). ufi restriction2 H6. awi H6. uf two_projections. aw. eee. am. am. am. am. nin (surjective_pr2 H4 H7). nin H8. ufi two_projections H8. awi H8. aw. ee. cut (inc x (set_of_increasing_mappings r (product2_order r' r''))). ir. exists x. split. uf restriction2. rww corresp_source. rww W_restriction2. ufi restriction2 H6. rwi corresp_target H6. ufi two_projections H9. rwii W_bl_function H9. aw. eee. bw. awi H6. ee. wri H9 H16. awi H16. wri H9 H19. awi H19. cp (Exercice1_6a H8 H10 H11). ee. red. ir. eee. app order_product2_order. bw. ir. rw product2_order_pr. red. wr H11. assert (inc x0 (substrate r)). order_tac. assert (inc y0 (substrate r)). order_tac. ee. app inc_W_target. ue. app inc_W_target. ue. wrr H24. wrr H24. red in H19. ee. app H33. wrr H25. wrr H25. red in H16. ee. app H33. am. am. am. am. app order_product2_order. app Exercice1_6b. aw. eee. Qed. Lemma Exercice1_6g: forall r r' r'', order r -> order r' -> order r'' -> order_isomorphism (two_projections_increasing r r' r'') (increasing_mappings_order r (product2_order r' r'')) (product2_order (increasing_mappings_order r r') (increasing_mappings_order r r'')). Proof. ir. assert (order (product2_order r' r'')). app order_product2_order. red. ee. app imo_order. app order_product2_order. app imo_order. app imo_order. app Exercice1_6f. uf two_projections_increasing. uf restriction2. aw. rww imo_substrate. bw. uf two_projections_increasing. uf restriction2. aw. rww imo_substrate. rww imo_substrate. app imo_order. app imo_order. uf two_projections_increasing. ir. ufi restriction2 H3. ufi restriction2 H4. rwi corresp_source H3. rwi corresp_source H4. assert (Ha:= H3). assert (Hb:= H4). awi H3. awi H4. ee. uf two_projections_increasing. cp (Exercice1_6b (a:=substrate r) (b:=substrate r') (c:=substrate r'')). cp (Exercice1_6e H H0 H1). rww W_restriction2. rww W_restriction2. clear H12. bwi H9. cp (Exercice1_6a H3 H8 H9). bwi H6. cp (Exercice1_6a H4 H5 H6). ee. uf two_projections. rww W_bl_function. rww W_bl_function. clear H11. rw product2_order_pr. uf product2_order_r. aw. do 3 rww imo_pr. set (f1 := first_projection x (substrate r) (substrate r')) in *. set (f2 := secnd_projection x (substrate r) (substrate r'')) in *. set (f3 := first_projection y (substrate r) (substrate r')) in *. set (f4 := secnd_projection y (substrate r) (substrate r'')) in *. repeat rww imo_substrate. cp (Exercice1_6d H H0 H1 H10). cp (Exercice1_6d H H0 H1 H7). ee. app iff_eq. (* direct *) ir. assert (source f1 = substrate r). uf f1. uf first_projection. aw. assert (source f2 = substrate r). uf f2. uf secnd_projection. aw. assert (source f3 = substrate r). uf f3. uf first_projection. aw. assert (source f4 = substrate r). uf f4. uf secnd_projection. aw. assert (target f1 = substrate r'). uf f1. uf first_projection. aw. assert (target f2 = substrate r''). uf f2. uf secnd_projection. aw. assert (target f3 = substrate r'). uf f3. uf first_projection. aw. assert (target f4 = substrate r''). uf f4. uf secnd_projection. aw. assert (inc f1 (set_of_increasing_mappings r r')). uf f1. aw. eee. assert (inc f3 (set_of_increasing_mappings r r')). uf f3. aw. eee. assert (inc f2 (set_of_increasing_mappings r r'')). uf f2. aw. eee. assert (inc f4 (set_of_increasing_mappings r r'')). uf f4. aw. eee. eee. red. eee. ir. rw H22. rw H17. red in H41. ee. cp (H48 _ H42). rwi product2_order_pr H49. red in H49. ee; am. am. am. red. eee. ir. rw H23. rw H18. red in H41. ee. cp (H48 _ H42). rwi product2_order_pr H49. red in H49. ee; am. am. am. (* converse *) ir. eee. red. eee. bw. bw. ir. red in H34. ee. cp (H45 _ H39). rwi H22 H46. rwi H17 H46. red in H32. ee. cp (H52 _ H39). rwi H23 H53. rwi H18 H53. rw product2_order_pr. red. ee. wr H9. app inc_W_target. ue. wr H6. app inc_W_target. ue. am. am. am. am. am. am. aw. eee. aw; eee. am. am. am. am. am. am. am. am. Qed. Lemma Exercice1_6h: forall f r r' r'', order r -> order r' -> nonempty (substrate r) -> nonempty (substrate r') -> increasing_fun f (product2_order r r') r'' -> ( (domain (source f)) = substrate r & (range (source f)) = substrate r'). Proof. ir. red in H3. ee. wr H6. bw. rw domain_product. tv. am. wr H6. bw. rw range_product. tv. am. Qed. Lemma Exercice1_6i: forall f x r r' r'', order r -> order r' -> nonempty (substrate r) -> nonempty (substrate r') -> increasing_fun f (product2_order r r') r'' -> inc x (substrate r) -> increasing_fun (second_partial_fun f x) r' r''. Proof. ir. nin (Exercice1_6h H H0 H1 H2 H3). red in H3. ee. assert (partial_fun_axioms f). red. ee. am. wr H9. bw. exists (substrate r). exists (substrate r'). tv. wri H5 H4. cp (function_spf H12 H4). red. eee. uf second_partial_fun. aw. sy; am. uf second_partial_fun. aw. ir. assert (inc (J x x0) (source f)). wr H9. bw. aw. eee. order_tac. assert (inc (J x y) (source f)). wr H9. bw. aw. eee. order_tac. rww W_spf. rww W_spf. app H11. rw product2_order_pr. red. bwi H9. rw H9. eee. aw. wrr order_reflexivity. wrr H5. aw. am. am. rw H6. order_tac. rw H6. order_tac. Qed. Lemma Exercice1_6j: forall f r r' r'', order r -> order r' -> nonempty (substrate r) -> nonempty (substrate r') -> increasing_fun f (product2_order r r') r'' -> (restriction2_axioms (second_partial_function f) (substrate r) (set_of_increasing_mappings r' r'')). Proof. ir. nin (Exercice1_6h H H0 H1 H2 H3). red in H3. ee. assert (partial_fun_axioms f). red. split. am. wr H8. bw. exists (substrate r). exists (substrate r'). tv. cp (function_spfa H11). red. ee. am. uf second_partial_fun. uf second_partial_fun. aw. uf second_partial_fun. uf second_partial_function. aw. ue. uf second_partial_fun. uf second_partial_function. aw. rw H5. uf set_of_increasing_mappings. wr H9. app Z_sub. red. ir. ufi image_by_fun H13. awi H13. nin H13. ee. cp (W_pr H12 H14). rwi W_spfa H15. wr H15. rw set_of_increasing_mappings_pr. eee. wr H15. app function_spf. rww H4. uf second_partial_fun. uf second_partial_fun. aw. uf second_partial_fun; uf second_partial_fun; aw. wr H15. app (Exercice1_6i (r:=r) (f:=f) (x:=x0)). red; intuition. tv. am. am. rww H4. Qed. Definition second_partial_map2 r r' r'':= BL (fun f=> restriction2 (second_partial_function f) (substrate r) (set_of_increasing_mappings r' r'')) (set_of_increasing_mappings (product2_order r r') r'') (set_of_increasing_mappings r (increasing_mappings_order r' r'')). Lemma Exercice1_6k: forall r r' r'', order r -> order r' -> order r'' -> nonempty (substrate r) -> nonempty (substrate r') -> transf_axioms (fun f=> restriction2 (second_partial_function f) (substrate r) (set_of_increasing_mappings r' r'')) (set_of_increasing_mappings (product2_order r r') r'') (set_of_increasing_mappings r (increasing_mappings_order r' r'')). Proof. ir. red. ir. awi H4. ee. cp (Exercice1_6j H H0 H2 H3 H7). nin (Exercice1_6h H H0 H2 H3 H7). assert (Ha: partial_fun_axioms c). red. split. am. rw H5. exists (substrate r). exists (substrate r'). tv. bw. set (g:= restriction2 (second_partial_function c) (substrate r) (set_of_increasing_mappings r' r'')). assert (is_function g). uf g. app function_restriction2. aw. eee. uf g. uf restriction2. aw. rww imo_substrate. uf g. uf restriction2. aw. red. eee. app imo_order. uf g. uf restriction2. aw. rww imo_substrate. uf g. uf restriction2. aw. ir. assert (inc x (substrate r)). order_tac. assert (inc y (substrate r)). order_tac. uf g. rww W_restriction2. rww W_restriction2. rww W_spfa. rww W_spfa. rww imo_pr. uf function_order_r. cp (Exercice1_6i H H0 H2 H3 H7 H14). cp (Exercice1_6i H H0 H2 H3 H7 H13). cp H15. red in H15. cp H16. red in H16. aw. eee. ir. rww W_spf. rww W_spf. assert (inc (J x i) (source c)). rw H5. bw. aw. eee. assert (inc (J y i) (source c)). rw H5. bw. aw. eee. red in H7. ee. app H36. rw product2_order_pr. red. ee; aw. eee. eee. order_tac. ue. ue. ue. ue. ue. ue. app imo_order. app order_product2_order. am. Qed. Lemma Exercice1_6l: forall r r' r'', order r -> order r' -> order r'' -> nonempty (substrate r) -> nonempty (substrate r') -> let f:= second_partial_map2 r r' r'' in (is_function f & source f = (set_of_increasing_mappings (product2_order r r') r'') & target f = (set_of_increasing_mappings r (increasing_mappings_order r' r'')) & bijective f). Proof. ir. uf f. uf second_partial_map2. aw. eee. app bl_function. app Exercice1_6k. split. (* injectivity *) app injective_bl_function. app Exercice1_6k. ir. awi H4. awi H5. ee. app funct_extensionality. ue. ue. assert (Ha: partial_fun_axioms u). red. split. am. rw H10. exists (substrate r). exists (substrate r'). bw. assert (Hb: partial_fun_axioms v). red. split. am. rw H7. exists (substrate r). exists (substrate r'). bw. cp (Exercice1_6h H H0 H2 H3 H9). cp (Exercice1_6h H H0 H2 H3 H12). ee. assert (second_partial_function u = second_partial_function v). app funct_extensionality. app function_spfa. app function_spfa. uf second_partial_function. aw. ue. uf second_partial_function. aw. rw H16. rww H15. rw H11; rww H8. ir. ufi second_partial_function H17. ufi BL H17. rwi corresp_source H17. rwi H14 H17. transitivity (W x (restriction2 (second_partial_function u) (substrate r) (set_of_increasing_mappings r' r''))). rww W_restriction2. app Exercice1_6j. rw H6. rww W_restriction2. app Exercice1_6j. clear H6. ir. bwi H10. rwi H10 H6. awi H6. ee. assert (inc (P x) (domain (source u))). ue. cp (W_spfa Ha H20). assert (inc (P x) (domain (source v))). rww H13. cp (W_spfa Hb H22). wri H17 H23. rwi H21 H23. assert (x = J (P x) (Q x)). sy. aw. rw H24. assert (inc (Q x) (range (source u))). ue. wr (W_spf Ha H20 H25). assert (inc (Q x) (range (source v))). ue. rwi H14 H20; wri H13 H20. wr (W_spf Hb H20 H26). rww H23. am. am. app order_product2_order. am. app order_product2_order. am. (* surjectivity *) app surjective_bl_function. app Exercice1_6k. ir. awi H4. nin H4. ee. assert (substrate (product2_order r r') = product (substrate r) (substrate r')). bw. assert (transf_axioms (fun z => W (Q z) (W (P z) y)) (product (substrate r)(substrate r')) (substrate r'')). red. ir. awi H9. ee. wri H5 H10. cp (inc_W_target H4 H10). rwi H6 H12. rwii imo_substrate H12. awii H12. ee. wr H14. app inc_W_target. ue. set (g:= BL (fun z => W (Q z) (W (P z) y)) (product (substrate r)(substrate r')) (substrate r'')). assert (is_function g). uf g. app bl_function. assert (increasing_fun g (product2_order r r') r''). red. eee. app order_product2_order. uf g. aw. uf g. aw. ir. assert (inc x (source g)). uf g. uf BL. rw corresp_source. wr H8. order_tac. assert (inc y0 (source g)). uf g. uf BL. rw corresp_source. wr H8. order_tac. uf g. ufi g H12. ufi BL H12. rwi corresp_source H12. ufi g H13. ufi BL H13. rwi corresp_source H13. rww W_bl_function. rww W_bl_function. awi H12. awi H13. ee. wri H5 H16. wri H5 H14. cp (inc_W_target H4 H16). cp (inc_W_target H4 H14). rwi H6 H18. rwi H6 H19. rwii imo_substrate H18. rwii imo_substrate H19. awii H18; awii H19; ee. rwi product2_order_pr H11. red in H11. ee. red in H22. ee. cp (H33 _ _ H28). red in H7. ee. cp (H39 _ _ H27). rwii imo_pr H40. ee. red in H42. ee. assert (gle r'' (W (Q x) (W (P x) y)) (W (Q x) (W (P y0) y))). app H48. order_tac. assert (source g = (product (substrate r) (substrate r'))). uf g. aw. assert (range (source g) = substrate r'). rw H12. rww range_product. assert (partial_fun_axioms g). red. split. am. rw H12. exists (substrate r). exists (substrate r'). tv. exists g. split. aw. ee. am. bw. uf g. aw. am. app order_product2_order. sy. app funct_extensionality. app function_restriction2. app Exercice1_6j. uf restriction2. aw. uf restriction2. aw. rwi imo_substrate H6. am. am. am. ir. rw W_restriction2. assert (Ha: inc x (domain (source g))). uf g. aw. nin H3. exists y0. aw. eee. fprops. rw W_spfa. cp (inc_W_target H4 H15). rwi H6 H16. rwii imo_substrate H16. awi H16. ee. app funct_extensionality. app function_spf. uf second_partial_fun. aw. rw H17. sy; am. uf second_partial_fun. rw H18. aw. uf g. aw. ir. rw W_spf. uf g. rw W_bl_function. aw. am. aw. eee. am. am. rww H13. ue. am. am. am. am. app Exercice1_6j. ue. am. app imo_order. Qed. Lemma Exercice1_6m: forall r r' r'', order r -> order r' -> order r'' -> nonempty (substrate r) -> nonempty (substrate r') -> order_isomorphism (second_partial_map2 r r' r'') (increasing_mappings_order (product2_order r r') r'') (increasing_mappings_order r (increasing_mappings_order r' r'')). Proof. ir. cp (Exercice1_6l H H0 H1 H2 H3). cbv zeta in H4. ee. assert (Hc:=Exercice1_6k H H0 H1 H2 H3). assert (Hd:order (increasing_mappings_order r' r'')). app imo_order. assert (Ha:order (product2_order r r')). app order_product2_order. red. eee. app imo_order. app imo_order. rww imo_substrate. rww imo_substrate. rw H5. ir. rww imo_pr. rww imo_pr. wr H6. cp H8; cp H9. awii H8; awii H9. ee. assert (Hu:partial_fun_axioms x). red. split. am. rw H15. bw. exists (substrate r). exists (substrate r'). tv. assert (Hv:partial_fun_axioms y). red. split. am. rw H12. bw. exists (substrate r). exists (substrate r'). tv. set (u1:= inc (W x (second_partial_map2 r r' r'')) (target (second_partial_map2 r r' r''))). assert u1. uf u1. app inc_W_target. ue. set (u2:= inc (W y (second_partial_map2 r r' r'')) (target (second_partial_map2 r r' r''))). assert u2. uf u2. app inc_W_target. ue. uf second_partial_map2. aw. cp H18; cp H19. ufi u1 H18; ufi u2 H19. unfold second_partial_map2 in H18, H19 |-*. rwii W_bl_function H18. rwii W_bl_function H19. awi H18. awi H19. cp (Exercice1_6j H H0 H2 H3 H17). cp (Exercice1_6j H H0 H2 H3 H14). app iff_eq. ir. ee. am. am. red. eee. ir. rww W_restriction2. rww W_restriction2. assert (inc i (domain (source x))). rw H30. bw. rww domain_product. assert (inc i (domain (source y))). rw H27. bw. rww domain_product. assert (range (source x) = substrate r'). rw H30. bw. rww range_product. assert (range (source y) = substrate r'). rw H27. bw. rww range_product. rww W_spfa. rww W_spfa. rw imo_pr. aw. aw. cp (Exercice1_6i H H0 H2 H3 H29 H39). cp (Exercice1_6i H H0 H2 H3 H32 H39). red in H44; red in H45. eee. red. eee. red. eee. red. eee. ir. rww W_spf. rww W_spf. nin H26. ee. app H62. bw. aw. eee. ue. ue. am. am. (* converse *) ir. eee. red. ir. eee. ir. bwi H33. awi H33. ee. assert(inc (P i) (domain (source x))). rw H15. bw. rww domain_product. assert(inc (P i) (domain (source y))). rw H12. wrr H15. assert(inc (Q i) (range (source x))). rw H15. bw. rww range_product. assert(inc (Q i) (range (source y))). rw H12. wrr H15. assert (i = J (P i) (Q i)). sy; app pair_recov. rw H40. red in H26. ee. cp (H46 _ H34). rwii W_restriction2 H47. rwii W_restriction2 H47. rwi W_spfa H47. rwi W_spfa H47. rwi imo_pr H47. ee. red in H49. ee. cp (H55 _ H35). rwii W_spf H56. rwii W_spf H56. am. am. am. am. am. am. am. am. am. am. am. am. Qed. Lemma constant_increasing: forall r r', order r -> order r' -> forall y (Hy: inc y (substrate r')), (inc (constant_function (substrate r) (substrate r') y Hy) (set_of_increasing_mappings r r')). Proof. ir. aw. eee. app function_constant_fun. uf constant_function. aw. uf constant_function. aw. red. eee. app function_constant_fun. uf constant_function. aw. uf constant_function. aw. ir. rw W_constant. rw W_constant. order_tac. order_tac. order_tac. Qed. Lemma constant_increasing1: forall r r', order r -> order r' -> nonempty (substrate r) -> forall y (Hy: inc y (substrate r')) y' (Hy': inc y' (substrate r')), gle r' y y' = gle (increasing_mappings_order r r') (constant_function (substrate r) (substrate r') y Hy) (constant_function (substrate r) (substrate r') y' Hy'). Proof. ir. rw imo_pr. app iff_eq. ir. eee. app constant_increasing. app constant_increasing. red. ee. app function_constant_fun. app function_constant_fun. uf constant_function. aw. uf constant_function. aw. uf constant_function. aw. uf constant_function. aw. ir. rw W_constant. rw W_constant. am. am. am. ir. ee. awi H4. nin H1. red in H4. ee. cp (H10 _ H1). rwi W_constant H11. rwi W_constant H11. am. am. am. am. am. Qed. Lemma Exercice1_6n: forall r r', order r -> order r'-> nonempty (substrate r) -> lattice r' = lattice (increasing_mappings_order r r'). Proof. ir. app iff_eq. clear H1. ir. uf lattice. split. app imo_order. rw imo_substrate. ir. assert (Hx:=H2). assert (Hy:=H3). awi H2. awi H3. set (E:=substrate r) in *; set (E':=substrate r') in *. assert (transf_axioms (fun i=> sup r' (W i x) (W i y)) E E'). red. ir. ee. assert (inc (W c x) E'). wr H9. app inc_W_target. ue. assert (inc (W c y) E'). wr H6. app inc_W_target. ue. nin (lattice_sup_pr H1 H11 H12). uf E'. order_tac. assert (transf_axioms (fun i=> inf r' (W i x) (W i y)) E E'). red. ir. ee. assert (inc (W c x) E'). wr H10. app inc_W_target. ue. assert (inc (W c y) E'). wr H7. app inc_W_target. ue. nin (lattice_inf_pr H1 H12 H13). uf E'. order_tac. set (f1:= BL (fun i=> sup r' (W i x) (W i y)) E E'). set (f2:= BL (fun i=> inf r' (W i x) (W i y)) E E'). assert (is_function f1). uf f1. app bl_function. assert (is_function f2). uf f2. app bl_function. assert (Ha:increasing_fun f1 r r'). red. eee. uf f1. aw. uf f1. aw. uf f1. ir. assert (inc x0 E). uf E. order_tac. assert (inc y0 E). uf E. order_tac. aw. set (a:= W x0 x). set (b:= W x0 y). assert (inc a E'). wr H12. uf a. app inc_W_target. ue. assert (inc b E'). wr H9. uf b. app inc_W_target. ue. set (a':= W y0 x). set (b':= W y0 y). assert (inc a' E'). wr H12. uf a'. app inc_W_target. ue. assert (inc b' E'). wr H9. uf b'. app inc_W_target. ue. cp (lattice_sup_pr H1 H17 H18). cp (lattice_sup_pr H1 H19 H20). ee. ap H26. assert (gle r' a a'). uf a; uf a'. red in H13. ee. app H31. order_tac. assert (gle r' b b'). uf b; uf b'. red in H10. ee. app H31. order_tac. assert (Hb:increasing_fun f2 r r'). red. eee. uf f2. aw. uf f2. aw. uf f2. ir. assert (inc x0 E). uf E. order_tac. assert (inc y0 E). uf E. order_tac. aw. set (a:= W x0 x). set (b:= W x0 y). assert (inc a E'). wr H12. uf a. app inc_W_target. ue. assert (inc b E'). wr H9. uf b. app inc_W_target. ue. set (a':= W y0 x). set (b':= W y0 y). assert (inc a' E'). wr H12. uf a'. app inc_W_target. ue. assert (inc b' E'). wr H9. uf b'. app inc_W_target. ue. cp (lattice_inf_pr H1 H17 H18). cp (lattice_inf_pr H1 H19 H20). ee. ap H24. assert (gle r' a a'). uf a; uf a'. red in H13. ee. app H31. order_tac. assert (gle r' b b'). uf b; uf b'. red in H10. ee. app H31. order_tac. assert (inc f1 (set_of_increasing_mappings r r')). aw. eee; uf f1; aw. assert (inc f2 (set_of_increasing_mappings r r')). aw. eee; uf f2; aw. ee. exists f1. rw least_upper_bound_pr. split. red. split. rww imo_substrate. ir. nin (doubleton_or H16); rw H17; rw imo_pr. eee. red. uf f1. eee. aw. aw. ir. aw. assert (inc (W i x) E'). wr H14. app inc_W_target. ue. assert (inc (W i y) E'). wr H11. app inc_W_target. ue. cp (lattice_sup_pr H1 H19 H20). nin H21. am. am. am. eee. red. uf f1. eee. aw. aw. ir. aw. assert (inc (W i x) E'). wr H14. app inc_W_target. ue. assert (inc (W i y) E'). wr H11. app inc_W_target. ue. cp (lattice_sup_pr H1 H19 H20). ee. am. am. am. ir. red in H16. nin H16. assert (inc x (doubleton x y)). fprops. cp (H17 _ H18). assert (inc y (doubleton x y)). fprops. cp (H17 _ H20). rw imo_pr. rwi imo_pr H19; rwi imo_pr H21. ee. am. am. aw. red. red in H23. eee. uf f1. aw. uf f1; aw. ir. red in H25; ee. cp (H38 _ H32). cp (H31 _ H32). uf f1. aw. assert (inc (W i x) E'). wr H14. app inc_W_target. ue. assert (inc (W i y) E'). wr H11. app inc_W_target. ue. cp (lattice_sup_pr H1 H41 H42). ee. app H45. am. am. am. am. am. am. am. am. am. am. app imo_order. rw imo_substrate. red. ir. nin (doubleton_or H16); rw H17; am. am. am. exists f2. aw. split. red. split. rww imo_substrate. ir. nin (doubleton_or H16); rw H17; rw imo_pr. eee. red. uf f2. eee. aw. aw. ir. aw. assert (inc (W i x) E'). wr H14. app inc_W_target. ue. assert (inc (W i y) E'). wr H11. app inc_W_target. ue. cp (lattice_inf_pr H1 H19 H20). nin H21. am. am. am. eee. red. uf f2. eee. aw. aw. ir. aw. assert (inc (W i x) E'). wr H14. app inc_W_target. ue. assert (inc (W i y) E'). wr H11. app inc_W_target. ue. cp (lattice_inf_pr H1 H19 H20). ee. am. am. am. ir. red in H16. nin H16. assert (inc x (doubleton x y)). fprops. cp (H17 _ H18). assert (inc y (doubleton x y)). fprops. cp (H17 _ H20). rw imo_pr. rwi imo_pr H19; rwi imo_pr H21. ee. am. am. aw. red. red in H23. eee. uf f2. aw. uf f2; aw. ir. red in H25; ee. cp (H38 _ H32). cp (H31 _ H32). uf f2. aw. assert (inc (W i x) E'). wr H14. app inc_W_target. ue. assert (inc (W i y) E'). wr H11. app inc_W_target. ue. cp (lattice_inf_pr H1 H41 H42). ee. app H45. am. am. am. am. am. am. am. am. am. am. app imo_order. rw imo_substrate. red. ir. nin (doubleton_or H16); rw H17; am. am. am. am. am. am. am. am. am. ir. red. split. am. assert (He:= H1). set (E:=substrate r) in *; set (E':=substrate r') in *. nin H1. ir. set (f1:= constant_function E E' x H3). set (f2:= constant_function E E' y0 H4). assert (inc f1 (set_of_increasing_mappings r r')). uf f1. uf E. uf E'. app constant_increasing. assert (inc f2 (set_of_increasing_mappings r r')). uf f2. uf E. uf E'. app constant_increasing. assert (inc f1 (substrate (increasing_mappings_order r r'))). rw imo_substrate. am. am. am. assert (inc f2 (substrate (increasing_mappings_order r r'))). rw imo_substrate. am. am. am. split. cp (lattice_sup_pr H2 H7 H8). set (f3:= sup (increasing_mappings_order r r') f1 f2) in *. ee. assert (inc f3 (substrate (increasing_mappings_order r r'))). app (inc_arg2_substrate H9). rwi imo_substrate H12. awi H12. ee. assert (inc (W y f3) E'). uf E'. wr H14. app inc_W_target. rww H13. exists (W y f3). rw least_upper_bound_pr. split. red. ee. am. ir. nin (doubleton_or H17); rw H18. rwi imo_pr H9. nin H9. nin H19. awi H20. red in H20. ee. cp (H26 _ H1). ufi f1 H27. rwi W_constant H27. am. am. am. am. rwi imo_pr H10. nin H10. nin H19. awi H20. red in H20. ee. cp (H26 _ H1). ufi f2 H27. rwi W_constant H27. am. am. am. am. ir. nin H17. set (f4:= constant_function E E' z H17). assert (inc f4 (set_of_increasing_mappings r r')). uf f4. uf E. uf E'. app constant_increasing. assert (gle (increasing_mappings_order r r') f1 f4). uf f1. uf f4. uf E. uf E'. wr constant_increasing1. app H18. fprops. am. am. am. assert (gle (increasing_mappings_order r r') f2 f4). uf f2. uf f4. uf E. uf E'. wr constant_increasing1. app H18. fprops. am. am. am. cp (H11 _ H20 H21). rwi imo_pr H22. ee. red in H24. ee. cp (H30 _ H1). ufi f4 H31. rwi W_constant H31. am. am. am. am. am. red. ir. nin (doubleton_or H17); rw H18; am. am. am. am. am. cp (lattice_inf_pr H2 H7 H8). set (f3:= inf (increasing_mappings_order r r') f1 f2) in *. ee. assert (inc f3 (substrate (increasing_mappings_order r r'))). app (inc_arg1_substrate H9). rwi imo_substrate H12. awi H12. ee. assert (inc (W y f3) E'). uf E'. wr H14. app inc_W_target. rww H13. exists (W y f3). aw. split. red. ee. am. ir. nin (doubleton_or H17); rw H18. rwi imo_pr H9. nin H9. nin H19. awi H20. red in H20. ee. cp (H26 _ H1). ufi f1 H27. rwi W_constant H27. am. am. am. am. rwi imo_pr H10. nin H10. nin H19. awi H20. red in H20. ee. cp (H26 _ H1). ufi f2 H27. rwi W_constant H27. am. am. am. am. ir. nin H17. set (f4:= constant_function E E' z H17). assert (inc f4 (set_of_increasing_mappings r r')). uf f4. uf E. uf E'. app constant_increasing. assert (gle (increasing_mappings_order r r') f4 f1). uf f1. uf f4. uf E. uf E'. wr constant_increasing1. app H18. fprops. am. am. am. assert (gle (increasing_mappings_order r r') f4 f2). uf f2. uf f4. uf E. uf E'. wr constant_increasing1. app H18. fprops. am. am. am. cp (H11 _ H20 H21). rwi imo_pr H22. ee. red in H24. ee. cp (H30 _ H1). ufi f4 H31. rwi W_constant H31. am. am. am. am. red. ir. nin (doubleton_or H17); rw H18; am. am. am. am. am. Qed. Lemma Exercice1_6o: forall r r', order r -> order r'-> nonempty (substrate r)-> total_order (increasing_mappings_order r r') -> total_order r'. Proof. ir. nin H2. red. ee. am. ir. set (E:=substrate r) in *; set (E':=substrate r') in *. nin H1. set (f1:= constant_function E E' x H4). set (f2:= constant_function E E' y H5). assert (substrate r = source f1). uf f1. uf constant_function. aw. assert (substrate r' = target f1). uf f1. uf constant_function. aw. assert (is_function f1). uf f1. app function_constant_fun. assert (inc f1 (set_of_increasing_mappings r r')). aw. eee. red. eee. uf f1. ir. rw W_constant. rw W_constant. wrr order_reflexivity. uf E; order_tac. uf E; order_tac. assert (substrate r = source f2). uf f2. uf constant_function. aw. assert (substrate r' = target f2). uf f2. uf constant_function. aw. assert (is_function f2). uf f2. app function_constant_fun. assert (inc f2 (set_of_increasing_mappings r r')). aw. eee. red. eee. uf f2. ir. rw W_constant. rw W_constant. wrr order_reflexivity. uf E; order_tac. uf E; order_tac. rwi imo_substrate H3. ufi gge H3. nin (H3 _ _ H9 H13); rwi imo_pr H14; ee. awi H16. red in H16. ee. cp (H22 _ H1). ufi f1 H23; ufi f2 H23. rwi W_constant H23. rwi W_constant H23. au. am. am. am. am. awi H16. red in H16. ee. cp (H22 _ H1). ufi f1 H23; ufi f2 H23. rwi W_constant H23. rwi W_constant H23. au. am. am. am. am. am. am. Qed. Lemma Exercice1_6p: forall r r', order r -> order r'-> is_singleton (substrate r') -> total_order (increasing_mappings_order r r'). Proof. ir. red. ee. app imo_order. rww imo_substrate. ir. left. rw imo_pr. eee. rwi set_of_increasing_mappings_pr H2. awi H3. ee. red. aw. eee. ir. assert (inc (W i x) (substrate r')). wr H8. app inc_W_target. ue. assert (inc (W i y) (substrate r')). wr H5. app inc_W_target. ue. nin H1. rwi H1 H11; rwi H1 H12. rw (singleton_eq H11). rw (singleton_eq H12). wrr order_reflexivity. rw H1. fprops. am. am. am. am. am. am. Qed. Lemma Exercice1_6q: forall r r', order r -> order r'-> is_singleton (substrate r) -> total_order r' -> total_order (increasing_mappings_order r r'). Proof. ir. red. ee. app imo_order. rww imo_substrate. nin H1. ir. uf gge. rww imo_pr. rww imo_pr. cp H3; cp H4; awii H3; awii H4. ee. assert (inc (W x y) (substrate r')). wr H8. app inc_W_target. ue. rw H1. fprops. assert (inc (W x x0) (substrate r')). wr H11. app inc_W_target. ue. rw H1. fprops. assert (forall u, inc u (substrate r) -> W u y = (W x y)). ir. rwi H1 H15. rww (singleton_eq H15). assert (forall u, inc u (substrate r) -> W u x0 = (W x x0)). ir. rwi H1 H16. rww (singleton_eq H16). nin H2. nin (H17 _ _ H13 H14). right. eee. red. eee. ir. rww H15. rww H16. left. eee. red. eee. ir. rww H15. rww H16. Qed. Lemma Exercice1_6r: forall r r', order r -> order r'-> total_order r -> total_order r' -> (exists a, exists b, substrate r' = doubleton a b) -> total_order (increasing_mappings_order r r'). Proof. ir. nin H3; nin H3. nin (equal_or_not x x0). app Exercice1_6p. rw H3. rw H4. rw doubleton_singleton. exists x0. tv. assert (exists u, exists v, substrate r' = doubleton u v & glt r' u v). red in H2. ee. assert (inc x (substrate r')). rw H3. fprops. assert (inc x0 (substrate r')). rw H3. fprops. nin (H5 _ _ H6 H7). exists x. exists x0. split. am. red. ee. am. am. exists x0. exists x. split. rw doubleton_symm. am. red. ee. am. intuition. clear H3. clear H4. nin H5. nin H3. nin H3. red. ee. app imo_order. rww imo_substrate. ir. cp H5; cp H6. awii H5; awii H6. ee. assert (forall u, inc u (substrate r) -> glt r' (W u x3) (W u y) -> (W u x3= x1 & W u y=x2)). ir. cp (inc_lt1_substrate H16). cp (inc_lt2_substrate H16). rwi H3 H17. rwi H3 H18. nin (doubleton_or H17); nin (doubleton_or H18);rwi H19 H16; rwi H20 H16. nin H16. elim H21. tv. eee. nin H16. order_tac. nin H16. elim H21. tv. nin (p_or_not_p (exists u, inc u (substrate r)& glt r' (W u x3) (W u y))). ir. nin H16. ee. left. rww imo_pr. ee. am. am. red. eee. ir. assert (inc (W i x3) (substrate r')). wr H13. app inc_W_target. ue. assert (inc (W i y) (substrate r')). wr H10. app inc_W_target. ue. nin H2. rwi H3 H19. rwi H3 H20. nin (doubleton_or H19). rw H22. nin (doubleton_or H20); rw H23. order_tac. rw H3. fprops. nin H4. am. rw H22. nin (doubleton_or H20); rw H23. cp (H15 _ H16 H17). nin H1. nin (H25 _ _ H16 H18). red in H11. ee. cp (H32 _ _ H26). wr H27. wrr H23. red in H26. nin H24. red in H14. ee. cp (H32 _ _ H26). rwi H22 H33; wri H27 H33. order_tac. order_tac. wr H10. wr H23. app inc_W_target. ue. right. red. rww imo_pr. eee. red. eee. ir. assert (inc (W i y) (substrate r')). wr H10. app inc_W_target. ue. assert (inc (W i x3) (substrate r')). wr H13. app inc_W_target. ue. nin H2. nin (H20 _ _ H18 H19). am. red in H21. nin (equal_or_not (W i y) (W i x3)). rw H22. order_tac. elim H16. exists i. split. am. red. split. am. intuition. Qed. Lemma Exercice1_6s: forall r r', order r -> order r'-> nonempty(substrate r) -> nonempty (substrate r') -> total_order (increasing_mappings_order r r') -> (is_singleton (substrate r') \/ (is_singleton (substrate r) & total_order r') \/ (total_order r' & total_order r & exists u , exists v, substrate r' = doubleton u v)). Proof. ir. cp (Exercice1_6o H H0 H1 H3). nin H2. nin (equal_or_not (substrate r') (singleton y)). ir. left. exists y. tv. ir. right. nin H1. nin (equal_or_not (substrate r) (singleton y0)). ir. left. split. exists y0. tv. am. ir. right. split. am. assert (exists y1, inc y1 (substrate r') & y1 <> y). app exists_proof. red. ir. elim H5. set_extens. nin (equal_or_not x y). rw H9. fprops. elim (H7 x). intuition. rw (singleton_eq H8). am. nin H7. assert (exists u, exists v, inc u (substrate r') & inc v (substrate r') & glt r' u v). ee. nin H4. nin (H9 _ _ H2 H7). exists y; exists x. eee. red. au. exists x; exists y. eee. red. eee. clear H2. clear H5. clear H7. nin H8. nin H2. assert (exists y1, inc y1 (substrate r) & y1 <> y0). app exists_proof. red. ir. elim H6. set_extens. nin (equal_or_not x2 y0). rw H8. fprops. elim (H5 x2). intuition. rw (singleton_eq H7). am. nin H5. rename y0 into a; rename x2 into a'; rename x0 into b; rename x1 into b'. assert (total_order r). red. ee. am. ir. set (f1:= fun u=> Yo (gle r u x0) b b'). assert (transf_axioms f1 (substrate r) (substrate r')). red. ir. uf f1. nin (p_or_not_p (gle r c x0)). ir. rww Y_if_rw. ir. rww Y_if_not_rw. assert (inc (BL f1 (substrate r) (substrate r')) (set_of_increasing_mappings r r')). aw. eee. app bl_function. red. eee. app bl_function. aw. aw. ir. assert (inc x1 (substrate r)). order_tac. assert (inc y1 (substrate r)). order_tac. rww W_bl_function. rww W_bl_function. uf f1. nin (p_or_not_p (gle r x1 x0)). ir. rww Y_if_rw. nin (p_or_not_p (gle r y1 x0)). ir. rww Y_if_rw. order_tac. ir. rww Y_if_not_rw. nin H9. am. rww Y_if_not_rw. nin (p_or_not_p (gle r y1 x0)). rww Y_if_rw. elim H16. order_tac. rww Y_if_not_rw. order_tac. set (f2:= fun u=> Yo (gle r u y0) b b'). assert (transf_axioms f2 (substrate r) (substrate r')). red. ir. uf f2. nin (p_or_not_p (gle r c y0)). rww Y_if_rw. rww Y_if_not_rw. assert (inc (BL f2 (substrate r) (substrate r')) (set_of_increasing_mappings r r')). aw. eee. app bl_function. red. eee. app bl_function. aw. aw. ir. assert (inc x1 (substrate r)). order_tac. assert (inc y1 (substrate r)). order_tac. rww W_bl_function. rww W_bl_function. uf f2. nin (p_or_not_p (gle r x1 y0)). rww Y_if_rw. nin (p_or_not_p (gle r y1 y0)). rww Y_if_rw. order_tac. rww Y_if_not_rw. nin H9. am. rww Y_if_not_rw. nin (p_or_not_p (gle r y1 y0)). rww Y_if_rw. elim H18. order_tac. rww Y_if_not_rw. order_tac. nin H3. rwi imo_substrate H16. nin (H16 _ _ H13 H15). rwi imo_pr H17. ee. red in H19. ee. awi H25. ufi f1 H25; ufi f2 H25. cp (H25 _ H11). awi H26. nin (p_or_not_p (gle r y0 x0)). ir. right. red. am. ir. rwi Y_if_not_rw H26. rwi Y_if_rw H26. nin H9. elim H28. order_tac. order_tac. am. am. am. am. am. am. am. red in H17. rwi imo_pr H17. ee. red in H19. ee. awi H25. ufi f1 H25; ufi f2 H25. cp (H25 _ H10). awi H26. nin (p_or_not_p (gle r x0 y0)). ir. left. am. ir. rwi Y_if_not_rw H26. rwi Y_if_rw H26. nin H9. elim H28. order_tac. order_tac. am. am. am. am. am. am. am. am. am. split. am. nin (equal_or_not (substrate r') (doubleton b b')). ir. exists b. exists b'. rw H8. tv. ir. assert (exists b'', inc b'' (substrate r') & b'' <> b & b'' <> b'). app exists_proof. red. ir. elim H8. set_extens. nin (equal_or_not x0 b). rw H11. fprops. nin (equal_or_not x0 b'). rw H12. fprops. elim (H9 x0). intuition. ee. nin (doubleton_or H10); rw H14; am. assert (exists u, exists v, exists w, inc u (substrate r') & inc v (substrate r') & inc w (substrate r') & glt r' u v & glt r' v w). nin H9. ee. red in H4. ee. nin (H15 _ _ H9 H2). exists x0; exists b; exists b'. eee. red. split. am. am. nin (H15 _ _ H9 H13). exists b; exists x0; exists b'. eee. red. red in H16. ee. am. intuition. red. ee. am. am. exists b; exists b'; exists x0. eee. red. ee. red in H17. am. intuition. clear H2. clear H9. clear H8. nin H10. nin H2. nin H2. ee. set (f:= constant_function (substrate r) (substrate r') x1 H8). assert (inc f (set_of_increasing_mappings r r')). uf f. aw. eee. app function_constant_fun. uf constant_function. aw. uf constant_function. aw. red. eee. app function_constant_fun. uf constant_function. aw. uf constant_function. aw. ir. rw W_constant. rw W_constant. order_tac. order_tac. order_tac. cp (total_order_lattice H7). cp (lattice_sup_pr H14 H1 H5). cp (lattice_inf_pr H14 H1 H5). set (u:=inf r a a'). assert (inc u (substrate r)). nin H16. order_tac. set (v:=sup r a a'). assert (inc v (substrate r)). nin H15. order_tac. assert (glt r u v). red. ee. order_tac. red. ir. ufi u H23; ufi v H23. rwi H23 H16. rwi H23 H19. assert (gle r a a'). order_tac. assert (gle r a' a). order_tac. elim H12. order_tac. set (g:= fun x=> Yo (gle r x u) x0 x2). assert (transf_axioms g (substrate r) (substrate r')). red. ir. uf g. nin (p_or_not_p (gle r c u)). ir. rww Y_if_rw. ir. rww Y_if_not_rw. assert (inc (BL g (substrate r) (substrate r')) (set_of_increasing_mappings r r')). aw. eee. app bl_function. red. eee. app bl_function. aw. aw. ir. assert (inc x3 (substrate r)). order_tac. assert (inc y0 (substrate r)). order_tac. rww W_bl_function. rww W_bl_function. uf g. nin (p_or_not_p (gle r x3 u)). ir. rww Y_if_rw. nin (p_or_not_p (gle r y0 u)). ir. rww Y_if_rw. order_tac. ir. rww Y_if_not_rw. nin H10. nin H11. order_tac. ir. rww Y_if_not_rw. nin (p_or_not_p (gle r y0 u)). ir. rww Y_if_rw. elim H28. order_tac. ir. rww Y_if_not_rw. order_tac. nin H3. rwi imo_substrate H22. nin (H22 _ _ H13 H21). rwi imo_pr H23. ee. red in H25. ee. awi H35. ufi f H35. awi H35. cp (H35 _ H17). awi H36. rwi W_constant H36. ufi g H36. rwi Y_if_rw H36. nin H10. elim H37. order_tac. order_tac. am. am. am. am. am. red in H23. rwi imo_pr H23. ee. red in H25. ee. awi H35. ufi f H35. awi H35. cp (H35 _ H18). awi H36. rwi W_constant H36. ufi g H36. rwi Y_if_not_rw H36. nin H11. elim H37. order_tac. red. ir. nin H19. elim H38. order_tac. am. am. am. am. am. am. am. Qed. (* Exercise 7 *) Definition comp_rel r:= fun x y => (gle r x y \/ gle r y x). Definition cr_equiv r := Exercice1.Sgraph (comp_rel r) (substrate r). Definition cr_component r := Exercice1.connected_comp (comp_rel r) (substrate r). Lemma cr_properties: forall r, order r -> (is_equivalence (cr_equiv r) & (forall x y, comp_rel r x y -> (inc x (substrate r)& inc y (substrate r))) & substrate (cr_equiv r) = substrate r & (forall x, inc x (substrate r) -> class (cr_equiv r) x = cr_component r x) & (forall x y, comp_rel r x y -> related (cr_equiv r) x y)). Proof. ir. uf cr_equiv. uf cr_component. assert (forall x y : Set, comp_rel r x y -> (inc x (substrate r) & inc y (substrate r))). ir. nin H0. split; order_tac. split; order_tac. assert (reflexive_r (comp_rel r) (substrate r)). red. ir. app iff_eq. uf comp_rel. ir. left. order_tac. ir. nin (H0 _ _ H1). am. assert (symmetric_r (comp_rel r)). red. uf comp_rel. ir. intuition. assert (forall x y, comp_rel r x y -> inc x (substrate r)). ir. nin (H0 _ _ H3). am. ee. app Exercice1.equivalence_Sgraph. am. app Exercice1.substrate_Sgraph. ir. app Exercice1.connected_comp_class. ir. red. uf Exercice1.Sgraph. uf graph_on. ee. Ztac. nin (H0 _ _ H4). fprops. red. aw. exists (Exercice1.chain_pair x y). au. Qed. Lemma Exercice1_7a: forall r x, right_directed r -> inc x (substrate r) -> cr_component r x = substrate r. Proof. ir. rwi right_directed_pr H. nin H. cp (cr_properties H). ee. wr H5. app extensionality. wr H4. app sub_class_substrate. red. ir. bw. nin (H1 _ _ H0 H7). nin H8. nin H9. assert (related (cr_equiv r) x x1). app H6. red. left. am. assert (related (cr_equiv r) x1 x0). app H6. red. right. am. apply transitivity_e with x1. am. am. nin H2. am. nin H2; am. Qed. Lemma Exercice1_7b: forall r x, left_directed r -> inc x (substrate r) -> cr_component r x = substrate r. Proof. ir. rwi left_directed_pr H. nin H. cp (cr_properties H). ee. wr H5. app extensionality. wr H4. app sub_class_substrate. red. ir. bw. nin (H1 _ _ H0 H7). nin H8. nin H9. assert (related (cr_equiv r) x x1). app H6. red. right. am. assert (related (cr_equiv r) x1 x0). app H6. red. left. am. apply transitivity_e with x1. am. am. nin H2. am. nin H2; am. Qed. Lemma Exercice1_7c: forall r r' f x y, increasing_fun f r r' -> decreasing_fun f r r' -> comp_rel r x y -> W x f = W y f. Proof. ir. assert (order r). nin H; eee. red in H. red in H0. ee. nin H1; cp (H12 _ _ H1); cp (H7 _ _ H1); red in H14; order_tac. Qed. Lemma Exercice1_7d: forall r r' f, increasing_fun f r r' -> decreasing_fun f r r' -> (exists x, inc x (substrate r) & cr_component r x = substrate r) -> (is_constant_function f). Proof. ir. assert (forall x y, comp_rel r x y -> W x f = W y f). ir. app (Exercice1_7c H H0 H2). red in H. ee. cp (cr_properties H3). nin H1. ee. red. ee. am. ir. assert (source f= substrate r). au. assert (forall u, inc u (source f) -> W u f = W x f). ir. rwi H16 H17. wri H13 H17. wri (H11 _ H1) H17. bwi H17. assert (related (cr_equiv r) u x). apply symmetricity. am. am. ufi cr_equiv H18. ufi Exercice1.Sgraph H18. ufi graph_on H18. red in H18. Ztac. nin H20. awi H20. ee. assert (forall x1, Exercice1.chained_r (comp_rel r) x1 -> Exercice1.chain_tail x1 = x -> W (Exercice1.chain_head x1) f = W x f). ir. nin x2. simpl. simpl in H24. simpl in H23. wr H24. app H2. simpl. simpl in H23. ee. cp (IHx2 H25 H24). simpl in H24. rww (H2 _ _ H23). wr (H23 _ H20 H22). ue. am. rw (H17 _ H14). rw (H17 _ H15). tv. Qed. Lemma Exercice1_7e: forall r r', order r -> order r' -> (exists u, exists v, inc u (substrate r') & inc v (substrate r') & u <>v) -> (exists x, inc x (substrate r) & cr_component r x <> substrate r) -> exists f, increasing_fun f r r' & decreasing_fun f r r' & ~(is_constant_function f). Proof. ir. nin H1. nin H1. nin H2. ee. cp (cr_properties H). ee. set (f:= (fun u => Yo (inc u (cr_component r x1)) x x0)). assert (transf_axioms f (substrate r) (substrate r')). red. ir. uf f. nin (p_or_not_p (inc c (cr_component r x1))). ir. rww Y_if_rw. ir. rww Y_if_not_rw. set (g:= BL f (substrate r) (substrate r')). assert (is_function g). uf g. app bl_function. assert (~ is_constant_function g). red. ir. red in H13. ee. assert (inc x1 (source g)). uf g. aw. elim H3. app extensionality. wr H9. wr H8. app sub_class_substrate. am. red. ir. assert (inc x2 (source g)). uf g. aw. cp (H14 _ _ H15 H17). ufi g H18. awi H18. ufi f H18. rwi Y_if_rw H18. nin (p_or_not_p (inc x2 (cr_component r x1))). tv. rwii Y_if_not_rw H18. contradiction. wrr H9. bw. app reflexivity_e. rww H8. nin H6; am. am. am. am. assert (forall a b, gle r a b -> (inc a (cr_component r x1) = inc b (cr_component r x1))). ir. wrr H9. bw. bw. assert (comp_rel r a b). red. left. am. cp (H10 _ _ H15). app iff_eq. ir. apply transitivity_e with a. am. am. am. ir. apply transitivity_e with b. am. am. app symmetricity. assert (forall a b, gle r a b -> W a g = W b g). ir. uf g. aw. uf f. nin (p_or_not_p (inc a (cr_component r x1))). ir. rww Y_if_rw. rww Y_if_rw. wrr (H14 _ _ H15). ir. rww Y_if_not_rw. rww Y_if_not_rw. wrr (H14 _ _ H15). order_tac. order_tac. assert (target g = substrate r'). uf g. aw. assert (source g = substrate r). uf g. aw. exists g. ee. red. eee. ir. rww (H15 _ _ H18). order_tac. wr H16. app inc_W_target. rw H17. order_tac. red. eee. ir. red. rww (H15 _ _ H18). order_tac. wr H16. app inc_W_target. rw H17. order_tac. am. Qed. (* Exercise 8 *) Lemma Exercice1_8: forall r r' f g, let A := Zo (substrate r) (fun z => W (W z f) g = z) in let B := Zo (substrate r') (fun z => W (W z g) f = z) in increasing_fun f r r' -> increasing_fun g r' r -> exists h, order_isomorphism h (induced_order r A)(induced_order r' B). Proof. ir. assert (forall x, inc x A -> inc (W x f) B). ir. ufi A H1. Ztac. clear H1. uf B. Ztac. red in H. ee. rw H6. app inc_W_target. wrr H5. rww H3. assert (forall x, inc x B -> inc (W x g) A). ir. ufi B H2. Ztac. clear H2. uf A. Ztac. red in H0. ee. rw H7. app inc_W_target. ue. ue. assert (Ha: source (restriction2 f A B) = A). uf restriction2. aw. assert (Hb: target (restriction2 f A B) = B). uf restriction2. aw. set (h:=restriction2 f A B). assert (restriction2_axioms f A B). red. red in H. ee. am. uf A. wr H5. app Z_sub. wr H6. uf B. app Z_sub. red. ir. ufi image_by_fun H8. awi H8. nin H8. nin H8. red in H9. wr (W_pr H H9). app H1. fprops. assert (is_function h). uf h. app function_restriction2. assert (injective h). red. split. am. uf h. rw Ha. intros x y Hx Hy. rww W_restriction2. rww W_restriction2. ir. ufi A Hx; ufi A Hy. Ztac. wr H7. clear Hy. Ztac. wr H5. sy; am. assert (surjective h). app surjective_pr6. uf h. rw Hb. rw Ha. ir. exists (W y g). ufi restriction2 H6. awi H6. split. app H2. rww W_restriction2. ufi B H6. Ztac. am. app H2. exists h. red in H. ee. assert (sub A (substrate r)). uf A. app Z_sub. assert (sub B (substrate r')). uf B. app Z_sub. red. ee. fprops. fprops. red; split;am. aw. aw. uf h. sy;am. uf h. rww Hb. aw. uf h. rw Ha. ir. rww W_restriction2. rww W_restriction2. app iff_eq. ir. awi H16. aw. app H11. app H1. app H1. am. am. ir. cp (H1 _ H14). cp (H1 _ H15). awi H16. aw. ufi A H14. Ztac. clear H14. ufi A H15. Ztac. wr H20. wr H21. red in H0; ee. app H26. am. am. Qed. (* Exercise 9 *) Lemma lattice_finite_sup1: forall r X x a, lattice r -> sub X (substrate r) -> least_upper_bound r X x -> inc a (substrate r) -> least_upper_bound r (tack_on X a) (sup r x a). Proof. ir. assert (order r). red in H;ee;am. assert (sub (tack_on X a) (substrate r)). app tack_on_sub. rwi (least_upper_bound_pr x H3 H0) H1. ee. assert (inc x (substrate r)). red in H1. ee. am. cp (lattice_sup_pr H H6 H2). ee. rw (least_upper_bound_pr (sup r x a) H3 H4). ee. red. ee. order_tac. ir. rwi tack_on_inc H10. nin H10. red in H1. ee. cp (H11 _ H10). order_tac. ue. ir. red in H10. app H9. ee. app H5. red. ee. am. ir. app H11. fprops. eee. Qed. Lemma lattice_finite_inf1: forall r X x a, lattice r -> sub X (substrate r) -> greatest_lower_bound r X x -> inc a (substrate r) -> greatest_lower_bound r (tack_on X a) (inf r x a). Proof. ir. assert (order r). red in H;ee;am. assert (sub (tack_on X a) (substrate r)). app tack_on_sub. rwi (greatest_lower_bound_pr x H3 H0) H1. ee. assert (inc x (substrate r)). red in H1. ee. am. cp (lattice_inf_pr H H6 H2). ee. rw (greatest_lower_bound_pr (inf r x a) H3 H4). ee. red. ee. app (inc_arg1_substrate H7). ir. rwi tack_on_inc H10. nin H10. red in H1. ee. cp (H11 _ H10). order_tac. ue. ir. red in H10. app H9. ee. app H5. red. ee. am. ir. app H11. fprops. eee. Qed. Lemma lattice_finite_sup2: forall r x, lattice r -> is_finite_set x -> nonempty x -> sub x (substrate r) -> has_supremum r x. Proof. ir. app (finite_set_induction2 (fun x => (sub x (substrate r))) (fun x => has_supremum r x)). ir. assert (inc a (substrate r)). app H3. fprops. wr (doubleton_singleton a). red in H. ee. nin (H5 _ _ H4 H4). am. ir. assert (inc b (substrate r)). app H5. fprops. assert (sub a (substrate r)). apply sub_trans with (tack_on a b). fprops. am. cp (H3 H7 H4). red in H8. nin H8. exists (sup r x0 b). ap (lattice_finite_sup1 H H7 H8 H6). Qed. Lemma lattice_finite_inf2: forall r x, lattice r -> is_finite_set x -> nonempty x -> sub x (substrate r) -> has_infimum r x. Proof. ir. app (finite_set_induction2 (fun x => (sub x (substrate r))) (fun x => has_infimum r x)). ir. assert (inc a (substrate r)). app H3. fprops. wr (doubleton_singleton a). red in H. ee. nin (H5 _ _ H4 H4). am. ir. assert (inc b (substrate r)). app H5. fprops. assert (sub a (substrate r)). apply sub_trans with (tack_on a b). fprops. am. cp (H3 H7 H4). red in H8. nin H8. exists (inf r x0 b). ap (lattice_finite_inf1 H H7 H8 H6). Qed. Lemma lattice_finite_sup3: forall r x y, lattice r -> is_finite_set x -> nonempty x -> sub x (substrate r) -> gle r (supremum r x) y = (forall z, inc z x -> gle r z y). Proof. ir. cp (lattice_finite_sup2 H H0 H1 H2). assert (order r). red in H; ee; am. cp (supremum_pr H4 H2 H3). ee. app iff_eq. ir. red in H5. ee. cp (H9 _ H8). order_tac. ir. app H6. red. ee. nin H1. cp (H7 _ H1). order_tac. am. Qed. Lemma lattice_finite_inf3: forall r x y, lattice r -> is_finite_set x -> nonempty x -> sub x (substrate r) -> gle r y (infimum r x) = (forall z, inc z x -> gle r y z). Proof. ir. cp (lattice_finite_inf2 H H0 H1 H2). assert (order r). red in H; ee; am. cp (infimum_pr H4 H2 H3). ee. app iff_eq. ir. red in H5. ee. cp (H9 _ H8). order_tac. ir. app H6. red. ee. nin H1. cp (H7 _ H1). order_tac. am. Qed. Lemma lattice_finite_sup4: forall r f y, lattice r -> fgraph f -> is_finite_set (domain f) -> nonempty (domain f) -> sub (range f) (substrate r) -> gle r (sup_graph r f) y = (forall z, inc z (domain f) -> gle r (V z f) y). Proof. ir. uf sup_graph. rw lattice_finite_sup3. app iff_eq. ir. app H4. fprops. ir. srwi H5. nin H5. ee. rw H6. app H4. am. am. app finite_range. nin H2. exists (V y0 f). fprops. am. Qed. Lemma lattice_finite_inf4: forall r f y, lattice r -> fgraph f -> is_finite_set (domain f) -> nonempty (domain f) -> sub (range f) (substrate r) -> gle r y (inf_graph r f) = (forall z, inc z (domain f) -> gle r y (V z f)). Proof. ir. uf inf_graph. rw lattice_finite_inf3. app iff_eq. ir. app H4. fprops. ir. srwi H5. nin H5. ee. rw H6. app H4. am. am. app finite_range. nin H2. exists (V y0 f). fprops. am. Qed. Lemma lattice_finite_sup5: forall r f, lattice r -> fgraph f -> is_finite_set (domain f) -> nonempty (domain f) -> sub (range f) (substrate r) -> inc (sup_graph r f) (substrate r). Proof. ir. assert (has_sup_graph r f). uf has_sup_graph. app lattice_finite_sup2. app finite_range. nin H2. exists (V y f). fprops. red in H. ee. cp (is_sup_graph_pr1 H H3 H4). rwi least_upper_bound_pr H6. ee. red in H6. ee. am. am. am. Qed. Lemma lattice_finite_inf5: forall r f, lattice r -> fgraph f -> is_finite_set (domain f) -> nonempty (domain f) -> sub (range f) (substrate r) -> inc (inf_graph r f) (substrate r). Proof. ir. assert (has_inf_graph r f). uf has_inf_graph. app lattice_finite_inf2. app finite_range. nin H2. exists (V y f). fprops. red in H. ee. cp (is_inf_graph_pr1 H H3 H4). rwi greatest_lower_bound_pr H6. ee. red in H6. ee. am. am. am. Qed. Lemma Exercice1_9: forall I1 I2 r f, lattice r -> fgraph f -> domain f = product I1 I2 -> is_finite_set I1 -> is_finite_set I2 -> nonempty I1 -> nonempty I2 -> sub (range f) (substrate r) -> gle r (sup_graph r (L I2 (fun j => inf_graph r (L I1 (fun i => V (J i j) f))))) (inf_graph r (L I1 (fun i => sup_graph r(L I2 (fun j => V (J i j) f))))). Proof. ir. assert (Ha: order r). red in H; ee; am. assert (Hb: forall i j, inc i I1 -> inc j I2 -> inc (V (J i j) f) (substrate r)). ir. app H6. app inc_V_range. rw H1. fprops. rw lattice_finite_sup4. ir. bwi H7. rw lattice_finite_inf4. ir. bwi H8. bw. apply order_transitivity with (V (J z0 z) f). am. set (fa:= (L I1 (fun i : Set => V (J i z) f))). assert (fgraph fa). uf fa. gprops. assert (is_finite_set (domain fa)). uf fa. bw. assert (nonempty (domain fa)). uf fa. bw. assert(sub (range fa) (substrate r)). uf fa. red. ir. srwi H12. nin H12. ee. bwi H12. bwi H13. rw H13. app Hb. am. gprops. cp (lattice_finite_inf4 (inf_graph r fa) H H9 H10 H11 H12). assert (gle r (inf_graph r fa) (inf_graph r fa)). order_tac. app lattice_finite_inf5. rwi H13 H14. assert (V (J z0 z) f = V z0 fa). uf fa. bw. tv. rw H15. app H14. uf fa. bw. set (fb:= L I2 (fun j : Set => V (J z0 j) f)). assert (fgraph fb). uf fb. gprops. assert (is_finite_set (domain fb)). uf fb. bw. assert (nonempty (domain fb)). uf fb. bw. assert(sub (range fb) (substrate r)). uf fb. red. ir. srwi H12. nin H12. ee. bwi H12. bwi H13. rw H13. app Hb. am. gprops. cp (lattice_finite_sup4 (sup_graph r fb) H H9 H10 H11 H12). assert (gle r (sup_graph r fb) (sup_graph r fb)). order_tac. app lattice_finite_sup5. rwi H13 H14. assert (V (J z0 z) f = V z fb). uf fb. bw. tv. rw H15. app H14. uf fb. bw. am. gprops. bw. bw. red. ir. rwi frange_inc_rw H8. nin H8. bwi H8. ee. rw H9. app lattice_finite_sup5. gprops. bw. bw. red. ir. srwi H10. nin H10. nin H10. bwi H10. bwi H11. rww H11. app Hb. am. gprops. ee;am. gprops. am. gprops. bw. bw. red. ir. srwi H7. nin H7. nin H7. bwi H7. bwi H8. rw H8. app lattice_finite_inf5. gprops. bw. bw. red. ir. srwi H9. nin H9. nin H9. bwi H9. bwi H10. rw H10. app Hb. am. gprops. ee;am. gprops. Qed. Lemma exercise1_10: forall r r' f, lattice r -> lattice r' -> is_function f -> substrate r = source f -> substrate r' = target f -> (increasing_fun f r r') = (forall x y, inc x (substrate r) -> inc y (substrate r) -> gle r' (W (inf r x y) f) (inf r' (W x f) (W y f))). Proof. ir. ap iff_eq. ir. red in H4; ee. destruct (lattice_inf_pr H H5 H6) as [Ha [Hb _ ]]. assert (inc (inf r x y) (source f)). wr H2. order_tac. rwi H2 H5; rwi H2 H6. cp (H11 _ _ Ha). cp (H11 _ _ Hb). simpl in H13; simpl in H14. cp (inc_W_target H4 H5). cp (inc_W_target H4 H6). wri H10 H15; wri H10 H16. cp (lattice_inf_pr H0 H15 H16). ee. app H19. ir. cp H; nin H. red. eee. nin H0;am. ir. rwi H2 H4. assert (inc x (source f)). wr H2. order_tac. assert (inc y (source f)). wr H2. order_tac. cp (H4 _ _ H8 H9). rwi (inf_comparable1 H H7) H10. cp (inc_W_target H1 H8). cp (inc_W_target H1 H9). wri H3 H11; wri H3 H12. destruct (lattice_inf_pr H0 H11 H12) as [_ [Hb _ ]]. nin H0. order_tac. Qed. Definition complete_lattice r := order r & forall X, sub X (substrate r) -> (has_supremum r X & has_infimum r X). Lemma exercise1_11a: forall r, complete_lattice r -> ((exists a, greatest_element r a) & (exists b, least_element r b)). Proof. ir. nin H. assert (sub emptyset (substrate r)). ap sub_emptyset_any. nin (H0 _ H1). nin H2. nin H3. rwi (least_upper_bound_emptyset x H) H2. rwi (greatest_lower_bound_emptyset x0 H) H3. split. exists x0; am. exists x;am. Qed. Lemma exercise1_11b: forall r, order r -> (forall X, sub X (substrate r) -> has_supremum r X) -> complete_lattice r. Proof. ir. red. split. am. ir. split. nin ( H0 _ H1). exists x. am. red. set (Z := (Zo (substrate r) (fun z => lower_bound r X z))). assert (sub Z (substrate r)). uf Z; app Z_sub. nin (H0 _ H2). exists x. rwi (least_upper_bound_pr x H H2) H3. nin H3. rw (greatest_lower_bound_pr x H H1). split. red. split. nin H3. am. ir. app H4. red. split. app H1. ir. ufi Z H6. Ztac. nin H8. app H9. ir. nin H3. app H6. uf Z. Ztac. nin H5. am. Qed. Lemma exercise1_11c: forall f g, axioms_product_order f g -> (forall i, inc i (domain f) -> complete_lattice (V i g)) -> complete_lattice (product_order f g). Proof. ir. app exercise1_11b. fprops. assert (Ha:substrate (product_order f g) = productb f). app substrate_product_order. assert (Hb:=H). red in H; ee. ir. set (Xi := fun i=> (image_by_fun (pr_i f i) X)). assert (forall i, inc i(domain f) -> sub (Xi i) (substrate (V i g))). ir. rw H4. uf Xi. assert (target (pr_i f i) = (V i f)). uf pr_i. aw. wr H7. uf image_by_fun. assert (is_function (pr_i f i)). app function_pri. apply sub_trans with (range (graph (pr_i f i))). app sub_image_by_graph. fprops. app range_correspondence. nin H8; am. am. set (v:= L (domain f) (fun i => supremum (V i g)(Xi i))). assert (forall i, inc i(domain f) -> least_upper_bound (V i g) (Xi i)(V i v)). ir. nin (H0 _ H7). cp (H6 _ H7). nin (H9 _ H10). uf v. bw. app supremum_pr1. assert (inc v (substrate (product_order f g))). rw Ha. rw productb_pr. ee. uf v. gprops. uf v. bw. ir. ufi v H8. bwi H8. nin (H7 _ H8). awi H9. Ztac. wrr H4. app H3. app Z_sub. am. exists v. rw least_upper_bound_pr. split. red. split. am. ir. rww related_product_order. rwi Ha H5; rwi Ha H8. ee. app H5. am. ir. cp (H7 _ H10). rwi least_upper_bound_pr H11. nin H11. red in H11. nin H11. app H13. uf Xi. uf image_by_fun. assert (is_function (pr_i f i)). app function_pri. aw. exists y. split. am. assert (inc y (productb f)). app H5. wr (W_pri H H10 H15). app defined_lem. uf pr_i. aw. fprops. app H6. ir. nin H9. rww related_product_order. eee. ir. cp (H7 _ H11). rwi least_upper_bound_pr H12. nin H12. app H13. red. split. rw H4. rwi Ha H9. rwi productb_pr H9. ee. app H15. ue. am. am. ir. ufi Xi H14. awi H14. ufi image_by_fun H14. awi H14. nin H14. ee. cp (H10 _ H14). rwi related_product_order H16. ee. assert (y = V i x). wr H15. wrr (W_pri (f:=f)). ue. app H18. am. app function_pri. uf pr_i. aw. ue. assert (is_function (pr_i f i)). app function_pri. app H3. app H6. app order_product_order. am. Qed. Lemma exercise1_11d: forall f g, axioms_product_order f g -> complete_lattice (product_order f g) -> (forall i, inc i (domain f) -> complete_lattice (V i g)). Proof. ir. cp (exercise1_11a H0). nin H2. clear H3. nin H2. nin H2. clear H3. awi H2; try am. assert (Ha:=H). nin H. ee. app exercise1_11b. app H5. ir. rwi productb_pr H2. ee. set (Y:= L(domain f) (fun j=> Yo(j=i) X (singleton (V j x)))). assert (sub (productb Y) (substrate (product_order f g))). aw. app productb_monotone1. uf Y. gprops. uf Y. bw. ir. uf Y. ufi Y H10. bwi H10. bw. nin (equal_or_not i0 i). rw Y_if_rw. wrr H6. ue. am. rw Y_if_not_rw. red. ir. rw (singleton_eq H12). app H9. ue. am. nin H0. cp (H11 _ H10). nin H12. nin H12. rwi least_upper_bound_pr H12. ee. red in H12. ee. awi H12. assert (W x0 (pr_i f i) = V i x0). rww W_pri. ee. exists (W x0 (pr_i f i)). rw least_upper_bound_pr. split. red. ee. rww H6. rw H16. rwi productb_pr H12. ee. app H18. rww H17. am. ir. assert (inc (L(domain f)(fun j=> (Yo (j = i) y (V j x)))) (productb Y)). rw productb_pr. bw. ee. gprops. uf Y. bw. ir. bw. nin (equal_or_not i0 i). rw H19. rw Y_if_rw. uf Y. bw. rw Y_if_rw. am. tv. tv. uf Y. bw. rw Y_if_not_rw. rw Y_if_not_rw. fprops. am. am. uf Y. gprops. cp (H15 _ H18). rwi related_product_order H19. ee. cp (H21 _ H1). bwi H22. rwi Y_if_rw H22. tv. rw H16. am. tv. am. am. ir. set (w:= (L(domain f)(fun j=> (Yo (j = i) z (V j x))))). assert (inc w (productb f)). uf w. rw productb_pr. bw. ee. gprops. tv. ir. bw. nin (equal_or_not i0 i). rw H19. rw Y_if_rw. red in H17. nin H17. wrr H6. tv. rw Y_if_not_rw. app H9. rww H8. am. am. assert (upper_bound (product_order f g) (productb Y) w). red. split. aw. ir. rw related_product_order. ee. awi H10. app H10. am. am. ir. uf w. bw. rwi productb_pr H19. ee. rwi H21 H22. ufi Y H22. bwi H22. cp (H22 _ H20). bwi H23. nin (equal_or_not i0 i). rw Y_if_rw. rwi Y_if_rw H23. red in H17. ee. rw H24. app H25. wrr H24. am. am. rw Y_if_not_rw. rwi Y_if_not_rw H23. rw (singleton_eq H23). wr order_reflexivity. rww H6. app H9. rww H8. app H5. am. am. am. uf Y; gprops. am. cp (H14 _ H19). rwi related_product_order H20. ee. rw H16. cp (H22 _ H1). ufi w H23. bwi H23. rwi Y_if_rw H23. am. tv. am. am. app H5. am. am. am. am. am. Qed. Definition greatest_induced r X x := greatest_element (induced_order r X) x. Definition least_induced r X x := least_element (induced_order r X) x. Lemma exercise1_11e: forall r f g, ordinal_sum_axioms1 r f g-> complete_lattice (ordinal_sum r f g) = (complete_lattice r & (forall j, sub j (substrate r) -> ~ (exists u, greatest_induced r j u) -> exists v, least_element (V (supremum r j) g) v) & (forall i x, inc i (substrate r) -> sub x (substrate (V i g)) -> (exists u, upper_bound (V i g) x u) -> nonempty x -> (exists u, least_upper_bound (V i g) x u)) & (forall i, inc i (substrate r) -> ~ (exists u, greatest_element (V i g) u ) -> exists v, least_induced r (Zo (substrate r) (fun j => glt r i j)) v & exists w, least_element (V v g) w)). Proof. ir. set (E:= substrate r). set (F:= disjoint_union f). nin H. rename H0 into Hb. assert (Hc:order (ordinal_sum r f g)). app order_ordinal_sum. assert (Ha: substrate (ordinal_sum r f g) = F). rww substrate_ordinal_sum. app iff_eq. ir. cp H. assert (Hd:forall i, inc i (domain f) -> exists y, inc y (V i f) & inc (J y i) (substrate (ordinal_sum r f g))). ir. nin (Hb _ H2). exists y. split. am. cp (inc_disjoint_union H2 H3). aw. set (k:= fun i => choose(fun y => inc y (V i f) & inc (J y i) (substrate (ordinal_sum r f g)))). assert (forall i, inc i (domain f) -> (inc (k i) (V i f)&inc (J (k i) i)F)). ir. wr Ha. uf k. app choose_pr. app Hd. nin H1. nin H0. assert (forall j, sub j E -> exists x, least_upper_bound (ordinal_sum r f g) (fun_image j (fun i => J (k i) i)) x & inc x F & least_upper_bound r j (Q x)). ir. ee. set (Y:= fun_image j (fun i=> J (k i) i)). assert (sub Y F). uf Y. red. ir. awi H11. nin H11. nin H11. assert (inc x0 (domain f)). wr H3. app H5. nin (H2 _ H13). ue. rwi Ha H4. nin (H4 _ H11). nin H12. exists x. split. am. awi H12. nin H12. assert (inc x F). red in H12. eee. eee. cp (du_index_pr H15). ee. aw. split. red. eee. ir. red in H12. ee. assert (inc (J (k y) y) Y). uf Y. aw. exists y. auto. cp (H20 _ H21). cp (related_ordinal_sum_order_id H H22). awi H23. am. ir. red in H19. ee. assert (inc (J (k z) z) F). rwi H3 H19. nin (H2 _ H19). am. assert (upper_bound (ordinal_sum r f g) Y (J (k z) z)). red. rw Ha. eee. ir. aw. eee. ufi Y H22. awi H22. nin H22. nin H22. cp (H20 _ H22). wr H23. aw. nin (equal_or_not x0 z). rw H25. right. split. tv. wrr order_reflexivity. rwi H3 H19. nin (H2 _ H19). rww H10. app H9. ue. uf glt. au. cp (H14 _ H22). cp (related_ordinal_sum_order_id H H23). awi H24. am. am. rww Ha. ee. app exercise1_11b. ir. nin (H5 _ H11). exists (Q x). ee. am. ir. nin (H5 _ H11). ee. assert (least_upper_bound r j ((supremum r j))). uf supremum. app choose_pr. exists (Q x). am. rw (supremum_unique H1 H16 H15). clear H16. assert (~ inc (Q x) j). red. ir. elim H12. exists (Q x). red. split. aw. aw. ir. aw. awi H15. ee. nin H15. app H19. am. am. exists (P x). awi H13. ee. cp (du_index_pr H14). ee. red. ee. rww H10. ir. rwi H10 H21. assert (inc (J x0 (Q x)) F). uf F. app inc_disjoint_union. assert (upper_bound (ordinal_sum r f g) (fun_image j (fun i : Set => J (k i) i)) (J x0 (Q x))). red. aw. ee. am. ir. aw. ee. awi H23. nin H23. ee. wr H24. app inc_disjoint_union. wr H3. app H11. assert (inc x1 (domain f)). wr H3. app H11. nin (H2 _ H25). am. am. awi H23. nin H23. nin H23. wr H24. aw. assert (x1 <> Q x). red. ir. elim H16. wrr H25. awi H15. nin H15. nin H15. left. split. app H27. am. am. am. cp (H17 _ H23). awi H24. ee. nin H26. nin H26. elim H27. tv. nin H26. am. am. am. am. red. ir. awi H17. rw Ha. nin H17. nin H17. wr H18. assert (inc x1 (domain f)). wr H3. app H11. nin (H2 _ H19). am. ir. assert (inc i (domain f)). wrr H3. rwi H10 H12. nin H13. assert (inc (J x0 i) F). uf F. app inc_disjoint_union. red in H13. ee. wrr H10. set (X := product x (singleton i)). assert (sub X F). red. ir. ufi X H17. awi H17. ee. assert (J (P x1) (Q x1)= x1). app pair_recov. wr H20. uf F. app inc_disjoint_union. rw H19. am. rw H19. app H12. rwi Ha H4. nin (H4 _ H17). nin H18. awi H18. ee. assert (upper_bound (ordinal_sum r f g) X (J x0 i)). red. ee. rw Ha. am. ir. aw. ee. app H17. am. ufi X H21. awi H21. ee. rw H23. right. split. tv. red in H13. ee. app H24. assert (Q x1 = i). cp (H20 _ H21). cp (related_ordinal_sum_order_id H H22). awi H23. nin H14. assert (inc (J y i) X). uf X. aw. ee. fprops. am. fprops. red in H18. ee. cp (H25 _ H24). cp (related_ordinal_sum_order_id H H26). awi H27. order_tac. exists (P x1). aw. nin H18. rwi Ha H18. cp (du_index_pr H18). ee. red. ee. rww H10. wrr H22. ir. assert (inc (J y i) X). uf X. aw. ee. fprops. am. fprops. cp (H23 _ H28). awi H29. ee. nin H31. red in H31. ee. elim H32. sy. am. nin H31. am. am. ir. red in H27. ee. rwi H10 H27. assert (inc (J z i) F). uf F. app inc_disjoint_union. assert (upper_bound (ordinal_sum r f g) X (J z i)). red. ee. rww Ha. ir. aw. ee. app H17. am. ufi X H30. awi H30. ee. rw H32. right. split. tv. app H28. cp (H20 _ H30). awi H31. ee. rwi H22 H33. nin H33. nin H33. elim H34. tv. nin H33. am. am. am. app H9. rww H10. am. rw Ha. am. am. ee. ir. set (X:= product (substrate (V i g)) (singleton i)). assert (sub X F). red. ir. uf F. ufi X H13. awi H13. ee. assert (J (P x) (Q x)= x). app pair_recov. wr H16. app inc_disjoint_union. rw H15. wrr H3. rww H15. wrr H10. wrr H3. rwi Ha H4. nin (H4 _ H13). nin H14. awi H14. nin H14. ufi E H11. rwi H3 H11. nin (Hb _ H11). assert (inc (J y i) X). uf X. aw. ee. fprops. rww H10. fprops. nin H14. set (Ii:=Zo E (fun j : Set => glt r i j)). assert (inc (Q x) Ii). cp (H19 _ H18). awi H20. ee. uf Ii. nin H22. Ztac. nin H22. uf E. order_tac. nin H22. elim H12. exists (P x). red. split. uf E. order_tac. ir. assert (inc (J x0 i) X). uf X. aw. eee. cp (H19 _ H25). awi H26. ee. nin H28. nin H28. elim H29. am. nin H28. am. am. am. exists (Q x). split. red. red. aw. split. am. ir. aw. ufi Ii H21. Ztac. ufi E H22. rwi H3 H22. nin (H2 _ H22). assert (upper_bound (ordinal_sum r f g) X (J (k x0) x0)). red. rw Ha. split. am. ir. aw. ee. app H13. am. left. ufi X H26. awi H26. ee. rww H28. cp (H16 _ H26). cp (related_ordinal_sum_order_id H H27). awi H28. am. uf Ii. app Z_sub. exists (P x). red. split. rwi Ha H14. cp (du_index_pr H14). ee. rww H10. rw H10. ir. assert (inc (J x0 (Q x)) F). uf F. app inc_disjoint_union. ufi Ii H20. Ztac. wr H3. am. assert (upper_bound (ordinal_sum r f g) X (J x0 (Q x))). red. rw Ha. split. am. ir. aw. ee. app H13. am. left. ufi X H23. awi H23. ee. rww H25. ufi Ii H20. Ztac. am. cp (H16 _ H23). awi H24. ee. nin H26. nin H26. elim H27. tv. nin H26. aw. am. ufi Ii H20. Ztac. wrr H3. am. rww Ha. ir. ee. app exercise1_11b. ir. rwi Ha H4. cp H. red in H. ee. set (j:= Zo E (fun i=> exists x, inc x X & i = Q x)). assert (sub j E). uf j. app Z_sub. nin H0. nin (H13 _ H12). clear H15. cp (supremum_pr1 H H12 H14). nin (p_or_not_p (exists u, greatest_induced r j u)). nin H16. rename x into k. red in H16. red in H16. awi H16; try am. ee. assert (Hd:forall z, upper_bound (ordinal_sum r f g) X z -> gle r k (Q z)). ir. red in H18. nin H18. rwi Ha H18. ufi j H16. Ztac. nin H21. nin H21. cp (H19 _ H21). rw H22. app (related_ordinal_sum_order_id H5 H23). assert (He: inc k (domain f)). wr H6. app H12. set (Xj:= Zo (substrate (V k g)) (fun y=> exists x, inc x X & y = P x & k = Q x)). assert (Hf:nonempty Xj). ufi j H16. Ztac. clear H16. nin H19. exists (P x). uf Xj. Ztac. ee. cp (du_index_pr (H4 _ H16)). ee. rw H11. rww H19. rww H19. exists x. eee. nin (p_or_not_p (exists u, upper_bound (V k g) Xj u)). ir. assert (sub Xj (substrate (V k g))). uf Xj. app Z_sub. nin (H2 _ _ (H12 _ H16) H19 H18). rwi least_upper_bound_pr H20. ee. assert (inc (J x k) F). uf F. app inc_disjoint_union. nin H20. wrr H11. exists (J x k). rw least_upper_bound_pr. split. red. aw. split. am. ir. aw. ee. app H4. am. cp (du_index_pr (H4 _ H23)). ee. assert (inc (Q y) j). uf j. Ztac. ee. uf E. rww H6. exists y. auto. cp (H17 _ H27). awi H28. nin (equal_or_not (Q y) k). right. split. am. nin H20. rw H29. app H30. uf Xj. Ztac. rww H11. wrr H29. exists y. auto. uf glt. au. am. am. ir. aw. cp (Hd _ H23). nin H23. rwi Ha H23. fold F. eee. nin (equal_or_not k (Q z)). right. split. am. app H21. red. nin (du_index_pr H23). nin H28. rw H11. split. rww H26. ir. ufi Xj H30. Ztac. nin H32. ee. cp (H25 _ H32). awi H35. ee. wri H34 H37. wri H26 H37. nin H37. nin H37. elim H38. tv. rw H33. nin H37. am. am. wr H6. app H12. uf glt. au. am. rw Ha. am. app H10. am. am. ir. nin (p_or_not_p (exists u, greatest_element (V k g) u)). ir. nin H19. elim H18. exists x. nin H19. red. split. am. ir. app H20. ufi Xj H21. Ztac. am. ir. nin (H3 _ (H12 _ H16) H19). nin H20. nin H21. red in H21. rwi H11 H21. nin H21. nin H20. awi H20. Ztac. clear H20. awi H23. assert (inc (J x0 x) F). uf F. app inc_disjoint_union. wrr H6. exists (J x0 x). rw least_upper_bound_pr. split. red. rw Ha. split. am. ir. aw. split. app H4. split. am. assert (inc (Q y) j). uf j. nin (du_index_pr (H4 _ H26)). ee. Ztac. uf E. rw H6. am. exists y. auto. cp (H17 _ H27). awi H28. left. nin H25. cp (order_transitivity H H28 H25). split. am. red. ir. rwi H31 H28. elim H29. order_tac. am. am. ir. cp (Hd _ H26). nin H26. rwi Ha H26. nin (du_index_pr H26). nin H30. aw. ee. am. am. assert (inc (Q z) (Zo E (fun j : Set => glt r k j))). Ztac. uf E. rww H6. split. am. red. ir. elim H18. exists (P z). split. rw H32. rww H11. ir. ufi Xj H33. Ztac. nin H35. ee. cp (H28 _ H35). awi H38. ee. nin H40. nin H40. elim H41. wr H32; wr H37; tv. nin H40. wri H37 H41. wri H36 H41. am. am. cp (H23 _ H32). awi H33. nin (equal_or_not x (Q z)). right. split. am. wri H34 H30. cp (H22 _ H30). awi H35. am. uf glt. au. clear H32. Ztac. am. am. rww Ha. am. app Z_sub. am. app Z_sub. wr H6. red in H20. red in H20. nin H20. awi H20. Ztac. am. am. app Z_sub. nin (H1 _ H12 H16). set (a:= supremum r j). assert (inc a (domain f)). nin H15. awi H15. Ztac. wrr H6. am. app Z_sub. set (b:= J x a). assert (inc b F). uf F. uf b. app inc_disjoint_union. wrr H11. nin H17; ee; am. exists b. rw least_upper_bound_pr. split. red. split. rww Ha. ir. aw. ee. app H4. am. left. uf b. aw. assert (inc (Q y) j). uf j. assert (inc y F). app H4. cp (du_index_pr H21). Ztac. uf E. rww H6. eee. exists y. eee. rwi least_upper_bound_pr H15. ee. nin H15. split. app H23. red. ir. rwi H24 H21. elim H16. exists a. red. red. split. aw. aw. ir. aw. app H23. am. am. aw. ir. nin H20. ee. rwi Ha H20. cp (du_index_pr H20). rwi least_upper_bound_pr H15. ee. assert (upper_bound r j (Q z)). red. ir. ee. rww H6. ir. ufi j H26. Ztac. nin H28. ee. rw H29. cp (H21 _ H28). ap (related_ordinal_sum_order_id H5 H30). cp (H25 _ H26). aw. eee. nin (equal_or_not (Q b) (Q z)). right. eee. uf b. aw. nin H17. app H29. rww H11. wri H28 H23. ufi b H23. awi H23. am. left. split. uf b. aw. am. am. am. am. rww Ha. Qed. Lemma exercise1_11f: forall r r', order r -> order r' -> nonempty (substrate r) -> complete_lattice (increasing_mappings_order r r') -> complete_lattice r'. Proof. ir. app exercise1_11b. ir. red. set (E:= substrate r). set (E':=substrate r'). set (Y:= Zo (substrate (increasing_mappings_order r r')) (fun f => exists y, exists Hy: inc y X, f = constant_function E E' y (H3 y Hy))). assert (sub Y (substrate (increasing_mappings_order r r'))). uf Y. app Z_sub. nin H2. nin (H5 _ H4). nin H6. awi H6. nin H6. nin H6. rwi imo_substrate H6. cp H1. nin H1. rwi set_of_increasing_mappings_pr H6. nin H6. ee. set (u:= W y x). assert (inc u E'). uf u. uf E'. wr H12. app inc_W_target. rww H11. exists u. aw. split. red. ee. am. ir. set (f:= constant_function E E' y0 (H3 y0 H15)). assert (inc f Y). uf Y. Ztac. rw imo_substrate. uf f. uf E; uf E'. app constant_increasing. am. am. exists y0. exists H15. tv. cp (H9 _ H16). rwi imo_pr H17. ee. red in H19. ee. cp (H25 _ H1). ufi f H26. awi H26. rwi W_constant H26. am. am. am. am. ir. red in H15. ee. set (f:= (constant_function (substrate r) (substrate r') z H15)). assert (inc f (set_of_increasing_mappings r r')). uf f. app constant_increasing. assert (upper_bound (increasing_mappings_order r r') Y f). red. ee. rww imo_substrate. ir. ufi Y H18. Ztac. nin H20. nin H20. rw H20. uf f. uf E. uf E'. wrr constant_increasing1. app H16. cp (H8 _ H18). rwi imo_pr H19. ee. red in H21. ee. cp (H27 _ H1). ufi f H28. awi H28. rwi W_constant H28. am. am. am. am. am. am. am. am. am. uf Y. app Z_sub. Qed. Lemma exercise1_11g: forall r r', order r -> order r' -> complete_lattice r' -> complete_lattice (increasing_mappings_order r r'). Proof. ir. app exercise1_11b. ir. app imo_order. rww imo_substrate. ir. set (E:=substrate r). set (E':=substrate r'). set (img:= fun x=> fun_image X (fun f => W x f)). assert (forall x, inc x E -> sub (img x) E'). ir. uf img. red. ir. awi H4. nin H4. ee. wr H5. cp (H2 _ H4). awi H6. ee. uf E'. wr H8. aw. app inc_W_target. rww H7. am. am. set (f:= fun x=> supremum r' (img x)). assert (forall x, inc x E -> least_upper_bound r' (img x) (f x)). ir. uf f. app supremum_pr1. app H3. nin H1. nin (H5 _ (H3 _ H4)). am. assert (transf_axioms f E E'). red. ee. ir. cp (H4 _ H5). awi H6. nin H6. red in H6. nin H6. am. am. app H3. assert (is_function (BL f E E')). app bl_function. assert (inc (BL f E E') (set_of_increasing_mappings r r')). aw. eee. red. eee. aw. aw. ir. assert (inc x E). uf E; order_tac. assert (inc y E). uf E; order_tac. rww W_bl_function. rww W_bl_function. cp (H4 _ H8). cp (H4 _ H9). awi H10. awi H11. ee. ap H13. red. red in H11. ee. am. ir. ufi img H15. awi H15. nin H15. ee. set (t:= W y x0). assert (inc t (img y)). uf img. aw. exists x0. split. am. tv. cp (H14 _ H17). cp (H2 _ H15). awi H19. ee. red in H22. ee. cp (H27 _ _ H7). wr H16. order_tac. ufi t H18. am. am. am. app H3. am. app H3. exists (BL f E E'). aw. split. red. split. rww imo_substrate. ir. rww imo_pr. ee. app H2. am. cp (H2 _ H8). awii H9. nin H9. ee. red. eee. aw. aw. ir. rww W_bl_function. cp (H4 _ H13). awi H14. ee. red in H14. ee. app H16. uf img. aw. ex_tac. am. app H3. ir. red in H8. ee. rww imo_pr. rwii imo_substrate H8. eee. awi H8. ee. red. eee. aw. aw. ir. rw W_bl_function. cp (H4 _ H13). awi H14. ee. app H15. red. split. wr H11. app inc_W_target. rww H10. ir. ufi img H16. awi H16. nin H16. nin H16. cp (H9 _ H16). rwi imo_pr H18. ee. red in H20. ee. cp (H26 _ H13). ue. am. am. am. app H3. am. am. am. am. app imo_order. rww imo_substrate. Qed. Lemma Exercise1_12: forall E f, is_function f -> source f = E -> target f = E -> complete_lattice (inclusion_suborder (Zo (powerset E) (fun X => sub (image_by_fun f X) X))). Proof. ir. set (F:=Zo (powerset E) (fun X : Set => sub (image_by_fun f X) X)). assert (order (inclusion_suborder F)). ap subinclusion_is_order. red. split. am. ir. rwi substrate_subinclusion_order H3. assert (sub F (powerset E)). uf F. ap Z_sub. assert (sub X (powerset E)). apply sub_trans with F. am. am. assert (Ha:sub (union X) E). red. ir. nin (union_exists H6). nin H7. cp (H5 _ H8). rwi powerset_inc_rw H9. app H9. set (v:= Yo (nonempty X) (intersection X) E). assert (Hb: sub v E). uf v. nin (p_or_not_p (nonempty X)). rww Y_if_rw. nin H6. cp (intersection_sub H6). apply sub_trans with y. am. app powerset_sub. app H5. rww Y_if_not_rw. fprops. assert (inc v F). uf F. Ztac. app powerset_inc. uf v. nin (p_or_not_p (nonempty X)). rww Y_if_rw. cp H6. nin H6. red. ir. awi H8. nin H8. nin H8. app intersection_inc. ir. cp (H3 _ H10). ufi F H11. Ztac. app H13. aw. exists x0. split. app (intersection_forall H8 H10). am. rw H0. app powerset_sub. am. rw H0. assert (v = intersection X). uf v. rww Y_if_rw. wrr H9. rww Y_if_not_rw. red. ir. awi H7. nin H7. ee. wr H8. wri H0 H7. wr H1. fprops. am. rw H0. fprops. assert (inc (union X) F). uf F. Ztac. app powerset_inc. red. ir. awi H7. nin H7. nin H7. nin (union_exists H7). nin H9. cp (H3 _ H10). ufi F H11. Ztac. apply union_inc with x1. app H13. aw. exists x0. split. am. am. rww H0. app powerset_sub. am. am. rww H0. split. red. exists (union X). app (union_is_sup1 H4 H3 H7). red. exists v. app (intersection_is_inf1 H4 H3 H6). Qed. Definition is_closure f r := increasing_fun f r r & (forall x, inc x (substrate r) -> gle r x (W x f)) & (forall x, inc x (substrate r) -> W (W x f) f = W x f). Definition set_of_invariants f := Zo (source f) (fun x => W x f = x). Definition set_of_upper_bounds F r x := Zo F (fun y => gle r x y). Lemma Exercise1_13a: forall f r x, is_closure f r -> let F := set_of_invariants f in inc x (source f) -> least_element (induced_order r (set_of_upper_bounds F r x)) (W x f). Proof. ir. red in H. ee. red in H. ee. assert (inc (W x f) F). uf F. uf set_of_invariants. Ztac. wr H5; rw H6; fprops. ap H2. rww H5. assert (sub (set_of_upper_bounds F r x) (substrate r)). uf set_of_upper_bounds. apply sub_trans with F. app Z_sub. uf F. uf set_of_invariants. wr H5. app Z_sub. uf least_element. aw. assert (inc (W x f) (set_of_upper_bounds F r x)). uf set_of_upper_bounds. Ztac. app H1. rww H5. split. am. ir. aw. ufi set_of_upper_bounds H11. Ztac. ufi F H12. ufi set_of_invariants H12. Ztac. wr H15. ap (H7 _ _ H13). Qed. Lemma Exercise1_13b: forall r G, order r -> sub G (substrate r) -> let g:= fun x => choose (fun y => least_element (induced_order r (set_of_upper_bounds G r x)) y) in (forall x, inc x (substrate r) -> exists y, least_element (induced_order r (set_of_upper_bounds G r x)) y) -> (is_closure (BL g (substrate r) (substrate r)) r & (G = set_of_invariants (BL g (substrate r) (substrate r)))). Proof. ir. assert (forall x, inc x (substrate r) -> least_element (induced_order r (set_of_upper_bounds G r x)) (g x)). ir. uf g. app choose_pr. app H1. clear H1. set (E:= substrate r) in *. assert (Ha:forall x, inc x E -> sub (set_of_upper_bounds G r x) E). ir. uf set_of_upper_bounds. apply sub_trans with G. app Z_sub. am. assert (Hb:forall x, inc x E -> (inc (g x) (set_of_upper_bounds G r x))). ir. cp (H2 _ H1). red in H3. ee. awi H3. am. am. app Ha. assert (Hc:forall x y, inc x E -> inc y (set_of_upper_bounds G r x) -> gle r (g x) y). ir. cp (H2 _ H1). red in H4. ee. awi H5. cp (H5 _ H3). awi H6. am. awi H4. am. am. app Ha. am. am. app Ha. assert (Hd:forall x, inc x E -> (inc (g x) E & inc (g x) G & gle r x (g x))). ir. cp (Hb _ H1). ufi set_of_upper_bounds H3. Ztac. ee. app H0. am. am. assert(He:forall x, inc x G -> (g x) = x). ir. assert (gle r (g x) x). app Hc. app H0. uf set_of_upper_bounds. Ztac. order_tac. app H0. assert (inc x E). app H0. cp (Hd _ H4). ee. order_tac. assert (transf_axioms g E E). red. ir. nin (Hd _ H1). am. assert (forall x, inc x E -> W x (BL g E E) = g x). ir. app W_bl_function. split. red. split. red. eee. app bl_function. aw. aw. ir. assert (inc y E). uf E. order_tac. assert (inc x E). uf E. order_tac. rw H3. rw H3. cp (Hd _ H5). ee. app Hc. uf set_of_upper_bounds. Ztac. order_tac. am. am. split. ir. rw H3. cp (Hd _ H4). ee; am. am. ir. set (y:= W x (BL g E E)). assert (y = g x). uf y. rww H3. cp (Hd _ H4). wri H5 H6. ee. rww H3. app He. uf set_of_invariants. set_extens. Ztac. aw. app H0. Ztac. rw H3. rww He. app H0. Ztac. awi H5. cp (Hd _ H5). rwi H3 H6. ee. wrr H6. am. Qed. Lemma Exercise1_13c: forall f r E, is_closure f r -> complete_lattice r -> let F := set_of_invariants f in sub E F -> nonempty E -> inc (infimum r E) F. Proof. ir. nin H0. nin H. ee. nin H. ee. assert (sub F (substrate r)). uf F. uf set_of_invariants. rw H8. app Z_sub. assert (sub E (substrate r)). apply sub_trans with F. am. am. nin (H3 _ H12). cp (infimum_pr1 H0 H12 H14). rwi greatest_lower_bound_pr H15. ee. set (y:= infimum r E) in *. assert (inc y (substrate r)). nin H15. am. assert (inc (W y f) (substrate r)). rw H9. app inc_W_target. wrr H8. assert (lower_bound r E (W y f)). red. split. am. ir. red in H15. ee. cp (H20 _ H19). cp (H10 _ _ H21). assert (inc y0 F). app H1. ufi F H23. ufi set_of_invariants H23. Ztac. wrr H25. cp (H16 _ H19). cp (H4 _ H17). uf F. uf set_of_invariants. Ztac. wrr H8. order_tac. am. am. Qed. Lemma Exercice1_14: forall A B R, let rho := fun X => Zo B (fun y => forall x,inc x X -> inc (J x y) R) in let sigma := fun Y => Zo A (fun x => forall y, inc y Y -> inc (J x y) R) in let fr:=BL rho (powerset A) (powerset B) in let fs:= BL sigma (powerset B) (powerset A) in let iA := inclusion_order A in let iB := inclusion_order B in sub R (product A B) -> ( decreasing_fun fr iA iB & decreasing_fun fs iB iA & is_closure (compose fs fr) iA & is_closure (compose fr fs) iB). Proof. ir. assert (transf_axioms rho (powerset A) (powerset B)). red. ir. uf rho. app powerset_inc. app Z_sub. assert (transf_axioms sigma (powerset B) (powerset A)). red. ir. uf sigma. app powerset_inc. app Z_sub. assert (is_function fr). uf fr. app bl_function. assert (is_function fs). uf fs. app bl_function. assert (composable fs fr). red. eee. uf fs. uf fr. aw. assert (composable fr fs). red. eee. uf fs. uf fr. aw. assert (is_function (compose fs fr)). app is_function_compose. assert (is_function (compose fr fs)). app is_function_compose. assert (forall u v, sub u v -> sub (rho v) (rho u)). ir. uf rho. red. ir. Ztac. clear H9. Ztac. assert (forall u v, sub u v -> sub (sigma v) (sigma u)). ir. uf sigma. red. ir. Ztac. clear H10. Ztac. assert (forall u v, sub u v -> sub (sigma (rho u)) (sigma (rho v))). ir. app H9. app H8. assert (forall u v, sub u v -> sub (rho (sigma u)) (rho (sigma v))). ir. app H8. app H9. assert (order iA). uf iA. app inclusion_is_order. assert (order iB). uf iB. app inclusion_is_order. assert (substrate iA = powerset A). uf iA. rww substrate_inclusion_order. assert (substrate iB = powerset B). uf iB. rww substrate_inclusion_order. split. red. eee. uf fr. aw. uf fr. aw. ir. ufi iA H16. awi H16. ee. uf fr. assert (inc x (powerset A)). app powerset_inc. assert (inc y (powerset A)). app powerset_inc. rww W_bl_function. rww W_bl_function. red. uf iB. aw. ufi iA H16. awi H16. ee. app powerset_sub. app H0. app powerset_sub. app H0. app H8. split. red. eee. uf fs. aw. uf fs. aw. ir. ufi iB H16. awi H16. ee. uf fs. assert (inc x (powerset B)). app powerset_inc. assert (inc y (powerset B)). app powerset_inc. rww W_bl_function. rww W_bl_function. red. uf iA. aw. ee. app powerset_sub. app H1. app powerset_sub. app H1. app H9. assert (Ha: forall x, sub x A -> sub x (sigma (rho x))). ir. red. ir. uf sigma. uf rho. Ztac. ir. Ztac. app H20. assert (Hb: forall x, sub x B -> sub x (rho (sigma x))). ir. red. ir. uf sigma. uf rho. Ztac. ir. Ztac. app H20. assert (increasing_fun (compose fs fr) iA iA). red. eee. uf fr. aw. uf fs. aw. ir. ufi iA H16. awi H16. ee. assert (inc x (powerset A)). app powerset_inc. assert (inc y (powerset A)). app powerset_inc. rww compose_W. rww compose_W. uf fr. aw. assert (inc (rho y) (powerset B)). app H0. assert (inc (rho x) (powerset B)). app H0. uf fs. aw. uf iA. aw. ee. app powerset_sub. app H1. app powerset_sub. app H1. app H10. uf fr. aw. uf fr. aw. assert (increasing_fun (compose fr fs) iB iB). red. eee. uf fs. aw. uf fr. aw. ir. ufi iB H17. awi H17. ee. assert (inc x (powerset B)). app powerset_inc. assert (inc y (powerset B)). app powerset_inc. rww compose_W. rww compose_W. uf fs. aw. assert (inc (sigma y) (powerset A)). app H1. assert (inc (sigma x) (powerset A)). app H1. uf fr. aw. uf iB. aw. ee. app powerset_sub. app H0. app powerset_sub. app H0. app H11. uf fs. aw. uf fs. aw. split. red. ee. am. ir. rww compose_W. uf fr. rwi H14 H18. rww W_bl_function. uf fs. assert (inc (rho x) (powerset B)). app H0. rww W_bl_function. uf iA. aw. ee. app powerset_sub. red. ir. ufi sigma H20. Ztac. am. app Ha. app powerset_sub. uf fr. aw. wrr H14. ir. set (y:= W x (compose fs fr)). rwi H14 H18. assert (y = sigma (rho x)). uf y. rw compose_W. uf fr. uf fs. aw. aw. app H0. am. uf fr. aw. assert (inc y (powerset A)). rw H19. app H1. app H0. assert (rho y = rho x). rw H19. app extensionality. app H8. app Ha. app powerset_sub. app Hb. app powerset_sub. app H0. assert (sigma (rho y) = sigma (rho x)). rww H21. rw compose_W. uf fr. uf fs. aw. sy. rww H22. aw. app H0. am. uf fr. aw. red. ee. am. ir. rww compose_W. uf fs. rwi H15 H18. rww W_bl_function. uf fr. assert (inc (sigma x) (powerset A)). app H1. rww W_bl_function. uf iB. aw. ee. app powerset_sub. red. ir. ufi rho H20. Ztac. am. app Hb. app powerset_sub. uf fs. aw. wrr H15. ir. set (y:= W x (compose fr fs)). rwi H15 H18. assert (y = rho (sigma x)). uf y. rw compose_W. uf fr. uf fs. aw. aw. app H1. am. uf fs. aw. assert (inc y (powerset B)). rw H19. app H0. app H1. assert (sigma y = sigma x). rw H19. app extensionality. app H9. app Hb. app powerset_sub. app Ha. app powerset_sub. app H1. assert (rho (sigma y) = rho (sigma x)). rww H21. rw compose_W. uf fr. uf fs. aw. sy. rww H22. aw. app H1. am. uf fs. aw. Qed. (* Exercice 15 *) Definition set_of_up_bounds r X := Zo (substrate r)(fun z => upper_bound r X z). Definition set_of_lo_bounds r X := Zo (substrate r)(fun z => lower_bound r X z). Definition set_of_uplo_bounds r X := set_of_lo_bounds r (set_of_up_bounds r X). Definition completion r:= Zo (powerset(substrate r)) (fun z => z = set_of_uplo_bounds r z). Definition completion_order r := inclusion_suborder (completion r). Lemma Exercise1_15a1: forall r A B, sub A B -> sub (set_of_up_bounds r B) (set_of_up_bounds r A). Proof. ir. uf set_of_up_bounds. red. ir. Ztac. clear H0. Ztac. nin H2. split. am. ir. ap (H2 _ (H _ H3)). Qed. Lemma Exercise1_15a2: forall r A B, sub A B -> sub (set_of_lo_bounds r B) (set_of_lo_bounds r A). Proof. ir. uf set_of_lo_bounds. red. ir. Ztac. clear H0. Ztac. nin H2. split. am. ir. ap (H2 _ (H _ H3)). Qed. Lemma Exercise1_15a3: forall r A B, sub A B -> sub (set_of_uplo_bounds r A) (set_of_uplo_bounds r B). Proof. ir. uf set_of_uplo_bounds. app Exercise1_15a2. app Exercise1_15a1. Qed. Lemma Exercise1_15a4: forall r A, sub A (substrate r) -> sub A (set_of_uplo_bounds r A). Proof. ir. red. ir. uf set_of_uplo_bounds. uf set_of_lo_bounds. Ztac. split. app H. ir. ufi set_of_up_bounds H1. Ztac. nin H3. app H4. Qed. Lemma Exercise1_15a5: forall r A, sub A (substrate r) -> set_of_lo_bounds r (set_of_up_bounds r (set_of_lo_bounds r A)) = (set_of_lo_bounds r A). Proof. ir. app extensionality. app Exercise1_15a2. uf set_of_up_bounds. red. ir. Ztac. split. app H. ir. ufi set_of_lo_bounds H1. Ztac. nin H3. app H4. assert (sub (set_of_lo_bounds r A) (substrate r)). uf set_of_lo_bounds. app Z_sub. ap (Exercise1_15a4 H0). Qed. Lemma Exercise1_15a6: forall r A, sub A (substrate r) -> set_of_uplo_bounds r (set_of_uplo_bounds r A) = (set_of_uplo_bounds r A). Proof. ir. uf set_of_uplo_bounds. rww Exercise1_15a5. uf set_of_up_bounds. app Z_sub. Qed. Lemma Exercise1_15a7: forall r A, sub A (substrate r) -> inc (set_of_uplo_bounds r A) (completion r). Proof. ir. uf completion. Ztac. uf set_of_uplo_bounds. uf set_of_lo_bounds. app powerset_inc. app Z_sub. rww Exercise1_15a6. Qed. Lemma Exercise1_15a8: forall r A, sub A (substrate r) -> inc (set_of_lo_bounds r A) (completion r). Proof. ir. uf completion. Ztac. uf set_of_lo_bounds. app powerset_inc. app Z_sub. uf set_of_uplo_bounds. rww Exercise1_15a5. Qed. Lemma Exercise1_15a9: forall r x y, order r -> inc x (substrate r) -> inc y (substrate r) -> (set_of_lo_bounds r (singleton x) = set_of_lo_bounds r (singleton y)) -> x = y. Proof. ir. assert (inc x (set_of_lo_bounds r (singleton y))). wr H2. uf set_of_lo_bounds. Ztac. split. am. ir. awi H3. rw H3. order_tac. assert (inc y (set_of_lo_bounds r (singleton x))). rw H2. uf set_of_lo_bounds. Ztac. split. am. ir. awi H4. rw H4. order_tac. ufi set_of_lo_bounds H3. ufi set_of_lo_bounds H4. Ztac. clear H4. Ztac. nin H6. nin H7. assert (inc x (singleton x)). fprops. cp (H8 _ H10). assert (inc y (singleton y)). fprops. cp (H9 _ H12). order_tac. Qed. Lemma Exercise1_15a10: forall r e, order r -> least_element r e -> least_element (completion_order r) (singleton e). Proof. ir. assert (Ha:order (completion_order r)). uf completion_order. fprops. assert (Hb: substrate (completion_order r) = completion r). uf completion_order. aw. assert (sub (singleton e) (substrate r)). red. ir. awi H1. rw H1. nin H0. am. assert(inc (singleton e) (completion r)). uf completion. Ztac. app powerset_inc. app extensionality. app Exercise1_15a4. red. ir. ufi set_of_uplo_bounds H2. ufi set_of_lo_bounds H2. Ztac. clear H2. nin H4. assert (least_element r x). nin H0. split. am. ir. app H4. uf set_of_up_bounds. Ztac. split. am. ir. awi H7. rw H7. app H5. rw (unique_least H H5 H0). fprops. split. ue. rw Hb. ir. uf completion_order. aw. eee. red. ir. awi H4. rw H4. ufi completion H3. Ztac. clear H3. rw H6. uf set_of_uplo_bounds. uf set_of_lo_bounds. Ztac. nin H0; am. split. nin H0. am. ir. nin H0. app H7. ufi set_of_up_bounds H3. Ztac. am. Qed. Lemma Exercise1_15a11: forall r, order r -> ~ (exists e, least_element r e) -> least_element (completion_order r) emptyset. Proof. ir. assert (Ha:order (completion_order r)). uf completion_order. fprops. assert (Hb: substrate (completion_order r) = completion r). uf completion_order. aw. assert (inc emptyset (completion r)). uf completion. Ztac. app powerset_inc. app sub_emptyset_any. sy. app is_emptyset. ir. dneg. exists y. ufi set_of_uplo_bounds H1. ufi set_of_lo_bounds H1. Ztac. clear H1. split. am. ir. nin H3. app H4. uf set_of_up_bounds. Ztac. split. am. ir. elim (emptyset_pr H5). split. ue. rw Hb. ir. uf completion_order. aw. eee. app sub_emptyset_any. Qed. Lemma Exercise1_15a12 : forall r, order r -> exists x, least_element (completion_order r) x. Proof. ir. nin (p_or_not_p (exists e, least_element r e)). nin H0. exists (singleton x). app Exercise1_15a10. exists emptyset. app Exercise1_15a11. Qed. Lemma Exercise1_15a13: forall r, order r -> greatest_element (completion_order r) (substrate r). Proof. ir. assert (inc (substrate r) (completion r)). uf completion. Ztac. app powerset_inc. fprops. app extensionality. app Exercise1_15a4. fprops. uf set_of_uplo_bounds. uf set_of_lo_bounds. red. ir. Ztac. am. split. uf completion_order. aw. uf completion_order. aw. ir. aw. eee. ufi completion H1. Ztac. app powerset_sub. Qed. Lemma Exercise1_15a14: forall r X, order r -> sub X (completion r) -> least_upper_bound (completion_order r) X (set_of_uplo_bounds r (union X)). Proof. ir. assert (Ha:order (completion_order r)). uf completion_order. fprops. assert (Hb: substrate (completion_order r) = completion r). uf completion_order. aw. set (v :=set_of_uplo_bounds r (union X)). assert (inc v (completion r)). uf v. app Exercise1_15a7. red. ir. nin (union_exists H1). nin H2. cp (H0 _ H3). ufi completion H4. Ztac. rwi powerset_inc_rw H5. app H5. rw least_upper_bound_pr. split. split. ue. ir. uf completion_order. aw. ee. app H0. am. uf v. cp (H0 _ H2). ufi completion H3. Ztac. rw H5. app Exercise1_15a3. app union_sub. ir. nin H2. uf completion_order. aw. rwi Hb H2. eee. uf v. ufi completion H2. Ztac. rwi powerset_inc_rw H4. rw H5. app Exercise1_15a3. red. ir. nin (union_exists H6). nin H7. cp (H3 _ H8). ufi completion_order H9. awi H9. ee. app H11. am. ue. Qed. Lemma Exercise1_15a15: forall r X, order r -> sub X (completion r) -> nonempty X -> greatest_lower_bound (completion_order r) X (set_of_uplo_bounds r (intersection X)). Proof. ir. assert (Ha:order (completion_order r)). uf completion_order. fprops. assert (Hb: substrate (completion_order r) = completion r). uf completion_order. aw. set (v :=set_of_uplo_bounds r (intersection X)). assert (inc v (completion r)). uf v. app Exercise1_15a7. red. ir. nin H1. cp (intersection_forall H2 H1). cp (H0 _ H1). ufi completion H4. Ztac. rwi powerset_inc_rw H5. app H5. rw greatest_lower_bound_pr. split. split. ue. ir. uf completion_order. aw. ee. am. app H0. uf v. cp (H0 _ H3). ufi completion H4. Ztac. rwi powerset_inc_rw H5. rw H6. app Exercise1_15a3. red. ir. ap (intersection_forall H7 H3). ir. red in H3. ee. rwi Hb H3. uf completion_order. aw. eee. ufi completion H3. Ztac. rw H6. uf v. app Exercise1_15a3. red. ir. app intersection_inc. ir. cp (H4 _ H8). ufi completion_order H9. awi H9. ee. app H11. am. ue. Qed. Lemma Exercise1_15a16: forall r, order r -> complete_lattice (completion_order r). Proof. ir. split. uf completion_order. fprops. ir. assert (Hb: substrate (completion_order r) = completion r). uf completion_order. aw. split. exists (set_of_uplo_bounds r (union X)). app Exercise1_15a14. ue. nin (emptyset_dichot X). rw H1. exists (substrate r). rww greatest_lower_bound_emptyset. app Exercise1_15a13. uf completion_order. fprops. exists (set_of_uplo_bounds r (intersection X)). app Exercise1_15a15. ue. Qed. Lemma Exercise1_15a17: forall r, order r -> transf_axioms (fun z => set_of_lo_bounds r (singleton z)) (substrate r) (substrate (completion_order r)). Proof. ir. uf completion_order. aw. red. ir. app Exercise1_15a8. red. ir. awi H1. ue. Qed. Lemma Exercise1_15a18: forall r, order r -> order_morphism (BL (fun z => set_of_lo_bounds r (singleton z)) (substrate r) (substrate (completion_order r))) r (completion_order r). Proof. ir. cp (Exercise1_15a17 H). ufi completion_order H0. awi H0. red. uf completion_order. aw. ee. am. fprops. app injective_bl_function. ir. app (Exercise1_15a9 H H1 H2 H3). bw. bw. simpl. ir. aw. ap iff_eq. ir. ee. app Exercise1_15a8. red. ir. awi H4. ue. app Exercise1_15a8. red. ir. awi H4. ue. red. ir. ufi set_of_lo_bounds H4. Ztac. nin H6. uf set_of_lo_bounds. clear H4. Ztac. split. am. ir. awi H4. rw H4. assert (inc x (singleton x)). fprops. cp (H7 _ H8). order_tac. ir. ee. assert (inc x (set_of_lo_bounds r (singleton x))). uf set_of_lo_bounds. Ztac. split. am. ir. awi H6. rw H6. order_tac. cp (H5 _ H6). ufi set_of_lo_bounds H7. Ztac. nin H9. app H10. fprops. Qed. Lemma Exercise1_15a19: forall r X x, order r -> sub X (substrate r) -> least_upper_bound r X x -> least_upper_bound (completion_order r) (fun_image X (fun z => set_of_lo_bounds r (singleton z))) (set_of_lo_bounds r (singleton x)). Proof. ir. set (Y:= fun_image X (fun z => set_of_lo_bounds r (singleton z))). assert (sub Y (completion r)). red. ir. ufi Y H2. awi H2. nin H2. nin H2. wr H3. app Exercise1_15a8. red. ir. awi H4. rw H4. app H0. cp (Exercise1_15a14 H H2). rwi least_upper_bound_pr H1. nin H1. set (t := set_of_lo_bounds r (singleton x)). assert (sub (union Y) t). red. ir. nin (union_exists H5). nin H6. ufi Y H7. awi H7. nin H7. nin H7. wri H8 H6. ufi set_of_lo_bounds H6. Ztac. nin H10. nin H1. cp (H12 _ H7). assert (inc x2 (singleton x2)). fprops. cp (H11 _ H14). uf t. uf set_of_lo_bounds. clear H6. Ztac. split. order_tac. ir. awi H6. rw H6. order_tac. assert (sub (set_of_uplo_bounds r (union Y)) t). assert (set_of_uplo_bounds r t = t). uf t. uf set_of_uplo_bounds. rww Exercise1_15a5. red. ir. awi H6. ue. nin H1. am. wr H6. app Exercise1_15a3. assert (t = set_of_uplo_bounds r (union Y)). app extensionality. red. ir. ufi t H7. ufi set_of_lo_bounds H7. Ztac. clear H7. nin H9. assert (inc x (singleton x)). fprops. cp (H9 _ H10). uf set_of_uplo_bounds. uf set_of_lo_bounds. Ztac. split. order_tac. ir. ufi set_of_up_bounds H12. Ztac. nin H14. cut (gle r x y). ir. order_tac. app H4. split. am. ir. app H15. assert (inc y0 (set_of_lo_bounds r (singleton y0))). uf set_of_lo_bounds. clear H12. Ztac. split. app H0. ir. awi H12. rw H12. order_tac. app H0. apply union_inc with (set_of_lo_bounds r (singleton y0)). am. uf Y. aw. exists y0. ee. am. tv. ue. am. am. Qed. Lemma Exercise1_15a20: forall r X x, order r -> sub X (substrate r) -> greatest_lower_bound r X x -> greatest_lower_bound (completion_order r) (fun_image X (fun z => set_of_lo_bounds r (singleton z))) (set_of_lo_bounds r (singleton x)). Proof. ir. set (Y:= fun_image X (fun z => set_of_lo_bounds r (singleton z))). assert (sub Y (completion r)). red. ir. ufi Y H2. awi H2. nin H2. nin H2. wr H3. app Exercise1_15a8. red. ir. awi H4. rw H4. app H0. nin (emptyset_dichot Y). assert (X = emptyset). app is_emptyset. ir. red. ir. elim (emptyset_pr (x:=set_of_lo_bounds r (singleton y))). wr H3. uf Y. aw. ex_tac. rwi H4 H1. rwi greatest_lower_bound_pr H1. nin H1. assert(set_of_lo_bounds r (singleton x) = substrate r). ap extensionality. uf set_of_lo_bounds. app Z_sub. red. ir. uf set_of_lo_bounds. Ztac. split. am. ir. awi H7. rw H7. app H5. split. am. ir. elim (emptyset_pr H8). rw H6. rw H3. rw greatest_lower_bound_emptyset. app Exercise1_15a13. uf completion_order. fprops. am. app sub_emptyset_any. cp (Exercise1_15a15 H H2 H3). rwii greatest_lower_bound_pr H1. nin H1. set (t := set_of_lo_bounds r (singleton x)). assert (sub t (intersection Y)). red. ir. app intersection_inc. ufi t H6. ufi set_of_lo_bounds H6. Ztac. clear H6. nin H8. ir. ufi Y H9. awi H9. nin H9. nin H9. wr H10. uf set_of_lo_bounds. Ztac. split. am. ir. awi H11. rw H11. assert (inc x (singleton x)). fprops. cp (H8 _ H12). nin H1. cp (H14 _ H9). order_tac. assert (sub t (set_of_uplo_bounds r (intersection Y))). assert (set_of_uplo_bounds r t = t). uf t. uf set_of_uplo_bounds. rww Exercise1_15a5. red. ir. awi H7. ue. nin H1. am. wr H7. app Exercise1_15a3. assert (t = set_of_uplo_bounds r (intersection Y)). app extensionality. red. ir. ufi set_of_uplo_bounds H8. ufi set_of_lo_bounds H8. Ztac. clear H8. nin H10. uf t. uf set_of_lo_bounds. Ztac. split. am. ir. awi H11. rw H11. app H10. uf set_of_up_bounds. Ztac. nin H1. am. split. nin H1. am. ir. app H5. split. nin H3. cp (intersection_forall H12 H3). ufi Y H3. awi H3. nin H3. nin H3. wri H14 H13. ufi set_of_lo_bounds H13. Ztac. am. ir. assert (inc (set_of_lo_bounds r (singleton y1)) Y). uf Y. aw. ex_tac. cp (intersection_forall H12 H14). ufi set_of_lo_bounds H15. Ztac. nin H17. app H18. fprops. ue. Qed. Lemma Exercise1_15b1: forall r X, order r -> sub X (substrate r) -> least_upper_bound (completion_order r) (fun_image X (fun z => set_of_lo_bounds r (singleton z))) (set_of_uplo_bounds r X). Proof. ir. set (Y:= fun_image X (fun z => set_of_lo_bounds r (singleton z))). assert (sub Y (completion r)). red. ir. ufi Y H1. awi H1. nin H1. nin H1. wr H2. ap Exercise1_15a8. red. ir. awi H3. rw H3. app H0. assert (set_of_uplo_bounds r (union Y) = set_of_uplo_bounds r X). uf set_of_uplo_bounds; uf set_of_lo_bounds; uf set_of_up_bounds. set_extens. Ztac. clear H2. nin H4. Ztac. split. am. ir. Ztac. clear H5. app H4. Ztac. nin H7. split. am. ir. nin (union_exists H8). nin H9. ufi Y H10. awi H10. nin H10. nin H10. wri H11 H9. ufi set_of_lo_bounds H9. Ztac. nin H13. cp (H7 _ H10). assert (inc x1 (singleton x1)). fprops. cp (H14 _ H16). order_tac. Ztac. clear H2. nin H4. Ztac. split. am. ir. Ztac. nin H7. app H4. clear H5. Ztac. split. am. ir. app H8. apply union_inc with (set_of_lo_bounds r (singleton y0)). uf set_of_lo_bounds. Ztac. split. app H0. ir. awi H9. rw H9. order_tac. app H0. uf Y. aw. ex_tac. wr H2. ap (Exercise1_15a14 H H1). Qed. Lemma Exercise1_15c: forall r, total_order r -> total_order (completion_order r). Proof. ir. nin H. split. uf completion_order. fprops. uf completion_order. assert (forall x a b, inc x (completion r) -> inc a x -> gle r b a -> inc b x). ir. ufi completion H1. Ztac. rw H5. uf set_of_uplo_bounds. uf set_of_lo_bounds. uf set_of_lo_bounds. clear H1. Ztac. order_tac. split. order_tac. ir. ufi set_of_up_bounds H1. Ztac. nin H7. cp (H8 _ H2). order_tac. aw. ir. uf gge. aw. nin (p_or_not_p (sub x y)). left. eee. right. eee. red. ir. nin (inc_or_not x0 x). am. elim H4. red. ir. nin (inc_or_not x1 y). am. assert (inc x0 (substrate r)). ufi completion H3. Ztac. rwi powerset_inc_rw H9. app H9. assert (inc x1 (substrate r)). ufi completion H2. Ztac. rwi powerset_inc_rw H10. app H10. nin (H0 _ _ H9 H10). cp (H1 _ _ _ H2 H7 H11). contradiction. cp (H1 _ _ _ H3 H5 H11). contradiction. Qed. (* Exercise 16 *) Definition distributive_lattice1 r := forall x y z, inc x (substrate r) ->inc y (substrate r) -> inc z (substrate r) -> sup r x (inf r y z) = inf r (sup r x y) (sup r x z). Definition distributive_lattice2 r := forall x y z, inc x (substrate r) ->inc y (substrate r) -> inc z (substrate r) -> inf r x (sup r y z) = sup r (inf r x y) (inf r x z). Definition distributive_lattice3 r := forall x y z, inc x (substrate r) ->inc y (substrate r) -> inc z (substrate r) -> sup r (inf r x y) (sup r (inf r y z) (inf r z x)) = inf r (sup r x y) (inf r (sup r y z) (sup r z x)). Definition distributive_lattice4 r := forall x y z, inc x (substrate r) ->inc y (substrate r) -> inc z (substrate r) -> gle r z x -> sup r z (inf r x y) = inf r x (sup r y z). Definition distributive_lattice5 r:= forall x y z, inc x (substrate r) ->inc y (substrate r) -> inc z (substrate r) -> gle r (inf r z (sup r x y)) (sup r x (inf r y z)). Definition distributive_lattice6 r := forall x y z, inc x (substrate r) ->inc y (substrate r) -> inc z (substrate r) -> inf r (sup r x y) (sup r z (inf r x y)) = sup r (inf r x y) (sup r (inf r y z) (inf r z x)). Lemma lattice_props: forall r, lattice r -> let E := substrate r in ( (forall x y, inc x E-> inc y E -> inc (sup r x y) E) & (forall x y, inc x E-> inc y E -> inc (inf r x y) E) & (forall x y, inc x E-> inc y E -> sup r x y = sup r y x) & (forall x y, inc x E-> inc y E -> inf r x y = inf r y x) & (forall x y, inc x E-> inc y E -> sup r (inf r x y) y = y) & (forall x y, inc x E-> inc y E -> inf r (sup r x y) y = y) & (forall x y z, inc x E-> inc y E -> inc z E -> sup r x (sup r y z) = sup r (sup r x y) z) & (forall x y z, inc x E-> inc y E -> inc z E -> inf r x (inf r y z) = inf r (inf r x y) z)). Proof. ir. uf E. ee. ir. cp (lattice_sup_pr H H0 H1). ee. order_tac. ir. cp (lattice_inf_pr H H0 H1). ee. order_tac. ir. cp (lattice_sup_pr H H0 H1). ee. cp (lattice_sup_pr H H1 H0). ee. nin H. assert (gle r (sup r x y) (sup r y x)). app H4. assert (gle r (sup r y x) (sup r x y)). app H7. order_tac. uf E. ir. cp (lattice_inf_pr H H0 H1). ee. cp (lattice_inf_pr H H1 H0). ee. nin H. assert (gle r (inf r x y) (inf r y x)). app H7. assert (gle r (inf r y x) (inf r x y)). app H4. order_tac. ir. app sup_comparable1. nin H. am. cp (lattice_inf_pr H H0 H1). ee. am. ir. cp (lattice_sup_pr H H0 H1). ee. assert (inc (sup r x y) (substrate r)). order_tac. cp (lattice_inf_pr H H5 H1). ee. assert (gle r y (inf r (sup r x y) y)). app H8. wrr order_reflexivity. nin H. am. nin H. order_tac. ir. cp (lattice_sup_pr H H0 H1). cp (lattice_sup_pr H H1 H2). ee. assert (inc (sup r y z) (substrate r)). order_tac. assert (inc (sup r x y) (substrate r)). order_tac. cp (lattice_sup_pr H H0 H9). cp (lattice_sup_pr H H10 H2). ee. nin H. assert (gle r (sup r x (sup r y z))(sup r (sup r x y) z)). ap H16. apply order_transitivity with(sup r x y). am. ap H3. ap H12. ap H6. apply order_transitivity with(sup r x y). am. ap H7. ap H12. ap H13. assert (gle r (sup r (sup r x y) z) (sup r x (sup r y z))). ap H14. ap H8. ap H11. apply order_transitivity with(sup r y z). am. ap H4. ap H15. apply order_transitivity with(sup r y z). am. ap H5. ap H15. order_tac. ir. cp (lattice_inf_pr H H0 H1). cp (lattice_inf_pr H H1 H2). ee. assert (inc (inf r y z) (substrate r)). order_tac. assert (inc (inf r x y) (substrate r)). order_tac. cp (lattice_inf_pr H H0 H9). cp (lattice_inf_pr H H10 H2). ee. nin H. assert (gle r (inf r x (inf r y z))(inf r (inf r x y) z)). ap H14. ap H8. ap H11. apply order_transitivity with(inf r y z). am. ap H15. ap H4. apply order_transitivity with(inf r y z). am. ap H15. ap H5. assert (gle r (inf r (inf r x y) z) (inf r x (inf r y z))). ap H16. apply order_transitivity with(inf r x y). am. ap H12. ap H3. ap H6. apply order_transitivity with(inf r x y). am. ap H12. ap H7. ap H13. order_tac. Qed. Lemma Exercise1_16a: forall r, lattice r -> ( (distributive_lattice1 r -> distributive_lattice3 r) & (distributive_lattice2 r -> distributive_lattice3 r)). Proof. ir. set (E:= substrate r). cp (lattice_props H). simpl in H0. fold E in H0. ee. rename H6 into sup_assoc. rename H7 into inf_assoc. ir. red in H6. red. fold E. fold E in H6. ir. set (sxy :=sup r x y). set (syz:=sup r y z). set (szx:= sup r z x). cp (H1 _ _ H8 H9). rw (H6 _ _ _ H10 H9 H7). rw (H2 _ _ H10 H7). rw (H6 _ _ _ H7 H8 H9). rw (H2 _ _ H7 H9). fold szx. fold sxy. rw (H4 _ _ H8 H9). assert (inf r z (inf r sxy szx) = inf r z sxy). cp (H0 _ _ H7 H8). cp (H0 _ _ H9 H7). uf sxy. uf szx. rw (H3 _ _ H11 H12). rw (inf_assoc _ _ _ H9 H12 H11). rw (H3 _ _ H9 H12). rw (H2 _ _ H9 H7). rw (H5 _ _ H7 H9). tv. rw H11. cp (H0 _ _ H7 H8). fold sxy in H12. rw (H6 _ _ _ (H1 _ _ H7 H8) H9 H12). rw (H2 _ _ (H1 _ _ H7 H8) H9). rw (H6 _ _ _ H9 H7 H8). fold szx. rw (H2 _ _ H9 H8). fold syz. rw (H2 _ _ (H1 _ _ H7 H8) H12). rw (H6 _ _ _ H12 H7 H8). assert (inf r (sup r sxy x) (sup r sxy y) = sxy). uf sxy. cp (lattice_sup_pr H H7 H8). ee. ufi sxy H12. assert (sup r (sup r x y) x = (sup r x y)). rw (H2 _ _ H12 H7). app sup_comparable1. nin H;am. rw H16. assert (sup r (sup r x y) y = (sup r x y)). rw (H2 _ _ H12 H8). app sup_comparable1. nin H;am. rw H17. app inf_comparable1. nin H. am. wrr order_reflexivity. nin H; am. rw H13. rw H3. ap uneq. rw H3. app refl_equal. uf szx. app H0. uf syz. app H0. app H1. uf szx. app H0. uf syz. app H0. am. rename H6 into sup_assoc. rename H7 into inf_assoc. ir. red in H6. red. fold E. fold E in H6. ir. set (sxy :=inf r x y). set (syz:=inf r y z). set (szx:= inf r z x). cp (H0 _ _ H8 H9). rw (H6 _ _ _ H10 H9 H7). rw (H3 _ _ H10 H7). rw (H6 _ _ _ H7 H8 H9). rw (H3 _ _ H7 H9). fold szx. fold sxy. rw (H5 _ _ H8 H9). assert (sup r z (sup r sxy szx) = sup r z sxy). cp (H1 _ _ H7 H8). cp (H1 _ _ H9 H7). uf sxy. uf szx. rw (H2 _ _ H11 H12). rw (sup_assoc _ _ _ H9 H12 H11). rw (H2 _ _ H9 H12). rw (H3 _ _ H9 H7). rw (H4 _ _ H7 H9). tv. rw H11. cp (H1 _ _ H7 H8). fold sxy in H12. rw (H6 _ _ _ (H0 _ _ H7 H8) H9 H12). rw (H3 _ _ (H0 _ _ H7 H8) H9). rw (H6 _ _ _ H9 H7 H8). fold szx. rw (H3 _ _ H9 H8). fold syz. rw (H3 _ _ (H0 _ _ H7 H8) H12). rw (H6 _ _ _ H12 H7 H8). assert (sup r (inf r sxy x) (inf r sxy y) = sxy). uf sxy. cp (lattice_inf_pr H H7 H8). ee. ufi sxy H12. assert (inf r (inf r x y) x = (inf r x y)). app inf_comparable1. nin H;am. rw H16. assert (inf r (inf r x y) y = (inf r x y)). app inf_comparable1. nin H;am. rw H17. app sup_comparable1. nin H. am. wrr order_reflexivity. nin H; am. rw H13. assert (sup r syz szx= sup r szx syz). rw H2. app refl_equal. uf syz. app H1. uf szx. app H1. rw H14. rw H2. app refl_equal. am. ap H0. uf szx. app H1. uf syz. app H1. Qed. Lemma Exercise1_16b: forall r, lattice r -> ( (distributive_lattice3 r -> distributive_lattice4 r) & (distributive_lattice3 r -> distributive_lattice1 r) & (distributive_lattice3 r -> distributive_lattice2 r)). Proof. ir. set (E:= substrate r). cp (lattice_props H). simpl in H0. fold E in H0. assert (distributive_lattice3 r -> distributive_lattice4 r). ee. rename H6 into sup_assoc. rename H7 into inf_assoc. ir. red in H6. red. fold E in H6. fold E. ir. cp (H6 _ _ _ H7 H8 H9). assert (inf r z x = z). app inf_comparable1. nin H; am. rwi H12 H11. assert (sup r z x = x). app sup_comparable1. nin H; am. rwi H13 H11. assert (sup r (inf r y z) z = z). app sup_comparable1. nin H; am. cp (lattice_inf_pr H H8 H9). ee. am. rwi H14 H11. assert (inf r (sup r y z) x = inf r x (sup r y z)). app H3. app H0. rwi H15 H11. rwi inf_assoc H11. assert (inf r (sup r x y) x = x). rww H3. app inf_comparable1. nin H; am. cp (lattice_sup_pr H H7 H8). ee; am. app H0. rwi H16 H11. wr H11. app H2. app H1. app H0. am. app H0. ee. am. ir. cp (H1 H9). rename H10 into HM. ufi distributive_lattice3 H9. ufi distributive_lattice4 HM. uf distributive_lattice1; fold E. fold E in H9. fold E in HM. ir. set (a:= sup r x (sup r (inf r x y) (sup r (inf r y z) (inf r z x)))). set (b:= sup r x (inf r (sup r x y) (inf r (sup r y z) (sup r z x)))). assert (a= b). uf a; uf b. rww H9. set (c:= sup r (inf r y z) (inf r z x)). assert (a = sup r x (sup r (inf r x y) c)). uf a. fold c. app refl_equal. rwi H7 H14. assert (sup r x (inf r x y) = x). rww H3. app sup_comparable1. nin H; am. cp (lattice_inf_pr H H10 H11). ee. am. app H2. rwi H15 H14. ufi c H14. clear H15. clear c. assert (sup r (inf r y z) (inf r z x) = sup r (inf r z x) (inf r y z)). app H3. app H2. app H2. rwi H15 H14. rwi H7 H14. assert (sup r x (inf r z x)=x). rww H3. app sup_comparable1. nin H;am. cp (lattice_inf_pr H H12 H10). ee. am. app H2. rwi H16 H14. wr H14. clear H15. clear H16. clear H14. rw H13. clear H13. clear a. set (c:= inf r (sup r y z) (sup r z x)) in *. assert (inc (sup r x y) E). app H0. assert (gle r x (sup r x y)). cp (lattice_sup_pr H H10 H11). ee. am. assert (inc c E). uf c. app H2. app H0. app H0. uf b. rw (HM _ _ _ H13 H15 H10 H14). uf c. set (d:= sup r (inf r (sup r y z) (sup r z x)) x). assert (d= sup r x (inf r (sup r y z) (sup r z x))). rww H3. assert (d= sup r x (inf r (sup r z x) (sup r y z))). rw H16. rw H4. app refl_equal. app H0. app H0. assert (inc (sup r z x) E). app H0. assert (inc (sup r y z) E). app H0. assert (gle r x (sup r z x)). cp (lattice_sup_pr H H12 H10). ee. am. rwi (HM _ _ _ H18 H19 H10 H20) H17. assert(d= (sup r z x)). rw H17. ap inf_comparable1. nin H; am. wr H7. set (e:= sup r z x). assert (inc e E). uf e. app H0. cp (lattice_sup_pr H H11 H21). ee. am. am. am. am. rw H21. rw (H3 _ _ H12 H10). app refl_equal. am. app H2. app H2. am. app H2. uf c. app H0. app H2. app H2. ir. cp (H1 H9). rename H10 into HM. ufi distributive_lattice3 H9. ufi distributive_lattice4 HM. uf distributive_lattice2; fold E. fold E in H9. fold E in HM. ir. set (a:= inf r x (sup r (inf r x y) (sup r (inf r y z) (inf r z x)))). set (b:= inf r x (inf r (sup r x y) (inf r (sup r y z) (sup r z x)))). assert (a= b). uf a; uf b. rww H9. set (c:= inf r (sup r y z) (sup r z x)). assert (b = inf r x (inf r (sup r x y) c)). uf b. fold c. app refl_equal. rwi H8 H14. assert (inf r x (sup r x y) = x). rww H4. rww H4. app inf_comparable1. nin H; am. cp (lattice_sup_pr H H10 H11). ee. am. app H0. app H0. rwi H15 H14. ufi c H14. clear H15. clear c. assert (inf r (sup r y z) (sup r z x) = inf r (sup r z x) (sup r y z)). app H4. app H0. app H0. rwi H15 H14. rwi H8 H14. assert (inf r x (sup r z x)=x). ap inf_comparable1. nin H;am. cp (lattice_sup_pr H H12 H10). ee. am. rwi H16 H14. wr H14. clear H15. clear H16. clear H14. wr H13. clear H13. clear b. set (c:= sup r (inf r y z) (inf r z x)) in *. assert (inc (inf r x y) E). app H2. assert (gle r (inf r x y) x). cp (lattice_inf_pr H H10 H11). ee. am. assert (inc c E). uf c. app H0. app H2. app H2. uf a. rw H3. wr (HM _ _ _ H10 H15 H13 H14). uf c. assert (inc (inf r y z) E). app H2. assert (gle r (inf r z x) x). cp (lattice_inf_pr H H12 H10). ee. am. assert (inc (inf r z x) E). app H2. wr (HM _ _ _ H10 H16 H18 H17). assert (sup r (inf r z x) (inf r x (inf r y z)) = inf r x z). rw (H4 _ _ H10 H12). rw (H4 _ _ H11 H12). rww H8. rw (H4 _ _ H10 H12). rw H3. ap sup_comparable1. nin H; am. cp (lattice_inf_pr H H18 H11). ee; am. am. app H2. rw H19. app refl_equal. am. am. am. app H0. app H0. am. app H0. uf c. app H2. app H0. app H0. Qed. Lemma Exercise1_16c: forall r, lattice r -> (distributive_lattice3 r = distributive_lattice5 r). Proof. ir. app iff_eq. ir. red. ir. cp (Exercise1_16b H). ee. cp (H6 H0). red in H7. rww H7. cp (lattice_props H). simpl in H8. set (E:= substrate r) in *. ee. rw (H11 _ _ H3 H2). set (b:=inf r y z). assert (inc b E). uf b; app H9. assert (inc (inf r z x) E). app H9. cp (lattice_sup_pr H H17 H16). ee. ap H20. assert (order r). nin H; am. assert (gle r (inf r z x) x). cp (lattice_inf_pr H H3 H1). ee. am. assert (gle r x (sup r x b)). cp (lattice_sup_pr H H1 H16). ee. am. order_tac. cp (lattice_sup_pr H H1 H16). ee. am. ir. cut (distributive_lattice1 r). ir. cp (Exercise1_16a H). ee. ap H2. am. red. red in H0. cp (lattice_props H). simpl in H1. ee. set (E:= substrate r) in *. ir. set (b:= sup r x (inf r y z)). assert (inc b E). uf b. app H1. app H2. assert (gle r b (inf r (sup r x y) (sup r x z))). assert (inc (sup r x y) E). app H1. assert (inc (sup r x z) E). app H1. cp (lattice_inf_pr H H13 H14). ee. ap H17. uf b. assert (inc (inf r y z) E). app H2. cp (lattice_sup_pr H H9 H18). ee. ap H21. cp (lattice_sup_pr H H9 H10). ee; am. assert (gle r (inf r y z) y). cp (lattice_inf_pr H H10 H11). ee; am. assert (gle r y (sup r x y)). cp (lattice_sup_pr H H9 H10). ee; am. nin H. order_tac. uf b. assert (inc (inf r y z) E). app H2. cp (lattice_sup_pr H H9 H18). ee. ap H21. cp (lattice_sup_pr H H9 H11). ee; am. assert (gle r (inf r y z) z). cp (lattice_inf_pr H H10 H11). ee; am. assert (gle r z (sup r x z)). cp (lattice_sup_pr H H9 H11). ee; am. nin H. order_tac. cp (H0 _ _ _ H9 H10 H11). cp (H0 _ _ _ H9 H11 (H1 _ _ H9 H10)). set (a:= inf r (sup r x y) (sup r x z)) in *. set (c:= inf r z (sup r x y)) in *. assert (sup r x b = b). ap sup_comparable1. nin H; am. uf b. cp (lattice_sup_pr H H9 (H2 _ _ H10 H11)). ee; am. assert (gle r (sup r x c) b). wr H16. assert (inc c E). uf c. app H2. app H1. cp (lattice_sup_pr H H9 H12). cp (lattice_sup_pr H H9 H17). ee. ap H21. ap H18. nin H. ap (order_transitivity H H14 H22). nin H. cp (order_transitivity H H15 H17). ap (order_antisymmetry H H13 H19). Qed. Lemma Exercise1_16d: forall r, lattice r -> (distributive_lattice3 r = distributive_lattice6 r). Proof. ir. cp (lattice_props H). ap iff_eq. ir. assert (distributive_lattice2 r). cp (Exercise1_16b H). ee. app H4. red in H2. red. simpl in H0. ir. ee. set (E:= substrate r) in *. cp (H0 _ _ H3 H4). cp (H6 _ _ H3 H4). rw (H2 _ _ _ H13 H5 H14). rw (H8 _ _ H13 H5). rw (H2 _ _ _ H5 H3 H4). rw (H8 _ _ H4 H5). assert (inf r (sup r x y) (inf r x y) = inf r x y). rww H8. ap inf_comparable1. nin H; am. cp (lattice_sup_pr H H3 H4). cp (lattice_inf_pr H H3 H4). ee. nin H. order_tac. rw H15. set (xy:=inf r x y). set (zx:=inf r z x). set (zy:=inf r z y). rw H7. ap uneq. rw H7. ap uneq. tv. uf zx. app H6. uf zy. app H6. app H0. uf zx. app H6. uf zy. app H6. app H14. ir. rww Exercise1_16c. red in H1. red. ir. simpl in H0. ir. ee. set (E:= substrate r) in *. set (a:=inf r z (sup r x y)). assert (a = inf r (sup r x y) z). rw H7. tv. app H0. am. assert (gle r a (inf r (sup r x y) (sup r z (inf r x y)))). rw H12. set (xy:= sup r x y). assert (inc xy E). uf xy. app H0. set (b:= inf r x y). assert (inc b E). uf b. app H5. assert (gle r z (sup r z b)). cp (lattice_sup_pr H H4 H14). ee. am. cp (lattice_inf_pr H H13 (H0 _ _ H4 H14)). ee. cp (lattice_inf_pr H H13 H4). ee. ap H18. am. nin H. order_tac. rwi H1 H13. set (c:= sup r (inf r x y) (sup r (inf r y z) (inf r z x))) in *. assert (gle r c (sup r x (inf r y z))). uf c. set (xy:= inf r x y). assert (inc xy E). uf xy. app H5. set (yz:= inf r y z). assert (inc yz E). uf yz. app H5. set (zx:= inf r z x). assert (inc zx E). uf zx. app H5. assert (inc (sup r yz zx) E). app H0. cp (lattice_inf_pr H H2 H3). fold xy in H18. ee. cp (lattice_inf_pr H H4 H2). fold zx in H21. ee. cp (lattice_sup_pr H H14 H17). ee. ap H26. cp (lattice_sup_pr H H2 H15). ee. nin H. order_tac. cp (lattice_sup_pr H H2 H15). ee. cp (lattice_sup_pr H H15 H16). ee. ap H32. am. nin H. ap (order_transitivity H H22 H27). nin H. ap (order_transitivity H H13 H14). am. am. am. Qed. Definition relatively_complemented r:= lattice r & (exists u, least_element r u) & (forall x y, gle r x y -> exists x', (inc x' (substrate r) &sup r x x' = y & inf r x x' = the_least_element r)). Definition boolean_lattice r:= relatively_complemented r & (exists u, greatest_element r u) & distributive_lattice3 r. Definition the_complement r x y:= choose (fun x' => (inc x' (substrate r) & sup r x x' = y & inf r x x' = the_least_element r)). Lemma the_complement_pr: forall r x y, relatively_complemented r -> gle r x y -> let x' := the_complement r x y in (inc x' (substrate r) & sup r x x' = y & inf r x x' = the_least_element r). Proof. ir. uf x'. uf the_complement. app choose_pr. red in H. ee. ir. app H2. Qed. Lemma least_greatest_pr: forall r, order r -> ((forall a, inc a (substrate r) -> (exists u, least_element r u) -> sup r (the_least_element r) a = a) & (forall a, inc a (substrate r) -> (exists u, greatest_element r u) -> inf r a (the_greatest_element r) = a) & (forall a, inc a (substrate r) -> (exists u, least_element r u) -> inf r (the_least_element r) a = (the_least_element r)) & (forall a, inc a (substrate r) -> (exists u, greatest_element r u)-> sup r a (the_greatest_element r) = (the_greatest_element r))). Proof. ir. ee. ir. app sup_comparable1. cp (the_least_element_pr H H1). nin H2. app H3. ir. app inf_comparable1. cp (the_greatest_element_pr H H1). nin H2. app H3. ir. app inf_comparable1. cp (the_least_element_pr H H1). nin H2. app H3. ir. app sup_comparable1. cp (the_greatest_element_pr H H1). nin H2. app H3. Qed. Lemma Exercice1_17a: forall r x y, relatively_complemented r -> distributive_lattice3 r -> gle r x y -> exists_unique (fun x' => (inc x' (substrate r) & sup r x x' = y & inf r x x' = the_least_element r)). Proof. ir. red. ee. exists (the_complement r x y). app the_complement_pr. ir. red in H. ee. cp (lattice_props H). simpl in H10. red in H0. cp (inc_arg1_substrate H1). cp (inc_arg2_substrate H1). set (E:= substrate r) in *. ee. cp (H0 _ _ _ H3 H11 H2). rwi H6 H20. rwi H14 H4. rwi H4 H20. rwi H7 H20. rwi H15 H5. rwi H5 H20. assert (order r). nin H; am. cp (least_greatest_pr H21). ee. assert (sup r (the_least_element r) (inf r x0 y0) = (inf r x0 y0)). rww H22. app H13. rwi H26 H20. rwi H26 H20. clear H26. set (sxy:=sup r x0 y0). assert (inc sxy E). uf sxy. app H10. assert (inf r y (sup r x0 y0) = (sup r x0 y0)). rw H15. app inf_comparable1. cp (lattice_sup_pr H H2 H3). ee. app H29. cp (lattice_sup_pr H H11 H2). ee. wr H6. am. cp (lattice_sup_pr H H3 H11). ee. wr H4. am. am. ufi sxy H26. am. rwi H27 H20. rwi H27 H20. cp (lattice_sup_pr H H2 H3). cp (lattice_inf_pr H H2 H3). ee. nin H. rwi H20 H29. rwi H20 H30. cp (order_transitivity H H28 H30). cp (order_transitivity H H32 H29). ap (order_antisymmetry H H35 H36). am. am. am. am. Qed. Definition standard_completion r x := the_complement r x (the_greatest_element r). Lemma standard_completion_pr: forall r x, let y := standard_completion r x in boolean_lattice r -> inc x (substrate r) -> (inc y (substrate r) & sup r x y = the_greatest_element r & inf r x y = the_least_element r). Proof. ir. uf y. uf standard_completion. app the_complement_pr. red in H; ee; am. red in H. ee. red in H. ee. red in H. ee. cp (the_greatest_element_pr H H1). red in H6. ee. app H7. Qed. Lemma standard_completion_involutive: forall r x, boolean_lattice r -> inc x (substrate r) -> standard_completion r (standard_completion r x) = x. Proof. ir. cp (standard_completion_pr H H0). ee. set (y:= standard_completion r x) in *. cp (standard_completion_pr H H1). ee. set (z:= standard_completion r y) in *. red in H; ee. assert (lattice r). red in H. ee. am. assert (gle r y (the_greatest_element r)). cp (lattice_sup_pr H9 H0 H1). ee. wr H2. am. nin (Exercice1_17a H H8 H10). app H12. eee. eee. cp (lattice_props H9). simpl in H13. eee. cp (lattice_props H9). simpl in H13. eee. Qed. Lemma standard_completion_monotone: forall r x y, boolean_lattice r -> gle r x y -> gle r (standard_completion r y) (standard_completion r x). Proof. ir. cp (inc_arg1_substrate H0). cp (inc_arg2_substrate H0). cp (standard_completion_pr H H1). cp (standard_completion_pr H H2). ee. set (a:= standard_completion r x) in *. set (b:= standard_completion r y) in *. set (c := inf r a b). red in H; ee. assert (lattice r). red in H; ee; am. nin (lattice_props H11). ee. assert (inf r y c = the_least_element r). uf c. rw (H15 _ _ H3 H4). rw H19. rw H6. ap inf_comparable1. nin H11; am. nin H. ee. nin H. cp (the_least_element_pr H H20). red in H23. ee. app H24. am. am. am. assert (sup r y c = sup r y a). uf c. assert (distributive_lattice1 r). cp (Exercise1_16b H11). ee. app H22. red in H21. rww H21. rw H5. nin H11. cp (least_greatest_pr H11). ee. app H24. app H12. assert (gle r (sup r x a) (sup r y a)). cp (lattice_sup_pr H11 H1 H3). cp (lattice_sup_pr H11 H2 H3). ee. ap H27. nin H11. order_tac. am. rwi H7 H22. assert (sup r y c = the_greatest_element r). rw H21. nin H11. nin (the_greatest_element_pr H11 H9). assert (inc (sup r y a) (substrate r)). app H12. cp (H25 _ H26). order_tac. cp (lattice_sup_pr H11 H2 H4). ee. rwi H5 H24. nin (Exercice1_17a H H10 H24). assert (c = b). ap H28. ee. uf c. app H13. am. am. ee. am. am. am. ufi c H29. cp (lattice_inf_pr H11 H3 H4). ee. wr H29. exact H30. Qed. Lemma Exercise1_17b: forall r, boolean_lattice r -> order_isomorphism (BL (standard_completion r) (substrate r)(substrate r)) r (opposite_order r). Proof. ir. assert (transf_axioms (standard_completion r) (substrate r) (substrate r)). red. ir. nin (standard_completion_pr H H0). am. assert (is_function (BL (standard_completion r) (substrate r) (substrate r))). app bl_function. assert (order r). red in H. ee. red in H; ee. nin H. am. red. ee. am. fprops. red. split. app injective_bl_function. ir. wr (standard_completion_involutive H H3). wr (standard_completion_involutive H H4). rww H5. app surjective_bl_function. ir. cp (standard_completion_involutive H H3). exists (standard_completion r y). split. app H0. am. aw. aw. aw. ir. rww W_bl_function. rww W_bl_function. aw. app iff_eq. ir. app standard_completion_monotone. ir. wr (standard_completion_involutive H H3). wr (standard_completion_involutive H H4). app standard_completion_monotone. Qed. Lemma Exercise1_17c: forall a, (boolean_lattice (inclusion_order a) & (forall x, inc x (powerset a) -> standard_completion (inclusion_order a) x = complement a x)). Proof. ir. set (r:=inclusion_order a). uf standard_completion. assert (substrate r = powerset a). uf r. rww substrate_inclusion_order. assert (order r). uf r. fprops. assert (least_element r emptyset). red. split. rw H. app powerset_inc. app sub_emptyset_any. ir. uf r. aw. rwi H H1. ee. app sub_emptyset_any. app powerset_sub. app sub_emptyset_any. assert (exists u, least_element r u). exists emptyset. am. assert (the_least_element r = emptyset). cp (the_least_element_pr H0 H2). app (unique_least H0 H3 H1). assert (greatest_element r a). red. split. rw H. app powerset_inc. fprops. ir. uf r. aw. rwi H H4. rwi powerset_inc_rw H4. ee. am. fprops. am. assert (exists u, greatest_element r u). exists a. am. assert (the_greatest_element r = a). cp (the_greatest_element_pr H0 H5). app (unique_greatest H0 H6 H4). rw H6. assert (forall x y, inc x (powerset a) -> inc y (powerset a) -> least_upper_bound r (doubleton x y) (union2 x y)). ir. assert (sub x a). app powerset_sub. assert (sub y a). app powerset_sub. assert (sub (union2 x y) a). red. ir. nin (union2_or H11). app H9. app H10. app least_upper_bound_doubleton. uf r. aw. eee. red. ir. app union2_first. uf r. aw. eee. red. ir. app union2_second. ir. uf r. aw. ee. am. ufi r H12. awi H12. ee. am. red. ir. nin(union2_or H14). ufi r H12. awi H12. ee. app H17. ufi r H13. awi H13. ee. app H17. assert (forall x y, inc x (powerset a) -> inc y (powerset a) -> greatest_lower_bound r (doubleton x y) (intersection2 x y)). ir. assert (sub x a). app powerset_sub. assert (sub y a). app powerset_sub. assert (sub (intersection2 x y) a). red. ir. app H10. ap (intersection2_first H12). app greatest_lower_bound_doubleton. uf r. aw. eee. app intersection2sub_first. uf r. aw. eee. app intersection2sub_second. red. ir. uf r. aw. ee. ufi r H13. awi H13. ee. am. red. ir. ufi r H13. awi H13. ee. ap H16. ap (intersection2_first H15). ufi r H13. awi H13. ufi r H14. awi H14. ee. red. ir. ap intersection2_inc. app H18. app H16. assert (forall x y, inc x (powerset a) ->inc y (powerset a) -> (has_supremum r (doubleton x y) & has_infimum r (doubleton x y))). ir. split. exists (union2 x y). app H7. exists (intersection2 x y). app H8. assert (forall x y, inc x (powerset a) ->inc y (powerset a) -> sup r x y = union2 x y). ir. cp (H7 _ _ H10 H11). assert (least_upper_bound r (doubleton x y) (sup r x y)). uf sup. app supremum_pr1. red. rw H. ir. nin (doubleton_or H13). rww H14. rww H14. exists (union2 x y). am. app (supremum_unique H0 H13 H12). assert (forall x y, inc x (powerset a) ->inc y (powerset a) -> inf r x y = intersection2 x y). ir. cp (H8 _ _ H11 H12). assert (greatest_lower_bound r (doubleton x y) (inf r x y)). uf inf. app infimum_pr1. red. rw H. ir. nin (doubleton_or H14). rww H15. rww H15. exists (intersection2 x y). am. app (infimum_unique H0 H14 H13). assert (Ha: lattice r). red. ee. am. rww H. assert (Hd:forall x y, sub x y -> sub y a -> (inc (complement y x) (substrate r) & sup r x (complement y x) = y & inf r x (complement y x) = the_least_element r)). ir. assert (sub x a). apply sub_trans with y. am. am. assert (inc (complement y x) (powerset a)). app powerset_inc. red. ir. srwi H15. ee. app H13. split. rw H. am. rw H3. split. rw H10. app union2_complement. app powerset_inc. am. rww H11. app intersection2_complement. app powerset_inc. assert (Hb:relatively_complemented r). red. ee. am. am. ir. ufi r H12. awi H12. exists (complement y x). app Hd. eee. eee. assert (Hc:distributive_lattice3 r). rww Exercise1_16c. red. ir. rwi H H12; rwi H H13; rwi H H14. rw (H10 _ _ H12 H13). assert (inc (union2 x y) (powerset a)). app powerset_inc. red. ir. nin (union2_or H15). assert (sub x a). app powerset_sub. app H17. assert (sub y a). app powerset_sub. app H17. rw (H11 _ _ H14 H15). rw (H11 _ _ H13 H14). assert (inc (intersection2 y z) (powerset a)). app powerset_inc. red. ir. assert (sub y a). app powerset_sub. app H17. ap (intersection2_first H16). rw (H10 _ _ H12 H16). uf r. aw. ee. assert (sub z a). app powerset_sub. red. ir. app H17. ap (intersection2_first H18). red. ir. nin (union2_or H17). assert (sub x a). app powerset_sub. app H19. assert (sub (intersection2 y z) a). app powerset_sub. app H19. red. ir. nin (intersection2_both H17). nin (union2_or H19). app union2_first. app union2_second. app intersection2_inc. split. red. eee. ir. assert (gle r x a). uf r. aw. rwi powerset_inc_rw H12. ee. am. fprops. am. nin (Exercice1_17a Hb Hc H13). ee. cp (the_complement_pr Hb H13). nin H16. app H15. eee. rw H3. wr H3. assert (sub x a). app powerset_sub. assert (sub a a). fprops. ap (Hd _ _ H18 H19). Qed. Lemma Exercise1_17d: forall r x y, boolean_lattice r -> inc x (substrate r) -> inc y (substrate r) -> let ys := (standard_completion r y) in (inf r y (sup r ys x) = inf r y x & sup r y (inf r ys x) = sup r y x & inf r ys (sup r y x) = inf r ys x & sup r ys (inf r y x) = sup r ys x). Proof. ir. cp H. nin H. nin H3. assert (order r). nin H. nin H. am. assert (lattice r). nin H; am. cp (lattice_props H6). simpl in H7. assert (forall x y, inc x (substrate r) -> inc y (substrate r) -> inf r y (sup r (standard_completion r y) x) = inf r y x). ee. ir. cp (standard_completion_pr H2 H16). ee. set (z:=standard_completion r y0) in *. red in H. ee. cp (Exercise1_16b H). ee. cp (H24 H4). red in H25. rww H25. rw H19. cp (least_greatest_pr H5). ee. rww H26. app H8. assert (forall x y, inc x (substrate r) -> inc y (substrate r) -> sup r y (inf r (standard_completion r y) x) = sup r y x). ee. ir. cp (standard_completion_pr H2 H17). ee. set (z:=standard_completion r y0) in *. red in H. ee. cp (Exercise1_16b H). ee. cp (H24 H4). red in H26. rww H26. rw H19. cp (least_greatest_pr H5). ee. rww H11. rww H28. app H7. cp (the_greatest_element_pr H5 H3). red in H31. ee. am. app H7. ee. uf ys. app H8. uf ys. app H9. cp (standard_completion_involutive H2 H1). fold ys in H17. wr H17. app H8. cp (standard_completion_pr H2 H1). ee. am. cp (standard_completion_involutive H2 H1). fold ys in H17. wr H17. app H9. cp (standard_completion_pr H2 H1). ee. am. Qed. Lemma Exercise1_17e: forall r f y, boolean_lattice r -> complete_lattice r -> inc y (substrate r) -> fgraph f -> sub (range f) (substrate r)-> inf r y (sup_graph r f) = sup_graph r (L (domain f) (fun x => inf r y (V x f))). Proof. ir. set (v:= inf r y (sup_graph r f)). set (g:= L (domain f) (fun x : Set => inf r y (V x f))). set (u:= sup_graph r g). red in H0. ee. assert (has_sup_graph r f). red. nin (H4 _ H3). am. assert (fgraph g). uf g. gprops. assert (sub (range g) (substrate r)). red. uf g. ir. rwi create_range H7. awi H7. nin H7. ee. wr H8. nin H. nin H. cp (lattice_props H). simpl in H11. ee. app H12. app H3. aw. exists x0. app fdefined_lem. nin H2. am. assert (has_sup_graph r g). nin (H4 _ H7). am. cp (is_sup_graph_pr1 H0 H3 H5). cp (is_sup_graph_pr (sup_graph r f) H0 H3 H2). ufi is_sup_graph H10. rwi H10 H9. clear H10. cp (is_sup_graph_pr1 H0 H7 H8). cp (is_sup_graph_pr (sup_graph r g) H0 H7 H6). ufi is_sup_graph H11. rwi H11 H10. clear H11. fold u in H10. ee. assert (lattice r). nin H. nin H. am. cp (lattice_inf_pr H15 H1 H9). fold v in H16. ee. assert (gle r u v). app H18. app H12. ir. uf g. ufi g H19. bwi H19. rw create_V_rewrite. assert (inc (V a f) (substrate r)). app H3. aw. exists a. ap fdefined_lem. am. am. nin H2;am. cp (lattice_inf_pr H15 H1 H20). ee. am. am. app H12. ir. uf g. ufi g H19. bwi H19. rw create_V_rewrite. assert (inc (V a f) (substrate r)). app H3. aw. exists a. ap fdefined_lem. am. am. nin H2;am. cp (lattice_inf_pr H15 H1 H20). ee. assert (gle r (V a f) (sup_graph r f)). app H13. order_tac. am. set (z:= sup r (standard_completion r y) u). assert (inf r y z = u). uf z. cp (Exercise1_17d H H10 H1). simpl in H20. ee. rw H20. cp (lattice_props H15). simpl in H24. ee. rw H27. app inf_comparable1. order_tac. am. am. assert (z = sup r (standard_completion r y) (sup_graph r f)). uf z. set (ys:= standard_completion r y) in *. uf u. assert (inc ys (substrate r)). cp (standard_completion_pr H H1). ee. am. cp (lattice_sup_pr H15 H21 H10). cp (lattice_sup_pr H15 H21 H9). ee. assert (gle r (sup r ys (sup_graph r g)) (sup r ys (sup_graph r f))). ap H27. ap H23. ap H12. cp (lattice_props H15). simpl in H28. ee. app H28. ir. assert (gle r (V a g) (sup_graph r f)). assert (gle r (V a g) (V a f)). uf g. ufi g H28. bwi H28. bw. assert (inc (V a f) (range f)). fprops. cp (H3 _ H29). cp (lattice_inf_pr H15 H1 H30). ee. am. ufi g H28. bwi H28. cp (H13 _ H28). order_tac. order_tac. assert (gle r (sup r ys (sup_graph r f)) (sup r ys (sup_graph r g))). ap H25. ap H22. ap H14. cp (lattice_props H15). simpl in H29. ee. app H29. ir. assert (gle r (inf r y (V a f)) (sup_graph r g)). assert (inf r y (V a f) = V a g). uf g. bw. rw H30. app H11. uf g. bw. assert (gle r (sup r ys (inf r y (V a f))) (sup r ys (sup_graph r g))). cp (lattice_sup_pr H15 H21 (inc_arg1_substrate H30)). cp (lattice_sup_pr H15 H21 (inc_arg2_substrate H30)). ee. ap H36. ap H32. order_tac. assert (inc (V a f) (range f)). fprops. cp (H3 _ H32). cp (Exercise1_17d H H33 H1). simpl in H34. fold ys in H34. ee. rwi H37 H31. cp (lattice_sup_pr H15 (inc_arg1_substrate H23) H33). ee. order_tac. order_tac. wr H20. rw H21. cp (Exercise1_17d H H9 H1). simpl in H22. ee. rw H22. fold v. app refl_equal. Qed. Lemma exercise1_11h: forall r, order r -> (forall X, sub X (substrate r) -> has_infimum r X) -> complete_lattice r. Proof. ir. red. split. am. ir. split. set (Z := (Zo (substrate r) (fun z => upper_bound r X z))). assert (sub Z (substrate r)). uf Z; app Z_sub. nin (H0 _ H2). exists x. rwi (greatest_lower_bound_pr x H H2) H3. nin H3. rw (least_upper_bound_pr x H H1). split. red. split. nin H3. am. ir. app H4. red. split. app H1. ir. ufi Z H6. Ztac. nin H8. app H9. ir. nin H3. app H6. uf Z. Ztac. nin H5. am. ap (H0 _ H1). Qed. Lemma exercise1_11i: forall r, complete_lattice r -> complete_lattice (opposite_order r). Proof. ir. red in H. ee. red. ee. fprops. rww substrate_opposite_order. ir. cp (H0 _ H1). nin H2. ee. nin H3. exists x. wrr inf_sup_opp. nin H2. exists x. wrr sup_inf_opp. Qed. Definition intersection_partition2 u v := Zo (intersection_covering2 u v) (fun z => nonempty z). Lemma intersection_is_partition2: forall u v x, partition u x -> partition v x -> partition (intersection_partition2 u v) x. Proof. ir. uf intersection_partition2. red. red in H; red in H0. ee. set_extens. nin (union_exists H5). ee. Ztac. rwi intersection_covering2_pr H8. nin H8; nin H8. ee. wr H. apply union_inc with x2. wri H11 H6. inter2tac. am. cp H5. wri H H5. wri H0 H6. nin (union_exists H5). nin (union_exists H6). ee. assert (inc x0 (intersection2 x1 x2)). app intersection2_inc. apply union_inc with (intersection2 x1 x2). am. Ztac. rw intersection_covering2_pr. exists x1. exists x2. eee. ex_tac. ir. Ztac. am. ir. Ztac. clear H6. Ztac. clear H5. rwi intersection_covering2_pr H6. rwi intersection_covering2_pr H7. nin H6. nin H5. nin H7. nin H6. ee. wr H12; wr H10. nin (H4 _ _ H5 H6). nin (H2 _ _ H7 H11). rw H13; rww H14. au. right. app disjoint_pr. ir. red in H14. elim (emptyset_pr (x:= u0)). wr H14. app intersection2_inc. inter2tac. inter2tac. right. app disjoint_pr. ir. red in H13. elim (emptyset_pr (x:= u0)). wr H13. app intersection2_inc. inter2tac. inter2tac. Qed. Lemma intersection_is_inf2: forall u v x, partition u x -> partition v x -> least_upper_bound (coarser x)(doubleton u v)(intersection_partition2 u v). Proof. ir. app least_upper_bound_doubleton. app coarser_order. rw related_coarser. ee. am. app intersection_is_partition2. red. ir. ufi intersection_partition2 H1. Ztac. rwi intersection_covering2_pr H2. nin H2. nin H2. ee. ex_tac. wr H5. app intersection2sub_first. rw related_coarser. ee. am. app intersection_is_partition2. red. ir. ufi intersection_partition2 H1. Ztac. rwi intersection_covering2_pr H2. nin H2. nin H2. ee. ex_tac. wr H5. app intersection2sub_second. intro. do 3 rw related_coarser. ir. ee. app intersection_is_partition2. am. red. ir. red in H6; red in H4. cp (H6 _ H7). cp (H4 _ H7). nin H8. nin H9. ee. assert (sub a (intersection2 x0 x1)). red. ir. app intersection2_inc. app H11. app H10. exists (intersection2 x0 x1). uf intersection_partition2. ee. Ztac. rw intersection_covering2_pr. exists x0; exists x1. eee. red in H5. ee. nin (H13 _ H7). cp (H12 _ H15). ex_tac. am. Qed. Definition intersection_partition f := complement (fun_image (productb f) intersectionb) (singleton (emptyset)). Lemma intersection_is_partition: forall f x, fgraph f -> (forall u, inc u (domain f) -> partition (V u f) x) -> nonempty (domain f) -> partition (intersection_partition f) x. Proof. ir. uf intersection_partition. red. ee. set_extens. nin (union_exists H2). ee. srwi H4. ee. awi H4. nin H4. ee. wri H6 H3. rwi productb_pr H4. ee. wri H7 H1. nin H1. cp (H8 _ H1). cp (intersectionb_forall H3 H1). wri H7 H0. cp (H0 _ H1). red in H11. ee. wr H11. union_tac. am. set (g := fun u => choose (fun v => inc x0 v & inc v (V u f))). assert (forall u, inc u (domain f) -> (inc x0 (g u) & inc (g u) (V u f))). ir. uf g. app choose_pr. cp (H0 _ H3). red in H4. ee. wri H4 H2. ap (union_exists H2). set (h:= L (domain f) g). assert (inc x0 (intersectionb h)). uf h. app intersectionb_inc. nin H1. exists (J y (V y h)). uf h. app fdefined_lem. gprops. bw. bw. ir. bw. nin (H3 _ H4). am. assert (inc h (productb f)). rw productb_pr. uf h. ee. gprops. bw. bw. ir. bw. nin (H3 _ H5). am. am. apply union_inc with (intersectionb h). am. srw. ee. aw. ex_tac. red. ir. rwi (singleton_eq H6) H4. elim (emptyset_pr H4). ir. srwi H2. ee. nin (emptyset_dichot a). elim H3. rw H4. fprops. am. ir. srwi H2. srwi H3. ee. awi H2. nin H2. awi H3. nin H3. ee. wr H7. wr H6. nin (equal_or_not x0 x1). rww H8. au. right. app disjoint_pr. ir. elim H8. rw (productb_extensionality H H2 H3). ir. rwi productb_pr H2. rwi productb_pr H3. ee. assert (inc u (V i x0)). app intersectionb_forall. ue. assert (inc u (V i x1)). app intersectionb_forall. ue. cp (H0 _ H11). red in H18. ee. rwi H14 H15. rwi H12 H13. nin (H20 _ _ (H15 _ H11) (H13 _ H11)). am. red in H21. elim (emptyset_pr (x:=u)). wr H21. app intersection2_inc. am. am. Qed. Lemma intersection_is_inf: forall f x, fgraph f -> (forall u, inc u (domain f) -> partition (V u f) x) -> nonempty (domain f) -> least_upper_bound (coarser x) (range f) (intersection_partition f). Proof. ir. rw least_upper_bound_pr. ee. red. rw substrate_coarser. split. rw set_of_partition_pr. app intersection_is_partition. ir. rwi frange_inc_rw H2. nin H2. ee. rw related_coarser. ee. rw H3. app H0. app intersection_is_partition. red. ir. ufi intersection_partition H4. srwi H4. ee. awi H4. nin H4. ee. rwi productb_pr H4. ee. exists (V x0 x1). ee. rw H3. app H8. ue. wr H6. red. ir. wri H7 H2. ap (intersectionb_forall H9 H2). am. am. ir. red in H2. ee. rwi substrate_coarser H2. rwi set_of_partition_pr H2. rw related_coarser. split. app intersection_is_partition. split. am. red. ir. set (g:= fun u => choose (fun z => inc z (V u f) & sub a z)). assert (forall u, inc u (domain f) -> (inc (g u) (V u f) & sub a (g u))). ir. assert (inc (V u f) (range f)). app inc_V_range. cp (H3 _ H6). rwi related_coarser H7. uf g. app choose_pr. ee. red in H9. app (H9 _ H4). assert (sub a (intersectionb (L (domain f) g))). red. ir. app intersectionb_inc. nin H1. exists (J y (V y (L (domain f) g))). app fdefined_lem. gprops. bw. bw. ir. bw. nin (H5 _ H7). app H9. exists (intersectionb (L (domain f) g)). split. uf intersection_partition. srw. ee. aw. exists (L (domain f) g). ee. rw productb_pr. ee. gprops. bw. bw. ir. bw. nin (H5 _ H7). am. am. tv. red. ir. rwi (singleton_eq H7) H6. red in H2. ee. nin (H8 _ H4). cp (H6 _ H10). elim (emptyset_pr H11). am. app coarser_order. rw substrate_coarser. red. ir. rw set_of_partition_pr. rwi frange_inc_rw H2. nin H2. ee. rw H3. app H0. am. Qed. Lemma exercise1_18a: forall E, complete_lattice (coarser E). Proof. ir. cp (coarser_order E). app exercise1_11b. ir. nin (emptyset_dichot X). rw H1. nin (emptyset_dichot E). red. exists emptyset. rw least_upper_bound_emptyset. red. assert (partition emptyset E). red. ee. rw H2. app is_emptyset. red. ir. nin (union_exists H3). ee. elim (emptyset_pr H5). ir. elim (emptyset_pr H3). ir. elim (emptyset_pr H3). rw substrate_coarser. ee. rw set_of_partition_pr. am. ir. rwi set_of_partition_pr H4. rw related_coarser. ee. am. am. red. ir. red in H4. ee. cp (H6 _ H5). nin H8. assert (inc y (union x)). union_tac. rwi H4 H9. rwi H2 H9. elim (emptyset_pr H9). am. red. exists (smallest_partition E). rw least_upper_bound_emptyset. red. rw substrate_coarser. ee. rw set_of_partition_pr. app partition_smallest. ir. rwi set_of_partition_pr H3. rw related_coarser. ee. app partition_smallest. am. red. ir. exists E. ee. uf smallest_partition. fprops. red in H3. ee. wr H3. app union_sub. am. set (f := diagonal X). assert (fgraph f). uf f. app fgraph_diagonal. assert (range f = X). uf f. app range_diagonal. assert (domain f = X). uf f. app domain_diagonal. wri H4 H1. assert (forall u, inc u (domain f) -> partition (V u f) E). rw H4. ir. uf f. rww V_diagonal. cp (H0 _ H5). rwi substrate_coarser H6. rwi set_of_partition_pr H6. am. cp (intersection_is_inf H2 H5 H1). rwi H3 H6. exists (intersection_partition f). am. Qed. Definition big_set3 E := (exists x, exists y, exists z, inc x E & inc y E & inc z E & x <> y & x <> z & y <> z). Lemma exercise1_18b: forall E, big_set3 E -> ~ (distributive_lattice1 (coarser E)). Proof. ir. destruct H as [ x [y [z H]]]. ee. set (o:= tack_on (doubleton x y) z). assert (forall i, inc i o = (i = x \/ i = y \/ i = z)). ir. ap iff_eq. uf o. ir. nin (tack_on_or H5). nin (doubleton_or H6); au. au. uf o. ir. nin H5. rw H5. app inc_tack_on_y. fprops. nin H5. rw H5. app inc_tack_on_y. fprops. rw H5. fprops. set (F:=complement E o). set (with_F := fun u => complement (tack_on u F) (singleton emptyset)). assert (forall u, ~ (inc emptyset u) -> F = emptyset -> with_F u = u). ir. uf with_F. set_extens. srwi H8. ee. nin (tack_on_or H8). am. elim H9. ue. ue. srw. ee. fprops. dneg. wrr (singleton_eq H9). assert (forall u, ~ (inc emptyset u) -> F <> emptyset -> with_F u = tack_on u F). ir. uf with_F. set_extens. srwi H9. ee; am. srw. ee. am. red. ir. rwi (singleton_eq H10) H9. nin (tack_on_or H9). contradiction. symmetry in H11. contradiction. assert (sub o E). red. ir. rwi H5 H8. nin H8. ue. nin H8; ue. assert (forall u, partition u o -> partition (with_F u) E). ir. red. ee. set_extens. nin (union_exists H10). ee. ufi with_F H12. srwi H12. ee. nin (tack_on_or H12). red in H9. ee. app H8. wr H9. union_tac. rwi H14 H11. ufi F H11. srwi H11. ee. am. nin H9. ee. assert (~ inc emptyset u). red. ir. cp (H11 _ H13). nin H14. elim (emptyset_pr H14). nin (equal_or_not F emptyset). rw H6. rw H9. nin (inc_or_not x0 o). am. elim (emptyset_pr (x:=x0)). wr H14. uf F. srw. eee. am. am. rww H7. nin (inc_or_not x0 o). wri H9 H15. nin (union_exists H15). ee. apply union_inc with x1. am. fprops. apply union_inc with F. uf F. srw. eee. fprops. uf with_F. ir. srwi H10. ee. nin (emptyset_dichot a). elim H11. rw H12. fprops. am. uf with_F. ir. srwi H10. srwi H11. ee. nin (equal_or_not a b). au. right. red in H9. ee. nin (tack_on_or H10). nin (tack_on_or H11). nin (H16 _ _ H17 H18). contradiction. am. app disjoint_pr. ir. assert (inc u0 o). wr H9. union_tac. assert (inc u0 F). ue. ufi F H22. srwi H22. ee. contradiction. nin (tack_on_or H11). app disjoint_pr. ir. assert (inc u0 o). wr H9. union_tac. assert (inc u0 F). ue. ufi F H22. srwi H22. ee. contradiction. elim H14. ue. assert (forall a b c, union2 (singleton a) (doubleton b c) = o -> a <> b -> a <> c -> b <> c -> partition (doubleton (singleton a) (doubleton b c)) o). ir. red. ee. am. ir. nin (doubleton_or H14). rw H15. fprops. rw H15. app nonempty_doubleton. ir. nin (doubleton_or H14); rw H16. nin (doubleton_or H15); rw H17; au. right. app disjoint_pr. ir. rwi (singleton_eq H18) H19. nin (doubleton_or H19). contradiction. contradiction. nin (doubleton_or H15); rw H17; au. right. app disjoint_pr. ir. rwi (singleton_eq H19) H18. nin (doubleton_or H18). contradiction. contradiction. assert (Ha:partition (doubleton (singleton x) (doubleton y z)) o). app H10. set_extens. rw H5. nin (union2_or H11). rw (singleton_eq H12). au. nin (doubleton_or H12); rw H13; au. rwi H5 H11. nin H11. rw H11. app union2_first. fprops. nin H11; rw H11; app union2_second; fprops. assert (Hb: partition (doubleton (singleton y) (doubleton x z)) o). app H10. set_extens. rw H5. nin (union2_or H11). rw (singleton_eq H12). au. nin (doubleton_or H12); rw H13; au. rwi H5 H11. nin H11. rw H11. app union2_second. fprops. nin H11; rw H11. app union2_first; fprops. app union2_second; fprops. intuition. assert (Hc:partition (doubleton (singleton z) (doubleton y x)) o). app H10. set_extens. rw H5. nin (union2_or H11). rw (singleton_eq H12). au. nin (doubleton_or H12); rw H13; au. rwi H5 H11. nin H11. rw H11. app union2_second; fprops. nin H11; rw H11. app union2_second; fprops. app union2_first; fprops. intuition. intuition. intuition. set (Px:= with_F (doubleton (singleton x) (doubleton y z))). assert (partition Px E). uf Px. app H9. set (Py:= with_F (doubleton (singleton y) (doubleton x z))). assert (partition Py E). uf Py. app H9. set (Pz:= with_F (doubleton (singleton z) (doubleton y x))). assert (partition Pz E). uf Pz. app H9. set (alpha:= with_F(largest_partition o)). assert (partition alpha E). uf alpha. app H9. app partition_largest. set (omega:= with_F(smallest_partition o)). assert (partition omega E). uf omega. app H9. app partition_smallest. uf o. exists z. fprops. red. ir. red in H16. rwi substrate_coarser H16. assert (inc Px (set_of_partition_set E)). rw set_of_partition_pr. am. assert (inc Py (set_of_partition_set E)). rw set_of_partition_pr. am. assert (inc Pz (set_of_partition_set E)). rw set_of_partition_pr. am. cp (H16 _ _ _ H17 H18 H19). clear H16; clear H17; clear H18; clear H19. assert (forall u v, partition u E -> partition v E -> sup (coarser E) u v = (intersection_partition2 u v)). ir. cp (intersection_is_inf2 H16 H17). set (w:= intersection_partition2 u v) in *. assert (has_supremum (coarser E) (doubleton u v)). exists w. am. uf sup. assert (sub (doubleton u v) (substrate (coarser E))). rw substrate_coarser. red. ir. rw set_of_partition_pr. nin (doubleton_or H21); rww H22. cp (supremum_pr1 (coarser_order E) H21 H19). app (supremum_unique (coarser_order E) H22 H18). assert (sup (coarser E) Px Py = alpha). rww H16. uf intersection_partition2. uf alpha. set_extens. Ztac. rwi intersection_covering2_pr H18. nin H18. nin H18. ufi Px H18; ufi Py H18. ee. ufi with_F H18. srwi H18. ufi with_F H21. srwi H21. ee. uf with_F. srw. ee. nin (tack_on_or H18). assert (inc x0 (largest_partition o)). rw largest_partition_pr. wri H22 H19. nin H19. nin (intersection2_both H19). nin (doubleton_or H25). exists x. ee. rw H5; au. wr H28. wr H22. set_extens. app intersection2_inc. rwi H28 H26. rwi H28 H29. awi H26. awi H29. rw H29. ue. inter2tac. nin (tack_on_or H21). nin (doubleton_or H29). exists y. ee. rw H5; au. wr H30. wr H22. set_extens. app intersection2_inc. rwi H30 H31. rwi H30 H27. awi H27. awi H31. rw H31. ue. inter2tac. exists z. ee. rw H5; au. wr H22. rw H28; rw H30. set_extens. rw (singleton_eq H31). app intersection2_inc. fprops. fprops. nin (intersection2_both H31). nin (doubleton_or H32). nin (doubleton_or H33). rwi H35 H34. contradiction. rw H35. fprops. rw H34. fprops. rwi H28 H26. rwi H29 H27. ufi F H27. srwi H27. ee. elim H30. rw H5. nin (doubleton_or H26); au. fprops. nin (tack_on_or H21). wri H22 H19. nin H19. rwi H25 H19. nin (intersection2_both H19). ufi F H27. srwi H27. ee. elim H29. rw H5. nin (doubleton_or H26). rwi H30 H28. awi H28. rw H28. au. rwi H30 H28. nin (doubleton_or H28); rw H31; au. rwi H25 H22. rwi H26 H22. rwi intersection2idem H22. wr H22. fprops. red. ir. awi H25. rwi H25 H19. nin H19. elim (emptyset_pr H19). Ztac. rw intersection_covering2_pr. ufi with_F H17. srwi H17. ee. nin (tack_on_or H17). rwi largest_partition_pr H19. nin H19. nin H19. rwi H5 H19. nin H19. exists x0. exists (doubleton x z). ee. uf Px. uf with_F. srw. ee. wr H19. rw H21. app inc_tack_on_y. fprops. am. uf Py. uf with_F. srw. ee. app inc_tack_on_y. fprops. red. ir. awi H22. elim (emptyset_pr (x:=x)). wr H22. fprops. wr H21. rw H19. set_extens. inter2tac. app intersection2_inc. awi H22. rw H22. fprops. nin H19. exists (doubleton y z). exists x0. ee. uf Px. uf with_F. srw. ee. app inc_tack_on_y. fprops. red. ir. awi H22. elim (emptyset_pr (x:=y)). wr H22. fprops. uf Py. wr H21. wr H19. uf with_F. srw. ee. app inc_tack_on_y. fprops. rww H21. wr H21. rw H19. set_extens. inter2tac. app intersection2_inc. awi H22. rw H22. fprops. rwi H19 H21. exists (doubleton y z). exists (doubleton x z). ee. uf Px. uf with_F. srw. ee. app inc_tack_on_y. fprops. red. ir. awi H22. elim (emptyset_pr (x:=y)). wr H22. fprops. uf Py. uf with_F. srw. ee. app inc_tack_on_y. fprops. red. ir. awi H22. elim (emptyset_pr (x:=x)). wr H22. fprops. wr H21. set_extens. nin (intersection2_both H22). nin (doubleton_or H23). nin (doubleton_or H24). rwi H26 H25. contradiction. rw H26. fprops. rw H25. fprops. awi H22. rw H22. app intersection2_inc. fprops. fprops. exists F. exists F. ee. uf Px. uf with_F. srw. ee. fprops. ue. uf Py. uf with_F. srw. ee. fprops. ue. rw H19. rw intersection2idem. tv. ufi with_F H17. srwi H17. ee. nin (emptyset_dichot x0). rwi H19 H18. elim H18. fprops. am. rwi H17 H20. clear H17. assert (sup (coarser E) Px Pz = alpha). rww H16. uf intersection_partition2. uf alpha. set_extens. Ztac. rwi intersection_covering2_pr H18. nin H18. nin H18. ufi Px H18; ufi Pz H18. ee. ufi with_F H18. srwi H18. ufi with_F H21. srwi H21. ee. uf with_F. srw. ee. nin (tack_on_or H18). assert (inc x0 (largest_partition o)). rw largest_partition_pr. wri H22 H19. nin H19. nin (intersection2_both H19). nin (doubleton_or H25). exists x. ee. rw H5; au. wr H28. wr H22. set_extens. app intersection2_inc. rwi H28 H26. rwi H28 H29. awi H26. awi H29. rw H29. ue. inter2tac. nin (tack_on_or H21). nin (doubleton_or H29). exists z. ee. rw H5; au. wr H30. wr H22. set_extens. app intersection2_inc. rwi H30 H31. rwi H30 H27. awi H27. awi H31. rw H31. ue. inter2tac. exists y. ee. rw H5; au. wr H22. rw H28; rw H30. set_extens. rw (singleton_eq H31). app intersection2_inc. fprops. fprops. nin (intersection2_both H31). nin (doubleton_or H32). rw H34. fprops. nin (doubleton_or H33). rw H35. fprops. rwi H35 H34. contradiction. rwi H28 H26. rwi H29 H27. ufi F H27. srwi H27. ee. elim H30. rw H5. nin (doubleton_or H26); au. fprops. nin (tack_on_or H21). wri H22 H19. nin H19. rwi H25 H19. nin (intersection2_both H19). ufi F H27. srwi H27. ee. elim H29. rw H5. nin (doubleton_or H26). rwi H30 H28. awi H28. rw H28. au. rwi H30 H28. nin (doubleton_or H28); rw H31; au. rwi H25 H22. rwi H26 H22. rwi intersection2idem H22. wr H22. fprops. red. ir. awi H25. rwi H25 H19. nin H19. elim (emptyset_pr H19). Ztac. rw intersection_covering2_pr. ufi with_F H17. srwi H17. ee. nin (tack_on_or H17). rwi largest_partition_pr H19. nin H19. nin H19. rwi H5 H19. nin H19. exists x0. exists (doubleton y x). ee. uf Px. uf with_F. srw. ee. wr H19. rw H21. app inc_tack_on_y. fprops. am. uf Pz. uf with_F. srw. ee. app inc_tack_on_y. fprops. red. ir. awi H22. elim (emptyset_pr (x:=x)). wr H22. fprops. wr H21. rw H19. set_extens. inter2tac. app intersection2_inc. awi H22. rw H22. fprops. nin H19. exists (doubleton y z). exists (doubleton y x). ee. uf Px. uf with_F. srw. ee. app inc_tack_on_y. fprops. red. ir. awi H22. elim (emptyset_pr (x:=y)). wr H22. fprops. uf Pz. uf with_F. srw. ee. app inc_tack_on_y. fprops. red. ir. awi H22. elim (emptyset_pr (x:=y)). wr H22. fprops. wr H21. wr H19. set_extens. nin (intersection2_both H22). nin (doubleton_or H23). rw H25. fprops. nin (doubleton_or H24). rw H26. fprops. rwi H26 H25. contradiction. app intersection2_inc. awi H22. rw H22. fprops. awi H22. rw H22. fprops. exists (doubleton y z). exists x0. ee. uf Px. uf with_F. srw. ee. app inc_tack_on_y. fprops. red. ir. awi H22. elim (emptyset_pr (x:=y)). wr H22. fprops. uf Pz. uf with_F. srw. ee. wr H21. rw H19. app inc_tack_on_y. fprops. am. wr H21. set_extens. inter2tac. app intersection2_inc. awi H22. rw H22. rw H19. fprops. exists F. exists F. ee. uf Px. uf with_F. srw. ee. fprops. ue. uf Pz. uf with_F. srw. ee. fprops. ue. rw H19. rw intersection2idem. tv. ufi with_F H17. srwi H17. ee. nin (emptyset_dichot x0). rwi H19 H18. elim H18. fprops. am. rwi H17 H20. clear H17. assert(inf (coarser E) alpha alpha = alpha). app inf_comparable1. app coarser_order. wrr order_reflexivity. rw substrate_coarser. rww set_of_partition_pr. app coarser_order. rwi H17 H20. clear H17. assert (forall u, partition u o -> gle (coarser E) omega (with_F u)). ir. rw related_coarser. ee. am. app H9. red. ir. ufi with_F H18. srwi H18. ee. nin (tack_on_or H18). exists o. uf omega. ee. uf with_F. srw. uf smallest_partition. ee. fprops. red. ir. awi H22. elim (emptyset_pr (x:=z)). wr H22. uf o. fprops. red in H17. ee. wr H17. app union_sub. exists F. ee. uf omega. uf with_F. srw. ee. fprops. ue. ue. set (w:= inf (coarser E) Py Pz) in *. assert (w = omega). assert (order (coarser E)). app coarser_order. assert (inc Py (substrate (coarser E))). rw substrate_coarser. rw set_of_partition_pr. am. assert (inc Pz (substrate (coarser E))). rw substrate_coarser. rw set_of_partition_pr. am. assert (has_infimum (coarser E) (doubleton Py Pz)). nin (exercise1_18a E). assert (sub (doubleton Py Pz) (substrate (coarser E))). red. ir. nin (doubleton_or H24); ue. nin (H23 _ H24). am. cp (inf_pr H18 H19 H21 H22). fold w in H23. ee. assert (gle (coarser E) omega w). app H25. uf Py. app H17. uf Pz. app H17. assert (gle (coarser E) w omega). rw related_coarser. rwi related_coarser H23. ee. am. am. rwi related_coarser H24. ee. red in H28; red in H30; red. ir. ufi omega H31. ufi smallest_partition H31. ufi with_F H31. srwi H31. ee. nin (tack_on_or H31). awi H33. assert (inc (doubleton x z) Py). uf Py. uf with_F. srw. ee. app inc_tack_on_y. fprops. red. ir. awi H34. elim (emptyset_pr (x:=x)). wr H34. fprops. assert (inc (doubleton y x) Pz). uf Pz. uf with_F. srw. ee. app inc_tack_on_y. fprops. red. ir. awi H35. elim (emptyset_pr (x:=x)). wr H35. fprops. nin (H28 _ H34). nin (H30 _ H35). ee. red in H24. ee. nin (H41 _ _ H36 H37). exists x0. ee. am. rw H33. red. ir. rwi H5 H43. nin H43. rw H43. app H39. fprops. nin H43. rw H43. rw H42. app H38. fprops. rw H43. app H39. fprops. elim (emptyset_pr (x:=x)). red in H42. wr H42. app intersection2_inc. app H39. fprops. app H38. fprops. assert (inc F Py). uf Py. uf with_F. srw. ee. fprops. ue. cp (H28 _ H34). rw H33. am. order_tac. rwi H18 H20. assert (sup (coarser E) omega Px = Px). app sup_comparable1. app coarser_order. uf Px. app H17. assert (sup (coarser E) Px omega = Px). uf sup. rw doubleton_symm. am. rwi H21 H20. assert (inc (doubleton y z) alpha). wr H20. uf Px. uf with_F. srw. ee. app inc_tack_on_y. fprops. red. ir. awi H22. elim (emptyset_pr (x:=y)). wr H22. fprops. ufi alpha H22. ufi with_F H22. srwi H22. ee. nin (tack_on_or H22). rwi largest_partition_pr H24. nin H24. ee. elim H4. assert (inc y (singleton x0)). rw H25. fprops. awi H26. rw H26. assert (inc z (singleton x0)). rw H25. fprops. awi H27. rw H27. tv. assert (inc y F). wr H24. fprops. ufi F H25. srwi H25. ee. elim H26. rw H5. au. Qed. Lemma exercise1_18c: forall E, largest_partition E = the_greatest_element (coarser E). Proof. ir. sy. app the_greatest_element_pr2. app coarser_order. red. rw substrate_coarser. ee. rw set_of_partition_pr. app partition_largest. ir. rwi set_of_partition_pr H. rw related_coarser. ee. am. app partition_largest. red. ir. rwi largest_partition_pr H0. nin H0. ee. red in H. ee. wri H H0. nin (union_exists H0). ee. ex_tac. wr H1. red. ir. awi H6. ue. Qed. Lemma exercise1_18d: forall E x y, let r := coarser E in gle r y x -> exists x', inc x' (substrate r) & inf r x x' = y & sup r x x' = largest_partition E. Proof. ir. ufi r H. rwi related_coarser H. ee. red in H1. set (X1 := Zo (powerset E) (fun z=> exists a, z = singleton a & ~(exists u, inc u x & a = rep u))). set (X2 := fun_image y (fun v => Zo E (fun z => exists u, inc u x & sub u v & z = rep u))). set (z:=union2 X1 X2). assert (partition z E). uf z. red. ee. set_extens. nin (union_exists H2). ee. nin (union2_or H4). ufi X1 H5. Ztac. nin H7. ee. rwi H7 H3. awi H3. rwi powerset_inc_rw H6. app H6. rw H7. ue. ufi X2 H5. awi H5. nin H5. ee. assert (sub x1 E). wr H6. app Z_sub. app H7. assert (Ha: inc x0 E). am. red in H0. ee. wri H0 H2. nin (union_exists H2). ee. nin (equal_or_not x0 (rep x1)). cp (H1 _ H6). nin H8. ee. set (w:= Zo E (fun z => exists u , inc u x & sub u x2 & z = rep u)). assert (inc x0 w). uf w. Ztac. exists x1. au. apply union_inc with w. am. app union2_second. uf X2. aw. exists x2. ee. am. tv. apply union_inc with (singleton x0). fprops. app union2_first. uf X1. Ztac. app powerset_inc. red. ir. awi H8. rw H8. am. exists x0. ee. tv. red. ir. nin H8. ee. assert (inc x0 x2). rw H9. ap nonempty_rep. app H3. nin (H4 _ _ H6 H8). elim H7. ue. red in H11. elim (emptyset_pr (x:= x0)). wr H11. app intersection2_inc. ir. nin (union_exists H2). ee. nin (doubleton_or H4). rwi H5 H3. ufi X1 H3. Ztac. nin H7. ee. rw H7. fprops. rwi H5 H3. ufi X2 H3. awi H3. nin H3. ee. wr H6. red in H; ee. nin (H7 _ H3). assert (inc y0 E). wr H. union_tac. red in H0. ee. wri H0 H10. nin (union_exists H10). ee. nin (H1 _ H14). ee. nin (H8 _ _ H3 H15). rw H17. cp (H11 _ H14). exists (rep x2). Ztac. wr H0. apply union_inc with x2. app nonempty_rep. am. exists x2. au. elim (emptyset_pr (x:=y0)). wr H17. app intersection2_inc. app H16. ir. nin (union2_or H2); nin (union2_or H3). ufi X1 H4; ufi X1 H5. Ztac; clear H5; Ztac. nin H7; nin H8. ee. rw H7; rw H8. nin (equal_or_not x1 x0). rw H11. au. right. app disjoint_pr. ir. app H11. awi H12; awi H13. ue. ufi X1 H4; Ztac. nin H7. ufi X2 H5. awi H5. nin H5. ee. right. wr H9. rw H7. app disjoint_pr. ir. Ztac. nin H13. awi H10. elim H8. wr H10. exists x2. eee. right. ufi X1 H5; Ztac. nin H7. ufi X2 H4. awi H4. nin H4. ee. app disjoint_pr. ir. elim H8. rwi H7 H11. awi H11. wr H11. wri H9 H10. Ztac. nin H13. exists x2. eee. ufi X2 H4; ufi X2 H5. awi H4; awi H5; nin H4; nin H5. ee. red in H. ee. nin (H9 _ _ H4 H5). left. wr H7; wr H6; rww H10. right. app disjoint_pr. ir. wri H7 H11. Ztac. wri H6 H12. Ztac. nin H14. nin H16. ee. red in H0. ee. elim (emptyset_pr (x:=u)). wr H10. app intersection2_inc. app H19. rw H20. app nonempty_rep. app H21. app H17. rw H18. app nonempty_rep. app H21. uf r. rw substrate_coarser. exists z. ee. rww set_of_partition_pr. assert (order (coarser E)). app coarser_order. assert (inc x (substrate (coarser E))). rw substrate_coarser. rw set_of_partition_pr. am. assert (inc z (substrate (coarser E))). rw substrate_coarser. rw set_of_partition_pr. am. assert (has_infimum (coarser E) (doubleton x z)). nin (exercise1_18a E). assert (sub (doubleton x z) (substrate (coarser E))). red. ir. nin (doubleton_or H8); ue. nin (H7 _ H8). am. cp (inf_pr H3 H4 H5 H6). set (w:= (inf (coarser E) x z)) in *. ee. assert (gle (coarser E) y w). app H9. rw related_coarser. au. rw related_coarser. eee. red. ir. ufi z H10. nin (union2_or H10). ufi X1 H11. Ztac. nin H13. ee. rw H13. assert (inc x0 E). rwi powerset_inc_rw H12. app H12. rw H13. fprops. red in H. ee. wri H H15. nin (union_exists H15). ee. ex_tac. red. ir. awi H20. rww H20. ufi X2 H11. awi H11. nin H11. ee. exists x0. ee. am. wr H12. red. ir. Ztac. nin H15. ee. rw H17. app H16. app nonempty_rep. red in H0. ee. app H18. assert (gle (coarser E) w y). rw related_coarser. rwi related_coarser H7. ee. am. am. rwi related_coarser H8. ee. red in H12. red in H14. red. ir. set (z1:= Zo E (fun z => exists u, inc u x & sub u a & z = rep u)). assert (inc z1 z). uf z. app union2_second. uf X2. aw. exists a. ee. am. tv. nin (H14 _ H16). exists x0. ee. am. red. ir. assert (inc x1 E). red in H. ee. wr H. union_tac. red in H0. ee. wri H0 H20. nin (union_exists H20). ee. nin (equal_or_not x1 (rep x2)). app H18. uf z1. Ztac. wrr H0. exists x2. ee. am. nin (H1 _ H24). ee. red in H. ee. nin (H29 _ _ H15 H26). rww H30. elim (emptyset_pr (x:=x1)). wr H30. app intersection2_inc. app H27. am. nin (H12 _ H24). ee. red in H7. ee. nin (H29 _ _ H17 H26). rw H30. app H27. elim (emptyset_pr (x:= rep x2)). wr H30. app intersection2_inc. app H18. uf z1. Ztac. wr H0. assert (inc (rep x2) x2). app nonempty_rep. app H21. union_tac. exists x2. ee. am. nin (H1 _ H24). ee. red in H. ee. nin (H34 _ _ H15 H31). rww H35. elim (emptyset_pr (x:=x1)). wr H35. app intersection2_inc. app H32. tv. app H27. app nonempty_rep. app H21. order_tac. assert (sup (coarser E) x z = (intersection_partition2 x z)). ir. cp (intersection_is_inf2 H0 H2). assert (has_supremum (coarser E) (doubleton x z)). exists (intersection_partition2 x z). am. uf sup. assert (sub (doubleton x z) (substrate (coarser E))). rw substrate_coarser. red. ir. rw set_of_partition_pr. nin (doubleton_or H5); rww H6. cp (supremum_pr1 (coarser_order E) H5 H4). app (supremum_unique (coarser_order E) H6 H3). rw H3. set_extens. rw largest_partition_pr. ufi intersection_partition2 H4. Ztac. clear H4. rwi intersection_covering2_pr H5. nin H5. nin H4. ee. ufi z H5. nin (union2_or H5). ufi X1 H8. Ztac. nin H10. ee. assert (inc x3 E). rwi powerset_inc_rw H9. ap H9. ue. exists x3. ee. am. wr H7. wr H10. set_extens. app intersection2_inc. rwi H10 H13. awi H13. rw H13. nin H6. wri H7 H6. nin (intersection2_both H6). rwi H10 H15. awi H15. ue. inter2tac. ufi X2 H8. awi H8. nin H8. ee. nin (H1 _ H4). ee. nin H6. wri H7 H6. nin (intersection2_both H6). exists y0. ee. red in H0. ee. wr H0. union_tac. set_extens. awi H14. rw H14. wr H7. app intersection2_inc. aw. wri H9 H13. Ztac. nin H16. wri H7 H14. nin (intersection2_both H14). wri H9 H18. Ztac. nin H20. ee. red in H0. ee. nin (H26 _ _ H16 H20). rw H22; rw H24; rww H27. assert (inc y0 x6). rw H24. app nonempty_rep. app H25. nin (H26 _ _ H4 H16). assert (inc x5 x7). rw H22. app nonempty_rep. app H25. nin (H26 _ _ H4 H20). elim (emptyset_pr (x:=x5)). wr H27. app intersection2_inc. ue. elim (emptyset_pr (x:=x5)). wr H31. app intersection2_inc. elim (emptyset_pr (x:=y0)). wr H29. app intersection2_inc. rwi largest_partition_pr H4. nin H4. ee. red in H0. ee. cp H4. wri H0 H8. nin (union_exists H8). ee. nin (H1 _ H10). ee. uf intersection_partition2. Ztac. rw intersection_covering2_pr. exists x2. nin (equal_or_not x1 (rep x2)). set (w:= Zo E (fun z => exists u, inc u x & sub u x3 & z = rep u)). exists w. ee. am. uf z. app union2_second. uf X2. aw. exists x3. eee. wr H5. set_extens. aw. nin (intersection2_both H14). ufi w H16. Ztac. nin H18. ee. rw H20. rw H13. nin (H7 _ _ H10 H18). ue. elim (emptyset_pr (x:=x4)). wr H21. app intersection2_inc. rw H20. app nonempty_rep. app H6. awi H14. rw H14. app intersection2_inc. uf w. Ztac. exists x2. au. exists (singleton x1). ee. am. uf z. app union2_first. uf X1. Ztac. app powerset_inc. red. ir. awi H14. rww H14. exists x1. ee. tv. red. ir. nin H14. ee. assert (inc x1 x4). rw H15. app nonempty_rep. app H6. nin (H7 _ _ H10 H14). elim H13. ue. elim (emptyset_pr (x:= x1)). wr H17. app intersection2_inc. wr H5. set_extens. inter2tac. app intersection2_inc. awi H14. ue. wr H5. fprops. Qed. (* Exercise 19 *) Definition without_gaps r := order r & (exists x, exists y, glt r x y) & (forall x y, glt r x y -> exists z, glt r x z & glt r z y). Lemma Exercise1_19a: forall r f g, ordinal_sum_axioms1 r f g -> (without_gaps (ordinal_sum r f g) = (((exists i, exists j, glt r i j) \/ (exists i, exists x, exists y, inc i (substrate r) & glt (V i g) x y)) & (forall i x y, inc i (substrate r) -> glt (V i g) x y -> without_gaps (V i g)) & (forall i j, glt r i j -> (exists k, glt r i k & glt r k j) \/ (forall u, ~ (maximal_element (V i g) u)) \/ (forall u, ~ (minimal_element (V j g) u))))). Proof. ir. nin H. rename H0 into Hw. cp H. red in H. ee. ap iff_eq. ir. red in H7. ee. nin H8. nin H8. nin H8. awi H8. ee. nin H12. left. exists (Q x). exists (Q x0). am. right. exists (Q x). exists (P x). exists (P x0). ee. nin (du_index_pr H8). ue. split. am. dneg. nin (du_index_pr H8). nin (du_index_pr H11). ee. app pair_extensionality. am. ir. red. ee. app H5. ue. exists x; exists y; am. ir. assert (glt (ordinal_sum r f g) (J x0 i) (J y0 i)). nin H12. split. aw. ee. app inc_disjoint_union. ue. wr H6. order_tac. ue. app inc_disjoint_union. ue. wr H6. order_tac. ue. right. ee. tv. am. dneg. inter2tac. nin (H9 _ _ H13). ee. nin H14. nin H15. awi H14. awi H15. ee. exists (P x1). nin H21. nin H19. order_tac. nin H19. rwi H19 H21. nin H21. elim H23. tv. nin H19. nin H21. rwi H21 H19. nin H19. elim H23. tv. nin H21. nin H19. ee. split. am. red. ir. elim H16. nin (du_index_pr H20). ee. app pair_extensionality. fprops. aw. aw. split. rww H21. red. ir. elim H17. nin (du_index_pr H20). ee. app pair_extensionality. fprops. aw. aw. am. am. ir. nin (p_or_not_p (forall u : Set, ~ maximal_element (V i g) u)). au. assert (exists u, maximal_element (V i g) u). app exists_proof. nin (p_or_not_p (forall u : Set, ~ minimal_element (V j g) u)). au. assert (exists u, minimal_element (V j g) u). app exists_proof. clear H11. clear H13. nin H12. nin H14. nin H11. nin H12. ee. assert (inc i (domain f)). wr H1. order_tac. assert (inc j (domain f)). wr H1. order_tac. assert (inc (J x i) (disjoint_union f)). app inc_disjoint_union. wrr H6. assert (inc (J x0 j) (disjoint_union f)). app inc_disjoint_union. wrr H6. assert (glt (ordinal_sum r f g) (J x i) (J x0 j)). split. aw. ee. am. am. au. nin H10. dneg. inter2tac. nin (H9 _ _ H19). ee. nin H20. nin H21. awi H20. awi H21. ee. nin H27. nin H25. left. exists (Q x1). au. nin H25. assert (inc (P x1) (substrate (V j g))). rwi H25 H28. cp (H5 _ H16). order_tac. rwi H25 H28. cp (H14 _ H28). elim H23. cp (du_index_pr H21). ee. app pair_extensionality. fprops. aw. aw. ee. assert (inc (P x1) (substrate (V i g))). cp (H5 _ H15). order_tac. cp (H13 _ H28). elim H22. cp (du_index_pr H26). ee. app pair_extensionality. fprops. sy. aw. aw. am. am. ir. ee. red. ee. fprops. nin H7. nin H7. nin H7. assert (inc x (domain f)). wr H1. order_tac. assert (inc x0 (domain f)). wr H1. order_tac. nin (Hw _ H10). nin (Hw _ H11). exists (J y x). exists (J y0 x0). split. aw. ee. app inc_disjoint_union. app inc_disjoint_union. au. nin H7. dneg. inter2tac. nin H7. nin H7. nin H7. ee. exists (J x0 x). exists (J x1 x). rwi H1 H7. split. aw. ee. app inc_disjoint_union. wr H6. order_tac. am. app inc_disjoint_union. wr H6. order_tac. am. right. nin H10. au. nin H10. dneg. inter2tac. ir. nin H10. awi H10. ee. nin H13. nin (H9 _ _ H13). nin H14. assert (inc x0 (domain f)). ee. wr H1. order_tac. nin (Hw _ H15). assert (inc (J y0 x0) (disjoint_union f)). ap inc_disjoint_union. am. am. exists (J y0 x0). ee. split. aw. eee. red. ir. nin H14. elim H20. rw H19. aw. split. aw. eee. red. ir. nin H18. elim H20. wr H19. aw. nin H14. cp (H14 (P x)). ufi maximal_element H15. assert (exists z, inc z (substrate (V (Q x) g)) & glt (V (Q x) g) (P x) z). app exists_proof. red. ir. elim H15. ee. cp (du_index_pr H10). ee. rww H6. ir. nin (equal_or_not x0 (P x)). am. elim (H16 x0). ee. order_tac. split. am. au. nin H16. assert (inc (J x0 (Q x)) (disjoint_union f)). app inc_disjoint_union. cp (du_index_pr H10). ee. am. ee. wrr H6. wr H1. order_tac. exists (J x0 (Q x)). ee; split. aw. ee. am. am. right. ee. tv. nin H18. am. nin H18. red. ir. elim H19. rw H20. aw. aw. ee. am. am. au. red. ir. nin H18. elim H11. wr H19. nin H13. elim H21. wr H19. aw. cp (H14 (P y)). ufi minimal_element H15. assert (exists z, inc z (substrate (V (Q y) g)) & glt (V (Q y) g) z (P y)). app exists_proof. red. ir. elim H15. ee. cp (du_index_pr H12). ee. rww H6. ir. nin (equal_or_not x0 (P y)). am. elim (H16 x0). ee. order_tac. split. am. au. nin H16. assert (inc (J x0 (Q y)) (disjoint_union f)). app inc_disjoint_union. cp (du_index_pr H12). ee. am. ee. wrr H6. wr H1. order_tac. exists (J x0 (Q y)). ee; split. aw. ee. am. am. au. nin H13. red. ir. elim H19. rw H20. aw. aw. ee. am. am. right. ee. tv. nin H18. am. nin H18. red. ir. elim H19. wr H20. aw. assert (glt (V (Q x) g) (P x) (P y)). red. ee. am. dneg. cp (du_index_pr H10); cp (du_index_pr H12);ee. app pair_extensionality. assert (inc (Q x) (substrate r)). cp (du_index_pr H10); ee. ue. nin (H8 _ _ _ H15 H14). ee. nin (H18 _ _ H14). assert (inc (J x0 (Q x)) (disjoint_union f)). app inc_disjoint_union. ue. nin H20. wr H6. order_tac. ue. exists (J x0 (Q x)). ee; split. aw. eee. right. ee. tv. nin H20. am. nin H20. red. ir. elim H23. rw H24. aw. aw. nin H22. eee. nin H22. red. ir. elim H23. wr H24. aw. am. Qed. Lemma Exercise1_19b: forall r f g, ordinal_sum_axioms1 r f g -> nonempty(substrate r) -> (forall i u, ~ (maximal_element (V i g) u)) -> (forall i, inc i (substrate r) -> without_gaps (V i g)) -> without_gaps (ordinal_sum r f g). Proof. ir. rww Exercise1_19a. ee. right. nin H0. cp (H2 _ H0). red in H3. ee. nin H4. nin H4. exists y. exists x. exists x0. eee. ir. app H2. ir. right. left. ir. app H1. Qed. Lemma Exercise1_19c: forall r f g, ordinal_sum_axioms1 r f g -> nonempty(substrate r) -> (forall i u, ~ (minimal_element (V i g) u)) -> (forall i, inc i (substrate r) -> without_gaps (V i g)) -> without_gaps (ordinal_sum r f g). Proof. ir. rww Exercise1_19a. ee. right. nin H0. cp (H2 _ H0). red in H3. ee. nin H4. nin H4. exists y. exists x. exists x0. eee. ir. app H2. ir. right. right. ir. app H1. Qed. Lemma Exercise1_19d: forall r f g, ordinal_sum_axioms1 r f g -> without_gaps r -> (forall i, inc i (substrate r) -> (without_gaps (V i g) \/ (forall x y, inc x (V i f) -> inc y (V i f) -> ~ (glt (V i g) x y)))) -> without_gaps (ordinal_sum r f g). Proof. ir. rww Exercise1_19a. ee. left. nin H0. ee. am. ir. nin (H1 _ H2). am. nin H. red in H. ee. assert (inc x (V i f)). wr H11. order_tac. ue. assert (inc y (V i f)). wr H11. order_tac. ue. cp (H4 _ _ H12 H13). contradiction. ir. red in H0. ee. au. Qed. (* exercise 20 *) Definition scattered r := order r & (forall x, sub x (substrate r) -> ~ (without_gaps (induced_order r x))). Lemma Exercise1_20a : forall r x, sub x (substrate r) -> scattered r -> scattered (induced_order r x). Proof. ir. red in H0; red. ee. fprops. ir. awii H2. wrr induced_trans. app H1. apply sub_trans with x. am. am. Qed. Lemma Exercise1_20b : forall r, worder r -> scattered r. Proof. ir. red. ee. nin H ;am. ir. red. ir. assert (worder (induced_order r x)). app worder_restriction. set (y:= induced_order r x) in *. red in H1. ee. nin H3. nin H3. assert (inc x0 (substrate y)). order_tac. assert (nonempty (substrate y)). exists x0. am. nin (worder_least H2 H6). set (z:= complement (substrate y) (singleton x2)). assert (sub z (substrate y)). uf z. app sub_complement. assert (inc x1 z). uf z. srw. ee. order_tac. aw. red. ir. rwi H9 H3. red in H7. ee. cp (H10 _ H5). order_tac. assert (nonempty z). exists x1. am. nin H2. nin (H11 _ H8 H10). assert (glt y x2 x3). nin H12. awi H12. nin H7. ufi z H12. srwi H12. nin H12. awi H15. split. app H14. au. am. am. nin (H4 _ _ H13). nin H14. assert (inc x4 z). uf z. srw. ee. order_tac. aw. nin H14. au. nin H12. awii H17. cp (H17 _ H16). awii H12. awii H18. order_tac. Qed. Lemma exists_proof2 : forall p : EEP, ~ (forall x y: Set, ~ p x y) -> (exists x, exists y, p x y). Proof. ir. nin (p_or_not_p (exists x, exists y, p x y)). am. elim H. ir. red. ir. elim H0. exists x. exists y. am. Qed. Lemma Exercise1_20c: forall r x y, scattered r -> glt r x y -> exists x', exists y', glt r x' y' & (forall z, ~ (glt r x' z & glt r z y')). Proof. ir. nin H. set (F:= interval_cc r x y). assert (sub F (substrate r)). uf F. uf interval_cc. app Z_sub. cp (H1 _ H2). app exists_proof2. dneg. clear H3. split. fprops. split. exists x; exists y. nin H0. split. aw. uf F. uf interval_cc. Ztac. order_tac. ee. order_tac. order_tac. am. uf F. uf interval_cc. Ztac. order_tac. ee. am. order_tac. order_tac. am. ir. app exists_proof. assert (glt r x0 y0). ap (related_induced_order2 H3). cp (H4 x0 y0). dneg. ee. am. ir. cp (H7 z). dneg. nin(related_induced_order4 H3). nin H9. nin H9; nin H12. assert (inc z F). uf F. uf interval_cc. Ztac. order_tac. ee. ufi F H10. ufi interval_cc H10. Ztac. ee. order_tac. ufi F H11. ufi interval_cc H11. Ztac. ee. order_tac. uf glt. ee. aw. am. aw. am. Qed. Definition cantor_tri_order:= lexicographic_order Bnat_order (L Bnat (fun _: Set => two_points)) (L Bnat (fun _: Set => example_worder)). Lemma cantor_tri_order_axioms: lexicographic_order_axioms Bnat_order (L Bnat (fun _: Set => two_points)) (L Bnat (fun _: Set => example_worder)). Proof. red. ee; bw;gprops. app worder_Bnat_order. app substrate_Bnat_order. ir. bw. nin example_is_worder. am. ir. bw. rww substrate_example_worder. Qed. Lemma cantor_tri_order_total : total_order cantor_tri_order. Proof. uf cantor_tri_order. app total_lexicographic_order. app cantor_tri_order_axioms. ir. bwi H. bw. app worder_total. app example_is_worder. rww inc_Bnat. Qed. Lemma related_cantor_tri_order: forall x x', related cantor_tri_order x x' = (inc x (productb (L Bnat (fun _: Set => two_points))) & inc x' (productb (L Bnat (fun _: Set => two_points))) & forall j, least_element (induced_order Bnat_order (Zo Bnat (fun i => V i x <> V i x'))) j -> (V j x = TPa & V j x' = TPb)). Proof. ir. uf cantor_tri_order. rw related_lexicographic_order. assert (Ha: forall j p, least_element (induced_order Bnat_order (Zo Bnat p)) j -> inc j Bnat). ir. red in H. ee. awi H. Ztac. am. nin worder_Bnat_order. am. rw substrate_Bnat_order. app Z_sub. ap iff_eq; ir; eee. ir. bwi H1. cp (H1 _ H2). bwi H3. nin H3. rwi example_worder_related H3. nin H3. nin H3. elim H4. ue. nin H3. nin H3. elim H4. ue. am. app (Ha _ _ H2). ir. bwi H2. cp (H1 _ H2). split. bw. rw example_worder_related. au. app (Ha _ _ H2). nin H3. rw H3; rw H4. fprops. app cantor_tri_order_axioms. Qed. Lemma glt_cantor_tri_order: forall x x', glt cantor_tri_order x x' = (inc x (productb (L Bnat (fun _: Set => two_points))) & inc x' (productb (L Bnat (fun _: Set => two_points))) & exists j, inc j Bnat & (forall i, inc i Bnat -> cardinal_lt i j -> V i x = V i x') & (V j x = TPa & V j x' = TPb)). Proof. ir. set (Z:=Zo Bnat (fun i : Set => V i x <> V i x')). assert (Hb: order Bnat_order). nin (worder_Bnat_order). am. assert (Hc:sub Z (substrate Bnat_order)). rw substrate_Bnat_order. uf Z. app Z_sub. assert (Ha: forall j p, least_element (induced_order Bnat_order (Zo Bnat p)) j -> inc j Bnat). ir. red in H. ee. awi H. Ztac. am. am. rw substrate_Bnat_order. app Z_sub. uf glt. rw related_cantor_tri_order. ap iff_eq; ir; eee. nin (emptyset_dichot Z). elim H0. assert (fgraph (L Bnat (fun _ : Set => two_points))). gprops. rw (productb_extensionality H4 H H1). bw. ir. nin (p_or_not_p (V i x = V i x')). am. elim (emptyset_pr (x:=i)). wr H3. uf Z. Ztac. fold Z in H2. nin (worder_Bnat_order). nin (H5 _ Hc H3). cp (H2 _ H6). exists x0. ee. app (Ha _ _ H6). ir. red in H6. awii H6. ee. nin (p_or_not_p (V i x = V i x')). am. assert (inc i Z). uf Z. Ztac. cp (H11 _ H13). cp (related_induced_order1 H14). assert (cardinal_le x0 i). rwi related_Bnat_order H15. red in H15. ee; am. co_tac. am. am. ir. nin H1. nin H1. nin H3. red in H2. fold Z in H2. awii H2. nin H2. assert (inc x0 Z). uf Z. Ztac. nin H4. rw H4; rw H6. fprops. cp (H5 _ H6). cp (related_induced_order1 H7). assert (cardinal_le j x0). rwi related_Bnat_order H8. red in H8. ee; am. assert (is_cardinal j). red in H9; eee. assert (is_cardinal x0). red in H9; eee. nin (cardinal_le_total_order2 H11 H10). assert (x0 = j). co_tac. wrr H13. assert (inc j Bnat). ufi Z H2. Ztac. am. cp (H3 _ H13 H12). ufi Z H2. Ztac. contradiction. nin H1. ee. red. ir. assert (TPa = TPb). wr H3; wr H4; wrr H5. fprops. Qed. Lemma Exercise1_20d: let r:= cantor_tri_order in forall x y, glt r x y -> exists x', exists y', glt r x' y' & (forall z, ~ (glt r x' z & glt r z y')). Proof. ir. cp H. ufi r H0. rwi glt_cantor_tri_order H0. ee. nin H2. ee. set (f:= fun i=> Yo (cardinal_le i x0) (V i x) TPb). set (g:= fun i=> Yo (cardinal_le i x0) (V i y) TPa). assert (fgraph (L Bnat (fun _ : Set => two_points))). gprops. assert (inc (L Bnat f) (productb (L Bnat (fun _ : Set => two_points)))). rww productb_pr. ee. gprops. bw. bw. ir. bw. uf f. nin (p_or_not_p (cardinal_le i x0)). rww Y_if_rw. rwi productb_pr H0. ee. bwi H9. rwi H9 H10. cp (H10 _ H7). bwi H11. am. am. gprops. rww Y_if_not_rw. fprops. assert (inc (L Bnat g) (productb (L Bnat (fun _ : Set => two_points)))). rww productb_pr. ee. gprops. bw. bw. ir. bw. uf g. nin (p_or_not_p (cardinal_le i x0)). rww Y_if_rw. rwi productb_pr H1. ee. bwi H10. rwi H10 H11. cp (H11 _ H8). bwi H12. am. am. gprops. rww Y_if_not_rw. fprops. exists (L Bnat f). exists (L Bnat g). ee. rw glt_cantor_tri_order. eee. exists x0. ee. am. ir. nin H10. uf f; uf g. bw. rww Y_if_rw. rww Y_if_rw. app H3. split; am. uf f. bw. rww Y_if_rw. fprops. uf g. bw. rww Y_if_rw. fprops. red. ir. ee. rwi glt_cantor_tri_order H9. rwi glt_cantor_tri_order H10. ee. nin H14. nin H12. ee. assert (is_cardinal x0). fprops. assert (is_cardinal x1). fprops. assert (is_cardinal x2). fprops. nin (cardinal_le_total_order2 H23 H21). nin (cardinal_le_total_order2 H22 H21). nin (cardinal_le_total_order1 H22 H23). assert (TPa = TPb). wr H16. wr H20. rww H26. fprops. nin H26. cp (H15 _ H14 H26). rwi H20 H27. assert (cardinal_lt x1 x0). co_tac. ufi g H27. bwi H27. cp (H3 _ H14 H28). nin H28. rwii Y_if_rw H27. ufi f H19. bwi H19. rwii Y_if_rw H19. wri H27 H29. rwi H19 H29. fprops. am. am. assert (cardinal_lt x2 x0). co_tac. ufi g H17. bwi H17. cp (H3 _ H12 H27). nin H27. rwii Y_if_rw H17. rwi H17 H28. cp (H18 _ H12 H26). rwi H16 H30. ufi f H30. bwi H30. rwii Y_if_rw H30. rwi H30 H28. fprops. am. am. ufi f H19. bwi H19. rwi Y_if_not_rw H19. fprops. red. ir. co_tac. am. ufi g H17. bwi H17. rwi Y_if_not_rw H17. fprops. red. ir. co_tac. am. Qed. Lemma Exercise1_20e: ~ (scattered cantor_tri_order). Proof. red. ir. red in H. ee. set (w:= productb (L Bnat (fun _ : Set => two_points))). assert (w = substrate cantor_tri_order). uf cantor_tri_order. rww substrate_lexicographic_order. app cantor_tri_order_axioms. set (all_a := Zo w (fun z => exists i, inc i Bnat & forall j, inc j Bnat -> cardinal_le i j -> V j z = TPa)). set (s:= complement w all_a). assert (sub s w). uf s. app sub_complement. set (zb:= L Bnat (fun _ : Set => TPb)). assert (inc zb s). uf s. srw. ee. uf w. rw productb_pr. uf zb. ee. gprops. bw. bw. ir. bw. fprops. gprops. red. ir. ufi all_a H3. Ztac. nin H5. ee. assert (cardinal_le x x). fprops. cp (H6 _ H5 H7). ufi zb H8. bwi H8. fprops. am. set (zab:= L Bnat (fun i : Set => Yo (i=card_zero) TPa TPb)). assert (inc zab s). uf s. srw. ee. uf w. rw productb_pr. uf zab. ee. gprops. bw. bw. ir. bw. nin (equal_or_not i card_zero). rww Y_if_rw. fprops. rww Y_if_not_rw. fprops. gprops. red. ir. ufi all_a H4. Ztac. nin H6. ee. assert (cardinal_le x (succ x)). app is_less_than_succ. fprops. assert (inc (succ x) Bnat). fprops. cp (H7 _ H9 H8). ufi zab H10. bwi H10. rwi Y_if_not_rw H10. fprops. app succ_nonzero. fprops. assert (glt cantor_tri_order zab zb). rw glt_cantor_tri_order. fold w. ee. ufi s H4. srwi H4. ee; am. ufi s H3. srwi H3. ee; am. exists card_zero. ee. fprops. ir. elim (zero_smallest1 H6). uf zab. bw. rww Y_if_rw. fprops. uf zb. bw. fprops. wri H1 H0. cp (H0 _ H2). elim H6. split. rwi H1 H2. fprops. ee. exists zab. exists zb. nin H5. split. aw. am. ir. cp (related_induced_order2 H7). nin H7. red in H7. red in H7. ufi induced_order H7. cp (intersection2_second H7). clear H7. ufi coarse H10. awi H10. ee. clear H7. rwi glt_cantor_tri_order H8. ee. nin H12. ee. set (u:= Zo Bnat (fun i => cardinal_lt x0 i & V i y = TPb)). nin (emptyset_dichot u). ufi s H11. srwi H11. nin H11. elim H17. uf all_a. Ztac. exists (succ x0). ee. fprops. ir. rwi lt_n_succ_le0 H19. assert (inc (V j y) two_points). ufi w H11. rwi productb_pr H11. ee. bwi H20. rwi H20 H21. cp (H21 _ H18). bwi H22. am. am. gprops. rwi two_points_pr H20. nin H20. am. elim (emptyset_pr (x:=j)). wr H16. uf u. Ztac. fprops. fprops. nin worder_Bnat_order. assert (sub u (substrate Bnat_order)). rw substrate_Bnat_order. uf u. app Z_sub. nin (H18 _ H19 H16). red in H20. awii H20. clear H18. nin H20. set (z:= L Bnat (fun i=> Yo (i= x1) TPa (V i y))). assert (inc z s). uf z. uf s. srw. ee. uf w. rw productb_pr. ee. gprops. bw. bw. ir. bw. nin (equal_or_not i x1). rww Y_if_rw. fprops. rw Y_if_not_rw. ufi s H11. srwi H11. nin H11. ufi w H11. rwi productb_pr H11. ee. bwi H24. rwi H24 H25. cp (H25 _ H21). bwi H26. am. am. gprops. am. gprops. red. ir. ufi all_a H21. Ztac. nin H23. ee. ufi s H11. srwi H11. nin H11. elim H25. uf all_a. clear H21. Ztac. assert (inc x1 Bnat). ufi u H18. srwi H18. Ztac. am. set (a:= succ (card_plus x1 x2)). assert (inc a Bnat). uf a. fprops. assert (cardinal_le x2 a). uf a. uf succ. assert (cardinal_le x2 (card_plus x1 x2)). rw card_plus_commutative. app sum_increasing3. fprops. fprops. assert (cardinal_le (card_plus x1 x2) (card_plus (card_plus x1 x2) card_one)). app sum_increasing3. fprops. fprops. co_tac. exists a. ee. am. ir. assert (cardinal_le x2 j). co_tac. cp (H24 _ H28 H30). bwi H31. rwi Y_if_not_rw H31. am. red. ir. rwi H32 H29. assert (cardinal_lt x1 a). uf a. rw lt_is_le_succ. app sum_increasing3. fprops. fprops. fprops. co_tac. am. assert (glt cantor_tri_order x z). rw glt_cantor_tri_order. fold w. ee. ufi s H10. srwi H10. ee; am. ufi s H21. srwi H21. ee; am. exists x0. ee. am. ir. rww H13. uf z. bw. rw Y_if_not_rw. tv. ufi u H18. Ztac. ee. assert (cardinal_lt i x1). co_tac. nin H27. am. am. uf z. bw. rw Y_if_not_rw. tv. ufi u H18. Ztac. ee. nin H23. am. assert (glt cantor_tri_order z y). rw glt_cantor_tri_order. fold w. ee. ufi s H21. srwi H21. ee. am. ufi s H11. srwi H11. ee; am. exists x1. ee. ufi u H18. Ztac. am. ir. uf z. bw. rw Y_if_not_rw. tv. nin H24. am. uf z. bw. rww Y_if_rw. ufi u H18. Ztac. am. ufi u H18. Ztac. ee. am. exists z. nin H22; nin H23. uf glt. ee. aw. am. aw. am. Qed. Lemma Exercise1_20f: forall r f g, ordinal_sum_axioms1 r f g -> scattered (ordinal_sum r f g) = (scattered r & forall i, inc i (domain f) -> scattered (V i g)). Proof. ir. cp H. nin H. cp H. red in H. ee. ap iff_eq. ir. ee. red. ee. am. ir. red. ir. set (w:= fun_image x (fun i => (J (rep (V i f)) i))). assert (sub w (substrate (ordinal_sum r f g))). red. ir. aw. ufi w H12. awi H12. nin H12. ee. wr H13. app inc_disjoint_union. ue. app nonempty_rep. app H1. ue. nin H9. cp (H13 _ H12). elim H14. red in H11. ee. red. ee. fprops. nin H15. nin H15. cp (related_induced_order4 H15). ee. cp (related_induced_order2 H15). assert (inc x0 (domain f)). wr H3. app H10. assert (inc x1 (domain f)). wr H3. app H10. exists (J (rep (V x0 f)) x0); exists (J (rep (V x1 f)) x1). split. aw. ee. app inc_disjoint_union. app nonempty_rep. app H1. app inc_disjoint_union. app nonempty_rep. app H1. au. uf w. aw. exists x0. eee. uf w. aw. exists x1. eee. nin H19. dneg. inter2tac. ir. cp (related_induced_order2 H17). nin H18. nin H17. ufi induced_order H17. cp (intersection2_second H17). ufi coarse H21. awi H21. ee. ufi w H22. awi H22. nin H22. ufi w H23. awi H23. nin H23. ee. awi H18. ee. nin H27. wri H25 H27. wri H24 H27. awi H27. assert (glt (induced_order r x) x1 x2). nin H27. split. aw. am. nin (H16 _ _ H28). assert (inc (J (rep (V x3 f)) x3) w). uf w. aw. nin H29. cp (related_induced_order4 H29). cp (related_induced_order4 H30). ee. exists x3. ee. am. tv. awi H12. nin H29. cp (related_induced_order2 H29). cp (related_induced_order2 H31). exists (J (rep (V x3 f)) x3). ee. uf glt. ee. aw. ee. am. app H12. wr H25. aw. au. wr H25. uf w. aw. exists x1. eee. wr H25. nin H32. dneg. inter2tac. split. aw. ee. app H12. am. wr H24. aw. au. wr H24. uf w. aw. exists x2. eee. wr H24. nin H33. dneg. inter2tac. am. elim H19. wr H25. wr H24. nin H27. wri H25 H27. wri H24 H27. awi H27. ue. am. ir. split. app H7. ir. set (y:= fun_image x (fun u => (J u i))). assert (sub y (substrate (ordinal_sum r f g))). aw. red. ir. ufi y H12. awi H12. nin H12. nin H12. wr H13. app inc_disjoint_union. wr H8. app H11. am. nin H9. cp (H13 _ H12). dneg. nin H15. nin H16. nin H16; nin H16. nin (related_induced_order4 H16). nin (related_induced_order2 H16). red. ee. fprops. exists (J x0 i). exists (J x1 i). split. aw. ee. app inc_disjoint_union. wrr H8. order_tac. app inc_disjoint_union. wrr H8. order_tac. right. au. uf y. aw. exists x0. au. uf y. aw. exists x1. au. dneg. inter2tac. ir. nin (related_induced_order4 H22). nin (related_induced_order2 H22). ufi y H23. awi H23. nin H23. ufi y H24. awi H24. nin H24. ee. wri H28 H25; wri H27 H25; awi H25. ee. nin H30. nin H30. elim H31. tv. nin H30. assert (glt (induced_order (V i g) x) x3 x4). split. aw. dneg. wr H28; wr H27. ue. nin (H17 _ _ H32). nin H33. nin (related_induced_order4 H33). nin (related_induced_order2 H33). nin (related_induced_order4 H34). nin (related_induced_order2 H34). assert (inc (J x5 i) y). uf y. aw. exists x5. au. assert (inc (J x5 i) (disjoint_union f)). app inc_disjoint_union. wr H8. order_tac. am. wr H28. wr H27. exists (J x5 i). uf glt. ee; aw. eee. uf y. aw. exists x3. au. clear H42. dneg. inter2tac. eee. uf y. aw. exists x4. au. dneg. inter2tac. am. ir. red. ee. fprops. ir. red. ir. set (ns := Zo (substrate r) (fun z => exists i, inc i x & z = Q i)). set (r' := induced_order r ns). assert (sub ns (substrate r)). uf ns. app Z_sub. assert (order r'). uf r'. fprops. assert (substrate r' = ns). uf r'. app substrate_induced_order. set (f':= L ns (fun i=> Zo (V i f) (fun u => inc (J u i) x))). assert (forall i, inc i ns -> sub (V i f') (V i f)). ir. uf f'. bw. app Z_sub. set (g' := L ns (fun i=> induced_order (V i g) (V i f'))). assert (nonempty x). nin H12. nin H17. nin H17. nin H17. nin (related_induced_order4 H17). exists x0. am. assert (ordinal_sum_axioms1 r' f' g'). uf f'; uf g'. split. red. ee. am. rw H15. bw. gprops. gprops. bw. bw. ir. bw. cp (H13 _ H18). rwi H3 H19. cp (H7 _ H19). app order_induced_order. rw H8. app H16. am. bw. ir. bw. cp (H13 _ H18). rwi H3 H19. aw. uf f'. bw. app H7. rww H8. app H16. bw. ir. bw. ufi ns H18. Ztac. nin H20. nin H20. clear H18. cp (H11 _ H20). awi H18. nin (du_index_pr H18). nin H23. exists (P x0). Ztac. rww H21. rww H21. aw. am. assert (x = substrate (ordinal_sum r' f' g')). aw. set_extens. awi H11. cp (H11 _ H19). nin (du_index_pr H20). nin H22. assert (x0 = J (P x0) (Q x0)). aw. assert (inc (Q x0) ns). uf ns. Ztac. rww H3. exists x0. au. rw H24. app inc_disjoint_union. uf f'. bw. uf f'. bw. Ztac. wrr H24. am. cp (du_index_pr H19). ee. ufi f' H21. bwi H21. Ztac. awi H24. am. am. ufi f' H20. bwi H20. am. nin H18. am. assert (induced_order (ordinal_sum r f g) x = (ordinal_sum r' f' g')). uf induced_order. set_extens. nin (intersection2_both H20). ufi coarse H22. awi H22. ee. set (a:= P x0) in *. set (b := Q x0) in *. assert (x0 = J a b). uf a; uf b; aw. rw H25. change (gle (ordinal_sum r' f' g') a b). aw. uf f'. rwi H25 H21. change (gle (ordinal_sum r f g) a b) in H21. awi H21. nin H21; nin H26. nin (du_index_pr H21). nin H29. nin (du_index_pr H26). nin H32. assert (a = J (P a) (Q a)). aw. assert (b = J (P b) (Q b)). aw. assert (inc (Q a) ns). uf ns. Ztac. rww H3. exists a. au. assert (inc (Q b) ns). uf ns. Ztac. rww H3. exists b. au. ee. rw H34. app inc_disjoint_union. bw. bw. Ztac. wrr H34. rw H35. app inc_disjoint_union. bw. bw. Ztac. wrr H35. uf r'; uf g'. nin H27. left. nin H27. split. aw. am. right. ee. am. red. red. bw. uf f'. bw. uf induced_order. app intersection2_inc. uf coarse. aw. ee. fprops. Ztac. wrr H34. Ztac. rw H27. am. rw H27. wrr H35. am. nin H18; am. assert (is_graph (ordinal_sum r' f' g')). app order_is_graph. nin H18; fprops. assert (is_pair x0). app H21. assert (x0 = J (P x0) (Q x0)). aw. rw H23. rwi H23 H20. change (gle (ordinal_sum r' f' g') (P x0) (Q x0)) in H20. assert (inc (P x0) (substrate (ordinal_sum r' f' g'))). order_tac. assert (inc (Q x0) (substrate (ordinal_sum r' f' g'))). order_tac. wri H19 H24; wri H19 H25. uf coarse. app intersection2_inc. set (a:= P x0) in *. set (b := Q x0) in *. change (gle (ordinal_sum r f g) a b). awi H20. ufi f' H20; ufi g' H20. ufi r' H20. ee. cp (du_index_pr H20). bwi H28. ee. cp (du_index_pr H26). bwi H31. ee. Ztac. clear H32. Ztac. clear H29. assert (a = J (P a) (Q a)). aw. assert (b = J (P b) (Q b)). aw. aw. ee. rw H29. app inc_disjoint_union. wr H3. app H13. rw H37. app inc_disjoint_union. wr H3. app H13. nin H27. left. ap (related_induced_order2 H27). right. ee. am. bwi H38. cp (related_induced_order1 H38). am. am. ee. am. ee; am. nin H18; am. aw. ee; am. rwi H20 H12. rwi (Exercise1_19a H18) H12. ee. assert (forall i x y, inc i (substrate r') -> ~ glt (V i g') x y). ir. nin (p_or_not_p (glt (V i g') x0 y)). cp (H21 _ _ _ H23 H24). assert (inc i (domain f)). wr H3. app H13. wrr H15. cp (H10 _ H26). ufi g' H25. bwi H25. nin H27. assert (sub (V i f') (substrate (V i g))). rww H8. app H16. wrr H15. elim (H28 _ H29). am. wrr H15. am. red in H9. ee. cp (H24 _ H13). elim H25. red. ee. fprops. nin H12. nin H12. nin H12. exists x0; exists x1. am. nin H12. nin H12. nin H12. ee. elim (H23 x0 x1 x2 H12). am. ir. fold r' in H26. fold r'. nin (H22 _ _ H26). am. nin H27. nin H18. assert (inc x0 (domain f')). uf f'. bw. wr H15. order_tac. nin (H28 _ H29). elim (H27 y0). split. red in H18. ee. rww H36. ir. nin (equal_or_not x1 y0). am. assert (glt (V x0 g') y0 x1). split. am. intuition. elim (H23 x0 y0 x1). rww H15. ufi f' H29. bwi H29. am. am. assert (inc y ns). wr H15. order_tac. assert (inc y (domain f')). uf f'. bw. nin H18. nin (H30 _ H29). elim (H27 y0). split. red in H18. ee. rww H37. ir. nin (equal_or_not x1 y0). am. assert (glt (V y g') x1 y0). split; am. elim (H23 y x1 y0). rww H15. am. Qed. (* exercice 21 *) Lemma Exercise1_2g: forall r s, weak_order_compatibility r s-> quotient_order_axiom r s -> total_order r -> let r' := (quotient_order r s) in forall x y, gle r' x y = (inc x (quotient s) & inc y (quotient s) & gle r (rep x) (rep y)). Proof. ir. rename H1 into Ht. red in H. red in H0. ee. assert (forall u, inc u (quotient s) -> inc (rep u) u). ir. app (inc_rep_itself H1 H4). uf r'. rw quotient_order_pr. uf quotient_order_r. ap iff_eq; ir; eee; ir;cp (H4 _ H5); cp (H4 _ H6). nin Ht. assert (inc (rep x) (substrate r)). wr H2. app (inc_in_quotient_substrate H1 H8 H5). assert (inc (rep y) (substrate r)). wr H2. app (inc_in_quotient_substrate H1 H9 H6). nin (H11 _ _ H13 H12). nin (H7 _ H8). ee. assert (gle r (rep y) x0). order_tac. assert (related s (rep y) x0). rwi inc_quotient H6. rwi is_class_rw H6. ee. wrr (H19 (rep y) x0 H9). am. am. cp (H0 _ _ _ H14 H16 H18). rwii related_rep_rep H19. rw H19. order_tac. am. assert (related s (rep x) x0). rw in_class_related. exists x. eee. wrr inc_quotient. am. nin (H3 _ _ _ H7 H11). exists x1. eee. rwi inc_quotient H6. rwi is_class_rw H6. ee. rww (H15 (rep y) x1 H10). am. am. Qed. Lemma Exercise1_2h: forall r s, weak_order_compatibility r s-> quotient_order_axiom r s -> total_order r -> total_order (quotient_order r s). Proof. ir. cp (Exercise1_2g H H0 H1). cp H. cp H1. red in H. red in H1. ee. cp (Exercise1_2d H6 H1 H7 H0). split. am. ir. rwi substrate_quotient_order H10. rwi substrate_quotient_order H11. assert (inc (rep x) (substrate r)). wr H7. app inc_rep_substrate. assert (inc (rep y) (substrate r)). wr H7. app inc_rep_substrate. nin (H5 _ _ H12 H13). left. rww Exercise1_2g. au. right. change (gle (quotient_order r s) y x). rww Exercise1_2g. au. am. am. am. am. am. am. Qed. Lemma Exercise1_2i: forall r s, let q := quotient s in let r' := (quotient_order r s) in let f' := L q (fun z: Set => z) in let g' := L q (fun z => induced_order r z) in let du := disjoint_union f' in let f := BL (fun x => J x (class s x)) (substrate r) du in weak_order_compatibility r s-> quotient_order_axiom r s -> total_order r -> (ordinal_sum_axioms1 r' f' g' & substrate (ordinal_sum r' f' g') = du & (forall x y, inc x (substrate r) -> inc y (substrate r) -> related s x y = related (equivalence_associated (second_proj du)) (W x f) (W y f)) & bijective f & order_isomorphism f r (ordinal_sum r' f' g')). Proof. ir. cp (Exercise1_2g H H0 H1). cp (Exercise1_2h H H0 H1). nin H. nin H4. nin H5. nin H1. assert (substrate r' = q). uf r'. rww substrate_quotient_order. assert (forall i, inc i q -> sub i (substrate r)). red. ir. wr H5. app (inc_in_quotient_substrate H4 H10 H9). assert (Ha: ordinal_sum_axioms1 r' f' g'). red. ee. red. ee. nin H3. am. uf f'. bw. uf f'. gprops. uf g'. gprops. uf f'; uf g'; bw. uf f'; uf g'. bw. ir. bw. cp (H9 _ H10). fprops. uf f'; uf g'. bw. ir. bw. cp (H9 _ H10). aw. uf f'. bw. ir. bw. app (non_empty_in_quotient H4 H10). assert (Hb: transf_axioms (fun x => J x (class s x)) (substrate r) du). red. ir. uf du. app inc_disjoint_union. uf f'. bw. uf q. app inc_class_quotient. ue. uf f'. bw. wri H5 H10. equiv_tac. wri H5 H10. uf q. gprops. assert (Hc: bijective f). uf f. app bijective_bl_function. ir. inter2tac. ir. ufi du H10. cp (du_index_pr H10). ufi f' H11 . bwi H11. ee. exists (P y). ee. wr H5. app (inc_in_quotient_substrate H4 H12 H11). app pair_extensionality. ee. fprops. aw. aw. sy. app is_class_pr. nin H11. am. ee. am. aw. nin Ha; am. ir. cp Exercise1_3a6. change (related s x y = related (E13_S f') (W x f) (W y f)). rw H12. uf E13_sF. fold du. ap iff_eq. ir. ee. assert (du = target f). uf f. aw. rw H14. app inc_W_target. fct_tac. uf f. aw. assert (du = target f). uf f. aw. rw H14. app inc_W_target. fct_tac. uf f. aw. uf f. aw. app related_class_eq1. ir. ee. ufi f H15. awi H15. rww related_class_eq. wri H5 H10. equiv_tac. am. am. am. am. am. red. nin Ha. eee. aw. uf f. aw. uf f. aw. uf f. aw. ir. aw. ap iff_eq. ir. ee. app Hb. app Hb. uf g'. bw. nin (p_or_not_p (related s x y)). right. ee. app related_class_eq1. aw. bw. assert (inc x (substrate r)). order_tac. wri H5 H16. equiv_tac. bw. left. split. uf r'. rw quotient_order_pr. red. ee. wri H5 H12. gprops. wri H5 H13. gprops. ir. assert (related s x x0). bwi H16. am. am. nin (H6 x y x0 H14 H17). exists x1. ee. bw. am. dneg. rww related_class_eq. wri H5 H12. equiv_tac. wri H5 H12. uf q. gprops. ir. ee. nin H16. nin H16. ufi r' H16. rwi quotient_order_pr H16. red in H16. ee. nin (H7 _ _ H12 H13). am. assert (inc x (class s x)). bw. wri H5 H12. equiv_tac. nin (H19 _ H21). ee. red in H0. assert (related s y x0). bwi H22. am. am. cp (H0 _ _ _ H20 H23 H24). elim H17. app related_class_eq1. equiv_tac. ee. ufi g' H17. bwi H17. ap (related_induced_order1 H17). uf q. wri H5 H12. gprops. Qed. Definition scattered_rel r x y := (gle r x y & scattered (induced_order r (interval_cc r x y))) \/ (gle r y x & scattered (induced_order r (interval_cc r y x))). Definition scattered_equiv r := graph_on (scattered_rel r) (substrate r). Lemma Exercise1_21a : forall r v, order r -> sub v (substrate r) -> scattered (induced_order r v) = (forall u, sub u v -> ~ without_gaps (induced_order r u)). Proof. ir. uf scattered. aw. ir. ap iff_eq. ir. ee. cp (H3 _ H2). wri (induced_trans H H2 H0) H4. am. ir. ee. fprops. ir. wr (induced_trans H H2 H0). app H1. Qed. Lemma Exercise1_21b : forall r u, total_order r -> sub u (substrate r) -> (exists x, exists y, glt (induced_order r u) x y)= (exists x, exists y, inc x u & inc y u & x<> y). Proof. ir. ap iff_eq; ir; nin H1; nin H1. nin (related_induced_order4 H1). nin (related_induced_order2 H1). exists x; exists x0. eee. nin H. ee. nin (H2 _ _ (H0 _ H1) (H0 _ H3)). exists x; exists x0. split. aw. am. exists x0; exists x. split. aw. au. Qed. Lemma Exercise1_21c : forall r u, total_order r -> sub u (substrate r) -> (forall a b, inc a u -> inc b u -> a = b) = ~ (exists x, exists y, glt (induced_order r u) x y). Proof. ir. rww Exercise1_21b. ap iff_eq. ir. red. ir. nin H2. nin H2. ee. elim H4. app H1. ir. nin (equal_or_not a b). am. elim H1. exists a; exists b; au. Qed. Lemma Exercise1_21d : forall r u, total_order r -> sub u (substrate r) -> (~ without_gaps (induced_order r u)) = ((forall a b, inc a u -> inc b u -> a = b) \/ (exists a, exists b, inc a u & inc b u & glt r a b & (forall z, inc z u -> gle r z a \/ gle r b z))). Proof. ir. ap iff_eq. ir. nin (p_or_not_p (forall a b : Set, inc a u -> inc b u -> a = b)). au. nin (p_or_not_p (exists a, exists b, inc a u & inc b u & glt r a b & (forall z : Set, inc z u -> gle r z a \/ gle r b z))). au. elim H1. red. ee. nin H. fprops. rwi (Exercise1_21c H H0) H2. app excluded_middle. ir. app exists_proof. dneg. nin (related_induced_order4 H4). cp (related_induced_order2 H4). exists x; exists y. eee. ir. nin H. nin (H10 _ _ (H0 _ H9) (H0 _ H6)). au. nin (H10 _ _ (H0 _ H7) (H0 _ H9)). au. nin (equal_or_not x z). left. rw H13. cp (H0 _ H9). order_tac. nin (equal_or_not z y). right. wr H14. cp (H0 _ H9). order_tac. elim (H5 z). uf glt. ee; aw. ir. red. ir. red in H2. ee. nin H1. rwi (Exercise1_21c H H0) H1. contradiction. nin H1. nin H1. ee. assert (glt (induced_order r u) x x0). nin H6. split. aw. am. nin (H4 _ _ H8). ee. nin (related_induced_order4 H9). cp (related_induced_order2 H9). nin (related_induced_order4 H10). cp (related_induced_order2 H10). nin H. nin (H7 _ H12). order_tac. order_tac. Qed. Lemma Exercise1_21e: forall r u a b, total_order r -> let v:= intersection2 u (interval_cc r a b) in sub u (substrate r) -> without_gaps (induced_order r u) -> (exists x, exists y, inc x v & inc y v & glt r x y) -> without_gaps (induced_order r v). Proof. ir. cp H. nin H. assert (sub v u). uf v. app intersection2sub_first. assert (sub v (interval_cc r a b)). uf v. app intersection2sub_second. assert (sub v (substrate r)). apply sub_trans with u. am. am. nin (p_or_not_p (without_gaps (induced_order r v))). am. assert (~ ~ without_gaps (induced_order r u)). red. ir. contradiction. rwi (Exercise1_21d H3 H7) H8. rwi (Exercise1_21d H3 H0) H9. elim H9. clear H9. nin H8. nin H2. nin H2. ee. nin H10. elim H11. app H8. right. nin H8. nin H8. ee. exists x; exists x0. ee. app H5. app H5. am. ir. nin (H4 _ _ (H0 _ H12) (H0 _ (H5 _ H8))). au. nin (equal_or_not x z). rw H14. left. order_tac. ap (H0 _ H12). nin (H4 _ _ (H0 _ (H5 _ H9)) (H0 _ H12)). au. nin (equal_or_not z x0). rw H16. right. order_tac. order_tac. assert (inc z v). uf v. app intersection2_inc. uf interval_cc. Ztac. ee. ufi v H8. cp (intersection2_second H8). ufi interval_cc H17. Ztac. change (gle r x z) in H13. ee; order_tac. ufi v H9. cp (intersection2_second H9). ufi interval_cc H17. Ztac. change (gle r z x0) in H15. ee; order_tac. app H11. Qed. Lemma Exercise1_21f: forall r a b u, total_order r -> sub u (substrate r) -> inc a u -> inc b u -> glt r a b -> without_gaps (induced_order r u) -> without_gaps (induced_order r (intersection2 u (interval_cc r a b))). Proof. ir. app Exercise1_21e. exists a. exists b. nin H. ee. app intersection2_inc. uf interval_cc. Ztac. nin H3. ee. order_tac. order_tac. am. app intersection2_inc. uf interval_cc. Ztac. nin H3. ee. am. order_tac. order_tac. am. Qed. Lemma Exercise1_21g : forall r x y z, total_order r -> gle r x y -> gle r y z-> scattered (induced_order r (interval_cc r x y)) -> scattered (induced_order r (interval_cc r y z)) -> scattered (induced_order r (interval_cc r x z)). Proof. ir. assert (gle r x z). nin H. order_tac. assert (Ha: order r). nin H; am. assert (Hb:sub (interval_cc r x z) (substrate r)). uf interval_cc. app Z_sub. assert (Hc:sub (interval_cc r x y) (substrate r)). uf interval_cc. app Z_sub. assert (Hd:sub (interval_cc r y z) (substrate r)). uf interval_cc. app Z_sub. rwii Exercise1_21a H2. rwii Exercise1_21a H3. rww Exercise1_21a. ir. assert (sub u (substrate r)). apply sub_trans with (interval_cc r x z). am. am. nin (p_or_not_p (forall a b : Set, inc a u -> inc b u -> a = b)). rww Exercise1_21d. au. rwi (Exercise1_21c H H6) H7. nin (excluded_middle H7). nin H8. set (u1 := intersection2 u (interval_cc r x y)). assert (sub u1 (interval_cc r x y)). uf u1. app intersection2sub_second. set (u2 := intersection2 u (interval_cc r y z)). assert (sub u2 (interval_cc r y z)). uf u2. app intersection2sub_second. cp (H2 _ H9); cp (H3 _ H10). nin (p_or_not_p (exists x2, exists y0 , inc x2 u1 & inc y0 u1 & glt r x2 y0)). red. ir. elim H11. uf u1. app Exercise1_21e. nin (p_or_not_p (exists x2, exists y0 , inc x2 u2 & inc y0 u2 & glt r x2 y0)). red. ir. elim H12. uf u2. app Exercise1_21e. assert (He:forall t, inc t u -> gle r t y -> inc t u1). ir. uf u1. app intersection2_inc. uf interval_cc. Ztac. ee. cp (H5 _ H15). ufi interval_cc H17. Ztac. ee. am. am. assert (Hf:forall t, inc t u -> gle r y t -> inc t u2). ir. uf u2. app intersection2_inc. uf interval_cc. Ztac. ee. am. cp (H5 _ H15). ufi interval_cc H17. Ztac. ee. am. red. ir. red in H15. ee. nin H16; nin H16. nin (H17 _ _ H16). ee. nin (related_induced_order4 H18). nin (related_induced_order2 H18). nin (related_induced_order4 H19). nin (related_induced_order2 H19). nin H. assert (inc y (substrate r)). order_tac. nin (H28 _ _ (H6 _ H21) H29). assert (inc x2 u1). app He. order_tac. assert (inc x4 u1). app He. elim H13. exists x2; exists x4. eee. split;am. change (gle r y x4) in H30. assert (inc x4 u2). app Hf. assert (inc x3 u2). app Hf. order_tac. elim H14. exists x4; exists x3. eee. split;am. Qed. Lemma Exercise1_21h : forall r, total_order r -> equivalence_re (scattered_rel r) (substrate r). Proof. ir. cp H. nin H. split. red. ee. red. uf scattered_rel. ir. nin H2; au. assert (Ha: forall a b x y, gle r x a -> gle r b y -> sub (interval_cc r a b) (interval_cc r x y)). red. uf interval_cc. ir. Ztac. clear H4. aw. Ztac. ee. order_tac. order_tac. assert (Hb: forall a b x y, gle r x a -> gle r b y -> sub (interval_cc r a b) (substrate (induced_order r (interval_cc r x y)))). ir. aw. app Ha. uf interval_cc. app Z_sub. red. uf scattered_rel. ir. nin H2; nin H3. left. ee. order_tac. app (Exercise1_21g H0 H2 H3 H5 H4). ee. assert (inc x (substrate r)). order_tac. assert (inc z (substrate r)). order_tac. nin (H1 _ _ H6 H7). left. split. am. assert (gle r x x). order_tac. cp (Hb _ _ _ _ H9 H3). cp (Exercise1_20a H10 H5). wrii induced_trans H11. app Ha. uf interval_cc. app Z_sub. right. split. am. assert (gle r z z). order_tac. cp (Hb _ _ _ _ H9 H2). cp (Exercise1_20a H10 H4). wrii induced_trans H11. app Ha. uf interval_cc. app Z_sub. ee. assert (inc x (substrate r)). order_tac. assert (inc z (substrate r)). order_tac. nin (H1 _ _ H6 H7). left. split. am. assert (gle r z z). order_tac. cp (Hb _ _ _ _ H2 H9). cp (Exercise1_20a H10 H4). wrii induced_trans H11. app Ha. uf interval_cc. app Z_sub. right. split. am. assert (gle r x x). order_tac. cp (Hb _ _ _ _ H3 H9). cp (Exercise1_20a H10 H5). wrii induced_trans H11. app Ha. uf interval_cc. app Z_sub. right. ee. order_tac. app (Exercise1_21g H0 H3 H2 H4 H5). red. ir. uf scattered_rel. app iff_eq. ir. left. set (int := interval_cc r y y). assert (forall t, inc t int -> t = y). ir. ufi int H3. ufi interval_cc H3. Ztac. ee. order_tac. assert (sub int (substrate r)). red. ir. rww (H3 _ H4). assert (gle r y y). order_tac. ee. am. rww Exercise1_21a. ir. rww Exercise1_21d. left. ir. rw (H3 _ (H6 _ H7)). rww (H3 _ (H6 _ H8)). apply sub_trans with int. am. am. ir. nin H2;ee;order_tac. Qed. Lemma Exercise1_21i : forall r, total_order r -> is_equivalence (scattered_equiv r). Proof. ir. assert (is_graph (scattered_equiv r)). uf scattered_equiv. app is_graph_graph_on. cp (Exercise1_21h H). assert (is_graph_of (scattered_equiv r) (scattered_rel r)). uf scattered_equiv. app equivalence_has_graph0. nin H1. app (equivalence_if_has_graph2 H0 H2 H1). Qed. Lemma Exercise1_21j : forall r, total_order r -> substrate (scattered_equiv r) = substrate r. Proof. ir. app extensionality. uf scattered_equiv. app substrate_graph_on. red. ir. assert (scattered_rel r x x). cp ( Exercise1_21h H). red in H1. ee. wrr H2. cp (Exercise1_21i H). assert (related (scattered_equiv r) x x). uf scattered_equiv. rww related_graph_on1. eee. substr_tac. Qed. Definition scattered_aux r x y := gle r x y & (forall u, sub u (interval_cc r x y) -> ((forall a b, inc a u -> inc b u -> a = b) \/ (exists a, exists b, inc a u & inc b u & glt r a b & (forall z, inc z u -> gle r z a \/ gle r b z)))). Lemma Exercise1_21k : forall r x y, total_order r -> gle r x y -> scattered (induced_order r (interval_cc r x y)) = (forall u, sub u (interval_cc r x y) -> ((forall a b, inc a u -> inc b u -> a = b) \/ (exists a, exists b, inc a u & inc b u & glt r a b & (forall z, inc z u -> gle r z a \/ gle r b z)))). Proof. ir. set (i:=interval_cc r x y). assert (sub i (substrate r)). uf i. uf interval_cc. app Z_sub. rww Exercise1_21a. assert (forall u, sub u i -> sub u (substrate r)). ir. ap (@sub_trans _ _ _ H2 H1). app iff_eq. ir. wr Exercise1_21d. app H3. am. app H2. ir. rww Exercise1_21d. app (H3 _ H4). app H2. nin H. am. Qed. Lemma Exercise1_21l : forall r x y, total_order r -> related (scattered_equiv r) x y= (scattered_aux r x y \/ scattered_aux r y x). Proof. ir. set (f := fun u => (forall a b, inc a u -> inc b u -> a = b) \/ (exists a, exists b, inc a u & inc b u & glt r a b & (forall z, inc z u -> gle r z a \/ gle r b z))). Proof. ir. uf scattered_aux. uf scattered_equiv. rw related_graph_on1. ap iff_eq. ir. ee. nin H2. ee. left. ee. am. rwii Exercise1_21k H3. right. ee. am. rwii Exercise1_21k H3. ir. nin H0. ee. nin H. order_tac. nin H; order_tac. left. ee. am. rww Exercise1_21k. ee. nin H. order_tac. nin H; order_tac. right. ee. am. rww Exercise1_21k. Qed. Lemma Exercise1_21m : forall r, total_order r -> weak_order_compatibility r (scattered_equiv r). Proof. ir. red. ee. nin H. app order_is_preorder. app Exercise1_21i. rww Exercise1_21j. ir. cp H. nin H. assert (inc y (substrate r)). order_tac. cp (Exercise1_21i H2). assert (inc y (substrate (scattered_equiv r))). rww Exercise1_21j. assert (related (scattered_equiv r) y y). equiv_tac. rwi (Exercise1_21l x x' H2) H1. nin H1. red in H1. ee. assert (inc x' (substrate r)). order_tac. nin (H3 _ _ H4 H9). exists x'. ee. rww Exercise1_21l. left. split. am. ir. app H8. red. ir. cp (H11 _ H12). uf interval_cc. ufi interval_cc H13. Ztac. clear H13. ee. Ztac. ee. order_tac. am. order_tac. exists y. ee. am. am. exists y. ee. am. nin H1. ee. order_tac. Qed. Lemma Exercise1_21n: forall r x, total_order r -> inc x (substrate r) -> scattered (induced_order r (class (scattered_equiv r) x)). Proof. ir. cp (Exercise1_21i H). cp (Exercise1_21j H). cp H. nin H. assert (sub (class (scattered_equiv r) x) (substrate r)). wr H2. app sub_class_substrate. red. ee. fprops. ir. red. ir. awii H6. rwii induced_order_trans H7. nin H7. ee. assert (forall u, inc u x0 -> related (scattered_equiv r) u x). ir. cp (H6 _ H10). bwi H11. equiv_tac. am. nin H8. nin H8. nin (related_induced_order4 H8). cp (H10 _ H11); cp (H10 _ H12). assert (related (scattered_equiv r) x x2). equiv_tac. assert (related (scattered_equiv r) x1 x2). equiv_tac. clear H13; clear H14; clear H15. rwi Exercise1_21l H16. nin H16. nin H13. set (u:=intersection2 x0 (interval_cc r x1 x2)). assert (sub u (interval_cc r x1 x2)). uf u. app intersection2sub_second. nin (H14 _ H15). assert (inc x1 u). uf u. app intersection2_inc. uf interval_cc. Ztac. ee. order_tac. order_tac. am. assert (inc x2 u). uf u. app intersection2_inc. uf interval_cc. Ztac. ee. am. order_tac. order_tac. nin H8. elim H19. app H16. nin H16. nin H16. ee. assert (glt (induced_order r x0) x3 x4). nin H18. split. aw. ufi u H16. inter2tac. ufi u H17. inter2tac. am. nin (H9 _ _ H20). ee. nin (related_induced_order4 H21). cp (related_induced_order2 H21). cp (related_induced_order2 H22). assert (inc x5 u). uf u. app intersection2_inc. uf interval_cc. Ztac. ee. ufi u H16. cp (intersection2_second H16). ufi interval_cc H27. Ztac. ee. nin H25. order_tac. ufi u H17. cp (intersection2_second H17). ufi interval_cc H27. Ztac. ee. nin H26. order_tac. nin (H19 _ H27). order_tac. order_tac. nin H13. cp (related_induced_order2 H8). order_tac. am. Qed. Lemma Exercise1_21o: forall r, total_order r -> quotient_order_axiom r (scattered_equiv r). Proof. red. ir. rwii Exercise1_21l H2. cp H. nin H. nin H2. rww Exercise1_21l. left. nin H2. split. am. ir. app H5. ufi interval_cc H6. uf interval_cc. red. ir. cp (H6 _ H7). Ztac. clear H8. Ztac. ee. am. order_tac. nin H2. assert (gle r y x). order_tac. assert (x = y). order_tac. rw H7. cp (Exercise1_21i H3). cp (Exercise1_21j H3). assert (inc y (substrate r)). order_tac. wri H9 H10. equiv_tac. Qed. Lemma Exercise1_21p: forall r, total_order r -> order (quotient_order r (scattered_equiv r)). Proof. ir. app Exercise1_2d. ap (Exercise1_21i H). nin H. am. ap (Exercise1_21j H). ap (Exercise1_21o H). Qed. Lemma Exercise1_2j: forall r s, weak_order_compatibility r s-> quotient_order_axiom r s -> total_order r -> let r' := (quotient_order r s) in forall x y, inc x (substrate r ) -> inc y (substrate r) -> ((gle r x y -> gle r' (class s x) (class s y)) & (glt r' (class s x) (class s y)) ->glt r x y). Proof. ir. assert (Ha:=Exercise1_2h H H0 H1). rename H1 into Ht. red in H. red in H0. assert (forall x y, inc x (substrate r ) -> inc y (substrate r) -> gle r x y -> gle r' (class s x) (class s y)). uf r'. ir. rw quotient_order_pr. uf quotient_order_r. nin H. nin H6. nin H7. wri H7 H1; wri H7 H4. ir. ee. gprops. gprops. ir. assert (inc x1 (substrate r)). wr H7. ap((sub_class_substrate H6 (x:= x0)) _ H9). nin Ht. rwi H7 H4. nin (H12 _ _ H10 H4). exists y0. ee. app inc_itself_class. rww H7. am. change (gle r y0 x1) in H13. assert (related s x0 x1). bwi H9. am. am. cp (H0 _ _ _ H5 H13 H14). exists x1. ee. wrr (related_class_eq1 H6 H15). order_tac. split. app H1. ir. nin H4. split. ir. ee. nin Ht. nin (H10 _ _ H2 H3). am. change (gle r y x) in H11. cp (H1 _ _ H3 H2 H11). elim H5. nin Ha. order_tac. dneg. rww H6. Qed. Lemma Exercise1_21q: forall r, total_order r -> let r' := quotient_order r (scattered_equiv r) in small_set (substrate r') \/ without_gaps r'. Proof. ir. assert (Ha:=Exercise1_21m H). assert (Hb:=Exercise1_21o H). cp (Exercise1_2g Ha Hb H). cp (Exercise1_2h Ha Hb H). set (s:= scattered_equiv r) in *. assert (Hc:forall a b, gle r a b -> gle r' (class s a) (class s b)). ir. cp H; nin H. assert (inc a (substrate r)). order_tac. assert (inc b (substrate r)). order_tac. nin (Exercise1_2j Ha Hb H3 H5 H6). app H7. simpl in H0. fold r' in H1. fold r' in H0. nin (p_or_not_p (small_set (substrate r'))). au. right. red. ee. nin H1; am. app exists_proof2. dneg. red. ir. nin H1. nin (equal_or_not u v). am. nin (H6 _ _ H4 H5). elim (H3 u v). split; am. elim (H3 v u). split. am. au. clear H2. ir. assert (Hf:=H2). red in Ha. ee. assert (substrate r' = quotient s). uf r'. rww substrate_quotient_order. nin H2. rwi H0 H2. ee. assert (~ (related s (rep x) (rep y))). dneg. rwii related_rep_rep H11. app exists_proof. dneg. uf s. uf scattered_equiv. rw related_graph_on1. ee. order_tac. order_tac. uf scattered_rel. left. ee. am. assert (sub (interval_cc r (rep x) (rep y)) (substrate r)). uf interval_cc. app Z_sub. assert (order r). nin H; am. rww Exercise1_21a. ir. red. ir. red in H16. ee. nin H17. nin H17. assert (Hx:class s (rep x) = x). app class_rep. assert (Hy:class s (rep y) = y). app class_rep. assert (forall a, inc a u -> inc a x \/ inc a y). ir. cp (H15 _ H19). ufi interval_cc H20. Ztac. ee. cp (Hc _ _ H22). rwi Hx H24. nin (equal_or_not x (class s a)). rw H25. left. app inc_itself_class. ue. assert (glt r' x (class s a)). split;am. cp (Hc _ _ H23). rwi Hy H27. nin (equal_or_not (class s a) y). wr H28. right. app inc_itself_class. ue. assert (glt r' (class s a) y). split;am. elim (H12 (class s a)). au. cp (class_rep H4 H2). cp (class_rep H4 H9). assert (inc (rep x) (substrate r)). order_tac. assert (inc (rep y) (substrate r)). order_tac. assert (Hu:forall a, inc a x -> x = class s a). ir. app is_class_pr. assert (Hv:forall a, inc a y -> y = class s a). ir. app is_class_pr. cp (Exercise1_21n H H22). ufi s H20. rwi H20 H24. clear H20. clear H22. cp (Exercise1_21n H H23). ufi s H21. rwi H21 H20. clear H21. clear H23. rwi Exercise1_21a H24. rwi Exercise1_21a H20. set (u1:= intersection2 u x). assert (sub u1 x). uf u1. app intersection2sub_second. cp (H24 _ H21). assert (sub u1 (substrate r)). red. ir. ufi u1 H23. cp (intersection2_first H23). cp (H15 _ H25). ufi interval_cc H26. Ztac. am. rwi (Exercise1_21d H H23) H22. clear H23. clear H21. set (u2:= intersection2 u y). assert (sub u2 y). uf u2. app intersection2sub_second. cp (H20 _ H21). assert (sub u2 (substrate r)). red. ir. ufi u2 H25. cp (intersection2_first H25). cp (H15 _ H26). ufi interval_cc H27. Ztac. am. rwi (Exercise1_21d H H25) H23. clear H25. clear H21. assert (~(exists a, exists b, inc a u1 & inc b u1 & glt r a b)). red. ir. nin H21. nin H21. ee. nin H22. nin H26. elim H27. app H22. nin H22. nin H22. ee. assert (glt (induced_order r u) x4 x5). nin H28. split. aw. ufi u1 H22. inter2tac. ufi u1 H27. inter2tac. am. nin (H18 _ _ H30). ee. nin (related_induced_order4 H31). cp (related_induced_order2 H31). nin (related_induced_order4 H32). cp (related_induced_order2 H32). assert (inc x6 u1). uf u1. app intersection2_inc. nin (H19 _ H36). am. nin H38. cp (Hc _ _ H38). wri (Hv _ H39) H41. ufi u1 H27. cp (intersection2_second H27). wri (Hu _ H42) H41. nin H1. order_tac. nin (H29 _ H39). order_tac. order_tac. assert (~(exists a, exists b, inc a u2 & inc b u2 & glt r a b)). red. ir. nin H25. nin H25. ee. nin H23. nin H27. elim H28. app H23. nin H23. nin H23. ee. assert (glt (induced_order r u) x4 x5). nin H29. split. aw. ufi u2 H23. inter2tac. ufi u2 H28. inter2tac. am. nin (H18 _ _ H31). ee. nin (related_induced_order4 H32). cp (related_induced_order2 H32). nin (related_induced_order4 H33). cp (related_induced_order2 H33). assert (inc x6 u2). uf u2. app intersection2_inc. nin (H19 _ H37). nin H36. cp (Hc _ _ H36). wri (Hu _ H40) H42. ufi u2 H23. cp (intersection2_second H23). wri (Hv _ H43) H42. nin H1. order_tac. am. nin (H30 _ H40). order_tac. order_tac. nin (H18 _ _ H17). ee. nin (related_induced_order4 H26). nin (related_induced_order4 H27). cp (related_induced_order2 H26). cp (related_induced_order2 H27). nin (H19 _ H28). assert (inc x0 u1). uf u1. app intersection2_inc. nin (H19 _ H29). assert (inc x2 u1). uf u1. app intersection2_inc. elim H21. exists x0; exists x2; ee; am. nin (H19 _ H31). assert (inc x1 u1). uf u1. app intersection2_inc. elim H21. exists x0; exists x1; ee; tv. nin H32. nin H. order_tac. assert (inc x2 u2). uf u2. app intersection2_inc. assert (inc x1 u2). uf u2. app intersection2_inc. elim H25. exists x2; exists x1; ee; tv. assert (inc x0 u2). uf u2. app intersection2_inc. nin (H19 _ H29). assert (inc x2 u1). uf u1. app intersection2_inc. nin (H19 _ H31). assert (inc x1 u1). uf u1. app intersection2_inc. elim H21. exists x2; exists x1; ee; tv. assert (inc x1 u2). uf u2. app intersection2_inc. elim H25. exists x0; exists x1; ee; tv. nin H32. nin H. order_tac. assert (inc x2 u2). uf u2. app intersection2_inc. elim H25. exists x0; exists x2; ee; tv. am. wr H5. red. ir. app (inc_in_quotient_substrate H4 H21 H9). am. wr H5. red. ir. app (inc_in_quotient_substrate H4 H21 H2). Qed. (* Exercise 22 *) Definition open_o r u:= sub u (substrate r) & forall x y, inc x u -> gle r x y -> inc y u. Definition open_r r u:= open_o r u & forall v, open_o r v -> sub u v -> cofinal_set (induced_order r v) u -> u = v. Lemma Exercise1_22a: forall r u1 u2, order r -> open_o r u1 -> open_o r u2 -> open_o r (union2 u1 u2). Proof. ir. red. nin H0. nin H1. split. red. ir. nin (union2_or H4). app H0. app H1. ir. nin (union2_or H4). app union2_first. app (H2 _ _ H6 H5). app union2_second. app (H3 _ _ H6 H5). Qed. Lemma Exercise1_22b: forall r u, order r -> nonempty u -> (forall x, inc x u -> open_o r x) -> open_o r (intersection u). Proof. ir. split. red. ir. nin H0. cp (H1 _ H0). nin H3. app H3. app (intersection_forall H2 H0). ir. app intersection_inc. ir. nin (H1 _ H4). assert (inc x y0). app (intersection_forall H2 H4). ap (H6 _ _ H7 H3). Qed. Lemma Exercise1_22c: forall r u, order r -> (forall x, inc x u -> open_o r x) -> open_o r (union u). Proof. ir. split. red. ir. nin (union_exists H1). nin H2. nin (H0 _ H3). app H4. ir. nin (union_exists H1). nin H3. nin (H0 _ H4). cp (H6 _ _ H3 H2). union_tac. Qed. Lemma Exercise1_22d: forall r x u1 u2, order r -> open_o r x -> open_r r u1 -> open_r r u2 -> sub x u1 -> sub x u2 -> cofinal_set (induced_order r u1) x -> cofinal_set (induced_order r u2) x -> u1 = u2. Proof. ir. set (u3:= union2 u1 u2). assert (open_o r u3). uf u3. app Exercise1_22a. nin H1. ee; am. nin H2. ee;am. nin H1. assert (sub u1 u3). uf u3. red. ir. inter2tac. assert (cofinal_set (induced_order r u3) u1). red. assert (sub u3 (substrate r)). nin H7. am. aw. ee. am. ir. assert (exists a, inc a x & gle r x0 a). ufi u3 H11. nin (union2_or H11). nin H5. awi H13. nin (H13 _ H12). exists x1. ee. am. ap (related_induced_order1 H15). am. nin H1; am. nin H6. awi H13. nin (H13 _ H12). exists x1. ee. am. ap (related_induced_order1 H15). am. nin H2. nin H2; am. nin H12. nin H12. exists x1. ee. app H3. aw. uf u3. app union2_first. app H3. cp (H8 _ H7 H9 H10). nin H0. assert (sub u2 u3). uf u3. red. ir. inter2tac. assert (cofinal_set (induced_order r u3) u2). red. assert (sub u3 (substrate r)). nin H7. am. aw. ee. am. ir. wri H11 H15. nin H5. awi H16. nin (H16 _ H15). nin H17. cp (related_induced_order1 H18). exists x1. ee. app H4. aw. app H9. app H13. app H4. am. nin H1;am. nin H2. rw H11. rww (H15 _ H7 H13 H14). Qed. Definition bar1_22 r u := union (Zo (powerset(substrate r)) (fun z => open_o r z & cofinal_set (induced_order r z) u)). Lemma Exercise1_22e: forall r u, order r -> open_o r u -> sub u (bar1_22 r u). Proof. ir. cp H0; nin H0. red. ir. uf bar1_22. apply union_inc with u. am. Ztac. app powerset_inc. ee. am. split. aw. fprops. aw. ir. exists x0. ee. am. aw. order_tac. app H0. Qed. Lemma Exercise1_22f: forall r u, order r -> open_o r u -> sub (bar1_22 r u) (substrate r). Proof. ir. uf bar1_22. red. ir. nin (union_exists H1). ee. Ztac. rwi powerset_inc_rw H4. app H4. Qed. Lemma Exercise1_22g: forall r u, order r -> open_o r u -> cofinal_set (induced_order r (bar1_22 r u)) u. Proof. ir. cp (Exercise1_22e H H0). cp (Exercise1_22f H H0). split. aw. aw. ir. ufi bar1_22 H3. nin (union_exists H3). nin H4. Ztac. ee. rwi powerset_inc_rw H6. nin H8. awii H9. nin (H9 _ H4). exists x1. ee. am. aw. ap (related_induced_order1 H11). app H1. Qed. Lemma Exercise1_22h: forall r u x, order r -> open_o r u -> inc x (substrate r) -> (forall y, gle r x y -> exists z, inc z u & gle r y z) -> inc x (bar1_22 r u). Proof. ir. assert (Ha:=Exercise1_22f H H0). set (t:=union2 (bar1_22 r u) (Zo (substrate r) (fun z=> gle r x z))). assert (sub t (substrate r)). red. uf t. ir. nin (union2_or H3). app Ha. Ztac. am. uf bar1_22. apply union_inc with t. uf t. app union2_second. Ztac. order_tac. Ztac. app powerset_inc. split. uf t. app Exercise1_22a. red. split. am. assert (open_o r (bar1_22 r u)). uf bar1_22. app Exercise1_22c. ir. Ztac. eee. nin H4. am. red. split. app Z_sub. ir. Ztac. clear H4. Ztac. order_tac. order_tac. split. aw. uf t. red. ir. app union2_first. app (Exercise1_22e H H0). aw. ir. ufi t H4. nin (union2_or H4). cp (Exercise1_22g H H0). nin H6. awi H7. cp (H7 _ H5). nin H8. ee. exists x1. ee. am. aw. ap (related_induced_order1 H9). uf t. app union2_first. nin (related_induced_order3 H9). am. am. am. Ztac. nin (H2 _ H7). exists x1. eee. aw. uf t. uf t. app union2_first. app (Exercise1_22e H H0). Qed. Lemma Exercise1_22i: forall r u, order r -> open_o r u -> open_r r (bar1_22 r u). Proof. ir. assert (sub (bar1_22 r u) (substrate r)). nin H0. uf bar1_22. red. ir. nin (union_exists H2). ee. Ztac. rwi powerset_inc_rw H5. app H5. assert (open_o r (bar1_22 r u)). uf bar1_22. app Exercise1_22c. ir. Ztac. eee. split. am. ir. app extensionality. red. ir. app Exercise1_22h. nin H3. app H3. ir. nin H5. awi H8. nin H3. cp (H9 _ _ H6 H7). nin (H8 _ H10). nin H11. cp (related_induced_order1 H12). nin (Exercise1_22g H H0). awi H15. nin (H15 _ H11). nin H16. cp (related_induced_order1 H17). ex_tac. order_tac. am. am. am. nin H3. am. Qed. Lemma Exercise1_22j: forall r u v, order r -> open_o r u -> open_o r v -> sub u v -> sub (bar1_22 r u) (bar1_22 r v). Proof. ir. red. ir. app Exercise1_22h. app (Exercise1_22f H H0). ir. assert (inc y (bar1_22 r u)). nin (Exercise1_22i H H0). nin H5. app (H7 _ _ H3 H4). nin (Exercise1_22g H H0). cp (Exercise1_22f H H0). awii H7. nin (H7 _ H5). nin H9. exists x0. ee. app H2. ap (related_induced_order1 H10). Qed. Lemma Exercise1_22k: forall r u v, order r -> open_o r u -> open_o r v -> intersection2 u v = emptyset -> intersection2 (bar1_22 r u) (bar1_22 r v) = emptyset. Proof. ir. app is_emptyset. ir. red. ir. nin (intersection2_both H3). ufi bar1_22 H4. nin (union_exists H4). nin H6. Ztac. clear H7. ee. nin H9. nin H7. awii H10. nin (H10 _ H6). nin H12. cp (related_induced_order1 H13). ufi bar1_22 H5. nin (union_exists H5). nin H15. Ztac. ee. nin H18. cp (H20 _ _ H15 H14). nin H19. awii H22. nin (H22 _ H21). nin H23. cp (related_induced_order1 H24). nin H0. cp (H26 _ _ H12 H25). elim (emptyset_pr (x:=x2)). wr H2. app intersection2_inc. Qed. Lemma Exercise1_22m: forall r, order r -> open_r r emptyset. Proof. ir. split. split. app sub_emptyset_any. ir. elim (emptyset_pr H0). ir. nin H2. nin H0. awii H3. nin (emptyset_dichot v). sy; am. nin H5. nin (H3 _ H5). nin H6. elim (emptyset_pr H6). Qed. Lemma Exercise1_22n: forall r, order r -> open_r r (substrate r). Proof. ir. split. split. fprops. ir. order_tac. ir. nin H0. app extensionality. Qed. Lemma Exercise1_22o: forall r, order r -> ~ (right_directed r) -> exists a, exists b, open_r r a & open_r r b & nonempty a & nonempty b& a <> b. Proof. ir. set (i:= fun x => Zo (substrate r) (fun z => gle r x z)). assert (forall x, inc x (substrate r) -> open_o r (i x)). ir. split. uf i. app Z_sub. uf i. ir. Ztac. clear H2. Ztac. order_tac. order_tac. assert (forall x, inc x (substrate r) -> open_r r (bar1_22 r (i x))). ir. app Exercise1_22i. app H1. assert (forall x, inc x (substrate r) -> inc x (bar1_22 r (i x))). ir. app Exercise1_22e. app H1. uf i. Ztac. order_tac. set (p := fun a b => inc a (substrate r) & inc b (substrate r) & intersection2 (i a) (i b) = emptyset). nin (p_or_not_p (exists a , exists b, p a b)). nin H4. nin H4. ufi p H4. ee. exists (bar1_22 r (i x)). exists (bar1_22 r (i x0)). cp (H1 _ H4). cp (H1 _ H5). cp (Exercise1_22k H H7 H8 H6). ee. app H2. app H2. exists x. app H3. exists x0. app H3. red. ir. rwi H10 H9. cp (H3 _ H5). elim (emptyset_pr (x:=x0)). wr H9. app intersection2_inc. elim H0. rw right_directed_pr. ee. am. ir. app exists_proof. dneg. exists x;exists y. uf p. eee. app is_emptyset. ir. cp (H7 y0). dneg. ufi i H9. nin (intersection2_both H9). Ztac. clear H11. Ztac. eee. Qed. Lemma Exercise1_22p: forall r x, order r -> open_r r x -> x = (bar1_22 r x) . Proof. ir. cp H0. nin H1. cp (Exercise1_22g H H1). nin (Exercise1_22i H H1). ap (H2 _ H4 (Exercise1_22e H H1) H3). Qed. Definition set_of_reg_open r := Zo (powerset (substrate r)) (fun z => open_r r z). Lemma Exercise1_22q: forall r, order r -> (exists a, exists b, a <> b & set_of_reg_open r = doubleton a b) = (nonempty (substrate r) & (right_directed r)). Proof. ir. ap iff_eq. ir. nin H0; nin H0. nin H0. assert (inc x (set_of_reg_open r)). rw H1. fprops. assert (inc x0 (set_of_reg_open r)). rw H1. fprops. ufi set_of_reg_open H2. Ztac. clear H2. ufi set_of_reg_open H3. Ztac. rwi powerset_inc_rw H4. rwi powerset_inc_rw H2. split. nin (emptyset_dichot x). nin (emptyset_dichot x0). elim H0. ue. nin H8. exists y. app H2. nin H7. exists y. app H4. nin (p_or_not_p (right_directed r)). am. cp (Exercise1_22o H H7). nin H8. nin H8. clear H3. ee. assert (inc x1 (set_of_reg_open r)). uf set_of_reg_open. Ztac. app powerset_inc. nin H3. nin H3. am. assert (inc x2 (set_of_reg_open r)). uf set_of_reg_open. Ztac. app powerset_inc. nin H8. nin H8. am. assert (inc emptyset (set_of_reg_open r)). uf set_of_reg_open. Ztac. app powerset_inc. app sub_emptyset_any. app Exercise1_22m. rwi H1 H14. rwi H1 H12. rwi H1 H13. nin H10; nin H9. nin (doubleton_or H14). nin (doubleton_or H13). elim (emptyset_pr (x:=y)). rww H15. wrr H16. nin (doubleton_or H12). elim (emptyset_pr (x:=y0)). rww H15. ue. elim H11. ue. nin (doubleton_or H13). nin (doubleton_or H12). elim H11. ue. elim (emptyset_pr (x:=y0)). rww H15. ue. elim (emptyset_pr (x:=y)). rww H15. ue. ir. ee. exists emptyset. exists (substrate r). ee. red. ir. nin H0. wri H2 H0. nin H0. elim x. sy. set_extens. nin (doubleton_or H2). rw H3. uf set_of_reg_open. Ztac. app powerset_inc. app sub_emptyset_any. app Exercise1_22m. rw H3. uf set_of_reg_open. Ztac. app powerset_inc. fprops. app Exercise1_22n. nin (emptyset_dichot x). rw H3. fprops. nin H3. ufi set_of_reg_open H2. Ztac. clear H2. cut (x = substrate r). ir. rw H2. fprops. rwi powerset_inc_rw H4. app extensionality. red. ir. assert (inc x0 (bar1_22 r x)). app Exercise1_22h. nin H5. am. ir. rwi right_directed_pr H1. nin H1. assert (inc y (substrate r)). app H4. assert (inc y0 (substrate r)). order_tac. nin (H7 _ _ H8 H9). exists x1. ee. nin H5. nin H5. app (H14 _ _ H3 H11). am. rww (Exercise1_22p H H5). Qed. Definition reg_open_order r := inclusion_suborder (set_of_reg_open r). Lemma Exercise1_22r: forall r u v, order r -> gle (reg_open_order r) u v = (open_r r u & open_r r v& sub u v). Proof. ir. uf reg_open_order. aw. uf set_of_reg_open. ap iff_eq. ir. ee. clear H1. Ztac. am. Ztac. am. am. ir. ee. Ztac. nin H0. nin H0. app powerset_inc. Ztac. nin H1. nin H1. app powerset_inc. am. Qed. Lemma Exercise1_22s: forall r, order r -> greatest_element (reg_open_order r) (substrate r). Proof. ir. cp (Exercise1_22n H). split. uf reg_open_order. aw. uf set_of_reg_open. Ztac. app powerset_inc. fprops. ir. rw Exercise1_22r. ufi reg_open_order H1. awi H1. ufi set_of_reg_open H1. Ztac. ee. am. am. nin H3. nin H3. am. am. Qed. Lemma Exercise1_22t: forall r, order r -> least_element (reg_open_order r) (emptyset). Proof. ir. cp (Exercise1_22m H). split. uf reg_open_order. aw. uf set_of_reg_open. Ztac. app powerset_inc. app sub_emptyset_any. ir. rw Exercise1_22r. ufi reg_open_order H1. awi H1. ufi set_of_reg_open H1. Ztac. ee. am. am. app sub_emptyset_any. am. Qed. Lemma inf_pr2: forall r x y z, order r -> gle r z x -> gle r z y -> (forall t, gle r t x -> gle r t y -> gle r t z) -> inf r x y = z. Proof. ir. cp (greatest_lower_bound_doubleton H H0 H1 H2). sy. uf inf. app infimum_pr2. red. ir. nin (doubleton_or H4); rw H5; order_tac. Qed. Lemma Exercise1_22u: forall r u v, order r -> open_r r u -> open_r r v -> inf (reg_open_order r) u v = bar1_22 r (intersection2 u v). Proof. ir. set (z := bar1_22 r (intersection2 u v)). assert (open_o r (intersection2 u v)). uf intersection2. app Exercise1_22b. app nonempty_doubleton. ir. nin H0; nin H1. nin (doubleton_or H2); ue. cp (Exercise1_22i H H2). cp (Exercise1_22e H H2). assert (gle (reg_open_order r) z u). uf z. rw Exercise1_22r. ee. am. am. cp (Exercise1_22p H H0). rw H5. app Exercise1_22j. wr H5. am. nin H0. am. wr H5. app intersection2sub_first. am. assert (gle (reg_open_order r) z v). uf z. rw Exercise1_22r. ee. am. am. cp (Exercise1_22p H H1). rw H6. app Exercise1_22j. wr H6. am. nin H1. am. wr H6. app intersection2sub_second. am. assert (forall t, gle (reg_open_order r) t u -> gle (reg_open_order r) t v -> gle (reg_open_order r) t z). ir. rwi Exercise1_22r H7. rwi Exercise1_22r H8. ee. rw Exercise1_22r. ee. am. am. apply sub_trans with (intersection2 u v). red. ir. app intersection2_inc. app H12. app H10. am. am. am. am. assert (order (reg_open_order r)). uf reg_open_order. fprops. ap (inf_pr2 H8 H5 H6 H7). Qed. Lemma Exercise1_22v: forall r X, order r -> sub X (substrate (reg_open_order r)) -> least_upper_bound (reg_open_order r) X (bar1_22 r (union X)). Proof. ir. assert (order (reg_open_order r)). uf reg_open_order. fprops. assert (substrate (reg_open_order r) = (set_of_reg_open r)). uf reg_open_order. aw. aw. assert (open_o r (union X)). app Exercise1_22c. ir. cp (H0 _ H3). rwi H2 H4. ufi set_of_reg_open H4. Ztac. nin H6. am. assert (open_r r (bar1_22 r (union X))). app Exercise1_22i. assert (inc (bar1_22 r (union X)) (set_of_reg_open r)). uf set_of_reg_open. Ztac. app powerset_inc. app Exercise1_22f. split. split. ue. ir. rw Exercise1_22r. rwi H2 H0. cp (H0 _ H6). ufi set_of_reg_open H7. Ztac. ee. am. am. apply sub_trans with (union X). app union_sub. app Exercise1_22e. am. ir. nin H6. rwi H2 H6. ufi set_of_reg_open H6. Ztac. rw Exercise1_22r. ee. am. am. rw (Exercise1_22p H H9). app Exercise1_22j. nin H9. am. red. ir. nin (union_exists H10). nin H11. cp (H7 _ H12). rwi Exercise1_22r H13. ee. app H15. am. am. Qed. Lemma Exercise1_22w: forall r u v, order r -> open_r r u -> open_r r v -> sup (reg_open_order r) u v = bar1_22 r (union2 u v). Proof. ir. assert (sub (doubleton u v) (substrate (reg_open_order r))). uf reg_open_order. aw. red. ir. uf set_of_reg_open. nin (doubleton_or H2); rw H3; Ztac; app powerset_inc. nin H0; nin H0; am. nin H1; nin H1; am. cp (Exercise1_22v H H2). uf sup. sy; app supremum_pr2. uf reg_open_order. fprops. Qed. Lemma Exercise1_22x: forall r, order r -> complete_lattice (reg_open_order r). Proof. ir. assert (order (reg_open_order r)). uf reg_open_order. fprops. app exercise1_11b. ir. exists (bar1_22 r (union X)). ap (Exercise1_22v H H1). Qed. Lemma Exercise1_22y: forall r, order r -> lattice (reg_open_order r). Proof. ir. cp (Exercise1_22x H). red. nin H0. ee. am. ir. assert (sub (doubleton x y) (substrate (reg_open_order r))). red. ir. nin (doubleton_or H4); rww H5. ap (H1 _ H4). Qed. Lemma Exercise1_22z: forall r, order r -> relatively_complemented (reg_open_order r). Proof. ir. cp (Exercise1_22t H). assert (order (reg_open_order r)). uf reg_open_order. fprops. assert (the_least_element (reg_open_order r) = emptyset). app the_least_element_pr2. red. ee. app Exercise1_22y. exists emptyset. am. rw H2. ir. rwi Exercise1_22r H3. ee. set (z:= Zo y (fun u => forall t, gle r u t -> ~ (inc t x))). assert (sub z (substrate r)). apply sub_trans with y. uf z; app Z_sub. nin H4. nin H4. am. assert (open_o r z). red. ee. am. ir. ufi z H7. Ztac. clear H7. uf z. Ztac. nin H4. nin H4. app (H11 _ _ H9 H8). ir. app H10. order_tac. cp (Exercise1_22i H H7). assert (open_o r (union2 x (bar1_22 r z))). app Exercise1_22a. nin H3; am. nin H8; am. exists (bar1_22 r z). split. uf reg_open_order. aw. uf set_of_reg_open. Ztac. app powerset_inc. nin H8; nin H8; am. split. rw Exercise1_22w. app extensionality. rw (Exercise1_22p H H4). app Exercise1_22j. nin H4; am. red. ir. nin (union2_or H10). app H5. rw (Exercise1_22p H H4). nin H4. assert (sub z y). uf z. app Z_sub. app (Exercise1_22j H H7 H4 H13). red. ir. app Exercise1_22h. nin H4. nin H4. app H4. ir. assert (inc y0 y). nin H4; nin H4. app (H13 _ _ H10 H11). nin (p_or_not_p (exists t, inc t x & gle r y0 t)). nin H13. nin H13. exists x1. ee. inter2tac. am. assert (inc y0 z). uf z. Ztac. ir. nin (inc_or_not t x). elim H13. exists t. au. am. exists y0. split. app union2_second. app (Exercise1_22e H H7). order_tac. order_tac. am. am. am. rw Exercise1_22u. assert (intersection2 x z = emptyset). app is_emptyset. ir. red. ir. nin (intersection2_both H10). ufi z H12. Ztac. assert (gle r y0 y0). order_tac. nin H3. nin H3. app H3. elim (H14 _ H15). am. cp (Exercise1_22p H H3). nin H3. cp (Exercise1_22k H H3 H7 H10). wri H11 H13. rw H13. sy. app Exercise1_22p. app Exercise1_22m. am. am. am. am. Qed. Lemma Exercise1_22A: forall r, order r -> boolean_lattice (reg_open_order r). Proof. ir. cp (Exercise1_22t H). assert (order (reg_open_order r)). uf reg_open_order. fprops. split. app Exercise1_22z. split. exists (substrate r). app (Exercise1_22s H). cp (Exercise1_22y H). rww (Exercise1_16c H2). red. ir. set (a1:= sup (reg_open_order r) x y). set (a2:= inf (reg_open_order r) y z). assert (substrate (reg_open_order r) = (set_of_reg_open r)). uf reg_open_order. aw. aw. assert (inc a1 (substrate (reg_open_order r))). uf a1. cp (lattice_sup_pr H2 H3 H4). ee. order_tac. assert (inc a2 (substrate (reg_open_order r))). uf a2. cp (lattice_inf_pr H2 H4 H5). ee. order_tac. rww (Exercise1_22r). ee. cp (lattice_inf_pr H2 H5 H7). ee. assert (inc (inf (reg_open_order r) z a1) (substrate (reg_open_order r))). order_tac. rwi H6 H12. ufi set_of_reg_open H12. Ztac. am. cp (lattice_sup_pr H2 H3 H8). ee. assert (inc (sup (reg_open_order r) x a2) (substrate (reg_open_order r))). order_tac. rwi H6 H12. ufi set_of_reg_open H12. Ztac. am. rwi H6 H3; rwi H6 H4; rwi H6 H5; rwi H6 H7;rwi H6 H8. ufi set_of_reg_open H3; ufi set_of_reg_open H4; ufi set_of_reg_open H5. ufi set_of_reg_open H7; ufi set_of_reg_open H8. Ztac. clear H8. Ztac. clear H7. Ztac. clear H5. Ztac. clear H4. Ztac. clear H3. clear H8; clear H9; clear H7; clear H5; clear H4. rw (Exercise1_22u H H12 H11). rw (Exercise1_22w H H14 H10). assert (a1 = bar1_22 r (union2 x y)). uf a1. rww (Exercise1_22w H H14 H13). assert (a2 = bar1_22 r (intersection2 y z)). uf a2. rww (Exercise1_22u H H13 H12). red. ir. ufi bar1_22 H5. nin (union_exists H5). clear H6. nin H7. Ztac. clear H7. clear H8. ee. assert (sub x1 (substrate r)). nin H7. am. clear H5. nin H8. awii H5. awii H8. app Exercise1_22h. app Exercise1_22a. nin H14. am. nin H10. am. nin H7. app H7. ir. assert (inc y0 x1). nin H7. app (H16 _ _ H6 H15). nin (H8 _ H16). nin H17. nin (intersection2_both H17). rwi H3 H20. ufi bar1_22 H20. nin (union_exists H20). clear H20. nin H21. Ztac. clear H21. ee. nin H23. rwi powerset_inc_rw H22. awii H23. awii H24. cp (related_induced_order1 H18). nin (H24 _ H20). nin H26. cp (related_induced_order1 H27). assert (gle r y0 x4). order_tac. exists x4. ee. nin (union2_or H26). app union2_first. app union2_second. rww H4. app Exercise1_22e. uf intersection2. app Exercise1_22b. app nonempty_doubleton. ir. nin (doubleton_or H31); rw H32. nin H13; am. nin H12; am. app intersection2_inc. nin H12. nin H12. app (H32 _ _ H19 H28). am. Qed. Lemma Exercise1_22B: forall r F x, order r -> cofinal_set r F -> open_r r x -> open_r (induced_order r F) (intersection2 x F). Proof. ir. assert (sub F (substrate r)). nin H0. am. assert (order (induced_order r F)). fprops. assert (open_o (induced_order r F) (intersection2 x F)). ir. red. ee. aw. app intersection2sub_second. ir. nin (intersection2_both H4). cp (related_induced_order1 H5). nin (related_induced_order3 H5). app intersection2_inc. nin H1. nin H1. ap (H12 _ _ H6 H8). red. split. am. ir. app extensionality. red. ir. assert (inc x0 F). nin H5. awi H5. app H5. am. am. app intersection2_inc. rw (Exercise1_22p H H1). cp (H2 _ H9). app Exercise1_22h. nin H1; am. ir. nin H0. assert (inc y (substrate r)). order_tac. nin (H12 _ H13). nin H14. nin H7. awii H16. assert (gle (induced_order r F) x0 x1). aw. order_tac. nin H5. cp (H18 _ _ H8 H17). nin (H16 _ H19). nin H20. exists x2. ee. inter2tac. cp (related_induced_order1 H21). cp (related_induced_order1 H22). order_tac. nin H5. am. Qed. Lemma Exercise1_22C: forall r F U U', order r -> cofinal_set r F -> open_r r U -> open_r r U' -> sub (intersection2 U F) (intersection2 U' F) -> sub U U'. Proof. ir. red. ir. rw (Exercise1_22p H H2). app Exercise1_22h. nin H2. am. nin H1; nin H1. app H1. ir. assert (inc y (substrate r)). order_tac. nin H0. nin (H7 _ H6). nin H8. nin H1. nin H1. assert (gle r x x0). order_tac. cp (H11 _ _ H4 H12). assert (inc x0 (intersection2 U' F)). app H3. app intersection2_inc. cp (intersection2_first H14). ex_tac. Qed. Lemma Exercise1_22D: forall r F, order r -> cofinal_set r F -> order_isomorphism (BL (fun z => intersection2 z F) (set_of_reg_open r) (set_of_reg_open (induced_order r F))) (reg_open_order r)(reg_open_order (induced_order r F)). Proof. ir. assert (sub F (substrate r)). nin H0. am. assert (order (induced_order r F)). fprops. set (r':= induced_order r F) in *. assert (substrate r' = F). uf r'. aw. assert (forall x, open_r r x -> open_r r' (intersection2 x F)). ir. app (Exercise1_22B H H0 H4). assert (transf_axioms (fun z => intersection2 z F) (set_of_reg_open r) (set_of_reg_open r')). red. ir. ufi set_of_reg_open H5. Ztac. clear H5. cp (H4 _ H7). uf set_of_reg_open. Ztac. app powerset_inc. nin H5. nin H5;am. set (g:=BL (fun z : Set => intersection2 z F) (set_of_reg_open r) (set_of_reg_open r')). assert (is_function g). uf g. app bl_function. assert (bijective g). uf g. app bijective_bl_function. ir. ufi set_of_reg_open H7. Ztac. ufi set_of_reg_open H8. Ztac. app extensionality. app (Exercise1_22C (U:=u) (U':=v) H H0). rw H9; fprops. app (Exercise1_22C (U:=v) (U':=u) H H0). rw H9; fprops. ir. ufi set_of_reg_open H7. Ztac. clear H7. clear H8. assert (sub y (substrate r)). nin H9. nin H7. apply sub_trans with F. ue. am. set (x1:= Zo (substrate r) (fun z => exists x, inc x y & gle r x z)). assert (sub x1 (substrate r)). uf x1. app Z_sub. assert (sub y x1). red. ir. uf x1. Ztac. exists x. ee. am. order_tac. app H7. assert (open_o r x1). red. ee. am. ir. ufi x1 H11. Ztac. clear H11. nin H14. uf x1. Ztac. order_tac. nin H11. ex_tac. order_tac. set (x2:= bar1_22 r x1). assert (sub y (intersection2 x2 F)). red. ir. app intersection2_inc. uf x2. app (Exercise1_22e H H11). app H10. nin H9. nin H9. wr H3. app H9. cp (Exercise1_22i H H11). exists (bar1_22 r x1). split. uf set_of_reg_open. Ztac. app powerset_inc. app Exercise1_22f. app extensionality. red. ir. nin (intersection2_both H14). rw (Exercise1_22p H2 H9). app Exercise1_22h. nin H9. am. ue. ir. ufi r' H17. cp (related_induced_order1 H17). assert (inc y0 (bar1_22 r x1)). nin H13. nin H13. ap (H20 _ _ H15 H18). ufi bar1_22 H19. nin (union_exists H19). nin H20. Ztac. nin H23. rwi powerset_inc_rw H22. nin H24. awii H25. nin (H25 _ H20). nin H26. cp (related_induced_order1 H27). ufi x1 H26. Ztac. nin H30. nin H30. nin H0. nin (H32 _ H29). nin H33. assert (gle r' x4 x5). uf r'. aw. order_tac. nin H9. nin H9. ue. exists x5. ee. nin H9. nin H9. app (H37 _ _ H30 H35). uf r'. aw. order_tac. nin (related_induced_order3 H17). am. red. ee. uf reg_open_order. fprops. uf reg_open_order. fprops. am. uf g. uf reg_open_order. aw. uf reg_open_order. uf g. aw. uf g. aw. ir. aw. uf reg_open_order. aw. ap iff_eq. ir. ee. app H5. app H5. red. ir. nin (intersection2_both H13). app intersection2_inc. app H12. ir. ee. am. am. ufi set_of_reg_open H8. Ztac. ufi set_of_reg_open H9. Ztac. app (Exercise1_22C H H0 H14 H16 H12). Qed. Lemma Exercise1_22E: forall r r' X X', order r -> order r' -> open_o r X -> open_o r' X' -> open_o (product2_order r r') (product X X'). Proof. ir. split. rww substrate_order_product2_order. app product_monotone. nin H1. am. nin H2. am. ir. rwi product2_order_pr H4. red in H4. ee. awi H3. ee. nin H1. cp (H10 _ _ H8 H6). nin H2. cp (H12 _ _ H9 H7). awi H5. aw. eee. Qed. Lemma Exercise1_22F: forall E X, let r := diagonal E in sub X E-> open_r r X. Proof. ir. assert (forall x y, gle r x y -> x = y). uf gle. uf related. uf r. ir. awi H0. ee. am. cp (diagonal_order E). split. split. uf r. rww substrate_diagonal. ir. wrr (H0 _ _ H3). ir. app extensionality. nin H4. awii H5. red. ir. nin (H5 _ H6). nin H7. awii H8. rww (H0 _ _ H8). nin (related_induced_order3 H8). am. nin H2; am. Qed. Lemma Exercise1_22G: forall E, let r := diagonal E in (product2_order r r = diagonal (product E E)). Proof. ir. assert (substrate r = E). uf r. rww substrate_diagonal. assert (forall x y, related r x y -> x = y). uf related. uf r. ir. awi H0. ee. am. uf product2_order. rw H. uf graph_on. set_extens. Ztac. nin H3. rwi H H3. ee. cp (H0 _ _ H5). cp (H0 _ _ H6). rw inc_diagonal. assert (is_pair x). awi H2; ee; am. ee. am. am. awi H3; awi H4. ee. app pair_extensionality. rwi inc_diagonal H1. ee. assert (Ha:= H2). awi H2. ee. Ztac. aw. ee. am. am. am. am. ue. ue. ue. red. rw H. wr H3. ee. am. am. uf r. red. rww inc_diagonal. ee. fprops. aw. aw. red. uf r. rww inc_diagonal. ee. fprops. aw. aw. Qed. (* Exercice 23 *) Definition set_of_nreg_open r := complement (set_of_reg_open r) (singleton emptyset). Definition set_of_nreg_order r := opposite_order (inclusion_suborder (set_of_nreg_open r)). Lemma Exercise1_23a: forall r X, inc X (set_of_nreg_open r) = (open_r r X & nonempty X). Proof. ir. uf set_of_nreg_open. ap iff_eq. ir. srwi H. nin H. ufi set_of_reg_open H. Ztac. ee. am. nin (emptyset_dichot X). elim H0. rw H3. fprops. am. ir. ee. srw. split. uf set_of_reg_open. Ztac. app powerset_inc. nin H; nin H; am. red. ir. awi H1. rwi H1 H0. nin H0. elim (emptyset_pr H0). Qed. Lemma Exercise1_23b: forall r X Y, order r -> gle (set_of_nreg_order r) X Y = (nonempty X & nonempty Y & open_r r X & open_r r Y & sub Y X). Proof. ir. uf set_of_nreg_order. aw. rw Exercise1_23a. rw Exercise1_23a. ap iff_eq. ir. eee. ir. eee. fprops. Qed. Definition canonical_reg_open r x := bar1_22 r (Zo (substrate r) (fun z => gle r x z)). Lemma Exercise1_23c: forall r x, order r -> open_o r (Zo (substrate r) (fun z => gle r x z)). Proof. ir. split. app Z_sub. ir. Ztac. clear H0. Ztac. order_tac. order_tac. Qed. Lemma Exercise1_23d1: forall r x, order r -> inc x (substrate r) -> inc x (canonical_reg_open r x). Proof. ir. uf canonical_reg_open. app Exercise1_22e. ap (Exercise1_23c x H). Ztac. order_tac. Qed. Lemma Exercise1_23d2: forall r x, order r -> inc x (substrate r) -> inc (canonical_reg_open r x) (set_of_nreg_open r). Proof. ir. rw Exercise1_23a. split. uf canonical_reg_open. cp (Exercise1_23c x H). app Exercise1_22i. exists x. app Exercise1_23d1. Qed. Lemma Exercise1_23e: forall r x y, order r -> inc x (substrate r) -> inc y (substrate r) -> (inc y (canonical_reg_open r x) = forall z, gle r y z -> exists t, gle r z t & gle r x t). Proof. ir. uf canonical_reg_open. app iff_eq. ir. ufi bar1_22 H2. nin (union_exists H2). nin H4. Ztac. nin H7. nin H8. awii H9. assert (inc z x0). nin H7. ap (H10 _ _ H4 H3). nin (H9 _ H10). nin H11. Ztac. cp (related_induced_order1 H12). exists x1. au. nin H7. am. ir. app Exercise1_22h. app Exercise1_23c. ir. nin (H2 _ H3). nin H4. exists x0. ee. Ztac. order_tac. am. Qed. Lemma Exercise1_23f: forall r x y, order r -> gle r x y -> gle (set_of_nreg_order r) (canonical_reg_open r x) (canonical_reg_open r y). Proof. ir. assert (inc x (substrate r)). order_tac. assert (inc y (substrate r)). order_tac. cp (Exercise1_23d2 H H1). cp (Exercise1_23d2 H H2). uf set_of_nreg_order. aw. ir. eee. uf canonical_reg_open. app Exercise1_22j. app Exercise1_23c. app Exercise1_23c. red. ir. Ztac. clear H5. Ztac. order_tac. fprops. Qed. Lemma Exercise1_23g: forall r, order r -> cofinal_set (set_of_nreg_order r) (fun_image (substrate r) (canonical_reg_open r)). Proof. ir. assert (substrate (set_of_nreg_order r) = set_of_nreg_open r). uf set_of_nreg_order. aw. fprops. red. split. red. ir. awi H1. nin H1. nin H1. rw H0. wr H2. app Exercise1_23d2. fprops. ir. rwi H0 H1. rwi Exercise1_23a H1. nin H1. nin H2. assert (inc y (substrate r)). nin H1. nin H1. app H1. cp (Exercise1_23d2 H H3). rwi Exercise1_23a H4. exists (canonical_reg_open r y). split. aw. exists y. split. am. tv. rww Exercise1_23b. eee. exists y. am. uf canonical_reg_open. rw (Exercise1_22p H H1). app Exercise1_22j. app Exercise1_23c. nin H1; am. red. ir. Ztac. nin H1. nin H1. ap (H10 _ _ H2 H8). Qed. Definition anti_directed r:= forall x y, inc x (substrate r) -> inc y (substrate r) -> (canonical_reg_open r x) = (canonical_reg_open r y) -> x = y. Lemma Exercise1_23h: forall r, order r -> let aux := (fun x y => forall z, gle r x z -> gle r y z -> False) in (anti_directed r) = ((forall x y, glt r x y -> exists z, (glt r x z & aux y z)) & forall x y, inc x (substrate r) -> inc y (substrate r) -> (gle r x y \/ gle r y x \/ (exists x', gle r x x' & aux x' y) \/ (exists y', gle r y y' & aux x y'))). Proof. ir. assert (forall x y, inc x (substrate r) -> inc y (substrate r) -> (inc y (canonical_reg_open r x) = forall z, gle r y z -> ~ (aux x z))). ir. rww Exercise1_23e. uf aux. ap iff_eq. ir. red. ir. nin (H2 _ H3). nin H5. app (H4 _ H6 H5). ir. cp (H2 _ H3). app exists_proof. dneg. ir. elim (H5 z0). au. assert (Hb:forall x y, gle r x y -> sub (canonical_reg_open r y) (canonical_reg_open r x)). ir. cp (Exercise1_23f H H1). rwi Exercise1_23b H2. ee; am. am. assert (Hc:forall x y, aux x y -> aux y x). uf aux. ir. app (H1 _ H3 H2). sy. app iff_eq. ir. red. ir. nin H1. assert (inc x (canonical_reg_open r x)). app (Exercise1_23d1). assert (inc y (canonical_reg_open r y)). app (Exercise1_23d1). nin (equal_or_not x y). am. nin (H5 _ _ H2 H3). assert (glt r x y). split;am. nin (H1 _ _ H10). nin H11. rwi H4 H6. rwi H0 H6. nin H11. elim (H6 _ H11). am. am. am. nin H9. assert (glt r y x). split;au. nin (H1 _ _ H10). nin H11. wri H4 H7. rwi H0 H7. nin H11. elim (H7 _ H11). am. am. am. nin H9. nin H9. nin H9. rwi H4 H6. rwi H0 H6. elim (H6 _ H9). app Hc. am. am. nin H9; nin H9. wri H4 H7. rwi H0 H7. elim (H7 _ H9). am. am. am. ir. assert (forall x y, inc x (substrate r) -> inc y (substrate r) -> (sub (canonical_reg_open r x) (canonical_reg_open r y) \/ (exists x' : Set, gle r x x' & aux x' y))). ir. nin (p_or_not_p (exists x' : Set, gle r x x' & aux x' y)). au. left. cp (Exercise1_23d2 H H3). rwi Exercise1_23a H5. nin H5. rw (Exercise1_22p H H5). uf canonical_reg_open. app Exercise1_22j. app Exercise1_23c. nin H5. am. red. ir. ufi canonical_reg_open H0. rw H0. Ztac. ir. nin (p_or_not_p (aux y z)). elim H4. exists z. ee. order_tac. app Hc. am. am. Ztac. am. cut ((forall x y : Set, glt r x y -> exists z : Set, glt r x z & aux y z)). ir. split. am. ir. nin (p_or_not_p (gle r x y)). left. am. right. nin (p_or_not_p (gle r y x)). left. am. right. nin (H2 _ _ H4 H5). nin (H2 _ _ H5 H4). assert (canonical_reg_open r x = canonical_reg_open r y). app extensionality. red in H1. rwi (H1 _ _ H4 H5 H10) H6. elim H6. order_tac. right. nin H9. nin H9. exists x0. ee. am. app Hc. left. am. ir. assert (inc x (substrate r)). order_tac. assert (inc y (substrate r)). order_tac. nin H3. nin (H2 _ _ H4 H5). cp (Exercise1_23f H H3). rwi Exercise1_23b H8. ee. assert (canonical_reg_open r x = canonical_reg_open r y). app extensionality. red in H1. rwi (H1 _ _ H4 H5 H13) H6. elim H6. tv. am. nin H7. exists x0. nin H7. split. split. am. red. ir. wri H9 H8. red in H8. elim (H8 y). am. order_tac. app Hc. Qed. Lemma Exercise1_23i: forall r x y, order r -> inc x y -> inc y (set_of_nreg_open r) -> gle (set_of_nreg_order r) y (canonical_reg_open r x). Proof. ir. rw Exercise1_23b. rwi Exercise1_23a H1. nin H1. assert (sub y (substrate r)). nin H1. nin H1. am. cp (Exercise1_23d2 H (H3 _ H0)). rwi Exercise1_23a H4. eee. uf canonical_reg_open. rw (Exercise1_22p H H1). app Exercise1_22j. app Exercise1_23c. nin H1; am. red. ir. Ztac. nin H1; nin H5. nin H1. ap (H10 _ _ H0 H8). am. Qed. Lemma Exercise1_23j: forall r, order r -> let r':= set_of_nreg_order r in (forall x y, inc x (substrate r') -> inc y (substrate r') -> (forall z, gle r' x z -> gle r' y z -> False) = (disjoint x y)). Proof. ir. assert (order r'). uf r'. uf set_of_nreg_order. fprops. assert (substrate r' = set_of_nreg_open r). uf r'. uf set_of_nreg_order. aw. fprops. ap iff_eq. ir. app disjoint_pr. ir. rwi H3 H0; rwi H3 H1. ap (H4 _ (Exercise1_23i H H5 H0) (Exercise1_23i H H6 H1)). ir. red in H4. ufi r' H5. rwi Exercise1_23b H5. ufi r' H6. rwi Exercise1_23b H6. ee. nin H11. elim (emptyset_pr (x:=y0)). wr H4. app intersection2_inc. app H14. app H10. am. am. Qed. Lemma Exercise1_23k: forall r, order r -> anti_directed (set_of_nreg_order r). Proof. ir. set (r':= set_of_nreg_order r). assert (order r'). uf r'. uf set_of_nreg_order. fprops. set (aux:=(fun x y => forall z, gle r' x z -> gle r' y z -> False)). assert (forall x y, inc x (substrate r') -> inc y (substrate r') -> (disjoint x y) -> aux x y). ir. uf aux. uf r'. rww Exercise1_23j. assert (Hv:substrate r' = set_of_nreg_open r). uf r'. uf set_of_nreg_order. aw. fprops. set (i := fun x y => bar1_22 r (Zo x (fun u => forall t, gle r u t -> ~ (inc t y)))). assert (forall x y, inc x (substrate r') -> inc y (substrate r') -> (open_r r (i x y) & sub (i x y) x & disjoint (i x y) y & (i x y = emptyset -> sub x y))). ir. set (z:= Zo x (fun u => forall t, gle r u t -> ~ (inc t y))). ufi r' H2; ufi r' H3. ufi set_of_nreg_order H2; ufi set_of_nreg_order H3. awi H2. awi H3. rwi Exercise1_23a H2. rwi Exercise1_23a H3. nin H2; nin H3. assert (sub z (substrate r)). apply sub_trans with x. uf z; app Z_sub. nin H2. nin H2. am. assert (open_o r z). red. ee. am. ir. ufi z H7. Ztac. clear H7. uf z. Ztac. nin H2. nin H2. app (H11 _ _ H9 H8). ir. app H10. order_tac. cp (Exercise1_22i H H7). uf i. fold z. split. am. split. rw (Exercise1_22p H H2). app Exercise1_22j. nin H2; am. uf z. app Z_sub. split. uf disjoint. rw (Exercise1_22p H H3). app Exercise1_22k. nin H3; am. app is_emptyset. red; ir. nin (intersection2_both H9). ufi z H10. Ztac. assert (gle r y0 y0). order_tac. nin H3. nin H3. app H3. elim (H13 _ H14). am. ir. red. ir. rw (Exercise1_22p H H3). app Exercise1_22h. nin H3;am. nin H2; nin H2; app H2. ir. assert (inc y0 x). nin H2. nin H2. app (H13 _ _ H10 H11). app exists_proof. red. ir. elim (emptyset_pr (x:= y0)). wr H9. app Exercise1_22e. uf z. Ztac. ir. nin (inc_or_not t y). elim (H13 t). au. am. fprops. fprops. assert (forall x y, inc x (substrate r') -> inc y (substrate r') -> (inc (i x y) (substrate r') \/ sub x y)). ir. cp (H2 _ _ H3 H4). ee. rw Hv. rw Exercise1_23a. nin (emptyset_dichot (i x y)). right. app H8. au. fprops. rw Exercise1_23h. split. ir. nin H4. assert (inc x (substrate r')). order_tac. assert (inc y (substrate r')). order_tac. ufi r' H4. rwi Exercise1_23b H4. ee. nin (H3 _ _ H6 H7). cp (H2 _ _ H6 H7). ee. exists (i x y). split. split. uf r'. rww Exercise1_23b. eee. rwi Hv H12. rwi Exercise1_23a H12. ee; am. red. ir. nin H8. elim (emptyset_pr (x:=y0)). red in H15. wr H15. app intersection2_inc. wr H17. app H11. app H1. app disjoint_symmetric. elim H5. app extensionality. am. ir. nin (H3 _ _ H4 H5). cp (H2 _ _ H4 H5). right. right. left. exists (i x y). split. uf r'. rww Exercise1_23b. rwi Hv H4. rwi Exercise1_23a H4. rwi Hv H6. rwi Exercise1_23a H6. ee; tv. ee. app H1. right; left. uf r'. rww Exercise1_23b. rwi Hv H4. rwi Exercise1_23a H4. rwi Hv H5. rwi Exercise1_23a H5. ee; tv. am. Qed. Lemma Exercise1_23l: forall r y, order r -> let r' := set_of_nreg_order r in inc y (set_of_nreg_open r') -> exists_unique (fun x => inc x (set_of_nreg_open r) & y = canonical_reg_open r' x). Proof. ir. split. set (E:=set_of_nreg_open r). set (E':=set_of_nreg_open r'). assert (order r'). uf r'. uf set_of_nreg_order. fprops. assert (E = substrate r'). uf r'. uf set_of_nreg_order. aw. fprops. assert (forall x t, inc x E -> inc t E -> inc t (canonical_reg_open r' x) = (forall u, gle r' t u -> nonempty (intersection2 u x))). ir. rww Exercise1_23e. ap iff_eq; ir; nin (H5 _ H6). ufi r' H7. rwi Exercise1_23b H7. rwi Exercise1_23b H7. ee. nin H9. exists y0. app intersection2_inc. app H16. app H12. am. am. nin (intersection2_both H7). assert (inc z (substrate r')). order_tac. ufi r' H10. ufi set_of_nreg_order H10; awi H10. cp (Exercise1_23i H H8 H10). rwi H2 H3. ufi r' H3. ufi set_of_nreg_order H3; awi H3. cp (Exercise1_23i H H9 H3). exists (canonical_reg_open r y0). au. fprops. fprops. ue. ue. assert (forall x t, inc x E -> inc t E -> inc t (canonical_reg_open r' x) = (forall a, inc a t -> (exists b, inc b x & gle r a b))). ir. rww H3. ap iff_eq. ir. cp (Exercise1_23i H H7 H5). cp (H6 _ H8). ufi canonical_reg_open H9. ufi E H4. rwi Exercise1_23a H4. nin H4. rwi (Exercise1_22p H H4) H9. nin (emptyset_dichot (intersection2 (Zo (substrate r) (fun z => gle r a z)) x)). nin H4. cp (Exercise1_23c a H). rwi (Exercise1_22k H H13 H4) H9. nin H9. elim (emptyset_pr H9). am. nin H11. nin (intersection2_both H11). Ztac. exists y0. au. ir. ufi r' H7. rwi Exercise1_23b H7. ee. nin H8. nin (H6 _ (H11 _ H8)). nin H12. nin H10. nin H10. cp (H15 _ _ H8 H13). exists x0. app intersection2_inc. am. assert (forall x t, nonempty x -> open_o r x -> inc t E -> inc t (canonical_reg_open r' (bar1_22 r x)) = (forall a, inc a t -> (exists b, inc b x & gle r a b))). ir. assert (inc (bar1_22 r x) E). uf E. rw Exercise1_23a. ee. app Exercise1_22i. nin H5. exists y0. app Exercise1_22e. rw (H4 _ _ H8 H7). ap iff_eq. ir. nin (H9 _ H10). ee. ufi bar1_22 H11. nin (union_exists H11). nin H13. Ztac. nin H16. nin H17. rwi powerset_inc_rw H15. awii H18. nin (H18 _ H13). nin H19. cp (related_induced_order1 H20). exists x2. split. am. order_tac. ir. nin (H9 _ H10). nin H11. exists x0. split. app Exercise1_22e. am. assert (Hu:forall x, inc x y -> open_r r x). ir. rwi Exercise1_23a H0. nin H0. nin H0. nin H0. cp (H0 _ H6). wri H2 H10. ufi E H10. rwi Exercise1_23a H10. nin H10. am. assert (sub (union y) (substrate r)). red. ir. nin (union_exists H6). nin H7. nin (Hu _ H8). nin H9. app H9. assert (open_o r (union y)). app Exercise1_22c. ir. nin (Hu _ H7). am. set (T:=Zo E (fun z => forall a, inc a z -> exists b, inc b (union y) & gle r a b)). set (y':= canonical_reg_open r' (bar1_22 r (union y))). assert (nonempty (union y)). rwi Exercise1_23a H0. nin H0. nin H8. nin H0. nin H0. cp (H0 _ H8). wri H2 H11. ufi E H11. rwi Exercise1_23a H11. nin H11. nin H12. exists y1. union_tac. assert (Hv:inc (bar1_22 r (union y)) E). uf E. rw Exercise1_23a. split. app Exercise1_22i. nin H8. exists y0. app Exercise1_22e. assert (T = y'). cut (forall t, inc t E -> (inc t T = inc t y')). ir. set_extens. wr H9. am. ufi T H10. Ztac. am. rw H9. am. ufi y' H10. assert (inc (bar1_22 r (union y)) (substrate r')). wr H2. am. cp (Exercise1_23d2 H1 H11). set (k := canonical_reg_open r' (bar1_22 r (union y))) in *. rwi Exercise1_23a H12. nin H12. nin H12. nin H12. rw H2. app H12. ir. uf y'. rw (H5 _ _ H8 H7 H9). ap iff_eq. ir. ufi T H10. Ztac. ap (H13 _ H11). ir. uf T. Ztac. assert (sub y y'). wr H9. uf T. red. ir. Ztac. rwi Exercise1_23a H0. nin H0. nin H0. nin H0. rw H2. ap (H0 _ H10). ir. exists a. assert (inc a (union y)). union_tac. split. am. order_tac. app H6. assert (open_o r' y'). uf y'. rwi H2 Hv. cp (Exercise1_23d2 H1 Hv). rwi Exercise1_23a H11. nin H11. nin 11. nin H11. am. am. assert (substrate (induced_order r' y') = y'). aw. wrr H2. wr H9. uf T. app Z_sub. assert (cofinal_set (induced_order r' y') y). split. rww H12. rw H12. ir. wri H9 H13. ufi T H13. Ztac. clear H13. (* set (T:= intersection(Zo (powerset (substrate r)) (fun x => open_o r x & forall a, inc a (union y) -> exists b, inc b x & gle r a b))). assert (open_o r T). uf T. app Exercise1_22b. exists (union y). Ztac. app powerset_inc. split. app Exercise1_22c. ir. nin (Hu _ H7). am. ir. exists a. split. am. order_tac. app H6. ir. Ztac. nin H9; am. *) (* assert (forall x y0 : Set, f x -> f y0 -> x = y0). uf f. ir. cp (Exercise1_23j H). red in H3. ee. fold r' in H3. assert (inc x (substrate r')). uf r'. uf set_of_nreg_order. aw. fprops. assert (inc y0 (substrate r')). uf r'. uf set_of_nreg_order. aw. fprops. cp (H3 _ _ H6 H7). app H8. wrr H5; wrr H4. red. eee. *) Abort. End Exercice2. Export Exercice2.