src/ZF/Constructible/Separation.thy
changeset 13363 c26eeb000470
parent 13352 3cd767f8d78b
child 13385 31df66ca0780
equal deleted inserted replaced
13362:cd7f9ea58338 13363:c26eeb000470
   454 apply (rule_tac env = "[xa,u,r,f,g,a,b]" in pair_iff_sats) 
   454 apply (rule_tac env = "[xa,u,r,f,g,a,b]" in pair_iff_sats) 
   455 apply (rule sep_rules | simp)+
   455 apply (rule sep_rules | simp)+
   456 apply (simp_all add: succ_Un_distrib [symmetric])
   456 apply (simp_all add: succ_Un_distrib [symmetric])
   457 done
   457 done
   458 
   458 
       
   459 
       
   460 subsection{*Instantiating the locale @{text M_axioms}*}
       
   461 text{*Separation (and Strong Replacement) for basic set-theoretic constructions
       
   462 such as intersection, Cartesian Product and image.*}
   459 
   463 
   460 ML
   464 ML
   461 {*
   465 {*
   462 val Inter_separation = thm "Inter_separation";
   466 val Inter_separation = thm "Inter_separation";
   463 val cartprod_separation = thm "cartprod_separation";
   467 val cartprod_separation = thm "cartprod_separation";
   479      converse_separation, restrict_separation, comp_separation, 
   483      converse_separation, restrict_separation, comp_separation, 
   480      pred_separation, Memrel_separation, funspace_succ_replacement, 
   484      pred_separation, Memrel_separation, funspace_succ_replacement, 
   481      well_ord_iso_separation, obase_separation,
   485      well_ord_iso_separation, obase_separation,
   482      obase_equals_separation, omap_replacement, is_recfun_separation]
   486      obase_equals_separation, omap_replacement, is_recfun_separation]
   483 
   487 
   484 fun axiomsL th =
   488 fun axioms_L th =
   485     kill_flex_triv_prems (m_axioms MRS (trivaxL th));
   489     kill_flex_triv_prems (m_axioms MRS (triv_axioms_L th));
   486 
   490 
   487 bind_thm ("cartprod_iff", axiomsL (thm "M_axioms.cartprod_iff"));
   491 bind_thm ("cartprod_iff", axioms_L (thm "M_axioms.cartprod_iff"));
   488 bind_thm ("cartprod_closed", axiomsL (thm "M_axioms.cartprod_closed"));
   492 bind_thm ("cartprod_closed", axioms_L (thm "M_axioms.cartprod_closed"));
   489 bind_thm ("sum_closed", axiomsL (thm "M_axioms.sum_closed"));
   493 bind_thm ("sum_closed", axioms_L (thm "M_axioms.sum_closed"));
   490 bind_thm ("M_converse_iff", axiomsL (thm "M_axioms.M_converse_iff"));
   494 bind_thm ("M_converse_iff", axioms_L (thm "M_axioms.M_converse_iff"));
   491 bind_thm ("converse_closed", axiomsL (thm "M_axioms.converse_closed"));
   495 bind_thm ("converse_closed", axioms_L (thm "M_axioms.converse_closed"));
   492 bind_thm ("converse_abs", axiomsL (thm "M_axioms.converse_abs"));
   496 bind_thm ("converse_abs", axioms_L (thm "M_axioms.converse_abs"));
   493 bind_thm ("image_closed", axiomsL (thm "M_axioms.image_closed"));
   497 bind_thm ("image_closed", axioms_L (thm "M_axioms.image_closed"));
   494 bind_thm ("vimage_abs", axiomsL (thm "M_axioms.vimage_abs"));
   498 bind_thm ("vimage_abs", axioms_L (thm "M_axioms.vimage_abs"));
   495 bind_thm ("vimage_closed", axiomsL (thm "M_axioms.vimage_closed"));
   499 bind_thm ("vimage_closed", axioms_L (thm "M_axioms.vimage_closed"));
   496 bind_thm ("domain_abs", axiomsL (thm "M_axioms.domain_abs"));
   500 bind_thm ("domain_abs", axioms_L (thm "M_axioms.domain_abs"));
   497 bind_thm ("domain_closed", axiomsL (thm "M_axioms.domain_closed"));
   501 bind_thm ("domain_closed", axioms_L (thm "M_axioms.domain_closed"));
   498 bind_thm ("range_abs", axiomsL (thm "M_axioms.range_abs"));
   502 bind_thm ("range_abs", axioms_L (thm "M_axioms.range_abs"));
   499 bind_thm ("range_closed", axiomsL (thm "M_axioms.range_closed"));
   503 bind_thm ("range_closed", axioms_L (thm "M_axioms.range_closed"));
   500 bind_thm ("field_abs", axiomsL (thm "M_axioms.field_abs"));
   504 bind_thm ("field_abs", axioms_L (thm "M_axioms.field_abs"));
   501 bind_thm ("field_closed", axiomsL (thm "M_axioms.field_closed"));
   505 bind_thm ("field_closed", axioms_L (thm "M_axioms.field_closed"));
   502 bind_thm ("relation_abs", axiomsL (thm "M_axioms.relation_abs"));
   506 bind_thm ("relation_abs", axioms_L (thm "M_axioms.relation_abs"));
   503 bind_thm ("function_abs", axiomsL (thm "M_axioms.function_abs"));
   507 bind_thm ("function_abs", axioms_L (thm "M_axioms.function_abs"));
   504 bind_thm ("apply_closed", axiomsL (thm "M_axioms.apply_closed"));
   508 bind_thm ("apply_closed", axioms_L (thm "M_axioms.apply_closed"));
   505 bind_thm ("apply_abs", axiomsL (thm "M_axioms.apply_abs"));
   509 bind_thm ("apply_abs", axioms_L (thm "M_axioms.apply_abs"));
   506 bind_thm ("typed_function_abs", axiomsL (thm "M_axioms.typed_function_abs"));
   510 bind_thm ("typed_function_abs", axioms_L (thm "M_axioms.typed_function_abs"));
   507 bind_thm ("injection_abs", axiomsL (thm "M_axioms.injection_abs"));
   511 bind_thm ("injection_abs", axioms_L (thm "M_axioms.injection_abs"));
   508 bind_thm ("surjection_abs", axiomsL (thm "M_axioms.surjection_abs"));
   512 bind_thm ("surjection_abs", axioms_L (thm "M_axioms.surjection_abs"));
   509 bind_thm ("bijection_abs", axiomsL (thm "M_axioms.bijection_abs"));
   513 bind_thm ("bijection_abs", axioms_L (thm "M_axioms.bijection_abs"));
   510 bind_thm ("M_comp_iff", axiomsL (thm "M_axioms.M_comp_iff"));
   514 bind_thm ("M_comp_iff", axioms_L (thm "M_axioms.M_comp_iff"));
   511 bind_thm ("comp_closed", axiomsL (thm "M_axioms.comp_closed"));
   515 bind_thm ("comp_closed", axioms_L (thm "M_axioms.comp_closed"));
   512 bind_thm ("composition_abs", axiomsL (thm "M_axioms.composition_abs"));
   516 bind_thm ("composition_abs", axioms_L (thm "M_axioms.composition_abs"));
   513 bind_thm ("restriction_is_function", axiomsL (thm "M_axioms.restriction_is_function"));
   517 bind_thm ("restriction_is_function", axioms_L (thm "M_axioms.restriction_is_function"));
   514 bind_thm ("restriction_abs", axiomsL (thm "M_axioms.restriction_abs"));
   518 bind_thm ("restriction_abs", axioms_L (thm "M_axioms.restriction_abs"));
   515 bind_thm ("M_restrict_iff", axiomsL (thm "M_axioms.M_restrict_iff"));
   519 bind_thm ("M_restrict_iff", axioms_L (thm "M_axioms.M_restrict_iff"));
   516 bind_thm ("restrict_closed", axiomsL (thm "M_axioms.restrict_closed"));
   520 bind_thm ("restrict_closed", axioms_L (thm "M_axioms.restrict_closed"));
   517 bind_thm ("Inter_abs", axiomsL (thm "M_axioms.Inter_abs"));
   521 bind_thm ("Inter_abs", axioms_L (thm "M_axioms.Inter_abs"));
   518 bind_thm ("Inter_closed", axiomsL (thm "M_axioms.Inter_closed"));
   522 bind_thm ("Inter_closed", axioms_L (thm "M_axioms.Inter_closed"));
   519 bind_thm ("Int_closed", axiomsL (thm "M_axioms.Int_closed"));
   523 bind_thm ("Int_closed", axioms_L (thm "M_axioms.Int_closed"));
   520 bind_thm ("finite_fun_closed", axiomsL (thm "M_axioms.finite_fun_closed"));
   524 bind_thm ("finite_fun_closed", axioms_L (thm "M_axioms.finite_fun_closed"));
   521 bind_thm ("is_funspace_abs", axiomsL (thm "M_axioms.is_funspace_abs"));
   525 bind_thm ("is_funspace_abs", axioms_L (thm "M_axioms.is_funspace_abs"));
   522 bind_thm ("succ_fun_eq2", axiomsL (thm "M_axioms.succ_fun_eq2"));
   526 bind_thm ("succ_fun_eq2", axioms_L (thm "M_axioms.succ_fun_eq2"));
   523 bind_thm ("funspace_succ", axiomsL (thm "M_axioms.funspace_succ"));
   527 bind_thm ("funspace_succ", axioms_L (thm "M_axioms.funspace_succ"));
   524 bind_thm ("finite_funspace_closed", axiomsL (thm "M_axioms.finite_funspace_closed"));
   528 bind_thm ("finite_funspace_closed", axioms_L (thm "M_axioms.finite_funspace_closed"));
   525 *}
   529 *}
   526 
   530 
   527 ML
   531 ML
   528 {*
   532 {*
   529 bind_thm ("is_recfun_equal", axiomsL (thm "M_axioms.is_recfun_equal"));  
   533 bind_thm ("is_recfun_equal", axioms_L (thm "M_axioms.is_recfun_equal"));  
   530 bind_thm ("is_recfun_cut", axiomsL (thm "M_axioms.is_recfun_cut")); 
   534 bind_thm ("is_recfun_cut", axioms_L (thm "M_axioms.is_recfun_cut")); 
   531 bind_thm ("is_recfun_functional", axiomsL (thm "M_axioms.is_recfun_functional"));
   535 bind_thm ("is_recfun_functional", axioms_L (thm "M_axioms.is_recfun_functional"));
   532 bind_thm ("is_recfun_relativize", axiomsL (thm "M_axioms.is_recfun_relativize"));
   536 bind_thm ("is_recfun_relativize", axioms_L (thm "M_axioms.is_recfun_relativize"));
   533 bind_thm ("is_recfun_restrict", axiomsL (thm "M_axioms.is_recfun_restrict"));
   537 bind_thm ("is_recfun_restrict", axioms_L (thm "M_axioms.is_recfun_restrict"));
   534 bind_thm ("univalent_is_recfun", axiomsL (thm "M_axioms.univalent_is_recfun"));
   538 bind_thm ("univalent_is_recfun", axioms_L (thm "M_axioms.univalent_is_recfun"));
   535 bind_thm ("exists_is_recfun_indstep", axiomsL (thm "M_axioms.exists_is_recfun_indstep"));
   539 bind_thm ("exists_is_recfun_indstep", axioms_L (thm "M_axioms.exists_is_recfun_indstep"));
   536 bind_thm ("wellfounded_exists_is_recfun", axiomsL (thm "M_axioms.wellfounded_exists_is_recfun"));
   540 bind_thm ("wellfounded_exists_is_recfun", axioms_L (thm "M_axioms.wellfounded_exists_is_recfun"));
   537 bind_thm ("wf_exists_is_recfun", axiomsL (thm "M_axioms.wf_exists_is_recfun")); 
   541 bind_thm ("wf_exists_is_recfun", axioms_L (thm "M_axioms.wf_exists_is_recfun")); 
   538 bind_thm ("is_recfun_abs", axiomsL (thm "M_axioms.is_recfun_abs"));
   542 bind_thm ("is_recfun_abs", axioms_L (thm "M_axioms.is_recfun_abs"));
   539 bind_thm ("irreflexive_abs", axiomsL (thm "M_axioms.irreflexive_abs"));  
   543 bind_thm ("irreflexive_abs", axioms_L (thm "M_axioms.irreflexive_abs"));  
   540 bind_thm ("transitive_rel_abs", axiomsL (thm "M_axioms.transitive_rel_abs"));  
   544 bind_thm ("transitive_rel_abs", axioms_L (thm "M_axioms.transitive_rel_abs"));  
   541 bind_thm ("linear_rel_abs", axiomsL (thm "M_axioms.linear_rel_abs"));  
   545 bind_thm ("linear_rel_abs", axioms_L (thm "M_axioms.linear_rel_abs"));  
   542 bind_thm ("wellordered_is_trans_on", axiomsL (thm "M_axioms.wellordered_is_trans_on")); 
   546 bind_thm ("wellordered_is_trans_on", axioms_L (thm "M_axioms.wellordered_is_trans_on")); 
   543 bind_thm ("wellordered_is_linear", axiomsL (thm "M_axioms.wellordered_is_linear")); 
   547 bind_thm ("wellordered_is_linear", axioms_L (thm "M_axioms.wellordered_is_linear")); 
   544 bind_thm ("wellordered_is_wellfounded_on", axiomsL (thm "M_axioms.wellordered_is_wellfounded_on")); 
   548 bind_thm ("wellordered_is_wellfounded_on", axioms_L (thm "M_axioms.wellordered_is_wellfounded_on")); 
   545 bind_thm ("wellfounded_imp_wellfounded_on", axiomsL (thm "M_axioms.wellfounded_imp_wellfounded_on")); 
   549 bind_thm ("wellfounded_imp_wellfounded_on", axioms_L (thm "M_axioms.wellfounded_imp_wellfounded_on")); 
   546 bind_thm ("wellfounded_on_subset_A", axiomsL (thm "M_axioms.wellfounded_on_subset_A"));
   550 bind_thm ("wellfounded_on_subset_A", axioms_L (thm "M_axioms.wellfounded_on_subset_A"));
   547 bind_thm ("wellfounded_on_iff_wellfounded", axiomsL (thm "M_axioms.wellfounded_on_iff_wellfounded"));
   551 bind_thm ("wellfounded_on_iff_wellfounded", axioms_L (thm "M_axioms.wellfounded_on_iff_wellfounded"));
   548 bind_thm ("wellfounded_on_imp_wellfounded", axiomsL (thm "M_axioms.wellfounded_on_imp_wellfounded"));
   552 bind_thm ("wellfounded_on_imp_wellfounded", axioms_L (thm "M_axioms.wellfounded_on_imp_wellfounded"));
   549 bind_thm ("wellfounded_on_field_imp_wellfounded", axiomsL (thm "M_axioms.wellfounded_on_field_imp_wellfounded"));
   553 bind_thm ("wellfounded_on_field_imp_wellfounded", axioms_L (thm "M_axioms.wellfounded_on_field_imp_wellfounded"));
   550 bind_thm ("wellfounded_iff_wellfounded_on_field", axiomsL (thm "M_axioms.wellfounded_iff_wellfounded_on_field"));
   554 bind_thm ("wellfounded_iff_wellfounded_on_field", axioms_L (thm "M_axioms.wellfounded_iff_wellfounded_on_field"));
   551 bind_thm ("wellfounded_induct", axiomsL (thm "M_axioms.wellfounded_induct")); 
   555 bind_thm ("wellfounded_induct", axioms_L (thm "M_axioms.wellfounded_induct")); 
   552 bind_thm ("wellfounded_on_induct", axiomsL (thm "M_axioms.wellfounded_on_induct")); 
   556 bind_thm ("wellfounded_on_induct", axioms_L (thm "M_axioms.wellfounded_on_induct")); 
   553 bind_thm ("wellfounded_on_induct2", axiomsL (thm "M_axioms.wellfounded_on_induct2")); 
   557 bind_thm ("wellfounded_on_induct2", axioms_L (thm "M_axioms.wellfounded_on_induct2")); 
   554 bind_thm ("linear_imp_relativized", axiomsL (thm "M_axioms.linear_imp_relativized")); 
   558 bind_thm ("linear_imp_relativized", axioms_L (thm "M_axioms.linear_imp_relativized")); 
   555 bind_thm ("trans_on_imp_relativized", axiomsL (thm "M_axioms.trans_on_imp_relativized")); 
   559 bind_thm ("trans_on_imp_relativized", axioms_L (thm "M_axioms.trans_on_imp_relativized")); 
   556 bind_thm ("wf_on_imp_relativized", axiomsL (thm "M_axioms.wf_on_imp_relativized")); 
   560 bind_thm ("wf_on_imp_relativized", axioms_L (thm "M_axioms.wf_on_imp_relativized")); 
   557 bind_thm ("wf_imp_relativized", axiomsL (thm "M_axioms.wf_imp_relativized")); 
   561 bind_thm ("wf_imp_relativized", axioms_L (thm "M_axioms.wf_imp_relativized")); 
   558 bind_thm ("well_ord_imp_relativized", axiomsL (thm "M_axioms.well_ord_imp_relativized")); 
   562 bind_thm ("well_ord_imp_relativized", axioms_L (thm "M_axioms.well_ord_imp_relativized")); 
   559 bind_thm ("order_isomorphism_abs", axiomsL (thm "M_axioms.order_isomorphism_abs"));  
   563 bind_thm ("order_isomorphism_abs", axioms_L (thm "M_axioms.order_isomorphism_abs"));  
   560 bind_thm ("pred_set_abs", axiomsL (thm "M_axioms.pred_set_abs"));  
   564 bind_thm ("pred_set_abs", axioms_L (thm "M_axioms.pred_set_abs"));  
   561 *}
   565 *}
   562 
   566 
   563 ML
   567 ML
   564 {*
   568 {*
   565 bind_thm ("pred_closed", axiomsL (thm "M_axioms.pred_closed"));  
   569 bind_thm ("pred_closed", axioms_L (thm "M_axioms.pred_closed"));  
   566 bind_thm ("membership_abs", axiomsL (thm "M_axioms.membership_abs"));  
   570 bind_thm ("membership_abs", axioms_L (thm "M_axioms.membership_abs"));  
   567 bind_thm ("M_Memrel_iff", axiomsL (thm "M_axioms.M_Memrel_iff"));
   571 bind_thm ("M_Memrel_iff", axioms_L (thm "M_axioms.M_Memrel_iff"));
   568 bind_thm ("Memrel_closed", axiomsL (thm "M_axioms.Memrel_closed"));  
   572 bind_thm ("Memrel_closed", axioms_L (thm "M_axioms.Memrel_closed"));  
   569 bind_thm ("wellordered_iso_predD", axiomsL (thm "M_axioms.wellordered_iso_predD"));
   573 bind_thm ("wellordered_iso_predD", axioms_L (thm "M_axioms.wellordered_iso_predD"));
   570 bind_thm ("wellordered_iso_pred_eq", axiomsL (thm "M_axioms.wellordered_iso_pred_eq"));
   574 bind_thm ("wellordered_iso_pred_eq", axioms_L (thm "M_axioms.wellordered_iso_pred_eq"));
   571 bind_thm ("wellfounded_on_asym", axiomsL (thm "M_axioms.wellfounded_on_asym"));
   575 bind_thm ("wellfounded_on_asym", axioms_L (thm "M_axioms.wellfounded_on_asym"));
   572 bind_thm ("wellordered_asym", axiomsL (thm "M_axioms.wellordered_asym"));
   576 bind_thm ("wellordered_asym", axioms_L (thm "M_axioms.wellordered_asym"));
   573 bind_thm ("ord_iso_pred_imp_lt", axiomsL (thm "M_axioms.ord_iso_pred_imp_lt"));
   577 bind_thm ("ord_iso_pred_imp_lt", axioms_L (thm "M_axioms.ord_iso_pred_imp_lt"));
   574 bind_thm ("obase_iff", axiomsL (thm "M_axioms.obase_iff"));
   578 bind_thm ("obase_iff", axioms_L (thm "M_axioms.obase_iff"));
   575 bind_thm ("omap_iff", axiomsL (thm "M_axioms.omap_iff"));
   579 bind_thm ("omap_iff", axioms_L (thm "M_axioms.omap_iff"));
   576 bind_thm ("omap_unique", axiomsL (thm "M_axioms.omap_unique"));
   580 bind_thm ("omap_unique", axioms_L (thm "M_axioms.omap_unique"));
   577 bind_thm ("omap_yields_Ord", axiomsL (thm "M_axioms.omap_yields_Ord"));
   581 bind_thm ("omap_yields_Ord", axioms_L (thm "M_axioms.omap_yields_Ord"));
   578 bind_thm ("otype_iff", axiomsL (thm "M_axioms.otype_iff"));
   582 bind_thm ("otype_iff", axioms_L (thm "M_axioms.otype_iff"));
   579 bind_thm ("otype_eq_range", axiomsL (thm "M_axioms.otype_eq_range"));
   583 bind_thm ("otype_eq_range", axioms_L (thm "M_axioms.otype_eq_range"));
   580 bind_thm ("Ord_otype", axiomsL (thm "M_axioms.Ord_otype"));
   584 bind_thm ("Ord_otype", axioms_L (thm "M_axioms.Ord_otype"));
   581 bind_thm ("domain_omap", axiomsL (thm "M_axioms.domain_omap"));
   585 bind_thm ("domain_omap", axioms_L (thm "M_axioms.domain_omap"));
   582 bind_thm ("omap_subset", axiomsL (thm "M_axioms.omap_subset")); 
   586 bind_thm ("omap_subset", axioms_L (thm "M_axioms.omap_subset")); 
   583 bind_thm ("omap_funtype", axiomsL (thm "M_axioms.omap_funtype")); 
   587 bind_thm ("omap_funtype", axioms_L (thm "M_axioms.omap_funtype")); 
   584 bind_thm ("wellordered_omap_bij", axiomsL (thm "M_axioms.wellordered_omap_bij"));
   588 bind_thm ("wellordered_omap_bij", axioms_L (thm "M_axioms.wellordered_omap_bij"));
   585 bind_thm ("omap_ord_iso", axiomsL (thm "M_axioms.omap_ord_iso"));
   589 bind_thm ("omap_ord_iso", axioms_L (thm "M_axioms.omap_ord_iso"));
   586 bind_thm ("Ord_omap_image_pred", axiomsL (thm "M_axioms.Ord_omap_image_pred"));
   590 bind_thm ("Ord_omap_image_pred", axioms_L (thm "M_axioms.Ord_omap_image_pred"));
   587 bind_thm ("restrict_omap_ord_iso", axiomsL (thm "M_axioms.restrict_omap_ord_iso"));
   591 bind_thm ("restrict_omap_ord_iso", axioms_L (thm "M_axioms.restrict_omap_ord_iso"));
   588 bind_thm ("obase_equals", axiomsL (thm "M_axioms.obase_equals")); 
   592 bind_thm ("obase_equals", axioms_L (thm "M_axioms.obase_equals")); 
   589 bind_thm ("omap_ord_iso_otype", axiomsL (thm "M_axioms.omap_ord_iso_otype"));
   593 bind_thm ("omap_ord_iso_otype", axioms_L (thm "M_axioms.omap_ord_iso_otype"));
   590 bind_thm ("obase_exists", axiomsL (thm "M_axioms.obase_exists"));
   594 bind_thm ("obase_exists", axioms_L (thm "M_axioms.obase_exists"));
   591 bind_thm ("omap_exists", axiomsL (thm "M_axioms.omap_exists"));
   595 bind_thm ("omap_exists", axioms_L (thm "M_axioms.omap_exists"));
   592 bind_thm ("otype_exists", axiomsL (thm "M_axioms.otype_exists"));
   596 bind_thm ("otype_exists", axioms_L (thm "M_axioms.otype_exists"));
   593 bind_thm ("omap_ord_iso_otype", axiomsL (thm "M_axioms.omap_ord_iso_otype"));
   597 bind_thm ("omap_ord_iso_otype", axioms_L (thm "M_axioms.omap_ord_iso_otype"));
   594 bind_thm ("ordertype_exists", axiomsL (thm "M_axioms.ordertype_exists"));
   598 bind_thm ("ordertype_exists", axioms_L (thm "M_axioms.ordertype_exists"));
   595 bind_thm ("relativized_imp_well_ord", axiomsL (thm "M_axioms.relativized_imp_well_ord")); 
   599 bind_thm ("relativized_imp_well_ord", axioms_L (thm "M_axioms.relativized_imp_well_ord")); 
   596 bind_thm ("well_ord_abs", axiomsL (thm "M_axioms.well_ord_abs"));  
   600 bind_thm ("well_ord_abs", axioms_L (thm "M_axioms.well_ord_abs"));  
   597 *}
   601 *}
   598 
   602 
   599 declare cartprod_closed [intro,simp]
   603 declare cartprod_closed [intro,simp]
   600 declare sum_closed [intro,simp]
   604 declare sum_closed [intro,simp]
   601 declare converse_closed [intro,simp]
   605 declare converse_closed [intro,simp]