diff --git a/etc/options b/etc/options --- a/etc/options +++ b/etc/options @@ -1,404 +1,413 @@ (* :mode=isabelle-options: *) section "Document Preparation" option browser_info : bool = false for build -- "generate theory browser information" option document : string = "" (standard "true") for build -- "build PDF document (enabled for \"pdf\" or \"true\", disabled for \"\" or \"false\")" option document_output : string = "" (standard "output") for build -- "document output directory" option document_echo : bool = false for build -- "inform about document file names during session presentation" option document_variants : string = "document" for build document -- "alternative document variants (separated by colons)" option document_tags : string = "" for document -- "default command tags (separated by commas)" option document_bibliography : bool = false for document -- "explicitly enable use of bibtex (default: according to presence of root.bib)" option document_build : string = "lualatex" (standard "build") for document -- "document build engine (e.g. build, lualatex, pdflatex)" option document_logo : string = "" (standard "_") for document -- "generate named instance of Isabelle logo (underscore means unnamed variant)" option document_heading_prefix : string = "isamarkup" (standard) for document -- "prefix for LaTeX macros generated from 'chapter', 'section' etc." option document_comment_latex : bool = false for document -- "use regular LaTeX version of comment.sty, instead of historic plain TeX version" option document_cite_commands : string = "cite,nocite,citet,citep" for document -- "antiquotation commands to determine the structure of bibliography" option thy_output_display : bool = false for content -- "indicate output as multi-line display-style material" option thy_output_break : bool = false for content -- "control line breaks in non-display material" option thy_output_cartouche : bool = false for content -- "indicate if the output should be delimited as cartouche" option thy_output_quotes : bool = false for content -- "indicate if the output should be delimited via double quotes" option thy_output_margin : int = 76 for content -- "right margin / page width for printing of display material" option thy_output_indent : int = 0 for content -- "indentation for pretty printing of display material" option thy_output_source : bool = false for content -- "print original source text rather than internal representation" option thy_output_source_cartouche : bool = false for content -- "print original source text rather than internal representation, preserve cartouches" option thy_output_modes : string = "" for content -- "additional print modes for document output (separated by commas)" section "Prover Output" option pide_reports : bool = true for content -- "enable reports of PIDE markup" option show_types : bool = false for content -- "show type constraints when printing terms" option show_sorts : bool = false for content -- "show sort constraints when printing types" option show_brackets : bool = false for content -- "show extra brackets when printing terms/types" option show_question_marks : bool = true for content -- "show leading question mark of schematic variables" option show_consts : bool = false for content -- "show constants with types when printing proof state" option show_main_goal : bool = false for content -- "show main goal when printing proof state" option goals_limit : int = 10 for content -- "maximum number of subgoals to be printed" option show_states : bool = false for content -- "show toplevel states even if outside of interactive mode" option names_long : bool = false for content -- "show fully qualified names" option names_short : bool = false for content -- "show base names only" option names_unique : bool = true for content -- "show partially qualified names, as required for unique name resolution" option eta_contract : bool = true for content -- "print terms in eta-contracted form" option print_mode : string = "" for content -- "additional print modes for prover output (separated by commas)" section "Parallel Processing and Timing" public option threads : int = 0 for build -- "maximum number of worker threads for prover process (0 = hardware max.)" option threads_trace : int = 0 -- "level of tracing information for multithreading" option threads_stack_limit : real = 0.25 -- "maximum stack size for worker threads (in giga words, 0 = unlimited)" public option parallel_limit : int = 0 for build -- "approximative limit for parallel tasks (0 = unlimited)" public option parallel_print : bool = true for build -- "parallel and asynchronous printing of results" public option parallel_proofs : int = 1 for build -- "level of parallel proof checking: 0, 1, 2" option parallel_subproofs_threshold : real = 0.01 for build -- "lower bound of timing estimate for forked nested proofs (seconds)" option command_timing_threshold : real = 0.1 for build -- "default threshold for persistent command timing (seconds)" public option timeout_scale : real = 1.0 for build -- "scale factor for timeout in Isabelle/ML and session build" +option context_data_timing : bool = false for build + -- "timing for context data operations" + +option context_theory_tracing : bool = false for build + -- "tracing of persistent theory values within ML heap" + +option context_proof_tracing : bool = false for build + -- "tracing of persistent Proof.context values within ML heap" + section "Detail of Proof Checking" option record_proofs : int = -1 for content -- "set level of proofterm recording: 0, 1, 2, negative means unchanged" option quick_and_dirty : bool = false for content -- "if true then some tools will OMIT some proofs" option skip_proofs : bool = false for content -- "skip over proofs (implicit 'sorry')" option strict_facts : bool = false for content -- "force lazy facts when defined in context" section "Global Session Parameters" option condition : string = "" for content -- "required environment variables for subsequent theories (separated by commas)" option timeout : real = 0 for build -- "timeout for session build job (seconds > 0)" option timeout_build : bool = true for build -- "observe timeout for session build" option process_policy : string = "" -- "command prefix for underlying process, notably ML with NUMA policy" option process_output_tail : int = 40 for build -- "build process output tail shown to user (in lines, 0 = unlimited)" option profiling : string = "" (standard "time") for build -- "ML profiling (possible values: time, time_thread, allocations)" option system_log : string = "" (standard "-") for build -- "output for system messages (log file or \"-\" for console progress)" option system_heaps : bool = false for build -- "store session heaps in $ISABELLE_HEAPS_SYSTEM, not $ISABELLE_HEAPS" section "ML System" option ML_print_depth : int = 20 for content -- "ML print depth for toplevel pretty-printing" public option ML_exception_trace : bool = false -- "ML exception trace for toplevel command execution" public option ML_exception_debugger : bool = false -- "ML debugger exception trace for toplevel command execution" public option ML_debugger : bool = false -- "ML debugger instrumentation for newly compiled code" public option ML_system_64 : bool = false for build -- "prefer native 64bit platform (change requires restart)" public option ML_system_apple : bool = true for build -- "prefer native Apple/ARM64 platform (change requires restart)" section "Build Process" option build_thorough : bool = false -- "observe dependencies on options with tag 'content' or 'document'" option build_hostname : string = "" -- "alternative hostname for build process (default $ISABELLE_HOSTNAME)" option build_engine : string = "" -- "alternative session build engine" option build_database_test : bool = false -- "expose state of build process via central database" section "Editor Session" public option editor_load_delay : real = 0.5 -- "delay for file load operations (new buffers etc.)" public option editor_input_delay : real = 0.2 -- "delay for user input (text edits, cursor movement etc.)" public option editor_generated_input_delay : real = 1.0 -- "delay for machine-generated input that may outperform user edits" public option editor_output_delay : real = 0.1 -- "delay for prover output (markup, common messages etc.)" public option editor_consolidate_delay : real = 1.0 -- "delay to consolidate status of command evaluation (execution forks)" public option editor_prune_delay : real = 15 -- "delay to prune history (delete old versions)" option editor_prune_size : int = 0 -- "retained size of pruned history (delete old versions)" public option editor_update_delay : real = 0.5 -- "delay for physical GUI updates" public option editor_reparse_limit : int = 10000 -- "maximum amount of reparsed text outside perspective" public option editor_tracing_messages : int = 1000 -- "initial number of tracing messages for each command transaction (0: unbounded)" public option editor_chart_delay : real = 3.0 -- "delay for chart repainting" public option editor_continuous_checking : bool = true -- "continuous checking of proof document (visible and required parts)" public option editor_output_state : bool = false -- "implicit output of proof state" public option editor_document_session : string = "" -- "session for interactive document preparation" public option editor_document_auto : bool = false -- "automatically build document when selected theories are finished" public option editor_document_delay : real = 2.0 -- "delay for document auto build" option editor_execution_delay : real = 0.02 -- "delay for start of execution process after document update (seconds)" option editor_syslog_limit : int = 100 -- "maximum amount of buffered syslog messages" section "Headless Session" option headless_consolidate_delay : real = 15 -- "delay to consolidate status of command evaluation (execution forks)" option headless_prune_delay : real = 60 -- "delay to prune history (delete old versions)" option headless_check_delay : real = 0.5 -- "delay for theory status check during PIDE processing (seconds)" option headless_check_limit : int = 0 -- "maximum number of theory status checks (0 = unlimited)" option headless_nodes_status_delay : real = -1 -- "delay for overall nodes status check during PIDE processing (seconds, disabled for < 0)" option headless_watchdog_timeout : real = 600 -- "watchdog timeout for PIDE processing of broken theories (seconds, 0 = disabled)" option headless_commit_cleanup_delay : real = 60 -- "delay for cleanup of already imported theories (seconds, 0 = disabled)" option headless_load_limit : real = 5.0 -- "limit in MiB for loaded theory files (0 = unlimited)" section "Miscellaneous Tools" public option find_theorems_limit : int = 40 -- "limit of displayed results" public option find_theorems_tactic_limit : int = 5 -- "limit of tactic search for 'solves' criterion" section "Completion" public option completion_limit : int = 40 -- "limit for completion within the formal context" public option completion_path_ignore : string = "*~:*.marks:*.orig:*.rej:.DS_Store" -- "glob patterns to ignore in file-system path completion (separated by colons)" section "Spell Checker" public option spell_checker : bool = true -- "enable spell-checker for prose words within document text, comments etc." public option spell_checker_dictionary : string = "en" -- "spell-checker dictionary name" public option spell_checker_include : string = "words,comment,comment1,comment2,comment3,ML_comment,SML_comment" -- "included markup elements for spell-checker (separated by commas)" public option spell_checker_exclude : string = "document_marker,antiquoted,raw_text" -- "excluded markup elements for spell-checker (separated by commas)" section "Secure Shell (OpenSSH)" public option ssh_batch_mode : bool = true -- "enable SSH batch mode (no user interaction)" public option ssh_multiplexing : bool = true -- "enable multiplexing of SSH sessions (ignored on Windows)" public option ssh_compression : bool = true -- "enable SSH compression" public option ssh_alive_interval : real = 30 -- "time interval to keep SSH server connection alive (seconds, ignore value < 0)" public option ssh_alive_count_max : int = 3 -- "maximum number of messages to keep SSH server connection alive (ignore value < 0)" section "Phabricator" option phabricator_version_arcanist : string = "4f70fcffa8a5393e210d64f237ffdaa256256d6a" -- "repository version for arcanist" option phabricator_version_phabricator : string = "193798385bd3a7f72dca255e44f8112f4f8fc155" -- "repository version for phabricator" section "Theory Export" option export_theory : bool = false for content -- "export theory content to Isabelle/Scala" option export_standard_proofs : bool = false for content -- "export standardized proof terms to Isabelle/Scala (not scalable)" option export_proofs : bool = false for content -- "export proof terms to Isabelle/Scala" option prune_proofs : bool = false for content -- "prune proof terms after export (do not store in Isabelle/ML)" section "Theory update" option update_inner_syntax_cartouches : bool = false for content update -- "update inner syntax (types, terms, etc.) to use cartouches" option update_mixfix_cartouches : bool = false for content update -- "update mixfix templates to use cartouches instead of double quotes" option update_control_cartouches : bool = false for content update -- "update antiquotations to use control symbol with cartouche argument" option update_path_cartouches : bool = false for content update -- "update file-system paths to use cartouches" option update_cite : bool = false for content update -- "update cite commands and antiquotations" section "Build Database" option build_database_server : bool = false for connection option build_database_user : string = "" for connection option build_database_password : string = "" for connection option build_database_name : string = "" for connection option build_database_host : string = "" for connection option build_database_port : int = 0 for connection option build_database_ssh_host : string = "" for connection option build_database_ssh_user : string = "" for connection option build_database_ssh_port : int = 0 for connection section "Build Log Database" option build_log_database_user : string = "" for connection option build_log_database_password : string = "" for connection option build_log_database_name : string = "" for connection option build_log_database_host : string = "" for connection option build_log_database_port : int = 0 for connection option build_log_ssh_host : string = "" for connection option build_log_ssh_user : string = "" for connection option build_log_ssh_port : int = 0 for connection option build_log_history : int = 30 -- "length of relevant history (in days)" option build_log_transaction_size : int = 1 -- "number of log files for each db update" section "Isabelle/Scala/ML system channel" option system_channel_address : string = "" for connection option system_channel_password : string = "" for connection section "Bash process execution server" option bash_process_debugging : bool = false for connection option bash_process_address : string = "" for connection option bash_process_password : string = "" for connection diff --git a/src/FOL/ex/Locale_Test/Locale_Test1.thy b/src/FOL/ex/Locale_Test/Locale_Test1.thy --- a/src/FOL/ex/Locale_Test/Locale_Test1.thy +++ b/src/FOL/ex/Locale_Test/Locale_Test1.thy @@ -1,844 +1,844 @@ (* Title: FOL/ex/Locale_Test/Locale_Test1.thy Author: Clemens Ballarin, TU Muenchen Test environment for the locale implementation. *) theory Locale_Test1 imports FOL begin typedecl int instance int :: \term\ .. consts plus :: \int => int => int\ (infixl \+\ 60) zero :: \int\ (\0\) minus :: \int => int\ (\- _\) axiomatization where int_assoc: \(x + y::int) + z = x + (y + z)\ and int_zero: \0 + x = x\ and int_minus: \(-x) + x = 0\ and int_minus2: \-(-x) = x\ section \Inference of parameter types\ locale param1 = fixes p print_locale! param1 locale param2 = fixes p :: \'b\ print_locale! param2 (* locale param_top = param2 r for r :: "'b :: {}" Fails, cannot generalise parameter. *) locale param3 = fixes p (infix \..\ 50) print_locale! param3 locale param4 = fixes p :: \'a => 'a => 'a\ (infix \..\ 50) print_locale! param4 subsection \Incremental type constraints\ locale constraint1 = fixes prod (infixl \**\ 65) assumes l_id: \x ** y = x\ assumes assoc: \(x ** y) ** z = x ** (y ** z)\ print_locale! constraint1 locale constraint2 = fixes p and q assumes \p = q\ print_locale! constraint2 section \Inheritance\ locale semi = fixes prod (infixl \**\ 65) assumes assoc: \(x ** y) ** z = x ** (y ** z)\ print_locale! semi thm semi_def locale lgrp = semi + fixes one and inv assumes lone: \one ** x = x\ and linv: \inv(x) ** x = one\ print_locale! lgrp thm lgrp_def lgrp_axioms_def locale add_lgrp = semi \(++)\ for sum (infixl \++\ 60) + fixes zero and neg assumes lzero: \zero ++ x = x\ and lneg: \neg(x) ++ x = zero\ print_locale! add_lgrp thm add_lgrp_def add_lgrp_axioms_def locale rev_lgrp = semi \%x y. y ++ x\ for sum (infixl \++\ 60) print_locale! rev_lgrp thm rev_lgrp_def locale hom = f: semi \f\ + g: semi \g\ for f and g print_locale! hom thm hom_def locale perturbation = semi + d: semi \%x y. delta(x) ** delta(y)\ for delta print_locale! perturbation thm perturbation_def locale pert_hom = d1: perturbation \f\ \d1\ + d2: perturbation \f\ \d2\ for f d1 d2 print_locale! pert_hom thm pert_hom_def text \Alternative expression, obtaining nicer names in \semi f\.\ locale pert_hom' = semi \f\ + d1: perturbation \f\ \d1\ + d2: perturbation \f\ \d2\ for f d1 d2 print_locale! pert_hom' thm pert_hom'_def section \Syntax declarations\ locale logic = fixes land (infixl \&&\ 55) and lnot (\-- _\ [60] 60) assumes assoc: \(x && y) && z = x && (y && z)\ and notnot: \-- (-- x) = x\ begin definition lor (infixl \||\ 50) where \x || y = --(-- x && -- y)\ end print_locale! logic locale use_decl = l: logic + semi \(||)\ print_locale! use_decl thm use_decl_def locale extra_type = fixes a :: \'a\ and P :: \'a => 'b => o\ begin definition test :: \'a => o\ where \test(x) \ (\b. P(x, b))\ end term \extra_type.test\ thm extra_type.test_def interpretation var?: extra_type \0\ \%x y. x = 0\ . thm var.test_def text \Under which circumstances notation remains active.\ ML \ fun check_syntax ctxt thm expected = let val obtained = Print_Mode.setmp [] (Thm.string_of_thm (Config.put show_markup false ctxt)) thm; in if obtained <> expected then error ("Theorem syntax '" ^ obtained ^ "' obtained, but '" ^ expected ^ "' expected.") else () end; \ declare [[show_hyps]] locale "syntax" = fixes p1 :: \'a => 'b\ and p2 :: \'b => o\ begin definition d1 :: \'a => o\ (\D1'(_')\) where \d1(x) \ \ p2(p1(x))\ definition d2 :: \'b => o\ (\D2'(_')\) where \d2(x) \ \ p2(x)\ thm d1_def d2_def end thm syntax.d1_def syntax.d2_def locale syntax' = "syntax" \p1\ \p2\ for p1 :: \'a => 'a\ and p2 :: \'a => o\ begin thm d1_def d2_def (* should print as "D1(?x) <-> ..." and "D2(?x) <-> ..." *) ML \ check_syntax \<^context> @{thm d1_def} "D1(?x) \ \ p2(p1(?x))"; check_syntax \<^context> @{thm d2_def} "D2(?x) \ \ p2(?x)"; \ end locale syntax'' = "syntax" \p3\ \p2\ for p3 :: \'a => 'b\ and p2 :: \'b => o\ begin thm d1_def d2_def (* should print as "d1(?x) <-> ..." and "D2(?x) <-> ..." *) ML \ check_syntax \<^context> @{thm d1_def} "d1(?x) \ \ p2(p3(?x))"; check_syntax \<^context> @{thm d2_def} "D2(?x) \ \ p2(?x)"; \ end section \Foundational versions of theorems\ thm logic.assoc thm logic.lor_def section \Defines\ locale logic_def = fixes land (infixl \&&\ 55) and lor (infixl \||\ 50) and lnot (\-- _\ [60] 60) assumes assoc: \(x && y) && z = x && (y && z)\ and notnot: \-- (-- x) = x\ defines \x || y == --(-- x && -- y)\ begin thm lor_def lemma \x || y = --(-- x && --y)\ by (unfold lor_def) (rule refl) end (* Inheritance of defines *) locale logic_def2 = logic_def begin lemma \x || y = --(-- x && --y)\ by (unfold lor_def) (rule refl) end section \Notes\ (* A somewhat arcane homomorphism example *) definition semi_hom where \semi_hom(prod, sum, h) \ (\x y. h(prod(x, y)) = sum(h(x), h(y)))\ lemma semi_hom_mult: \semi_hom(prod, sum, h) \ h(prod(x, y)) = sum(h(x), h(y))\ by (simp add: semi_hom_def) locale semi_hom_loc = prod: semi \prod\ + sum: semi \sum\ for prod and sum and h + assumes semi_homh: \semi_hom(prod, sum, h)\ notes semi_hom_mult = semi_hom_mult [OF semi_homh] thm semi_hom_loc.semi_hom_mult -(* unspecified, attribute not applied in backgroud theory !!! *) +(* unspecified, attribute not applied in background theory !!! *) lemma (in semi_hom_loc) \h(prod(x, y)) = sum(h(x), h(y))\ by (rule semi_hom_mult) (* Referring to facts from within a context specification *) lemma assumes x: \P \ P\ notes y = x shows \True\ .. section \Theorem statements\ lemma (in lgrp) lcancel: \x ** y = x ** z \ y = z\ proof assume \x ** y = x ** z\ then have \inv(x) ** x ** y = inv(x) ** x ** z\ by (simp add: assoc) then show \y = z\ by (simp add: lone linv) qed simp print_locale! lgrp locale rgrp = semi + fixes one and inv assumes rone: \x ** one = x\ and rinv: \x ** inv(x) = one\ begin lemma rcancel: \y ** x = z ** x \ y = z\ proof assume \y ** x = z ** x\ then have \y ** (x ** inv(x)) = z ** (x ** inv(x))\ by (simp add: assoc [symmetric]) then show \y = z\ by (simp add: rone rinv) qed simp end print_locale! rgrp subsection \Patterns\ lemma (in rgrp) assumes \y ** x = z ** x\ (is \?a\) shows \y = z\ (is \?t\) proof - txt \Weird proof involving patterns from context element and conclusion.\ { assume \?a\ then have \y ** (x ** inv(x)) = z ** (x ** inv(x))\ by (simp add: assoc [symmetric]) then have \?t\ by (simp add: rone rinv) } note x = this show \?t\ by (rule x [OF \?a\]) qed section \Interpretation between locales: sublocales\ sublocale lgrp < right?: rgrp print_facts proof unfold_locales { fix x have \inv(x) ** x ** one = inv(x) ** x\ by (simp add: linv lone) then show \x ** one = x\ by (simp add: assoc lcancel) } note rone = this { fix x have \inv(x) ** x ** inv(x) = inv(x) ** one\ by (simp add: linv lone rone) then show \x ** inv(x) = one\ by (simp add: assoc lcancel) } qed (* effect on printed locale *) print_locale! lgrp (* use of derived theorem *) lemma (in lgrp) \y ** x = z ** x \ y = z\ apply (rule rcancel) done (* circular interpretation *) sublocale rgrp < left: lgrp proof unfold_locales { fix x have \one ** (x ** inv(x)) = x ** inv(x)\ by (simp add: rinv rone) then show \one ** x = x\ by (simp add: assoc [symmetric] rcancel) } note lone = this { fix x have \inv(x) ** (x ** inv(x)) = one ** inv(x)\ by (simp add: rinv lone rone) then show \inv(x) ** x = one\ by (simp add: assoc [symmetric] rcancel) } qed (* effect on printed locale *) print_locale! rgrp print_locale! lgrp (* Duality *) locale order = fixes less :: \'a => 'a => o\ (infix \<<\ 50) assumes refl: \x << x\ and trans: \[| x << y; y << z |] ==> x << z\ sublocale order < dual: order \%x y. y << x\ apply unfold_locales apply (rule refl) apply (blast intro: trans) done print_locale! order (* Only two instances of order. *) locale order' = fixes less :: \'a => 'a => o\ (infix \<<\ 50) assumes refl: \x << x\ and trans: \[| x << y; y << z |] ==> x << z\ locale order_with_def = order' begin definition greater :: \'a => 'a => o\ (infix \>>\ 50) where \x >> y \ y << x\ end sublocale order_with_def < dual: order' \(>>)\ apply unfold_locales unfolding greater_def apply (rule refl) apply (blast intro: trans) done print_locale! order_with_def (* Note that decls come after theorems that make use of them. *) (* locale with many parameters --- interpretations generate alternating group A5 *) locale A5 = fixes A and B and C and D and E assumes eq: \A \ B \ C \ D \ E\ sublocale A5 < 1: A5 _ _ \D\ \E\ \C\ print_facts using eq apply (blast intro: A5.intro) done sublocale A5 < 2: A5 \C\ _ \E\ _ \A\ print_facts using eq apply (blast intro: A5.intro) done sublocale A5 < 3: A5 \B\ \C\ \A\ _ _ print_facts using eq apply (blast intro: A5.intro) done (* Any even permutation of parameters is subsumed by the above. *) print_locale! A5 (* Free arguments of instance *) locale trivial = fixes P and Q :: \o\ assumes Q: \P \ P \ Q\ begin lemma Q_triv: \Q\ using Q by fast end sublocale trivial < x: trivial \x\ _ apply unfold_locales using Q by fast print_locale! trivial context trivial begin thm x.Q [where ?x = \True\] end sublocale trivial < y: trivial \Q\ \Q\ by unfold_locales (* Succeeds since previous interpretation is more general. *) print_locale! trivial (* No instance for y created (subsumed). *) subsection \Sublocale, then interpretation in theory\ interpretation int?: lgrp \(+)\ \0\ \minus\ proof unfold_locales qed (rule int_assoc int_zero int_minus)+ thm int.assoc int.semi_axioms interpretation int2?: semi \(+)\ by unfold_locales (* subsumed, thm int2.assoc not generated *) ML \(Global_Theory.get_thms \<^theory> "int2.assoc"; raise Fail "thm int2.assoc was generated") handle ERROR _ => ([]:thm list);\ thm int.lone int.right.rone (* the latter comes through the sublocale relation *) subsection \Interpretation in theory, then sublocale\ interpretation fol: logic \(+)\ \minus\ by unfold_locales (rule int_assoc int_minus2)+ locale logic2 = fixes land (infixl \&&\ 55) and lnot (\-- _\ [60] 60) assumes assoc: \(x && y) && z = x && (y && z)\ and notnot: \-- (-- x) = x\ begin definition lor (infixl \||\ 50) where \x || y = --(-- x && -- y)\ end sublocale logic < two: logic2 by unfold_locales (rule assoc notnot)+ thm fol.two.assoc subsection \Declarations and sublocale\ locale logic_a = logic locale logic_b = logic sublocale logic_a < logic_b by unfold_locales subsection \Interpretation\ subsection \Rewrite morphism\ locale logic_o = fixes land (infixl \&&\ 55) and lnot (\-- _\ [60] 60) assumes assoc_o: \(x && y) && z \ x && (y && z)\ and notnot_o: \-- (-- x) \ x\ begin definition lor_o (infixl \||\ 50) where \x || y \ --(-- x && -- y)\ end interpretation x: logic_o \(\)\ \Not\ rewrites bool_logic_o: \x.lor_o(x, y) \ x \ y\ proof - show bool_logic_o: \PROP logic_o((\), Not)\ by unfold_locales fast+ show \logic_o.lor_o((\), Not, x, y) \ x \ y\ by (unfold logic_o.lor_o_def [OF bool_logic_o]) fast qed thm x.lor_o_def bool_logic_o lemma lor_triv: \z \ z\ .. lemma (in logic_o) lor_triv: \x || y \ x || y\ by fast thm lor_triv [where z = \True\] (* Check strict prefix. *) x.lor_triv subsection \Rewrite morphisms in expressions\ interpretation y: logic_o \(\)\ \Not\ rewrites bool_logic_o2: \logic_o.lor_o((\), Not, x, y) \ x \ y\ proof - show bool_logic_o: \PROP logic_o((\), Not)\ by unfold_locales fast+ show \logic_o.lor_o((\), Not, x, y) \ x \ y\ unfolding logic_o.lor_o_def [OF bool_logic_o] by fast qed subsection \Inheritance of rewrite morphisms\ locale reflexive = fixes le :: \'a => 'a => o\ (infix \\\ 50) assumes refl: \x \ x\ begin definition less (infix \\\ 50) where \x \ y \ x \ y \ x \ y\ end axiomatization gle :: \'a => 'a => o\ and gless :: \'a => 'a => o\ and gle' :: \'a => 'a => o\ and gless' :: \'a => 'a => o\ where grefl: \gle(x, x)\ and gless_def: \gless(x, y) \ gle(x, y) \ x \ y\ and grefl': \gle'(x, x)\ and gless'_def: \gless'(x, y) \ gle'(x, y) \ x \ y\ text \Setup\ locale mixin = reflexive begin lemmas less_thm = less_def end interpretation le: mixin \gle\ rewrites \reflexive.less(gle, x, y) \ gless(x, y)\ proof - show \mixin(gle)\ by unfold_locales (rule grefl) note reflexive = this[unfolded mixin_def] show \reflexive.less(gle, x, y) \ gless(x, y)\ by (simp add: reflexive.less_def[OF reflexive] gless_def) qed text \Rewrite morphism propagated along the locale hierarchy\ locale mixin2 = mixin begin lemmas less_thm2 = less_def end interpretation le: mixin2 \gle\ by unfold_locales thm le.less_thm2 (* rewrite morphism applied *) lemma \gless(x, y) \ gle(x, y) \ x \ y\ by (rule le.less_thm2) text \Rewrite morphism does not leak to a side branch.\ locale mixin3 = reflexive begin lemmas less_thm3 = less_def end interpretation le: mixin3 \gle\ by unfold_locales thm le.less_thm3 (* rewrite morphism not applied *) lemma \reflexive.less(gle, x, y) \ gle(x, y) \ x \ y\ by (rule le.less_thm3) text \Rewrite morphism only available in original context\ locale mixin4_base = reflexive locale mixin4_mixin = mixin4_base interpretation le: mixin4_mixin \gle\ rewrites \reflexive.less(gle, x, y) \ gless(x, y)\ proof - show \mixin4_mixin(gle)\ by unfold_locales (rule grefl) note reflexive = this[unfolded mixin4_mixin_def mixin4_base_def mixin_def] show \reflexive.less(gle, x, y) \ gless(x, y)\ by (simp add: reflexive.less_def[OF reflexive] gless_def) qed locale mixin4_copy = mixin4_base begin lemmas less_thm4 = less_def end locale mixin4_combined = le1?: mixin4_mixin \le'\ + le2?: mixin4_copy \le\ for le' le begin lemmas less_thm4' = less_def end interpretation le4: mixin4_combined \gle'\ \gle\ by unfold_locales (rule grefl') thm le4.less_thm4' (* rewrite morphism not applied *) lemma \reflexive.less(gle, x, y) \ gle(x, y) \ x \ y\ by (rule le4.less_thm4') text \Inherited rewrite morphism applied to new theorem\ locale mixin5_base = reflexive locale mixin5_inherited = mixin5_base interpretation le5: mixin5_base \gle\ rewrites \reflexive.less(gle, x, y) \ gless(x, y)\ proof - show \mixin5_base(gle)\ by unfold_locales note reflexive = this[unfolded mixin5_base_def mixin_def] show \reflexive.less(gle, x, y) \ gless(x, y)\ by (simp add: reflexive.less_def[OF reflexive] gless_def) qed interpretation le5: mixin5_inherited \gle\ by unfold_locales lemmas (in mixin5_inherited) less_thm5 = less_def thm le5.less_thm5 (* rewrite morphism applied *) lemma \gless(x, y) \ gle(x, y) \ x \ y\ by (rule le5.less_thm5) text \Rewrite morphism pushed down to existing inherited locale\ locale mixin6_base = reflexive locale mixin6_inherited = mixin5_base interpretation le6: mixin6_base \gle\ by unfold_locales interpretation le6: mixin6_inherited \gle\ by unfold_locales interpretation le6: mixin6_base \gle\ rewrites \reflexive.less(gle, x, y) \ gless(x, y)\ proof - show \mixin6_base(gle)\ by unfold_locales note reflexive = this[unfolded mixin6_base_def mixin_def] show \reflexive.less(gle, x, y) \ gless(x, y)\ by (simp add: reflexive.less_def[OF reflexive] gless_def) qed lemmas (in mixin6_inherited) less_thm6 = less_def thm le6.less_thm6 (* mixin applied *) lemma \gless(x, y) \ gle(x, y) \ x \ y\ by (rule le6.less_thm6) text \Existing rewrite morphism inherited through sublocale relation\ locale mixin7_base = reflexive locale mixin7_inherited = reflexive interpretation le7: mixin7_base \gle\ rewrites \reflexive.less(gle, x, y) \ gless(x, y)\ proof - show \mixin7_base(gle)\ by unfold_locales note reflexive = this[unfolded mixin7_base_def mixin_def] show \reflexive.less(gle, x, y) \ gless(x, y)\ by (simp add: reflexive.less_def[OF reflexive] gless_def) qed interpretation le7: mixin7_inherited \gle\ by unfold_locales lemmas (in mixin7_inherited) less_thm7 = less_def thm le7.less_thm7 (* before, rewrite morphism not applied *) lemma \reflexive.less(gle, x, y) \ gle(x, y) \ x \ y\ by (rule le7.less_thm7) sublocale mixin7_inherited < mixin7_base by unfold_locales lemmas (in mixin7_inherited) less_thm7b = less_def thm le7.less_thm7b (* after, mixin applied *) lemma \gless(x, y) \ gle(x, y) \ x \ y\ by (rule le7.less_thm7b) text \This locale will be interpreted in later theories.\ locale mixin_thy_merge = le: reflexive \le\ + le': reflexive \le'\ for le le' subsection \Rewrite morphisms in sublocale\ text \Simulate a specification of left groups where unit and inverse are defined rather than specified. This is possible, but not in FOL, due to the lack of a selection operator.\ axiomatization glob_one and glob_inv where glob_lone: \prod(glob_one(prod), x) = x\ and glob_linv: \prod(glob_inv(prod, x), x) = glob_one(prod)\ locale dgrp = semi begin definition one where \one = glob_one(prod)\ lemma lone: \one ** x = x\ unfolding one_def by (rule glob_lone) definition inv where \inv(x) = glob_inv(prod, x)\ lemma linv: \inv(x) ** x = one\ unfolding one_def inv_def by (rule glob_linv) end sublocale lgrp < def?: dgrp rewrites one_equation: \dgrp.one(prod) = one\ and inv_equation: \dgrp.inv(prod, x) = inv(x)\ proof - show \dgrp(prod)\ by unfold_locales from this interpret d: dgrp . \ \Unit\ have \dgrp.one(prod) = glob_one(prod)\ by (rule d.one_def) also have \... = glob_one(prod) ** one\ by (simp add: rone) also have \... = one\ by (simp add: glob_lone) finally show \dgrp.one(prod) = one\ . \ \Inverse\ then have \dgrp.inv(prod, x) ** x = inv(x) ** x\ by (simp add: glob_linv d.linv linv) then show \dgrp.inv(prod, x) = inv(x)\ by (simp add: rcancel) qed print_locale! lgrp context lgrp begin text \Equations stored in target\ lemma \dgrp.one(prod) = one\ by (rule one_equation) lemma \dgrp.inv(prod, x) = inv(x)\ by (rule inv_equation) text \Rewrite morphisms applied\ lemma \one = glob_one(prod)\ by (rule one_def) lemma \inv(x) = glob_inv(prod, x)\ by (rule inv_def) end text \Interpreted versions\ lemma \0 = glob_one ((+))\ by (rule int.def.one_def) lemma \- x = glob_inv((+), x)\ by (rule int.def.inv_def) text \Roundup applies rewrite morphisms at declaration level in DFS tree\ locale roundup = fixes x assumes true: \x \ True\ sublocale roundup \ sub: roundup \x\ rewrites \x \ True \ True\ apply unfold_locales apply (simp add: true) done lemma (in roundup) \True \ True \ True\ by (rule sub.true) section \Interpretation in named contexts\ locale container begin interpretation "private": roundup \True\ by unfold_locales rule lemmas true_copy = private.true end context container begin ML \(Context.>> (fn generic => let val context = Context.proof_of generic val _ = Proof_Context.get_thms context "private.true" in generic end); raise Fail "thm private.true was persisted") handle ERROR _ => ([]:thm list);\ thm true_copy end section \Interpretation in proofs\ notepad begin interpret "local": lgrp \(+)\ \0\ \minus\ by unfold_locales (* subsumed *) { fix zero :: \int\ assume \!!x. zero + x = x\ \!!x. (-x) + x = zero\ then interpret local_fixed: lgrp \(+)\ \zero\ \minus\ by unfold_locales thm local_fixed.lone } assume \!!x zero. zero + x = x\ \!!x zero. (-x) + x = zero\ then interpret local_free: lgrp \(+)\ \zero\ \minus\ for zero by unfold_locales thm local_free.lone [where ?zero = \0\] end notepad begin { fix pand and pnot and por assume passoc: \\x y z. pand(pand(x, y), z) \ pand(x, pand(y, z))\ and pnotnot: \\x. pnot(pnot(x)) \ x\ and por_def: \\x y. por(x, y) \ pnot(pand(pnot(x), pnot(y)))\ interpret loc: logic_o \pand\ \pnot\ rewrites por_eq: \\x y. logic_o.lor_o(pand, pnot, x, y) \ por(x, y)\ (* FIXME *) proof - show logic_o: \PROP logic_o(pand, pnot)\ using passoc pnotnot by unfold_locales fix x y show \logic_o.lor_o(pand, pnot, x, y) \ por(x, y)\ by (unfold logic_o.lor_o_def [OF logic_o]) (rule por_def [symmetric]) qed print_interps logic_o have \\x y. por(x, y) \ pnot(pand(pnot(x), pnot(y)))\ by (rule loc.lor_o_def) } end end diff --git a/src/HOL/Analysis/Abstract_Topological_Spaces.thy b/src/HOL/Analysis/Abstract_Topological_Spaces.thy --- a/src/HOL/Analysis/Abstract_Topological_Spaces.thy +++ b/src/HOL/Analysis/Abstract_Topological_Spaces.thy @@ -1,3531 +1,3468 @@ (* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) section \Various Forms of Topological Spaces\ theory Abstract_Topological_Spaces imports Lindelof_Spaces Locally Sum_Topology FSigma begin subsection\Connected topological spaces\ lemma connected_space_eq_frontier_eq_empty: "connected_space X \ (\S. S \ topspace X \ X frontier_of S = {} \ S = {} \ S = topspace X)" by (meson clopenin_eq_frontier_of connected_space_clopen_in) lemma connected_space_frontier_eq_empty: "connected_space X \ S \ topspace X \ (X frontier_of S = {} \ S = {} \ S = topspace X)" by (meson connected_space_eq_frontier_eq_empty frontier_of_empty frontier_of_topspace) lemma connectedin_eq_subset_separated_union: "connectedin X C \ C \ topspace X \ (\S T. separatedin X S T \ C \ S \ T \ C \ S \ C \ T)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using connectedin_subset_topspace connectedin_subset_separated_union by blast next assume ?rhs then show ?lhs by (metis closure_of_subset connectedin_separation dual_order.eq_iff inf.orderE separatedin_def sup.boundedE) qed lemma connectedin_clopen_cases: "\connectedin X C; closedin X T; openin X T\ \ C \ T \ disjnt C T" by (metis Diff_eq_empty_iff Int_empty_right clopenin_eq_frontier_of connectedin_Int_frontier_of disjnt_def) lemma connected_space_quotient_map_image: "\quotient_map X X' q; connected_space X\ \ connected_space X'" by (metis connectedin_continuous_map_image connectedin_topspace quotient_imp_continuous_map quotient_imp_surjective_map) lemma connected_space_retraction_map_image: "\retraction_map X X' r; connected_space X\ \ connected_space X'" using connected_space_quotient_map_image retraction_imp_quotient_map by blast lemma connectedin_imp_perfect_gen: assumes X: "t1_space X" and S: "connectedin X S" and nontriv: "\a. S = {a}" shows "S \ X derived_set_of S" unfolding derived_set_of_def proof (intro subsetI CollectI conjI strip) show XS: "x \ topspace X" if "x \ S" for x using that S connectedin by fastforce show "\y. y \ x \ y \ S \ y \ T" if "x \ S" and "x \ T \ openin X T" for x T proof - have opeXx: "openin X (topspace X - {x})" by (meson X openin_topspace t1_space_openin_delete_alt) moreover have "S \ T \ (topspace X - {x})" using XS that(2) by auto moreover have "(topspace X - {x}) \ S \ {}" by (metis Diff_triv S connectedin double_diff empty_subsetI inf_commute insert_subsetI nontriv that(1)) ultimately show ?thesis using that connectedinD [OF S, of T "topspace X - {x}"] by blast qed qed lemma connectedin_imp_perfect: "\Hausdorff_space X; connectedin X S; \a. S = {a}\ \ S \ X derived_set_of S" by (simp add: Hausdorff_imp_t1_space connectedin_imp_perfect_gen) -proposition connected_space_prod_topology: - "connected_space(prod_topology X Y) \ - topspace(prod_topology X Y) = {} \ connected_space X \ connected_space Y" (is "?lhs=?rhs") -proof (cases "topspace(prod_topology X Y) = {}") - case True - then show ?thesis - using connected_space_topspace_empty by blast -next - case False - then have nonempty: "topspace X \ {}" "topspace Y \ {}" - by force+ - show ?thesis - proof - assume ?lhs - then show ?rhs - by (meson connected_space_quotient_map_image nonempty quotient_map_fst quotient_map_snd) - next - assume ?rhs - then have conX: "connected_space X" and conY: "connected_space Y" - using False by blast+ - have False - if "openin (prod_topology X Y) U" and "openin (prod_topology X Y) V" - and UV: "topspace X \ topspace Y \ U \ V" "U \ V = {}" - and "U \ {}" and "V \ {}" - for U V - proof - - have Usub: "U \ topspace X \ topspace Y" and Vsub: "V \ topspace X \ topspace Y" - using that by (metis openin_subset topspace_prod_topology)+ - obtain a b where ab: "(a,b) \ U" and a: "a \ topspace X" and b: "b \ topspace Y" - using \U \ {}\ Usub by auto - have "\ topspace X \ topspace Y \ U" - using Usub Vsub \U \ V = {}\ \V \ {}\ by auto - then obtain x y where x: "x \ topspace X" and y: "y \ topspace Y" and "(x,y) \ U" - by blast - have oX: "openin X {x \ topspace X. (x,y) \ U}" "openin X {x \ topspace X. (x,y) \ V}" - and oY: "openin Y {y \ topspace Y. (a,y) \ U}" "openin Y {y \ topspace Y. (a,y) \ V}" - by (force intro: openin_continuous_map_preimage [where Y = "prod_topology X Y"] - simp: that continuous_map_pairwise o_def x y a)+ - have 1: "topspace Y \ {y \ topspace Y. (a,y) \ U} \ {y \ topspace Y. (a,y) \ V}" - using a that(3) by auto - have 2: "{y \ topspace Y. (a,y) \ U} \ {y \ topspace Y. (a,y) \ V} = {}" - using that(4) by auto - have 3: "{y \ topspace Y. (a,y) \ U} \ {}" - using ab b by auto - have 4: "{y \ topspace Y. (a,y) \ V} \ {}" - proof - - show ?thesis - using connected_spaceD [OF conX oX] UV \(x,y) \ U\ a x y - disjoint_iff_not_equal by blast - qed - show ?thesis - using connected_spaceD [OF conY oY 1 2 3 4] by auto - qed - then show ?lhs - unfolding connected_space_def topspace_prod_topology by blast - qed -qed - -lemma connectedin_Times: - "connectedin (prod_topology X Y) (S \ T) \ - S = {} \ T = {} \ connectedin X S \ connectedin Y T" - by (force simp: connectedin_def subtopology_Times connected_space_prod_topology) - subsection\The notion of "separated between" (complement of "connected between)"\ definition separated_between where "separated_between X S T \ \U V. openin X U \ openin X V \ U \ V = topspace X \ disjnt U V \ S \ U \ T \ V" lemma separated_between_alt: "separated_between X S T \ (\U V. closedin X U \ closedin X V \ U \ V = topspace X \ disjnt U V \ S \ U \ T \ V)" unfolding separated_between_def by (metis separatedin_open_sets separation_closedin_Un_gen subtopology_topspace separatedin_closed_sets separation_openin_Un_gen) lemma separated_between: "separated_between X S T \ (\U. closedin X U \ openin X U \ S \ U \ T \ topspace X - U)" unfolding separated_between_def closedin_def disjnt_def by (smt (verit, del_insts) Diff_cancel Diff_disjoint Diff_partition Un_Diff Un_Diff_Int openin_subset) lemma separated_between_mono: "\separated_between X S T; S' \ S; T' \ T\ \ separated_between X S' T'" by (meson order.trans separated_between) lemma separated_between_refl: "separated_between X S S \ S = {}" unfolding separated_between_def by (metis Un_empty_right disjnt_def disjnt_empty2 disjnt_subset2 disjnt_sym le_iff_inf openin_empty openin_topspace) lemma separated_between_sym: "separated_between X S T \ separated_between X T S" by (metis disjnt_sym separated_between_alt sup_commute) lemma separated_between_imp_subset: "separated_between X S T \ S \ topspace X \ T \ topspace X" by (metis le_supI1 le_supI2 separated_between_def) lemma separated_between_empty: "(separated_between X {} S \ S \ topspace X) \ (separated_between X S {} \ S \ topspace X)" by (metis Diff_empty bot.extremum closedin_empty openin_empty separated_between separated_between_imp_subset separated_between_sym) lemma separated_between_Un: "separated_between X S (T \ U) \ separated_between X S T \ separated_between X S U" by (auto simp: separated_between) lemma separated_between_Un': "separated_between X (S \ T) U \ separated_between X S U \ separated_between X T U" by (simp add: separated_between_Un separated_between_sym) lemma separated_between_imp_disjoint: "separated_between X S T \ disjnt S T" by (meson disjnt_iff separated_between_def subsetD) lemma separated_between_imp_separatedin: "separated_between X S T \ separatedin X S T" by (meson separated_between_def separatedin_mono separatedin_open_sets) lemma separated_between_full: assumes "S \ T = topspace X" shows "separated_between X S T \ disjnt S T \ closedin X S \ openin X S \ closedin X T \ openin X T" proof - have "separated_between X S T \ separatedin X S T" by (simp add: separated_between_imp_separatedin) then show ?thesis unfolding separated_between_def by (metis assms separation_closedin_Un_gen separation_openin_Un_gen subset_refl subtopology_topspace) qed lemma separated_between_eq_separatedin: "S \ T = topspace X \ (separated_between X S T \ separatedin X S T)" by (simp add: separated_between_full separatedin_full) lemma separated_between_pointwise_left: assumes "compactin X S" shows "separated_between X S T \ (S = {} \ T \ topspace X) \ (\x \ S. separated_between X {x} T)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using separated_between_imp_subset separated_between_mono by fastforce next assume R: ?rhs then have "T \ topspace X" by (meson equals0I separated_between_imp_subset) show ?lhs proof - obtain U where U: "\x \ S. openin X (U x)" "\x \ S. \V. openin X V \ U x \ V = topspace X \ disjnt (U x) V \ {x} \ U x \ T \ V" using R unfolding separated_between_def by metis then have "S \ \(U ` S)" by blast then obtain K where "finite K" "K \ S" and K: "S \ (\i\K. U i)" using assms U unfolding compactin_def by (smt (verit) finite_subset_image imageE) show ?thesis unfolding separated_between proof (intro conjI exI) have "\x. x \ K \ closedin X (U x)" by (smt (verit) \K \ S\ Diff_cancel U(2) Un_Diff Un_Diff_Int disjnt_def openin_closedin_eq subsetD) then show "closedin X (\ (U ` K))" by (metis (mono_tags, lifting) \finite K\ closedin_Union finite_imageI image_iff) show "openin X (\ (U ` K))" using U(1) \K \ S\ by blast show "S \ \ (U ` K)" by (simp add: K) have "\x i. \x \ T; i \ K; x \ U i\ \ False" by (meson U(2) \K \ S\ disjnt_iff subsetD) then show "T \ topspace X - \ (U ` K)" using \T \ topspace X\ by auto qed qed qed lemma separated_between_pointwise_right: "compactin X T \ separated_between X S T \ (T = {} \ S \ topspace X) \ (\y \ T. separated_between X S {y})" by (meson separated_between_pointwise_left separated_between_sym) lemma separated_between_closure_of: "S \ topspace X \ separated_between X (X closure_of S) T \ separated_between X S T" by (meson closure_of_minimal_eq separated_between_alt) lemma separated_between_closure_of': "T \ topspace X \ separated_between X S (X closure_of T) \ separated_between X S T" by (meson separated_between_closure_of separated_between_sym) lemma separated_between_closure_of_eq: "separated_between X S T \ S \ topspace X \ separated_between X (X closure_of S) T" by (metis separated_between_closure_of separated_between_imp_subset) lemma separated_between_closure_of_eq': "separated_between X S T \ T \ topspace X \ separated_between X S (X closure_of T)" by (metis separated_between_closure_of' separated_between_imp_subset) lemma separated_between_frontier_of_eq': "separated_between X S T \ T \ topspace X \ disjnt S T \ separated_between X S (X frontier_of T)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis interior_of_union_frontier_of separated_between_Un separated_between_closure_of_eq' separated_between_imp_disjoint) next assume R: ?rhs then obtain U where U: "closedin X U" "openin X U" "S \ U" "X closure_of T - X interior_of T \ topspace X - U" by (metis frontier_of_def separated_between) show ?lhs proof (rule separated_between_mono [of _ S "X closure_of T"]) have "separated_between X S T" unfolding separated_between proof (intro conjI exI) show "S \ U - T" "T \ topspace X - (U - T)" using R U(3) by (force simp: disjnt_iff)+ have "T \ X closure_of T" by (simp add: R closure_of_subset) then have *: "U - T = U - X interior_of T" using U(4) interior_of_subset by fastforce then show "closedin X (U - T)" by (simp add: U(1) closedin_diff) have "U \ X frontier_of T = {}" using U(4) frontier_of_def by fastforce then show "openin X (U - T)" by (metis * Diff_Un U(2) Un_Diff_Int Un_Int_eq(1) closedin_closure_of interior_of_union_frontier_of openin_diff sup_bot_right) qed then show "separated_between X S (X closure_of T)" by (simp add: R separated_between_closure_of') qed (auto simp: R closure_of_subset) qed lemma separated_between_frontier_of_eq: "separated_between X S T \ S \ topspace X \ disjnt S T \ separated_between X (X frontier_of S) T" by (metis disjnt_sym separated_between_frontier_of_eq' separated_between_sym) lemma separated_between_frontier_of: "\S \ topspace X; disjnt S T\ \ (separated_between X (X frontier_of S) T \ separated_between X S T)" using separated_between_frontier_of_eq by blast lemma separated_between_frontier_of': "\T \ topspace X; disjnt S T\ \ (separated_between X S (X frontier_of T) \ separated_between X S T)" using separated_between_frontier_of_eq' by auto lemma connected_space_separated_between: "connected_space X \ (\S T. separated_between X S T \ S = {} \ T = {})" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Diff_cancel connected_space_clopen_in separated_between subset_empty) next assume ?rhs then show ?lhs by (meson connected_space_eq_not_separated separated_between_eq_separatedin) qed lemma connected_space_imp_separated_between_trivial: "connected_space X \ (separated_between X S T \ S = {} \ T \ topspace X \ S \ topspace X \ T = {})" by (metis connected_space_separated_between separated_between_empty) subsection\Connected components\ lemma connected_component_of_subtopology_eq: "connected_component_of (subtopology X U) a = connected_component_of X a \ connected_component_of_set X a \ U" by (force simp: connected_component_of_set connectedin_subtopology connected_component_of_def fun_eq_iff subset_iff) lemma connected_components_of_subtopology: assumes "C \ connected_components_of X" "C \ U" shows "C \ connected_components_of (subtopology X U)" proof - obtain a where a: "connected_component_of_set X a \ U" and "a \ topspace X" and Ceq: "C = connected_component_of_set X a" using assms by (force simp: connected_components_of_def) then have "a \ U" by (simp add: connected_component_of_refl in_mono) then have "connected_component_of_set X a = connected_component_of_set (subtopology X U) a" by (metis a connected_component_of_subtopology_eq) then show ?thesis by (simp add: Ceq \a \ U\ \a \ topspace X\ connected_component_in_connected_components_of) qed thm connected_space_iff_components_eq lemma open_in_finite_connected_components: assumes "finite(connected_components_of X)" "C \ connected_components_of X" shows "openin X C" proof - have "closedin X (topspace X - C)" by (metis DiffD1 assms closedin_Union closedin_connected_components_of complement_connected_components_of_Union finite_Diff) then show ?thesis by (simp add: assms connected_components_of_subset openin_closedin) qed thm connected_component_of_eq_overlap lemma connected_components_of_disjoint: assumes "C \ connected_components_of X" "C' \ connected_components_of X" shows "(disjnt C C' \ (C \ C'))" proof - have "C \ {}" using nonempty_connected_components_of assms by blast with assms show ?thesis by (metis disjnt_self_iff_empty pairwiseD pairwise_disjoint_connected_components_of) qed lemma connected_components_of_overlap: "\C \ connected_components_of X; C' \ connected_components_of X\ \ C \ C' \ {} \ C = C'" by (meson connected_components_of_disjoint disjnt_def) lemma pairwise_separated_connected_components_of: "pairwise (separatedin X) (connected_components_of X)" by (simp add: closedin_connected_components_of connected_components_of_disjoint pairwiseI separatedin_closed_sets) lemma finite_connected_components_of_finite: "finite(topspace X) \ finite(connected_components_of X)" by (simp add: Union_connected_components_of finite_UnionD) lemma connected_component_of_unique: "\x \ C; connectedin X C; \C'. x \ C' \ connectedin X C' \ C' \ C\ \ connected_component_of_set X x = C" by (meson connected_component_of_maximal connectedin_connected_component_of subsetD subset_antisym) lemma closedin_connected_component_of_subtopology: "\C \ connected_components_of (subtopology X s); X closure_of C \ s\ \ closedin X C" by (metis closedin_Int_closure_of closedin_connected_components_of closure_of_eq inf.absorb_iff2) lemma connected_component_of_discrete_topology: "connected_component_of_set (discrete_topology U) x = (if x \ U then {x} else {})" by (simp add: locally_path_connected_space_discrete_topology flip: path_component_eq_connected_component_of) lemma connected_components_of_discrete_topology: "connected_components_of (discrete_topology U) = (\x. {x}) ` U" by (simp add: connected_component_of_discrete_topology connected_components_of_def) lemma connected_component_of_continuous_image: "\continuous_map X Y f; connected_component_of X x y\ \ connected_component_of Y (f x) (f y)" by (meson connected_component_of_def connectedin_continuous_map_image image_eqI) lemma homeomorphic_map_connected_component_of: assumes "homeomorphic_map X Y f" and x: "x \ topspace X" shows "connected_component_of_set Y (f x) = f ` (connected_component_of_set X x)" proof - obtain g where g: "continuous_map X Y f" "continuous_map Y X g " "\x. x \ topspace X \ g (f x) = x" "\y. y \ topspace Y \ f (g y) = y" using assms(1) homeomorphic_map_maps homeomorphic_maps_def by fastforce show ?thesis using connected_component_in_topspace [of Y] x g connected_component_of_continuous_image [of X Y f] connected_component_of_continuous_image [of Y X g] by force qed lemma homeomorphic_map_connected_components_of: assumes "homeomorphic_map X Y f" shows "connected_components_of Y = (image f) ` (connected_components_of X)" proof - have "topspace Y = f ` topspace X" by (metis assms homeomorphic_imp_surjective_map) with homeomorphic_map_connected_component_of [OF assms] show ?thesis by (auto simp: connected_components_of_def image_iff) qed lemma connected_component_of_pair: "connected_component_of_set (prod_topology X Y) (x,y) = connected_component_of_set X x \ connected_component_of_set Y y" proof (cases "x \ topspace X \ y \ topspace Y") case True show ?thesis proof (rule connected_component_of_unique) show "(x, y) \ connected_component_of_set X x \ connected_component_of_set Y y" using True by (simp add: connected_component_of_refl) show "connectedin (prod_topology X Y) (connected_component_of_set X x \ connected_component_of_set Y y)" by (metis connectedin_Times connectedin_connected_component_of) show "C \ connected_component_of_set X x \ connected_component_of_set Y y" if "(x, y) \ C \ connectedin (prod_topology X Y) C" for C using that unfolding connected_component_of_def apply clarsimp by (metis (no_types) connectedin_continuous_map_image continuous_map_fst continuous_map_snd fst_conv imageI snd_conv) qed next case False then show ?thesis by (metis Sigma_empty1 Sigma_empty2 connected_component_of_eq_empty mem_Sigma_iff topspace_prod_topology) qed lemma connected_components_of_prod_topology: "connected_components_of (prod_topology X Y) = {C \ D |C D. C \ connected_components_of X \ D \ connected_components_of Y}" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" apply (clarsimp simp: connected_components_of_def) by (metis (no_types) connected_component_of_pair imageI) next show "?rhs \ ?lhs" using connected_component_of_pair by (fastforce simp: connected_components_of_def) qed lemma connected_component_of_product_topology: "connected_component_of_set (product_topology X I) x = (if x \ extensional I then PiE I (\i. connected_component_of_set (X i) (x i)) else {})" (is "?lhs = If _ ?R _") proof (cases "x \ topspace(product_topology X I)") case True have "?lhs = (\\<^sub>E i\I. connected_component_of_set (X i) (x i))" if xX: "\i. i\I \ x i \ topspace (X i)" and ext: "x \ extensional I" proof (rule connected_component_of_unique) show "x \ ?R" by (simp add: PiE_iff connected_component_of_refl local.ext xX) show "connectedin (product_topology X I) ?R" by (simp add: connectedin_PiE connectedin_connected_component_of) show "C \ ?R" if "x \ C \ connectedin (product_topology X I) C" for C proof - have "C \ extensional I" using PiE_def connectedin_subset_topspace that by fastforce have "\y. y \ C \ y \ (\ i\I. connected_component_of_set (X i) (x i))" apply (simp add: connected_component_of_def Pi_def) by (metis connectedin_continuous_map_image continuous_map_product_projection imageI that) then show ?thesis using PiE_def \C \ extensional I\ by fastforce qed qed with True show ?thesis by (simp add: PiE_iff) next case False then show ?thesis apply (simp add: PiE_iff) by (smt (verit) Collect_empty_eq False PiE_eq_empty_iff PiE_iff connected_component_of_eq_empty) qed lemma connected_components_of_product_topology: "connected_components_of (product_topology X I) = {PiE I B |B. \i \ I. B i \ connected_components_of(X i)}" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" by (auto simp: connected_components_of_def connected_component_of_product_topology PiE_iff) show "?rhs \ ?lhs" proof fix F assume "F \ ?rhs" then obtain B where Feq: "F = Pi\<^sub>E I B" and "\i\I. \x\topspace (X i). B i = connected_component_of_set (X i) x" by (force simp: connected_components_of_def connected_component_of_product_topology image_iff) then obtain f where f: "\i. i \ I \ f i \ topspace (X i) \ B i = connected_component_of_set (X i) (f i)" by metis then have "(\i\I. f i) \ ((\\<^sub>E i\I. topspace (X i)) \ extensional I)" by simp with f show "F \ ?lhs" unfolding Feq connected_components_of_def connected_component_of_product_topology image_iff by (smt (verit, del_insts) PiE_cong restrict_PiE_iff restrict_apply' restrict_extensional topspace_product_topology) qed qed subsection \Monotone maps (in the general topological sense)\ definition monotone_map where "monotone_map X Y f == f ` (topspace X) \ topspace Y \ (\y \ topspace Y. connectedin X {x \ topspace X. f x = y})" lemma monotone_map: "monotone_map X Y f \ f ` (topspace X) \ topspace Y \ (\y. connectedin X {x \ topspace X. f x = y})" apply (simp add: monotone_map_def) by (metis (mono_tags, lifting) connectedin_empty [of X] Collect_empty_eq image_subset_iff) lemma monotone_map_in_subtopology: "monotone_map X (subtopology Y S) f \ monotone_map X Y f \ f ` (topspace X) \ S" by (smt (verit, del_insts) le_inf_iff monotone_map topspace_subtopology) lemma monotone_map_from_subtopology: assumes "monotone_map X Y f" "\x y. x \ topspace X \ y \ topspace X \ x \ S \ f x = f y \ y \ S" shows "monotone_map (subtopology X S) Y f" using assms unfolding monotone_map_def connectedin_subtopology by (smt (verit, del_insts) Collect_cong Collect_empty_eq IntE IntI connectedin_empty image_subset_iff mem_Collect_eq subsetI topspace_subtopology) lemma monotone_map_restriction: "monotone_map X Y f \ {x \ topspace X. f x \ v} = u \ monotone_map (subtopology X u) (subtopology Y v) f" by (smt (verit, best) IntI Int_Collect image_subset_iff mem_Collect_eq monotone_map monotone_map_from_subtopology topspace_subtopology) lemma injective_imp_monotone_map: assumes "f ` topspace X \ topspace Y" "inj_on f (topspace X)" shows "monotone_map X Y f" unfolding monotone_map_def proof (intro conjI assms strip) fix y assume "y \ topspace Y" then have "{x \ topspace X. f x = y} = {} \ (\a \ topspace X. {x \ topspace X. f x = y} = {a})" using assms(2) unfolding inj_on_def by blast then show "connectedin X {x \ topspace X. f x = y}" by (metis (no_types, lifting) connectedin_empty connectedin_sing) qed lemma embedding_imp_monotone_map: "embedding_map X Y f \ monotone_map X Y f" by (metis (no_types) embedding_map_def homeomorphic_eq_everything_map inf.absorb_iff2 injective_imp_monotone_map topspace_subtopology) lemma section_imp_monotone_map: "section_map X Y f \ monotone_map X Y f" by (simp add: embedding_imp_monotone_map section_imp_embedding_map) lemma homeomorphic_imp_monotone_map: "homeomorphic_map X Y f \ monotone_map X Y f" by (meson section_and_retraction_eq_homeomorphic_map section_imp_monotone_map) proposition connected_space_monotone_quotient_map_preimage: assumes f: "monotone_map X Y f" "quotient_map X Y f" and "connected_space Y" shows "connected_space X" proof (rule ccontr) assume "\ connected_space X" then obtain U V where "openin X U" "openin X V" "U \ V = {}" "U \ {}" "V \ {}" and topUV: "topspace X \ U \ V" by (auto simp: connected_space_def) then have UVsub: "U \ topspace X" "V \ topspace X" by (auto simp: openin_subset) have "\ connected_space Y" unfolding connected_space_def not_not proof (intro exI conjI) show "topspace Y \ f`U \ f`V" by (metis f(2) image_Un quotient_imp_surjective_map subset_Un_eq topUV) show "f`U \ {}" by (simp add: \U \ {}\) show "(f`V) \ {}" by (simp add: \V \ {}\) have *: "y \ f ` V" if "y \ f ` U" for y proof - have \
: "connectedin X {x \ topspace X. f x = y}" using f(1) monotone_map by fastforce show ?thesis using connectedinD [OF \
\openin X U\ \openin X V\] UVsub topUV \U \ V = {}\ that by (force simp: disjoint_iff) qed then show "f`U \ f`V = {}" by blast show "openin Y (f`U)" using f \openin X U\ topUV * unfolding quotient_map_saturated_open by force show "openin Y (f`V)" using f \openin X V\ topUV * unfolding quotient_map_saturated_open by force qed then show False by (simp add: assms) qed lemma connectedin_monotone_quotient_map_preimage: assumes "monotone_map X Y f" "quotient_map X Y f" "connectedin Y C" "openin Y C \ closedin Y C" shows "connectedin X {x \ topspace X. f x \ C}" proof - have "connected_space (subtopology X {x \ topspace X. f x \ C})" proof - have "connected_space (subtopology Y C)" using \connectedin Y C\ connectedin_def by blast moreover have "quotient_map (subtopology X {a \ topspace X. f a \ C}) (subtopology Y C) f" by (simp add: assms quotient_map_restriction) ultimately show ?thesis using \monotone_map X Y f\ connected_space_monotone_quotient_map_preimage monotone_map_restriction by blast qed then show ?thesis by (simp add: connectedin_def) qed lemma monotone_open_map: assumes "continuous_map X Y f" "open_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "monotone_map X Y f \ (\C. connectedin Y C \ connectedin X {x \ topspace X. f x \ C})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding connectedin_def proof (intro strip conjI) fix C assume C: "C \ topspace Y \ connected_space (subtopology Y C)" show "connected_space (subtopology X {x \ topspace X. f x \ C})" proof (rule connected_space_monotone_quotient_map_preimage) show "monotone_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" by (simp add: L monotone_map_restriction) show "quotient_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" proof (rule continuous_open_imp_quotient_map) show "continuous_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" using assms continuous_map_from_subtopology continuous_map_in_subtopology by fastforce qed (use open_map_restriction assms in fastforce)+ qed (simp add: C) qed auto next assume ?rhs then have "\y. connectedin Y {y} \ connectedin X {x \ topspace X. f x = y}" by (smt (verit) Collect_cong singletonD singletonI) then show ?lhs by (simp add: fim monotone_map_def) qed lemma monotone_closed_map: assumes "continuous_map X Y f" "closed_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "monotone_map X Y f \ (\C. connectedin Y C \ connectedin X {x \ topspace X. f x \ C})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding connectedin_def proof (intro strip conjI) fix C assume C: "C \ topspace Y \ connected_space (subtopology Y C)" show "connected_space (subtopology X {x \ topspace X. f x \ C})" proof (rule connected_space_monotone_quotient_map_preimage) show "monotone_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" by (simp add: L monotone_map_restriction) show "quotient_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" proof (rule continuous_closed_imp_quotient_map) show "continuous_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" using assms continuous_map_from_subtopology continuous_map_in_subtopology by fastforce qed (use closed_map_restriction assms in fastforce)+ qed (simp add: C) qed auto next assume ?rhs then have "\y. connectedin Y {y} \ connectedin X {x \ topspace X. f x = y}" by (smt (verit) Collect_cong singletonD singletonI) then show ?lhs by (simp add: fim monotone_map_def) qed subsection\Other countability properties\ definition second_countable where "second_countable X \ \\. countable \ \ (\V \ \. openin X V) \ (\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U))" definition first_countable where "first_countable X \ \x \ topspace X. \\. countable \ \ (\V \ \. openin X V) \ (\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U))" definition separable_space where "separable_space X \ \C. countable C \ C \ topspace X \ X closure_of C = topspace X" lemma second_countable: "second_countable X \ (\\. countable \ \ openin X = arbitrary union_of (\x. x \ \))" by (smt (verit) openin_topology_base_unique second_countable_def) lemma second_countable_subtopology: assumes "second_countable X" shows "second_countable (subtopology X S)" proof - obtain \ where \: "countable \" "\V. V \ \ \ openin X V" "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) show ?thesis unfolding second_countable_def proof (intro exI conjI) show "\V\((\)S) ` \. openin (subtopology X S) V" using openin_subtopology_Int2 \ by blast show "\U x. openin (subtopology X S) U \ x \ U \ (\V\((\)S) ` \. x \ V \ V \ U)" using \ unfolding openin_subtopology by (smt (verit, del_insts) IntI image_iff inf_commute inf_le1 subset_iff) qed (use \ in auto) qed lemma second_countable_discrete_topology: "second_countable(discrete_topology U) \ countable U" (is "?lhs=?rhs") proof assume L: ?lhs then obtain \ where \: "countable \" "\V. V \ \ \ V \ U" "\W x. W \ U \ x \ W \ (\V \ \. x \ V \ V \ W)" by (auto simp: second_countable_def) then have "{x} \ \" if "x \ U" for x by (metis empty_subsetI insertCI insert_subset subset_antisym that) then show ?rhs by (smt (verit) countable_subset image_subsetI \countable \\ countable_image_inj_on [OF _ inj_singleton]) next assume ?rhs then show ?lhs unfolding second_countable_def by (rule_tac x="(\x. {x}) ` U" in exI) auto qed lemma second_countable_open_map_image: assumes "continuous_map X Y f" "open_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "second_countable X" shows "second_countable Y" proof - have openXYf: "\U. openin X U \ openin Y (f ` U)" using assms by (auto simp: open_map_def) obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) show ?thesis unfolding second_countable_def proof (intro exI conjI strip) fix V y assume V: "openin Y V \ y \ V" then obtain x where "x \ topspace X" and x: "f x = y" by (metis fim image_iff openin_subset subsetD) then obtain W where "W\\" "x \ W" "W \ {x \ topspace X. f x \ V}" using * [of "{x \ topspace X. f x \ V}" x] V assms openin_continuous_map_preimage by force then show "\W \ (image f) ` \. y \ W \ W \ V" using x by auto qed (use \ openXYf in auto) qed lemma homeomorphic_space_second_countability: "X homeomorphic_space Y \ (second_countable X \ second_countable Y)" by (meson homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym second_countable_open_map_image) lemma second_countable_retraction_map_image: "\retraction_map X Y r; second_countable X\ \ second_countable Y" using hereditary_imp_retractive_property homeomorphic_space_second_countability second_countable_subtopology by blast lemma second_countable_imp_first_countable: "second_countable X \ first_countable X" by (metis first_countable_def second_countable_def) lemma first_countable_subtopology: assumes "first_countable X" shows "first_countable (subtopology X S)" unfolding first_countable_def proof fix x assume "x \ topspace (subtopology X S)" then obtain \ where "countable \" and \: "\V. V \ \ \ openin X V" "\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms first_countable_def by force show "\\. countable \ \ (\V\\. openin (subtopology X S) V) \ (\U. openin (subtopology X S) U \ x \ U \ (\V\\. x \ V \ V \ U))" proof (intro exI conjI strip) show "countable (((\)S) ` \)" using \countable \\ by blast show "openin (subtopology X S) V" if "V \ ((\)S) ` \" for V using \ openin_subtopology_Int2 that by fastforce show "\V\((\)S) ` \. x \ V \ V \ U" if "openin (subtopology X S) U \ x \ U" for U using that \(2) by (clarsimp simp: openin_subtopology) (meson le_infI2) qed qed lemma first_countable_discrete_topology: "first_countable (discrete_topology U)" unfolding first_countable_def topspace_discrete_topology openin_discrete_topology proof fix x assume "x \ U" show "\\. countable \ \ (\V\\. V \ U) \ (\Ua. Ua \ U \ x \ Ua \ (\V\\. x \ V \ V \ Ua))" using \x \ U\ by (rule_tac x="{{x}}" in exI) auto qed lemma first_countable_open_map_image: assumes "continuous_map X Y f" "open_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "first_countable X" shows "first_countable Y" unfolding first_countable_def proof fix y assume "y \ topspace Y" have openXYf: "\U. openin X U \ openin Y (f ` U)" using assms by (auto simp: open_map_def) then obtain x where x: "x \ topspace X" "f x = y" by (metis \y \ topspace Y\ fim imageE) obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms x first_countable_def by force show "\\. countable \ \ (\V\\. openin Y V) \ (\U. openin Y U \ y \ U \ (\V\\. y \ V \ V \ U))" proof (intro exI conjI strip) fix V assume "openin Y V \ y \ V" then have "\W\\. x \ W \ W \ {x \ topspace X. f x \ V}" using * [of "{x \ topspace X. f x \ V}"] assms openin_continuous_map_preimage x by fastforce then show "\V' \ (image f) ` \. y \ V' \ V' \ V" using image_mono x by auto qed (use \ openXYf in force)+ qed lemma homeomorphic_space_first_countability: "X homeomorphic_space Y \ first_countable X \ first_countable Y" by (meson first_countable_open_map_image homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym) lemma first_countable_retraction_map_image: "\retraction_map X Y r; first_countable X\ \ first_countable Y" using first_countable_subtopology hereditary_imp_retractive_property homeomorphic_space_first_countability by blast lemma separable_space_open_subset: assumes "separable_space X" "openin X S" shows "separable_space (subtopology X S)" proof - obtain C where C: "countable C" "C \ topspace X" "X closure_of C = topspace X" by (meson assms separable_space_def) then have "\x T. \x \ topspace X; x \ T; openin (subtopology X S) T\ \ \y. y \ S \ y \ C \ y \ T" by (smt (verit) \openin X S\ in_closure_of openin_open_subtopology subsetD) with C \openin X S\ show ?thesis unfolding separable_space_def by (rule_tac x="S \ C" in exI) (force simp: in_closure_of) qed lemma separable_space_continuous_map_image: assumes "separable_space X" "continuous_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "separable_space Y" proof - have cont: "\S. f ` (X closure_of S) \ Y closure_of f ` S" by (simp add: assms continuous_map_image_closure_subset) obtain C where C: "countable C" "C \ topspace X" "X closure_of C = topspace X" by (meson assms separable_space_def) then show ?thesis unfolding separable_space_def by (metis cont fim closure_of_subset_topspace countable_image image_mono subset_antisym) qed lemma separable_space_quotient_map_image: "\quotient_map X Y q; separable_space X\ \ separable_space Y" by (meson quotient_imp_continuous_map quotient_imp_surjective_map separable_space_continuous_map_image) lemma separable_space_retraction_map_image: "\retraction_map X Y r; separable_space X\ \ separable_space Y" using retraction_imp_quotient_map separable_space_quotient_map_image by blast lemma homeomorphic_separable_space: "X homeomorphic_space Y \ (separable_space X \ separable_space Y)" by (meson homeomorphic_eq_everything_map homeomorphic_maps_map homeomorphic_space_def separable_space_continuous_map_image) lemma separable_space_discrete_topology: "separable_space(discrete_topology U) \ countable U" by (metis countable_Int2 discrete_topology_closure_of dual_order.refl inf.orderE separable_space_def topspace_discrete_topology) lemma second_countable_imp_separable_space: assumes "second_countable X" shows "separable_space X" proof - obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) obtain c where c: "\V. \V \ \; V \ {}\ \ c V \ V" by (metis all_not_in_conv) then have **: "\x. x \ topspace X \ x \ X closure_of c ` (\ - {{}})" using * by (force simp: closure_of_def) show ?thesis unfolding separable_space_def proof (intro exI conjI) show "countable (c ` (\-{{}}))" using \(1) by blast show "(c ` (\-{{}})) \ topspace X" using \(2) c openin_subset by fastforce show "X closure_of (c ` (\-{{}})) = topspace X" by (meson ** closure_of_subset_topspace subsetI subset_antisym) qed qed lemma second_countable_imp_Lindelof_space: assumes "second_countable X" shows "Lindelof_space X" unfolding Lindelof_space_def proof clarify fix \ assume "\U \ \. openin X U" and UU: "\\ = topspace X" obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) define \' where "\' = {B \ \. \U. U \ \ \ B \ U}" have \': "countable \'" "\\' = \\" using \ using "*" \\U\\. openin X U\ by (fastforce simp: \'_def)+ have "\b. \U. b \ \' \ U \ \ \ b \ U" by (simp add: \'_def) then obtain G where G: "\b. b \ \' \ G b \ \ \ b \ G b" by metis with \' UU show "\\. countable \ \ \ \ \ \ \\ = topspace X" by (rule_tac x="G ` \'" in exI) fastforce qed subsection \Neigbourhood bases EXTRAS\ (* Neigbourhood bases (useful for "local" properties of various kind). *) lemma openin_topology_neighbourhood_base_unique: "openin X = arbitrary union_of P \ (\u. P u \ openin X u) \ neighbourhood_base_of P X" by (smt (verit, best) open_neighbourhood_base_of openin_topology_base_unique) lemma neighbourhood_base_at_topology_base: " openin X = arbitrary union_of b \ (neighbourhood_base_at x P X \ (\w. b w \ x \ w \ (\u v. openin X u \ P v \ x \ u \ u \ v \ v \ w)))" apply (simp add: neighbourhood_base_at_def) by (smt (verit, del_insts) openin_topology_base_unique subset_trans) lemma neighbourhood_base_of_unlocalized: assumes "\S t. P S \ openin X t \ (t \ {}) \ t \ S \ P t" shows "neighbourhood_base_of P X \ (\x \ topspace X. \u v. openin X u \ P v \ x \ u \ u \ v \ v \ topspace X)" apply (simp add: neighbourhood_base_of_def) by (smt (verit, ccfv_SIG) assms empty_iff neighbourhood_base_at_unlocalized) lemma neighbourhood_base_at_discrete_topology: "neighbourhood_base_at x P (discrete_topology u) \ x \ u \ P {x}" apply (simp add: neighbourhood_base_at_def) by (smt (verit) empty_iff empty_subsetI insert_subset singletonI subsetD subset_singletonD) lemma neighbourhood_base_of_discrete_topology: "neighbourhood_base_of P (discrete_topology u) \ (\x \ u. P {x})" apply (simp add: neighbourhood_base_of_def) using neighbourhood_base_at_discrete_topology[of _ P u] by (metis empty_subsetI insert_subset neighbourhood_base_at_def openin_discrete_topology singletonI) lemma second_countable_neighbourhood_base_alt: "second_countable X \ (\\. countable \ \ (\V \ \. openin X V) \ neighbourhood_base_of (\A. A\\) X)" by (metis (full_types) openin_topology_neighbourhood_base_unique second_countable) lemma first_countable_neighbourhood_base_alt: "first_countable X \ (\x \ topspace X. \\. countable \ \ (\V \ \. openin X V) \ neighbourhood_base_at x (\V. V \ \) X)" unfolding first_countable_def apply (intro ball_cong refl ex_cong conj_cong) by (metis (mono_tags, lifting) open_neighbourhood_base_at) lemma second_countable_neighbourhood_base: "second_countable X \ (\\. countable \ \ neighbourhood_base_of (\V. V \ \) X)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using second_countable_neighbourhood_base_alt by blast next assume ?rhs then obtain \ where "countable \" and \: "\W x. openin X W \ x \ W \ (\U. openin X U \ (\V. V \ \ \ x \ U \ U \ V \ V \ W))" by (metis neighbourhood_base_of) then show ?lhs unfolding second_countable_neighbourhood_base_alt neighbourhood_base_of apply (rule_tac x="(\u. X interior_of u) ` \" in exI) by (smt (verit, best) interior_of_eq interior_of_mono countable_image image_iff openin_interior_of) qed lemma first_countable_neighbourhood_base: "first_countable X \ (\x \ topspace X. \\. countable \ \ neighbourhood_base_at x (\V. V \ \) X)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis first_countable_neighbourhood_base_alt) next assume R: ?rhs show ?lhs unfolding first_countable_neighbourhood_base_alt proof fix x assume "x \ topspace X" with R obtain \ where "countable \" and \: "neighbourhood_base_at x (\V. V \ \) X" by blast then show "\\. countable \ \ Ball \ (openin X) \ neighbourhood_base_at x (\V. V \ \) X" unfolding neighbourhood_base_at_def apply (rule_tac x="(\u. X interior_of u) ` \" in exI) by (smt (verit, best) countable_image image_iff interior_of_eq interior_of_mono openin_interior_of) qed qed subsection\$T_0$ spaces and the Kolmogorov quotient\ definition t0_space where "t0_space X \ \x \ topspace X. \y \ topspace X. x \ y \ (\U. openin X U \ (x \ U \ y \ U))" lemma t0_space_expansive: "\topspace Y = topspace X; \U. openin X U \ openin Y U\ \ t0_space X \ t0_space Y" by (metis t0_space_def) lemma t1_imp_t0_space: "t1_space X \ t0_space X" by (metis t0_space_def t1_space_def) lemma t1_eq_symmetric_t0_space_alt: "t1_space X \ t0_space X \ (\x \ topspace X. \y \ topspace X. x \ X closure_of {y} \ y \ X closure_of {x})" apply (simp add: t0_space_def t1_space_def closure_of_def) by (smt (verit, best) openin_topspace) lemma t1_eq_symmetric_t0_space: "t1_space X \ t0_space X \ (\x y. x \ X closure_of {y} \ y \ X closure_of {x})" by (auto simp: t1_eq_symmetric_t0_space_alt in_closure_of) lemma Hausdorff_imp_t0_space: "Hausdorff_space X \ t0_space X" by (simp add: Hausdorff_imp_t1_space t1_imp_t0_space) lemma t0_space: "t0_space X \ (\x \ topspace X. \y \ topspace X. x \ y \ (\C. closedin X C \ (x \ C \ y \ C)))" unfolding t0_space_def by (metis Diff_iff closedin_def openin_closedin_eq) lemma homeomorphic_t0_space: assumes "X homeomorphic_space Y" shows "t0_space X \ t0_space Y" proof - obtain f where f: "homeomorphic_map X Y f" and F: "inj_on f (topspace X)" and "topspace Y = f ` topspace X" by (metis assms homeomorphic_imp_injective_map homeomorphic_imp_surjective_map homeomorphic_space) with inj_on_image_mem_iff [OF F] show ?thesis apply (simp add: t0_space_def homeomorphic_eq_everything_map continuous_map_def open_map_def inj_on_def) by (smt (verit) mem_Collect_eq openin_subset) qed lemma t0_space_closure_of_sing: "t0_space X \ (\x \ topspace X. \y \ topspace X. X closure_of {x} = X closure_of {y} \ x = y)" by (simp add: t0_space_def closure_of_def set_eq_iff) (smt (verit)) lemma t0_space_discrete_topology: "t0_space (discrete_topology S)" by (simp add: Hausdorff_imp_t0_space) lemma t0_space_subtopology: "t0_space X \ t0_space (subtopology X U)" by (simp add: t0_space_def openin_subtopology) (metis Int_iff) lemma t0_space_retraction_map_image: "\retraction_map X Y r; t0_space X\ \ t0_space Y" using hereditary_imp_retractive_property homeomorphic_t0_space t0_space_subtopology by blast lemma XY: "{x}\{y} = {(x,y)}" by simp lemma t0_space_prod_topologyI: "\t0_space X; t0_space Y\ \ t0_space (prod_topology X Y)" by (simp add: t0_space_closure_of_sing closure_of_Times closure_of_eq_empty_gen times_eq_iff flip: XY insert_Times_insert) lemma t0_space_prod_topology_iff: "t0_space (prod_topology X Y) \ topspace (prod_topology X Y) = {} \ t0_space X \ t0_space Y" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Sigma_empty1 Sigma_empty2 retraction_map_fst retraction_map_snd t0_space_retraction_map_image topspace_prod_topology) qed (metis empty_iff t0_space_def t0_space_prod_topologyI) proposition t0_space_product_topology: "t0_space (product_topology X I) \ topspace(product_topology X I) = {} \ (\i \ I. t0_space (X i))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (meson retraction_map_product_projection t0_space_retraction_map_image) next assume R: ?rhs show ?lhs proof (cases "topspace(product_topology X I) = {}") case True then show ?thesis by (simp add: t0_space_def) next case False show ?thesis unfolding t0_space proof (intro strip) fix x y assume x: "x \ topspace (product_topology X I)" and y: "y \ topspace (product_topology X I)" and "x \ y" then obtain i where "i \ I" "x i \ y i" by (metis PiE_ext topspace_product_topology) then have "t0_space (X i)" using False R by blast then obtain U where "closedin (X i) U" "(x i \ U \ y i \ U)" by (metis t0_space PiE_mem \i \ I\ \x i \ y i\ topspace_product_topology x y) with \i \ I\ x y show "\U. closedin (product_topology X I) U \ (x \ U) = (y \ U)" by (rule_tac x="PiE I (\j. if j = i then U else topspace(X j))" in exI) (simp add: closedin_product_topology PiE_iff) qed qed qed subsection \Kolmogorov quotients\ definition Kolmogorov_quotient where "Kolmogorov_quotient X \ \x. @y. \U. openin X U \ (y \ U \ x \ U)" lemma Kolmogorov_quotient_in_open: "openin X U \ (Kolmogorov_quotient X x \ U \ x \ U)" by (smt (verit, ccfv_SIG) Kolmogorov_quotient_def someI_ex) lemma Kolmogorov_quotient_in_topspace: "Kolmogorov_quotient X x \ topspace X \ x \ topspace X" by (simp add: Kolmogorov_quotient_in_open) lemma Kolmogorov_quotient_in_closed: "closedin X C \ (Kolmogorov_quotient X x \ C \ x \ C)" unfolding closedin_def by (meson DiffD2 DiffI Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace in_mono) lemma continuous_map_Kolmogorov_quotient: "continuous_map X X (Kolmogorov_quotient X)" using Kolmogorov_quotient_in_open openin_subopen openin_subset by (fastforce simp: continuous_map_def Kolmogorov_quotient_in_topspace) lemma open_map_Kolmogorov_quotient_explicit: "openin X U \ Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \ U" using Kolmogorov_quotient_in_open openin_subset by fastforce lemma open_map_Kolmogorov_quotient_gen: "open_map (subtopology X S) (subtopology X (image (Kolmogorov_quotient X) S)) (Kolmogorov_quotient X)" proof (clarsimp simp: open_map_def openin_subtopology_alt image_iff) fix U assume "openin X U" then have "Kolmogorov_quotient X ` (S \ U) = Kolmogorov_quotient X ` S \ U" using Kolmogorov_quotient_in_open [of X U] by auto then show "\V. openin X V \ Kolmogorov_quotient X ` (S \ U) = Kolmogorov_quotient X ` S \ V" using \openin X U\ by blast qed lemma open_map_Kolmogorov_quotient: "open_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" by (metis open_map_Kolmogorov_quotient_gen subtopology_topspace) lemma closed_map_Kolmogorov_quotient_explicit: "closedin X U \ Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \ U" using closedin_subset by (fastforce simp: Kolmogorov_quotient_in_closed) lemma closed_map_Kolmogorov_quotient_gen: "closed_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" using Kolmogorov_quotient_in_closed by (force simp: closed_map_def closedin_subtopology_alt image_iff) lemma closed_map_Kolmogorov_quotient: "closed_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" by (metis closed_map_Kolmogorov_quotient_gen subtopology_topspace) lemma quotient_map_Kolmogorov_quotient_gen: "quotient_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" proof (intro continuous_open_imp_quotient_map) show "continuous_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" by (simp add: continuous_map_Kolmogorov_quotient continuous_map_from_subtopology continuous_map_in_subtopology image_mono) show "open_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" using open_map_Kolmogorov_quotient_gen by blast show "Kolmogorov_quotient X ` topspace (subtopology X S) = topspace (subtopology X (Kolmogorov_quotient X ` S))" by (force simp: Kolmogorov_quotient_in_open) qed lemma quotient_map_Kolmogorov_quotient: "quotient_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" by (metis quotient_map_Kolmogorov_quotient_gen subtopology_topspace) lemma Kolmogorov_quotient_eq: "Kolmogorov_quotient X x = Kolmogorov_quotient X y \ (\U. openin X U \ (x \ U \ y \ U))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Kolmogorov_quotient_in_open) next assume ?rhs then show ?lhs by (simp add: Kolmogorov_quotient_def) qed lemma Kolmogorov_quotient_eq_alt: "Kolmogorov_quotient X x = Kolmogorov_quotient X y \ (\U. closedin X U \ (x \ U \ y \ U))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Kolmogorov_quotient_in_closed) next assume ?rhs then show ?lhs by (smt (verit) Diff_iff Kolmogorov_quotient_eq closedin_topspace in_mono openin_closedin_eq) qed lemma Kolmogorov_quotient_continuous_map: assumes "continuous_map X Y f" "t0_space Y" and x: "x \ topspace X" shows "f (Kolmogorov_quotient X x) = f x" using assms unfolding continuous_map_def t0_space_def by (smt (verit, ccfv_SIG) Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace x mem_Collect_eq) lemma t0_space_Kolmogorov_quotient: "t0_space (subtopology X (Kolmogorov_quotient X ` topspace X))" apply (clarsimp simp: t0_space_def ) by (smt (verit, best) Kolmogorov_quotient_eq imageE image_eqI open_map_Kolmogorov_quotient open_map_def) lemma Kolmogorov_quotient_id: "t0_space X \ x \ topspace X \ Kolmogorov_quotient X x = x" by (metis Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace t0_space_def) lemma Kolmogorov_quotient_idemp: "Kolmogorov_quotient X (Kolmogorov_quotient X x) = Kolmogorov_quotient X x" by (simp add: Kolmogorov_quotient_eq Kolmogorov_quotient_in_open) lemma retraction_maps_Kolmogorov_quotient: "retraction_maps X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X) id" unfolding retraction_maps_def continuous_map_in_subtopology using Kolmogorov_quotient_idemp continuous_map_Kolmogorov_quotient by force lemma retraction_map_Kolmogorov_quotient: "retraction_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" using retraction_map_def retraction_maps_Kolmogorov_quotient by blast lemma retract_of_space_Kolmogorov_quotient_image: "Kolmogorov_quotient X ` topspace X retract_of_space X" proof - have "continuous_map X X (Kolmogorov_quotient X)" by (simp add: continuous_map_Kolmogorov_quotient) then have "Kolmogorov_quotient X ` topspace X \ topspace X" by (simp add: continuous_map_image_subset_topspace) then show ?thesis by (meson retract_of_space_retraction_maps retraction_maps_Kolmogorov_quotient) qed lemma Kolmogorov_quotient_lift_exists: assumes "S \ topspace X" "t0_space Y" and f: "continuous_map (subtopology X S) Y f" obtains g where "continuous_map (subtopology X (image (Kolmogorov_quotient X) S)) Y g" "\x. x \ S \ g(Kolmogorov_quotient X x) = f x" proof - have "\x y. \x \ S; y \ S; Kolmogorov_quotient X x = Kolmogorov_quotient X y\ \ f x = f y" using assms apply (simp add: Kolmogorov_quotient_eq t0_space_def continuous_map_def Int_absorb1 openin_subtopology) by (smt (verit, del_insts) Int_iff mem_Collect_eq) then obtain g where g: "continuous_map (subtopology X (Kolmogorov_quotient X ` S)) Y g" "g ` (topspace X \ Kolmogorov_quotient X ` S) = f ` S" "\x. x \ S \ g (Kolmogorov_quotient X x) = f x" using quotient_map_lift_exists [OF quotient_map_Kolmogorov_quotient_gen [of X S] f] by (metis assms(1) topspace_subtopology topspace_subtopology_subset) show ?thesis proof qed (use g in auto) qed subsection\Closed diagonals and graphs\ lemma Hausdorff_space_closedin_diagonal: "Hausdorff_space X \ closedin (prod_topology X X) ((\x. (x,x)) ` topspace X)" proof - have \
: "((\x. (x, x)) ` topspace X) \ topspace X \ topspace X" by auto show ?thesis apply (simp add: closedin_def openin_prod_topology_alt Hausdorff_space_def disjnt_iff \
) apply (intro all_cong1 imp_cong ex_cong1 conj_cong refl) by (force dest!: openin_subset)+ qed lemma closed_map_diag_eq: "closed_map X (prod_topology X X) (\x. (x,x)) \ Hausdorff_space X" proof - have "section_map X (prod_topology X X) (\x. (x, x))" unfolding section_map_def retraction_maps_def by (smt (verit) continuous_map_fst continuous_map_of_fst continuous_map_on_empty continuous_map_pairwise fst_conv fst_diag_fst snd_diag_fst) then have "embedding_map X (prod_topology X X) (\x. (x, x))" by (rule section_imp_embedding_map) then show ?thesis using Hausdorff_space_closedin_diagonal embedding_imp_closed_map_eq by blast qed lemma closedin_continuous_maps_eq: assumes "Hausdorff_space Y" and f: "continuous_map X Y f" and g: "continuous_map X Y g" shows "closedin X {x \ topspace X. f x = g x}" proof - have \
:"{x \ topspace X. f x = g x} = {x \ topspace X. (f x,g x) \ ((\y.(y,y)) ` topspace Y)}" using f continuous_map_image_subset_topspace by fastforce show ?thesis unfolding \
proof (intro closedin_continuous_map_preimage) show "continuous_map X (prod_topology Y Y) (\x. (f x, g x))" by (simp add: continuous_map_pairedI f g) show "closedin (prod_topology Y Y) ((\y. (y, y)) ` topspace Y)" using Hausdorff_space_closedin_diagonal assms by blast qed qed lemma retract_of_space_imp_closedin: assumes "Hausdorff_space X" and S: "S retract_of_space X" shows "closedin X S" proof - obtain r where r: "continuous_map X (subtopology X S) r" "\x\S. r x = x" using assms by (meson retract_of_space_def) then have \
: "S = {x \ topspace X. r x = x}" using S retract_of_space_imp_subset by (force simp: continuous_map_def) show ?thesis unfolding \
using r continuous_map_into_fulltopology assms by (force intro: closedin_continuous_maps_eq) qed lemma homeomorphic_maps_graph: "homeomorphic_maps X (subtopology (prod_topology X Y) ((\x. (x, f x)) ` (topspace X))) (\x. (x, f x)) fst \ continuous_map X Y f" (is "?lhs=?rhs") proof assume ?lhs then have h: "homeomorphic_map X (subtopology (prod_topology X Y) ((\x. (x, f x)) ` topspace X)) (\x. (x, f x))" by (auto simp: homeomorphic_maps_map) have "f = snd \ (\x. (x, f x))" by force then show ?rhs by (metis (no_types, lifting) h continuous_map_in_subtopology continuous_map_snd_of homeomorphic_eq_everything_map) next assume ?rhs then show ?lhs unfolding homeomorphic_maps_def by (smt (verit, ccfv_threshold) continuous_map_eq continuous_map_subtopology_fst embedding_map_def embedding_map_graph homeomorphic_eq_everything_map image_cong image_iff prod.collapse prod.inject) qed subsection \ KC spaces, those where all compact sets are closed.\ definition kc_space where "kc_space X \ \S. compactin X S \ closedin X S" lemma kc_space_euclidean: "kc_space (euclidean :: 'a::metric_space topology)" by (simp add: compact_imp_closed kc_space_def) lemma kc_space_expansive: "\kc_space X; topspace Y = topspace X; \U. openin X U \ openin Y U\ \ kc_space Y" by (meson compactin_contractive kc_space_def topology_finer_closedin) lemma compactin_imp_closedin_gen: "\kc_space X; compactin X S\ \ closedin X S" using kc_space_def by blast lemma Hausdorff_imp_kc_space: "Hausdorff_space X \ kc_space X" by (simp add: compactin_imp_closedin kc_space_def) lemma kc_imp_t1_space: "kc_space X \ t1_space X" by (simp add: finite_imp_compactin kc_space_def t1_space_closedin_finite) lemma kc_space_subtopology: "kc_space X \ kc_space(subtopology X S)" by (metis closedin_Int_closure_of closure_of_eq compactin_subtopology inf.absorb2 kc_space_def) lemma kc_space_discrete_topology: "kc_space(discrete_topology U)" using Hausdorff_space_discrete_topology compactin_imp_closedin kc_space_def by blast lemma kc_space_continuous_injective_map_preimage: assumes "kc_space Y" "continuous_map X Y f" and injf: "inj_on f (topspace X)" shows "kc_space X" unfolding kc_space_def proof (intro strip) fix S assume S: "compactin X S" have "S = {x \ topspace X. f x \ f ` S}" using S compactin_subset_topspace inj_onD [OF injf] by blast with assms S show "closedin X S" by (metis (no_types, lifting) Collect_cong closedin_continuous_map_preimage compactin_imp_closedin_gen image_compactin) qed lemma kc_space_retraction_map_image: assumes "retraction_map X Y r" "kc_space X" shows "kc_space Y" proof - obtain s where s: "continuous_map X Y r" "continuous_map Y X s" "\x. x \ topspace Y \ r (s x) = x" using assms by (force simp: retraction_map_def retraction_maps_def) then have inj: "inj_on s (topspace Y)" by (metis inj_on_def) show ?thesis unfolding kc_space_def proof (intro strip) fix S assume S: "compactin Y S" have "S = {x \ topspace Y. s x \ s ` S}" using S compactin_subset_topspace inj_onD [OF inj] by blast with assms S show "closedin Y S" by (meson compactin_imp_closedin_gen inj kc_space_continuous_injective_map_preimage s(2)) qed qed lemma homeomorphic_kc_space: "X homeomorphic_space Y \ kc_space X \ kc_space Y" by (meson homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym kc_space_continuous_injective_map_preimage) lemma compact_kc_eq_maximal_compact_space: assumes "compact_space X" shows "kc_space X \ (\Y. topspace Y = topspace X \ (\S. openin X S \ openin Y S) \ compact_space Y \ Y = X)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis closedin_compact_space compactin_contractive kc_space_def topology_eq topology_finer_closedin) next assume R: ?rhs show ?lhs unfolding kc_space_def proof (intro strip) fix S assume S: "compactin X S" define Y where "Y \ topology (arbitrary union_of (finite intersection_of (\A. A = topspace X - S \ openin X A) relative_to (topspace X)))" have "topspace Y = topspace X" by (auto simp: Y_def) have "openin X T \ openin Y T" for T by (simp add: Y_def arbitrary_union_of_inc finite_intersection_of_inc openin_subbase openin_subset relative_to_subset) have "compact_space Y" proof (rule Alexander_subbase_alt) show "\\'. finite \' \ \' \ \ \ topspace X \ \ \'" if \: "\ \ insert (topspace X - S) (Collect (openin X))" and sub: "topspace X \ \\" for \ proof - consider "\ \ Collect (openin X)" | \ where "\ = insert (topspace X - S) \" "\ \ Collect (openin X)" using \ by (metis insert_Diff subset_insert_iff) then show ?thesis proof cases case 1 then show ?thesis by (metis assms compact_space_alt mem_Collect_eq subsetD that(2)) next case 2 then have "S \ \\" using S sub compactin_subset_topspace by blast with 2 obtain \ where "finite \ \ \ \ \ \ S \ \\" using S unfolding compactin_def by (metis Ball_Collect) with 2 show ?thesis by (rule_tac x="insert (topspace X - S) \" in exI) blast qed qed qed (auto simp: Y_def) have "Y = X" using R \\S. openin X S \ openin Y S\ \compact_space Y\ \topspace Y = topspace X\ by blast moreover have "openin Y (topspace X - S)" by (simp add: Y_def arbitrary_union_of_inc finite_intersection_of_inc openin_subbase relative_to_subset) ultimately show "closedin X S" unfolding closedin_def using S compactin_subset_topspace by blast qed qed lemma continuous_imp_closed_map_gen: "\compact_space X; kc_space Y; continuous_map X Y f\ \ closed_map X Y f" by (meson closed_map_def closedin_compact_space compactin_imp_closedin_gen image_compactin) lemma kc_space_compact_subtopologies: "kc_space X \ (\K. compactin X K \ kc_space(subtopology X K))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (auto simp: kc_space_def closedin_closed_subtopology compactin_subtopology) next assume R: ?rhs show ?lhs unfolding kc_space_def proof (intro strip) fix K assume K: "compactin X K" then have "K \ topspace X" by (simp add: compactin_subset_topspace) moreover have "X closure_of K \ K" proof fix x assume x: "x \ X closure_of K" have "kc_space (subtopology X K)" by (simp add: R \compactin X K\) have "compactin X (insert x K)" by (metis K x compactin_Un compactin_sing in_closure_of insert_is_Un) then show "x \ K" by (metis R x K Int_insert_left_if1 closedin_Int_closure_of compact_imp_compactin_subtopology insertCI kc_space_def subset_insertI) qed ultimately show "closedin X K" using closure_of_subset_eq by blast qed qed lemma kc_space_compact_prod_topology: assumes "compact_space X" shows "kc_space(prod_topology X X) \ Hausdorff_space X" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding closed_map_diag_eq [symmetric] proof (intro continuous_imp_closed_map_gen) show "continuous_map X (prod_topology X X) (\x. (x, x))" by (intro continuous_intros) qed (use L assms in auto) next assume ?rhs then show ?lhs by (simp add: Hausdorff_imp_kc_space Hausdorff_space_prod_topology) qed lemma kc_space_prod_topology: "kc_space(prod_topology X X) \ (\K. compactin X K \ Hausdorff_space(subtopology X K))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis compactin_subspace kc_space_compact_prod_topology kc_space_subtopology subtopology_Times) next assume R: ?rhs have "kc_space (subtopology (prod_topology X X) L)" if "compactin (prod_topology X X) L" for L proof - define K where "K \ fst ` L \ snd ` L" have "L \ K \ K" by (force simp: K_def) have "compactin X K" by (metis K_def compactin_Un continuous_map_fst continuous_map_snd image_compactin that) then have "Hausdorff_space (subtopology X K)" by (simp add: R) then have "kc_space (prod_topology (subtopology X K) (subtopology X K))" by (simp add: \compactin X K\ compact_space_subtopology kc_space_compact_prod_topology) then have "kc_space (subtopology (prod_topology (subtopology X K) (subtopology X K)) L)" using kc_space_subtopology by blast then show ?thesis using \L \ K \ K\ subtopology_Times subtopology_subtopology by (metis (no_types, lifting) Sigma_cong inf.absorb_iff2) qed then show ?lhs using kc_space_compact_subtopologies by blast qed lemma kc_space_prod_topology_alt: "kc_space(prod_topology X X) \ kc_space X \ (\K. compactin X K \ Hausdorff_space(subtopology X K))" using Hausdorff_imp_kc_space kc_space_compact_subtopologies kc_space_prod_topology by blast proposition kc_space_prod_topology_left: assumes X: "kc_space X" and Y: "Hausdorff_space Y" shows "kc_space (prod_topology X Y)" unfolding kc_space_def proof (intro strip) fix K assume K: "compactin (prod_topology X Y) K" then have "K \ topspace X \ topspace Y" using compactin_subset_topspace topspace_prod_topology by blast moreover have "\T. openin (prod_topology X Y) T \ (a,b) \ T \ T \ (topspace X \ topspace Y) - K" if ab: "(a,b) \ (topspace X \ topspace Y) - K" for a b proof - have "compactin Y {b}" using that by force moreover have "compactin Y {y \ topspace Y. (a,y) \ K}" proof - have "compactin (prod_topology X Y) (K \ {a} \ topspace Y)" using that compact_Int_closedin [OF K] by (simp add: X closedin_prod_Times_iff compactin_imp_closedin_gen) moreover have "subtopology (prod_topology X Y) (K \ {a} \ topspace Y) homeomorphic_space subtopology Y {y \ topspace Y. (a, y) \ K}" unfolding homeomorphic_space_def homeomorphic_maps_def using that apply (rule_tac x="snd" in exI) apply (rule_tac x="Pair a" in exI) by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology continuous_map_subtopology_snd continuous_map_paired) ultimately show ?thesis by (simp add: compactin_subspace homeomorphic_compact_space) qed moreover have "disjnt {b} {y \ topspace Y. (a,y) \ K}" using ab by force ultimately obtain V U where VU: "openin Y V" "openin Y U" "{b} \ V" "{y \ topspace Y. (a,y) \ K} \ U" "disjnt V U" using Hausdorff_space_compact_separation [OF Y] by blast define V' where "V' \ topspace Y - U" have W: "closedin Y V'" "{y \ topspace Y. (a,y) \ K} \ topspace Y - V'" "disjnt V (topspace Y - V')" using VU by (auto simp: V'_def disjnt_iff) with VU obtain "V \ topspace Y" "V' \ topspace Y" by (meson closedin_subset openin_closedin_eq) then obtain "b \ V" "disjnt {y \ topspace Y. (a,y) \ K} V'" "V \ V'" using VU unfolding disjnt_iff V'_def by force define C where "C \ image fst (K \ {z \ topspace(prod_topology X Y). snd z \ V'})" have "closedin (prod_topology X Y) {z \ topspace (prod_topology X Y). snd z \ V'}" using closedin_continuous_map_preimage \closedin Y V'\ continuous_map_snd by blast then have "compactin X C" unfolding C_def by (meson K compact_Int_closedin continuous_map_fst image_compactin) then have "closedin X C" using assms by (auto simp: kc_space_def) show ?thesis proof (intro exI conjI) show "openin (prod_topology X Y) ((topspace X - C) \ V)" by (simp add: VU \closedin X C\ openin_diff openin_prod_Times_iff) have "a \ C" using VU by (auto simp: C_def V'_def) then show "(a, b) \ (topspace X - C) \ V" using \a \ C\ \b \ V\ that by blast show "(topspace X - C) \ V \ topspace X \ topspace Y - K" using \V \ V'\ \V \ topspace Y\ apply (simp add: C_def ) by (smt (verit, ccfv_threshold) DiffE DiffI IntI SigmaE SigmaI image_eqI mem_Collect_eq prod.sel(1) snd_conv subset_iff) qed qed ultimately show "closedin (prod_topology X Y) K" by (metis surj_pair closedin_def openin_subopen topspace_prod_topology) qed lemma kc_space_prod_topology_right: "\Hausdorff_space X; kc_space Y\ \ kc_space (prod_topology X Y)" using kc_space_prod_topology_left homeomorphic_kc_space homeomorphic_space_prod_topology_swap by blast subsection \Regular spaces\ text \Regular spaces are *not* a priori assumed to be Hausdorff or $T_1$\ definition regular_space where "regular_space X \ \C a. closedin X C \ a \ topspace X - C \ (\U V. openin X U \ openin X V \ a \ U \ C \ V \ disjnt U V)" lemma homeomorphic_regular_space_aux: assumes hom: "X homeomorphic_space Y" and X: "regular_space X" shows "regular_space Y" proof - obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g" and fg: "(\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" using assms X homeomorphic_maps_map homeomorphic_space_def by fastforce show ?thesis unfolding regular_space_def proof clarify fix C a assume Y: "closedin Y C" "a \ topspace Y" and "a \ C" then obtain "closedin X (g ` C)" "g a \ topspace X" "g a \ g ` C" using \closedin Y C\ hmg homeomorphic_map_closedness_eq by (smt (verit, ccfv_SIG) fg homeomorphic_imp_surjective_map image_iff in_mono) then obtain S T where ST: "openin X S" "openin X T" "g a \ S" "g`C \ T" "disjnt S T" using X unfolding regular_space_def by (metis DiffI) then have "openin Y (f`S)" "openin Y (f`T)" by (meson hmf homeomorphic_map_openness_eq)+ moreover have "a \ f`S \ C \ f`T" using ST by (smt (verit, best) Y closedin_subset fg image_eqI subset_iff) moreover have "disjnt (f`S) (f`T)" using ST by (smt (verit, ccfv_SIG) disjnt_iff fg image_iff openin_subset subsetD) ultimately show "\U V. openin Y U \ openin Y V \ a \ U \ C \ V \ disjnt U V" by metis qed qed lemma homeomorphic_regular_space: "X homeomorphic_space Y \ (regular_space X \ regular_space Y)" by (meson homeomorphic_regular_space_aux homeomorphic_space_sym) lemma regular_space: "regular_space X \ (\C a. closedin X C \ a \ topspace X - C \ (\U. openin X U \ a \ U \ disjnt C (X closure_of U)))" unfolding regular_space_def proof (intro all_cong1 imp_cong refl ex_cong1) fix C a U assume C: "closedin X C \ a \ topspace X - C" show "(\V. openin X U \ openin X V \ a \ U \ C \ V \ disjnt U V) \ (openin X U \ a \ U \ disjnt C (X closure_of U))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (smt (verit, best) disjnt_iff in_closure_of subsetD) next assume R: ?rhs then have "disjnt U (topspace X - X closure_of U)" by (meson DiffD2 closure_of_subset disjnt_iff openin_subset subsetD) moreover have "C \ topspace X - X closure_of U" by (meson C DiffI R closedin_subset disjnt_iff subset_eq) ultimately show ?lhs using R by (rule_tac x="topspace X - X closure_of U" in exI) auto qed qed lemma neighbourhood_base_of_closedin: "neighbourhood_base_of (closedin X) X \ regular_space X" (is "?lhs=?rhs") proof - have "?lhs \ (\W x. openin X W \ x \ W \ (\U. openin X U \ (\V. closedin X V \ x \ U \ U \ V \ V \ W)))" by (simp add: neighbourhood_base_of) also have "\ \ (\W x. closedin X W \ x \ topspace X - W \ (\U. openin X U \ (\V. closedin X V \ x \ U \ U \ V \ V \ topspace X - W)))" by (smt (verit) Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) also have "\ \ ?rhs" proof - have \
: "(\V. closedin X V \ x \ U \ U \ V \ V \ topspace X - W) \ (\V. openin X V \ x \ U \ W \ V \ disjnt U V)" (is "?lhs=?rhs") if "openin X U" "closedin X W" "x \ topspace X" "x \ W" for W U x proof assume ?lhs with \closedin X W\ show ?rhs unfolding closedin_def by (smt (verit) Diff_mono disjnt_Diff1 double_diff subset_eq) next assume ?rhs with \openin X U\ show ?lhs unfolding openin_closedin_eq disjnt_def by (smt (verit) Diff_Diff_Int Diff_disjoint Diff_eq_empty_iff Int_Diff inf.orderE) qed show ?thesis unfolding regular_space_def by (intro all_cong1 ex_cong1 imp_cong refl) (metis \
DiffE) qed finally show ?thesis . qed lemma regular_space_discrete_topology: "regular_space (discrete_topology S)" using neighbourhood_base_of_closedin neighbourhood_base_of_discrete_topology by fastforce lemma regular_space_subtopology: "regular_space X \ regular_space (subtopology X S)" unfolding regular_space_def openin_subtopology_alt closedin_subtopology_alt disjnt_iff by clarsimp (smt (verit, best) inf.orderE inf_le1 le_inf_iff) lemma regular_space_retraction_map_image: "\retraction_map X Y r; regular_space X\ \ regular_space Y" using hereditary_imp_retractive_property homeomorphic_regular_space regular_space_subtopology by blast lemma regular_t0_imp_Hausdorff_space: "\regular_space X; t0_space X\ \ Hausdorff_space X" apply (clarsimp simp: regular_space_def t0_space Hausdorff_space_def) by (metis disjnt_sym subsetD) lemma regular_t0_eq_Hausdorff_space: "regular_space X \ (t0_space X \ Hausdorff_space X)" using Hausdorff_imp_t0_space regular_t0_imp_Hausdorff_space by blast lemma regular_t1_imp_Hausdorff_space: "\regular_space X; t1_space X\ \ Hausdorff_space X" by (simp add: regular_t0_imp_Hausdorff_space t1_imp_t0_space) lemma regular_t1_eq_Hausdorff_space: "regular_space X \ t1_space X \ Hausdorff_space X" using regular_t0_imp_Hausdorff_space t1_imp_t0_space t1_or_Hausdorff_space by blast lemma compact_Hausdorff_imp_regular_space: assumes "compact_space X" "Hausdorff_space X" shows "regular_space X" unfolding regular_space_def proof clarify fix S a assume "closedin X S" and "a \ topspace X" and "a \ S" then show "\U V. openin X U \ openin X V \ a \ U \ S \ V \ disjnt U V" using assms unfolding Hausdorff_space_compact_sets by (metis closedin_compact_space compactin_sing disjnt_empty1 insert_subset disjnt_insert1) qed lemma regular_space_topspace_empty: "topspace X = {} \ regular_space X" by (simp add: Hausdorff_space_topspace_empty compact_Hausdorff_imp_regular_space compact_space_topspace_empty) lemma neighbourhood_base_of_closed_Hausdorff_space: "regular_space X \ Hausdorff_space X \ neighbourhood_base_of (\C. closedin X C \ Hausdorff_space(subtopology X C)) X" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (simp add: Hausdorff_space_subtopology neighbourhood_base_of_closedin) next assume ?rhs then show ?lhs by (metis (mono_tags, lifting) Hausdorff_space_closed_neighbourhood neighbourhood_base_of neighbourhood_base_of_closedin openin_topspace) qed lemma locally_compact_imp_kc_eq_Hausdorff_space: "neighbourhood_base_of (compactin X) X \ kc_space X \ Hausdorff_space X" by (metis Hausdorff_imp_kc_space kc_imp_t1_space kc_space_def neighbourhood_base_of_closedin neighbourhood_base_of_mono regular_t1_imp_Hausdorff_space) lemma regular_space_compact_closed_separation: assumes X: "regular_space X" and S: "compactin X S" and T: "closedin X T" and "disjnt S T" shows "\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V" proof (cases "S={}") case True then show ?thesis by (meson T closedin_def disjnt_empty1 empty_subsetI openin_empty openin_topspace) next case False then have "\U V. x \ S \ openin X U \ openin X V \ x \ U \ T \ V \ disjnt U V" for x using assms unfolding regular_space_def by (smt (verit) Diff_iff compactin_subset_topspace disjnt_iff subsetD) then obtain U V where UV: "\x. x \ S \ openin X (U x) \ openin X (V x) \ x \ (U x) \ T \ (V x) \ disjnt (U x) (V x)" by metis then obtain \ where "finite \" "\ \ U ` S" "S \ \ \" using S unfolding compactin_def by (smt (verit) UN_iff image_iff subsetI) then obtain K where "finite K" "K \ S" and K: "S \ \(U ` K)" by (metis finite_subset_image) show ?thesis proof (intro exI conjI) show "openin X (\(U ` K))" using \K \ S\ UV by blast show "openin X (\(V ` K))" using False K UV \K \ S\ \finite K\ by blast show "S \ \(U ` K)" by (simp add: K) show "T \ \(V ` K)" using UV \K \ S\ by blast show "disjnt (\(U ` K)) (\(V ` K))" by (smt (verit) Inter_iff UN_E UV \K \ S\ disjnt_iff image_eqI subset_iff) qed qed lemma regular_space_compact_closed_sets: "regular_space X \ (\S T. compactin X S \ closedin X T \ disjnt S T \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using regular_space_compact_closed_separation by fastforce next assume R: ?rhs show ?lhs unfolding regular_space proof clarify fix S x assume "closedin X S" and "x \ topspace X" and "x \ S" then obtain U V where "openin X U \ openin X V \ {x} \ U \ S \ V \ disjnt U V" by (metis R compactin_sing disjnt_empty1 disjnt_insert1) then show "\U. openin X U \ x \ U \ disjnt S (X closure_of U)" by (smt (verit, best) disjnt_iff in_closure_of insert_subset subsetD) qed qed lemma regular_space_prod_topology: "regular_space (prod_topology X Y) \ topspace X = {} \ topspace Y = {} \ regular_space X \ regular_space Y" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis regular_space_retraction_map_image retraction_map_fst retraction_map_snd) next assume R: ?rhs show ?lhs proof (cases "topspace X = {} \ topspace Y = {}") case True then show ?thesis by (simp add: regular_space_topspace_empty) next case False then have "regular_space X" "regular_space Y" using R by auto show ?thesis unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of proof clarify fix W x y assume W: "openin (prod_topology X Y) W" and "(x,y) \ W" then obtain U V where U: "openin X U" "x \ U" and V: "openin Y V" "y \ V" and "U \ V \ W" by (metis openin_prod_topology_alt) obtain D1 C1 where 1: "openin X D1" "closedin X C1" "x \ D1" "D1 \ C1" "C1 \ U" by (metis \regular_space X\ U neighbourhood_base_of neighbourhood_base_of_closedin) obtain D2 C2 where 2: "openin Y D2" "closedin Y C2" "y \ D2" "D2 \ C2" "C2 \ V" by (metis \regular_space Y\ V neighbourhood_base_of neighbourhood_base_of_closedin) show "\U V. openin (prod_topology X Y) U \ closedin (prod_topology X Y) V \ (x,y) \ U \ U \ V \ V \ W" proof (intro conjI exI) show "openin (prod_topology X Y) (D1 \ D2)" by (simp add: 1 2 openin_prod_Times_iff) show "closedin (prod_topology X Y) (C1 \ C2)" by (simp add: 1 2 closedin_prod_Times_iff) qed (use 1 2 \U \ V \ W\ in auto) qed qed qed lemma regular_space_product_topology: "regular_space (product_topology X I) \ topspace (product_topology X I) = {} \ (\i \ I. regular_space (X i))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (meson regular_space_retraction_map_image retraction_map_product_projection) next assume R: ?rhs show ?lhs proof (cases "topspace(product_topology X I) = {}") case True then show ?thesis by (simp add: regular_space_topspace_empty) next case False then obtain x where x: "x \ topspace (product_topology X I)" by blast define \ where "\ \ {Pi\<^sub>E I U |U. finite {i \ I. U i \ topspace (X i)} \ (\i\I. openin (X i) (U i))}" have oo: "openin (product_topology X I) = arbitrary union_of (\W. W \ \)" by (simp add: \_def openin_product_topology product_topology_base_alt) have "\U V. openin (product_topology X I) U \ closedin (product_topology X I) V \ x \ U \ U \ V \ V \ Pi\<^sub>E I W" if fin: "finite {i \ I. W i \ topspace (X i)}" and opeW: "\k. k \ I \ openin (X k) (W k)" and x: "x \ PiE I W" for W x proof - have "\i. i \ I \ \U V. openin (X i) U \ closedin (X i) V \ x i \ U \ U \ V \ V \ W i" by (metis False PiE_iff R neighbourhood_base_of neighbourhood_base_of_closedin opeW x) then obtain U C where UC: "\i. i \ I \ openin (X i) (U i) \ closedin (X i) (C i) \ x i \ U i \ U i \ C i \ C i \ W i" by metis define PI where "PI \ \V. PiE I (\i. if W i = topspace(X i) then topspace(X i) else V i)" show ?thesis proof (intro conjI exI) have "\i\I. W i \ topspace (X i) \ openin (X i) (U i)" using UC by force with fin show "openin (product_topology X I) (PI U)" by (simp add: Collect_mono_iff PI_def openin_PiE_gen rev_finite_subset) show "closedin (product_topology X I) (PI C)" by (simp add: UC closedin_product_topology PI_def) show "x \ PI U" using UC x by (fastforce simp: PI_def) show "PI U \ PI C" by (smt (verit) UC Orderings.order_eq_iff PiE_mono PI_def) show "PI C \ Pi\<^sub>E I W" by (simp add: UC PI_def subset_PiE) qed qed then have "neighbourhood_base_of (closedin (product_topology X I)) (product_topology X I)" unfolding neighbourhood_base_of_topology_base [OF oo] by (force simp: \_def) then show ?thesis by (simp add: neighbourhood_base_of_closedin) qed qed lemma closed_map_paired_gen_aux: assumes "regular_space Y" and f: "closed_map Z X f" and g: "closed_map Z Y g" and clo: "\y. y \ topspace X \ closedin Z {x \ topspace Z. f x = y}" and contg: "continuous_map Z Y g" shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" unfolding closed_map_def proof (intro strip) fix C assume "closedin Z C" then have "C \ topspace Z" by (simp add: closedin_subset) have "f ` topspace Z \ topspace X" "g ` topspace Z \ topspace Y" by (simp_all add: assms closed_map_imp_subset_topspace) show "closedin (prod_topology X Y) ((\x. (f x, g x)) ` C)" unfolding closedin_def topspace_prod_topology proof (intro conjI) have "closedin Y (g ` C)" using \closedin Z C\ assms(3) closed_map_def by blast with assms show "(\x. (f x, g x)) ` C \ topspace X \ topspace Y" using \C \ topspace Z\ \f ` topspace Z \ topspace X\ continuous_map_closedin subsetD by fastforce have *: "\T. openin (prod_topology X Y) T \ (y1,y2) \ T \ T \ topspace X \ topspace Y - (\x. (f x, g x)) ` C" if "(y1,y2) \ (\x. (f x, g x)) ` C" and y1: "y1 \ topspace X" and y2: "y2 \ topspace Y" for y1 y2 proof - define A where "A \ topspace Z - (C \ {x \ topspace Z. f x = y1})" have A: "openin Z A" "{x \ topspace Z. g x = y2} \ A" using that \closedin Z C\ clo that(2) by (auto simp: A_def) obtain V0 where "openin Y V0 \ y2 \ V0" and UA: "{x \ topspace Z. g x \ V0} \ A" using g A y2 unfolding closed_map_fibre_neighbourhood by blast then obtain V V' where VV: "openin Y V \ closedin Y V' \ y2 \ V \ V \ V'" and "V' \ V0" by (metis (no_types, lifting) \regular_space Y\ neighbourhood_base_of neighbourhood_base_of_closedin) with UA have subA: "{x \ topspace Z. g x \ V'} \ A" by blast show ?thesis proof - define B where "B \ topspace Z - (C \ {x \ topspace Z. g x \ V'})" have "openin Z B" using VV \closedin Z C\ contg by (fastforce simp: B_def continuous_map_closedin) have "{x \ topspace Z. f x = y1} \ B" using A_def subA by (auto simp: A_def B_def) then obtain U where "openin X U" "y1 \ U" and subB: "{x \ topspace Z. f x \ U} \ B" using \openin Z B\ y1 f unfolding closed_map_fibre_neighbourhood by meson show ?thesis proof (intro conjI exI) show "openin (prod_topology X Y) (U \ V)" by (metis VV \openin X U\ openin_prod_Times_iff) show "(y1, y2) \ U \ V" by (simp add: VV \y1 \ U\) show "U \ V \ topspace X \ topspace Y - (\x. (f x, g x)) ` C" using VV \C \ topspace Z\ \openin X U\ subB by (force simp: image_iff B_def subset_iff dest: openin_subset) qed qed qed then show "openin (prod_topology X Y) (topspace X \ topspace Y - (\x. (f x, g x)) ` C)" by (smt (verit, ccfv_threshold) Diff_iff SigmaE openin_subopen) qed qed lemma closed_map_paired_gen: assumes f: "closed_map Z X f" and g: "closed_map Z Y g" and D: "(regular_space X \ continuous_map Z X f \ (\z \ topspace Y. closedin Z {x \ topspace Z. g x = z}) \ regular_space Y \ continuous_map Z Y g \ (\y \ topspace X. closedin Z {x \ topspace Z. f x = y}))" (is "?RSX \ ?RSY") shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" using D proof assume RSX: ?RSX have eq: "(\x. (f x, g x)) = (\(x,y). (y,x)) \ (\x. (g x, f x))" by auto show ?thesis unfolding eq proof (rule closed_map_compose) show "closed_map Z (prod_topology Y X) (\x. (g x, f x))" using RSX closed_map_paired_gen_aux f g by fastforce show "closed_map (prod_topology Y X) (prod_topology X Y) (\(x, y). (y, x))" using homeomorphic_imp_closed_map homeomorphic_map_swap by blast qed qed (blast intro: assms closed_map_paired_gen_aux) lemma closed_map_paired: assumes "closed_map Z X f" and contf: "continuous_map Z X f" "closed_map Z Y g" and contg: "continuous_map Z Y g" and D: "t1_space X \ regular_space Y \ regular_space X \ t1_space Y" shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" proof (rule closed_map_paired_gen) show "regular_space X \ continuous_map Z X f \ (\z\topspace Y. closedin Z {x \ topspace Z. g x = z}) \ regular_space Y \ continuous_map Z Y g \ (\y\topspace X. closedin Z {x \ topspace Z. f x = y})" using D contf contg by (smt (verit, del_insts) Collect_cong closedin_continuous_map_preimage t1_space_closedin_singleton singleton_iff) qed (use assms in auto) lemma closed_map_pairwise: assumes "closed_map Z X (fst \ f)" "continuous_map Z X (fst \ f)" "closed_map Z Y (snd \ f)" "continuous_map Z Y (snd \ f)" "t1_space X \ regular_space Y \ regular_space X \ t1_space Y" shows "closed_map Z (prod_topology X Y) f" proof - have "closed_map Z (prod_topology X Y) (\a. ((fst \ f) a, (snd \ f) a))" using assms closed_map_paired by blast then show ?thesis by auto qed lemma tube_lemma_right: assumes W: "openin (prod_topology X Y) W" and C: "compactin Y C" and x: "x \ topspace X" and subW: "{x} \ C \ W" shows "\U V. openin X U \ openin Y V \ x \ U \ C \ V \ U \ V \ W" proof (cases "C = {}") case True with x show ?thesis by auto next case False have "\U V. openin X U \ openin Y V \ x \ U \ y \ V \ U \ V \ W" if "y \ C" for y using W openin_prod_topology_alt subW subsetD that by fastforce then obtain U V where UV: "\y. y \ C \ openin X (U y) \ openin Y (V y) \ x \ U y \ y \ V y \ U y \ V y \ W" by metis then obtain D where D: "finite D" "D \ C" "C \ \ (V ` D)" using compactinD [OF C, of "V`C"] by (smt (verit) UN_I finite_subset_image imageE subsetI) show ?thesis proof (intro exI conjI) show "openin X (\ (U ` D))" "openin Y (\ (V ` D))" using D False UV by blast+ show "x \ \ (U ` D)" "C \ \ (V ` D)" "\ (U ` D) \ \ (V ` D) \ W" using D UV by force+ qed qed lemma closed_map_fst: assumes "compact_space Y" shows "closed_map (prod_topology X Y) X fst" proof - have *: "{x \ topspace X \ topspace Y. fst x \ U} = U \ topspace Y" if "U \ topspace X" for U using that by force have **: "\U y. \openin (prod_topology X Y) U; y \ topspace X; {x \ topspace X \ topspace Y. fst x = y} \ U\ \ \V. openin X V \ y \ V \ V \ topspace Y \ U" using tube_lemma_right[of X Y _ "topspace Y"] assms compact_space_def by force show ?thesis unfolding closed_map_fibre_neighbourhood by (force simp: * openin_subset cong: conj_cong intro: **) qed lemma closed_map_snd: assumes "compact_space X" shows "closed_map (prod_topology X Y) Y snd" proof - have "snd = fst o prod.swap" by force moreover have "closed_map (prod_topology X Y) Y (fst o prod.swap)" proof (rule closed_map_compose) show "closed_map (prod_topology X Y) (prod_topology Y X) prod.swap" by (metis (no_types, lifting) homeomorphic_imp_closed_map homeomorphic_map_eq homeomorphic_map_swap prod.swap_def split_beta) show "closed_map (prod_topology Y X) Y fst" by (simp add: closed_map_fst assms) qed ultimately show ?thesis by metis qed lemma closed_map_paired_closed_map_right: "\closed_map X Y f; regular_space X; \y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}\ \ closed_map X (prod_topology X Y) (\x. (x, f x))" by (rule closed_map_paired_gen [OF closed_map_id, unfolded id_def]) auto lemma closed_map_paired_closed_map_left: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "closed_map X (prod_topology Y X) (\x. (f x, x))" proof - have eq: "(\x. (f x, x)) = (\(x,y). (y,x)) \ (\x. (x, f x))" by auto show ?thesis unfolding eq proof (rule closed_map_compose) show "closed_map X (prod_topology X Y) (\x. (x, f x))" by (simp add: assms closed_map_paired_closed_map_right) show "closed_map (prod_topology X Y) (prod_topology Y X) (\(x, y). (y, x))" using homeomorphic_imp_closed_map homeomorphic_map_swap by blast qed qed lemma closed_map_imp_closed_graph: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "closedin (prod_topology X Y) ((\x. (x, f x)) ` topspace X)" using assms closed_map_def closed_map_paired_closed_map_right by blast lemma proper_map_paired_closed_map_right: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "proper_map X (prod_topology X Y) (\x. (x, f x))" by (simp add: assms closed_injective_imp_proper_map inj_on_def closed_map_paired_closed_map_right) lemma proper_map_paired_closed_map_left: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "proper_map X (prod_topology Y X) (\x. (f x, x))" by (simp add: assms closed_injective_imp_proper_map inj_on_def closed_map_paired_closed_map_left) proposition regular_space_continuous_proper_map_image: assumes "regular_space X" and contf: "continuous_map X Y f" and pmapf: "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "regular_space Y" unfolding regular_space_def proof clarify fix C y assume "closedin Y C" and "y \ topspace Y" and "y \ C" have "closed_map X Y f" "(\y \ topspace Y. compactin X {x \ topspace X. f x = y})" using pmapf proper_map_def by force+ moreover have "closedin X {z \ topspace X. f z \ C}" using \closedin Y C\ contf continuous_map_closedin by fastforce moreover have "disjnt {z \ topspace X. f z = y} {z \ topspace X. f z \ C}" using \y \ C\ disjnt_iff by blast ultimately obtain U V where UV: "openin X U" "openin X V" "{z \ topspace X. f z = y} \ U" "{z \ topspace X. f z \ C} \ V" and dUV: "disjnt U V" using \y \ topspace Y\ \regular_space X\ unfolding regular_space_compact_closed_sets by meson have *: "\U T. openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ (\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U)" using \closed_map X Y f\ unfolding closed_map_preimage_neighbourhood by blast obtain V1 where V1: "openin Y V1 \ y \ V1" and sub1: "{x \ topspace X. f x \ V1} \ U" using * [of U "{y}"] UV \y \ topspace Y\ by auto moreover obtain V2 where "openin Y V2 \ C \ V2" and sub2: "{x \ topspace X. f x \ V2} \ V" by (smt (verit, ccfv_SIG) * UV \closedin Y C\ closedin_subset mem_Collect_eq subset_iff) moreover have "disjnt V1 V2" proof - have "\x. \\x. x \ U \ x \ V; x \ V1; x \ V2\ \ False" by (smt (verit) V1 fim image_iff mem_Collect_eq openin_subset sub1 sub2 subsetD) with dUV show ?thesis by (auto simp: disjnt_iff) qed ultimately show "\U V. openin Y U \ openin Y V \ y \ U \ C \ V \ disjnt U V" by meson qed lemma regular_space_perfect_map_image: "\regular_space X; perfect_map X Y f\ \ regular_space Y" by (meson perfect_map_def regular_space_continuous_proper_map_image) proposition regular_space_perfect_map_image_eq: assumes "Hausdorff_space X" and perf: "perfect_map X Y f" shows "regular_space X \ regular_space Y" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using perf regular_space_perfect_map_image by blast next assume R: ?rhs have "continuous_map X Y f" "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" using perf by (auto simp: perfect_map_def) then have "closed_map X Y f" and preYf: "(\y \ topspace Y. compactin X {x \ topspace X. f x = y})" by (simp_all add: proper_map_def) show ?lhs unfolding regular_space_def proof clarify fix C x assume "closedin X C" and "x \ topspace X" and "x \ C" obtain U1 U2 where "openin X U1" "openin X U2" "{x} \ U1" and "disjnt U1 U2" and subV: "C \ {z \ topspace X. f z = f x} \ U2" proof (rule Hausdorff_space_compact_separation [of X "{x}" "C \ {z \ topspace X. f z = f x}", OF \Hausdorff_space X\]) show "compactin X {x}" by (simp add: \x \ topspace X\) show "compactin X (C \ {z \ topspace X. f z = f x})" using \closedin X C\ fim \x \ topspace X\ closed_Int_compactin preYf by fastforce show "disjnt {x} (C \ {z \ topspace X. f z = f x})" using \x \ C\ by force qed have "closedin Y (f ` (C - U2))" using \closed_map X Y f\ \closedin X C\ \openin X U2\ closed_map_def by blast moreover have "f x \ topspace Y - f ` (C - U2)" using \closedin X C\ \continuous_map X Y f\ \x \ topspace X\ closedin_subset continuous_map_def subV by fastforce ultimately obtain V1 V2 where VV: "openin Y V1" "openin Y V2" "f x \ V1" and subV2: "f ` (C - U2) \ V2" and "disjnt V1 V2" by (meson R regular_space_def) show "\U U'. openin X U \ openin X U' \ x \ U \ C \ U' \ disjnt U U'" proof (intro exI conjI) show "openin X (U1 \ {x \ topspace X. f x \ V1})" using VV(1) \continuous_map X Y f\ \openin X U1\ continuous_map by fastforce show "openin X (U2 \ {x \ topspace X. f x \ V2})" using VV(2) \continuous_map X Y f\ \openin X U2\ continuous_map by fastforce show "x \ U1 \ {x \ topspace X. f x \ V1}" using VV(3) \x \ topspace X\ \{x} \ U1\ by auto show "C \ U2 \ {x \ topspace X. f x \ V2}" using \closedin X C\ closedin_subset subV2 by auto show "disjnt (U1 \ {x \ topspace X. f x \ V1}) (U2 \ {x \ topspace X. f x \ V2})" using \disjnt U1 U2\ \disjnt V1 V2\ by (auto simp: disjnt_iff) qed qed qed subsection\Locally compact spaces\ definition locally_compact_space where "locally_compact_space X \ \x \ topspace X. \U K. openin X U \ compactin X K \ x \ U \ U \ K" lemma homeomorphic_locally_compact_spaceD: assumes X: "locally_compact_space X" and "X homeomorphic_space Y" shows "locally_compact_space Y" proof - obtain f where hmf: "homeomorphic_map X Y f" using assms homeomorphic_space by blast then have eq: "topspace Y = f ` (topspace X)" by (simp add: homeomorphic_imp_surjective_map) have "\V K. openin Y V \ compactin Y K \ f x \ V \ V \ K" if "x \ topspace X" "openin X U" "compactin X K" "x \ U" "U \ K" for x U K using that by (meson hmf homeomorphic_map_compactness_eq homeomorphic_map_openness_eq image_mono image_eqI) with X show ?thesis by (smt (verit) eq image_iff locally_compact_space_def) qed lemma homeomorphic_locally_compact_space: assumes "X homeomorphic_space Y" shows "locally_compact_space X \ locally_compact_space Y" by (meson assms homeomorphic_locally_compact_spaceD homeomorphic_space_sym) lemma locally_compact_space_retraction_map_image: assumes "retraction_map X Y r" and X: "locally_compact_space X" shows "locally_compact_space Y" proof - obtain s where s: "retraction_maps X Y r s" using assms retraction_map_def by blast obtain T where "T retract_of_space X" and Teq: "T = s ` topspace Y" using retraction_maps_section_image1 s by blast then obtain r where r: "continuous_map X (subtopology X T) r" "\x\T. r x = x" by (meson retract_of_space_def) have "locally_compact_space (subtopology X T)" unfolding locally_compact_space_def openin_subtopology_alt proof clarsimp fix x assume "x \ topspace X" "x \ T" obtain U K where UK: "openin X U \ compactin X K \ x \ U \ U \ K" by (meson X \x \ topspace X\ locally_compact_space_def) then have "compactin (subtopology X T) (r ` K) \ T \ U \ r ` K" by (smt (verit) IntD1 image_compactin image_iff inf_le2 r subset_iff) then show "\U. openin X U \ (\K. compactin (subtopology X T) K \ x \ U \ T \ U \ K)" using UK by auto qed with Teq show ?thesis using homeomorphic_locally_compact_space retraction_maps_section_image2 s by blast qed lemma compact_imp_locally_compact_space: "compact_space X \ locally_compact_space X" using compact_space_def locally_compact_space_def by blast lemma neighbourhood_base_imp_locally_compact_space: "neighbourhood_base_of (compactin X) X \ locally_compact_space X" by (metis locally_compact_space_def neighbourhood_base_of openin_topspace) lemma locally_compact_imp_neighbourhood_base: assumes loc: "locally_compact_space X" and reg: "regular_space X" shows "neighbourhood_base_of (compactin X) X" unfolding neighbourhood_base_of proof clarify fix W x assume "openin X W" and "x \ W" then obtain U K where "openin X U" "compactin X K" "x \ U" "U \ K" by (metis loc locally_compact_space_def openin_subset subsetD) moreover have "openin X (U \ W) \ x \ U \ W" using \openin X W\ \x \ W\ \openin X U\ \x \ U\ by blast then have "\u' v. openin X u' \ closedin X v \ x \ u' \ u' \ v \ v \ U \ v \ W" using reg by (metis le_infE neighbourhood_base_of neighbourhood_base_of_closedin) then show "\U V. openin X U \ compactin X V \ x \ U \ U \ V \ V \ W" by (meson \U \ K\ \compactin X K\ closed_compactin subset_trans) qed lemma Hausdorff_regular: "\Hausdorff_space X; neighbourhood_base_of (compactin X) X\ \ regular_space X" by (metis compactin_imp_closedin neighbourhood_base_of_closedin neighbourhood_base_of_mono) lemma locally_compact_Hausdorff_imp_regular_space: assumes loc: "locally_compact_space X" and "Hausdorff_space X" shows "regular_space X" unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of proof clarify fix W x assume "openin X W" and "x \ W" then have "x \ topspace X" using openin_subset by blast then obtain U K where "openin X U" "compactin X K" and UK: "x \ U" "U \ K" by (meson loc locally_compact_space_def) with \Hausdorff_space X\ have "regular_space (subtopology X K)" using Hausdorff_space_subtopology compact_Hausdorff_imp_regular_space compact_space_subtopology by blast then have "\U' V'. openin (subtopology X K) U' \ closedin (subtopology X K) V' \ x \ U' \ U' \ V' \ V' \ K \ W" unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of by (meson IntI \U \ K\ \openin X W\ \x \ U\ \x \ W\ openin_subtopology_Int2 subsetD) then obtain V C where "openin X V" "closedin X C" and VC: "x \ K \ V" "K \ V \ K \ C" "K \ C \ K \ W" by (metis Int_commute closedin_subtopology openin_subtopology) show "\U V. openin X U \ closedin X V \ x \ U \ U \ V \ V \ W" proof (intro conjI exI) show "openin X (U \ V)" using \openin X U\ \openin X V\ by blast show "closedin X (K \ C)" using \closedin X C\ \compactin X K\ compactin_imp_closedin \Hausdorff_space X\ by blast qed (use UK VC in auto) qed lemma locally_compact_space_neighbourhood_base: "Hausdorff_space X \ regular_space X \ locally_compact_space X \ neighbourhood_base_of (compactin X) X" by (metis locally_compact_imp_neighbourhood_base locally_compact_Hausdorff_imp_regular_space neighbourhood_base_imp_locally_compact_space) lemma locally_compact_Hausdorff_or_regular: "locally_compact_space X \ (Hausdorff_space X \ regular_space X) \ locally_compact_space X \ regular_space X" using locally_compact_Hausdorff_imp_regular_space by blast lemma locally_compact_space_compact_closedin: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ (\x \ topspace X. \U K. openin X U \ compactin X K \ closedin X K \ x \ U \ U \ K)" using locally_compact_Hausdorff_or_regular unfolding locally_compact_space_def by (metis assms closed_compactin inf.absorb_iff2 le_infE neighbourhood_base_of neighbourhood_base_of_closedin) lemma locally_compact_space_compact_closure_of: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ (\x \ topspace X. \U. openin X U \ compactin X (X closure_of U) \ x \ U)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis assms closed_compactin closedin_closure_of closure_of_eq closure_of_mono locally_compact_space_compact_closedin) next assume ?rhs then show ?lhs by (meson closure_of_subset locally_compact_space_def openin_subset) qed lemma locally_compact_space_neighbourhood_base_closedin: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ neighbourhood_base_of (\C. compactin X C \ closedin X C) X" (is "?lhs=?rhs") proof assume L: ?lhs then have "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space by blast with L have "neighbourhood_base_of (compactin X) X" by (simp add: locally_compact_imp_neighbourhood_base) with \regular_space X\ show ?rhs by (smt (verit, ccfv_threshold) closed_compactin neighbourhood_base_of subset_trans neighbourhood_base_of_closedin) next assume ?rhs then show ?lhs using neighbourhood_base_imp_locally_compact_space neighbourhood_base_of_mono by blast qed lemma locally_compact_space_neighbourhood_base_closure_of: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ neighbourhood_base_of (\T. compactin X (X closure_of T)) X" (is "?lhs=?rhs") proof assume L: ?lhs then have "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space by blast with L have "neighbourhood_base_of (\A. compactin X A \ closedin X A) X" using locally_compact_space_neighbourhood_base_closedin by blast then show ?rhs by (simp add: closure_of_closedin neighbourhood_base_of_mono) next assume ?rhs then show ?lhs unfolding locally_compact_space_def neighbourhood_base_of by (meson closure_of_subset openin_topspace subset_trans) qed lemma locally_compact_space_neighbourhood_base_open_closure_of: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ neighbourhood_base_of (\U. openin X U \ compactin X (X closure_of U)) X" (is "?lhs=?rhs") proof assume L: ?lhs then have "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space by blast then have "neighbourhood_base_of (\T. compactin X (X closure_of T)) X" using L locally_compact_space_neighbourhood_base_closure_of by auto with L show ?rhs unfolding neighbourhood_base_of by (meson closed_compactin closure_of_closure_of closure_of_eq closure_of_mono subset_trans) next assume ?rhs then show ?lhs unfolding locally_compact_space_def neighbourhood_base_of by (meson closure_of_subset openin_topspace subset_trans) qed lemma locally_compact_space_compact_closed_compact: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ (\K. compactin X K \ (\U L. openin X U \ compactin X L \ closedin X L \ K \ U \ U \ L))" (is "?lhs=?rhs") proof assume L: ?lhs then obtain U L where UL: "\x \ topspace X. openin X (U x) \ compactin X (L x) \ closedin X (L x) \ x \ U x \ U x \ L x" unfolding locally_compact_space_compact_closedin [OF assms] by metis show ?rhs proof clarify fix K assume "compactin X K" then have "K \ topspace X" by (simp add: compactin_subset_topspace) then have *: "(\U\U ` K. openin X U) \ K \ \ (U ` K)" using UL by blast with \compactin X K\ obtain KF where KF: "finite KF" "KF \ K" "K \ \(U ` KF)" by (metis compactinD finite_subset_image) show "\U L. openin X U \ compactin X L \ closedin X L \ K \ U \ U \ L" proof (intro conjI exI) show "openin X (\ (U ` KF))" using "*" \KF \ K\ by fastforce show "compactin X (\ (L ` KF))" by (smt (verit) UL \K \ topspace X\ KF compactin_Union finite_imageI imageE subset_iff) show "closedin X (\ (L ` KF))" by (smt (verit) UL \K \ topspace X\ KF closedin_Union finite_imageI imageE subsetD) qed (use UL \K \ topspace X\ KF in auto) qed next assume ?rhs then show ?lhs by (metis compactin_sing insert_subset locally_compact_space_def) qed lemma locally_compact_regular_space_neighbourhood_base: "locally_compact_space X \ regular_space X \ neighbourhood_base_of (\C. compactin X C \ closedin X C) X" using locally_compact_space_neighbourhood_base_closedin neighbourhood_base_of_closedin neighbourhood_base_of_mono by blast lemma locally_compact_kc_space: "neighbourhood_base_of (compactin X) X \ kc_space X \ locally_compact_space X \ Hausdorff_space X" using Hausdorff_imp_kc_space locally_compact_imp_kc_eq_Hausdorff_space locally_compact_space_neighbourhood_base by blast lemma locally_compact_kc_space_alt: "neighbourhood_base_of (compactin X) X \ kc_space X \ locally_compact_space X \ Hausdorff_space X \ regular_space X" using Hausdorff_regular locally_compact_kc_space by blast lemma locally_compact_kc_imp_regular_space: "\neighbourhood_base_of (compactin X) X; kc_space X\ \ regular_space X" using Hausdorff_regular locally_compact_imp_kc_eq_Hausdorff_space by blast lemma kc_locally_compact_space: "kc_space X \ neighbourhood_base_of (compactin X) X \ locally_compact_space X \ Hausdorff_space X \ regular_space X" using Hausdorff_regular locally_compact_kc_space by blast lemma locally_compact_space_closed_subset: assumes loc: "locally_compact_space X" and "closedin X S" shows "locally_compact_space (subtopology X S)" proof (clarsimp simp: locally_compact_space_def) fix x assume x: "x \ topspace X" "x \ S" then obtain U K where UK: "openin X U \ compactin X K \ x \ U \ U \ K" by (meson loc locally_compact_space_def) show "\U. openin (subtopology X S) U \ (\K. compactin (subtopology X S) K \ x \ U \ U \ K)" proof (intro conjI exI) show "openin (subtopology X S) (S \ U)" by (simp add: UK openin_subtopology_Int2) show "compactin (subtopology X S) (S \ K)" by (simp add: UK assms(2) closed_Int_compactin compactin_subtopology) qed (use UK x in auto) qed lemma locally_compact_space_open_subset: assumes reg: "regular_space X" and loc: "locally_compact_space X" and "openin X S" shows "locally_compact_space (subtopology X S)" proof (clarsimp simp: locally_compact_space_def) fix x assume x: "x \ topspace X" "x \ S" then obtain U K where UK: "openin X U" "compactin X K" "x \ U" "U \ K" by (meson loc locally_compact_space_def) have "openin X (U \ S)" by (simp add: UK \openin X S\ openin_Int) with UK reg x obtain V C where VC: "openin X V" "closedin X C" "x \ V" "V \ C" "C \ U" "C \ S" by (metis IntI le_inf_iff neighbourhood_base_of neighbourhood_base_of_closedin) show "\U. openin (subtopology X S) U \ (\K. compactin (subtopology X S) K \ x \ U \ U \ K)" proof (intro conjI exI) show "openin (subtopology X S) V" using VC by (meson \openin X S\ openin_open_subtopology order_trans) show "compactin (subtopology X S) (C \ K)" using UK VC closed_Int_compactin compactin_subtopology by fastforce qed (use UK VC x in auto) qed lemma locally_compact_space_discrete_topology: "locally_compact_space (discrete_topology U)" by (simp add: neighbourhood_base_imp_locally_compact_space neighbourhood_base_of_discrete_topology) lemma locally_compact_space_continuous_open_map_image: "\continuous_map X X' f; open_map X X' f; f ` topspace X = topspace X'; locally_compact_space X\ \ locally_compact_space X'" unfolding locally_compact_space_def open_map_def by (smt (verit, ccfv_SIG) image_compactin image_iff image_mono) lemma locally_compact_subspace_openin_closure_of: assumes "Hausdorff_space X" and S: "S \ topspace X" and loc: "locally_compact_space (subtopology X S)" shows "openin (subtopology X (X closure_of S)) S" unfolding openin_subopen [where S=S] proof clarify fix a assume "a \ S" then obtain T K where *: "openin X T" "compactin X K" "K \ S" "a \ S" "a \ T" "S \ T \ K" using loc unfolding locally_compact_space_def by (metis IntE S compactin_subtopology inf_commute openin_subtopology topspace_subtopology_subset) have "T \ X closure_of S \ X closure_of (T \ S)" by (simp add: "*"(1) openin_Int_closure_of_subset) also have "... \ S" using * \Hausdorff_space X\ by (metis closure_of_minimal compactin_imp_closedin order.trans inf_commute) finally have "T \ X closure_of S \ T \ S" by simp then have "openin (subtopology X (X closure_of S)) (T \ S)" unfolding openin_subtopology using \openin X T\ S closure_of_subset by fastforce with * show "\T. openin (subtopology X (X closure_of S)) T \ a \ T \ T \ S" by blast qed lemma locally_compact_subspace_closed_Int_openin: "\Hausdorff_space X \ S \ topspace X \ locally_compact_space(subtopology X S)\ \ \C U. closedin X C \ openin X U \ C \ U = S" by (metis closedin_closure_of inf_commute locally_compact_subspace_openin_closure_of openin_subtopology) lemma locally_compact_subspace_open_in_closure_of_eq: assumes "Hausdorff_space X" and loc: "locally_compact_space X" shows "openin (subtopology X (X closure_of S)) S \ S \ topspace X \ locally_compact_space(subtopology X S)" (is "?lhs=?rhs") proof assume L: ?lhs then obtain "S \ topspace X" "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space openin_subset by fastforce then have "locally_compact_space (subtopology (subtopology X (X closure_of S)) S)" by (simp add: L loc locally_compact_space_closed_subset locally_compact_space_open_subset regular_space_subtopology) then show ?rhs by (metis L inf.orderE inf_commute le_inf_iff openin_subset subtopology_subtopology topspace_subtopology) next assume ?rhs then show ?lhs using assms locally_compact_subspace_openin_closure_of by blast qed lemma locally_compact_subspace_closed_Int_openin_eq: assumes "Hausdorff_space X" and loc: "locally_compact_space X" shows "(\C U. closedin X C \ openin X U \ C \ U = S) \ S \ topspace X \ locally_compact_space(subtopology X S)" (is "?lhs=?rhs") proof assume L: ?lhs then obtain C U where "closedin X C" "openin X U" and Seq: "S = C \ U" by blast then have "C \ topspace X" by (simp add: closedin_subset) have "locally_compact_space (subtopology (subtopology X C) (topspace (subtopology X C) \ U))" proof (rule locally_compact_space_open_subset) show "regular_space (subtopology X C)" by (simp add: \Hausdorff_space X\ loc locally_compact_Hausdorff_imp_regular_space regular_space_subtopology) show "locally_compact_space (subtopology X C)" by (simp add: \closedin X C\ loc locally_compact_space_closed_subset) show "openin (subtopology X C) (topspace (subtopology X C) \ U)" by (simp add: \openin X U\ Int_left_commute inf_commute openin_Int openin_subtopology_Int2) qed then show ?rhs by (metis Seq \C \ topspace X\ inf.coboundedI1 subtopology_subtopology subtopology_topspace) next assume ?rhs then show ?lhs using assms locally_compact_subspace_closed_Int_openin by blast qed lemma dense_locally_compact_openin_Hausdorff_space: "\Hausdorff_space X; S \ topspace X; X closure_of S = topspace X; locally_compact_space (subtopology X S)\ \ openin X S" by (metis locally_compact_subspace_openin_closure_of subtopology_topspace) lemma locally_compact_space_prod_topology: "locally_compact_space (prod_topology X Y) \ topspace (prod_topology X Y) = {} \ locally_compact_space X \ locally_compact_space Y" (is "?lhs=?rhs") proof (cases "topspace (prod_topology X Y) = {}") case True then show ?thesis unfolding locally_compact_space_def by blast next case False then obtain w z where wz: "w \ topspace X" "z \ topspace Y" by auto show ?thesis proof assume L: ?lhs then show ?rhs by (metis wz empty_iff locally_compact_space_retraction_map_image retraction_map_fst retraction_map_snd) next assume R: ?rhs show ?lhs unfolding locally_compact_space_def proof clarsimp fix x y assume "x \ topspace X" and "y \ topspace Y" obtain U C where "openin X U" "compactin X C" "x \ U" "U \ C" by (meson False R \x \ topspace X\ locally_compact_space_def) obtain V D where "openin Y V" "compactin Y D" "y \ V" "V \ D" by (meson False R \y \ topspace Y\ locally_compact_space_def) show "\U. openin (prod_topology X Y) U \ (\K. compactin (prod_topology X Y) K \ (x, y) \ U \ U \ K)" proof (intro exI conjI) show "openin (prod_topology X Y) (U \ V)" by (simp add: \openin X U\ \openin Y V\ openin_prod_Times_iff) show "compactin (prod_topology X Y) (C \ D)" by (simp add: \compactin X C\ \compactin Y D\ compactin_Times) show "(x, y) \ U \ V" by (simp add: \x \ U\ \y \ V\) show "U \ V \ C \ D" by (simp add: Sigma_mono \U \ C\ \V \ D\) qed qed qed qed lemma locally_compact_space_product_topology: "locally_compact_space(product_topology X I) \ topspace(product_topology X I) = {} \ finite {i \ I. \ compact_space(X i)} \ (\i \ I. locally_compact_space(X i))" (is "?lhs=?rhs") proof (cases "topspace (product_topology X I) = {}") case True then show ?thesis by (simp add: locally_compact_space_def) next case False show ?thesis proof assume L: ?lhs obtain z where z: "z \ topspace (product_topology X I)" using False by auto with L z obtain U C where "openin (product_topology X I) U" "compactin (product_topology X I) C" "z \ U" "U \ C" by (meson locally_compact_space_def) then obtain V where finV: "finite {i \ I. V i \ topspace (X i)}" and "\i \ I. openin (X i) (V i)" and "z \ PiE I V" "PiE I V \ U" by (auto simp: openin_product_topology_alt) have "compact_space (X i)" if "i \ I" "V i = topspace (X i)" for i proof - have "compactin (X i) ((\x. x i) ` C)" using \compactin (product_topology X I) C\ image_compactin by (metis continuous_map_product_projection \i \ I\) moreover have "V i \ (\x. x i) ` C" proof - have "V i \ (\x. x i) ` Pi\<^sub>E I V" by (metis \z \ Pi\<^sub>E I V\ empty_iff image_projection_PiE order_refl \i \ I\) also have "\ \ (\x. x i) ` C" using \U \ C\ \Pi\<^sub>E I V \ U\ by blast finally show ?thesis . qed ultimately show ?thesis by (metis closed_compactin closedin_topspace compact_space_def that(2)) qed with finV have "finite {i \ I. \ compact_space (X i)}" by (metis (mono_tags, lifting) mem_Collect_eq finite_subset subsetI) moreover have "locally_compact_space (X i)" if "i \ I" for i by (meson False L locally_compact_space_retraction_map_image retraction_map_product_projection that) ultimately show ?rhs by metis next assume R: ?rhs show ?lhs unfolding locally_compact_space_def proof clarsimp fix z assume z: "z \ (\\<^sub>E i\I. topspace (X i))" have "\U C. openin (X i) U \ compactin (X i) C \ z i \ U \ U \ C \ (compact_space(X i) \ U = topspace(X i) \ C = topspace(X i))" if "i \ I" for i using that R z unfolding locally_compact_space_def compact_space_def by (metis (no_types, lifting) False PiE_mem openin_topspace set_eq_subset) then obtain U C where UC: "\i. i \ I \ openin (X i) (U i) \ compactin (X i) (C i) \ z i \ U i \ U i \ C i \ (compact_space(X i) \ U i = topspace(X i) \ C i = topspace(X i))" by metis show "\U. openin (product_topology X I) U \ (\K. compactin (product_topology X I) K \ z \ U \ U \ K)" proof (intro exI conjI) show "openin (product_topology X I) (Pi\<^sub>E I U)" by (smt (verit) Collect_cong False R UC compactin_subspace openin_PiE_gen subset_antisym subtopology_topspace) show "compactin (product_topology X I) (Pi\<^sub>E I C)" by (simp add: UC compactin_PiE) qed (use UC z in blast)+ qed qed qed lemma locally_compact_space_sum_topology: "locally_compact_space (sum_topology X I) \ (\i \ I. locally_compact_space (X i))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis closed_map_component_injection embedding_map_imp_homeomorphic_space embedding_map_component_injection embedding_imp_closed_map_eq homeomorphic_locally_compact_space locally_compact_space_closed_subset) next assume R: ?rhs show ?lhs unfolding locally_compact_space_def proof clarsimp fix i y assume "i \ I" and y: "y \ topspace (X i)" then obtain U K where UK: "openin (X i) U" "compactin (X i) K" "y \ U" "U \ K" using R by (fastforce simp: locally_compact_space_def) then show "\U. openin (sum_topology X I) U \ (\K. compactin (sum_topology X I) K \ (i, y) \ U \ U \ K)" by (metis \i \ I\ continuous_map_component_injection image_compactin image_mono imageI open_map_component_injection open_map_def) qed qed proposition quotient_map_prod_right: assumes loc: "locally_compact_space Z" and reg: "Hausdorff_space Z \ regular_space Z" and f: "quotient_map X Y f" shows "quotient_map (prod_topology Z X) (prod_topology Z Y) (\(x,y). (x,f y))" proof - define h where "h \ (\(x::'a,y). (x,f y))" have "continuous_map (prod_topology Z X) Y (f o snd)" by (simp add: continuous_map_of_snd f quotient_imp_continuous_map) then have cmh: "continuous_map (prod_topology Z X) (prod_topology Z Y) h" by (simp add: h_def continuous_map_paired split_def continuous_map_fst o_def) have fim: "f ` topspace X = topspace Y" by (simp add: f quotient_imp_surjective_map) moreover have "openin (prod_topology Z X) {u \ topspace Z \ topspace X. h u \ W} \ openin (prod_topology Z Y) W" (is "?lhs=?rhs") if W: "W \ topspace Z \ topspace Y" for W proof define S where "S \ {u \ topspace Z \ topspace X. h u \ W}" assume ?lhs then have L: "openin (prod_topology Z X) S" using S_def by blast have "\T. openin (prod_topology Z Y) T \ (x0, z0) \ T \ T \ W" if \
: "(x0,z0) \ W" for x0 z0 proof - have x0: "x0 \ topspace Z" using W that by blast obtain y0 where y0: "y0 \ topspace X" "f y0 = z0" by (metis W fim imageE insert_absorb insert_subset mem_Sigma_iff \
) then have "(x0, y0) \ S" by (simp add: S_def h_def that x0) have "continuous_map Z (prod_topology Z X) (\x. (x, y0))" by (simp add: continuous_map_paired y0) with openin_continuous_map_preimage [OF _ L] have ope_ZS: "openin Z {x \ topspace Z. (x,y0) \ S}" by blast obtain U U' where "openin Z U" "compactin Z U'" "closedin Z U'" "x0 \ U" "U \ U'" "U' \ {x \ topspace Z. (x,y0) \ S}" using loc ope_ZS x0 \(x0, y0) \ S\ by (force simp: locally_compact_space_neighbourhood_base_closedin [OF reg] neighbourhood_base_of) then have D: "U' \ {y0} \ S" by (auto simp: ) define V where "V \ {z \ topspace Y. U' \ {y \ topspace X. f y = z} \ S}" have "z0 \ V" using D y0 Int_Collect fim by (fastforce simp: h_def V_def S_def) have "openin X {x \ topspace X. f x \ V} \ openin Y V" using f unfolding V_def quotient_map_def subset_iff by (smt (verit, del_insts) Collect_cong mem_Collect_eq) moreover have "openin X {x \ topspace X. f x \ V}" proof - let ?Z = "subtopology Z U'" have *: "{x \ topspace X. f x \ V} = topspace X - snd ` (U' \ topspace X - S)" by (force simp: V_def S_def h_def simp flip: fim) have "compact_space ?Z" using \compactin Z U'\ compactin_subspace by auto moreover have "closedin (prod_topology ?Z X) (U' \ topspace X - S)" by (simp add: L \closedin Z U'\ closedin_closed_subtopology closedin_diff closedin_prod_Times_iff prod_topology_subtopology(1)) ultimately show ?thesis using "*" closed_map_snd closed_map_def by fastforce qed ultimately have "openin Y V" by metis show ?thesis proof (intro conjI exI) show "openin (prod_topology Z Y) (U \ V)" by (simp add: openin_prod_Times_iff \openin Z U\ \openin Y V\) show "(x0, z0) \ U \ V" by (simp add: \x0 \ U\ \z0 \ V\) show "U \ V \ W" using \U \ U'\ by (force simp: V_def S_def h_def simp flip: fim) qed qed with openin_subopen show ?rhs by force next assume ?rhs then show ?lhs using openin_continuous_map_preimage cmh by fastforce qed ultimately show ?thesis by (fastforce simp: image_iff quotient_map_def h_def) qed lemma quotient_map_prod_left: assumes loc: "locally_compact_space Z" and reg: "Hausdorff_space Z \ regular_space Z" and f: "quotient_map X Y f" shows "quotient_map (prod_topology X Z) (prod_topology Y Z) (\(x,y). (f x,y))" proof - have "(\(x,y). (f x,y)) = prod.swap \ (\(x,y). (x,f y)) \ prod.swap" by force then show ?thesis apply (rule ssubst) proof (intro quotient_map_compose) show "quotient_map (prod_topology X Z) (prod_topology Z X) prod.swap" "quotient_map (prod_topology Z Y) (prod_topology Y Z) prod.swap" using homeomorphic_map_def homeomorphic_map_swap quotient_map_eq by fastforce+ show "quotient_map (prod_topology Z X) (prod_topology Z Y) (\(x, y). (x, f y))" by (simp add: f loc quotient_map_prod_right reg) qed qed lemma locally_compact_space_perfect_map_preimage: assumes "locally_compact_space X'" and f: "perfect_map X X' f" shows "locally_compact_space X" unfolding locally_compact_space_def proof (intro strip) fix x assume x: "x \ topspace X" then obtain U K where "openin X' U" "compactin X' K" "f x \ U" "U \ K" using assms unfolding locally_compact_space_def perfect_map_def by (metis (no_types, lifting) continuous_map_closedin) show "\U K. openin X U \ compactin X K \ x \ U \ U \ K" proof (intro exI conjI) have "continuous_map X X' f" using f perfect_map_def by blast then show "openin X {x \ topspace X. f x \ U}" by (simp add: \openin X' U\ continuous_map) show "compactin X {x \ topspace X. f x \ K}" using \compactin X' K\ f perfect_imp_proper_map proper_map_alt by blast qed (use x \f x \ U\ \U \ K\ in auto) qed subsection\Special characterizations of classes of functions into and out of R\ lemma monotone_map_into_euclideanreal_alt: assumes "continuous_map X euclideanreal f" shows "(\k. is_interval k \ connectedin X {x \ topspace X. f x \ k}) \ connected_space X \ monotone_map X euclideanreal f" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof show "connected_space X" using L connected_space_subconnected by blast have "connectedin X {x \ topspace X. f x \ {y}}" for y by (metis L is_interval_1 nle_le singletonD) then show "monotone_map X euclideanreal f" by (simp add: monotone_map) qed next assume R: ?rhs then have *: False if "a < b" "closedin X U" "closedin X V" "U \ {}" "V \ {}" "disjnt U V" and UV: "{x \ topspace X. f x \ {a..b}} = U \ V" and dis: "disjnt U {x \ topspace X. f x = b}" "disjnt V {x \ topspace X. f x = a}" for a b U V proof - define E1 where "E1 \ U \ {x \ topspace X. f x \ {c. c \ a}}" define E2 where "E2 \ V \ {x \ topspace X. f x \ {c. b \ c}}" have "closedin X {x \ topspace X. f x \ a}" "closedin X {x \ topspace X. b \ f x}" using assms continuous_map_upper_lower_semicontinuous_le by blast+ then have "closedin X E1" "closedin X E2" unfolding E1_def E2_def using that by auto moreover have "E1 \ E2 = {}" unfolding E1_def E2_def using \a \disjnt U V\ dis UV by (simp add: disjnt_def set_eq_iff) (smt (verit)) have "topspace X \ E1 \ E2" unfolding E1_def E2_def using UV by fastforce have "E1 = {} \ E2 = {}" using R connected_space_closedin using \E1 \ E2 = {}\ \closedin X E1\ \closedin X E2\ \topspace X \ E1 \ E2\ by blast then show False using E1_def E2_def \U \ {}\ \V \ {}\ by fastforce qed show ?lhs proof (intro strip) fix K :: "real set" assume "is_interval K" have False if "a \ K" "b \ K" and clo: "closedin X U" "closedin X V" and UV: "{x. x \ topspace X \ f x \ K} \ U \ V" "U \ V \ {x. x \ topspace X \ f x \ K} = {}" and nondis: "\ disjnt U {x. x \ topspace X \ f x = a}" "\ disjnt V {x. x \ topspace X \ f x = b}" for a b U V proof - have "\y. connectedin X {x. x \ topspace X \ f x = y}" using R monotone_map by fastforce then have **: False if "p \ U \ q \ V \ f p = f q \ f q \ K" for p q unfolding connectedin_closedin using \a \ K\ \b \ K\ UV clo that by (smt (verit, ccfv_threshold) closedin_subset disjoint_iff mem_Collect_eq subset_iff) consider "a < b" | "a = b" | "b < a" by linarith then show ?thesis proof cases case 1 define W where "W \ {x \ topspace X. f x \ {a..b}}" have "closedin X W" unfolding W_def by (metis (no_types) assms closed_real_atLeastAtMost closed_closedin continuous_map_closedin) show ?thesis proof (rule * [OF 1 , of "U \ W" "V \ W"]) show "closedin X (U \ W)" "closedin X (V \ W)" using \closedin X W\ clo by auto show "U \ W \ {}" "V \ W \ {}" using nondis 1 by (auto simp: disjnt_iff W_def) show "disjnt (U \ W) (V \ W)" using \is_interval K\ unfolding is_interval_1 disjnt_iff W_def by (metis (mono_tags, lifting) \a \ K\ \b \ K\ ** Int_Collect atLeastAtMost_iff) have "\x. \x \ topspace X; a \ f x; f x \ b\ \ x \ U \ x \ V" using \a \ K\ \b \ K\ \is_interval K\ UV unfolding is_interval_1 disjnt_iff by blast then show "{x \ topspace X. f x \ {a..b}} = U \ W \ V \ W" by (auto simp: W_def) show "disjnt (U \ W) {x \ topspace X. f x = b}" "disjnt (V \ W) {x \ topspace X. f x = a}" using ** \a \ K\ \b \ K\ nondis by (force simp: disjnt_iff)+ qed next case 2 then show ?thesis using ** nondis \b \ K\ by (force simp add: disjnt_iff) next case 3 define W where "W \ {x \ topspace X. f x \ {b..a}}" have "closedin X W" unfolding W_def by (metis (no_types) assms closed_real_atLeastAtMost closed_closedin continuous_map_closedin) show ?thesis proof (rule * [OF 3, of "V \ W" "U \ W"]) show "closedin X (U \ W)" "closedin X (V \ W)" using \closedin X W\ clo by auto show "U \ W \ {}" "V \ W \ {}" using nondis 3 by (auto simp: disjnt_iff W_def) show "disjnt (V \ W) (U \ W)" using \is_interval K\ unfolding is_interval_1 disjnt_iff W_def by (metis (mono_tags, lifting) \a \ K\ \b \ K\ ** Int_Collect atLeastAtMost_iff) have "\x. \x \ topspace X; b \ f x; f x \ a\ \ x \ U \ x \ V" using \a \ K\ \b \ K\ \is_interval K\ UV unfolding is_interval_1 disjnt_iff by blast then show "{x \ topspace X. f x \ {b..a}} = V \ W \ U \ W" by (auto simp: W_def) show "disjnt (V \ W) {x \ topspace X. f x = a}" "disjnt (U \ W) {x \ topspace X. f x = b}" using ** \a \ K\ \b \ K\ nondis by (force simp: disjnt_iff)+ qed qed qed then show "connectedin X {x \ topspace X. f x \ K}" unfolding connectedin_closedin disjnt_iff by blast qed qed lemma monotone_map_into_euclideanreal: "\connected_space X; continuous_map X euclideanreal f\ \ monotone_map X euclideanreal f \ (\k. is_interval k \ connectedin X {x \ topspace X. f x \ k})" by (simp add: monotone_map_into_euclideanreal_alt) lemma monotone_map_euclideanreal_alt: "(\I::real set. is_interval I \ is_interval {x::real. x \ S \ f x \ I}) \ is_interval S \ (mono_on S f \ antimono_on S f)" (is "?lhs=?rhs") proof assume L [rule_format]: ?lhs show ?rhs proof show "is_interval S" using L is_interval_1 by auto have False if "a \ S" "b \ S" "c \ S" "a f c < f b \ f a > f b \ f c > f b" for a b c using d proof assume "f a < f b \ f c < f b" then show False using L [of "{y. y < f b}"] unfolding is_interval_1 by (smt (verit, best) mem_Collect_eq that) next assume "f b < f a \ f b < f c" then show False using L [of "{y. y > f b}"] unfolding is_interval_1 by (smt (verit, best) mem_Collect_eq that) qed then show "mono_on S f \ monotone_on S (\) (\) f" unfolding monotone_on_def by (smt (verit)) qed next assume ?rhs then show ?lhs unfolding is_interval_1 monotone_on_def by simp meson qed lemma monotone_map_euclideanreal: fixes S :: "real set" shows "\is_interval S; continuous_on S f\ \ monotone_map (top_of_set S) euclideanreal f \ (mono_on S f \ monotone_on S (\) (\) f)" using monotone_map_euclideanreal_alt by (simp add: monotone_map_into_euclideanreal connectedin_subtopology is_interval_connected_1) lemma injective_eq_monotone_map: fixes f :: "real \ real" assumes "is_interval S" "continuous_on S f" shows "inj_on f S \ strict_mono_on S f \ strict_antimono_on S f" by (metis assms injective_imp_monotone_map monotone_map_euclideanreal strict_antimono_iff_antimono strict_mono_iff_mono top_greatest topspace_euclidean topspace_euclidean_subtopology) subsection\Normal spaces including Urysohn's lemma and the Tietze extension theorem\ definition normal_space where "normal_space X \ \S T. closedin X S \ closedin X T \ disjnt S T \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V)" lemma normal_space_retraction_map_image: assumes r: "retraction_map X Y r" and X: "normal_space X" shows "normal_space Y" unfolding normal_space_def proof clarify fix S T assume "closedin Y S" and "closedin Y T" and "disjnt S T" obtain r' where r': "retraction_maps X Y r r'" using r retraction_map_def by blast have "closedin X {x \ topspace X. r x \ S}" "closedin X {x \ topspace X. r x \ T}" using closedin_continuous_map_preimage \closedin Y S\ \closedin Y T\ r' by (auto simp: retraction_maps_def) moreover have "disjnt {x \ topspace X. r x \ S} {x \ topspace X. r x \ T}" using \disjnt S T\ by (auto simp: disjnt_def) ultimately obtain U V where UV: "openin X U \ openin X V \ {x \ topspace X. r x \ S} \ U \ {x \ topspace X. r x \ T} \ V" "disjnt U V" by (meson X normal_space_def) show "\U V. openin Y U \ openin Y V \ S \ U \ T \ V \ disjnt U V" proof (intro exI conjI) show "openin Y {x \ topspace Y. r' x \ U}" "openin Y {x \ topspace Y. r' x \ V}" using openin_continuous_map_preimage UV r' by (auto simp: retraction_maps_def) show "S \ {x \ topspace Y. r' x \ U}" "T \ {x \ topspace Y. r' x \ V}" using openin_continuous_map_preimage UV r' \closedin Y S\ \closedin Y T\ by (auto simp add: closedin_def continuous_map_closedin retraction_maps_def subset_iff) show "disjnt {x \ topspace Y. r' x \ U} {x \ topspace Y. r' x \ V}" using \disjnt U V\ by (auto simp: disjnt_def) qed qed lemma homeomorphic_normal_space: "X homeomorphic_space Y \ normal_space X \ normal_space Y" unfolding homeomorphic_space_def by (meson homeomorphic_imp_retraction_maps homeomorphic_maps_sym normal_space_retraction_map_image retraction_map_def) lemma normal_space: "normal_space X \ (\S T. closedin X S \ closedin X T \ disjnt S T \ (\U. openin X U \ S \ U \ disjnt T (X closure_of U)))" proof - have "(\V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V) \ openin X U \ S \ U \ disjnt T (X closure_of U)" (is "?lhs=?rhs") if "closedin X S" "closedin X T" "disjnt S T" for S T U proof show "?lhs \ ?rhs" by (smt (verit, best) disjnt_iff in_closure_of subsetD) assume R: ?rhs then have "(U \ S) \ (topspace X - X closure_of U) = {}" by (metis Diff_eq_empty_iff Int_Diff Int_Un_eq(4) closure_of_subset inf.orderE openin_subset) moreover have "T \ topspace X - X closure_of U" by (meson DiffI R closedin_subset disjnt_iff subsetD subsetI that(2)) ultimately show ?lhs by (metis R closedin_closure_of closedin_def disjnt_def sup.orderE) qed then show ?thesis unfolding normal_space_def by meson qed lemma normal_space_alt: "normal_space X \ (\S U. closedin X S \ openin X U \ S \ U \ (\V. openin X V \ S \ V \ X closure_of V \ U))" proof - have "\V. openin X V \ S \ V \ X closure_of V \ U" if "\T. closedin X T \ disjnt S T \ (\U. openin X U \ S \ U \ disjnt T (X closure_of U))" "closedin X S" "openin X U" "S \ U" for S U using that by (smt (verit) Diff_eq_empty_iff Int_Diff closure_of_subset_topspace disjnt_def inf.orderE inf_commute openin_closedin_eq) moreover have "\U. openin X U \ S \ U \ disjnt T (X closure_of U)" if "\U. openin X U \ S \ U \ (\V. openin X V \ S \ V \ X closure_of V \ U)" and "closedin X S" "closedin X T" "disjnt S T" for S T using that by (smt (verit) Diff_Diff_Int Diff_eq_empty_iff Int_Diff closedin_def disjnt_def inf.absorb_iff2 inf.orderE) ultimately show ?thesis by (fastforce simp: normal_space) qed lemma normal_space_closures: "normal_space X \ (\S T. S \ topspace X \ T \ topspace X \ disjnt (X closure_of S) (X closure_of T) \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V))" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" by (meson closedin_closure_of closure_of_subset normal_space_def order.trans) show "?rhs \ ?lhs" by (metis closedin_subset closure_of_eq normal_space_def) qed lemma normal_space_disjoint_closures: "normal_space X \ (\S T. closedin X S \ closedin X T \ disjnt S T \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt (X closure_of U) (X closure_of V)))" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" by (metis closedin_closure_of normal_space) show "?rhs \ ?lhs" by (smt (verit) closure_of_subset disjnt_iff normal_space openin_subset subset_eq) qed lemma normal_space_dual: "normal_space X \ (\U V. openin X U \ openin X V \ U \ V = topspace X \ (\S T. closedin X S \ closedin X T \ S \ U \ T \ V \ S \ T = topspace X))" (is "_ = ?rhs") proof - have "normal_space X \ (\U V. closedin X U \ closedin X V \ disjnt U V \ (\S T. \ (openin X S \ openin X T \ \ (U \ S \ V \ T \ disjnt S T))))" unfolding normal_space_def by meson also have "... \ (\U V. openin X U \ openin X V \ disjnt (topspace X - U) (topspace X - V) \ (\S T. \ (openin X S \ openin X T \ \ (topspace X - U \ S \ topspace X - V \ T \ disjnt S T))))" by (auto simp: all_closedin) also have "... \ ?rhs" proof - have *: "disjnt (topspace X - U) (topspace X - V) \ U \ V = topspace X" if "U \ topspace X" "V \ topspace X" for U V using that by (auto simp: disjnt_iff) show ?thesis using ex_closedin * apply (simp add: ex_closedin * [OF openin_subset openin_subset] cong: conj_cong) apply (intro all_cong1 ex_cong1 imp_cong refl) by (smt (verit, best) "*" Diff_Diff_Int Diff_subset Diff_subset_conv inf.orderE inf_commute openin_subset sup_commute) qed finally show ?thesis . qed lemma normal_t1_imp_Hausdorff_space: assumes "normal_space X" "t1_space X" shows "Hausdorff_space X" unfolding Hausdorff_space_def proof clarify fix x y assume xy: "x \ topspace X" "y \ topspace X" "x \ y" then have "disjnt {x} {y}" by (auto simp: disjnt_iff) then show "\U V. openin X U \ openin X V \ x \ U \ y \ V \ disjnt U V" using assms xy closedin_t1_singleton normal_space_def by (metis singletonI subsetD) qed lemma normal_t1_eq_Hausdorff_space: "normal_space X \ t1_space X \ Hausdorff_space X" using normal_t1_imp_Hausdorff_space t1_or_Hausdorff_space by blast lemma normal_t1_imp_regular_space: "\normal_space X; t1_space X\ \ regular_space X" by (metis compactin_imp_closedin normal_space_def normal_t1_eq_Hausdorff_space regular_space_compact_closed_sets) lemma compact_Hausdorff_or_regular_imp_normal_space: "\compact_space X; Hausdorff_space X \ regular_space X\ \ normal_space X" by (metis Hausdorff_space_compact_sets closedin_compact_space normal_space_def regular_space_compact_closed_sets) lemma normal_space_discrete_topology: "normal_space(discrete_topology U)" by (metis discrete_topology_closure_of inf_le2 normal_space_alt) lemma normal_space_fsigmas: "normal_space X \ (\S T. fsigma_in X S \ fsigma_in X T \ separatedin X S T \ (\U B. openin X U \ openin X B \ S \ U \ T \ B \ disjnt U B))" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof clarify fix S T assume "fsigma_in X S" then obtain C where C: "\n. closedin X (C n)" "\n. C n \ C (Suc n)" "\ (range C) = S" by (meson fsigma_in_ascending) assume "fsigma_in X T" then obtain D where D: "\n. closedin X (D n)" "\n. D n \ D (Suc n)" "\ (range D) = T" by (meson fsigma_in_ascending) assume "separatedin X S T" have "\n. disjnt (D n) (X closure_of S)" by (metis D(3) \separatedin X S T\ disjnt_Union1 disjnt_def rangeI separatedin_def) then have "\n. \V V'. openin X V \ openin X V' \ D n \ V \ X closure_of S \ V' \ disjnt V V'" by (metis D(1) L closedin_closure_of normal_space_def) then obtain V V' where V: "\n. openin X (V n)" and "\n. openin X (V' n)" "\n. disjnt (V n) (V' n)" and DV: "\n. D n \ V n" and subV': "\n. X closure_of S \ V' n" by metis then have VV: "V' n \ X closure_of V n = {}" for n using openin_Int_closure_of_eq_empty [of X "V' n" "V n"] by (simp add: Int_commute disjnt_def) have "\n. disjnt (C n) (X closure_of T)" by (metis C(3) \separatedin X S T\ disjnt_Union1 disjnt_def rangeI separatedin_def) then have "\n. \U U'. openin X U \ openin X U' \ C n \ U \ X closure_of T \ U' \ disjnt U U'" by (metis C(1) L closedin_closure_of normal_space_def) then obtain U U' where U: "\n. openin X (U n)" and "\n. openin X (U' n)" "\n. disjnt (U n) (U' n)" and CU: "\n. C n \ U n" and subU': "\n. X closure_of T \ U' n" by metis then have UU: "U' n \ X closure_of U n = {}" for n using openin_Int_closure_of_eq_empty [of X "U' n" "U n"] by (simp add: Int_commute disjnt_def) show "\U B. openin X U \ openin X B \ S \ U \ T \ B \ disjnt U B" proof (intro conjI exI) have "\S n. closedin X (\m\n. X closure_of V m)" by (force intro: closedin_Union) then show "openin X (\n. U n - (\m\n. X closure_of V m))" using U by blast have "\S n. closedin X (\m\n. X closure_of U m)" by (force intro: closedin_Union) then show "openin X (\n. V n - (\m\n. X closure_of U m))" using V by blast have "S \ topspace X" by (simp add: \fsigma_in X S\ fsigma_in_subset) then show "S \ (\n. U n - (\m\n. X closure_of V m))" apply (clarsimp simp: Ball_def) by (metis VV C(3) CU IntI UN_E closure_of_subset empty_iff subV' subsetD) have "T \ topspace X" by (simp add: \fsigma_in X T\ fsigma_in_subset) then show "T \ (\n. V n - (\m\n. X closure_of U m))" apply (clarsimp simp: Ball_def) by (metis UU D(3) DV IntI UN_E closure_of_subset empty_iff subU' subsetD) have "\x m n. \x \ U n; x \ V m; \k\m. x \ X closure_of U k\ \ \k\n. x \ X closure_of V k" by (meson U V closure_of_subset nat_le_linear openin_subset subsetD) then show "disjnt (\n. U n - (\m\n. X closure_of V m)) (\n. V n - (\m\n. X closure_of U m))" by (force simp: disjnt_iff) qed qed next show "?rhs \ ?lhs" by (simp add: closed_imp_fsigma_in normal_space_def separatedin_closed_sets) qed lemma normal_space_fsigma_subtopology: assumes "normal_space X" "fsigma_in X S" shows "normal_space (subtopology X S)" unfolding normal_space_fsigmas proof clarify fix T U assume "fsigma_in (subtopology X S) T" and "fsigma_in (subtopology X S) U" and TU: "separatedin (subtopology X S) T U" then obtain A B where "openin X A \ openin X B \ T \ A \ U \ B \ disjnt A B" by (metis assms fsigma_in_fsigma_subtopology normal_space_fsigmas separatedin_subtopology) then show "\A B. openin (subtopology X S) A \ openin (subtopology X S) B \ T \ A \ U \ B \ disjnt A B" using TU by (force simp add: separatedin_subtopology openin_subtopology_alt disjnt_iff) qed lemma normal_space_closed_subtopology: assumes "normal_space X" "closedin X S" shows "normal_space (subtopology X S)" by (simp add: assms closed_imp_fsigma_in normal_space_fsigma_subtopology) lemma normal_space_continuous_closed_map_image: assumes "normal_space X" and contf: "continuous_map X Y f" and clof: "closed_map X Y f" and fim: "f ` topspace X = topspace Y" shows "normal_space Y" unfolding normal_space_def proof clarify fix S T assume "closedin Y S" and "closedin Y T" and "disjnt S T" have "closedin X {x \ topspace X. f x \ S}" "closedin X {x \ topspace X. f x \ T}" using \closedin Y S\ \closedin Y T\ closedin_continuous_map_preimage contf by auto moreover have "disjnt {x \ topspace X. f x \ S} {x \ topspace X. f x \ T}" using \disjnt S T\ by (auto simp: disjnt_iff) ultimately obtain U V where "closedin X U" "closedin X V" and subXU: "{x \ topspace X. f x \ S} \ topspace X - U" and subXV: "{x \ topspace X. f x \ T} \ topspace X - V" and dis: "disjnt (topspace X - U) (topspace X -V)" using \normal_space X\ by (force simp add: normal_space_def ex_openin) have "closedin Y (f ` U)" "closedin Y (f ` V)" using \closedin X U\ \closedin X V\ clof closed_map_def by blast+ moreover have "S \ topspace Y - f ` U" using \closedin Y S\ \closedin X U\ subXU by (force dest: closedin_subset) moreover have "T \ topspace Y - f ` V" using \closedin Y T\ \closedin X V\ subXV by (force dest: closedin_subset) moreover have "disjnt (topspace Y - f ` U) (topspace Y - f ` V)" using fim dis by (force simp add: disjnt_iff) ultimately show "\U V. openin Y U \ openin Y V \ S \ U \ T \ V \ disjnt U V" by (force simp add: ex_openin) qed subsection \Hereditary topological properties\ definition hereditarily where "hereditarily P X \ \S. S \ topspace X \ P(subtopology X S)" lemma hereditarily: "hereditarily P X \ (\S. P(subtopology X S))" by (metis Int_lower1 hereditarily_def subtopology_restrict) lemma hereditarily_mono: "\hereditarily P X; \x. P x \ Q x\ \ hereditarily Q X" by (simp add: hereditarily) lemma hereditarily_inc: "hereditarily P X \ P X" by (metis hereditarily subtopology_topspace) lemma hereditarily_subtopology: "hereditarily P X \ hereditarily P (subtopology X S)" by (simp add: hereditarily subtopology_subtopology) lemma hereditarily_normal_space_continuous_closed_map_image: assumes X: "hereditarily normal_space X" and contf: "continuous_map X Y f" and clof: "closed_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "hereditarily normal_space Y" unfolding hereditarily_def proof (intro strip) fix T assume "T \ topspace Y" then have nx: "normal_space (subtopology X {x \ topspace X. f x \ T})" by (meson X hereditarily) moreover have "continuous_map (subtopology X {x \ topspace X. f x \ T}) (subtopology Y T) f" by (simp add: contf continuous_map_from_subtopology continuous_map_in_subtopology image_subset_iff) moreover have "closed_map (subtopology X {x \ topspace X. f x \ T}) (subtopology Y T) f" by (simp add: clof closed_map_restriction) ultimately show "normal_space (subtopology Y T)" using fim normal_space_continuous_closed_map_image by fastforce qed lemma homeomorphic_hereditarily_normal_space: "X homeomorphic_space Y \ (hereditarily normal_space X \ hereditarily normal_space Y)" by (meson hereditarily_normal_space_continuous_closed_map_image homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym) lemma hereditarily_normal_space_retraction_map_image: "\retraction_map X Y r; hereditarily normal_space X\ \ hereditarily normal_space Y" by (smt (verit) hereditarily_subtopology hereditary_imp_retractive_property homeomorphic_hereditarily_normal_space) subsection\Limits in a topological space\ lemma limitin_const_iff: assumes "t1_space X" "\ trivial_limit F" shows "limitin X (\k. a) l F \ l \ topspace X \ a = l" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using assms unfolding limitin_def t1_space_def by (metis eventually_const openin_topspace) next assume ?rhs then show ?lhs using assms by (auto simp: limitin_def t1_space_def) qed lemma compactin_sequence_with_limit: assumes lim: "limitin X \ l sequentially" and "S \ range \" and SX: "S \ topspace X" shows "compactin X (insert l S)" unfolding compactin_def proof (intro conjI strip) show "insert l S \ topspace X" by (meson SX insert_subset lim limitin_topspace) fix \ assume \
: "Ball \ (openin X) \ insert l S \ \ \" have "\V. finite V \ V \ \ \ (\t \ V. l \ t) \ S \ \ V" if *: "\x \ S. \T \ \. x \ T" and "T \ \" "l \ T" for T proof - obtain V where V: "\x. x \ S \ V x \ \ \ x \ V x" using * by metis obtain N where N: "\n. N \ n \ \ n \ T" by (meson "\
" \T \ \\ \l \ T\ lim limitin_sequentially) show ?thesis proof (intro conjI exI) have "x \ T" if "x \ S" and "\A. (\x \ S. (\n\N. x \ \ n) \ A \ V x) \ x \ A" for x by (metis (no_types) N V that assms(2) imageE nle_le subsetD) then show "S \ \ (insert T (V ` (S \ \ ` {0..N})))" by force qed (use V that in auto) qed then show "\\. finite \ \ \ \ \ \ insert l S \ \ \" by (smt (verit, best) Union_iff \
insert_subset subsetD) qed lemma limitin_Hausdorff_unique: assumes "limitin X f l1 F" "limitin X f l2 F" "\ trivial_limit F" "Hausdorff_space X" shows "l1 = l2" proof (rule ccontr) assume "l1 \ l2" with assms obtain U V where "openin X U" "openin X V" "l1 \ U" "l2 \ V" "disjnt U V" by (metis Hausdorff_space_def limitin_topspace) then have "eventually (\x. f x \ U) F" "eventually (\x. f x \ V) F" using assms by (fastforce simp: limitin_def)+ then have "\x. f x \ U \ f x \ V" using assms eventually_elim2 filter_eq_iff by fastforce with assms \disjnt U V\ show False by (meson disjnt_iff) qed lemma limitin_kc_unique: assumes "kc_space X" and lim1: "limitin X f l1 sequentially" and lim2: "limitin X f l2 sequentially" shows "l1 = l2" proof (rule ccontr) assume "l1 \ l2" define A where "A \ insert l1 (range f - {l2})" have "l1 \ topspace X" using lim1 limitin_def by fastforce moreover have "compactin X (insert l1 (topspace X \ (range f - {l2})))" by (meson Diff_subset compactin_sequence_with_limit inf_le1 inf_le2 lim1 subset_trans) ultimately have "compactin X (topspace X \ A)" by (simp add: A_def) then have OXA: "openin X (topspace X - A)" by (metis Diff_Diff_Int Diff_subset \kc_space X\ kc_space_def openin_closedin_eq) have "l2 \ topspace X - A" using \l1 \ l2\ A_def lim2 limitin_topspace by fastforce then have "\\<^sub>F x in sequentially. f x = l2" using limitinD [OF lim2 OXA] by (auto simp: A_def eventually_conj_iff) then show False using limitin_transform_eventually [OF _ lim1] limitin_const_iff [OF kc_imp_t1_space trivial_limit_sequentially] using \l1 \ l2\ \kc_space X\ by fastforce qed lemma limitin_closedin: assumes lim: "limitin X f l F" and "closedin X S" and ev: "eventually (\x. f x \ S) F" "\ trivial_limit F" shows "l \ S" proof (rule ccontr) assume "l \ S" have "\\<^sub>F x in F. f x \ topspace X - S" by (metis Diff_iff \l \ S\ \closedin X S\ closedin_def lim limitin_def) with ev eventually_elim2 trivial_limit_def show False by force qed end diff --git a/src/HOL/Analysis/Abstract_Topology.thy b/src/HOL/Analysis/Abstract_Topology.thy --- a/src/HOL/Analysis/Abstract_Topology.thy +++ b/src/HOL/Analysis/Abstract_Topology.thy @@ -1,4986 +1,4994 @@ (* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) section \Operators involving abstract topology\ theory Abstract_Topology imports Complex_Main "HOL-Library.Set_Idioms" "HOL-Library.FuncSet" begin subsection \General notion of a topology as a value\ definition\<^marker>\tag important\ istopology :: "('a set \ bool) \ bool" where "istopology L \ (\S T. L S \ L T \ L (S \ T)) \ (\\. (\K\\. L K) \ L (\\))" typedef\<^marker>\tag important\ 'a topology = "{L::('a set) \ bool. istopology L}" morphisms "openin" "topology" unfolding istopology_def by blast lemma istopology_openin[iff]: "istopology(openin U)" using openin[of U] by blast lemma istopology_open[iff]: "istopology open" by (auto simp: istopology_def) lemma topology_inverse' [simp]: "istopology U \ openin (topology U) = U" using topology_inverse[unfolded mem_Collect_eq] . lemma topology_inverse_iff: "istopology U \ openin (topology U) = U" by (metis istopology_openin topology_inverse') lemma topology_eq: "T1 = T2 \ (\S. openin T1 S \ openin T2 S)" proof assume "T1 = T2" then show "\S. openin T1 S \ openin T2 S" by simp next assume H: "\S. openin T1 S \ openin T2 S" then have "openin T1 = openin T2" by (simp add: fun_eq_iff) then have "topology (openin T1) = topology (openin T2)" by simp then show "T1 = T2" unfolding openin_inverse . qed text\The "universe": the union of all sets in the topology.\ definition "topspace T = \{S. openin T S}" subsubsection \Main properties of open sets\ proposition openin_clauses: fixes U :: "'a topology" shows "openin U {}" "\S T. openin U S \ openin U T \ openin U (S\T)" "\K. (\S \ K. openin U S) \ openin U (\K)" using openin[of U] unfolding istopology_def by auto lemma openin_subset: "openin U S \ S \ topspace U" unfolding topspace_def by blast lemma openin_empty[simp]: "openin U {}" by (rule openin_clauses) lemma openin_Int[intro]: "openin U S \ openin U T \ openin U (S \ T)" by (rule openin_clauses) lemma openin_Union[intro]: "(\S. S \ K \ openin U S) \ openin U (\K)" using openin_clauses by blast lemma openin_Un[intro]: "openin U S \ openin U T \ openin U (S \ T)" using openin_Union[of "{S,T}" U] by auto lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (force simp: openin_Union topspace_def) lemma openin_subopen: "openin U S \ (\x \ S. \T. openin U T \ x \ T \ T \ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs by auto next assume H: ?rhs let ?t = "\{T. openin U T \ T \ S}" have "openin U ?t" by (force simp: openin_Union) also have "?t = S" using H by auto finally show "openin U S" . qed lemma openin_INT [intro]: assumes "finite I" "\i. i \ I \ openin T (U i)" shows "openin T ((\i \ I. U i) \ topspace T)" using assms by (induct, auto simp: inf_sup_aci(2) openin_Int) lemma openin_INT2 [intro]: assumes "finite I" "I \ {}" "\i. i \ I \ openin T (U i)" shows "openin T (\i \ I. U i)" proof - have "(\i \ I. U i) \ topspace T" using \I \ {}\ openin_subset[OF assms(3)] by auto then show ?thesis using openin_INT[of _ _ U, OF assms(1) assms(3)] by (simp add: inf.absorb2 inf_commute) qed lemma openin_Inter [intro]: assumes "finite \" "\ \ {}" "\X. X \ \ \ openin T X" shows "openin T (\\)" by (metis (full_types) assms openin_INT2 image_ident) lemma openin_Int_Inter: assumes "finite \" "openin T U" "\X. X \ \ \ openin T X" shows "openin T (U \ \\)" using openin_Inter [of "insert U \"] assms by auto subsubsection \Closed sets\ definition\<^marker>\tag important\ closedin :: "'a topology \ 'a set \ bool" where "closedin U S \ S \ topspace U \ openin U (topspace U - S)" lemma closedin_subset: "closedin U S \ S \ topspace U" by (metis closedin_def) lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def) lemma closedin_topspace[intro, simp]: "closedin U (topspace U)" by (simp add: closedin_def) lemma closedin_Un[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" by (auto simp: Diff_Un closedin_def) lemma Diff_Inter[intro]: "A - \S = \{A - s|s. s\S}" by auto lemma closedin_Union: assumes "finite S" "\T. T \ S \ closedin U T" shows "closedin U (\S)" using assms by induction auto lemma closedin_Inter[intro]: assumes Ke: "K \ {}" and Kc: "\S. S \K \ closedin U S" shows "closedin U (\K)" using Ke Kc unfolding closedin_def Diff_Inter by auto lemma closedin_INT[intro]: assumes "A \ {}" "\x. x \ A \ closedin U (B x)" shows "closedin U (\x\A. B x)" using assms by blast lemma closedin_Int[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" using closedin_Inter[of "{S,T}" U] by auto lemma openin_closedin_eq: "openin U S \ S \ topspace U \ closedin U (topspace U - S)" by (metis Diff_subset closedin_def double_diff equalityD1 openin_subset) lemma topology_finer_closedin: "topspace X = topspace Y \ (\S. openin Y S \ openin X S) \ (\S. closedin Y S \ closedin X S)" by (metis closedin_def openin_closedin_eq) lemma openin_closedin: "S \ topspace U \ (openin U S \ closedin U (topspace U - S))" by (simp add: openin_closedin_eq) lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)" by (metis Int_Diff cT closedin_def inf.orderE oS openin_Int openin_subset) lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)" by (metis Int_Diff cT closedin_Int closedin_subset inf.orderE oS openin_closedin_eq) lemma all_openin: "(\U. openin X U \ P U) \ (\U. closedin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) lemma all_closedin: "(\U. closedin X U \ P U) \ (\U. openin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_subset inf.absorb_iff2 openin_closedin_eq) lemma ex_openin: "(\U. openin X U \ P U) \ (\U. closedin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) lemma ex_closedin: "(\U. closedin X U \ P U) \ (\U. openin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_subset inf.absorb_iff2 openin_closedin_eq) subsection\The discrete topology\ definition discrete_topology where "discrete_topology U \ topology (\S. S \ U)" lemma openin_discrete_topology [simp]: "openin (discrete_topology U) S \ S \ U" proof - have "istopology (\S. S \ U)" by (auto simp: istopology_def) then show ?thesis by (simp add: discrete_topology_def topology_inverse') qed lemma topspace_discrete_topology [simp]: "topspace(discrete_topology U) = U" by (meson openin_discrete_topology openin_subset openin_topspace order_refl subset_antisym) lemma closedin_discrete_topology [simp]: "closedin (discrete_topology U) S \ S \ U" by (simp add: closedin_def) lemma discrete_topology_unique: "discrete_topology U = X \ topspace X = U \ (\x \ U. openin X {x})" (is "?lhs = ?rhs") proof assume R: ?rhs then have "openin X S" if "S \ U" for S using openin_subopen subsetD that by fastforce then show ?lhs by (metis R openin_discrete_topology openin_subset topology_eq) qed auto lemma discrete_topology_unique_alt: "discrete_topology U = X \ topspace X \ U \ (\x \ U. openin X {x})" using openin_subset by (auto simp: discrete_topology_unique) lemma subtopology_eq_discrete_topology_empty: "X = discrete_topology {} \ topspace X = {}" using discrete_topology_unique [of "{}" X] by auto lemma subtopology_eq_discrete_topology_sing: "X = discrete_topology {a} \ topspace X = {a}" by (metis discrete_topology_unique openin_topspace singletonD) subsection \Subspace topology\ definition\<^marker>\tag important\ subtopology :: "'a topology \ 'a set \ 'a topology" where "subtopology U V = topology (\T. \S. T = S \ V \ openin U S)" lemma istopology_subtopology: "istopology (\T. \S. T = S \ V \ openin U S)" (is "istopology ?L") proof - have "?L {}" by blast { fix A B assume A: "?L A" and B: "?L B" from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \ V" and Sb: "openin U Sb" "B = Sb \ V" by blast have "A \ B = (Sa \ Sb) \ V" "openin U (Sa \ Sb)" using Sa Sb by blast+ then have "?L (A \ B)" by blast } moreover { fix K assume K: "K \ Collect ?L" have th0: "Collect ?L = (\S. S \ V) ` Collect (openin U)" by blast from K[unfolded th0 subset_image_iff] obtain Sk where Sk: "Sk \ Collect (openin U)" "K = (\S. S \ V) ` Sk" by blast have "\K = (\Sk) \ V" using Sk by auto moreover have "openin U (\Sk)" using Sk by (auto simp: subset_eq) ultimately have "?L (\K)" by blast } ultimately show ?thesis unfolding subset_eq mem_Collect_eq istopology_def by auto qed lemma openin_subtopology: "openin (subtopology U V) S \ (\T. openin U T \ S = T \ V)" unfolding subtopology_def topology_inverse'[OF istopology_subtopology] by auto lemma openin_subtopology_Int: "openin X S \ openin (subtopology X T) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_Int2: "openin X T \ openin (subtopology X S) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_diff_closed: "\S \ topspace X; closedin X T\ \ openin (subtopology X S) (S - T)" unfolding closedin_def openin_subtopology by (rule_tac x="topspace X - T" in exI) auto lemma openin_relative_to: "(openin X relative_to S) = openin (subtopology X S)" by (force simp: relative_to_def openin_subtopology) lemma topspace_subtopology [simp]: "topspace (subtopology U V) = topspace U \ V" by (auto simp: topspace_def openin_subtopology) lemma topspace_subtopology_subset: "S \ topspace X \ topspace(subtopology X S) = S" by (simp add: inf.absorb_iff2) lemma closedin_subtopology: "closedin (subtopology U V) S \ (\T. closedin U T \ S = T \ V)" unfolding closedin_def topspace_subtopology by (auto simp: openin_subtopology) +lemma closedin_subset_topspace: + "\closedin X S; S \ T\ \ closedin (subtopology X T) S" + using closedin_subtopology by fastforce + lemma closedin_relative_to: "(closedin X relative_to S) = closedin (subtopology X S)" by (force simp: relative_to_def closedin_subtopology) lemma openin_subtopology_refl: "openin (subtopology U V) V \ V \ topspace U" unfolding openin_subtopology by auto (metis IntD1 in_mono openin_subset) lemma subtopology_subtopology: "subtopology (subtopology X S) T = subtopology X (S \ T)" proof - have eq: "\T'. (\S'. T' = S' \ T \ (\T. openin X T \ S' = T \ S)) = (\Sa. T' = Sa \ (S \ T) \ openin X Sa)" by (metis inf_assoc) have "subtopology (subtopology X S) T = topology (\Ta. \Sa. Ta = Sa \ T \ openin (subtopology X S) Sa)" by (simp add: subtopology_def) also have "\ = subtopology X (S \ T)" by (simp add: openin_subtopology eq) (simp add: subtopology_def) finally show ?thesis . qed lemma openin_subtopology_alt: "openin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (openin X)" by (simp add: image_iff inf_commute openin_subtopology) lemma closedin_subtopology_alt: "closedin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (closedin X)" by (simp add: image_iff inf_commute closedin_subtopology) lemma subtopology_superset: assumes UV: "topspace U \ V" shows "subtopology U V = U" proof - { fix S have "openin U S" if "openin U T" "S = T \ V" for T by (metis Int_subset_iff assms inf.orderE openin_subset that) then have "(\T. openin U T \ S = T \ V) \ openin U S" by (metis assms inf.orderE inf_assoc openin_subset) } then show ?thesis unfolding topology_eq openin_subtopology by blast qed lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U" by (simp add: subtopology_superset) lemma subtopology_UNIV[simp]: "subtopology U UNIV = U" by (simp add: subtopology_superset) lemma subtopology_restrict: "subtopology X (topspace X \ S) = subtopology X S" by (metis subtopology_subtopology subtopology_topspace) lemma openin_subtopology_empty: "openin (subtopology U {}) S \ S = {}" by (metis Int_empty_right openin_empty openin_subtopology) lemma closedin_subtopology_empty: "closedin (subtopology U {}) S \ S = {}" by (metis Int_empty_right closedin_empty closedin_subtopology) lemma closedin_subtopology_refl [simp]: "closedin (subtopology U X) X \ X \ topspace U" by (metis closedin_def closedin_topspace inf.absorb_iff2 le_inf_iff topspace_subtopology) lemma closedin_topspace_empty: "topspace T = {} \ (closedin T S \ S = {})" by (simp add: closedin_def) lemma open_in_topspace_empty: "topspace X = {} \ openin X S \ S = {}" by (simp add: openin_closedin_eq) lemma openin_imp_subset: "openin (subtopology U S) T \ T \ S" by (metis Int_iff openin_subtopology subsetI) lemma closedin_imp_subset: "closedin (subtopology U S) T \ T \ S" by (simp add: closedin_def) lemma openin_open_subtopology: "openin X S \ openin (subtopology X S) T \ openin X T \ T \ S" by (metis inf.orderE openin_Int openin_imp_subset openin_subtopology) lemma closedin_closed_subtopology: "closedin X S \ (closedin (subtopology X S) T \ closedin X T \ T \ S)" by (metis closedin_Int closedin_imp_subset closedin_subtopology inf.orderE) lemma closedin_trans_full: "\closedin (subtopology X U) S; closedin X U\ \ closedin X S" using closedin_closed_subtopology by blast lemma openin_subtopology_Un: "\openin (subtopology X T) S; openin (subtopology X U) S\ \ openin (subtopology X (T \ U)) S" by (simp add: openin_subtopology) blast lemma closedin_subtopology_Un: "\closedin (subtopology X T) S; closedin (subtopology X U) S\ \ closedin (subtopology X (T \ U)) S" by (simp add: closedin_subtopology) blast lemma openin_trans_full: "\openin (subtopology X U) S; openin X U\ \ openin X S" by (simp add: openin_open_subtopology) subsection \The canonical topology from the underlying type class\ abbreviation\<^marker>\tag important\ euclidean :: "'a::topological_space topology" where "euclidean \ topology open" abbreviation top_of_set :: "'a::topological_space set \ 'a topology" where "top_of_set \ subtopology (topology open)" lemma open_openin: "open S \ openin euclidean S" by simp declare open_openin [symmetric, simp] lemma topspace_euclidean [simp]: "topspace euclidean = UNIV" by (force simp: topspace_def) lemma topspace_euclidean_subtopology[simp]: "topspace (top_of_set S) = S" by (simp) lemma closed_closedin: "closed S \ closedin euclidean S" by (simp add: closed_def closedin_def Compl_eq_Diff_UNIV) declare closed_closedin [symmetric, simp] lemma openin_subtopology_self [simp]: "openin (top_of_set S) S" by (metis openin_topspace topspace_euclidean_subtopology) subsubsection\The most basic facts about the usual topology and metric on R\ abbreviation euclideanreal :: "real topology" where "euclideanreal \ topology open" subsection \Basic "localization" results are handy for connectedness.\ lemma openin_open: "openin (top_of_set U) S \ (\T. open T \ (S = U \ T))" by (auto simp: openin_subtopology) lemma openin_Int_open: "\openin (top_of_set U) S; open T\ \ openin (top_of_set U) (S \ T)" by (metis open_Int Int_assoc openin_open) lemma openin_open_Int[intro]: "open S \ openin (top_of_set U) (U \ S)" by (auto simp: openin_open) lemma open_openin_trans[trans]: "open S \ open T \ T \ S \ openin (top_of_set S) T" by (metis Int_absorb1 openin_open_Int) lemma open_subset: "S \ T \ open S \ openin (top_of_set T) S" by (auto simp: openin_open) lemma closedin_closed: "closedin (top_of_set U) S \ (\T. closed T \ S = U \ T)" by (simp add: closedin_subtopology Int_ac) lemma closedin_closed_Int: "closed S \ closedin (top_of_set U) (U \ S)" by (metis closedin_closed) lemma closed_subset: "S \ T \ closed S \ closedin (top_of_set T) S" by (auto simp: closedin_closed) lemma closedin_closed_subset: "\closedin (top_of_set U) V; T \ U; S = V \ T\ \ closedin (top_of_set T) S" by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE) lemma finite_imp_closedin: fixes S :: "'a::t1_space set" shows "\finite S; S \ T\ \ closedin (top_of_set T) S" by (simp add: finite_imp_closed closed_subset) lemma closedin_singleton [simp]: fixes a :: "'a::t1_space" shows "closedin (top_of_set U) {a} \ a \ U" using closedin_subset by (force intro: closed_subset) lemma openin_euclidean_subtopology_iff: fixes S U :: "'a::metric_space set" shows "openin (top_of_set U) S \ S \ U \ (\x\S. \e>0. \x'\U. dist x' x < e \ x'\ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs unfolding openin_open open_dist by blast next define T where "T = {x. \a\S. \d>0. (\y\U. dist y a < d \ y \ S) \ dist x a < d}" have 1: "\x\T. \e>0. \y. dist y x < e \ y \ T" unfolding T_def apply clarsimp apply (rule_tac x="d - dist x a" in exI) by (metis add_0_left dist_commute dist_triangle_lt less_diff_eq) assume ?rhs then have 2: "S = U \ T" unfolding T_def by auto (metis dist_self) from 1 2 show ?lhs unfolding openin_open open_dist by fast qed lemma connected_openin: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_def openin_open disjoint_iff_not_equal by blast lemma connected_openin_eq: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_openin by (metis (no_types, lifting) Un_subset_iff openin_imp_subset subset_antisym) lemma connected_closedin: "connected S \ (\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp add: connected_closed closedin_closed) next assume R: ?rhs then show ?lhs proof (clarsimp simp add: connected_closed closedin_closed) fix A B assume s_sub: "S \ A \ B" "B \ S \ {}" and disj: "A \ B \ S = {}" and cl: "closed A" "closed B" have "S - A = B \ S" using Diff_subset_conv Un_Diff_Int disj s_sub(1) by auto then show "A \ S = {}" by (metis Int_Diff_Un Int_Diff_disjoint R cl closedin_closed_Int dual_order.refl inf_commute s_sub(2)) qed qed lemma connected_closedin_eq: "connected S \ \(\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_closedin by (metis Un_subset_iff closedin_imp_subset subset_antisym) text \These "transitivity" results are handy too\ lemma openin_trans[trans]: "openin (top_of_set T) S \ openin (top_of_set U) T \ openin (top_of_set U) S" by (metis openin_Int_open openin_open) lemma openin_open_trans: "openin (top_of_set T) S \ open T \ open S" by (auto simp: openin_open intro: openin_trans) lemma closedin_trans[trans]: "closedin (top_of_set T) S \ closedin (top_of_set U) T \ closedin (top_of_set U) S" by (auto simp: closedin_closed closed_Inter Int_assoc) lemma closedin_closed_trans: "closedin (top_of_set T) S \ closed T \ closed S" by (auto simp: closedin_closed intro: closedin_trans) lemma openin_subtopology_Int_subset: "\openin (top_of_set u) (u \ S); v \ u\ \ openin (top_of_set v) (v \ S)" by (auto simp: openin_subtopology) lemma openin_open_eq: "open s \ (openin (top_of_set s) t \ open t \ t \ s)" using open_subset openin_open_trans openin_subset by fastforce subsection\Derived set (set of limit points)\ definition derived_set_of :: "'a topology \ 'a set \ 'a set" (infixl "derived'_set'_of" 80) where "X derived_set_of S \ {x \ topspace X. (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))}" lemma derived_set_of_restrict [simp]: "X derived_set_of (topspace X \ S) = X derived_set_of S" by (simp add: derived_set_of_def) (metis openin_subset subset_iff) lemma in_derived_set_of: "x \ X derived_set_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))" by (simp add: derived_set_of_def) lemma derived_set_of_subset_topspace: "X derived_set_of S \ topspace X" by (auto simp add: derived_set_of_def) lemma derived_set_of_subtopology: "(subtopology X U) derived_set_of S = U \ (X derived_set_of (U \ S))" by (simp add: derived_set_of_def openin_subtopology) blast lemma derived_set_of_subset_subtopology: "(subtopology X S) derived_set_of T \ S" by (simp add: derived_set_of_subtopology) lemma derived_set_of_empty [simp]: "X derived_set_of {} = {}" by (auto simp: derived_set_of_def) lemma derived_set_of_mono: "S \ T \ X derived_set_of S \ X derived_set_of T" unfolding derived_set_of_def by blast lemma derived_set_of_Un: "X derived_set_of (S \ T) = X derived_set_of S \ X derived_set_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (clarsimp simp: in_derived_set_of) (metis IntE IntI openin_Int) show "?rhs \ ?lhs" by (simp add: derived_set_of_mono) qed lemma derived_set_of_Union: "finite \ \ X derived_set_of (\\) = (\S \ \. X derived_set_of S)" proof (induction \ rule: finite_induct) case (insert S \) then show ?case by (simp add: derived_set_of_Un) qed auto lemma derived_set_of_topspace: "X derived_set_of (topspace X) = {x \ topspace X. \ openin X {x}}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (auto simp: in_derived_set_of) show "?rhs \ ?lhs" by (clarsimp simp: in_derived_set_of) (metis openin_closedin_eq openin_subopen singletonD subset_eq) qed lemma discrete_topology_unique_derived_set: "discrete_topology U = X \ topspace X = U \ X derived_set_of U = {}" by (auto simp: discrete_topology_unique derived_set_of_topspace) lemma subtopology_eq_discrete_topology_eq: "subtopology X U = discrete_topology U \ U \ topspace X \ U \ X derived_set_of U = {}" using discrete_topology_unique_derived_set [of U "subtopology X U"] by (auto simp: eq_commute derived_set_of_subtopology) lemma subtopology_eq_discrete_topology: "S \ topspace X \ S \ X derived_set_of S = {} \ subtopology X S = discrete_topology S" by (simp add: subtopology_eq_discrete_topology_eq) lemma subtopology_eq_discrete_topology_gen: assumes "S \ X derived_set_of S = {}" shows "subtopology X S = discrete_topology(topspace X \ S)" proof - have "subtopology X S = subtopology X (topspace X \ S)" by (simp add: subtopology_restrict) then show ?thesis using assms by (simp add: inf.assoc subtopology_eq_discrete_topology_eq) qed lemma subtopology_discrete_topology [simp]: "subtopology (discrete_topology U) S = discrete_topology(U \ S)" proof - have "(\T. \Sa. T = Sa \ S \ Sa \ U) = (\Sa. Sa \ U \ Sa \ S)" by force then show ?thesis by (simp add: subtopology_def) (simp add: discrete_topology_def) qed lemma openin_Int_derived_set_of_subset: "openin X S \ S \ X derived_set_of T \ X derived_set_of (S \ T)" by (auto simp: derived_set_of_def) lemma openin_Int_derived_set_of_eq: assumes "openin X S" shows "S \ X derived_set_of T = S \ X derived_set_of (S \ T)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: assms openin_Int_derived_set_of_subset) show "?rhs \ ?lhs" by (metis derived_set_of_mono inf_commute inf_le1 inf_mono order_refl) qed subsection\ Closure with respect to a topological space\ definition closure_of :: "'a topology \ 'a set \ 'a set" (infixr "closure'_of" 80) where "X closure_of S \ {x \ topspace X. \T. x \ T \ openin X T \ (\y \ S. y \ T)}" lemma closure_of_restrict: "X closure_of S = X closure_of (topspace X \ S)" unfolding closure_of_def using openin_subset by blast lemma in_closure_of: "x \ X closure_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y. y \ S \ y \ T))" by (auto simp: closure_of_def) lemma closure_of: "X closure_of S = topspace X \ (S \ X derived_set_of S)" by (fastforce simp: in_closure_of in_derived_set_of) lemma closure_of_alt: "X closure_of S = topspace X \ S \ X derived_set_of S" using derived_set_of_subset_topspace [of X S] unfolding closure_of_def in_derived_set_of by safe (auto simp: in_derived_set_of) lemma derived_set_of_subset_closure_of: "X derived_set_of S \ X closure_of S" by (fastforce simp: closure_of_def in_derived_set_of) lemma closure_of_subtopology: "(subtopology X U) closure_of S = U \ (X closure_of (U \ S))" unfolding closure_of_def topspace_subtopology openin_subtopology by safe (metis (full_types) IntI Int_iff inf.commute)+ lemma closure_of_empty [simp]: "X closure_of {} = {}" by (simp add: closure_of_alt) lemma closure_of_topspace [simp]: "X closure_of topspace X = topspace X" by (simp add: closure_of) lemma closure_of_UNIV [simp]: "X closure_of UNIV = topspace X" by (simp add: closure_of) lemma closure_of_subset_topspace: "X closure_of S \ topspace X" by (simp add: closure_of) lemma closure_of_subset_subtopology: "(subtopology X S) closure_of T \ S" by (simp add: closure_of_subtopology) lemma closure_of_mono: "S \ T \ X closure_of S \ X closure_of T" by (fastforce simp add: closure_of_def) lemma closure_of_subtopology_subset: "(subtopology X U) closure_of S \ (X closure_of S)" unfolding closure_of_subtopology by clarsimp (meson closure_of_mono contra_subsetD inf.cobounded2) lemma closure_of_subtopology_mono: "T \ U \ (subtopology X T) closure_of S \ (subtopology X U) closure_of S" unfolding closure_of_subtopology by auto (meson closure_of_mono inf_mono subset_iff) lemma closure_of_Un [simp]: "X closure_of (S \ T) = X closure_of S \ X closure_of T" by (simp add: Un_assoc Un_left_commute closure_of_alt derived_set_of_Un inf_sup_distrib1) lemma closure_of_Union: "finite \ \ X closure_of (\\) = (\S \ \. X closure_of S)" by (induction \ rule: finite_induct) auto lemma closure_of_subset: "S \ topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_Int: "topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_eq: "S \ topspace X \ X closure_of S \ S \ closedin X S" proof - have "openin X (topspace X - S)" if "\x. \x \ topspace X; \T. x \ T \ openin X T \ S \ T \ {}\ \ x \ S" apply (subst openin_subopen) by (metis Diff_iff Diff_mono Diff_triv inf.commute openin_subset order_refl that) then show ?thesis by (auto simp: closedin_def closure_of_def disjoint_iff_not_equal) qed lemma closure_of_eq: "X closure_of S = S \ closedin X S" by (metis closure_of_subset closure_of_subset_eq closure_of_subset_topspace subset_antisym) lemma closedin_contains_derived_set: "closedin X S \ X derived_set_of S \ S \ S \ topspace X" proof (intro iffI conjI) show "closedin X S \ X derived_set_of S \ S" using closure_of_eq derived_set_of_subset_closure_of by fastforce show "closedin X S \ S \ topspace X" using closedin_subset by blast show "X derived_set_of S \ S \ S \ topspace X \ closedin X S" by (metis closure_of closure_of_eq inf.absorb_iff2 sup.orderE) qed lemma derived_set_subset_gen: "X derived_set_of S \ S \ closedin X (topspace X \ S)" by (simp add: closedin_contains_derived_set derived_set_of_subset_topspace) lemma derived_set_subset: "S \ topspace X \ (X derived_set_of S \ S \ closedin X S)" by (simp add: closedin_contains_derived_set) lemma closedin_derived_set: "closedin (subtopology X T) S \ S \ topspace X \ S \ T \ (\x. x \ X derived_set_of S \ x \ T \ x \ S)" by (auto simp: closedin_contains_derived_set derived_set_of_subtopology Int_absorb1) lemma closedin_Int_closure_of: "closedin (subtopology X S) T \ S \ X closure_of T = T" by (metis Int_left_absorb closure_of_eq closure_of_subtopology) lemma closure_of_closedin: "closedin X S \ X closure_of S = S" by (simp add: closure_of_eq) lemma closure_of_eq_diff: "X closure_of S = topspace X - \{T. openin X T \ disjnt S T}" by (auto simp: closure_of_def disjnt_iff) lemma closedin_closure_of [simp]: "closedin X (X closure_of S)" unfolding closure_of_eq_diff by blast lemma closure_of_closure_of [simp]: "X closure_of (X closure_of S) = X closure_of S" by (simp add: closure_of_eq) lemma closure_of_hull: assumes "S \ topspace X" shows "X closure_of S = (closedin X) hull S" by (metis assms closedin_closure_of closure_of_eq closure_of_mono closure_of_subset hull_unique) lemma closure_of_minimal: "\S \ T; closedin X T\ \ (X closure_of S) \ T" by (metis closure_of_eq closure_of_mono) lemma closure_of_minimal_eq: "\S \ topspace X; closedin X T\ \ (X closure_of S) \ T \ S \ T" by (meson closure_of_minimal closure_of_subset subset_trans) lemma closure_of_unique: "\S \ T; closedin X T; \T'. \S \ T'; closedin X T'\ \ T \ T'\ \ X closure_of S = T" by (meson closedin_closure_of closedin_subset closure_of_minimal closure_of_subset eq_iff order.trans) lemma closure_of_eq_empty_gen: "X closure_of S = {} \ disjnt (topspace X) S" unfolding disjnt_def closure_of_restrict [where S=S] using closure_of by fastforce lemma closure_of_eq_empty: "S \ topspace X \ X closure_of S = {} \ S = {}" using closure_of_subset by fastforce lemma openin_Int_closure_of_subset: assumes "openin X S" shows "S \ X closure_of T \ X closure_of (S \ T)" proof - have "S \ X derived_set_of T = S \ X derived_set_of (S \ T)" by (meson assms openin_Int_derived_set_of_eq) moreover have "S \ (S \ T) = S \ T" by fastforce ultimately show ?thesis by (metis closure_of_alt inf.cobounded2 inf_left_commute inf_sup_distrib1) qed lemma closure_of_openin_Int_closure_of: assumes "openin X S" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" proof show "X closure_of (S \ X closure_of T) \ X closure_of (S \ T)" by (simp add: assms closure_of_minimal openin_Int_closure_of_subset) next show "X closure_of (S \ T) \ X closure_of (S \ X closure_of T)" by (metis Int_subset_iff assms closure_of_alt closure_of_mono inf_mono openin_subset subset_refl sup.coboundedI1) qed lemma openin_Int_closure_of_eq: assumes "openin X S" shows "S \ X closure_of T = S \ X closure_of (S \ T)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: assms openin_Int_closure_of_subset) show "?rhs \ ?lhs" by (metis closure_of_mono inf_commute inf_le1 inf_mono order_refl) qed lemma openin_Int_closure_of_eq_empty: assumes "openin X S" shows "S \ X closure_of T = {} \ S \ T = {}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding disjoint_iff by (meson assms in_closure_of in_mono openin_subset) show "?rhs \ ?lhs" by (simp add: assms openin_Int_closure_of_eq) qed lemma closure_of_openin_Int_superset: "openin X S \ S \ X closure_of T \ X closure_of (S \ T) = X closure_of S" by (metis closure_of_openin_Int_closure_of inf.orderE) lemma closure_of_openin_subtopology_Int_closure_of: assumes S: "openin (subtopology X U) S" and "T \ U" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" (is "?lhs = ?rhs") proof obtain S0 where S0: "openin X S0" "S = S0 \ U" using assms by (auto simp: openin_subtopology) then show "?lhs \ ?rhs" proof - have "S0 \ X closure_of T = S0 \ X closure_of (S0 \ T)" by (meson S0(1) openin_Int_closure_of_eq) moreover have "S0 \ T = S0 \ U \ T" using \T \ U\ by fastforce ultimately have "S \ X closure_of T \ X closure_of (S \ T)" using S0(2) by auto then show ?thesis by (meson closedin_closure_of closure_of_minimal) qed next show "?rhs \ ?lhs" proof - have "T \ S \ T \ X derived_set_of T" by force then show ?thesis by (smt (verit, del_insts) Int_iff in_closure_of inf.orderE openin_subset subsetI) qed qed lemma closure_of_subtopology_open: "openin X U \ S \ U \ (subtopology X U) closure_of S = U \ X closure_of S" by (metis closure_of_subtopology inf_absorb2 openin_Int_closure_of_eq) lemma discrete_topology_closure_of: "(discrete_topology U) closure_of S = U \ S" by (metis closedin_discrete_topology closure_of_restrict closure_of_unique discrete_topology_unique inf_sup_ord(1) order_refl) text\ Interior with respect to a topological space. \ definition interior_of :: "'a topology \ 'a set \ 'a set" (infixr "interior'_of" 80) where "X interior_of S \ {x. \T. openin X T \ x \ T \ T \ S}" lemma interior_of_restrict: "X interior_of S = X interior_of (topspace X \ S)" using openin_subset by (auto simp: interior_of_def) lemma interior_of_eq: "(X interior_of S = S) \ openin X S" unfolding interior_of_def using openin_subopen by blast lemma interior_of_openin: "openin X S \ X interior_of S = S" by (simp add: interior_of_eq) lemma interior_of_empty [simp]: "X interior_of {} = {}" by (simp add: interior_of_eq) lemma interior_of_topspace [simp]: "X interior_of (topspace X) = topspace X" by (simp add: interior_of_eq) lemma openin_interior_of [simp]: "openin X (X interior_of S)" unfolding interior_of_def using openin_subopen by fastforce lemma interior_of_interior_of [simp]: "X interior_of X interior_of S = X interior_of S" by (simp add: interior_of_eq) lemma interior_of_subset: "X interior_of S \ S" by (auto simp: interior_of_def) lemma interior_of_subset_closure_of: "X interior_of S \ X closure_of S" by (metis closure_of_subset_Int dual_order.trans interior_of_restrict interior_of_subset) lemma subset_interior_of_eq: "S \ X interior_of S \ openin X S" by (metis interior_of_eq interior_of_subset subset_antisym) lemma interior_of_mono: "S \ T \ X interior_of S \ X interior_of T" by (auto simp: interior_of_def) lemma interior_of_maximal: "\T \ S; openin X T\ \ T \ X interior_of S" by (auto simp: interior_of_def) lemma interior_of_maximal_eq: "openin X T \ T \ X interior_of S \ T \ S" by (meson interior_of_maximal interior_of_subset order_trans) lemma interior_of_unique: "\T \ S; openin X T; \T'. \T' \ S; openin X T'\ \ T' \ T\ \ X interior_of S = T" by (simp add: interior_of_maximal_eq interior_of_subset subset_antisym) lemma interior_of_subset_topspace: "X interior_of S \ topspace X" by (simp add: openin_subset) lemma interior_of_subset_subtopology: "(subtopology X S) interior_of T \ S" by (meson openin_imp_subset openin_interior_of) lemma interior_of_Int: "X interior_of (S \ T) = X interior_of S \ X interior_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: interior_of_mono) show "?rhs \ ?lhs" by (meson inf_mono interior_of_maximal interior_of_subset openin_Int openin_interior_of) qed lemma interior_of_Inter_subset: "X interior_of (\\) \ (\S \ \. X interior_of S)" by (simp add: INT_greatest Inf_lower interior_of_mono) lemma union_interior_of_subset: "X interior_of S \ X interior_of T \ X interior_of (S \ T)" by (simp add: interior_of_mono) lemma interior_of_eq_empty: "X interior_of S = {} \ (\T. openin X T \ T \ S \ T = {})" by (metis bot.extremum_uniqueI interior_of_maximal interior_of_subset openin_interior_of) lemma interior_of_eq_empty_alt: "X interior_of S = {} \ (\T. openin X T \ T \ {} \ T - S \ {})" by (auto simp: interior_of_eq_empty) lemma interior_of_Union_openin_subsets: "\{T. openin X T \ T \ S} = X interior_of S" by (rule interior_of_unique [symmetric]) auto lemma interior_of_complement: "X interior_of (topspace X - S) = topspace X - X closure_of S" by (auto simp: interior_of_def closure_of_def) lemma interior_of_closure_of: "X interior_of S = topspace X - X closure_of (topspace X - S)" unfolding interior_of_complement [symmetric] by (metis Diff_Diff_Int interior_of_restrict) lemma closure_of_interior_of: "X closure_of S = topspace X - X interior_of (topspace X - S)" by (simp add: interior_of_complement Diff_Diff_Int closure_of) lemma closure_of_complement: "X closure_of (topspace X - S) = topspace X - X interior_of S" unfolding interior_of_def closure_of_def by (blast dest: openin_subset) lemma interior_of_eq_empty_complement: "X interior_of S = {} \ X closure_of (topspace X - S) = topspace X" using interior_of_subset_topspace [of X S] closure_of_complement by fastforce lemma closure_of_eq_topspace: "X closure_of S = topspace X \ X interior_of (topspace X - S) = {}" using closure_of_subset_topspace [of X S] interior_of_complement by fastforce lemma interior_of_subtopology_subset: "U \ X interior_of S \ (subtopology X U) interior_of S" by (auto simp: interior_of_def openin_subtopology) lemma interior_of_subtopology_subsets: "T \ U \ T \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis inf.absorb_iff2 interior_of_subtopology_subset subtopology_subtopology) lemma interior_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis dual_order.trans inf.orderE inf_commute interior_of_subset interior_of_subtopology_subsets) lemma interior_of_subtopology_open: assumes "openin X U" shows "(subtopology X U) interior_of S = U \ X interior_of S" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (meson assms interior_of_maximal interior_of_subset le_infI openin_interior_of openin_open_subtopology) show "?rhs \ ?lhs" by (simp add: interior_of_subtopology_subset) qed lemma dense_intersects_open: "X closure_of S = topspace X \ (\T. openin X T \ T \ {} \ S \ T \ {})" proof - have "X closure_of S = topspace X \ (topspace X - X interior_of (topspace X - S) = topspace X)" by (simp add: closure_of_interior_of) also have "\ \ X interior_of (topspace X - S) = {}" by (simp add: closure_of_complement interior_of_eq_empty_complement) also have "\ \ (\T. openin X T \ T \ {} \ S \ T \ {})" unfolding interior_of_eq_empty_alt using openin_subset by fastforce finally show ?thesis . qed lemma interior_of_closedin_union_empty_interior_of: assumes "closedin X S" and disj: "X interior_of T = {}" shows "X interior_of (S \ T) = X interior_of S" proof - have "X closure_of (topspace X - T) = topspace X" by (metis Diff_Diff_Int disj closure_of_eq_topspace closure_of_restrict interior_of_closure_of) then show ?thesis unfolding interior_of_closure_of by (metis Diff_Un Diff_subset assms(1) closedin_def closure_of_openin_Int_superset) qed lemma interior_of_union_eq_empty: "closedin X S \ (X interior_of (S \ T) = {} \ X interior_of S = {} \ X interior_of T = {})" by (metis interior_of_closedin_union_empty_interior_of le_sup_iff subset_empty union_interior_of_subset) lemma discrete_topology_interior_of [simp]: "(discrete_topology U) interior_of S = U \ S" by (simp add: interior_of_restrict [of _ S] interior_of_eq) subsection \Frontier with respect to topological space \ definition frontier_of :: "'a topology \ 'a set \ 'a set" (infixr "frontier'_of" 80) where "X frontier_of S \ X closure_of S - X interior_of S" lemma frontier_of_closures: "X frontier_of S = X closure_of S \ X closure_of (topspace X - S)" by (metis Diff_Diff_Int closure_of_complement closure_of_subset_topspace double_diff frontier_of_def interior_of_subset_closure_of) lemma interior_of_union_frontier_of [simp]: "X interior_of S \ X frontier_of S = X closure_of S" by (simp add: frontier_of_def interior_of_subset_closure_of subset_antisym) lemma frontier_of_restrict: "X frontier_of S = X frontier_of (topspace X \ S)" by (metis closure_of_restrict frontier_of_def interior_of_restrict) lemma closedin_frontier_of: "closedin X (X frontier_of S)" by (simp add: closedin_Int frontier_of_closures) lemma frontier_of_subset_topspace: "X frontier_of S \ topspace X" by (simp add: closedin_frontier_of closedin_subset) lemma frontier_of_subset_subtopology: "(subtopology X S) frontier_of T \ S" by (metis (no_types) closedin_derived_set closedin_frontier_of) lemma frontier_of_subtopology_subset: "U \ (subtopology X U) frontier_of S \ (X frontier_of S)" proof - have "U \ X interior_of S - subtopology X U interior_of S = {}" by (simp add: interior_of_subtopology_subset) moreover have "X closure_of S \ subtopology X U closure_of S = subtopology X U closure_of S" by (meson closure_of_subtopology_subset inf.absorb_iff2) ultimately show ?thesis unfolding frontier_of_def by blast qed lemma frontier_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X T) frontier_of S \ (subtopology X U) frontier_of S" by (simp add: frontier_of_def Diff_mono closure_of_subtopology_mono interior_of_subtopology_mono) lemma clopenin_eq_frontier_of: "closedin X S \ openin X S \ S \ topspace X \ X frontier_of S = {}" proof (cases "S \ topspace X") case True then show ?thesis by (metis Diff_eq_empty_iff closure_of_eq closure_of_subset_eq frontier_of_def interior_of_eq interior_of_subset interior_of_union_frontier_of sup_bot_right) next case False then show ?thesis by (simp add: frontier_of_closures openin_closedin_eq) qed lemma frontier_of_eq_empty: "S \ topspace X \ (X frontier_of S = {} \ closedin X S \ openin X S)" by (simp add: clopenin_eq_frontier_of) lemma frontier_of_openin: "openin X S \ X frontier_of S = X closure_of S - S" by (metis (no_types) frontier_of_def interior_of_eq) lemma frontier_of_openin_straddle_Int: assumes "openin X U" "U \ X frontier_of S \ {}" shows "U \ S \ {}" "U - S \ {}" proof - have "U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}" using assms by (simp add: frontier_of_closures) then show "U \ S \ {}" using assms openin_Int_closure_of_eq_empty by fastforce show "U - S \ {}" proof - have "\A. X closure_of (A - S) \ U \ {}" using \U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}\ by blast then have "\ U \ S" by (metis Diff_disjoint Diff_eq_empty_iff Int_Diff assms(1) inf_commute openin_Int_closure_of_eq_empty) then show ?thesis by blast qed qed lemma frontier_of_subset_closedin: "closedin X S \ (X frontier_of S) \ S" using closure_of_eq frontier_of_def by fastforce lemma frontier_of_empty [simp]: "X frontier_of {} = {}" by (simp add: frontier_of_def) lemma frontier_of_topspace [simp]: "X frontier_of topspace X = {}" by (simp add: frontier_of_def) lemma frontier_of_subset_eq: assumes "S \ topspace X" shows "(X frontier_of S) \ S \ closedin X S" proof show "X frontier_of S \ S \ closedin X S" by (metis assms closure_of_subset_eq interior_of_subset interior_of_union_frontier_of le_sup_iff) show "closedin X S \ X frontier_of S \ S" by (simp add: frontier_of_subset_closedin) qed lemma frontier_of_complement: "X frontier_of (topspace X - S) = X frontier_of S" by (metis Diff_Diff_Int closure_of_restrict frontier_of_closures inf_commute) lemma frontier_of_disjoint_eq: assumes "S \ topspace X" shows "((X frontier_of S) \ S = {} \ openin X S)" proof assume "X frontier_of S \ S = {}" then have "closedin X (topspace X - S)" using assms closure_of_subset frontier_of_def interior_of_eq interior_of_subset by fastforce then show "openin X S" using assms by (simp add: openin_closedin) next show "openin X S \ X frontier_of S \ S = {}" by (simp add: Diff_Diff_Int closedin_def frontier_of_openin inf.absorb_iff2 inf_commute) qed lemma frontier_of_disjoint_eq_alt: "S \ (topspace X - X frontier_of S) \ openin X S" proof (cases "S \ topspace X") case True show ?thesis using True frontier_of_disjoint_eq by auto next case False then show ?thesis by (meson Diff_subset openin_subset subset_trans) qed lemma frontier_of_Int: "X frontier_of (S \ T) = X closure_of (S \ T) \ (X frontier_of S \ X frontier_of T)" proof - have *: "U \ S \ U \ T \ U \ (S \ A \ T \ B) = U \ (A \ B)" for U S T A B :: "'a set" by blast show ?thesis by (simp add: frontier_of_closures closure_of_mono Diff_Int * flip: closure_of_Un) qed lemma frontier_of_Int_subset: "X frontier_of (S \ T) \ X frontier_of S \ X frontier_of T" by (simp add: frontier_of_Int) lemma frontier_of_Int_closedin: assumes "closedin X S" "closedin X T" shows "X frontier_of(S \ T) = X frontier_of S \ T \ S \ X frontier_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using assms by (force simp add: frontier_of_Int closedin_Int closure_of_closedin) show "?rhs \ ?lhs" using assms frontier_of_subset_closedin by (auto simp add: frontier_of_Int closedin_Int closure_of_closedin) qed lemma frontier_of_Un_subset: "X frontier_of(S \ T) \ X frontier_of S \ X frontier_of T" by (metis Diff_Un frontier_of_Int_subset frontier_of_complement) lemma frontier_of_Union_subset: "finite \ \ X frontier_of (\\) \ (\T \ \. X frontier_of T)" proof (induction \ rule: finite_induct) case (insert A \) then show ?case using frontier_of_Un_subset by fastforce qed simp lemma frontier_of_frontier_of_subset: "X frontier_of (X frontier_of S) \ X frontier_of S" by (simp add: closedin_frontier_of frontier_of_subset_closedin) lemma frontier_of_subtopology_open: "openin X U \ (subtopology X U) frontier_of S = U \ X frontier_of S" by (simp add: Diff_Int_distrib closure_of_subtopology_open frontier_of_def interior_of_subtopology_open) lemma discrete_topology_frontier_of [simp]: "(discrete_topology U) frontier_of S = {}" by (simp add: Diff_eq discrete_topology_closure_of frontier_of_closures) subsection\Locally finite collections\ definition locally_finite_in where "locally_finite_in X \ \ (\\ \ topspace X) \ (\x \ topspace X. \V. openin X V \ x \ V \ finite {U \ \. U \ V \ {}})" lemma finite_imp_locally_finite_in: "\finite \; \\ \ topspace X\ \ locally_finite_in X \" by (auto simp: locally_finite_in_def) lemma locally_finite_in_subset: assumes "locally_finite_in X \" "\ \ \" shows "locally_finite_in X \" proof - have "finite (\ \ {U. U \ V \ {}}) \ finite (\ \ {U. U \ V \ {}})" for V by (meson \\ \ \\ finite_subset inf_le1 inf_le2 le_inf_iff subset_trans) then show ?thesis using assms unfolding locally_finite_in_def Int_def by fastforce qed lemma locally_finite_in_refinement: assumes \: "locally_finite_in X \" and f: "\S. S \ \ \ f S \ S" shows "locally_finite_in X (f ` \)" proof - show ?thesis unfolding locally_finite_in_def proof safe fix x assume "x \ topspace X" then obtain V where "openin X V" "x \ V" "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast moreover have "{U \ \. f U \ V \ {}} \ {U \ \. U \ V \ {}}" for V using f by blast ultimately have "finite {U \ \. f U \ V \ {}}" using finite_subset by blast moreover have "f ` {U \ \. f U \ V \ {}} = {U \ f ` \. U \ V \ {}}" by blast ultimately have "finite {U \ f ` \. U \ V \ {}}" by (metis (no_types, lifting) finite_imageI) then show "\V. openin X V \ x \ V \ finite {U \ f ` \. U \ V \ {}}" using \openin X V\ \x \ V\ by blast next show "\x xa. \xa \ \; x \ f xa\ \ x \ topspace X" by (meson Sup_upper \ f locally_finite_in_def subset_iff) qed qed lemma locally_finite_in_subtopology: assumes \: "locally_finite_in X \" "\\ \ S" shows "locally_finite_in (subtopology X S) \" unfolding locally_finite_in_def proof safe fix x assume x: "x \ topspace (subtopology X S)" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def topspace_subtopology by blast show "\V. openin (subtopology X S) V \ x \ V \ finite {U \ \. U \ V \ {}}" proof (intro exI conjI) show "openin (subtopology X S) (S \ V)" by (simp add: \openin X V\ openin_subtopology_Int2) have "{U \ \. U \ (S \ V) \ {}} \ {U \ \. U \ V \ {}}" by auto with fin show "finite {U \ \. U \ (S \ V) \ {}}" using finite_subset by auto show "x \ S \ V" using x \x \ V\ by (simp) qed next show "\x A. \x \ A; A \ \\ \ x \ topspace (subtopology X S)" using assms unfolding locally_finite_in_def topspace_subtopology by blast qed lemma closedin_locally_finite_Union: assumes clo: "\S. S \ \ \ closedin X S" and \: "locally_finite_in X \" shows "closedin X (\\)" using \ unfolding locally_finite_in_def closedin_def proof clarify show "openin X (topspace X - \\)" proof (subst openin_subopen, clarify) fix x assume "x \ topspace X" and "x \ \\" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast let ?T = "V - \{S \ \. S \ V \ {}}" show "\T. openin X T \ x \ T \ T \ topspace X - \\" proof (intro exI conjI) show "openin X ?T" by (metis (no_types, lifting) fin \openin X V\ clo closedin_Union mem_Collect_eq openin_diff) show "x \ ?T" using \x \ \\\ \x \ V\ by auto show "?T \ topspace X - \\" using \openin X V\ openin_subset by auto qed qed qed lemma locally_finite_in_closure: assumes \: "locally_finite_in X \" shows "locally_finite_in X ((\S. X closure_of S) ` \)" using \ unfolding locally_finite_in_def proof (intro conjI; clarsimp) fix x A assume "x \ X closure_of A" then show "x \ topspace X" by (meson in_closure_of) next fix x assume "x \ topspace X" and "\\ \ topspace X" then obtain V where V: "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast have eq: "{y \ f ` \. Q y} = f ` {x. x \ \ \ Q(f x)}" for f and Q :: "'a set \ bool" by blast have eq2: "{A \ \. X closure_of A \ V \ {}} = {A \ \. A \ V \ {}}" using openin_Int_closure_of_eq_empty V by blast have "finite {U \ (closure_of) X ` \. U \ V \ {}}" by (simp add: eq eq2 fin) with V show "\V. openin X V \ x \ V \ finite {U \ (closure_of) X ` \. U \ V \ {}}" by blast qed lemma closedin_Union_locally_finite_closure: "locally_finite_in X \ \ closedin X (\((\S. X closure_of S) ` \))" by (metis (mono_tags) closedin_closure_of closedin_locally_finite_Union imageE locally_finite_in_closure) lemma closure_of_Union_subset: "\((\S. X closure_of S) ` \) \ X closure_of (\\)" by (simp add: SUP_le_iff Sup_upper closure_of_mono) lemma closure_of_locally_finite_Union: assumes "locally_finite_in X \" shows "X closure_of (\\) = \((\S. X closure_of S) ` \)" proof (rule closure_of_unique) show "\ \ \ \ ((closure_of) X ` \)" using assms by (simp add: SUP_upper2 Sup_le_iff closure_of_subset locally_finite_in_def) show "closedin X (\ ((closure_of) X ` \))" using assms by (simp add: closedin_Union_locally_finite_closure) show "\T'. \\ \ \ T'; closedin X T'\ \ \ ((closure_of) X ` \) \ T'" by (simp add: Sup_le_iff closure_of_minimal) qed subsection\<^marker>\tag important\ \Continuous maps\ text \We will need to deal with continuous maps in terms of topologies and not in terms of type classes, as defined below.\ definition continuous_map where "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" lemma continuous_map: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" by (auto simp: continuous_map_def) lemma continuous_map_image_subset_topspace: "continuous_map X Y f \ f ` (topspace X) \ topspace Y" by (auto simp: continuous_map_def) lemma continuous_map_on_empty: "topspace X = {} \ continuous_map X Y f" by (auto simp: continuous_map_def) lemma continuous_map_closedin: "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" proof - have "(\U. openin Y U \ openin X {x \ topspace X. f x \ U}) = (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" if "\x. x \ topspace X \ f x \ topspace Y" proof - have eq: "{x \ topspace X. f x \ topspace Y \ f x \ C} = (topspace X - {x \ topspace X. f x \ C})" for C using that by blast show ?thesis proof (intro iffI allI impI) fix C assume "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and "closedin Y C" then show "closedin X {x \ topspace X. f x \ C}" by (auto simp add: closedin_def eq) next fix U assume "\C. closedin Y C \ closedin X {x \ topspace X. f x \ C}" and "openin Y U" then show "openin X {x \ topspace X. f x \ U}" by (auto simp add: openin_closedin_eq eq) qed qed then show ?thesis by (auto simp: continuous_map_def) qed lemma openin_continuous_map_preimage: "\continuous_map X Y f; openin Y U\ \ openin X {x \ topspace X. f x \ U}" by (simp add: continuous_map_def) lemma closedin_continuous_map_preimage: "\continuous_map X Y f; closedin Y C\ \ closedin X {x \ topspace X. f x \ C}" by (simp add: continuous_map_closedin) lemma openin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "openin X U" "openin Y V" shows "openin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) openin_closedin_eq by fastforce show ?thesis unfolding eq using assms openin_continuous_map_preimage by fastforce qed lemma closedin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "closedin X U" "closedin Y V" shows "closedin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) closedin_def by fastforce show ?thesis unfolding eq using assms closedin_continuous_map_preimage by fastforce qed lemma continuous_map_image_closure_subset: assumes "continuous_map X Y f" shows "f ` (X closure_of S) \ Y closure_of f ` S" proof - have *: "f ` (topspace X) \ topspace Y" by (meson assms continuous_map) have "X closure_of T \ {x \ X closure_of T. f x \ Y closure_of (f ` T)}" if "T \ topspace X" for T proof (rule closure_of_minimal) show "T \ {x \ X closure_of T. f x \ Y closure_of f ` T}" using closure_of_subset * that by (fastforce simp: in_closure_of) next show "closedin X {x \ X closure_of T. f x \ Y closure_of f ` T}" using assms closedin_continuous_map_preimage_gen by fastforce qed then show ?thesis by (smt (verit, ccfv_threshold) assms continuous_map image_eqI image_subset_iff in_closure_of mem_Collect_eq) qed lemma continuous_map_subset_aux1: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_image_closure_subset by blast lemma continuous_map_subset_aux2: assumes "\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S" shows "continuous_map X Y f" unfolding continuous_map_closedin proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "f x \ topspace Y" using assms closure_of_subset_topspace by fastforce next fix C assume "closedin Y C" then show "closedin X {x \ topspace X. f x \ C}" proof (clarsimp simp flip: closure_of_subset_eq, intro conjI) fix x assume x: "x \ X closure_of {x \ topspace X. f x \ C}" and "C \ topspace Y" and "Y closure_of C \ C" show "x \ topspace X" by (meson x in_closure_of) have "{a \ topspace X. f a \ C} \ topspace X" by simp moreover have "Y closure_of f ` {a \ topspace X. f a \ C} \ C" by (simp add: \closedin Y C\ closure_of_minimal image_subset_iff) ultimately show "f x \ C" using x assms by blast qed qed lemma continuous_map_eq_image_closure_subset: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_alt: "continuous_map X Y f \ (\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_gen: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 continuous_map_image_subset_topspace by metis lemma continuous_map_closure_preimage_subset: "continuous_map X Y f \ X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" unfolding continuous_map_closedin by (rule closure_of_minimal) (use in_closure_of in \fastforce+\) lemma continuous_map_frontier_frontier_preimage_subset: assumes "continuous_map X Y f" shows "X frontier_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y frontier_of T}" proof - have eq: "topspace X - {x \ topspace X. f x \ T} = {x \ topspace X. f x \ topspace Y - T}" using assms unfolding continuous_map_def by blast have "X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" by (simp add: assms continuous_map_closure_preimage_subset) moreover have "X closure_of (topspace X - {x \ topspace X. f x \ T}) \ {x \ topspace X. f x \ Y closure_of (topspace Y - T)}" using continuous_map_closure_preimage_subset [OF assms] eq by presburger ultimately show ?thesis by (auto simp: frontier_of_closures) qed lemma topology_finer_continuous_id: assumes "topspace X = topspace Y" shows "(\S. openin X S \ openin Y S) \ continuous_map Y X id" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding continuous_map_def using assms openin_subopen openin_subset by fastforce show "?rhs \ ?lhs" unfolding continuous_map_def using assms openin_subopen topspace_def by fastforce qed lemma continuous_map_const [simp]: "continuous_map X Y (\x. C) \ topspace X = {} \ C \ topspace Y" proof (cases "topspace X = {}") case False show ?thesis proof (cases "C \ topspace Y") case True with openin_subopen show ?thesis by (auto simp: continuous_map_def) next case False then show ?thesis unfolding continuous_map_def by fastforce qed qed (auto simp: continuous_map_on_empty) declare continuous_map_const [THEN iffD2, continuous_intros] lemma continuous_map_compose [continuous_intros]: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" shows "continuous_map X X'' (g \ f)" unfolding continuous_map_def proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "(g \ f) x \ topspace X''" using assms unfolding continuous_map_def by force next fix U assume "openin X'' U" have eq: "{x \ topspace X. (g \ f) x \ U} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U}}" by auto (meson f continuous_map_def) show "openin X {x \ topspace X. (g \ f) x \ U}" unfolding eq using assms unfolding continuous_map_def using \openin X'' U\ by blast qed lemma continuous_map_eq: assumes "continuous_map X X' f" and "\x. x \ topspace X \ f x = g x" shows "continuous_map X X' g" proof - have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. g x \ U}" for U using assms by auto show ?thesis using assms by (simp add: continuous_map_def eq) qed lemma restrict_continuous_map [simp]: "topspace X \ S \ continuous_map X X' (restrict f S) \ continuous_map X X' f" by (auto simp: elim!: continuous_map_eq) lemma continuous_map_in_subtopology: "continuous_map X (subtopology X' S) f \ continuous_map X X' f \ f ` (topspace X) \ S" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof - have "\A. f ` (X closure_of A) \ subtopology X' S closure_of f ` A" by (meson L continuous_map_image_closure_subset) then show ?thesis by (metis (no_types) closure_of_subset_subtopology closure_of_subtopology_subset closure_of_topspace continuous_map_eq_image_closure_subset order.trans) qed next assume R: ?rhs then have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. f x \ U \ f x \ S}" for U by auto show ?lhs using R unfolding continuous_map by (auto simp: openin_subtopology eq) qed lemma continuous_map_from_subtopology: "continuous_map X X' f \ continuous_map (subtopology X S) X' f" by (auto simp: continuous_map openin_subtopology) lemma continuous_map_into_fulltopology: "continuous_map X (subtopology X' T) f \ continuous_map X X' f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_into_subtopology: "\continuous_map X X' f; f ` topspace X \ T\ \ continuous_map X (subtopology X' T) f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_from_subtopology_mono: "\continuous_map (subtopology X T) X' f; S \ T\ \ continuous_map (subtopology X S) X' f" by (metis inf.absorb_iff2 continuous_map_from_subtopology subtopology_subtopology) lemma continuous_map_from_discrete_topology [simp]: "continuous_map (discrete_topology U) X f \ f ` U \ topspace X" by (auto simp: continuous_map_def) lemma continuous_map_iff_continuous [simp]: "continuous_map (top_of_set S) euclidean g = continuous_on S g" by (fastforce simp add: continuous_map openin_subtopology continuous_on_open_invariant) lemma continuous_map_iff_continuous2 [simp]: "continuous_map euclidean euclidean g = continuous_on UNIV g" by (metis continuous_map_iff_continuous subtopology_UNIV) lemma continuous_map_openin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X (topspace X \ f -` U))" by (auto simp: continuous_map_def vimage_def Int_def) lemma continuous_map_closedin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. closedin Y U \ closedin X (topspace X \ f -` U))" by (auto simp: continuous_map_closedin vimage_def Int_def) lemma continuous_map_square_root: "continuous_map euclideanreal euclideanreal sqrt" by (simp add: continuous_at_imp_continuous_on isCont_real_sqrt) lemma continuous_map_sqrt [continuous_intros]: "continuous_map X euclideanreal f \ continuous_map X euclideanreal (\x. sqrt(f x))" by (meson continuous_map_compose continuous_map_eq continuous_map_square_root o_apply) lemma continuous_map_id [simp, continuous_intros]: "continuous_map X X id" unfolding continuous_map_def using openin_subopen topspace_def by fastforce declare continuous_map_id [unfolded id_def, simp, continuous_intros] lemma continuous_map_id_subt [simp]: "continuous_map (subtopology X S) X id" by (simp add: continuous_map_from_subtopology) declare continuous_map_id_subt [unfolded id_def, simp] lemma\<^marker>\tag important\ continuous_map_alt: "continuous_map T1 T2 f = ((\U. openin T2 U \ openin T1 (f -` U \ topspace T1)) \ f ` topspace T1 \ topspace T2)" by (auto simp: continuous_map_def vimage_def image_def Collect_conj_eq inf_commute) lemma continuous_map_open [intro]: "continuous_map T1 T2 f \ openin T2 U \ openin T1 (f-`U \ topspace(T1))" unfolding continuous_map_alt by auto lemma continuous_map_preimage_topspace [intro]: assumes "continuous_map T1 T2 f" shows "f-`(topspace T2) \ topspace T1 = topspace T1" using assms unfolding continuous_map_def by auto subsection\Open and closed maps (not a priori assumed continuous)\ definition open_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "open_map X1 X2 f \ \U. openin X1 U \ openin X2 (f ` U)" definition closed_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "closed_map X1 X2 f \ \U. closedin X1 U \ closedin X2 (f ` U)" lemma open_map_imp_subset_topspace: "open_map X1 X2 f \ f ` (topspace X1) \ topspace X2" unfolding open_map_def by (simp add: openin_subset) lemma open_map_on_empty: "topspace X = {} \ open_map X Y f" by (metis empty_iff imageE in_mono open_map_def openin_subopen openin_subset) lemma closed_map_on_empty: "topspace X = {} \ closed_map X Y f" by (simp add: closed_map_def closedin_topspace_empty) lemma closed_map_const: "closed_map X Y (\x. c) \ topspace X = {} \ closedin Y {c}" by (metis closed_map_def closed_map_on_empty closedin_empty closedin_topspace image_constant_conv) lemma open_map_imp_subset: "\open_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" by (meson order_trans open_map_imp_subset_topspace subset_image_iff) lemma topology_finer_open_id: "(\S. openin X S \ openin X' S) \ open_map X X' id" unfolding open_map_def by auto lemma open_map_id: "open_map X X id" unfolding open_map_def by auto lemma open_map_eq: "\open_map X X' f; \x. x \ topspace X \ f x = g x\ \ open_map X X' g" unfolding open_map_def by (metis image_cong openin_subset subset_iff) lemma open_map_inclusion_eq: "open_map (subtopology X S) X id \ openin X (topspace X \ S)" by (metis openin_topspace openin_trans_full subtopology_restrict topology_finer_open_id topspace_subtopology) lemma open_map_inclusion: "openin X S \ open_map (subtopology X S) X id" by (simp add: open_map_inclusion_eq openin_Int) lemma open_map_compose: "\open_map X X' f; open_map X' X'' g\ \ open_map X X'' (g \ f)" by (metis (no_types, lifting) image_comp open_map_def) lemma closed_map_imp_subset_topspace: "closed_map X1 X2 f \ f ` (topspace X1) \ topspace X2" by (simp add: closed_map_def closedin_subset) lemma closed_map_imp_subset: "\closed_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" using closed_map_imp_subset_topspace by blast lemma topology_finer_closed_id: "(\S. closedin X S \ closedin X' S) \ closed_map X X' id" by (simp add: closed_map_def) lemma closed_map_id: "closed_map X X id" by (simp add: closed_map_def) lemma closed_map_eq: "\closed_map X X' f; \x. x \ topspace X \ f x = g x\ \ closed_map X X' g" unfolding closed_map_def by (metis image_cong closedin_subset subset_iff) lemma closed_map_compose: "\closed_map X X' f; closed_map X' X'' g\ \ closed_map X X'' (g \ f)" by (metis (no_types, lifting) closed_map_def image_comp) lemma closed_map_inclusion_eq: "closed_map (subtopology X S) X id \ closedin X (topspace X \ S)" proof - have *: "closedin X (T \ S)" if "closedin X (S \ topspace X)" "closedin X T" for T by (smt (verit, best) closedin_Int closure_of_subset_eq inf_sup_aci le_iff_inf that) then show ?thesis by (fastforce simp add: closed_map_def Int_commute closedin_subtopology_alt intro: *) qed lemma closed_map_inclusion: "closedin X S \ closed_map (subtopology X S) X id" by (simp add: closed_map_inclusion_eq closedin_Int) lemma open_map_into_subtopology: "\open_map X X' f; f ` topspace X \ S\ \ open_map X (subtopology X' S) f" unfolding open_map_def openin_subtopology using openin_subset by fastforce lemma closed_map_into_subtopology: "\closed_map X X' f; f ` topspace X \ S\ \ closed_map X (subtopology X' S) f" unfolding closed_map_def closedin_subtopology using closedin_subset by fastforce lemma open_map_into_discrete_topology: "open_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding open_map_def openin_discrete_topology using openin_subset by blast lemma closed_map_into_discrete_topology: "closed_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding closed_map_def closedin_discrete_topology using closedin_subset by blast lemma bijective_open_imp_closed_map: "\open_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ closed_map X X' f" unfolding open_map_def closed_map_def closedin_def by auto (metis Diff_subset inj_on_image_set_diff) lemma bijective_closed_imp_open_map: "\closed_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ open_map X X' f" unfolding closed_map_def open_map_def openin_closedin_eq by auto (metis Diff_subset inj_on_image_set_diff) lemma open_map_from_subtopology: "\open_map X X' f; openin X U\ \ open_map (subtopology X U) X' f" unfolding open_map_def openin_subtopology_alt by blast lemma closed_map_from_subtopology: "\closed_map X X' f; closedin X U\ \ closed_map (subtopology X U) X' f" unfolding closed_map_def closedin_subtopology_alt by blast lemma open_map_restriction: assumes f: "open_map X X' f" and U: "{x \ topspace X. f x \ V} = U" shows "open_map (subtopology X U) (subtopology X' V) f" unfolding open_map_def proof clarsimp fix W assume "openin (subtopology X U) W" then obtain T where "openin X T" "W = T \ U" by (meson openin_subtopology) with f U have "f ` W = (f ` T) \ V" unfolding open_map_def openin_closedin_eq by auto then show "openin (subtopology X' V) (f ` W)" by (metis \openin X T\ f open_map_def openin_subtopology_Int) qed lemma closed_map_restriction: assumes f: "closed_map X X' f" and U: "{x \ topspace X. f x \ V} = U" shows "closed_map (subtopology X U) (subtopology X' V) f" unfolding closed_map_def proof clarsimp fix W assume "closedin (subtopology X U) W" then obtain T where "closedin X T" "W = T \ U" by (meson closedin_subtopology) with f U have "f ` W = (f ` T) \ V" unfolding closed_map_def closedin_def by auto then show "closedin (subtopology X' V) (f ` W)" by (metis \closedin X T\ closed_map_def closedin_subtopology f) qed lemma closed_map_closure_of_image: "closed_map X Y f \ f ` topspace X \ topspace Y \ (\S. S \ topspace X \ Y closure_of (f ` S) \ image f (X closure_of S))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (simp add: closed_map_def closed_map_imp_subset_topspace closure_of_minimal closure_of_subset image_mono) next assume ?rhs then show ?lhs by (metis closed_map_def closed_map_into_discrete_topology closure_of_eq closure_of_subset_eq topspace_discrete_topology) qed lemma open_map_interior_of_image_subset: "open_map X Y f \ (\S. image f (X interior_of S) \ Y interior_of (f ` S))" by (metis image_mono interior_of_eq interior_of_maximal interior_of_subset open_map_def openin_interior_of subset_antisym) lemma open_map_interior_of_image_subset_alt: "open_map X Y f \ (\S\topspace X. f ` (X interior_of S) \ Y interior_of f ` S)" by (metis interior_of_eq open_map_def open_map_interior_of_image_subset openin_subset subset_interior_of_eq) lemma open_map_interior_of_image_subset_gen: "open_map X Y f \ (f ` topspace X \ topspace Y \ (\S. f ` (X interior_of S) \ Y interior_of f ` S))" by (meson open_map_imp_subset_topspace open_map_interior_of_image_subset) lemma open_map_preimage_neighbourhood: "open_map X Y f \ (f ` topspace X \ topspace Y \ (\U T. closedin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ (\V. closedin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U)))" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof (intro conjI strip) show "f ` topspace X \ topspace Y" by (simp add: \open_map X Y f\ open_map_imp_subset_topspace) next fix U T assume UT: "closedin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U" have "closedin Y (topspace Y - f ` (topspace X - U))" by (meson UT L open_map_def openin_closedin_eq openin_diff openin_topspace) with UT show "\V. closedin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" using image_iff by auto qed next assume R: ?rhs show ?lhs unfolding open_map_def proof (intro strip) fix U assume "openin X U" have "{x \ topspace X. f x \ topspace Y - f ` U} \ topspace X - U" by blast then obtain V where V: "closedin Y V" and sub: "topspace Y - f ` U \ V" "{x \ topspace X. f x \ V} \ topspace X - U" using R \openin X U\ by (meson Diff_subset openin_closedin_eq) then have "f ` U \ topspace Y - V" using R \openin X U\ openin_subset by fastforce with sub have "f ` U = topspace Y - V" by auto then show "openin Y (f ` U)" using V(1) by auto qed qed lemma closed_map_preimage_neighbourhood: "closed_map X Y f \ image f (topspace X) \ topspace Y \ (\U T. openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ (\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U))" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof (intro conjI strip) show "f ` topspace X \ topspace Y" by (simp add: L closed_map_imp_subset_topspace) next fix U T assume UT: "openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U" then have "openin Y (topspace Y - f ` (topspace X - U))" by (meson L closed_map_def closedin_def closedin_diff closedin_topspace) then show "\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" using UT image_iff by auto qed next assume R: ?rhs show ?lhs unfolding closed_map_def proof (intro strip) fix U assume "closedin X U" have "{x \ topspace X. f x \ topspace Y - f ` U} \ topspace X - U" by blast then obtain V where V: "openin Y V" and sub: "topspace Y - f ` U \ V" "{x \ topspace X. f x \ V} \ topspace X - U" using R Diff_subset \closedin X U\ unfolding closedin_def by (smt (verit, ccfv_threshold) Collect_mem_eq Collect_mono_iff) then have "f ` U \ topspace Y - V" using R \closedin X U\ closedin_subset by fastforce with sub have "f ` U = topspace Y - V" by auto with V show "closedin Y (f ` U)" by auto qed qed lemma closed_map_fibre_neighbourhood: "closed_map X Y f \ f ` (topspace X) \ topspace Y \ (\U y. openin X U \ y \ topspace Y \ {x \ topspace X. f x = y} \ U \ (\V. openin Y V \ y \ V \ {x \ topspace X. f x \ V} \ U))" unfolding closed_map_preimage_neighbourhood proof (intro conj_cong refl all_cong1) fix U assume "f ` topspace X \ topspace Y" show "(\T. openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ (\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U)) = (\y. openin X U \ y \ topspace Y \ {x \ topspace X. f x = y} \ U \ (\V. openin Y V \ y \ V \ {x \ topspace X. f x \ V} \ U))" (is "(\T. ?P T) \ (\y. ?Q y)") proof assume L [rule_format]: "\T. ?P T" show "\y. ?Q y" proof fix y show "?Q y" using L [of "{y}"] by blast qed next assume R: "\y. ?Q y" show "\T. ?P T" proof (cases "openin X U") case True note [[unify_search_bound=3]] obtain V where V: "\y. \y \ topspace Y; {x \ topspace X. f x = y} \ U\ \ openin Y (V y) \ y \ V y \ {x \ topspace X. f x \ V y} \ U" using R by (simp add: True) meson show ?thesis proof clarify fix T assume "openin X U" and "T \ topspace Y" and "{x \ topspace X. f x \ T} \ U" with V show "\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" by (rule_tac x="\y\T. V y" in exI) fastforce qed qed auto qed qed lemma open_map_in_subtopology: "openin Y S \ (open_map X (subtopology Y S) f \ open_map X Y f \ f ` (topspace X) \ S)" by (metis le_inf_iff open_map_def open_map_imp_subset_topspace open_map_into_subtopology openin_trans_full topspace_subtopology) lemma open_map_from_open_subtopology: "\openin Y S; open_map X (subtopology Y S) f\ \ open_map X Y f" using open_map_in_subtopology by blast lemma closed_map_in_subtopology: "closedin Y S \ closed_map X (subtopology Y S) f \ (closed_map X Y f \ f ` topspace X \ S)" by (metis closed_map_def closed_map_imp_subset_topspace closed_map_into_subtopology closedin_closed_subtopology closedin_subset topspace_subtopology_subset) lemma closed_map_from_closed_subtopology: "\closedin Y S; closed_map X (subtopology Y S) f\ \ closed_map X Y f" using closed_map_in_subtopology by blast lemma closed_map_from_composition_left: assumes cmf: "closed_map X Z (g \ f)" and contf: "continuous_map X Y f" and fim: "f ` topspace X = topspace Y" shows "closed_map Y Z g" unfolding closed_map_def proof (intro strip) fix U assume "closedin Y U" then have "closedin X {x \ topspace X. f x \ U}" using contf closedin_continuous_map_preimage by blast then have "closedin Z ((g \ f) ` {x \ topspace X. f x \ U})" using cmf closed_map_def by blast moreover have "\y. y \ U \ \x \ topspace X. f x \ U \ g y = g (f x)" by (smt (verit, ccfv_SIG) \closedin Y U\ closedin_subset fim image_iff subsetD) then have "(g \ f) ` {x \ topspace X. f x \ U} = g`U" by auto ultimately show "closedin Z (g ` U)" by metis qed text \identical proof as the above\ lemma open_map_from_composition_left: assumes cmf: "open_map X Z (g \ f)" and contf: "continuous_map X Y f" and fim: "f ` topspace X = topspace Y" shows "open_map Y Z g" unfolding open_map_def proof (intro strip) fix U assume "openin Y U" then have "openin X {x \ topspace X. f x \ U}" using contf openin_continuous_map_preimage by blast then have "openin Z ((g \ f) ` {x \ topspace X. f x \ U})" using cmf open_map_def by blast moreover have "\y. y \ U \ \x \ topspace X. f x \ U \ g y = g (f x)" by (smt (verit, ccfv_SIG) \openin Y U\ openin_subset fim image_iff subsetD) then have "(g \ f) ` {x \ topspace X. f x \ U} = g`U" by auto ultimately show "openin Z (g ` U)" by metis qed lemma closed_map_from_composition_right: assumes cmf: "closed_map X Z (g \ f)" "f ` topspace X \ topspace Y" "continuous_map Y Z g" "inj_on g (topspace Y)" shows "closed_map X Y f" unfolding closed_map_def proof (intro strip) fix C assume "closedin X C" have "\y c. \y \ topspace Y; g y = g (f c); c \ C\ \ y \ f ` C" using \closedin X C\ assms closedin_subset inj_onD by fastforce then have "f ` C = {x \ topspace Y. g x \ (g \ f) ` C}" using \closedin X C\ assms(2) closedin_subset by fastforce moreover have "closedin Z ((g \ f) ` C)" using \closedin X C\ cmf closed_map_def by blast ultimately show "closedin Y (f ` C)" using assms(3) closedin_continuous_map_preimage by fastforce qed text \identical proof as the above\ lemma open_map_from_composition_right: assumes "open_map X Z (g \ f)" "f ` topspace X \ topspace Y" "continuous_map Y Z g" "inj_on g (topspace Y)" shows "open_map X Y f" unfolding open_map_def proof (intro strip) fix C assume "openin X C" have "\y c. \y \ topspace Y; g y = g (f c); c \ C\ \ y \ f ` C" using \openin X C\ assms openin_subset inj_onD by fastforce then have "f ` C = {x \ topspace Y. g x \ (g \ f) ` C}" using \openin X C\ assms(2) openin_subset by fastforce moreover have "openin Z ((g \ f) ` C)" using \openin X C\ assms(1) open_map_def by blast ultimately show "openin Y (f ` C)" using assms(3) openin_continuous_map_preimage by fastforce qed subsection\Quotient maps\ definition quotient_map where "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (openin X {x. x \ topspace X \ f x \ U} \ openin X' U))" lemma quotient_map_eq: assumes "quotient_map X X' f" "\x. x \ topspace X \ f x = g x" shows "quotient_map X X' g" by (smt (verit) Collect_cong assms image_cong quotient_map_def) lemma quotient_map_compose: assumes f: "quotient_map X X' f" and g: "quotient_map X' X'' g" shows "quotient_map X X'' (g \ f)" unfolding quotient_map_def proof (intro conjI allI impI) show "(g \ f) ` topspace X = topspace X''" using assms by (simp only: image_comp [symmetric]) (simp add: quotient_map_def) next fix U'' assume U'': "U'' \ topspace X''" define U' where "U' \ {y \ topspace X'. g y \ U''}" have "U' \ topspace X'" by (auto simp add: U'_def) then have U': "openin X {x \ topspace X. f x \ U'} = openin X' U'" using assms unfolding quotient_map_def by simp have "{x \ topspace X. f x \ topspace X' \ g (f x) \ U''} = {x \ topspace X. (g \ f) x \ U''}" using f quotient_map_def by fastforce then show "openin X {x \ topspace X. (g \ f) x \ U''} = openin X'' U''" by (smt (verit, best) Collect_cong U' U'_def U'' g mem_Collect_eq quotient_map_def) qed lemma quotient_map_from_composition: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" and gf: "quotient_map X X'' (g \ f)" shows "quotient_map X' X'' g" unfolding quotient_map_def proof (intro conjI allI impI) show "g ` topspace X' = topspace X''" using assms unfolding continuous_map_def quotient_map_def by fastforce next fix U'' :: "'c set" assume U'': "U'' \ topspace X''" have eq: "{x \ topspace X. g (f x) \ U''} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U''}}" using continuous_map_def f by fastforce show "openin X' {x \ topspace X'. g x \ U''} = openin X'' U''" using assms unfolding continuous_map_def quotient_map_def by (metis (mono_tags, lifting) Collect_cong U'' comp_apply eq) qed lemma quotient_imp_continuous_map: "quotient_map X X' f \ continuous_map X X' f" by (simp add: continuous_map openin_subset quotient_map_def) lemma quotient_imp_surjective_map: "quotient_map X X' f \ f ` (topspace X) = topspace X'" by (simp add: quotient_map_def) lemma quotient_map_closedin: "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (closedin X {x. x \ topspace X \ f x \ U} \ closedin X' U))" proof - have eq: "(topspace X - {x \ topspace X. f x \ U'}) = {x \ topspace X. f x \ topspace X' \ f x \ U'}" if "f ` topspace X = topspace X'" "U' \ topspace X'" for U' using that by auto have "(\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U) = (\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U)" if "f ` topspace X = topspace X'" proof (rule iffI; intro allI impI subsetI) fix U' assume *[rule_format]: "\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U" and U': "U' \ topspace X'" show "closedin X {x \ topspace X. f x \ U'} = closedin X' U'" using U' by (auto simp add: closedin_def simp flip: * [of "topspace X' - U'"] eq [OF that]) next fix U' :: "'b set" assume *[rule_format]: "\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U" and U': "U' \ topspace X'" show "openin X {x \ topspace X. f x \ U'} = openin X' U'" using U' by (auto simp add: openin_closedin_eq simp flip: * [of "topspace X' - U'"] eq [OF that]) qed then show ?thesis unfolding quotient_map_def by force qed lemma continuous_open_imp_quotient_map: assumes "continuous_map X X' f" and om: "open_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - { fix U assume U: "U \ topspace X'" and "openin X {x \ topspace X. f x \ U}" then have ope: "openin X' (f ` {x \ topspace X. f x \ U})" using om unfolding open_map_def by blast then have "openin X' U" using U feq by (subst openin_subopen) force } moreover have "openin X {x \ topspace X. f x \ U}" if "U \ topspace X'" and "openin X' U" for U using that assms unfolding continuous_map_def by blast ultimately show ?thesis unfolding quotient_map_def using assms by blast qed lemma continuous_closed_imp_quotient_map: assumes "continuous_map X X' f" and om: "closed_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - have "f ` {x \ topspace X. f x \ U} = U" if "U \ topspace X'" for U using that feq by auto with assms show ?thesis unfolding quotient_map_closedin closed_map_def continuous_map_closedin by auto qed lemma continuous_open_quotient_map: "\continuous_map X X' f; open_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_open_imp_quotient_map quotient_map_def) lemma continuous_closed_quotient_map: "\continuous_map X X' f; closed_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_closed_imp_quotient_map quotient_map_def) lemma injective_quotient_map: assumes "inj_on f (topspace X)" shows "quotient_map X X' f \ continuous_map X X' f \ open_map X X' f \ closed_map X X' f \ f ` (topspace X) = topspace X'" (is "?lhs = ?rhs") proof assume L: ?lhs have om: "open_map X X' f" proof (clarsimp simp add: open_map_def) fix U assume "openin X U" then have "U \ topspace X" by (simp add: openin_subset) moreover have "{x \ topspace X. f x \ f ` U} = U" using \U \ topspace X\ assms inj_onD by fastforce ultimately show "openin X' (f ` U)" using L unfolding quotient_map_def by (metis (no_types, lifting) Collect_cong \openin X U\ image_mono) qed then have "closed_map X X' f" by (simp add: L assms bijective_open_imp_closed_map quotient_imp_surjective_map) then show ?rhs using L om by (simp add: quotient_imp_continuous_map quotient_imp_surjective_map) next assume ?rhs then show ?lhs by (simp add: continuous_closed_imp_quotient_map) qed lemma continuous_compose_quotient_map: assumes f: "quotient_map X X' f" and g: "continuous_map X X'' (g \ f)" shows "continuous_map X' X'' g" unfolding quotient_map_def continuous_map_def proof (intro conjI ballI allI impI) show "\x'. x' \ topspace X' \ g x' \ topspace X''" using assms unfolding quotient_map_def by (metis (no_types, opaque_lifting) continuous_map_image_subset_topspace image_comp image_subset_iff) next fix U'' :: "'c set" assume U'': "openin X'' U''" have "f ` topspace X = topspace X'" by (simp add: f quotient_imp_surjective_map) then have eq: "{x \ topspace X. f x \ topspace X' \ g (f x) \ U} = {x \ topspace X. g (f x) \ U}" for U by auto have "openin X {x \ topspace X. f x \ topspace X' \ g (f x) \ U''}" unfolding eq using U'' g openin_continuous_map_preimage by fastforce then have *: "openin X {x \ topspace X. f x \ {x \ topspace X'. g x \ U''}}" by auto show "openin X' {x \ topspace X'. g x \ U''}" using f unfolding quotient_map_def by (metis (no_types) Collect_subset *) qed lemma continuous_compose_quotient_map_eq: "quotient_map X X' f \ continuous_map X X'' (g \ f) \ continuous_map X' X'' g" using continuous_compose_quotient_map continuous_map_compose quotient_imp_continuous_map by blast lemma quotient_map_compose_eq: "quotient_map X X' f \ quotient_map X X'' (g \ f) \ quotient_map X' X'' g" by (meson continuous_compose_quotient_map_eq quotient_imp_continuous_map quotient_map_compose quotient_map_from_composition) lemma quotient_map_restriction: assumes quo: "quotient_map X Y f" and U: "{x \ topspace X. f x \ V} = U" and disj: "openin Y V \ closedin Y V" shows "quotient_map (subtopology X U) (subtopology Y V) f" using disj proof assume V: "openin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: openin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin X {x \ topspace X. f x \ U} = openin Y U" using quo unfolding quotient_map_def by auto have "openin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_def proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "openin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = openin (subtopology Y V) Y'" using U V Y \openin X U\ \Y' \ topspace Y\ \Y' \ V\ by (simp add: openin_open_subtopology eq) (auto simp: openin_closedin_eq) qed next assume V: "closedin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: closedin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ closedin X {x \ topspace X. f x \ U} = closedin Y U" using quo unfolding quotient_map_closedin by auto have "closedin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_closedin proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "closedin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = closedin (subtopology Y V) Y'" using U V Y \closedin X U\ \Y' \ topspace Y\ \Y' \ V\ by (simp add: closedin_closed_subtopology eq) (auto simp: closedin_def) qed qed lemma quotient_map_saturated_open: "quotient_map X Y f \ continuous_map X Y f \ f ` (topspace X) = topspace Y \ (\U. openin X U \ {x \ topspace X. f x \ f ` U} \ U \ openin Y (f ` U))" (is "?lhs = ?rhs") proof assume L: ?lhs then have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin Y U = openin X {x \ topspace X. f x \ U}" unfolding quotient_map_def by auto show ?rhs proof (intro conjI allI impI) show "continuous_map X Y f" by (simp add: L quotient_imp_continuous_map) show "f ` topspace X = topspace Y" by (simp add: fim) next fix U :: "'a set" assume U: "openin X U \ {x \ topspace X. f x \ f ` U} \ U" then have sub: "f ` U \ topspace Y" and eq: "{x \ topspace X. f x \ f ` U} = U" using fim openin_subset by fastforce+ show "openin Y (f ` U)" by (simp add: sub Y eq U) qed next assume ?rhs then have YX: "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and fim: "f ` topspace X = topspace Y" and XY: "\U. \openin X U; {x \ topspace X. f x \ f ` U} \ U\ \ openin Y (f ` U)" by (auto simp: quotient_map_def continuous_map_def) show ?lhs proof (simp add: quotient_map_def fim, intro allI impI iffI) fix U :: "'b set" assume "U \ topspace Y" and X: "openin X {x \ topspace X. f x \ U}" have feq: "f ` {x \ topspace X. f x \ U} = U" using \U \ topspace Y\ fim by auto show "openin Y U" using XY [OF X] by (simp add: feq) next fix U :: "'b set" assume "U \ topspace Y" and Y: "openin Y U" show "openin X {x \ topspace X. f x \ U}" by (metis YX [OF Y]) qed qed lemma quotient_map_saturated_closed: "quotient_map X Y f \ continuous_map X Y f \ f ` (topspace X) = topspace Y \ (\U. closedin X U \ {x \ topspace X. f x \ f ` U} \ U \ closedin Y (f ` U))" (is "?lhs = ?rhs") proof assume L: ?lhs then obtain fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ closedin Y U = closedin X {x \ topspace X. f x \ U}" by (simp add: quotient_map_closedin) show ?rhs proof (intro conjI allI impI) show "continuous_map X Y f" by (simp add: L quotient_imp_continuous_map) show "f ` topspace X = topspace Y" by (simp add: fim) next fix U :: "'a set" assume U: "closedin X U \ {x \ topspace X. f x \ f ` U} \ U" then have sub: "f ` U \ topspace Y" and eq: "{x \ topspace X. f x \ f ` U} = U" using fim closedin_subset by fastforce+ show "closedin Y (f ` U)" by (simp add: sub Y eq U) qed next assume ?rhs then obtain YX: "\U. closedin Y U \ closedin X {x \ topspace X. f x \ U}" and fim: "f ` topspace X = topspace Y" and XY: "\U. \closedin X U; {x \ topspace X. f x \ f ` U} \ U\ \ closedin Y (f ` U)" by (simp add: continuous_map_closedin) show ?lhs proof (simp add: quotient_map_closedin fim, intro allI impI iffI) fix U :: "'b set" assume "U \ topspace Y" and X: "closedin X {x \ topspace X. f x \ U}" have feq: "f ` {x \ topspace X. f x \ U} = U" using \U \ topspace Y\ fim by auto show "closedin Y U" using XY [OF X] by (simp add: feq) next fix U :: "'b set" assume "U \ topspace Y" and Y: "closedin Y U" show "closedin X {x \ topspace X. f x \ U}" by (metis YX [OF Y]) qed qed lemma quotient_map_onto_image: assumes "f ` topspace X \ topspace Y" and U: "\U. U \ topspace Y \ openin X {x \ topspace X. f x \ U} = openin Y U" shows "quotient_map X (subtopology Y (f ` topspace X)) f" unfolding quotient_map_def topspace_subtopology proof (intro conjI strip) fix U assume "U \ topspace Y \ f ` topspace X" with U have "openin X {x \ topspace X. f x \ U} \ \x. openin Y x \ U = f ` topspace X \ x" by fastforce moreover have "\x. openin Y x \ U = f ` topspace X \ x \ openin X {x \ topspace X. f x \ U}" by (metis (mono_tags, lifting) Collect_cong IntE IntI U image_eqI openin_subset) ultimately show "openin X {x \ topspace X. f x \ U} = openin (subtopology Y (f ` topspace X)) U" by (force simp: openin_subtopology_alt image_iff) qed (use assms in auto) lemma quotient_map_lift_exists: assumes f: "quotient_map X Y f" and h: "continuous_map X Z h" and "\x y. \x \ topspace X; y \ topspace X; f x = f y\ \ h x = h y" obtains g where "continuous_map Y Z g" "g ` topspace Y = h ` topspace X" "\x. x \ topspace X \ g(f x) = h x" proof - obtain g where g: "\x. x \ topspace X \ h x = g(f x)" using function_factors_left_gen[of "\x. x \ topspace X" f h] assms by blast show ?thesis proof show "g ` topspace Y = h ` topspace X" using f g by (force dest!: quotient_imp_surjective_map) show "continuous_map Y Z g" by (smt (verit) f g h continuous_compose_quotient_map_eq continuous_map_eq o_def) qed (simp add: g) qed subsection\ Separated Sets\ definition separatedin :: "'a topology \ 'a set \ 'a set \ bool" where "separatedin X S T \ S \ topspace X \ T \ topspace X \ S \ X closure_of T = {} \ T \ X closure_of S = {}" lemma separatedin_empty [simp]: "separatedin X S {} \ S \ topspace X" "separatedin X {} S \ S \ topspace X" by (simp_all add: separatedin_def) lemma separatedin_refl [simp]: "separatedin X S S \ S = {}" by (metis closure_of_subset empty_subsetI inf.orderE separatedin_def) lemma separatedin_sym: "separatedin X S T \ separatedin X T S" by (auto simp: separatedin_def) lemma separatedin_imp_disjoint: "separatedin X S T \ disjnt S T" by (meson closure_of_subset disjnt_def disjnt_subset2 separatedin_def) lemma separatedin_mono: "\separatedin X S T; S' \ S; T' \ T\ \ separatedin X S' T'" unfolding separatedin_def using closure_of_mono by blast lemma separatedin_open_sets: "\openin X S; openin X T\ \ separatedin X S T \ disjnt S T" unfolding disjnt_def separatedin_def by (auto simp: openin_Int_closure_of_eq_empty openin_subset) lemma separatedin_closed_sets: "\closedin X S; closedin X T\ \ separatedin X S T \ disjnt S T" unfolding closure_of_eq disjnt_def separatedin_def by (metis closedin_def closure_of_eq inf_commute) lemma separatedin_subtopology: "separatedin (subtopology X U) S T \ S \ U \ T \ U \ separatedin X S T" by (auto simp: separatedin_def closure_of_subtopology Int_ac disjoint_iff elim!: inf.orderE) lemma separatedin_discrete_topology: "separatedin (discrete_topology U) S T \ S \ U \ T \ U \ disjnt S T" by (metis openin_discrete_topology separatedin_def separatedin_open_sets topspace_discrete_topology) lemma separated_eq_distinguishable: "separatedin X {x} {y} \ x \ topspace X \ y \ topspace X \ (\U. openin X U \ x \ U \ (y \ U)) \ (\v. openin X v \ y \ v \ (x \ v))" by (force simp: separatedin_def closure_of_def) lemma separatedin_Un [simp]: "separatedin X S (T \ U) \ separatedin X S T \ separatedin X S U" "separatedin X (S \ T) U \ separatedin X S U \ separatedin X T U" by (auto simp: separatedin_def) lemma separatedin_Union: "finite \ \ separatedin X S (\\) \ S \ topspace X \ (\T \ \. separatedin X S T)" "finite \ \ separatedin X (\\) S \ (\T \ \. separatedin X S T) \ S \ topspace X" by (auto simp: separatedin_def closure_of_Union) lemma separatedin_openin_diff: "\openin X S; openin X T\ \ separatedin X (S - T) (T - S)" unfolding separatedin_def by (metis Diff_Int_distrib2 Diff_disjoint Diff_empty Diff_mono empty_Diff empty_subsetI openin_Int_closure_of_eq_empty openin_subset) lemma separatedin_closedin_diff: assumes "closedin X S" "closedin X T" shows "separatedin X (S - T) (T - S)" proof - have "S - T \ topspace X" "T - S \ topspace X" using assms closedin_subset by auto with assms show ?thesis by (simp add: separatedin_def Diff_Int_distrib2 closure_of_minimal inf_absorb2) qed lemma separation_closedin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ closedin (subtopology X (S \ T)) S \ closedin (subtopology X (S \ T)) T" by (auto simp add: separatedin_def closedin_Int_closure_of disjnt_iff dest: closure_of_subset) lemma separation_openin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ openin (subtopology X (S \ T)) S \ openin (subtopology X (S \ T)) T" unfolding openin_closedin_eq topspace_subtopology separation_closedin_Un_gen disjnt_def by (auto simp: Diff_triv Int_commute Un_Diff inf_absorb1 topspace_def) lemma separatedin_full: "S \ T = topspace X \ separatedin X S T \ disjnt S T \ closedin X S \ openin X S \ closedin X T \ openin X T" by (metis separatedin_open_sets separation_closedin_Un_gen separation_openin_Un_gen subtopology_topspace) subsection\Homeomorphisms\ text\(1-way and 2-way versions may be useful in places)\ definition homeomorphic_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "homeomorphic_map X Y f \ quotient_map X Y f \ inj_on f (topspace X)" definition homeomorphic_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "homeomorphic_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" lemma homeomorphic_map_eq: "\homeomorphic_map X Y f; \x. x \ topspace X \ f x = g x\ \ homeomorphic_map X Y g" by (meson homeomorphic_map_def inj_on_cong quotient_map_eq) lemma homeomorphic_maps_eq: "\homeomorphic_maps X Y f g; \x. x \ topspace X \ f x = f' x; \y. y \ topspace Y \ g y = g' y\ \ homeomorphic_maps X Y f' g'" unfolding homeomorphic_maps_def by (metis continuous_map_eq continuous_map_eq_image_closure_subset_gen image_subset_iff) lemma homeomorphic_maps_sym: "homeomorphic_maps X Y f g \ homeomorphic_maps Y X g f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_maps_id: "homeomorphic_maps X Y id id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have "topspace X = topspace Y" by (auto simp: homeomorphic_maps_def continuous_map_def) with L show ?rhs unfolding homeomorphic_maps_def by (metis topology_finer_continuous_id topology_eq) next assume ?rhs then show ?lhs unfolding homeomorphic_maps_def by auto qed lemma homeomorphic_map_id [simp]: "homeomorphic_map X Y id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have eq: "topspace X = topspace Y" by (auto simp: homeomorphic_map_def continuous_map_def quotient_map_def) then have "\S. openin X S \ openin Y S" by (meson L homeomorphic_map_def injective_quotient_map topology_finer_open_id) then show ?rhs using L unfolding homeomorphic_map_def by (metis eq quotient_imp_continuous_map topology_eq topology_finer_continuous_id) next assume ?rhs then show ?lhs unfolding homeomorphic_map_def by (simp add: closed_map_id continuous_closed_imp_quotient_map) qed lemma homeomorphic_map_compose: assumes "homeomorphic_map X Y f" "homeomorphic_map Y X'' g" shows "homeomorphic_map X X'' (g \ f)" proof - have "inj_on g (f ` topspace X)" by (metis (no_types) assms homeomorphic_map_def quotient_imp_surjective_map) then show ?thesis using assms by (meson comp_inj_on homeomorphic_map_def quotient_map_compose_eq) qed lemma homeomorphic_maps_compose: "homeomorphic_maps X Y f h \ homeomorphic_maps Y X'' g k \ homeomorphic_maps X X'' (g \ f) (h \ k)" unfolding homeomorphic_maps_def by (auto simp: continuous_map_compose; simp add: continuous_map_def) lemma homeomorphic_eq_everything_map: "homeomorphic_map X Y f \ continuous_map X Y f \ open_map X Y f \ closed_map X Y f \ f ` (topspace X) = topspace Y \ inj_on f (topspace X)" unfolding homeomorphic_map_def by (force simp: injective_quotient_map intro: injective_quotient_map) lemma homeomorphic_imp_continuous_map: "homeomorphic_map X Y f \ continuous_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_open_map: "homeomorphic_map X Y f \ open_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_closed_map: "homeomorphic_map X Y f \ closed_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_surjective_map: "homeomorphic_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_injective_map: "homeomorphic_map X Y f \ inj_on f (topspace X)" by (simp add: homeomorphic_eq_everything_map) lemma bijective_open_imp_homeomorphic_map: "\continuous_map X Y f; open_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: homeomorphic_map_def continuous_open_imp_quotient_map) lemma bijective_closed_imp_homeomorphic_map: "\continuous_map X Y f; closed_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: continuous_closed_quotient_map homeomorphic_map_def) lemma open_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "open_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y open_map_def continuous_map_def eq) qed lemma closed_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "closed_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "closedin X U" for U using closedin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y closed_map_def continuous_map_closedin eq) qed lemma homeomorphic_maps_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f \ homeomorphic_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "continuous_map Y X g" "\x\topspace X. g (f x) = x" "\x'\topspace Y. f (g x') = x'" by (auto simp: homeomorphic_maps_def) show ?rhs proof (intro conjI bijective_open_imp_homeomorphic_map L) show "open_map X Y f" using L using open_eq_continuous_inverse_map [of concl: X Y f g] by (simp add: continuous_map_def) show "open_map Y X g" using L using open_eq_continuous_inverse_map [of concl: Y X g f] by (simp add: continuous_map_def) show "f ` topspace X = topspace Y" "g ` topspace Y = topspace X" using L by (force simp: continuous_map_closedin)+ show "inj_on f (topspace X)" "inj_on g (topspace Y)" using L unfolding inj_on_def by metis+ qed next assume ?rhs then show ?lhs by (auto simp: homeomorphic_maps_def homeomorphic_imp_continuous_map) qed lemma homeomorphic_maps_imp_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f" using homeomorphic_maps_map by blast lemma homeomorphic_map_maps: "homeomorphic_map X Y f \ (\g. homeomorphic_maps X Y f g)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "open_map X Y f" "closed_map X Y f" "f ` (topspace X) = topspace Y" "inj_on f (topspace X)" by (auto simp: homeomorphic_eq_everything_map) have X: "\x. x \ topspace X \ f x \ topspace Y \ inv_into (topspace X) f (f x) = x" using L by auto have Y: "\y. y \ topspace Y \ inv_into (topspace X) f y \ topspace X \ f (inv_into (topspace X) f y) = y" by (simp add: L f_inv_into_f inv_into_into) have "homeomorphic_maps X Y f (inv_into (topspace X) f)" unfolding homeomorphic_maps_def proof (intro conjI L) show "continuous_map Y X (inv_into (topspace X) f)" by (simp add: L X Y flip: open_eq_continuous_inverse_map [where f=f]) next show "\x\topspace X. inv_into (topspace X) f (f x) = x" "\y\topspace Y. f (inv_into (topspace X) f y) = y" using X Y by auto qed then show ?rhs by metis next assume ?rhs then show ?lhs using homeomorphic_maps_map by blast qed lemma homeomorphic_maps_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_maps X X f f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_map_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_map X X f" using homeomorphic_maps_involution homeomorphic_maps_map by blast lemma homeomorphic_map_openness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "openin Y (f ` U) \ openin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "openin X U \ openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast show "openin Y (f ` U) = openin X U" proof assume L: "openin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "openin X U" by (metis L homeomorphic_imp_open_map open_map_def g) next assume "openin X U" then show "openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast qed qed lemma homeomorphic_map_closedness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "closedin Y (f ` U) \ closedin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "closedin X U \ closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast show "closedin Y (f ` U) = closedin X U" proof assume L: "closedin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "closedin X U" by (metis L homeomorphic_imp_closed_map closed_map_def g) next assume "closedin X U" then show "closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast qed qed lemma homeomorphic_map_openness_eq: "homeomorphic_map X Y f \ openin X U \ U \ topspace X \ openin Y (f ` U)" by (meson homeomorphic_map_openness openin_closedin_eq) lemma homeomorphic_map_closedness_eq: "homeomorphic_map X Y f \ closedin X U \ U \ topspace X \ closedin Y (f ` U)" by (meson closedin_subset homeomorphic_map_closedness) lemma all_openin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. openin Y V \ P V) \ (\U. openin X U \ P(f ` U))" by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_openness openin_subset subset_image_iff) lemma all_closedin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. closedin Y V \ P V) \ (\U. closedin X U \ P(f ` U))" (is "?lhs = ?rhs") by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_closedness closedin_subset subset_image_iff) lemma homeomorphic_map_derived_set_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y derived_set_of (f ` S) = f ` (X derived_set_of S)" proof - have fim: "f ` (topspace X) = topspace Y" and inj: "inj_on f (topspace X)" using hom by (auto simp: homeomorphic_eq_everything_map) have iff: "(\T. x \ T \ openin X T \ (\y. y \ x \ y \ S \ y \ T)) = (\T. T \ topspace Y \ f x \ T \ openin Y T \ (\y. y \ f x \ y \ f ` S \ y \ T))" if "x \ topspace X" for x proof - have \
: "(x \ T \ openin X T) = (T \ topspace X \ f x \ f ` T \ openin Y (f ` T))" for T by (meson hom homeomorphic_map_openness_eq inj inj_on_image_mem_iff that) moreover have "(\y. y \ x \ y \ S \ y \ T) = (\y. y \ f x \ y \ f ` S \ y \ f ` T)" (is "?lhs = ?rhs") if "T \ topspace X \ f x \ f ` T \ openin Y (f ` T)" for T by (smt (verit, del_insts) S \x \ topspace X\ image_iff inj inj_on_def subsetD that) ultimately show ?thesis by (auto simp flip: fim simp: all_subset_image) qed have *: "\T = f ` S; \x. x \ S \ P x \ Q(f x)\ \ {y. y \ T \ Q y} = f ` {x \ S. P x}" for T S P Q by auto show ?thesis unfolding derived_set_of_def by (rule *) (use fim iff openin_subset in force)+ qed lemma homeomorphic_map_closure_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y closure_of (f ` S) = f ` (X closure_of S)" unfolding closure_of using homeomorphic_imp_surjective_map [OF hom] S by (auto simp: in_derived_set_of homeomorphic_map_derived_set_of [OF assms]) lemma homeomorphic_map_interior_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y interior_of (f ` S) = f ` (X interior_of S)" proof - { fix y assume "y \ topspace Y" and "y \ Y closure_of (topspace Y - f ` S)" then have "y \ f ` (topspace X - X closure_of (topspace X - S))" using homeomorphic_eq_everything_map [THEN iffD1, OF hom] homeomorphic_map_closure_of [OF hom] by (metis DiffI Diff_subset S closure_of_subset_topspace inj_on_image_set_diff) } moreover { fix x assume "x \ topspace X" then have "f x \ topspace Y" using hom homeomorphic_imp_surjective_map by blast } moreover { fix x assume "x \ topspace X" and "x \ X closure_of (topspace X - S)" and "f x \ Y closure_of (topspace Y - f ` S)" then have "False" using homeomorphic_map_closure_of [OF hom] hom unfolding homeomorphic_eq_everything_map by (metis Diff_subset S closure_of_subset_topspace inj_on_image_mem_iff inj_on_image_set_diff) } ultimately show ?thesis by (auto simp: interior_of_closure_of) qed lemma homeomorphic_map_frontier_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y frontier_of (f ` S) = f ` (X frontier_of S)" unfolding frontier_of_def proof (intro equalityI subsetI DiffI) fix y assume "y \ Y closure_of f ` S - Y interior_of f ` S" then show "y \ f ` (X closure_of S - X interior_of S)" using S hom homeomorphic_map_closure_of homeomorphic_map_interior_of by fastforce next fix y assume "y \ f ` (X closure_of S - X interior_of S)" then show "y \ Y closure_of f ` S" using S hom homeomorphic_map_closure_of by fastforce next fix x assume "x \ f ` (X closure_of S - X interior_of S)" then obtain y where y: "x = f y" "y \ X closure_of S" "y \ X interior_of S" by blast then show "x \ Y interior_of f ` S" using S hom homeomorphic_map_interior_of y(1) unfolding homeomorphic_map_def by (smt (verit, ccfv_SIG) in_closure_of inj_on_image_mem_iff interior_of_subset_topspace) qed lemma homeomorphic_maps_subtopologies: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_maps_subtopologies_alt: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) \ T; g ` (topspace Y \ T) \ S\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_map_subtopologies: "\homeomorphic_map X Y f; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_map (subtopology X S) (subtopology Y T) f" by (meson homeomorphic_map_maps homeomorphic_maps_subtopologies) lemma homeomorphic_map_subtopologies_alt: assumes hom: "homeomorphic_map X Y f" and S: "\x. \x \ topspace X; f x \ topspace Y\ \ f x \ T \ x \ S" shows "homeomorphic_map (subtopology X S) (subtopology Y T) f" proof - have "homeomorphic_maps (subtopology X S) (subtopology Y T) f g" if "homeomorphic_maps X Y f g" for g proof (rule homeomorphic_maps_subtopologies [OF that]) have "f ` (topspace X \ S) \ topspace Y \ T" using S hom homeomorphic_imp_surjective_map by fastforce then show "f ` (topspace X \ S) = topspace Y \ T" using that unfolding homeomorphic_maps_def continuous_map_def by (smt (verit, del_insts) Int_iff S image_iff subsetI subset_antisym) qed then show ?thesis using hom by (meson homeomorphic_map_maps) qed subsection\Relation of homeomorphism between topological spaces\ definition homeomorphic_space (infixr "homeomorphic'_space" 50) where "X homeomorphic_space Y \ \f g. homeomorphic_maps X Y f g" lemma homeomorphic_space_refl: "X homeomorphic_space X" by (meson homeomorphic_maps_id homeomorphic_space_def) lemma homeomorphic_space_sym: "X homeomorphic_space Y \ Y homeomorphic_space X" unfolding homeomorphic_space_def by (metis homeomorphic_maps_sym) lemma homeomorphic_space_trans [trans]: "\X1 homeomorphic_space X2; X2 homeomorphic_space X3\ \ X1 homeomorphic_space X3" unfolding homeomorphic_space_def by (metis homeomorphic_maps_compose) lemma homeomorphic_space: "X homeomorphic_space Y \ (\f. homeomorphic_map X Y f)" by (simp add: homeomorphic_map_maps homeomorphic_space_def) lemma homeomorphic_maps_imp_homeomorphic_space: "homeomorphic_maps X Y f g \ X homeomorphic_space Y" unfolding homeomorphic_space_def by metis lemma homeomorphic_map_imp_homeomorphic_space: "homeomorphic_map X Y f \ X homeomorphic_space Y" unfolding homeomorphic_map_maps using homeomorphic_space_def by blast lemma homeomorphic_empty_space: "X homeomorphic_space Y \ topspace X = {} \ topspace Y = {}" by (metis homeomorphic_imp_surjective_map homeomorphic_space image_is_empty) lemma homeomorphic_empty_space_eq: assumes "topspace X = {}" shows "X homeomorphic_space Y \ topspace Y = {}" unfolding homeomorphic_maps_def homeomorphic_space_def by (metis assms continuous_map_on_empty continuous_map_closedin ex_in_conv) subsection\Connected topological spaces\ definition connected_space :: "'a topology \ bool" where "connected_space X \ \(\E1 E2. openin X E1 \ openin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" definition connectedin :: "'a topology \ 'a set \ bool" where "connectedin X S \ S \ topspace X \ connected_space (subtopology X S)" lemma connected_spaceD: "\connected_space X; openin X U; openin X V; topspace X \ U \ V; U \ V = {}; U \ {}; V \ {}\ \ False" by (auto simp: connected_space_def) lemma connectedin_subset_topspace: "connectedin X S \ S \ topspace X" by (simp add: connectedin_def) lemma connectedin_topspace: "connectedin X (topspace X) \ connected_space X" by (simp add: connectedin_def) lemma connected_space_subtopology: "connectedin X S \ connected_space (subtopology X S)" by (simp add: connectedin_def) lemma connectedin_subtopology: "connectedin (subtopology X S) T \ connectedin X T \ T \ S" by (force simp: connectedin_def subtopology_subtopology inf_absorb2) lemma connected_space_eq: "connected_space X \ (\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_space_def by (metis openin_Un openin_subset subset_antisym) lemma connected_space_closedin: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then have "\E1 E2. \openin X E1; E1 \ E2 = {}; topspace X \ E1 \ E2; openin X E2\ \ E1 = {} \ E2 = {}" by (simp add: connected_space_def) then show ?rhs unfolding connected_space_def by (metis disjnt_def separatedin_closed_sets separation_openin_Un_gen subtopology_superset) next assume R: ?rhs then show ?lhs unfolding connected_space_def by (metis Diff_triv Int_commute separatedin_openin_diff separation_closedin_Un_gen subtopology_superset) qed lemma connected_space_closedin_eq: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" by (metis closedin_Un closedin_def connected_space_closedin subset_antisym) lemma connected_space_clopen_in: "connected_space X \ (\T. openin X T \ closedin X T \ T = {} \ T = topspace X)" proof - have eq: "openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ P \ E2 = topspace X - E1 \ openin X E1 \ openin X E2 \ P" for E1 E2 P using openin_subset by blast show ?thesis unfolding connected_space_eq eq closedin_def by (auto simp: openin_closedin_eq) qed lemma connectedin: "connectedin X S \ S \ topspace X \ (\E1 E2. openin X E1 \ openin X E2 \ S \ E1 \ E2 \ E1 \ E2 \ S = {} \ E1 \ S \ {} \ E2 \ S \ {})" (is "?lhs = ?rhs") proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_def openin_subtopology topspace_subtopology * by (intro conj_cong arg_cong [where f=Not] ex_cong1; blast dest!: openin_subset) qed lemma connectedinD: "\connectedin X S; openin X E1; openin X E2; S \ E1 \ E2; E1 \ E2 \ S = {}; E1 \ S \ {}; E2 \ S \ {}\ \ False" by (meson connectedin) lemma connectedin_iff_connected [simp]: "connectedin euclidean S \ connected S" by (simp add: connected_def connectedin) lemma connectedin_closedin: "connectedin X S \ S \ topspace X \ \(\E1 E2. closedin X E1 \ closedin X E2 \ S \ (E1 \ E2) \ (E1 \ E2 \ S = {}) \ \(E1 \ S = {}) \ \(E2 \ S = {}))" proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_closedin closedin_subtopology topspace_subtopology * by (intro conj_cong arg_cong [where f=Not] ex_cong1; blast dest!: openin_subset) qed lemma connectedin_empty [simp]: "connectedin X {}" by (simp add: connectedin) lemma connected_space_topspace_empty: "topspace X = {} \ connected_space X" using connectedin_topspace by fastforce lemma connectedin_sing [simp]: "connectedin X {a} \ a \ topspace X" by (simp add: connectedin) lemma connectedin_absolute [simp]: "connectedin (subtopology X S) S \ connectedin X S" by (simp add: connectedin_subtopology) lemma connectedin_Union: assumes \: "\S. S \ \ \ connectedin X S" and ne: "\\ \ {}" shows "connectedin X (\\)" proof - have "\\ \ topspace X" using \ by (simp add: Union_least connectedin_def) moreover have False if "openin X E1" "openin X E2" and cover: "\\ \ E1 \ E2" and disj: "E1 \ E2 \ \\ = {}" and overlap1: "E1 \ \\ \ {}" and overlap2: "E2 \ \\ \ {}" for E1 E2 proof - have disjS: "E1 \ E2 \ S = {}" if "S \ \" for S using Diff_triv that disj by auto have coverS: "S \ E1 \ E2" if "S \ \" for S using that cover by blast have "\ \ {}" using overlap1 by blast obtain a where a: "\U. U \ \ \ a \ U" using ne by force with \\ \ {}\ have "a \ \\" by blast then consider "a \ E1" | "a \ E2" using \\\ \ E1 \ E2\ by auto then show False proof cases case 1 then obtain b S where "b \ E2" "b \ S" "S \ \" using overlap2 by blast then show ?thesis using "1" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) next case 2 then obtain b S where "b \ E1" "b \ S" "S \ \" using overlap1 by blast then show ?thesis using "2" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) qed qed ultimately show ?thesis unfolding connectedin by blast qed lemma connectedin_Un: "\connectedin X S; connectedin X T; S \ T \ {}\ \ connectedin X (S \ T)" using connectedin_Union [of "{S,T}"] by auto lemma connected_space_subconnected: "connected_space X \ (\x \ topspace X. \y \ topspace X. \S. connectedin X S \ x \ S \ y \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using connectedin_topspace by blast next assume R [rule_format]: ?rhs have False if "openin X U" "openin X V" and disj: "U \ V = {}" and cover: "topspace X \ U \ V" and "U \ {}" "V \ {}" for U V proof - obtain u v where "u \ U" "v \ V" using \U \ {}\ \V \ {}\ by auto then obtain T where "u \ T" "v \ T" and T: "connectedin X T" using R [of u v] that by (meson \openin X U\ \openin X V\ subsetD openin_subset) then show False using that unfolding connectedin by (metis IntI \u \ U\ \v \ V\ empty_iff inf_bot_left subset_trans) qed then show ?lhs by (auto simp: connected_space_def) qed lemma connectedin_intermediate_closure_of: assumes "connectedin X S" "S \ T" "T \ X closure_of S" shows "connectedin X T" proof - have S: "S \ topspace X" and T: "T \ topspace X" using assms by (meson closure_of_subset_topspace dual_order.trans)+ have \
: "\E1 E2. \openin X E1; openin X E2; E1 \ S = {} \ E2 \ S = {}\ \ E1 \ T = {} \ E2 \ T = {}" using assms unfolding disjoint_iff by (meson in_closure_of subsetD) then show ?thesis using assms unfolding connectedin closure_of_subset_topspace S T by (metis Int_empty_right T dual_order.trans inf.orderE inf_left_commute) qed lemma connectedin_closure_of: "connectedin X S \ connectedin X (X closure_of S)" by (meson closure_of_subset connectedin_def connectedin_intermediate_closure_of subset_refl) lemma connectedin_separation: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ C1 \ X closure_of C2 = {} \ C2 \ X closure_of C1 = {})" unfolding connectedin_def connected_space_closedin_eq closedin_Int_closure_of topspace_subtopology apply (intro conj_cong refl arg_cong [where f=Not]) apply (intro ex_cong1 iffI, blast) using closure_of_subset_Int by force lemma connectedin_eq_not_separated: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" unfolding separatedin_def by (metis connectedin_separation sup.boundedE) lemma connectedin_eq_not_separated_subset: "connectedin X S \ S \ topspace X \ (\C1 C2. S \ C1 \ C2 \ S \ C1 \ {} \ S \ C2 \ {} \ separatedin X C1 C2)" proof - have "\C1 C2. S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" if "\C1 C2. C1 \ C2 = S \ C1 = {} \ C2 = {} \ \ separatedin X C1 C2" proof (intro allI) fix C1 C2 show "S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" using that [of "S \ C1" "S \ C2"] by (auto simp: separatedin_mono) qed then show ?thesis by (metis Un_Int_eq(1) Un_Int_eq(2) connectedin_eq_not_separated order_refl) qed lemma connected_space_eq_not_separated: "connected_space X \ (\C1 C2. C1 \ C2 = topspace X \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" by (simp add: connectedin_eq_not_separated flip: connectedin_topspace) lemma connected_space_eq_not_separated_subset: "connected_space X \ (\C1 C2. topspace X \ C1 \ C2 \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" by (metis connected_space_eq_not_separated le_sup_iff separatedin_def subset_antisym) lemma connectedin_subset_separated_union: "\connectedin X C; separatedin X S T; C \ S \ T\ \ C \ S \ C \ T" unfolding connectedin_eq_not_separated_subset by blast lemma connectedin_nonseparated_union: assumes "connectedin X S" "connectedin X T" "\separatedin X S T" shows "connectedin X (S \ T)" proof - have "\C1 C2. \T \ C1 \ C2; S \ C1 \ C2\ \ S \ C1 = {} \ T \ C1 = {} \ S \ C2 = {} \ T \ C2 = {} \ \ separatedin X C1 C2" using assms unfolding connectedin_eq_not_separated_subset by (metis (no_types, lifting) assms connectedin_subset_separated_union inf.orderE separatedin_empty(1) separatedin_mono separatedin_sym) then show ?thesis unfolding connectedin_eq_not_separated_subset by (simp add: assms connectedin_subset_topspace Int_Un_distrib2) qed lemma connected_space_closures: "connected_space X \ (\e1 e2. e1 \ e2 = topspace X \ X closure_of e1 \ X closure_of e2 = {} \ e1 \ {} \ e2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding connected_space_closedin_eq by (metis Un_upper1 Un_upper2 closedin_closure_of closure_of_Un closure_of_eq_empty closure_of_topspace) next assume ?rhs then show ?lhs unfolding connected_space_closedin_eq by (metis closure_of_eq) qed lemma connectedin_Int_frontier_of: assumes "connectedin X S" "S \ T \ {}" "S - T \ {}" shows "S \ X frontier_of T \ {}" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) moreover have "S - (topspace X \ T) \ {}" using assms(3) by blast moreover have "S \ topspace X \ T \ {}" using assms connectedin by fastforce moreover have False if "S \ T \ {}" "S - T \ {}" "T \ topspace X" "S \ X frontier_of T = {}" for T proof - have null: "S \ (X closure_of T - X interior_of T) = {}" using that unfolding frontier_of_def by blast have "X interior_of T \ (topspace X - X closure_of T) \ S = {}" by (metis Diff_disjoint inf_bot_left interior_of_Int interior_of_complement interior_of_empty) moreover have "S \ X interior_of T \ (topspace X - X closure_of T)" using that \S \ topspace X\ null by auto moreover have "S \ X interior_of T \ {}" using closure_of_subset that(1) that(3) null by fastforce ultimately have "S \ X interior_of (topspace X - T) = {}" by (metis "*" inf_commute interior_of_complement openin_interior_of) then have "topspace (subtopology X S) \ X interior_of T = S" using \S \ topspace X\ interior_of_complement null by fastforce then show ?thesis using that by (metis Diff_eq_empty_iff inf_le2 interior_of_subset subset_trans) qed ultimately show ?thesis by (metis Int_lower1 frontier_of_restrict inf_assoc) qed lemma connectedin_continuous_map_image: assumes f: "continuous_map X Y f" and "connectedin X S" shows "connectedin Y (f ` S)" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) show ?thesis unfolding connectedin connected_space_def proof (intro conjI notI; clarify) show "f x \ topspace Y" if "x \ S" for x using \S \ topspace X\ continuous_map_image_subset_topspace f that by blast next fix U V let ?U = "{x \ topspace X. f x \ U}" let ?V = "{x \ topspace X. f x \ V}" assume UV: "openin Y U" "openin Y V" "f ` S \ U \ V" "U \ V \ f ` S = {}" "U \ f ` S \ {}" "V \ f ` S \ {}" then have 1: "?U \ ?V \ S = {}" by auto have 2: "openin X ?U" "openin X ?V" using \openin Y U\ \openin Y V\ continuous_map f by fastforce+ show "False" using * [of ?U ?V] UV \S \ topspace X\ by (auto simp: 1 2) qed qed +lemma connected_space_quotient_map_image: + "\quotient_map X X' q; connected_space X\ \ connected_space X'" + by (metis connectedin_continuous_map_image connectedin_topspace quotient_imp_continuous_map quotient_imp_surjective_map) + lemma homeomorphic_connected_space: "X homeomorphic_space Y \ connected_space X \ connected_space Y" unfolding homeomorphic_space_def homeomorphic_maps_def by (metis connected_space_subconnected connectedin_continuous_map_image connectedin_topspace continuous_map_image_subset_topspace image_eqI image_subset_iff) lemma homeomorphic_map_connectedness: assumes f: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "connectedin Y (f ` U) \ connectedin X U" proof - have 1: "f ` U \ topspace Y \ U \ topspace X" using U f homeomorphic_imp_surjective_map by blast moreover have "connected_space (subtopology Y (f ` U)) \ connected_space (subtopology X U)" proof (rule homeomorphic_connected_space) have "f ` U \ topspace Y" by (simp add: U 1) then have "topspace Y \ f ` U = f ` U" by (simp add: subset_antisym) then show "subtopology Y (f ` U) homeomorphic_space subtopology X U" by (metis U f homeomorphic_map_imp_homeomorphic_space homeomorphic_map_subtopologies homeomorphic_space_sym inf.absorb_iff2) qed ultimately show ?thesis by (auto simp: connectedin_def) qed lemma homeomorphic_map_connectedness_eq: "homeomorphic_map X Y f \ connectedin X U \ U \ topspace X \ connectedin Y (f ` U)" using homeomorphic_map_connectedness connectedin_subset_topspace by metis lemma connectedin_discrete_topology: "connectedin (discrete_topology U) S \ S \ U \ (\a. S \ {a})" proof (cases "S \ U") case True show ?thesis proof (cases "S = {}") case False moreover have "connectedin (discrete_topology U) S \ (\a. S = {a})" proof show "connectedin (discrete_topology U) S \ \a. S = {a}" using False connectedin_Int_frontier_of insert_Diff by fastforce qed (use True in auto) ultimately show ?thesis by auto qed simp next case False then show ?thesis by (simp add: connectedin_def) qed lemma connected_space_discrete_topology: "connected_space (discrete_topology U) \ (\a. U \ {a})" by (metis connectedin_discrete_topology connectedin_topspace order_refl topspace_discrete_topology) subsection\Compact sets\ definition compactin where "compactin X S \ S \ topspace X \ (\\. (\U \ \. openin X U) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\))" definition compact_space where "compact_space X \ compactin X (topspace X)" lemma compact_space_alt: "compact_space X \ (\\. (\U \ \. openin X U) \ topspace X \ \\ \ (\\. finite \ \ \ \ \ \ topspace X \ \\))" by (simp add: compact_space_def compactin_def) lemma compact_space: "compact_space X \ (\\. (\U \ \. openin X U) \ \\ = topspace X \ (\\. finite \ \ \ \ \ \ \\ = topspace X))" unfolding compact_space_alt using openin_subset by fastforce lemma compactinD: "\compactin X S; \U. U \ \ \ openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" by (auto simp: compactin_def) lemma compactin_euclidean_iff [simp]: "compactin euclidean S \ compact S" by (simp add: compact_eq_Heine_Borel compactin_def) meson lemma compactin_absolute [simp]: "compactin (subtopology X S) S \ compactin X S" proof - have eq: "(\U \ \. \Y. openin X Y \ U = Y \ S) \ \ \ (\Y. Y \ S) ` {y. openin X y}" for \ by auto show ?thesis by (auto simp: compactin_def openin_subtopology eq imp_conjL all_subset_image ex_finite_subset_image) qed lemma compactin_subspace: "compactin X S \ S \ topspace X \ compact_space (subtopology X S)" unfolding compact_space_def topspace_subtopology by (metis compactin_absolute compactin_def inf.absorb2) lemma compact_space_subtopology: "compactin X S \ compact_space (subtopology X S)" by (simp add: compactin_subspace) lemma compactin_subtopology: "compactin (subtopology X S) T \ compactin X T \ T \ S" by (metis compactin_subspace inf.absorb_iff2 le_inf_iff subtopology_subtopology topspace_subtopology) lemma compactin_subset_topspace: "compactin X S \ S \ topspace X" by (simp add: compactin_subspace) lemma compactin_contractive: "\compactin X' S; topspace X' = topspace X; \U. openin X U \ openin X' U\ \ compactin X S" by (simp add: compactin_def) lemma finite_imp_compactin: "\S \ topspace X; finite S\ \ compactin X S" by (metis compactin_subspace compact_space finite_UnionD inf.absorb_iff2 order_refl topspace_subtopology) lemma compactin_empty [iff]: "compactin X {}" by (simp add: finite_imp_compactin) lemma compact_space_topspace_empty: "topspace X = {} \ compact_space X" by (simp add: compact_space_def) lemma finite_imp_compactin_eq: "finite S \ (compactin X S \ S \ topspace X)" using compactin_subset_topspace finite_imp_compactin by blast lemma compactin_sing [simp]: "compactin X {a} \ a \ topspace X" by (simp add: finite_imp_compactin_eq) lemma closed_compactin: assumes XK: "compactin X K" and "C \ K" and XC: "closedin X C" shows "compactin X C" unfolding compactin_def proof (intro conjI allI impI) show "C \ topspace X" by (simp add: XC closedin_subset) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ C \ \\" have "(\U\insert (topspace X - C) \. openin X U)" using XC \ by blast moreover have "K \ \(insert (topspace X - C) \)" using \ XK compactin_subset_topspace by fastforce ultimately obtain \ where "finite \" "\ \ insert (topspace X - C) \" "K \ \\" using assms unfolding compactin_def by metis moreover have "openin X (topspace X - C)" using XC by auto ultimately show "\\. finite \ \ \ \ \ \ C \ \\" using \C \ K\ by (rule_tac x="\ - {topspace X - C}" in exI) auto qed lemma closedin_compact_space: "\compact_space X; closedin X S\ \ compactin X S" by (simp add: closed_compactin closedin_subset compact_space_def) lemma compact_Int_closedin: assumes "compactin X S" "closedin X T" shows "compactin X (S \ T)" proof - have "compactin (subtopology X S) (S \ T)" by (metis assms closedin_compact_space closedin_subtopology compactin_subspace inf_commute) then show ?thesis by (simp add: compactin_subtopology) qed lemma closed_Int_compactin: "\closedin X S; compactin X T\ \ compactin X (S \ T)" by (metis compact_Int_closedin inf_commute) lemma compactin_Un: assumes S: "compactin X S" and T: "compactin X T" shows "compactin X (S \ T)" unfolding compactin_def proof (intro conjI allI impI) show "S \ T \ topspace X" using assms by (auto simp: compactin_def) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ S \ T \ \\" with S obtain \ where \: "finite \" "\ \ \" "S \ \\" unfolding compactin_def by (meson sup.bounded_iff) obtain \ where "finite \" "\ \ \" "T \ \\" using \ T unfolding compactin_def by (meson sup.bounded_iff) with \ show "\\. finite \ \ \ \ \ \ S \ T \ \\" by (rule_tac x="\ \ \" in exI) auto qed lemma compactin_Union: "\finite \; \S. S \ \ \ compactin X S\ \ compactin X (\\)" by (induction rule: finite_induct) (simp_all add: compactin_Un) lemma compactin_subtopology_imp_compact: assumes "compactin (subtopology X S) K" shows "compactin X K" using assms proof (clarsimp simp add: compactin_def) fix \ define \ where "\ \ (\U. U \ S) ` \" assume "K \ topspace X" and "K \ S" and "\x\\. openin X x" and "K \ \\" then have "\V \ \. openin (subtopology X S) V" "K \ \\" unfolding \_def by (auto simp: openin_subtopology) moreover assume "\\. (\x\\. openin (subtopology X S) x) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson then have \: "\U. U \ \ \ V = U \ S" if "V \ \" for V unfolding \_def using that by blast let ?\ = "(\F. @U. U \ \ \ F = U \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using someI_ex [OF \] by blast show "K \ \?\" proof clarsimp fix x assume "x \ K" then show "\V \ \. x \ (SOME U. U \ \ \ V = U \ S)" using \K \ \\\ someI_ex [OF \] by (metis (no_types, lifting) IntD1 Union_iff subsetCE) qed qed qed lemma compact_imp_compactin_subtopology: assumes "compactin X K" "K \ S" shows "compactin (subtopology X S) K" using assms proof (clarsimp simp add: compactin_def) fix \ :: "'a set set" define \ where "\ \ {V. openin X V \ (\U \ \. U = V \ S)}" assume "K \ S" and "K \ topspace X" and "\U\\. openin (subtopology X S) U" and "K \ \\" then have "\V \ \. openin X V" "K \ \\" unfolding \_def by (fastforce simp: subset_eq openin_subtopology)+ moreover assume "\\. (\U\\. openin X U) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson let ?\ = "(\F. F \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using \_def \\ \ \\ by blast show "K \ \?\" using \K \ \\\ assms(2) by auto qed qed proposition compact_space_fip: "compact_space X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" (is "_ = ?rhs") proof (cases "topspace X = {}") case True then show ?thesis unfolding compact_space_def by (metis Sup_bot_conv(1) closedin_topspace_empty compactin_empty finite.emptyI finite_UnionD order_refl) next case False show ?thesis proof safe fix \ :: "'a set set" assume * [rule_format]: "\\. finite \ \ \ \ \ \ \\ \ {}" define \ where "\ \ (\S. topspace X - S) ` \" assume clo: "\C\\. closedin X C" and [simp]: "\\ = {}" then have "\V \ \. openin X V" "topspace X \ \\" by (auto simp: \_def) moreover assume [unfolded compact_space_alt, rule_format, of \]: "compact_space X" ultimately obtain \ where \: "finite \" "\ \ \" "topspace X \ topspace X - \\" by (auto simp: ex_finite_subset_image \_def) moreover have "\ \ {}" using \ \topspace X \ {}\ by blast ultimately show "False" using * [of \] by auto (metis Diff_iff Inter_iff clo closedin_def subsetD) next assume R [rule_format]: ?rhs show "compact_space X" unfolding compact_space_alt proof clarify fix \ :: "'a set set" define \ where "\ \ (\S. topspace X - S) ` \" assume "\C\\. openin X C" and "topspace X \ \\" with \topspace X \ {}\ have *: "\V \ \. closedin X V" "\ \ {}" by (auto simp: \_def) show "\\. finite \ \ \ \ \ \ topspace X \ \\" proof (rule ccontr; simp) assume "\\\\. finite \ \ \ topspace X \ \\" then have "\\. finite \ \ \ \ \ \ \\ \ {}" by (simp add: \_def all_finite_subset_image) with \topspace X \ \\\ show False using R [of \] * by (simp add: \_def) qed qed qed qed corollary compactin_fip: "compactin X S \ S \ topspace X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof (cases "S = {}") case False show ?thesis proof (cases "S \ topspace X") case True then have "compactin X S \ (\\. \ \ (\T. S \ T) ` {T. closedin X T} \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" by (simp add: compact_space_fip compactin_subspace closedin_subtopology image_def subset_eq Int_commute imp_conjL) also have "\ = (\\\Collect (closedin X). (\\. finite \ \ \ \ (\) S ` \ \ \\ \ {}) \ \ ((\) S ` \) \ {})" by (simp add: all_subset_image) also have "\ = (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof - have eq: "((\\. finite \ \ \ \ \ \ \ ((\) S ` \) \ {}) \ \ ((\) S ` \) \ {}) \ ((\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" for \ by simp (use \S \ {}\ in blast) show ?thesis unfolding imp_conjL [symmetric] all_finite_subset_image eq by blast qed finally show ?thesis using True by simp qed (simp add: compactin_subspace) qed force corollary compact_space_imp_nest: fixes C :: "nat \ 'a set" assumes "compact_space X" and clo: "\n. closedin X (C n)" and ne: "\n. C n \ {}" and dec: "decseq C" shows "(\n. C n) \ {}" proof - let ?\ = "range (\n. \m \ n. C m)" have "closedin X A" if "A \ ?\" for A using that clo by auto moreover have "(\n\K. \m \ n. C m) \ {}" if "finite K" for K proof - obtain n where "\k. k \ K \ k \ n" using Max.coboundedI \finite K\ by blast with dec have "C n \ (\n\K. \m \ n. C m)" unfolding decseq_def by blast with ne [of n] show ?thesis by blast qed ultimately show ?thesis using \compact_space X\ [unfolded compact_space_fip, rule_format, of ?\] by (simp add: all_finite_subset_image INT_extend_simps UN_atMost_UNIV del: INT_simps) qed lemma compactin_discrete_topology: "compactin (discrete_topology X) S \ S \ X \ finite S" (is "?lhs = ?rhs") proof (intro iffI conjI) assume L: ?lhs then show "S \ X" by (auto simp: compactin_def) have *: "\\. Ball \ (openin (discrete_topology X)) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\)" using L by (auto simp: compactin_def) show "finite S" using * [of "(\x. {x}) ` X"] \S \ X\ by clarsimp (metis UN_singleton finite_subset_image infinite_super) next assume ?rhs then show ?lhs by (simp add: finite_imp_compactin) qed lemma compact_space_discrete_topology: "compact_space(discrete_topology X) \ finite X" by (simp add: compactin_discrete_topology compact_space_def) lemma compact_space_imp_Bolzano_Weierstrass: assumes "compact_space X" "infinite S" "S \ topspace X" shows "X derived_set_of S \ {}" proof assume X: "X derived_set_of S = {}" then have "closedin X S" by (simp add: closedin_contains_derived_set assms) then have "compactin X S" by (rule closedin_compact_space [OF \compact_space X\]) with X show False by (metis \infinite S\ compactin_subspace compact_space_discrete_topology inf_bot_right subtopology_eq_discrete_topology_eq) qed lemma compactin_imp_Bolzano_Weierstrass: "\compactin X S; infinite T \ T \ S\ \ S \ X derived_set_of T \ {}" using compact_space_imp_Bolzano_Weierstrass [of "subtopology X S"] by (simp add: compactin_subspace derived_set_of_subtopology inf_absorb2) lemma compact_closure_of_imp_Bolzano_Weierstrass: "\compactin X (X closure_of S); infinite T; T \ S; T \ topspace X\ \ X derived_set_of T \ {}" using closure_of_mono closure_of_subset compactin_imp_Bolzano_Weierstrass by fastforce lemma discrete_compactin_eq_finite: "S \ X derived_set_of S = {} \ compactin X S \ S \ topspace X \ finite S" by (meson compactin_imp_Bolzano_Weierstrass finite_imp_compactin_eq order_refl) lemma discrete_compact_space_eq_finite: "X derived_set_of (topspace X) = {} \ (compact_space X \ finite(topspace X))" by (metis compact_space_discrete_topology discrete_topology_unique_derived_set) lemma image_compactin: assumes cpt: "compactin X S" and cont: "continuous_map X Y f" shows "compactin Y (f ` S)" unfolding compactin_def proof (intro conjI allI impI) show "f ` S \ topspace Y" using compactin_subset_topspace cont continuous_map_image_subset_topspace cpt by blast next fix \ :: "'b set set" assume \: "Ball \ (openin Y) \ f ` S \ \\" define \ where "\ \ (\U. {x \ topspace X. f x \ U}) ` \" have "S \ topspace X" and *: "\\. \\U\\. openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" using cpt by (auto simp: compactin_def) obtain \ where \: "finite \" "\ \ \" "S \ \\" proof - have 1: "\U\\. openin X U" unfolding \_def using \ cont[unfolded continuous_map] by blast have 2: "S \ \\" unfolding \_def using compactin_subset_topspace cpt \ by fastforce show thesis using * [OF 1 2] that by metis qed have "\v \ \. \U. U \ \ \ v = {x \ topspace X. f x \ U}" using \_def by blast then obtain U where U: "\v \ \. U v \ \ \ v = {x \ topspace X. f x \ U v}" by metis show "\\. finite \ \ \ \ \ \ f ` S \ \\" proof (intro conjI exI) show "finite (U ` \)" by (simp add: \finite \\) next show "U ` \ \ \" using \\ \ \\ U by auto next show "f ` S \ \ (U ` \)" using \(2-3) U UnionE subset_eq U by fastforce qed qed lemma homeomorphic_compact_space: assumes "X homeomorphic_space Y" shows "compact_space X \ compact_space Y" using homeomorphic_space_sym by (metis assms compact_space_def homeomorphic_eq_everything_map homeomorphic_space image_compactin) lemma homeomorphic_map_compactness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "compactin Y (f ` U) \ compactin X U" proof - have "f ` U \ topspace Y" using hom U homeomorphic_imp_surjective_map by blast moreover have "homeomorphic_map (subtopology X U) (subtopology Y (f ` U)) f" using U hom homeomorphic_imp_surjective_map by (blast intro: homeomorphic_map_subtopologies) then have "compact_space (subtopology Y (f ` U)) = compact_space (subtopology X U)" using homeomorphic_compact_space homeomorphic_map_imp_homeomorphic_space by blast ultimately show ?thesis by (simp add: compactin_subspace U) qed lemma homeomorphic_map_compactness_eq: "homeomorphic_map X Y f \ compactin X U \ U \ topspace X \ compactin Y (f ` U)" by (meson compactin_subset_topspace homeomorphic_map_compactness) subsection\Embedding maps\ definition embedding_map where "embedding_map X Y f \ homeomorphic_map X (subtopology Y (f ` (topspace X))) f" lemma embedding_map_eq: "\embedding_map X Y f; \x. x \ topspace X \ f x = g x\ \ embedding_map X Y g" unfolding embedding_map_def by (metis homeomorphic_map_eq image_cong) lemma embedding_map_compose: assumes "embedding_map X X' f" "embedding_map X' X'' g" shows "embedding_map X X'' (g \ f)" proof - have hm: "homeomorphic_map X (subtopology X' (f ` topspace X)) f" "homeomorphic_map X' (subtopology X'' (g ` topspace X')) g" using assms by (auto simp: embedding_map_def) then obtain C where "g ` topspace X' \ C = (g \ f) ` topspace X" by (metis homeomorphic_imp_surjective_map image_comp image_mono inf.absorb_iff2 topspace_subtopology) then have "homeomorphic_map (subtopology X' (f ` topspace X)) (subtopology X'' ((g \ f) ` topspace X)) g" by (metis hm homeomorphic_imp_surjective_map homeomorphic_map_subtopologies image_comp subtopology_subtopology topspace_subtopology) then show ?thesis unfolding embedding_map_def using hm(1) homeomorphic_map_compose by blast qed lemma surjective_embedding_map: "embedding_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (force simp: embedding_map_def homeomorphic_eq_everything_map) lemma embedding_map_in_subtopology: "embedding_map X (subtopology Y S) f \ embedding_map X Y f \ f ` (topspace X) \ S" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding embedding_map_def by (metis continuous_map_in_subtopology homeomorphic_imp_continuous_map inf_absorb2 subtopology_subtopology) qed (simp add: embedding_map_def inf.absorb_iff2 subtopology_subtopology) lemma injective_open_imp_embedding_map: "\continuous_map X Y f; open_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def by (simp add: continuous_map_in_subtopology continuous_open_quotient_map eq_iff homeomorphic_map_def open_map_imp_subset open_map_into_subtopology) lemma injective_closed_imp_embedding_map: "\continuous_map X Y f; closed_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def by (simp add: closed_map_imp_subset closed_map_into_subtopology continuous_closed_quotient_map continuous_map_in_subtopology dual_order.eq_iff homeomorphic_map_def) lemma embedding_map_imp_homeomorphic_space: "embedding_map X Y f \ X homeomorphic_space (subtopology Y (f ` (topspace X)))" unfolding embedding_map_def using homeomorphic_space by blast lemma embedding_imp_closed_map: "\embedding_map X Y f; closedin Y (f ` topspace X)\ \ closed_map X Y f" unfolding closed_map_def by (auto simp: closedin_closed_subtopology embedding_map_def homeomorphic_map_closedness_eq) lemma embedding_imp_closed_map_eq: "embedding_map X Y f \ (closed_map X Y f \ closedin Y (f ` topspace X))" using closed_map_def embedding_imp_closed_map by blast subsection\Retraction and section maps\ definition retraction_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "retraction_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace Y. f(g x) = x)" definition section_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "section_map X Y f \ \g. retraction_maps Y X g f" definition retraction_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "retraction_map X Y f \ \g. retraction_maps X Y f g" lemma retraction_maps_eq: "\retraction_maps X Y f g; \x. x \ topspace X \ f x = f' x; \x. x \ topspace Y \ g x = g' x\ \ retraction_maps X Y f' g'" unfolding retraction_maps_def by (metis (no_types, lifting) continuous_map_def continuous_map_eq) lemma section_map_eq: "\section_map X Y f; \x. x \ topspace X \ f x = g x\ \ section_map X Y g" unfolding section_map_def using retraction_maps_eq by blast lemma retraction_map_eq: "\retraction_map X Y f; \x. x \ topspace X \ f x = g x\ \ retraction_map X Y g" unfolding retraction_map_def using retraction_maps_eq by blast lemma homeomorphic_imp_retraction_maps: "homeomorphic_maps X Y f g \ retraction_maps X Y f g" by (simp add: homeomorphic_maps_def retraction_maps_def) lemma section_and_retraction_eq_homeomorphic_map: "section_map X Y f \ retraction_map X Y f \ homeomorphic_map X Y f" (is "?lhs = ?rhs") proof assume ?lhs then obtain g where "homeomorphic_maps X Y f g" unfolding homeomorphic_maps_def retraction_map_def section_map_def by (smt (verit, best) continuous_map_def retraction_maps_def) then show ?rhs using homeomorphic_map_maps by blast next assume ?rhs then show ?lhs unfolding retraction_map_def section_map_def by (meson homeomorphic_imp_retraction_maps homeomorphic_map_maps homeomorphic_maps_sym) qed lemma section_imp_embedding_map: "section_map X Y f \ embedding_map X Y f" unfolding section_map_def embedding_map_def homeomorphic_map_maps retraction_maps_def homeomorphic_maps_def by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology) lemma retraction_imp_quotient_map: assumes "retraction_map X Y f" shows "quotient_map X Y f" unfolding quotient_map_def proof (intro conjI subsetI allI impI) show "f ` topspace X = topspace Y" using assms by (force simp: retraction_map_def retraction_maps_def continuous_map_def) next fix U assume U: "U \ topspace Y" have "openin Y U" if "\x\topspace Y. g x \ topspace X" "\x\topspace Y. f (g x) = x" "openin Y {x \ topspace Y. g x \ {x \ topspace X. f x \ U}}" for g using openin_subopen U that by fastforce then show "openin X {x \ topspace X. f x \ U} = openin Y U" using assms by (auto simp: retraction_map_def retraction_maps_def continuous_map_def) qed lemma retraction_maps_compose: "\retraction_maps X Y f f'; retraction_maps Y Z g g'\ \ retraction_maps X Z (g \ f) (f' \ g')" by (clarsimp simp: retraction_maps_def continuous_map_compose) (simp add: continuous_map_def) lemma retraction_map_compose: "\retraction_map X Y f; retraction_map Y Z g\ \ retraction_map X Z (g \ f)" by (meson retraction_map_def retraction_maps_compose) lemma section_map_compose: "\section_map X Y f; section_map Y Z g\ \ section_map X Z (g \ f)" by (meson retraction_maps_compose section_map_def) lemma surjective_section_eq_homeomorphic_map: "section_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (meson section_and_retraction_eq_homeomorphic_map section_imp_embedding_map surjective_embedding_map) lemma surjective_retraction_or_section_map: "f ` (topspace X) = topspace Y \ retraction_map X Y f \ section_map X Y f \ retraction_map X Y f" using section_and_retraction_eq_homeomorphic_map surjective_section_eq_homeomorphic_map by fastforce lemma retraction_imp_surjective_map: "retraction_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: retraction_imp_quotient_map quotient_imp_surjective_map) lemma section_imp_injective_map: "\section_map X Y f; x \ topspace X; y \ topspace X\ \ f x = f y \ x = y" by (metis (mono_tags, opaque_lifting) retraction_maps_def section_map_def) lemma retraction_maps_to_retract_maps: "retraction_maps X Y r s \ retraction_maps X (subtopology X (s ` (topspace Y))) (s \ r) id" unfolding retraction_maps_def by (auto simp: continuous_map_compose continuous_map_into_subtopology continuous_map_from_subtopology) subsection \Continuity\ lemma continuous_on_open: "continuous_on S f \ (\T. openin (top_of_set (f ` S)) T \ openin (top_of_set S) (S \ f -` T))" unfolding continuous_on_open_invariant openin_open Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_closed: "continuous_on S f \ (\T. closedin (top_of_set (f ` S)) T \ closedin (top_of_set S) (S \ f -` T))" unfolding continuous_on_closed_invariant closedin_closed Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_imp_closedin: assumes "continuous_on S f" "closedin (top_of_set (f ` S)) T" shows "closedin (top_of_set S) (S \ f -` T)" using assms continuous_on_closed by blast lemma continuous_map_subtopology_eu [simp]: "continuous_map (top_of_set S) (subtopology euclidean T) h \ continuous_on S h \ h ` S \ T" by (simp add: continuous_map_in_subtopology) lemma continuous_map_euclidean_top_of_set: assumes eq: "f -` S = UNIV" and cont: "continuous_on UNIV f" shows "continuous_map euclidean (top_of_set S) f" by (simp add: cont continuous_map_into_subtopology eq image_subset_iff_subset_vimage) subsection\<^marker>\tag unimportant\ \Half-global and completely global cases\ lemma continuous_openin_preimage_gen: assumes "continuous_on S f" "open T" shows "openin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "openin (top_of_set (f ` S)) (T \ f ` S)" using openin_open_Int[of T "f ` S", OF assms(2)] unfolding openin_open by auto then show ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_closedin_preimage: assumes "continuous_on S f" and "closed T" shows "closedin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "closedin (top_of_set (f ` S)) (T \ f ` S)" using closedin_closed_Int[of T "f ` S", OF assms(2)] by (simp add: Int_commute) then show ?thesis using assms(1)[unfolded continuous_on_closed, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_openin_preimage_eq: "continuous_on S f \ (\T. open T \ openin (top_of_set S) (S \ f -` T))" by (metis Int_commute continuous_on_open_invariant open_openin openin_subtopology) lemma continuous_closedin_preimage_eq: "continuous_on S f \ (\T. closed T \ closedin (top_of_set S) (S \ f -` T))" by (metis Int_commute closedin_closed continuous_on_closed_invariant) lemma continuous_open_preimage: assumes contf: "continuous_on S f" and "open S" "open T" shows "open (S \ f -` T)" proof- obtain U where "open U" "(S \ f -` T) = S \ U" using continuous_openin_preimage_gen[OF contf \open T\] unfolding openin_open by auto then show ?thesis using open_Int[of S U, OF \open S\] by auto qed lemma continuous_closed_preimage: assumes contf: "continuous_on S f" and "closed S" "closed T" shows "closed (S \ f -` T)" proof- obtain U where "closed U" "(S \ f -` T) = S \ U" using continuous_closedin_preimage[OF contf \closed T\] unfolding closedin_closed by auto then show ?thesis using closed_Int[of S U, OF \closed S\] by auto qed lemma continuous_open_vimage: "open S \ (\x. continuous (at x) f) \ open (f -` S)" by (metis continuous_on_eq_continuous_within open_vimage) lemma continuous_closed_vimage: "closed S \ (\x. continuous (at x) f) \ closed (f -` S)" by (simp add: closed_vimage continuous_on_eq_continuous_within) lemma Times_in_interior_subtopology: assumes "(x, y) \ U" "openin (top_of_set (S \ T)) U" obtains V W where "openin (top_of_set S) V" "x \ V" "openin (top_of_set T) W" "y \ W" "(V \ W) \ U" proof - from assms obtain E where "open E" "U = S \ T \ E" "(x, y) \ E" "x \ S" "y \ T" by (auto simp: openin_open) from open_prod_elim[OF \open E\ \(x, y) \ E\] obtain E1 E2 where "open E1" "open E2" "(x, y) \ E1 \ E2" "E1 \ E2 \ E" by blast show ?thesis proof show "openin (top_of_set S) (E1 \ S)" "openin (top_of_set T) (E2 \ T)" using \open E1\ \open E2\ by (auto simp: openin_open) show "x \ E1 \ S" "y \ E2 \ T" using \(x, y) \ E1 \ E2\ \x \ S\ \y \ T\ by auto show "(E1 \ S) \ (E2 \ T) \ U" using \E1 \ E2 \ E\ \U = _\ by auto qed qed lemma closedin_Times: "closedin (top_of_set S) S' \ closedin (top_of_set T) T' \ closedin (top_of_set (S \ T)) (S' \ T')" unfolding closedin_closed using closed_Times by blast lemma openin_Times: "openin (top_of_set S) S' \ openin (top_of_set T) T' \ openin (top_of_set (S \ T)) (S' \ T')" unfolding openin_open using open_Times by blast lemma openin_Times_eq: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" shows "openin (top_of_set (S \ T)) (S' \ T') \ S' = {} \ T' = {} \ openin (top_of_set S) S' \ openin (top_of_set T) T'" (is "?lhs = ?rhs") proof (cases "S' = {} \ T' = {}") case True then show ?thesis by auto next case False then obtain x y where "x \ S'" "y \ T'" by blast show ?thesis proof assume ?lhs have "openin (top_of_set S) S'" proof (subst openin_subopen, clarify) show "\U. openin (top_of_set S) U \ x \ U \ U \ S'" if "x \ S'" for x using that \y \ T'\ Times_in_interior_subtopology [OF _ \?lhs\, of x y] by simp (metis mem_Sigma_iff subsetD subsetI) qed moreover have "openin (top_of_set T) T'" proof (subst openin_subopen, clarify) show "\U. openin (top_of_set T) U \ y \ U \ U \ T'" if "y \ T'" for y using that \x \ S'\ Times_in_interior_subtopology [OF _ \?lhs\, of x y] by simp (metis mem_Sigma_iff subsetD subsetI) qed ultimately show ?rhs by simp next assume ?rhs with False show ?lhs by (simp add: openin_Times) qed qed lemma Lim_transform_within_openin: assumes f: "(f \ l) (at a within T)" and "openin (top_of_set T) S" "a \ S" and eq: "\x. \x \ S; x \ a\ \ f x = g x" shows "(g \ l) (at a within T)" proof - have "\\<^sub>F x in at a within T. x \ T \ x \ a" by (simp add: eventually_at_filter) moreover from \openin _ _\ obtain U where "open U" "S = T \ U" by (auto simp: openin_open) then have "a \ U" using \a \ S\ by auto from topological_tendstoD[OF tendsto_ident_at \open U\ \a \ U\] have "\\<^sub>F x in at a within T. x \ U" by auto ultimately have "\\<^sub>F x in at a within T. f x = g x" by eventually_elim (auto simp: \S = _\ eq) with f show ?thesis by (rule Lim_transform_eventually) qed lemma continuous_on_open_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. openin (top_of_set T) U \ openin (top_of_set S) (S \ f -` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (clarsimp simp add: continuous_openin_preimage_eq openin_open) (metis Int_assoc assms image_subset_iff_subset_vimage inf.absorb_iff1) next assume R [rule_format]: ?rhs show ?lhs proof (clarsimp simp add: continuous_openin_preimage_eq) fix U::"'a set" assume "open U" then have "openin (top_of_set S) (S \ f -` (U \ T))" by (metis R inf_commute openin_open) then show "openin (top_of_set S) (S \ f -` U)" by (metis Int_assoc Int_commute assms image_subset_iff_subset_vimage inf.absorb_iff2 vimage_Int) qed qed lemma continuous_openin_preimage: "\continuous_on S f; f ` S \ T; openin (top_of_set T) U\ \ openin (top_of_set S) (S \ f -` U)" by (simp add: continuous_on_open_gen) lemma continuous_on_closed_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. closedin (top_of_set T) U \ closedin (top_of_set S) (S \ f -` U))" proof - have *: "U \ T \ S \ f -` (T - U) = S - (S \ f -` U)" for U using assms by blast then show ?thesis unfolding continuous_on_open_gen [OF assms] by (metis closedin_def inf.cobounded1 openin_closedin_eq topspace_euclidean_subtopology) qed lemma continuous_closedin_preimage_gen: assumes "continuous_on S f" "f ` S \ T" "closedin (top_of_set T) U" shows "closedin (top_of_set S) (S \ f -` U)" using assms continuous_on_closed_gen by blast lemma continuous_transform_within_openin: assumes "continuous (at a within T) f" and "openin (top_of_set T) S" "a \ S" and eq: "\x. x \ S \ f x = g x" shows "continuous (at a within T) g" using assms by (simp add: Lim_transform_within_openin continuous_within) subsection\<^marker>\tag important\ \The topology generated by some (open) subsets\ text \In the definition below of a generated topology, the \Empty\ case is not necessary, as it follows from \UN\ taking for \K\ the empty set. However, it is convenient to have, and is never a problem in proofs, so I prefer to write it down explicitly. We do not require \UNIV\ to be an open set, as this will not be the case in applications. (We are thinking of a topology on a subset of \UNIV\, the remaining part of \UNIV\ being irrelevant.)\ inductive generate_topology_on for S where Empty: "generate_topology_on S {}" | Int: "generate_topology_on S a \ generate_topology_on S b \ generate_topology_on S (a \ b)" | UN: "(\k. k \ K \ generate_topology_on S k) \ generate_topology_on S (\K)" | Basis: "s \ S \ generate_topology_on S s" lemma istopology_generate_topology_on: "istopology (generate_topology_on S)" unfolding istopology_def by (auto intro: generate_topology_on.intros) text \The basic property of the topology generated by a set \S\ is that it is the smallest topology containing all the elements of \S\:\ lemma generate_topology_on_coarsest: assumes T: "istopology T" "\s. s \ S \ T s" and gen: "generate_topology_on S s0" shows "T s0" using gen by (induct rule: generate_topology_on.induct) (use T in \auto simp: istopology_def\) abbreviation\<^marker>\tag unimportant\ topology_generated_by::"('a set set) \ ('a topology)" where "topology_generated_by S \ topology (generate_topology_on S)" lemma openin_topology_generated_by_iff: "openin (topology_generated_by S) s \ generate_topology_on S s" using topology_inverse'[OF istopology_generate_topology_on[of S]] by simp lemma openin_topology_generated_by: "openin (topology_generated_by S) s \ generate_topology_on S s" using openin_topology_generated_by_iff by auto lemma topology_generated_by_topspace [simp]: "topspace (topology_generated_by S) = (\S)" proof { fix s assume "openin (topology_generated_by S) s" then have "generate_topology_on S s" by (rule openin_topology_generated_by) then have "s \ (\S)" by (induct, auto) } then show "topspace (topology_generated_by S) \ (\S)" unfolding topspace_def by auto next have "generate_topology_on S (\S)" using generate_topology_on.UN[OF generate_topology_on.Basis, of S S] by simp then show "(\S) \ topspace (topology_generated_by S)" unfolding topspace_def using openin_topology_generated_by_iff by auto qed lemma topology_generated_by_Basis: "s \ S \ openin (topology_generated_by S) s" by (simp add: Basis openin_topology_generated_by_iff) lemma generate_topology_on_Inter: "\finite \; \K. K \ \ \ generate_topology_on \ K; \ \ {}\ \ generate_topology_on \ (\\)" by (induction \ rule: finite_induct; force intro: generate_topology_on.intros) subsection\Topology bases and sub-bases\ lemma istopology_base_alt: "istopology (arbitrary union_of P) \ (\S T. (arbitrary union_of P) S \ (arbitrary union_of P) T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_def) (blast intro: arbitrary_union_of_Union) lemma istopology_base_eq: "istopology (arbitrary union_of P) \ (\S T. P S \ P T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_base_alt arbitrary_union_of_Int_eq) lemma istopology_base: "(\S T. \P S; P T\ \ P(S \ T)) \ istopology (arbitrary union_of P)" by (simp add: arbitrary_def istopology_base_eq union_of_inc) lemma openin_topology_base_unique: "openin X = arbitrary union_of P \ (\V. P V \ openin X V) \ (\U x. openin X U \ x \ U \ (\V. P V \ x \ V \ V \ U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: union_of_def arbitrary_def) next assume R: ?rhs then have *: "\\\Collect P. \\ = S" if "openin X S" for S using that by (rule_tac x="{V. P V \ V \ S}" in exI) fastforce from R show ?lhs by (fastforce simp add: union_of_def arbitrary_def intro: *) qed lemma topology_base_unique: assumes "\S. P S \ openin X S" "\U x. \openin X U; x \ U\ \ \B. P B \ x \ B \ B \ U" shows "topology (arbitrary union_of P) = X" proof - have "X = topology (openin X)" by (simp add: openin_inverse) also from assms have "openin X = arbitrary union_of P" by (subst openin_topology_base_unique) auto finally show ?thesis .. qed lemma topology_bases_eq_aux: "\(arbitrary union_of P) S; \U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U\ \ (arbitrary union_of Q) S" by (metis arbitrary_union_of_alt arbitrary_union_of_idempot) lemma topology_bases_eq: "\\U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U; \V x. \Q V; x \ V\ \ \U. P U \ x \ U \ U \ V\ \ topology (arbitrary union_of P) = topology (arbitrary union_of Q)" by (fastforce intro: arg_cong [where f=topology] elim: topology_bases_eq_aux) lemma istopology_subbase: "istopology (arbitrary union_of (finite intersection_of P relative_to S))" by (simp add: finite_intersection_of_Int istopology_base relative_to_Int) lemma openin_subbase: "openin (topology (arbitrary union_of (finite intersection_of B relative_to U))) S \ (arbitrary union_of (finite intersection_of B relative_to U)) S" by (simp add: istopology_subbase topology_inverse') lemma topspace_subbase [simp]: "topspace(topology (arbitrary union_of (finite intersection_of B relative_to U))) = U" (is "?lhs = _") proof show "?lhs \ U" by (metis arbitrary_union_of_relative_to openin_subbase openin_topspace relative_to_imp_subset) show "U \ ?lhs" by (metis arbitrary_union_of_inc finite_intersection_of_empty inf.orderE istopology_subbase openin_subset relative_to_inc subset_UNIV topology_inverse') qed lemma minimal_topology_subbase: assumes X: "\S. P S \ openin X S" and "openin X U" and S: "openin(topology(arbitrary union_of (finite intersection_of P relative_to U))) S" shows "openin X S" proof - have "(arbitrary union_of (finite intersection_of P relative_to U)) S" using S openin_subbase by blast with X \openin X U\ show ?thesis by (force simp add: union_of_def intersection_of_def relative_to_def intro: openin_Int_Inter) qed lemma istopology_subbase_UNIV: "istopology (arbitrary union_of (finite intersection_of P))" by (simp add: istopology_base finite_intersection_of_Int) lemma generate_topology_on_eq: "generate_topology_on S = arbitrary union_of finite' intersection_of (\x. x \ S)" (is "?lhs = ?rhs") proof (intro ext iffI) fix A assume "?lhs A" then show "?rhs A" proof induction case (Int a b) then show ?case by (metis (mono_tags, lifting) istopology_base_alt finite'_intersection_of_Int istopology_base) next case (UN K) then show ?case by (simp add: arbitrary_union_of_Union) next case (Basis s) then show ?case by (simp add: Sup_upper arbitrary_union_of_inc finite'_intersection_of_inc relative_to_subset) qed auto next fix A assume "?rhs A" then obtain \ where \: "\T. T \ \ \ \\. finite' \ \ \ \ S \ \\ = T" and eq: "A = \\" unfolding union_of_def intersection_of_def by auto show "?lhs A" unfolding eq proof (rule generate_topology_on.UN) fix T assume "T \ \" with \ obtain \ where "finite' \" "\ \ S" "\\ = T" by blast have "generate_topology_on S (\\)" proof (rule generate_topology_on_Inter) show "finite \" "\ \ {}" by (auto simp: \finite' \\) show "\K. K \ \ \ generate_topology_on S K" by (metis \\ \ S\ generate_topology_on.simps subset_iff) qed then show "generate_topology_on S T" using \\\ = T\ by blast qed qed lemma continuous_on_generated_topo_iff: "continuous_map T1 (topology_generated_by S) f \ ((\U. U \ S \ openin T1 (f-`U \ topspace(T1))) \ (f`(topspace T1) \ (\ S)))" unfolding continuous_map_alt topology_generated_by_topspace proof (auto simp add: topology_generated_by_Basis) assume H: "\U. U \ S \ openin T1 (f -` U \ topspace T1)" fix U assume "openin (topology_generated_by S) U" then have "generate_topology_on S U" by (rule openin_topology_generated_by) then show "openin T1 (f -` U \ topspace T1)" proof (induct) fix a b assume H: "openin T1 (f -` a \ topspace T1)" "openin T1 (f -` b \ topspace T1)" have "f -` (a \ b) \ topspace T1 = (f-`a \ topspace T1) \ (f-`b \ topspace T1)" by auto then show "openin T1 (f -` (a \ b) \ topspace T1)" using H by auto next fix K assume H: "openin T1 (f -` k \ topspace T1)" if "k\ K" for k define L where "L = {f -` k \ topspace T1|k. k \ K}" have *: "openin T1 l" if "l \L" for l using that H unfolding L_def by auto have "openin T1 (\L)" using openin_Union[OF *] by simp moreover have "(\L) = (f -` \K \ topspace T1)" unfolding L_def by auto ultimately show "openin T1 (f -` \K \ topspace T1)" by simp qed (auto simp add: H) qed lemma continuous_on_generated_topo: assumes "\U. U \S \ openin T1 (f-`U \ topspace(T1))" "f`(topspace T1) \ (\ S)" shows "continuous_map T1 (topology_generated_by S) f" using assms continuous_on_generated_topo_iff by blast subsection\Continuity via bases/subbases, hence upper and lower semicontinuity\ lemma continuous_map_into_topology_base: assumes P: "openin Y = arbitrary union_of P" and f: "\x. x \ topspace X \ f x \ topspace Y" and ope: "\U. P U \ openin X {x \ topspace X. f x \ U}" shows "continuous_map X Y f" proof - have *: "\\. (\t. t \ \ \ P t) \ openin X {x \ topspace X. \U\\. f x \ U}" by (smt (verit) Ball_Collect ope mem_Collect_eq openin_subopen) show ?thesis using P by (auto simp: continuous_map_def arbitrary_def union_of_def intro!: f *) qed lemma continuous_map_into_topology_base_eq: assumes P: "openin Y = arbitrary union_of P" shows "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. P U \ openin X {x \ topspace X. f x \ U})" (is "?lhs=?rhs") proof assume L: ?lhs then have "\x. x \ topspace X \ f x \ topspace Y" by (meson continuous_map_def) moreover have "\U. P U \ openin X {x \ topspace X. f x \ U}" using L assms continuous_map openin_topology_base_unique by fastforce ultimately show ?rhs by auto qed (simp add: assms continuous_map_into_topology_base) lemma continuous_map_into_topology_subbase: fixes U P defines "Y \ topology(arbitrary union_of (finite intersection_of P relative_to U))" assumes f: "\x. x \ topspace X \ f x \ topspace Y" and ope: "\U. P U \ openin X {x \ topspace X. f x \ U}" shows "continuous_map X Y f" proof (intro continuous_map_into_topology_base) show "openin Y = arbitrary union_of (finite intersection_of P relative_to U)" unfolding Y_def using istopology_subbase topology_inverse' by blast show "openin X {x \ topspace X. f x \ V}" if \
: "(finite intersection_of P relative_to U) V" for V proof - define finv where "finv \ \V. {x \ topspace X. f x \ V}" obtain \ where \: "finite \" "\V. V \ \ \ P V" "{x \ topspace X. f x \ V} = (\V \ insert U \. finv V)" using \
by (fastforce simp: finv_def intersection_of_def relative_to_def) show ?thesis unfolding \ proof (intro openin_Inter ope) have U: "U = topspace Y" unfolding Y_def using topspace_subbase by fastforce fix V assume V: "V \ finv ` insert U \" with U f have "openin X {x \ topspace X. f x \ U}" by (auto simp: openin_subopen [of X "Collect _"]) then show "openin X V" using V \(2) ope by (fastforce simp: finv_def) qed (use \finite \\ in auto) qed qed (use f in auto) lemma continuous_map_into_topology_subbase_eq: assumes "Y = topology(arbitrary union_of (finite intersection_of P relative_to U))" shows "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. P U \ openin X {x \ topspace X. f x \ U})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof (intro conjI strip) show "\x. x \ topspace X \ f x \ topspace Y" using L continuous_map_def by fastforce fix V assume "P V" have "U = topspace Y" unfolding assms using topspace_subbase by fastforce then have eq: "{x \ topspace X. f x \ V} = {x \ topspace X. f x \ U \ V}" using L by (auto simp: continuous_map) have "openin Y (U \ V)" unfolding assms openin_subbase by (meson \P V\ arbitrary_union_of_inc finite_intersection_of_inc relative_to_inc) show "openin X {x \ topspace X. f x \ V}" using L \openin Y (U \ V)\ continuous_map eq by fastforce qed next show "?rhs \ ?lhs" unfolding assms by (intro continuous_map_into_topology_subbase) auto qed lemma subbase_subtopology_euclidean: fixes U :: "'a::order_topology set" shows "topology (arbitrary union_of (finite intersection_of (\x. x \ range greaterThan \ range lessThan) relative_to U)) = subtopology euclidean U" proof - have "\V. (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)) V \ x \ V \ V \ W" if "open W" "x \ W" for W and x::'a using \open W\ [unfolded open_generated_order] \x \ W\ proof (induct rule: generate_topology.induct) case UNIV then show ?case using finite_intersection_of_empty by blast next case (Int a b) then show ?case by (meson Int_iff finite_intersection_of_Int inf_mono) next case (UN K) then show ?case by (meson Union_iff subset_iff) next case (Basis s) then show ?case by (metis (no_types, lifting) Un_iff finite_intersection_of_inc order_refl) qed moreover have "\V::'a set. (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)) V \ open V" by (force simp: intersection_of_def subset_iff) ultimately have *: "openin (euclidean::'a topology) = (arbitrary union_of (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)))" by (smt (verit, best) openin_topology_base_unique open_openin) show ?thesis unfolding subtopology_def arbitrary_union_of_relative_to [symmetric] apply (simp add: relative_to_def flip: *) by (metis Int_commute) qed lemma continuous_map_upper_lower_semicontinuous_lt_gen: fixes U :: "'a::order_topology set" shows "continuous_map X (subtopology euclidean U) f \ (\x \ topspace X. f x \ U) \ (\a. openin X {x \ topspace X. f x > a}) \ (\a. openin X {x \ topspace X. f x < a})" by (auto simp: continuous_map_into_topology_subbase_eq [OF subbase_subtopology_euclidean [symmetric, of U]] greaterThan_def lessThan_def image_iff simp flip: all_simps) lemma continuous_map_upper_lower_semicontinuous_lt: fixes f :: "'a \ 'b::order_topology" shows "continuous_map X euclidean f \ (\a. openin X {x \ topspace X. f x > a}) \ (\a. openin X {x \ topspace X. f x < a})" using continuous_map_upper_lower_semicontinuous_lt_gen [where U=UNIV] by auto lemma Int_Collect_imp_eq: "A \ {x. x\A \ P x} = {x\A. P x}" by blast lemma continuous_map_upper_lower_semicontinuous_le_gen: shows "continuous_map X (subtopology euclideanreal U) f \ (\x \ topspace X. f x \ U) \ (\a. closedin X {x \ topspace X. f x \ a}) \ (\a. closedin X {x \ topspace X. f x \ a})" unfolding continuous_map_upper_lower_semicontinuous_lt_gen by (auto simp: closedin_def Diff_eq Compl_eq not_le Int_Collect_imp_eq) lemma continuous_map_upper_lower_semicontinuous_le: "continuous_map X euclideanreal f \ (\a. closedin X {x \ topspace X. f x \ a}) \ (\a. closedin X {x \ topspace X. f x \ a})" using continuous_map_upper_lower_semicontinuous_le_gen [where U=UNIV] by auto lemma continuous_map_upper_lower_semicontinuous_lte_gen: "continuous_map X (subtopology euclideanreal U) f \ (\x \ topspace X. f x \ U) \ (\a. openin X {x \ topspace X. f x < a}) \ (\a. closedin X {x \ topspace X. f x \ a})" unfolding continuous_map_upper_lower_semicontinuous_lt_gen by (auto simp: closedin_def Diff_eq Compl_eq not_le Int_Collect_imp_eq) lemma continuous_map_upper_lower_semicontinuous_lte: "continuous_map X euclideanreal f \ (\a. openin X {x \ topspace X. f x < a}) \ (\a. closedin X {x \ topspace X. f x \ a})" using continuous_map_upper_lower_semicontinuous_lte_gen [where U=UNIV] by auto subsection\<^marker>\tag important\ \Pullback topology\ text \Pulling back a topology by map gives again a topology. \subtopology\ is a special case of this notion, pulling back by the identity. We introduce the general notion as we will need it to define the strong operator topology on the space of continuous linear operators, by pulling back the product topology on the space of all functions.\ text \\pullback_topology A f T\ is the pullback of the topology \T\ by the map \f\ on the set \A\.\ definition\<^marker>\tag important\ pullback_topology::"('a set) \ ('a \ 'b) \ ('b topology) \ ('a topology)" where "pullback_topology A f T = topology (\S. \U. openin T U \ S = f-`U \ A)" lemma istopology_pullback_topology: "istopology (\S. \U. openin T U \ S = f-`U \ A)" unfolding istopology_def proof (auto) fix K assume "\S\K. \U. openin T U \ S = f -` U \ A" then have "\U. \S\K. openin T (U S) \ S = f-`(U S) \ A" by (rule bchoice) then obtain U where U: "\S\K. openin T (U S) \ S = f-`(U S) \ A" by blast define V where "V = (\S\K. U S)" have "openin T V" "\K = f -` V \ A" unfolding V_def using U by auto then show "\V. openin T V \ \K = f -` V \ A" by auto qed lemma openin_pullback_topology: "openin (pullback_topology A f T) S \ (\U. openin T U \ S = f-`U \ A)" unfolding pullback_topology_def topology_inverse'[OF istopology_pullback_topology] by auto lemma topspace_pullback_topology: "topspace (pullback_topology A f T) = f-`(topspace T) \ A" by (auto simp add: topspace_def openin_pullback_topology) proposition continuous_map_pullback [intro]: assumes "continuous_map T1 T2 g" shows "continuous_map (pullback_topology A f T1) T2 (g o f)" unfolding continuous_map_alt proof (auto) fix U::"'b set" assume "openin T2 U" then have "openin T1 (g-`U \ topspace T1)" using assms unfolding continuous_map_alt by auto have "(g o f)-`U \ topspace (pullback_topology A f T1) = (g o f)-`U \ A \ f-`(topspace T1)" unfolding topspace_pullback_topology by auto also have "... = f-`(g-`U \ topspace T1) \ A " by auto also have "openin (pullback_topology A f T1) (...)" unfolding openin_pullback_topology using \openin T1 (g-`U \ topspace T1)\ by auto finally show "openin (pullback_topology A f T1) ((g \ f) -` U \ topspace (pullback_topology A f T1))" by auto next fix x assume "x \ topspace (pullback_topology A f T1)" then have "f x \ topspace T1" unfolding topspace_pullback_topology by auto then show "g (f x) \ topspace T2" using assms unfolding continuous_map_def by auto qed proposition continuous_map_pullback' [intro]: assumes "continuous_map T1 T2 (f o g)" "topspace T1 \ g-`A" shows "continuous_map T1 (pullback_topology A f T2) g" unfolding continuous_map_alt proof (auto) fix U assume "openin (pullback_topology A f T2) U" then have "\V. openin T2 V \ U = f-`V \ A" unfolding openin_pullback_topology by auto then obtain V where "openin T2 V" "U = f-`V \ A" by blast then have "g -` U \ topspace T1 = g-`(f-`V \ A) \ topspace T1" by blast also have "... = (f o g)-`V \ (g-`A \ topspace T1)" by auto also have "... = (f o g)-`V \ topspace T1" using assms(2) by auto also have "openin T1 (...)" using assms(1) \openin T2 V\ by auto finally show "openin T1 (g -` U \ topspace T1)" by simp next fix x assume "x \ topspace T1" have "(f o g) x \ topspace T2" using assms(1) \x \ topspace T1\ unfolding continuous_map_def by auto then have "g x \ f-`(topspace T2)" unfolding comp_def by blast moreover have "g x \ A" using assms(2) \x \ topspace T1\ by blast ultimately show "g x \ topspace (pullback_topology A f T2)" unfolding topspace_pullback_topology by blast qed subsection\Proper maps (not a priori assumed continuous) \ definition proper_map where "proper_map X Y f \ closed_map X Y f \ (\y \ topspace Y. compactin X {x \ topspace X. f x = y})" lemma proper_imp_closed_map: "proper_map X Y f \ closed_map X Y f" by (simp add: proper_map_def) lemma proper_map_imp_subset_topspace: "proper_map X Y f \ f ` (topspace X) \ topspace Y" by (simp add: closed_map_imp_subset_topspace proper_map_def) lemma closed_injective_imp_proper_map: assumes f: "closed_map X Y f" and inj: "inj_on f (topspace X)" shows "proper_map X Y f" unfolding proper_map_def proof (clarsimp simp: f) show "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y using inj_on_eq_iff [OF inj] that proof - have "{x \ topspace X. f x = y} = {} \ (\a \ topspace X. {x \ topspace X. f x = y} = {a})" using inj_on_eq_iff [OF inj] by auto then show ?thesis using that by (metis (no_types, lifting) compactin_empty compactin_sing) qed qed lemma injective_imp_proper_eq_closed_map: "inj_on f (topspace X) \ (proper_map X Y f \ closed_map X Y f)" using closed_injective_imp_proper_map proper_imp_closed_map by blast lemma homeomorphic_imp_proper_map: "homeomorphic_map X Y f \ proper_map X Y f" by (simp add: closed_injective_imp_proper_map homeomorphic_eq_everything_map) lemma compactin_proper_map_preimage: assumes f: "proper_map X Y f" and "compactin Y K" shows "compactin X {x. x \ topspace X \ f x \ K}" proof - have "f ` (topspace X) \ topspace Y" by (simp add: f proper_map_imp_subset_topspace) have *: "\y. y \ topspace Y \ compactin X {x \ topspace X. f x = y}" using f by (auto simp: proper_map_def) show ?thesis unfolding compactin_def proof clarsimp show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x \ K} \ \\" if \: "\U\\. openin X U" and sub: "{x \ topspace X. f x \ K} \ \\" for \ proof - have "\y \ K. \\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" proof fix y assume "y \ K" then have "compactin X {x \ topspace X. f x = y}" by (metis "*" \compactin Y K\ compactin_subspace subsetD) with \y \ K\ show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" unfolding compactin_def using \ sub by fastforce qed then obtain \ where \: "\y. y \ K \ finite (\ y) \ \ y \ \ \ {x \ topspace X. f x = y} \ \(\ y)" by (metis (full_types)) define F where "F \ \y. topspace Y - f ` (topspace X - \(\ y))" have "\\. finite \ \ \ \ F ` K \ K \ \\" proof (rule compactinD [OF \compactin Y K\]) have "\x. x \ K \ closedin Y (f ` (topspace X - \(\ x)))" using f unfolding proper_map_def closed_map_def by (meson \ \ openin_Union openin_closedin_eq subsetD) then show "openin Y U" if "U \ F ` K" for U using that by (auto simp: F_def) show "K \ \(F ` K)" using \ \compactin Y K\ unfolding F_def compactin_def by fastforce qed then obtain J where "finite J" "J \ K" and J: "K \ \(F ` J)" by (auto simp: ex_finite_subset_image) show ?thesis unfolding F_def proof (intro exI conjI) show "finite (\(\ ` J))" using \ \J \ K\ \finite J\ by blast show "\(\ ` J) \ \" using \ \J \ K\ by blast show "{x \ topspace X. f x \ K} \ \(\(\ ` J))" using J \J \ K\ unfolding F_def by auto qed qed qed qed lemma compact_space_proper_map_preimage: assumes f: "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "compact_space Y" shows "compact_space X" proof - have eq: "topspace X = {x \ topspace X. f x \ topspace Y}" using fim by blast moreover have "compactin Y (topspace Y)" using \compact_space Y\ compact_space_def by auto ultimately show ?thesis unfolding compact_space_def using eq f compactin_proper_map_preimage by fastforce qed lemma proper_map_alt: "proper_map X Y f \ closed_map X Y f \ (\K. compactin Y K \ compactin X {x. x \ topspace X \ f x \ K})" proof (intro iffI conjI allI impI) show "compactin X {x \ topspace X. f x \ K}" if "proper_map X Y f" and "compactin Y K" for K using that by (simp add: compactin_proper_map_preimage) show "proper_map X Y f" if f: "closed_map X Y f \ (\K. compactin Y K \ compactin X {x \ topspace X. f x \ K})" proof - have "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y proof - have "compactin X {x \ topspace X. f x \ {y}}" using f compactin_sing that by fastforce then show ?thesis by auto qed with f show ?thesis by (auto simp: proper_map_def) qed qed (simp add: proper_imp_closed_map) lemma proper_map_on_empty: "topspace X = {} \ proper_map X Y f" by (auto simp: proper_map_def closed_map_on_empty) lemma proper_map_id [simp]: "proper_map X X id" proof (clarsimp simp: proper_map_alt closed_map_id) fix K assume K: "compactin X K" then have "{a \ topspace X. a \ K} = K" by (simp add: compactin_subspace subset_antisym subset_iff) then show "compactin X {a \ topspace X. a \ K}" using K by auto qed lemma proper_map_compose: assumes "proper_map X Y f" "proper_map Y Z g" shows "proper_map X Z (g \ f)" proof - have "closed_map X Y f" and f: "\K. compactin Y K \ compactin X {x \ topspace X. f x \ K}" and "closed_map Y Z g" and g: "\K. compactin Z K \ compactin Y {x \ topspace Y. g x \ K}" using assms by (auto simp: proper_map_alt) show ?thesis unfolding proper_map_alt proof (intro conjI allI impI) show "closed_map X Z (g \ f)" using \closed_map X Y f\ \closed_map Y Z g\ closed_map_compose by blast have "{x \ topspace X. g (f x) \ K} = {x \ topspace X. f x \ {b \ topspace Y. g b \ K}}" for K using \closed_map X Y f\ closed_map_imp_subset_topspace by blast then show "compactin X {x \ topspace X. (g \ f) x \ K}" if "compactin Z K" for K using f [OF g [OF that]] by auto qed qed lemma proper_map_const: "proper_map X Y (\x. c) \ compact_space X \ (topspace X = {} \ closedin Y {c})" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: compact_space_topspace_empty proper_map_on_empty) next case False have *: "compactin X {x \ topspace X. c = y}" if "compact_space X" for y using that unfolding compact_space_def by (metis (mono_tags, lifting) compactin_empty empty_subsetI mem_Collect_eq subsetI subset_antisym) then show ?thesis using closed_compactin closedin_subset by (force simp: False proper_map_def closed_map_const compact_space_def) qed lemma proper_map_inclusion: "S \ topspace X \ proper_map (subtopology X S) X id \ closedin X S \ (\k. compactin X k \ compactin X (S \ k))" by (metis closed_Int_compactin closed_map_inclusion_eq inf.absorb_iff2 inj_on_id injective_imp_proper_eq_closed_map) subsection\Perfect maps (proper, continuous and surjective)\ definition perfect_map where "perfect_map X Y f \ continuous_map X Y f \ proper_map X Y f \ f ` (topspace X) = topspace Y" lemma homeomorphic_imp_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f" by (simp add: homeomorphic_eq_everything_map homeomorphic_imp_proper_map perfect_map_def) lemma perfect_imp_quotient_map: "perfect_map X Y f \ quotient_map X Y f" by (simp add: continuous_closed_imp_quotient_map perfect_map_def proper_map_def) lemma homeomorphic_eq_injective_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f \ inj_on f (topspace X)" using homeomorphic_imp_perfect_map homeomorphic_map_def perfect_imp_quotient_map by blast lemma perfect_injective_eq_homeomorphic_map: "perfect_map X Y f \ inj_on f (topspace X) \ homeomorphic_map X Y f" by (simp add: homeomorphic_eq_injective_perfect_map) lemma perfect_map_id [simp]: "perfect_map X X id" by (simp add: homeomorphic_imp_perfect_map) lemma perfect_map_compose: "\perfect_map X Y f; perfect_map Y Z g\ \ perfect_map X Z (g \ f)" by (meson continuous_map_compose perfect_imp_quotient_map perfect_map_def proper_map_compose quotient_map_compose_eq quotient_map_def) lemma perfect_imp_continuous_map: "perfect_map X Y f \ continuous_map X Y f" using perfect_map_def by blast lemma perfect_imp_closed_map: "perfect_map X Y f \ closed_map X Y f" by (simp add: perfect_map_def proper_map_def) lemma perfect_imp_proper_map: "perfect_map X Y f \ proper_map X Y f" by (simp add: perfect_map_def) lemma perfect_imp_surjective_map: "perfect_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: perfect_map_def) end diff --git a/src/HOL/Analysis/Abstract_Topology_2.thy b/src/HOL/Analysis/Abstract_Topology_2.thy --- a/src/HOL/Analysis/Abstract_Topology_2.thy +++ b/src/HOL/Analysis/Abstract_Topology_2.thy @@ -1,1642 +1,1691 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) section \Abstract Topology 2\ theory Abstract_Topology_2 imports Elementary_Topology Abstract_Topology "HOL-Library.Indicator_Function" begin text \Combination of Elementary and Abstract Topology\ lemma approachable_lt_le2: "(\(d::real) > 0. \x. Q x \ f x < d \ P x) \ (\d>0. \x. f x \ d \ Q x \ P x)" by (meson dense less_eq_real_def order_le_less_trans) lemma triangle_lemma: fixes x y z :: real assumes x: "0 \ x" and y: "0 \ y" and z: "0 \ z" and xy: "x\<^sup>2 \ y\<^sup>2 + z\<^sup>2" shows "x \ y + z" proof - have "y\<^sup>2 + z\<^sup>2 \ y\<^sup>2 + 2 * y * z + z\<^sup>2" using z y by simp with xy have th: "x\<^sup>2 \ (y + z)\<^sup>2" by (simp add: power2_eq_square field_simps) from y z have yz: "y + z \ 0" by arith from power2_le_imp_le[OF th yz] show ?thesis . qed lemma isCont_indicator: fixes x :: "'a::t2_space" shows "isCont (indicator A :: 'a \ real) x = (x \ frontier A)" proof auto fix x assume cts_at: "isCont (indicator A :: 'a \ real) x" and fr: "x \ frontier A" with continuous_at_open have 1: "\V::real set. open V \ indicator A x \ V \ (\U::'a set. open U \ x \ U \ (\y\U. indicator A y \ V))" by auto show False proof (cases "x \ A") assume x: "x \ A" hence "indicator A x \ ({0<..<2} :: real set)" by simp with 1 obtain U where U: "open U" "x \ U" "\y\U. indicator A y \ ({0<..<2} :: real set)" using open_greaterThanLessThan by metis hence "\y\U. indicator A y > (0::real)" unfolding greaterThanLessThan_def by auto hence "U \ A" using indicator_eq_0_iff by force hence "x \ interior A" using U interiorI by auto thus ?thesis using fr unfolding frontier_def by simp next assume x: "x \ A" hence "indicator A x \ ({-1<..<1} :: real set)" by simp with 1 obtain U where U: "open U" "x \ U" "\y\U. indicator A y \ ({-1<..<1} :: real set)" using 1 open_greaterThanLessThan by metis hence "\y\U. indicator A y < (1::real)" unfolding greaterThanLessThan_def by auto hence "U \ -A" by auto hence "x \ interior (-A)" using U interiorI by auto thus ?thesis using fr interior_complement unfolding frontier_def by auto qed next assume nfr: "x \ frontier A" hence "x \ interior A \ x \ interior (-A)" by (auto simp: frontier_def closure_interior) thus "isCont ((indicator A)::'a \ real) x" proof assume int: "x \ interior A" then obtain U where U: "open U" "x \ U" "U \ A" unfolding interior_def by auto hence "\y\U. indicator A y = (1::real)" unfolding indicator_def by auto hence "continuous_on U (indicator A)" by (simp add: indicator_eq_1_iff) thus ?thesis using U continuous_on_eq_continuous_at by auto next assume ext: "x \ interior (-A)" then obtain U where U: "open U" "x \ U" "U \ -A" unfolding interior_def by auto then have "continuous_on U (indicator A)" using continuous_on_topological by (auto simp: subset_iff) thus ?thesis using U continuous_on_eq_continuous_at by auto qed qed lemma islimpt_closure: "\S \ T; \x. \x islimpt S; x \ T\ \ x \ S\ \ S = T \ closure S" using closure_def by fastforce lemma closedin_limpt: "closedin (top_of_set T) S \ S \ T \ (\x. x islimpt S \ x \ T \ x \ S)" proof - have "\U x. \closed U; S = T \ U; x islimpt S; x \ T\ \ x \ S" by (meson IntI closed_limpt inf_le2 islimpt_subset) then show ?thesis by (metis closed_closure closedin_closed closedin_imp_subset islimpt_closure) qed lemma closedin_closed_eq: "closed S \ closedin (top_of_set S) T \ closed T \ T \ S" by (meson closedin_limpt closed_subset closedin_closed_trans) lemma connected_closed_set: "closed S \ connected S \ (\A B. closed A \ closed B \ A \ {} \ B \ {} \ A \ B = S \ A \ B = {})" unfolding connected_closedin_eq closedin_closed_eq connected_closedin_eq by blast text \If a connnected set is written as the union of two nonempty closed sets, then these sets have to intersect.\ lemma connected_as_closed_union: assumes "connected C" "C = A \ B" "closed A" "closed B" "A \ {}" "B \ {}" shows "A \ B \ {}" by (metis assms closed_Un connected_closed_set) lemma closedin_subset_trans: "closedin (top_of_set U) S \ S \ T \ T \ U \ closedin (top_of_set T) S" by (meson closedin_limpt subset_iff) lemma openin_subset_trans: "openin (top_of_set U) S \ S \ T \ T \ U \ openin (top_of_set T) S" by (auto simp: openin_open) lemma closedin_compact: "\compact S; closedin (top_of_set S) T\ \ compact T" by (metis closedin_closed compact_Int_closed) lemma closedin_compact_eq: fixes S :: "'a::t2_space set" shows "compact S \ (closedin (top_of_set S) T \ compact T \ T \ S)" by (metis closedin_imp_subset closedin_compact closed_subset compact_imp_closed) subsection \Closure\ lemma euclidean_closure_of [simp]: "euclidean closure_of S = closure S" by (auto simp: closure_of_def closure_def islimpt_def) lemma closure_openin_Int_closure: assumes ope: "openin (top_of_set U) S" and "T \ U" shows "closure(S \ closure T) = closure(S \ T)" proof obtain V where "open V" and S: "S = U \ V" using ope using openin_open by metis show "closure (S \ closure T) \ closure (S \ T)" proof (clarsimp simp: S) fix x assume "x \ closure (U \ V \ closure T)" then have "V \ closure T \ A \ x \ closure A" for A by (metis closure_mono subsetD inf.coboundedI2 inf_assoc) then have "x \ closure (T \ V)" by (metis \open V\ closure_closure inf_commute open_Int_closure_subset) then show "x \ closure (U \ V \ T)" by (metis \T \ U\ inf.absorb_iff2 inf_assoc inf_commute) qed next show "closure (S \ T) \ closure (S \ closure T)" by (meson Int_mono closure_mono closure_subset order_refl) qed corollary infinite_openin: fixes S :: "'a :: t1_space set" shows "\openin (top_of_set U) S; x \ S; x islimpt U\ \ infinite S" by (clarsimp simp add: openin_open islimpt_eq_acc_point inf_commute) lemma closure_Int_ballI: assumes "\U. \openin (top_of_set S) U; U \ {}\ \ T \ U \ {}" shows "S \ closure T" proof (clarsimp simp: closure_iff_nhds_not_empty) fix x and A and V assume "x \ S" "V \ A" "open V" "x \ V" "T \ A = {}" then have "openin (top_of_set S) (A \ V \ S)" by (simp add: inf_absorb2 openin_subtopology_Int) moreover have "A \ V \ S \ {}" using \x \ V\ \V \ A\ \x \ S\ by auto ultimately show False using \T \ A = {}\ assms by fastforce qed subsection \Frontier\ lemma euclidean_interior_of [simp]: "euclidean interior_of S = interior S" by (auto simp: interior_of_def interior_def) lemma euclidean_frontier_of [simp]: "euclidean frontier_of S = frontier S" by (auto simp: frontier_of_def frontier_def) lemma connected_Int_frontier: "\connected S; S \ T \ {}; S - T \ {}\ \ S \ frontier T \ {}" apply (simp add: frontier_interiors connected_openin, safe) apply (drule_tac x="S \ interior T" in spec, safe) apply (drule_tac [2] x="S \ interior (-T)" in spec) apply (auto simp: disjoint_eq_subset_Compl dest: interior_subset [THEN subsetD]) done subsection \Compactness\ lemma openin_delete: fixes a :: "'a :: t1_space" shows "openin (top_of_set u) S \ openin (top_of_set u) (S - {a})" by (metis Int_Diff open_delete openin_open) lemma compact_eq_openin_cover: "compact S \ (\C. (\c\C. openin (top_of_set S) c) \ S \ \C \ (\D\C. finite D \ S \ \D))" proof safe fix C assume "compact S" and "\c\C. openin (top_of_set S) c" and "S \ \C" then have "\c\{T. open T \ S \ T \ C}. open c" and "S \ \{T. open T \ S \ T \ C}" unfolding openin_open by force+ with \compact S\ obtain D where "D \ {T. open T \ S \ T \ C}" and "finite D" and "S \ \D" by (meson compactE) then have "image (\T. S \ T) D \ C \ finite (image (\T. S \ T) D) \ S \ \(image (\T. S \ T) D)" by auto then show "\D\C. finite D \ S \ \D" .. next assume 1: "\C. (\c\C. openin (top_of_set S) c) \ S \ \C \ (\D\C. finite D \ S \ \D)" show "compact S" proof (rule compactI) fix C let ?C = "image (\T. S \ T) C" assume "\t\C. open t" and "S \ \C" then have "(\c\?C. openin (top_of_set S) c) \ S \ \?C" unfolding openin_open by auto with 1 obtain D where "D \ ?C" and "finite D" and "S \ \D" by metis let ?D = "inv_into C (\T. S \ T) ` D" have "?D \ C \ finite ?D \ S \ \?D" proof (intro conjI) from \D \ ?C\ show "?D \ C" by (fast intro: inv_into_into) from \finite D\ show "finite ?D" by (rule finite_imageI) from \S \ \D\ show "S \ \?D" by (metis \D \ (\) S ` C\ image_inv_into_cancel inf_Sup le_infE) qed then show "\D\C. finite D \ S \ \D" .. qed qed subsection \Continuity\ lemma interior_image_subset: assumes "inj f" "\x. continuous (at x) f" shows "interior (f ` S) \ f ` (interior S)" proof fix x assume "x \ interior (f ` S)" then obtain T where as: "open T" "x \ T" "T \ f ` S" .. then have "x \ f ` S" by auto then obtain y where y: "y \ S" "x = f y" by auto have "open (f -` T)" using assms \open T\ by (simp add: continuous_at_imp_continuous_on open_vimage) moreover have "y \ vimage f T" using \x = f y\ \x \ T\ by simp moreover have "vimage f T \ S" using \T \ image f S\ \inj f\ unfolding inj_on_def subset_eq by auto ultimately have "y \ interior S" .. with \x = f y\ show "x \ f ` interior S" .. qed subsection\<^marker>\tag unimportant\ \Equality of continuous functions on closure and related results\ lemma continuous_closedin_preimage_constant: fixes f :: "_ \ 'b::t1_space" shows "continuous_on S f \ closedin (top_of_set S) {x \ S. f x = a}" using continuous_closedin_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq) lemma continuous_closed_preimage_constant: fixes f :: "_ \ 'b::t1_space" shows "continuous_on S f \ closed S \ closed {x \ S. f x = a}" using continuous_closed_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq) lemma continuous_constant_on_closure: fixes f :: "_ \ 'b::t1_space" assumes "continuous_on (closure S) f" and "\x. x \ S \ f x = a" and "x \ closure S" shows "f x = a" using continuous_closed_preimage_constant[of "closure S" f a] assms closure_minimal[of S "{x \ closure S. f x = a}"] closure_subset unfolding subset_eq by auto lemma image_closure_subset: assumes contf: "continuous_on (closure S) f" and "closed T" and "(f ` S) \ T" shows "f ` (closure S) \ T" proof - have "S \ {x \ closure S. f x \ T}" using assms(3) closure_subset by auto moreover have "closed (closure S \ f -` T)" using continuous_closed_preimage[OF contf] \closed T\ by auto ultimately have "closure S = (closure S \ f -` T)" using closure_minimal[of S "(closure S \ f -` T)"] by auto then show ?thesis by auto qed subsection\<^marker>\tag unimportant\ \A function constant on a set\ definition constant_on (infixl "(constant'_on)" 50) where "f constant_on A \ \y. \x\A. f x = y" lemma constant_on_subset: "\f constant_on A; B \ A\ \ f constant_on B" unfolding constant_on_def by blast lemma injective_not_constant: fixes S :: "'a::{perfect_space} set" shows "\open S; inj_on f S; f constant_on S\ \ S = {}" unfolding constant_on_def by (metis equals0I inj_on_contraD islimpt_UNIV islimpt_def) lemma constant_on_compose: assumes "f constant_on A" shows "g \ f constant_on A" using assms by (auto simp: constant_on_def) lemma not_constant_onI: "f x \ f y \ x \ A \ y \ A \ \f constant_on A" unfolding constant_on_def by metis lemma constant_onE: assumes "f constant_on S" and "\x. x\S \ f x = g x" shows "g constant_on S" using assms unfolding constant_on_def by simp lemma constant_on_closureI: fixes f :: "_ \ 'b::t1_space" assumes cof: "f constant_on S" and contf: "continuous_on (closure S) f" shows "f constant_on (closure S)" using continuous_constant_on_closure [OF contf] cof unfolding constant_on_def by metis subsection\<^marker>\tag unimportant\ \Continuity relative to a union.\ lemma continuous_on_Un_local: "\closedin (top_of_set (S \ T)) S; closedin (top_of_set (S \ T)) T; continuous_on S f; continuous_on T f\ \ continuous_on (S \ T) f" unfolding continuous_on closedin_limpt by (metis Lim_trivial_limit Lim_within_Un Un_iff trivial_limit_within) lemma continuous_on_cases_local: "\closedin (top_of_set (S \ T)) S; closedin (top_of_set (S \ T)) T; continuous_on S f; continuous_on T g; \x. \x \ S \ \P x \ x \ T \ P x\ \ f x = g x\ \ continuous_on (S \ T) (\x. if P x then f x else g x)" by (rule continuous_on_Un_local) (auto intro: continuous_on_eq) lemma continuous_on_cases_le: fixes h :: "'a :: topological_space \ real" assumes "continuous_on {x \ S. h x \ a} f" and "continuous_on {x \ S. a \ h x} g" and h: "continuous_on S h" and "\x. \x \ S; h x = a\ \ f x = g x" shows "continuous_on S (\x. if h x \ a then f(x) else g(x))" proof - have S: "S = (S \ h -` atMost a) \ (S \ h -` atLeast a)" by force have 1: "closedin (top_of_set S) (S \ h -` atMost a)" by (rule continuous_closedin_preimage [OF h closed_atMost]) have 2: "closedin (top_of_set S) (S \ h -` atLeast a)" by (rule continuous_closedin_preimage [OF h closed_atLeast]) have [simp]: "S \ h -` {..a} = {x \ S. h x \ a}" "S \ h -` {a..} = {x \ S. a \ h x}" by auto have "continuous_on (S \ h -` {..a} \ S \ h -` {a..}) (\x. if h x \ a then f x else g x)" by (intro continuous_on_cases_local) (use 1 2 S assms in auto) then show ?thesis using S by force qed lemma continuous_on_cases_1: fixes S :: "real set" assumes "continuous_on {t \ S. t \ a} f" and "continuous_on {t \ S. a \ t} g" and "a \ S \ f a = g a" shows "continuous_on S (\t. if t \ a then f(t) else g(t))" using assms by (auto intro: continuous_on_cases_le [where h = id, simplified]) subsection\<^marker>\tag unimportant\\Inverse function property for open/closed maps\ lemma continuous_on_inverse_open_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "\x. x \ S \ g (f x) = x" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" shows "continuous_on T g" proof - from imf injf have gTS: "g ` T = S" by force from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U by force show ?thesis by (simp add: continuous_on_open [of T g] gTS) (metis openin_imp_subset fU oo) qed lemma continuous_on_inverse_closed_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "\x. x \ S \ g(f x) = x" and oo: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" shows "continuous_on T g" proof - from imf injf have gTS: "g ` T = S" by force from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U by force show ?thesis by (simp add: continuous_on_closed [of T g] gTS) (metis closedin_imp_subset fU oo) qed lemma homeomorphism_injective_open_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "inj_on f S" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" obtains g where "homeomorphism S T f g" proof have "continuous_on T (inv_into S f)" by (metis contf continuous_on_inverse_open_map imf injf inv_into_f_f oo) with imf injf contf show "homeomorphism S T f (inv_into S f)" by (auto simp: homeomorphism_def) qed lemma homeomorphism_injective_closed_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "inj_on f S" and oo: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" obtains g where "homeomorphism S T f g" proof have "continuous_on T (inv_into S f)" by (metis contf continuous_on_inverse_closed_map imf injf inv_into_f_f oo) with imf injf contf show "homeomorphism S T f (inv_into S f)" by (auto simp: homeomorphism_def) qed lemma homeomorphism_imp_open_map: assumes hom: "homeomorphism S T f g" and oo: "openin (top_of_set S) U" shows "openin (top_of_set T) (f ` U)" proof - from hom oo have [simp]: "f ` U = T \ g -` U" using openin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) from hom have "continuous_on T g" unfolding homeomorphism_def by blast moreover have "g ` T = S" by (metis hom homeomorphism_def) ultimately show ?thesis by (simp add: continuous_on_open oo) qed lemma homeomorphism_imp_closed_map: assumes hom: "homeomorphism S T f g" and oo: "closedin (top_of_set S) U" shows "closedin (top_of_set T) (f ` U)" proof - from hom oo have [simp]: "f ` U = T \ g -` U" using closedin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) from hom have "continuous_on T g" unfolding homeomorphism_def by blast moreover have "g ` T = S" by (metis hom homeomorphism_def) ultimately show ?thesis by (simp add: continuous_on_closed oo) qed subsection\<^marker>\tag unimportant\ \Seperability\ lemma subset_second_countable: obtains \ :: "'a:: second_countable_topology set set" where "countable \" "{} \ \" "\C. C \ \ \ openin(top_of_set S) C" "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" proof - obtain \ :: "'a set set" where "countable \" and opeB: "\C. C \ \ \ openin(top_of_set S) C" and \: "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" proof - obtain \ :: "'a set set" where "countable \" and ope: "\C. C \ \ \ open C" and \: "\S. open S \ \U. U \ \ \ S = \U" by (metis univ_second_countable that) show ?thesis proof show "countable ((\C. S \ C) ` \)" by (simp add: \countable \\) show "\C. C \ (\) S ` \ \ openin (top_of_set S) C" using ope by auto show "\T. openin (top_of_set S) T \ \\\(\) S ` \. T = \\" by (metis \ image_mono inf_Sup openin_open) qed qed show ?thesis proof show "countable (\ - {{}})" using \countable \\ by blast show "\C. \C \ \ - {{}}\ \ openin (top_of_set S) C" by (simp add: \\C. C \ \ \ openin (top_of_set S) C\) show "\\\\ - {{}}. T = \\" if "openin (top_of_set S) T" for T using \ [OF that] apply clarify by (rule_tac x="\ - {{}}" in exI, auto) qed auto qed lemma Lindelof_openin: fixes \ :: "'a::second_countable_topology set set" assumes "\S. S \ \ \ openin (top_of_set U) S" obtains \' where "\' \ \" "countable \'" "\\' = \\" proof - have "\S. S \ \ \ \T. open T \ S = U \ T" using assms by (simp add: openin_open) then obtain tf where tf: "\S. S \ \ \ open (tf S) \ (S = U \ tf S)" by metis have [simp]: "\\'. \' \ \ \ \\' = U \ \(tf ` \')" using tf by fastforce obtain \ where "countable \ \ \ \ tf ` \" "\\ = \(tf ` \)" using tf by (force intro: Lindelof [of "tf ` \"]) then obtain \' where \': "\' \ \" "countable \'" "\\' = \\" by (clarsimp simp add: countable_subset_image) then show ?thesis .. qed subsection\<^marker>\tag unimportant\\Closed Maps\ lemma continuous_imp_closed_map: fixes f :: "'a::t2_space \ 'b::t2_space" assumes "closedin (top_of_set S) U" "continuous_on S f" "f ` S = T" "compact S" shows "closedin (top_of_set T) (f ` U)" by (metis assms closedin_compact_eq compact_continuous_image continuous_on_subset subset_image_iff) lemma closed_map_restrict: assumes cloU: "closedin (top_of_set (S \ f -` T')) U" and cc: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" and "T' \ T" shows "closedin (top_of_set T') (f ` U)" proof - obtain V where "closed V" "U = S \ f -` T' \ V" using cloU by (auto simp: closedin_closed) with cc [of "S \ V"] \T' \ T\ show ?thesis by (fastforce simp add: closedin_closed) qed subsection\<^marker>\tag unimportant\\Open Maps\ lemma open_map_restrict: assumes opeU: "openin (top_of_set (S \ f -` T')) U" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" and "T' \ T" shows "openin (top_of_set T') (f ` U)" proof - obtain V where "open V" "U = S \ f -` T' \ V" using opeU by (auto simp: openin_open) with oo [of "S \ V"] \T' \ T\ show ?thesis by (fastforce simp add: openin_open) qed subsection\<^marker>\tag unimportant\\Quotient maps\ lemma quotient_map_imp_continuous_open: assumes T: "f ` S \ T" and ope: "\U. U \ T \ (openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U)" shows "continuous_on S f" proof - have [simp]: "S \ f -` f ` S = S" by auto show ?thesis by (meson T continuous_on_open_gen ope openin_imp_subset) qed lemma quotient_map_imp_continuous_closed: assumes T: "f ` S \ T" and ope: "\U. U \ T \ (closedin (top_of_set S) (S \ f -` U) \ closedin (top_of_set T) U)" shows "continuous_on S f" proof - have [simp]: "S \ f -` f ` S = S" by auto show ?thesis by (meson T closedin_imp_subset continuous_on_closed_gen ope) qed lemma open_map_imp_quotient_map: assumes contf: "continuous_on S f" and T: "T \ f ` S" and ope: "\T. openin (top_of_set S) T \ openin (top_of_set (f ` S)) (f ` T)" shows "openin (top_of_set S) (S \ f -` T) = openin (top_of_set (f ` S)) T" proof - have "T = f ` (S \ f -` T)" using T by blast then show ?thesis using "ope" contf continuous_on_open by metis qed lemma closed_map_imp_quotient_map: assumes contf: "continuous_on S f" and T: "T \ f ` S" and ope: "\T. closedin (top_of_set S) T \ closedin (top_of_set (f ` S)) (f ` T)" shows "openin (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" (is "?lhs = ?rhs") proof assume ?lhs then have *: "closedin (top_of_set S) (S - (S \ f -` T))" using closedin_diff by fastforce have [simp]: "(f ` S - f ` (S - (S \ f -` T))) = T" using T by blast show ?rhs using ope [OF *, unfolded closedin_def] by auto next assume ?rhs with contf show ?lhs by (auto simp: continuous_on_open) qed lemma continuous_right_inverse_imp_quotient_map: assumes contf: "continuous_on S f" and imf: "f ` S \ T" and contg: "continuous_on T g" and img: "g ` T \ S" and fg [simp]: "\y. y \ T \ f(g y) = y" and U: "U \ T" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U" (is "?lhs = ?rhs") proof - have f: "\Z. openin (top_of_set (f ` S)) Z \ openin (top_of_set S) (S \ f -` Z)" and g: "\Z. openin (top_of_set (g ` T)) Z \ openin (top_of_set T) (T \ g -` Z)" using contf contg by (auto simp: continuous_on_open) show ?thesis proof have "T \ g -` (g ` T \ (S \ f -` U)) = {x \ T. f (g x) \ U}" using imf img by blast also have "... = U" using U by auto finally have eq: "T \ g -` (g ` T \ (S \ f -` U)) = U" . assume ?lhs then have *: "openin (top_of_set (g ` T)) (g ` T \ (S \ f -` U))" by (meson img openin_Int openin_subtopology_Int_subset openin_subtopology_self) show ?rhs using g [OF *] eq by auto next assume rhs: ?rhs show ?lhs by (metis f fg image_eqI image_subset_iff imf img openin_subopen openin_subtopology_self openin_trans rhs) qed qed lemma continuous_left_inverse_imp_quotient_map: assumes "continuous_on S f" and "continuous_on (f ` S) g" and "\x. x \ S \ g(f x) = x" and "U \ f ` S" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set (f ` S)) U" using assms by (intro continuous_right_inverse_imp_quotient_map) auto lemma continuous_imp_quotient_map: fixes f :: "'a::t2_space \ 'b::t2_space" assumes "continuous_on S f" "f ` S = T" "compact S" "U \ T" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U" by (simp add: assms closed_map_imp_quotient_map continuous_imp_closed_map) subsection\<^marker>\tag unimportant\\Pasting lemmas for functions, for of casewise definitions\ subsubsection\on open sets\ lemma pasting_lemma: assumes ope: "\i. i \ I \ openin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" unfolding continuous_map_openin_preimage_eq proof (intro conjI allI impI) show "g ` topspace X \ topspace Y" using g cont continuous_map_image_subset_topspace by fastforce next fix U assume Y: "openin Y U" have T: "T i \ topspace X" if "i \ I" for i using ope by (simp add: openin_subset that) have *: "topspace X \ g -` U = (\i \ I. T i \ f i -` U)" using f g T by fastforce have "\i. i \ I \ openin X (T i \ f i -` U)" using cont unfolding continuous_map_openin_preimage_eq by (metis Y T inf.commute inf_absorb1 ope topspace_subtopology openin_trans_full) then show "openin X (topspace X \ g -` U)" by (auto simp: *) qed lemma pasting_lemma_exists: assumes X: "topspace X \ (\i \ I. T i)" and ope: "\i. i \ I \ openin X (T i)" and cont: "\i. i \ I \ continuous_map (subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof let ?h = "\x. f (SOME i. i \ I \ x \ T i) x" show "continuous_map X Y ?h" apply (rule pasting_lemma [OF ope cont]) apply (blast intro: f)+ by (metis (no_types, lifting) UN_E X subsetD someI_ex) show "f (SOME i. i \ I \ x \ T i) x = f i x" if "i \ I" "x \ topspace X \ T i" for i x by (metis (no_types, lifting) IntD2 IntI f someI_ex that) qed lemma pasting_lemma_locally_finite: assumes fin: "\x. x \ topspace X \ \V. openin X V \ x \ V \ finite {i \ I. T i \ V \ {}}" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" unfolding continuous_map_closedin_preimage_eq proof (intro conjI allI impI) show "g ` topspace X \ topspace Y" using g cont continuous_map_image_subset_topspace by fastforce next fix U assume Y: "closedin Y U" have T: "T i \ topspace X" if "i \ I" for i using clo by (simp add: closedin_subset that) have *: "topspace X \ g -` U = (\i \ I. T i \ f i -` U)" using f g T by fastforce have cTf: "\i. i \ I \ closedin X (T i \ f i -` U)" using cont unfolding continuous_map_closedin_preimage_eq topspace_subtopology by (simp add: Int_absorb1 T Y clo closedin_closed_subtopology) have sub: "{Z \ (\i. T i \ f i -` U) ` I. Z \ V \ {}} \ (\i. T i \ f i -` U) ` {i \ I. T i \ V \ {}}" for V by auto have 1: "(\i\I. T i \ f i -` U) \ topspace X" using T by blast then have "locally_finite_in X ((\i. T i \ f i -` U) ` I)" unfolding locally_finite_in_def using finite_subset [OF sub] fin by force then show "closedin X (topspace X \ g -` U)" by (smt (verit, best) * cTf closedin_locally_finite_Union image_iff) qed subsubsection\Likewise on closed sets, with a finiteness assumption\ lemma pasting_lemma_closed: assumes fin: "finite I" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" using pasting_lemma_locally_finite [OF _ clo cont f g] fin by auto lemma pasting_lemma_exists_locally_finite: assumes fin: "\x. x \ topspace X \ \V. openin X V \ x \ V \ finite {i \ I. T i \ V \ {}}" and X: "topspace X \ \(T ` I)" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof show "continuous_map X Y (\x. f(@i. i \ I \ x \ T i) x)" apply (rule pasting_lemma_locally_finite [OF fin]) apply (blast intro: assms)+ by (metis (no_types, lifting) UN_E X set_rev_mp someI_ex) next fix x i assume "i \ I" and "x \ topspace X \ T i" then show "f (SOME i. i \ I \ x \ T i) x = f i x" by (metis (mono_tags, lifting) IntE IntI f someI2) qed lemma pasting_lemma_exists_closed: assumes fin: "finite I" and X: "topspace X \ \(T ` I)" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof show "continuous_map X Y (\x. f (SOME i. i \ I \ x \ T i) x)" apply (rule pasting_lemma_closed [OF \finite I\ clo cont]) apply (blast intro: f)+ by (metis (mono_tags, lifting) UN_iff X someI_ex subset_iff) next fix x i assume "i \ I" "x \ topspace X \ T i" then show "f (SOME i. i \ I \ x \ T i) x = f i x" by (metis (no_types, lifting) IntD2 IntI f someI_ex) qed lemma continuous_map_cases: assumes f: "continuous_map (subtopology X (X closure_of {x. P x})) Y f" and g: "continuous_map (subtopology X (X closure_of {x. \ P x})) Y g" and fg: "\x. x \ X frontier_of {x. P x} \ f x = g x" shows "continuous_map X Y (\x. if P x then f x else g x)" proof (rule pasting_lemma_closed) let ?f = "\b. if b then f else g" let ?g = "\x. if P x then f x else g x" let ?T = "\b. if b then X closure_of {x. P x} else X closure_of {x. ~P x}" show "finite {True,False}" by auto have eq: "topspace X - Collect P = topspace X \ {x. \ P x}" by blast show "?f i x = ?f j x" if "i \ {True,False}" "j \ {True,False}" and x: "x \ topspace X \ ?T i \ ?T j" for i j x proof - have "f x = g x" if "i" "\ j" by (smt (verit, best) Diff_Diff_Int closure_of_interior_of closure_of_restrict eq fg frontier_of_closures interior_of_complement that x) moreover have "g x = f x" if "x \ X closure_of {x. \ P x}" "x \ X closure_of Collect P" "\ i" "j" for x by (metis IntI closure_of_restrict eq fg frontier_of_closures that) ultimately show ?thesis using that by (auto simp flip: closure_of_restrict) qed show "\j. j \ {True,False} \ x \ ?T j \ (if P x then f x else g x) = ?f j x" if "x \ topspace X" for x by simp (metis in_closure_of mem_Collect_eq that) qed (auto simp: f g) lemma continuous_map_cases_alt: assumes f: "continuous_map (subtopology X (X closure_of {x \ topspace X. P x})) Y f" and g: "continuous_map (subtopology X (X closure_of {x \ topspace X. ~P x})) Y g" and fg: "\x. x \ X frontier_of {x \ topspace X. P x} \ f x = g x" shows "continuous_map X Y (\x. if P x then f x else g x)" apply (rule continuous_map_cases) using assms apply (simp_all add: Collect_conj_eq closure_of_restrict [symmetric] frontier_of_restrict [symmetric]) done lemma continuous_map_cases_function: assumes contp: "continuous_map X Z p" and contf: "continuous_map (subtopology X {x \ topspace X. p x \ Z closure_of U}) Y f" and contg: "continuous_map (subtopology X {x \ topspace X. p x \ Z closure_of (topspace Z - U)}) Y g" and fg: "\x. \x \ topspace X; p x \ Z frontier_of U\ \ f x = g x" shows "continuous_map X Y (\x. if p x \ U then f x else g x)" proof (rule continuous_map_cases_alt) show "continuous_map (subtopology X (X closure_of {x \ topspace X. p x \ U})) Y f" proof (rule continuous_map_from_subtopology_mono) let ?T = "{x \ topspace X. p x \ Z closure_of U}" show "continuous_map (subtopology X ?T) Y f" by (simp add: contf) show "X closure_of {x \ topspace X. p x \ U} \ ?T" by (rule continuous_map_closure_preimage_subset [OF contp]) qed show "continuous_map (subtopology X (X closure_of {x \ topspace X. p x \ U})) Y g" proof (rule continuous_map_from_subtopology_mono) let ?T = "{x \ topspace X. p x \ Z closure_of (topspace Z - U)}" show "continuous_map (subtopology X ?T) Y g" by (simp add: contg) have "X closure_of {x \ topspace X. p x \ U} \ X closure_of {x \ topspace X. p x \ topspace Z - U}" by (smt (verit, del_insts) DiffI mem_Collect_eq subset_iff closure_of_mono continuous_map_closedin contp) then show "X closure_of {x \ topspace X. p x \ U} \ ?T" by (rule order_trans [OF _ continuous_map_closure_preimage_subset [OF contp]]) qed next show "f x = g x" if "x \ X frontier_of {x \ topspace X. p x \ U}" for x using that continuous_map_frontier_frontier_preimage_subset [OF contp, of U] fg by blast qed subsection \Retractions\ definition\<^marker>\tag important\ retraction :: "('a::topological_space) set \ 'a set \ ('a \ 'a) \ bool" where "retraction S T r \ T \ S \ continuous_on S r \ r ` S \ T \ (\x\T. r x = x)" definition\<^marker>\tag important\ retract_of (infixl "retract'_of" 50) where "T retract_of S \ (\r. retraction S T r)" lemma retraction_idempotent: "retraction S T r \ x \ S \ r (r x) = r x" unfolding retraction_def by auto text \Preservation of fixpoints under (more general notion of) retraction\ lemma invertible_fixpoint_property: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" assumes contt: "continuous_on T i" and "i ` T \ S" and contr: "continuous_on S r" and "r ` S \ T" and ri: "\y. y \ T \ r (i y) = y" and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" and contg: "continuous_on T g" and "g ` T \ T" obtains y where "y \ T" and "g y = y" proof - have "\x\S. (i \ g \ r) x = x" proof (rule FP) show "continuous_on S (i \ g \ r)" by (meson contt contr assms(4) contg assms(8) continuous_on_compose continuous_on_subset) show "(i \ g \ r) ` S \ S" using assms(2,4,8) by force qed then obtain x where x: "x \ S" "(i \ g \ r) x = x" .. then have *: "g (r x) \ T" using assms(4,8) by auto have "r ((i \ g \ r) x) = r x" using x by auto then show ?thesis using "*" ri that by auto qed lemma homeomorphic_fixpoint_property: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" assumes "S homeomorphic T" shows "(\f. continuous_on S f \ f ` S \ S \ (\x\S. f x = x)) \ (\g. continuous_on T g \ g ` T \ T \ (\y\T. g y = y))" (is "?lhs = ?rhs") proof - obtain r i where r: "\x\S. i (r x) = x" "r ` S = T" "continuous_on S r" "\y\T. r (i y) = y" "i ` T = S" "continuous_on T i" using assms unfolding homeomorphic_def homeomorphism_def by blast show ?thesis proof assume ?lhs with r show ?rhs by (metis invertible_fixpoint_property[of T i S r] order_refl) next assume ?rhs with r show ?lhs by (metis invertible_fixpoint_property[of S r T i] order_refl) qed qed lemma retract_fixpoint_property: fixes f :: "'a::topological_space \ 'b::topological_space" and S :: "'a set" assumes "T retract_of S" and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" and contg: "continuous_on T g" and "g ` T \ T" obtains y where "y \ T" and "g y = y" proof - obtain h where "retraction S T h" using assms(1) unfolding retract_of_def .. then show ?thesis unfolding retraction_def using invertible_fixpoint_property[OF continuous_on_id _ _ _ _ FP] by (metis assms(4) contg image_ident that) qed lemma retraction: "retraction S T r \ T \ S \ continuous_on S r \ r ` S = T \ (\x \ T. r x = x)" by (force simp: retraction_def) lemma retractionE: \ \yields properties normalized wrt. simp -- less likely to loop\ assumes "retraction S T r" obtains "T = r ` S" "r ` S \ S" "continuous_on S r" "\x. x \ S \ r (r x) = r x" proof (rule that) from retraction [of S T r] assms have "T \ S" "continuous_on S r" "r ` S = T" and "\x \ T. r x = x" by simp_all then show "T = r ` S" "r ` S \ S" "continuous_on S r" by simp_all from \\x \ T. r x = x\ have "r x = x" if "x \ T" for x using that by simp with \r ` S = T\ show "r (r x) = r x" if "x \ S" for x using that by auto qed lemma retract_ofE: \ \yields properties normalized wrt. simp -- less likely to loop\ assumes "T retract_of S" obtains r where "T = r ` S" "r ` S \ S" "continuous_on S r" "\x. x \ S \ r (r x) = r x" proof - from assms obtain r where "retraction S T r" by (auto simp add: retract_of_def) with that show thesis by (auto elim: retractionE) qed lemma retract_of_imp_extensible: assumes "S retract_of T" and "continuous_on S f" and "f ` S \ U" obtains g where "continuous_on T g" "g ` T \ U" "\x. x \ S \ g x = f x" proof - from \S retract_of T\ obtain r where "retraction T S r" by (auto simp add: retract_of_def) show thesis by (rule that [of "f \ r"]) (use \continuous_on S f\ \f ` S \ U\ \retraction T S r\ in \auto simp: continuous_on_compose2 retraction\) qed lemma idempotent_imp_retraction: assumes "continuous_on S f" and "f ` S \ S" and "\x. x \ S \ f(f x) = f x" shows "retraction S (f ` S) f" by (simp add: assms retraction) lemma retraction_subset: assumes "retraction S T r" and "T \ s'" and "s' \ S" shows "retraction s' T r" unfolding retraction_def by (metis assms continuous_on_subset image_mono retraction) lemma retract_of_subset: assumes "T retract_of S" and "T \ s'" and "s' \ S" shows "T retract_of s'" by (meson assms retract_of_def retraction_subset) lemma retraction_refl [simp]: "retraction S S (\x. x)" by (simp add: retraction) lemma retract_of_refl [iff]: "S retract_of S" unfolding retract_of_def retraction_def using continuous_on_id by blast lemma retract_of_imp_subset: "S retract_of T \ S \ T" by (simp add: retract_of_def retraction_def) lemma retract_of_empty [simp]: "({} retract_of S) \ S = {}" "(S retract_of {}) \ S = {}" by (auto simp: retract_of_def retraction_def) lemma retract_of_singleton [iff]: "({x} retract_of S) \ x \ S" unfolding retract_of_def retraction_def by force lemma retraction_comp: "\retraction S T f; retraction T U g\ \ retraction S U (g \ f)" by (smt (verit, best) comp_apply continuous_on_compose image_comp retraction subset_iff) lemma retract_of_trans [trans]: assumes "S retract_of T" and "T retract_of U" shows "S retract_of U" using assms by (auto simp: retract_of_def intro: retraction_comp) lemma closedin_retract: fixes S :: "'a :: t2_space set" assumes "S retract_of T" shows "closedin (top_of_set T) S" proof - obtain r where r: "S \ T" "continuous_on T r" "r ` T \ S" "\x. x \ S \ r x = x" using assms by (auto simp: retract_of_def retraction_def) have "S = {x\T. x = r x}" using r by auto also have "\ = T \ ((\x. (x, r x)) -` ({y. \x. y = (x, x)}))" unfolding vimage_def mem_Times_iff fst_conv snd_conv using r by auto also have "closedin (top_of_set T) \" by (rule continuous_closedin_preimage) (auto intro!: closed_diagonal continuous_on_Pair r) finally show ?thesis . qed lemma closedin_self [simp]: "closedin (top_of_set S) S" by simp lemma retract_of_closed: fixes S :: "'a :: t2_space set" shows "\closed T; S retract_of T\ \ closed S" by (metis closedin_retract closedin_closed_eq) lemma retract_of_compact: "\compact T; S retract_of T\ \ compact S" by (metis compact_continuous_image retract_of_def retraction) lemma retract_of_connected: "\connected T; S retract_of T\ \ connected S" by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) lemma retraction_openin_vimage_iff: assumes r: "retraction S T r" and "U \ T" shows "openin (top_of_set S) (S \ r -` U) \ openin (top_of_set T) U" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using r retraction_def retractionE by (smt (verit, best) continuous_right_inverse_imp_quotient_map retraction_subset \U \ T\) show "?rhs \ ?lhs" by (meson continuous_openin_preimage r retraction_def) qed lemma retract_of_Times: "\S retract_of s'; T retract_of t'\ \ (S \ T) retract_of (s' \ t')" apply (simp add: retract_of_def retraction_def Sigma_mono, clarify) apply (rename_tac f g) apply (rule_tac x="\z. ((f \ fst) z, (g \ snd) z)" in exI) apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ done subsection\Retractions on a topological space\ definition retract_of_space :: "'a set \ 'a topology \ bool" (infix "retract'_of'_space" 50) where "S retract_of_space X \ S \ topspace X \ (\r. continuous_map X (subtopology X S) r \ (\x \ S. r x = x))" lemma retract_of_space_retraction_maps: "S retract_of_space X \ S \ topspace X \ (\r. retraction_maps X (subtopology X S) r id)" by (auto simp: retract_of_space_def retraction_maps_def) lemma retract_of_space_section_map: "S retract_of_space X \ S \ topspace X \ section_map (subtopology X S) X id" unfolding retract_of_space_def retraction_maps_def section_map_def by (auto simp: continuous_map_from_subtopology) lemma retract_of_space_imp_subset: "S retract_of_space X \ S \ topspace X" by (simp add: retract_of_space_def) lemma retract_of_space_topspace: "topspace X retract_of_space X" using retract_of_space_def by force lemma retract_of_space_empty [simp]: "{} retract_of_space X \ topspace X = {}" by (auto simp: continuous_map_def retract_of_space_def) lemma retract_of_space_singleton [simp]: "{a} retract_of_space X \ a \ topspace X" proof - have "continuous_map X (subtopology X {a}) (\x. a) \ (\x. a) a = a" if "a \ topspace X" using that by simp then show ?thesis by (force simp: retract_of_space_def) qed lemma retract_of_space_trans: assumes "S retract_of_space X" "T retract_of_space subtopology X S" shows "T retract_of_space X" using assms apply (simp add: retract_of_space_retraction_maps) by (metis id_comp inf.absorb_iff2 retraction_maps_compose subtopology_subtopology) lemma retract_of_space_subtopology: assumes "S retract_of_space X" "S \ U" shows "S retract_of_space subtopology X U" using assms apply (clarsimp simp: retract_of_space_def) by (metis continuous_map_from_subtopology inf.absorb2 subtopology_subtopology) lemma retract_of_space_clopen: assumes "openin X S" "closedin X S" "S = {} \ topspace X = {}" shows "S retract_of_space X" proof (cases "S = {}") case False then obtain a where "a \ S" by blast show ?thesis unfolding retract_of_space_def proof (intro exI conjI) show "S \ topspace X" by (simp add: assms closedin_subset) have "continuous_map X X (\x. if x \ S then x else a)" proof (rule continuous_map_cases) show "continuous_map (subtopology X (X closure_of {x. x \ S})) X (\x. x)" by (simp add: continuous_map_from_subtopology) show "continuous_map (subtopology X (X closure_of {x. x \ S})) X (\x. a)" using \S \ topspace X\ \a \ S\ by force show "x = a" if "x \ X frontier_of {x. x \ S}" for x using assms that clopenin_eq_frontier_of by fastforce qed then show "continuous_map X (subtopology X S) (\x. if x \ S then x else a)" using \S \ topspace X\ \a \ S\ by (auto simp: continuous_map_in_subtopology) qed auto qed (use assms in auto) lemma retract_of_space_disjoint_union: assumes "openin X S" "openin X T" and ST: "disjnt S T" "S \ T = topspace X" and "S = {} \ topspace X = {}" shows "S retract_of_space X" proof (rule retract_of_space_clopen) have "S \ T = {}" by (meson ST disjnt_def) then have "S = topspace X - T" using ST by auto then show "closedin X S" using \openin X T\ by blast qed (auto simp: assms) lemma retraction_maps_section_image1: assumes "retraction_maps X Y r s" shows "s ` (topspace Y) retract_of_space X" unfolding retract_of_space_section_map proof show "s ` topspace Y \ topspace X" using assms continuous_map_image_subset_topspace retraction_maps_def by blast show "section_map (subtopology X (s ` topspace Y)) X id" unfolding section_map_def using assms retraction_maps_to_retract_maps by blast qed lemma retraction_maps_section_image2: "retraction_maps X Y r s \ subtopology X (s ` (topspace Y)) homeomorphic_space Y" using embedding_map_imp_homeomorphic_space homeomorphic_space_sym section_imp_embedding_map section_map_def by blast lemma hereditary_imp_retractive_property: assumes "\X S. P X \ P(subtopology X S)" "\X X'. X homeomorphic_space X' \ (P X \ Q X')" assumes "retraction_map X X' r" "P X" shows "Q X'" by (meson assms retraction_map_def retraction_maps_section_image2) subsection\Paths and path-connectedness\ definition pathin :: "'a topology \ (real \ 'a) \ bool" where "pathin X g \ continuous_map (subtopology euclideanreal {0..1}) X g" lemma pathin_compose: "\pathin X g; continuous_map X Y f\ \ pathin Y (f \ g)" by (simp add: continuous_map_compose pathin_def) lemma pathin_subtopology: "pathin (subtopology X S) g \ pathin X g \ (\x \ {0..1}. g x \ S)" by (auto simp: pathin_def continuous_map_in_subtopology) lemma pathin_const: "pathin X (\x. a) \ a \ topspace X" by (simp add: pathin_def) lemma path_start_in_topspace: "pathin X g \ g 0 \ topspace X" by (force simp: pathin_def continuous_map) lemma path_finish_in_topspace: "pathin X g \ g 1 \ topspace X" by (force simp: pathin_def continuous_map) lemma path_image_subset_topspace: "pathin X g \ g ` ({0..1}) \ topspace X" by (force simp: pathin_def continuous_map) definition path_connected_space :: "'a topology \ bool" where "path_connected_space X \ \x \ topspace X. \ y \ topspace X. \g. pathin X g \ g 0 = x \ g 1 = y" definition path_connectedin :: "'a topology \ 'a set \ bool" where "path_connectedin X S \ S \ topspace X \ path_connected_space(subtopology X S)" lemma path_connectedin_absolute [simp]: "path_connectedin (subtopology X S) S \ path_connectedin X S" by (simp add: path_connectedin_def subtopology_subtopology) lemma path_connectedin_subset_topspace: "path_connectedin X S \ S \ topspace X" by (simp add: path_connectedin_def) lemma path_connectedin_subtopology: "path_connectedin (subtopology X S) T \ path_connectedin X T \ T \ S" by (auto simp: path_connectedin_def subtopology_subtopology inf.absorb2) lemma path_connectedin: "path_connectedin X S \ S \ topspace X \ (\x \ S. \y \ S. \g. pathin X g \ g ` {0..1} \ S \ g 0 = x \ g 1 = y)" unfolding path_connectedin_def path_connected_space_def pathin_def continuous_map_in_subtopology by (intro conj_cong refl ball_cong) (simp_all add: inf.absorb_iff2) lemma path_connectedin_topspace: "path_connectedin X (topspace X) \ path_connected_space X" by (simp add: path_connectedin_def) lemma path_connected_imp_connected_space: assumes "path_connected_space X" shows "connected_space X" proof - have *: "\S. connectedin X S \ g 0 \ S \ g 1 \ S" if "pathin X g" for g proof (intro exI conjI) have "continuous_map (subtopology euclideanreal {0..1}) X g" using connectedin_absolute that by (simp add: pathin_def) then show "connectedin X (g ` {0..1})" by (rule connectedin_continuous_map_image) auto qed auto show ?thesis using assms by (auto intro: * simp add: path_connected_space_def connected_space_subconnected Ball_def) qed lemma path_connectedin_imp_connectedin: "path_connectedin X S \ connectedin X S" by (simp add: connectedin_def path_connected_imp_connected_space path_connectedin_def) lemma path_connected_space_topspace_empty: "topspace X = {} \ path_connected_space X" by (simp add: path_connected_space_def) lemma path_connectedin_empty [simp]: "path_connectedin X {}" by (simp add: path_connectedin) lemma path_connectedin_singleton [simp]: "path_connectedin X {a} \ a \ topspace X" proof show "path_connectedin X {a} \ a \ topspace X" by (simp add: path_connectedin) show "a \ topspace X \ path_connectedin X {a}" unfolding path_connectedin using pathin_const by fastforce qed lemma path_connectedin_continuous_map_image: assumes f: "continuous_map X Y f" and S: "path_connectedin X S" shows "path_connectedin Y (f ` S)" proof - have fX: "f ` (topspace X) \ topspace Y" by (metis f continuous_map_image_subset_topspace) show ?thesis unfolding path_connectedin proof (intro conjI ballI; clarify?) fix x assume "x \ S" show "f x \ topspace Y" by (meson S fX \x \ S\ image_subset_iff path_connectedin_subset_topspace set_mp) next fix x y assume "x \ S" and "y \ S" then obtain g where g: "pathin X g" "g ` {0..1} \ S" "g 0 = x" "g 1 = y" using S by (force simp: path_connectedin) show "\g. pathin Y g \ g ` {0..1} \ f ` S \ g 0 = f x \ g 1 = f y" proof (intro exI conjI) show "pathin Y (f \ g)" using \pathin X g\ f pathin_compose by auto qed (use g in auto) qed qed lemma path_connectedin_discrete_topology: "path_connectedin (discrete_topology U) S \ S \ U \ (\a. S \ {a})" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (meson connectedin_discrete_topology path_connectedin_imp_connectedin) show "?rhs \ ?lhs" using subset_singletonD by fastforce qed lemma path_connected_space_discrete_topology: "path_connected_space (discrete_topology U) \ (\a. U \ {a})" by (metis path_connectedin_discrete_topology path_connectedin_topspace path_connected_space_topspace_empty subset_singletonD topspace_discrete_topology) lemma homeomorphic_path_connected_space_imp: "\path_connected_space X; X homeomorphic_space Y\ \ path_connected_space Y" unfolding homeomorphic_space_def homeomorphic_maps_def by (metis (no_types, opaque_lifting) continuous_map_closedin continuous_map_image_subset_topspace imageI order_class.order.antisym path_connectedin_continuous_map_image path_connectedin_topspace subsetI) lemma homeomorphic_path_connected_space: "X homeomorphic_space Y \ path_connected_space X \ path_connected_space Y" by (meson homeomorphic_path_connected_space_imp homeomorphic_space_sym) lemma homeomorphic_map_path_connectedness: assumes "homeomorphic_map X Y f" "U \ topspace X" shows "path_connectedin Y (f ` U) \ path_connectedin X U" unfolding path_connectedin_def proof (intro conj_cong homeomorphic_path_connected_space) show "(f ` U \ topspace Y) = (U \ topspace X)" using assms homeomorphic_imp_surjective_map by blast next assume "U \ topspace X" show "subtopology Y (f ` U) homeomorphic_space subtopology X U" using assms unfolding homeomorphic_eq_everything_map by (metis (no_types, opaque_lifting) assms homeomorphic_map_subtopologies homeomorphic_space homeomorphic_space_sym image_mono inf.absorb_iff2) qed lemma homeomorphic_map_path_connectedness_eq: "homeomorphic_map X Y f \ path_connectedin X U \ U \ topspace X \ path_connectedin Y (f ` U)" by (meson homeomorphic_map_path_connectedness path_connectedin_def) subsection\Connected components\ definition connected_component_of :: "'a topology \ 'a \ 'a \ bool" where "connected_component_of X x y \ \T. connectedin X T \ x \ T \ y \ T" abbreviation connected_component_of_set where "connected_component_of_set X x \ Collect (connected_component_of X x)" definition connected_components_of :: "'a topology \ ('a set) set" where "connected_components_of X \ connected_component_of_set X ` topspace X" lemma connected_component_in_topspace: "connected_component_of X x y \ x \ topspace X \ y \ topspace X" by (meson connected_component_of_def connectedin_subset_topspace in_mono) lemma connected_component_of_refl: "connected_component_of X x x \ x \ topspace X" by (meson connected_component_in_topspace connected_component_of_def connectedin_sing insertI1) lemma connected_component_of_sym: "connected_component_of X x y \ connected_component_of X y x" by (meson connected_component_of_def) lemma connected_component_of_trans: "\connected_component_of X x y; connected_component_of X y z\ \ connected_component_of X x z" unfolding connected_component_of_def using connectedin_Un by blast lemma connected_component_of_mono: "\connected_component_of (subtopology X S) x y; S \ T\ \ connected_component_of (subtopology X T) x y" by (metis connected_component_of_def connectedin_subtopology inf.absorb_iff2 subtopology_subtopology) lemma connected_component_of_set: "connected_component_of_set X x = {y. \T. connectedin X T \ x \ T \ y \ T}" by (meson connected_component_of_def) lemma connected_component_of_subset_topspace: "connected_component_of_set X x \ topspace X" using connected_component_in_topspace by force lemma connected_component_of_eq_empty: "connected_component_of_set X x = {} \ (x \ topspace X)" using connected_component_in_topspace connected_component_of_refl by fastforce lemma connected_space_iff_connected_component: "connected_space X \ (\x \ topspace X. \y \ topspace X. connected_component_of X x y)" by (simp add: connected_component_of_def connected_space_subconnected) lemma connected_space_imp_connected_component_of: "\connected_space X; a \ topspace X; b \ topspace X\ \ connected_component_of X a b" by (simp add: connected_space_iff_connected_component) lemma connected_space_connected_component_set: "connected_space X \ (\x \ topspace X. connected_component_of_set X x = topspace X)" using connected_component_of_subset_topspace connected_space_iff_connected_component by fastforce lemma connected_component_of_maximal: "\connectedin X S; x \ S\ \ S \ connected_component_of_set X x" by (meson Ball_Collect connected_component_of_def) lemma connected_component_of_equiv: "connected_component_of X x y \ x \ topspace X \ y \ topspace X \ connected_component_of X x = connected_component_of X y" apply (simp add: connected_component_in_topspace fun_eq_iff) by (meson connected_component_of_refl connected_component_of_sym connected_component_of_trans) lemma connected_component_of_disjoint: "disjnt (connected_component_of_set X x) (connected_component_of_set X y) \ ~(connected_component_of X x y)" using connected_component_of_equiv unfolding disjnt_iff by force lemma connected_component_of_eq: "connected_component_of X x = connected_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ x \ topspace X \ y \ topspace X \ connected_component_of X x y" by (metis Collect_empty_eq_bot connected_component_of_eq_empty connected_component_of_equiv) lemma connectedin_connected_component_of: "connectedin X (connected_component_of_set X x)" proof - have "connected_component_of_set X x = \ {T. connectedin X T \ x \ T}" by (auto simp: connected_component_of_def) then show ?thesis by (metis (no_types, lifting) InterI connectedin_Union emptyE mem_Collect_eq) qed lemma Union_connected_components_of: "\(connected_components_of X) = topspace X" unfolding connected_components_of_def using connected_component_in_topspace connected_component_of_refl by fastforce lemma connected_components_of_maximal: "\C \ connected_components_of X; connectedin X S; ~disjnt C S\ \ S \ C" unfolding connected_components_of_def disjnt_def apply clarify by (metis Int_emptyI connected_component_of_def connected_component_of_trans mem_Collect_eq) lemma pairwise_disjoint_connected_components_of: "pairwise disjnt (connected_components_of X)" unfolding connected_components_of_def pairwise_def by (smt (verit, best) connected_component_of_disjoint connected_component_of_eq imageE) lemma complement_connected_components_of_Union: "C \ connected_components_of X \ topspace X - C = \ (connected_components_of X - {C})" by (metis Union_connected_components_of bot.extremum ccpo_Sup_singleton diff_Union_pairwise_disjoint insert_subset pairwise_disjoint_connected_components_of) lemma nonempty_connected_components_of: "C \ connected_components_of X \ C \ {}" unfolding connected_components_of_def by (metis (no_types, lifting) connected_component_of_eq_empty imageE) lemma connected_components_of_subset: "C \ connected_components_of X \ C \ topspace X" using Union_connected_components_of by fastforce lemma connectedin_connected_components_of: assumes "C \ connected_components_of X" shows "connectedin X C" proof - have "C \ connected_component_of_set X ` topspace X" using assms connected_components_of_def by blast then show ?thesis using connectedin_connected_component_of by fastforce qed lemma connected_component_in_connected_components_of: "connected_component_of_set X a \ connected_components_of X \ a \ topspace X" by (metis (no_types, lifting) connected_component_of_eq_empty connected_components_of_def image_iff) lemma connected_space_iff_components_eq: "connected_space X \ (\C \ connected_components_of X. \C' \ connected_components_of X. C = C')" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: connected_components_of_def connected_space_connected_component_set) show "?rhs \ ?lhs" by (metis Union_connected_components_of Union_iff connected_space_subconnected connectedin_connected_components_of) qed lemma connected_components_of_eq_empty: "connected_components_of X = {} \ topspace X = {}" by (simp add: connected_components_of_def) lemma connected_components_of_empty_space: "topspace X = {} \ connected_components_of X = {}" by (simp add: connected_components_of_eq_empty) lemma connected_components_of_subset_sing: "connected_components_of X \ {S} \ connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: connected_components_of_empty_space connected_space_topspace_empty) next case False then have "\connected_components_of X \ {S}\ \ topspace X = S" by (metis Sup_empty Union_connected_components_of ccpo_Sup_singleton subset_singleton_iff) with False show ?thesis unfolding connected_components_of_def by (metis connected_space_connected_component_set empty_iff image_subset_iff insert_iff) qed lemma connected_space_iff_components_subset_singleton: "connected_space X \ (\a. connected_components_of X \ {a})" by (simp add: connected_components_of_subset_sing) lemma connected_components_of_eq_singleton: "connected_components_of X = {S} \ connected_space X \ topspace X \ {} \ S = topspace X" by (metis ccpo_Sup_singleton connected_components_of_subset_sing insert_not_empty subset_singleton_iff) lemma connected_components_of_connected_space: "connected_space X \ connected_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: connected_components_of_eq_empty connected_components_of_eq_singleton) lemma exists_connected_component_of_superset: assumes "connectedin X S" and ne: "topspace X \ {}" shows "\C. C \ connected_components_of X \ S \ C" proof (cases "S = {}") case True then show ?thesis using ne connected_components_of_def by blast next case False then show ?thesis by (meson all_not_in_conv assms(1) connected_component_in_connected_components_of connected_component_of_maximal connectedin_subset_topspace in_mono) qed lemma closedin_connected_components_of: assumes "C \ connected_components_of X" shows "closedin X C" proof - obtain x where "x \ topspace X" and x: "C = connected_component_of_set X x" using assms by (auto simp: connected_components_of_def) have "connected_component_of_set X x \ topspace X" by (simp add: connected_component_of_subset_topspace) moreover have "X closure_of connected_component_of_set X x \ connected_component_of_set X x" proof (rule connected_component_of_maximal) show "connectedin X (X closure_of connected_component_of_set X x)" by (simp add: connectedin_closure_of connectedin_connected_component_of) show "x \ X closure_of connected_component_of_set X x" by (simp add: \x \ topspace X\ closure_of connected_component_of_refl) qed ultimately show ?thesis using closure_of_subset_eq x by auto qed lemma closedin_connected_component_of: "closedin X (connected_component_of_set X x)" by (metis closedin_connected_components_of closedin_empty connected_component_in_connected_components_of connected_component_of_eq_empty) lemma connected_component_of_eq_overlap: "connected_component_of_set X x = connected_component_of_set X y \ (x \ topspace X) \ (y \ topspace X) \ ~(connected_component_of_set X x \ connected_component_of_set X y = {})" using connected_component_of_equiv by fastforce lemma connected_component_of_nonoverlap: "connected_component_of_set X x \ connected_component_of_set X y = {} \ (x \ topspace X) \ (y \ topspace X) \ ~(connected_component_of_set X x = connected_component_of_set X y)" by (metis connected_component_of_eq_empty connected_component_of_eq_overlap inf.idem) lemma connected_component_of_overlap: "~(connected_component_of_set X x \ connected_component_of_set X y = {}) \ x \ topspace X \ y \ topspace X \ connected_component_of_set X x = connected_component_of_set X y" by (meson connected_component_of_nonoverlap) subsection\Combining theorems for continuous functions into the reals\ text \The homeomorphism between the real line and the open interval $(-1,1)$\ lemma continuous_map_real_shrink: "continuous_map euclideanreal (top_of_set {-1<..<1}) (\x. x / (1 + \x\))" proof - have "continuous_on UNIV (\x::real. x / (1 + \x\))" by (intro continuous_intros) auto then show ?thesis by (auto simp: continuous_map_in_subtopology divide_simps) qed lemma continuous_on_real_grow: "continuous_on {-1<..<1} (\x::real. x / (1 - \x\))" by (intro continuous_intros) auto lemma real_grow_shrink: fixes x::real shows "x / (1 + \x\) / (1 - \x / (1 + \x\)\) = x" by (force simp: divide_simps) lemma homeomorphic_maps_real_shrink: "homeomorphic_maps euclideanreal (subtopology euclideanreal {-1<..<1}) (\x. x / (1 + \x\)) (\y. y / (1 - \y\))" by (force simp: homeomorphic_maps_def continuous_map_real_shrink continuous_on_real_grow divide_simps) +lemma real_shrink_Galois: + fixes x::real + shows "(x / (1 + \x\) = y) \ (\y\ < 1 \ y / (1 - \y\) = x)" + using real_grow_shrink by (fastforce simp add: distrib_left) + +lemma real_shrink_lt: + fixes x::real + shows "x / (1 + \x\) < y / (1 + \y\) \ x < y" + using zero_less_mult_iff [of x y] by (auto simp: field_simps abs_if not_less) + +lemma real_shrink_le: + fixes x::real + shows "x / (1 + \x\) \ y / (1 + \y\) \ x \ y" + by (meson linorder_not_le real_shrink_lt) + +lemma real_shrink_grow: + fixes x::real + shows "\x\ < 1 \ x / (1 - \x\) / (1 + \x / (1 - \x\)\) = x" + using real_shrink_Galois by blast + +lemma continuous_shrink: + "continuous_on UNIV (\x::real. x / (1 + \x\))" + by (intro continuous_intros) auto + +lemma strict_mono_shrink: + "strict_mono (\x::real. x / (1 + \x\))" + by (simp add: monotoneI real_shrink_lt) + +lemma shrink_range: "(\x::real. x / (1 + \x\)) ` S \ {-1<..<1}" + by (auto simp: divide_simps) + +text \Note: connected sets of real numbers are the same thing as intervals\ +lemma connected_shrink: + fixes S :: "real set" + shows "connected ((\x. x / (1 + \x\)) ` S) \ connected S" (is "?lhs = ?rhs") +proof + assume "?lhs" + then have "connected ((\x. x / (1 - \x\)) ` (\x. x / (1 + \x\)) ` S)" + by (metis continuous_on_real_grow shrink_range connected_continuous_image + continuous_on_subset) + then show "?rhs" + using real_grow_shrink by (force simp add: image_comp) +next + assume ?rhs + then show ?lhs + using connected_continuous_image + by (metis continuous_on_subset continuous_shrink subset_UNIV) +qed + end \ No newline at end of file diff --git a/src/HOL/Analysis/Elementary_Metric_Spaces.thy b/src/HOL/Analysis/Elementary_Metric_Spaces.thy --- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy +++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy @@ -1,3209 +1,3212 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) section \Elementary Metric Spaces\ theory Elementary_Metric_Spaces imports Abstract_Topology_2 Metric_Arith begin subsection \Open and closed balls\ definition\<^marker>\tag important\ ball :: "'a::metric_space \ real \ 'a set" where "ball x e = {y. dist x y < e}" definition\<^marker>\tag important\ cball :: "'a::metric_space \ real \ 'a set" where "cball x e = {y. dist x y \ e}" definition\<^marker>\tag important\ sphere :: "'a::metric_space \ real \ 'a set" where "sphere x e = {y. dist x y = e}" lemma mem_ball [simp, metric_unfold]: "y \ ball x e \ dist x y < e" by (simp add: ball_def) lemma mem_cball [simp, metric_unfold]: "y \ cball x e \ dist x y \ e" by (simp add: cball_def) lemma mem_sphere [simp]: "y \ sphere x e \ dist x y = e" by (simp add: sphere_def) lemma ball_trivial [simp]: "ball x 0 = {}" by auto lemma cball_trivial [simp]: "cball x 0 = {x}" by auto lemma sphere_trivial [simp]: "sphere x 0 = {x}" by auto lemma disjoint_ballI: "dist x y \ r+s \ ball x r \ ball y s = {}" using dist_triangle_less_add not_le by fastforce lemma disjoint_cballI: "dist x y > r + s \ cball x r \ cball y s = {}" by (metis add_mono disjoint_iff_not_equal dist_triangle2 dual_order.trans leD mem_cball) lemma sphere_empty [simp]: "r < 0 \ sphere a r = {}" for a :: "'a::metric_space" by auto lemma centre_in_ball [simp]: "x \ ball x e \ 0 < e" by simp lemma centre_in_cball [simp]: "x \ cball x e \ 0 \ e" by simp lemma ball_subset_cball [simp, intro]: "ball x e \ cball x e" by (simp add: subset_eq) lemma mem_ball_imp_mem_cball: "x \ ball y e \ x \ cball y e" by auto lemma sphere_cball [simp,intro]: "sphere z r \ cball z r" by force lemma cball_diff_sphere: "cball a r - sphere a r = ball a r" by auto lemma subset_ball[intro]: "d \ e \ ball x d \ ball x e" by auto lemma subset_cball[intro]: "d \ e \ cball x d \ cball x e" by auto lemma mem_ball_leI: "x \ ball y e \ e \ f \ x \ ball y f" by auto lemma mem_cball_leI: "x \ cball y e \ e \ f \ x \ cball y f" by auto lemma cball_trans: "y \ cball z b \ x \ cball y a \ x \ cball z (b + a)" by metric lemma ball_max_Un: "ball a (max r s) = ball a r \ ball a s" by auto lemma ball_min_Int: "ball a (min r s) = ball a r \ ball a s" by auto lemma cball_max_Un: "cball a (max r s) = cball a r \ cball a s" by auto lemma cball_min_Int: "cball a (min r s) = cball a r \ cball a s" by auto lemma cball_diff_eq_sphere: "cball a r - ball a r = sphere a r" by auto lemma open_ball [intro, simp]: "open (ball x e)" proof - have "open (dist x -` {.. (\x\S. \e>0. ball x e \ S)" by (simp add: open_dist subset_eq Ball_def dist_commute) lemma openI [intro?]: "(\x. x\S \ \e>0. ball x e \ S) \ open S" by (auto simp: open_contains_ball) lemma openE[elim?]: assumes "open S" "x\S" obtains e where "e>0" "ball x e \ S" using assms unfolding open_contains_ball by auto lemma open_contains_ball_eq: "open S \ x\S \ (\e>0. ball x e \ S)" by (metis open_contains_ball subset_eq centre_in_ball) lemma ball_eq_empty[simp]: "ball x e = {} \ e \ 0" unfolding mem_ball set_eq_iff by (simp add: not_less) metric lemma ball_empty: "e \ 0 \ ball x e = {}" by simp lemma closed_cball [iff]: "closed (cball x e)" proof - have "closed (dist x -` {..e})" by (intro closed_vimage closed_atMost continuous_intros) also have "dist x -` {..e} = cball x e" by auto finally show ?thesis . qed lemma open_contains_cball: "open S \ (\x\S. \e>0. cball x e \ S)" proof - { fix x and e::real assume "x\S" "e>0" "ball x e \ S" then have "\d>0. cball x d \ S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto) } moreover { fix x and e::real assume "x\S" "e>0" "cball x e \ S" then have "\d>0. ball x d \ S" using mem_ball_imp_mem_cball by blast } ultimately show ?thesis unfolding open_contains_ball by auto qed lemma open_contains_cball_eq: "open S \ (\x. x \ S \ (\e>0. cball x e \ S))" by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball) lemma eventually_nhds_ball: "d > 0 \ eventually (\x. x \ ball z d) (nhds z)" by (rule eventually_nhds_in_open) simp_all lemma eventually_at_ball: "d > 0 \ eventually (\t. t \ ball z d \ t \ A) (at z within A)" unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute) lemma eventually_at_ball': "d > 0 \ eventually (\t. t \ ball z d \ t \ z \ t \ A) (at z within A)" unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute) lemma at_within_ball: "e > 0 \ dist x y < e \ at y within ball x e = at y" by (subst at_within_open) auto lemma atLeastAtMost_eq_cball: fixes a b::real shows "{a .. b} = cball ((a + b)/2) ((b - a)/2)" by (auto simp: dist_real_def field_simps) lemma cball_eq_atLeastAtMost: fixes a b::real shows "cball a b = {a - b .. a + b}" by (auto simp: dist_real_def) lemma greaterThanLessThan_eq_ball: fixes a b::real shows "{a <..< b} = ball ((a + b)/2) ((b - a)/2)" by (auto simp: dist_real_def field_simps) lemma ball_eq_greaterThanLessThan: fixes a b::real shows "ball a b = {a - b <..< a + b}" by (auto simp: dist_real_def) lemma interior_ball [simp]: "interior (ball x e) = ball x e" by (simp add: interior_open) lemma cball_eq_empty [simp]: "cball x e = {} \ e < 0" by (smt (verit, best) Diff_empty ball_eq_empty cball_diff_sphere centre_in_ball centre_in_cball sphere_empty) lemma cball_empty [simp]: "e < 0 \ cball x e = {}" by simp lemma cball_sing: fixes x :: "'a::metric_space" shows "e = 0 \ cball x e = {x}" by simp lemma ball_divide_subset: "d \ 1 \ ball x (e/d) \ ball x e" by (metis ball_eq_empty div_by_1 frac_le linear subset_ball zero_less_one) lemma ball_divide_subset_numeral: "ball x (e / numeral w) \ ball x e" using ball_divide_subset one_le_numeral by blast lemma cball_divide_subset: "d \ 1 \ cball x (e/d) \ cball x e" by (smt (verit, best) cball_empty div_by_1 frac_le subset_cball zero_le_divide_iff) lemma cball_divide_subset_numeral: "cball x (e / numeral w) \ cball x e" using cball_divide_subset one_le_numeral by blast lemma cball_scale: assumes "a \ 0" shows "(\x. a *\<^sub>R x) ` cball c r = cball (a *\<^sub>R c :: 'a :: real_normed_vector) (\a\ * r)" proof - have 1: "(\x. a *\<^sub>R x) ` cball c r \ cball (a *\<^sub>R c) (\a\ * r)" if "a \ 0" for a r and c :: 'a proof safe fix x assume x: "x \ cball c r" have "dist (a *\<^sub>R c) (a *\<^sub>R x) = norm (a *\<^sub>R c - a *\<^sub>R x)" by (auto simp: dist_norm) also have "a *\<^sub>R c - a *\<^sub>R x = a *\<^sub>R (c - x)" by (simp add: algebra_simps) finally show "a *\<^sub>R x \ cball (a *\<^sub>R c) (\a\ * r)" using that x by (auto simp: dist_norm) qed have "cball (a *\<^sub>R c) (\a\ * r) = (\x. a *\<^sub>R x) ` (\x. inverse a *\<^sub>R x) ` cball (a *\<^sub>R c) (\a\ * r)" unfolding image_image using assms by simp also have "\ \ (\x. a *\<^sub>R x) ` cball (inverse a *\<^sub>R (a *\<^sub>R c)) (\inverse a\ * (\a\ * r))" using assms by (intro image_mono 1) auto also have "\ = (\x. a *\<^sub>R x) ` cball c r" using assms by (simp add: algebra_simps) finally have "cball (a *\<^sub>R c) (\a\ * r) \ (\x. a *\<^sub>R x) ` cball c r" . moreover from assms have "(\x. a *\<^sub>R x) ` cball c r \ cball (a *\<^sub>R c) (\a\ * r)" by (intro 1) auto ultimately show ?thesis by blast qed lemma ball_scale: assumes "a \ 0" shows "(\x. a *\<^sub>R x) ` ball c r = ball (a *\<^sub>R c :: 'a :: real_normed_vector) (\a\ * r)" proof - have 1: "(\x. a *\<^sub>R x) ` ball c r \ ball (a *\<^sub>R c) (\a\ * r)" if "a \ 0" for a r and c :: 'a proof safe fix x assume x: "x \ ball c r" have "dist (a *\<^sub>R c) (a *\<^sub>R x) = norm (a *\<^sub>R c - a *\<^sub>R x)" by (auto simp: dist_norm) also have "a *\<^sub>R c - a *\<^sub>R x = a *\<^sub>R (c - x)" by (simp add: algebra_simps) finally show "a *\<^sub>R x \ ball (a *\<^sub>R c) (\a\ * r)" using that x by (auto simp: dist_norm) qed have "ball (a *\<^sub>R c) (\a\ * r) = (\x. a *\<^sub>R x) ` (\x. inverse a *\<^sub>R x) ` ball (a *\<^sub>R c) (\a\ * r)" unfolding image_image using assms by simp also have "\ \ (\x. a *\<^sub>R x) ` ball (inverse a *\<^sub>R (a *\<^sub>R c)) (\inverse a\ * (\a\ * r))" using assms by (intro image_mono 1) auto also have "\ = (\x. a *\<^sub>R x) ` ball c r" using assms by (simp add: algebra_simps) finally have "ball (a *\<^sub>R c) (\a\ * r) \ (\x. a *\<^sub>R x) ` ball c r" . moreover from assms have "(\x. a *\<^sub>R x) ` ball c r \ ball (a *\<^sub>R c) (\a\ * r)" by (intro 1) auto ultimately show ?thesis by blast qed lemma frequently_atE: fixes x :: "'a :: metric_space" assumes "frequently P (at x within s)" shows "\f. filterlim f (at x within s) sequentially \ (\n. P (f n))" proof - have "\y. y \ s \ (ball x (1 / real (Suc n)) - {x}) \ P y" for n proof - have "\z\s. z \ x \ dist z x < (1 / real (Suc n)) \ P z" by (metis assms divide_pos_pos frequently_at of_nat_0_less_iff zero_less_Suc zero_less_one) then show ?thesis by (auto simp: dist_commute conj_commute) qed then obtain f where f: "\n. f n \ s \ (ball x (1 / real (Suc n)) - {x}) \ P (f n)" by metis have "filterlim f (nhds x) sequentially" unfolding tendsto_iff proof clarify fix e :: real assume e: "e > 0" then obtain n where n: "Suc n > 1 / e" by (meson le_nat_floor lessI not_le) have "dist (f k) x < e" if "k \ n" for k proof - have "dist (f k) x < 1 / real (Suc k)" using f[of k] by (auto simp: dist_commute) also have "\ \ 1 / real (Suc n)" using that by (intro divide_left_mono) auto also have "\ < e" using n e by (simp add: field_simps) finally show ?thesis . qed thus "\\<^sub>F k in sequentially. dist (f k) x < e" unfolding eventually_at_top_linorder by blast qed moreover have "f n \ x" for n using f[of n] by auto ultimately have "filterlim f (at x within s) sequentially" using f by (auto simp: filterlim_at) with f show ?thesis by blast qed subsection \Limit Points\ lemma islimpt_approachable: fixes x :: "'a::metric_space" shows "x islimpt S \ (\e>0. \x'\S. x' \ x \ dist x' x < e)" unfolding islimpt_iff_eventually eventually_at by fast lemma islimpt_approachable_le: "x islimpt S \ (\e>0. \x'\ S. x' \ x \ dist x' x \ e)" for x :: "'a::metric_space" unfolding islimpt_approachable using approachable_lt_le2 [where f="\y. dist y x" and P="\y. y \ S \ y = x" and Q="\x. True"] by auto lemma limpt_of_limpts: "x islimpt {y. y islimpt S} \ x islimpt S" for x :: "'a::metric_space" by (metis islimpt_def islimpt_eq_acc_point mem_Collect_eq) lemma closed_limpts: "closed {x::'a::metric_space. x islimpt S}" using closed_limpt limpt_of_limpts by blast lemma limpt_of_closure: "x islimpt closure S \ x islimpt S" for x :: "'a::metric_space" by (auto simp: closure_def islimpt_Un dest: limpt_of_limpts) lemma islimpt_eq_infinite_ball: "x islimpt S \ (\e>0. infinite(S \ ball x e))" unfolding islimpt_eq_acc_point by (metis open_ball Int_commute Int_mono finite_subset open_contains_ball_eq subset_eq) lemma islimpt_eq_infinite_cball: "x islimpt S \ (\e>0. infinite(S \ cball x e))" unfolding islimpt_eq_infinite_ball by (metis open_ball ball_subset_cball centre_in_ball finite_Int inf.absorb_iff2 inf_assoc open_contains_cball_eq) subsection \Perfect Metric Spaces\ lemma perfect_choose_dist: "0 < r \ \a. a \ x \ dist a x < r" for x :: "'a::{perfect_space,metric_space}" using islimpt_UNIV [of x] by (simp add: islimpt_approachable) lemma cball_eq_sing: fixes x :: "'a::{metric_space,perfect_space}" shows "cball x e = {x} \ e = 0" by (smt (verit, best) open_ball ball_eq_empty ball_subset_cball cball_empty cball_trivial not_open_singleton subset_singleton_iff) subsection \Finite and discrete\ lemma finite_ball_include: fixes a :: "'a::metric_space" assumes "finite S" shows "\e>0. S \ ball a e" using assms proof induction case (insert x S) then obtain e0 where "e0>0" and e0:"S \ ball a e0" by auto define e where "e = max e0 (2 * dist a x)" have "e>0" unfolding e_def using \e0>0\ by auto moreover have "insert x S \ ball a e" using e0 \e>0\ unfolding e_def by auto ultimately show ?case by auto qed (auto intro: zero_less_one) lemma finite_set_avoid: fixes a :: "'a::metric_space" assumes "finite S" shows "\d>0. \x\S. x \ a \ d \ dist a x" using assms proof induction case (insert x S) then obtain d where "d > 0" and d: "\x\S. x \ a \ d \ dist a x" by blast show ?case by (smt (verit, del_insts) dist_pos_lt insert.IH insert_iff) qed (auto intro: zero_less_one) lemma discrete_imp_closed: fixes S :: "'a::metric_space set" assumes e: "0 < e" and d: "\x \ S. \y \ S. dist y x < e \ y = x" shows "closed S" proof - have False if C: "\e. e>0 \ \x'\S. x' \ x \ dist x' x < e" for x proof - from e have e2: "e/2 > 0" by arith from C[OF e2] obtain y where y: "y \ S" "y \ x" "dist y x < e/2" by blast from e2 y(2) have mp: "min (e/2) (dist x y) > 0" by simp from d y C[OF mp] show ?thesis by metric qed then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a]) qed lemma discrete_imp_not_islimpt: assumes e: "0 < e" and d: "\x y. x \ S \ y \ S \ dist y x < e \ y = x" shows "\ x islimpt S" proof assume "x islimpt S" hence "x islimpt S - {x}" by (meson islimpt_punctured) moreover from assms have "closed (S - {x})" by (intro discrete_imp_closed) auto ultimately show False unfolding closed_limpt by blast qed subsection \Interior\ lemma mem_interior: "x \ interior S \ (\e>0. ball x e \ S)" using open_contains_ball_eq [where S="interior S"] by (simp add: open_subset_interior) lemma mem_interior_cball: "x \ interior S \ (\e>0. cball x e \ S)" by (meson ball_subset_cball interior_subset mem_interior open_contains_cball open_interior subset_trans) +lemma ball_iff_cball: "(\r>0. ball x r \ U) \ (\r>0. cball x r \ U)" + by (meson mem_interior mem_interior_cball) + subsection \Frontier\ lemma frontier_straddle: fixes a :: "'a::metric_space" shows "a \ frontier S \ (\e>0. (\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e))" unfolding frontier_def closure_interior by (auto simp: mem_interior subset_eq ball_def) subsection \Limits\ proposition Lim: "(f \ l) net \ trivial_limit net \ (\e>0. eventually (\x. dist (f x) l < e) net)" by (auto simp: tendsto_iff trivial_limit_eq) text \Show that they yield usual definitions in the various cases.\ proposition Lim_within_le: "(f \ l)(at a within S) \ (\e>0. \d>0. \x\S. 0 < dist x a \ dist x a \ d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at_le) proposition Lim_within: "(f \ l) (at a within S) \ (\e >0. \d>0. \x \ S. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at) corollary Lim_withinI [intro?]: assumes "\e. e > 0 \ \d>0. \x \ S. 0 < dist x a \ dist x a < d \ dist (f x) l \ e" shows "(f \ l) (at a within S)" unfolding Lim_within by (smt (verit, ccfv_SIG) assms zero_less_dist_iff) proposition Lim_at: "(f \ l) (at a) \ (\e >0. \d>0. \x. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at) lemma Lim_transform_within_set: fixes a :: "'a::metric_space" and l :: "'b::metric_space" shows "\(f \ l) (at a within S); eventually (\x. x \ S \ x \ T) (at a)\ \ (f \ l) (at a within T)" by (simp add: eventually_at Lim_within) (smt (verit, best)) text \Another limit point characterization.\ lemma limpt_sequential_inj: fixes x :: "'a::metric_space" shows "x islimpt S \ (\f. (\n::nat. f n \ S - {x}) \ inj f \ (f \ x) sequentially)" (is "?lhs = ?rhs") proof assume ?lhs then have "\e>0. \x'\S. x' \ x \ dist x' x < e" by (force simp: islimpt_approachable) then obtain y where y: "\e. e>0 \ y e \ S \ y e \ x \ dist (y e) x < e" by metis define f where "f \ rec_nat (y 1) (\n fn. y (min (inverse(2 ^ (Suc n))) (dist fn x)))" have [simp]: "f 0 = y 1" and fSuc: "f(Suc n) = y (min (inverse(2 ^ (Suc n))) (dist (f n) x))" for n by (simp_all add: f_def) have f: "f n \ S \ (f n \ x) \ dist (f n) x < inverse(2 ^ n)" for n proof (induction n) case 0 show ?case by (simp add: y) next case (Suc n) then show ?case by (smt (verit, best) fSuc dist_pos_lt inverse_positive_iff_positive y zero_less_power) qed show ?rhs proof (intro exI conjI allI) show "\n. f n \ S - {x}" using f by blast have "dist (f n) x < dist (f m) x" if "m < n" for m n using that proof (induction n) case 0 then show ?case by simp next case (Suc n) then consider "m < n" | "m = n" using less_Suc_eq by blast then show ?case proof cases assume "m < n" have "dist (f(Suc n)) x = dist (y (min (inverse(2 ^ (Suc n))) (dist (f n) x))) x" by (simp add: fSuc) also have "\ < dist (f n) x" by (metis dist_pos_lt f min.strict_order_iff min_less_iff_conj y) also have "\ < dist (f m) x" using Suc.IH \m < n\ by blast finally show ?thesis . next assume "m = n" then show ?case by (smt (verit, best) dist_pos_lt f fSuc y) qed qed then show "inj f" by (metis less_irrefl linorder_injI) have "\e n. \0 < e; nat \1 / e\ \ n\ \ dist (f n) x < e" apply (rule less_trans [OF f [THEN conjunct2, THEN conjunct2]]) by (simp add: divide_simps order_le_less_trans) then show "f \ x" using lim_sequentially by blast qed next assume ?rhs then show ?lhs by (fastforce simp add: islimpt_approachable lim_sequentially) qed lemma Lim_dist_ubound: assumes "\(trivial_limit net)" and "(f \ l) net" and "eventually (\x. dist a (f x) \ e) net" shows "dist a l \ e" using assms by (fast intro: tendsto_le tendsto_intros) subsection \Continuity\ text\Derive the epsilon-delta forms, which we often use as "definitions"\ proposition continuous_within_eps_delta: "continuous (at x within s) f \ (\e>0. \d>0. \x'\ s. dist x' x < d --> dist (f x') (f x) < e)" unfolding continuous_within and Lim_within by fastforce corollary continuous_at_eps_delta: "continuous (at x) f \ (\e > 0. \d > 0. \x'. dist x' x < d \ dist (f x') (f x) < e)" using continuous_within_eps_delta [of x UNIV f] by simp lemma continuous_at_right_real_increasing: fixes f :: "real \ real" assumes nondecF: "\x y. x \ y \ f x \ f y" shows "continuous (at_right a) f \ (\e>0. \d>0. f (a + d) - f a < e)" apply (simp add: greaterThan_def dist_real_def continuous_within Lim_within_le) apply (intro all_cong ex_cong) by (smt (verit, best) nondecF) lemma continuous_at_left_real_increasing: assumes nondecF: "\ x y. x \ y \ f x \ ((f y) :: real)" shows "(continuous (at_left (a :: real)) f) \ (\e > 0. \delta > 0. f a - f (a - delta) < e)" apply (simp add: lessThan_def dist_real_def continuous_within Lim_within_le) apply (intro all_cong ex_cong) by (smt (verit) nondecF) text\Versions in terms of open balls.\ lemma continuous_within_ball: "continuous (at x within S) f \ (\e > 0. \d > 0. f ` (ball x d \ S) \ ball (f x) e)" (is "?lhs = ?rhs") proof assume ?lhs { fix e :: real assume "e > 0" then obtain d where "d>0" and d: "\y\S. 0 < dist y x \ dist y x < d \ dist (f y) (f x) < e" using \?lhs\[unfolded continuous_within Lim_within] by auto { fix y assume "y \ f ` (ball x d \ S)" then have "y \ ball (f x) e" using d \e > 0\ by (auto simp: dist_commute) } then have "\d>0. f ` (ball x d \ S) \ ball (f x) e" using \d > 0\ by blast } then show ?rhs by auto next assume ?rhs then show ?lhs apply (simp add: continuous_within Lim_within ball_def subset_eq) by (metis (mono_tags, lifting) Int_iff dist_commute mem_Collect_eq) qed lemma continuous_at_ball: "continuous (at x) f \ (\e>0. \d>0. f ` (ball x d) \ ball (f x) e)" apply (simp add: continuous_at Lim_at subset_eq Ball_def Bex_def image_iff) by (smt (verit, ccfv_threshold) dist_commute dist_self) text\Define setwise continuity in terms of limits within the set.\ lemma continuous_on_iff: "continuous_on s f \ (\x\s. \e>0. \d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e)" unfolding continuous_on_def Lim_within by (metis dist_pos_lt dist_self) lemma continuous_within_E: assumes "continuous (at x within S) f" "e>0" obtains d where "d>0" "\x'. \x'\ S; dist x' x \ d\ \ dist (f x') (f x) < e" using assms unfolding continuous_within_eps_delta by (metis dense order_le_less_trans) lemma continuous_onI [intro?]: assumes "\x e. \e > 0; x \ S\ \ \d>0. \x'\S. dist x' x < d \ dist (f x') (f x) \ e" shows "continuous_on S f" apply (simp add: continuous_on_iff, clarify) apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) done text\Some simple consequential lemmas.\ lemma continuous_onE: assumes "continuous_on s f" "x\s" "e>0" obtains d where "d>0" "\x'. \x' \ s; dist x' x \ d\ \ dist (f x') (f x) < e" using assms unfolding continuous_on_iff by (metis dense order_le_less_trans) text\The usual transformation theorems.\ lemma continuous_transform_within: fixes f g :: "'a::metric_space \ 'b::topological_space" assumes "continuous (at x within s) f" and "0 < d" and "x \ s" and "\x'. \x' \ s; dist x' x < d\ \ f x' = g x'" shows "continuous (at x within s) g" using assms unfolding continuous_within by (force intro: Lim_transform_within) subsection \Closure and Limit Characterization\ lemma closure_approachable: fixes S :: "'a::metric_space set" shows "x \ closure S \ (\e>0. \y\S. dist y x < e)" using dist_self by (force simp: closure_def islimpt_approachable) lemma closure_approachable_le: fixes S :: "'a::metric_space set" shows "x \ closure S \ (\e>0. \y\S. dist y x \ e)" unfolding closure_approachable using dense by force lemma closure_approachableD: assumes "x \ closure S" "e>0" shows "\y\S. dist x y < e" using assms unfolding closure_approachable by (auto simp: dist_commute) lemma closed_approachable: fixes S :: "'a::metric_space set" shows "closed S \ (\e>0. \y\S. dist y x < e) \ x \ S" by (metis closure_closed closure_approachable) lemma closure_contains_Inf: fixes S :: "real set" assumes "S \ {}" "bdd_below S" shows "Inf S \ closure S" proof - have *: "\x\S. Inf S \ x" using cInf_lower[of _ S] assms by metis { fix e :: real assume "e > 0" then have "Inf S < Inf S + e" by simp with assms obtain x where "x \ S" "x < Inf S + e" using cInf_lessD by blast with * have "\x\S. dist x (Inf S) < e" using dist_real_def by force } then show ?thesis unfolding closure_approachable by auto qed lemma closure_contains_Sup: fixes S :: "real set" assumes "S \ {}" "bdd_above S" shows "Sup S \ closure S" proof - have *: "\x\S. x \ Sup S" using cSup_upper[of _ S] assms by metis { fix e :: real assume "e > 0" then have "Sup S - e < Sup S" by simp with assms obtain x where "x \ S" "Sup S - e < x" using less_cSupE by blast with * have "\x\S. dist x (Sup S) < e" using dist_real_def by force } then show ?thesis unfolding closure_approachable by auto qed lemma not_trivial_limit_within_ball: "\ trivial_limit (at x within S) \ (\e>0. S \ ball x e - {x} \ {})" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs proof - { fix e :: real assume "e > 0" then obtain y where "y \ S - {x}" and "dist y x < e" using \?lhs\ not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto then have "y \ S \ ball x e - {x}" unfolding ball_def by (simp add: dist_commute) then have "S \ ball x e - {x} \ {}" by blast } then show ?thesis by auto qed show ?lhs if ?rhs proof - { fix e :: real assume "e > 0" then obtain y where "y \ S \ ball x e - {x}" using \?rhs\ by blast then have "y \ S - {x}" and "dist y x < e" unfolding ball_def by (simp_all add: dist_commute) then have "\y \ S - {x}. dist y x < e" by auto } then show ?thesis using not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto qed qed subsection \Boundedness\ (* FIXME: This has to be unified with BSEQ!! *) definition\<^marker>\tag important\ (in metric_space) bounded :: "'a set \ bool" where "bounded S \ (\x e. \y\S. dist x y \ e)" lemma bounded_subset_cball: "bounded S \ (\e x. S \ cball x e \ 0 \ e)" unfolding bounded_def subset_eq by auto (meson order_trans zero_le_dist) lemma bounded_any_center: "bounded S \ (\e. \y\S. dist a y \ e)" unfolding bounded_def by auto (metis add.commute add_le_cancel_right dist_commute dist_triangle_le) lemma bounded_iff: "bounded S \ (\a. \x\S. norm x \ a)" unfolding bounded_any_center [where a=0] by (simp add: dist_norm) lemma bdd_above_norm: "bdd_above (norm ` X) \ bounded X" by (simp add: bounded_iff bdd_above_def) lemma bounded_norm_comp: "bounded ((\x. norm (f x)) ` S) = bounded (f ` S)" by (simp add: bounded_iff) lemma boundedI: assumes "\x. x \ S \ norm x \ B" shows "bounded S" using assms bounded_iff by blast lemma bounded_empty [simp]: "bounded {}" by (simp add: bounded_def) lemma bounded_subset: "bounded T \ S \ T \ bounded S" by (metis bounded_def subset_eq) lemma bounded_interior[intro]: "bounded S \ bounded(interior S)" by (metis bounded_subset interior_subset) lemma bounded_closure[intro]: assumes "bounded S" shows "bounded (closure S)" proof - from assms obtain x and a where a: "\y\S. dist x y \ a" unfolding bounded_def by auto { fix y assume "y \ closure S" then obtain f where f: "\n. f n \ S" "(f \ y) sequentially" unfolding closure_sequential by auto have "\n. f n \ S \ dist x (f n) \ a" using a by simp then have "eventually (\n. dist x (f n) \ a) sequentially" by (simp add: f(1)) then have "dist x y \ a" using Lim_dist_ubound f(2) trivial_limit_sequentially by blast } then show ?thesis unfolding bounded_def by auto qed lemma bounded_closure_image: "bounded (f ` closure S) \ bounded (f ` S)" by (simp add: bounded_subset closure_subset image_mono) lemma bounded_cball[simp,intro]: "bounded (cball x e)" unfolding bounded_def using mem_cball by blast lemma bounded_ball[simp,intro]: "bounded (ball x e)" by (metis ball_subset_cball bounded_cball bounded_subset) lemma bounded_Un[simp]: "bounded (S \ T) \ bounded S \ bounded T" by (auto simp: bounded_def) (metis Un_iff bounded_any_center le_max_iff_disj) lemma bounded_Union[intro]: "finite F \ \S\F. bounded S \ bounded (\F)" by (induct rule: finite_induct[of F]) auto lemma bounded_UN [intro]: "finite A \ \x\A. bounded (B x) \ bounded (\x\A. B x)" by auto lemma bounded_insert [simp]: "bounded (insert x S) \ bounded S" proof - have "\y\{x}. dist x y \ 0" by simp then have "bounded {x}" unfolding bounded_def by fast then show ?thesis by (metis insert_is_Un bounded_Un) qed lemma bounded_subset_ballI: "S \ ball x r \ bounded S" by (meson bounded_ball bounded_subset) lemma bounded_subset_ballD: assumes "bounded S" shows "\r. 0 < r \ S \ ball x r" proof - obtain e::real and y where "S \ cball y e" "0 \ e" using assms by (auto simp: bounded_subset_cball) then show ?thesis by (intro exI[where x="dist x y + e + 1"]) metric qed lemma finite_imp_bounded [intro]: "finite S \ bounded S" by (induct set: finite) simp_all lemma bounded_Int[intro]: "bounded S \ bounded T \ bounded (S \ T)" by (metis Int_lower1 Int_lower2 bounded_subset) lemma bounded_diff[intro]: "bounded S \ bounded (S - T)" by (metis Diff_subset bounded_subset) lemma bounded_dist_comp: assumes "bounded (f ` S)" "bounded (g ` S)" shows "bounded ((\x. dist (f x) (g x)) ` S)" proof - from assms obtain M1 M2 where *: "dist (f x) undefined \ M1" "dist undefined (g x) \ M2" if "x \ S" for x by (auto simp: bounded_any_center[of _ undefined] dist_commute) have "dist (f x) (g x) \ M1 + M2" if "x \ S" for x using *[OF that] by metric then show ?thesis by (auto intro!: boundedI) qed lemma bounded_Times: assumes "bounded s" "bounded t" shows "bounded (s \ t)" proof - obtain x y a b where "\z\s. dist x z \ a" "\z\t. dist y z \ b" using assms [unfolded bounded_def] by auto then have "\z\s \ t. dist (x, y) z \ sqrt (a\<^sup>2 + b\<^sup>2)" by (auto simp: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono) then show ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto qed subsection \Compactness\ lemma compact_imp_bounded: assumes "compact U" shows "bounded U" proof - have "compact U" "\x\U. open (ball x 1)" "U \ (\x\U. ball x 1)" using assms by auto then obtain D where D: "D \ U" "finite D" "U \ (\x\D. ball x 1)" by (metis compactE_image) from \finite D\ have "bounded (\x\D. ball x 1)" by (simp add: bounded_UN) then show "bounded U" using \U \ (\x\D. ball x 1)\ by (rule bounded_subset) qed lemma continuous_on_compact_bound: assumes "compact A" "continuous_on A f" obtains B where "B \ 0" "\x. x \ A \ norm (f x) \ B" proof - have "compact (f ` A)" by (metis assms compact_continuous_image) then obtain B where "\x\A. norm (f x) \ B" by (auto dest!: compact_imp_bounded simp: bounded_iff) hence "max B 0 \ 0" and "\x\A. norm (f x) \ max B 0" by auto thus ?thesis using that by blast qed lemma closure_Int_ball_not_empty: assumes "S \ closure T" "x \ S" "r > 0" shows "T \ ball x r \ {}" using assms centre_in_ball closure_iff_nhds_not_empty by blast lemma compact_sup_maxdistance: fixes S :: "'a::metric_space set" assumes "compact S" and "S \ {}" shows "\x\S. \y\S. \u\S. \v\S. dist u v \ dist x y" proof - have "compact (S \ S)" using \compact S\ by (intro compact_Times) moreover have "S \ S \ {}" using \S \ {}\ by auto moreover have "continuous_on (S \ S) (\x. dist (fst x) (snd x))" by (intro continuous_at_imp_continuous_on ballI continuous_intros) ultimately show ?thesis using continuous_attains_sup[of "S \ S" "\x. dist (fst x) (snd x)"] by auto qed text \ If \A\ is a compact subset of an open set \B\ in a metric space, then there exists an \\ > 0\ such that the Minkowski sum of \A\ with an open ball of radius \\\ is also a subset of \B\. \ lemma compact_subset_open_imp_ball_epsilon_subset: assumes "compact A" "open B" "A \ B" obtains e where "e > 0" "(\x\A. ball x e) \ B" proof - have "\x\A. \e. e > 0 \ ball x e \ B" using assms unfolding open_contains_ball by blast then obtain e where e: "\x. x \ A \ e x > 0" "\x. x \ A \ ball x (e x) \ B" by metis define C where "C = e ` A" obtain X where X: "X \ A" "finite X" "A \ (\c\X. ball c (e c / 2))" using assms(1) proof (rule compactE_image) show "open (ball x (e x / 2))" if "x \ A" for x by simp show "A \ (\c\A. ball c (e c / 2))" using e by auto qed auto define e' where "e' = Min (insert 1 ((\x. e x / 2) ` X))" have "e' > 0" unfolding e'_def using e X by (subst Min_gr_iff) auto have e': "e' \ e x / 2" if "x \ X" for x using that X unfolding e'_def by (intro Min.coboundedI) auto show ?thesis proof show "e' > 0" by fact next show "(\x\A. ball x e') \ B" proof clarify fix x y assume xy: "x \ A" "y \ ball x e'" from xy(1) X obtain z where z: "z \ X" "x \ ball z (e z / 2)" by auto have "dist y z \ dist x y + dist z x" by (metis dist_commute dist_triangle) also have "dist z x < e z / 2" using xy z by auto also have "dist x y < e'" using xy by auto also have "\ \ e z / 2" using z by (intro e') auto finally have "y \ ball z (e z)" by (simp add: dist_commute) also have "\ \ B" using z X by (intro e) auto finally show "y \ B" . qed qed qed lemma compact_subset_open_imp_cball_epsilon_subset: assumes "compact A" "open B" "A \ B" obtains e where "e > 0" "(\x\A. cball x e) \ B" proof - obtain e where "e > 0" and e: "(\x\A. ball x e) \ B" using compact_subset_open_imp_ball_epsilon_subset [OF assms] by blast then have "(\x\A. cball x (e / 2)) \ (\x\A. ball x e)" by auto with \0 < e\ that show ?thesis by (metis e half_gt_zero_iff order_trans) qed subsubsection\Totally bounded\ lemma cauchy_def: "Cauchy S \ (\e>0. \N. \m n. m \ N \ n \ N \ dist (S m) (S n) < e)" unfolding Cauchy_def by metis proposition seq_compact_imp_totally_bounded: assumes "seq_compact S" shows "\e>0. \k. finite k \ k \ S \ S \ (\x\k. ball x e)" proof - { fix e::real assume "e > 0" assume *: "\k. finite k \ k \ S \ \ S \ (\x\k. ball x e)" let ?Q = "\x n r. r \ S \ (\m < (n::nat). \ (dist (x m) r < e))" have "\x. \n::nat. ?Q x n (x n)" proof (rule dependent_wellorder_choice) fix n x assume "\y. y < n \ ?Q x y (x y)" then have "\ S \ (\x\x ` {0..S" "z \ (\x\x ` {0..r. ?Q x n r" using z by auto qed simp then obtain x where "\n::nat. x n \ S" and x:"\n m. m < n \ \ (dist (x m) (x n) < e)" by blast then obtain l r where "l \ S" and r:"strict_mono r" and "((x \ r) \ l) sequentially" using assms by (metis seq_compact_def) then have "Cauchy (x \ r)" using LIMSEQ_imp_Cauchy by auto then obtain N::nat where "\m n. N \ m \ N \ n \ dist ((x \ r) m) ((x \ r) n) < e" unfolding cauchy_def using \e > 0\ by blast then have False using x[of "r N" "r (N+1)"] r by (auto simp: strict_mono_def) } then show ?thesis by metis qed subsubsection\Heine-Borel theorem\ proposition seq_compact_imp_Heine_Borel: fixes S :: "'a :: metric_space set" assumes "seq_compact S" shows "compact S" proof - from seq_compact_imp_totally_bounded[OF \seq_compact S\] obtain f where f: "\e>0. finite (f e) \ f e \ S \ S \ (\x\f e. ball x e)" unfolding choice_iff' .. define K where "K = (\(x, r). ball x r) ` ((\e \ \ \ {0 <..}. f e) \ \)" have "countably_compact S" using \seq_compact S\ by (rule seq_compact_imp_countably_compact) then show "compact S" proof (rule countably_compact_imp_compact) show "countable K" unfolding K_def using f by (auto intro: countable_finite countable_subset countable_rat intro!: countable_image countable_SIGMA countable_UN) show "\b\K. open b" by (auto simp: K_def) next fix T x assume T: "open T" "x \ T" and x: "x \ S" from openE[OF T] obtain e where "0 < e" "ball x e \ T" by auto then have "0 < e/2" "ball x (e/2) \ T" by auto from Rats_dense_in_real[OF \0 < e/2\] obtain r where "r \ \" "0 < r" "r < e/2" by auto from f[rule_format, of r] \0 < r\ \x \ S\ obtain k where "k \ f r" "x \ ball k r" by auto from \r \ \\ \0 < r\ \k \ f r\ have "ball k r \ K" by (auto simp: K_def) then show "\b\K. x \ b \ b \ S \ T" proof (rule bexI[rotated], safe) fix y assume "y \ ball k r" with \r < e/2\ \x \ ball k r\ have "dist x y < e" by (intro dist_triangle_half_r [of k _ e]) (auto simp: dist_commute) with \ball x e \ T\ show "y \ T" by auto next show "x \ ball k r" by fact qed qed qed proposition compact_eq_seq_compact_metric: "compact (S :: 'a::metric_space set) \ seq_compact S" using compact_imp_seq_compact seq_compact_imp_Heine_Borel by blast proposition compact_def: \ \this is the definition of compactness in HOL Light\ "compact (S :: 'a::metric_space set) \ (\f. (\n. f n \ S) \ (\l\S. \r::nat\nat. strict_mono r \ (f \ r) \ l))" unfolding compact_eq_seq_compact_metric seq_compact_def by auto subsubsection \Complete the chain of compactness variants\ proposition compact_eq_Bolzano_Weierstrass: fixes S :: "'a::metric_space set" shows "compact S \ (\T. infinite T \ T \ S \ (\x \ S. x islimpt T))" using Bolzano_Weierstrass_imp_seq_compact Heine_Borel_imp_Bolzano_Weierstrass compact_eq_seq_compact_metric by blast proposition Bolzano_Weierstrass_imp_bounded: "(\T. \infinite T; T \ S\ \ (\x \ S. x islimpt T)) \ bounded S" using compact_imp_bounded unfolding compact_eq_Bolzano_Weierstrass by metis subsection \Banach fixed point theorem\ theorem banach_fix:\ \TODO: rename to \Banach_fix\\ assumes s: "complete s" "s \ {}" and c: "0 \ c" "c < 1" and f: "f ` s \ s" and lipschitz: "\x\s. \y\s. dist (f x) (f y) \ c * dist x y" shows "\!x\s. f x = x" proof - from c have "1 - c > 0" by simp from s(2) obtain z0 where z0: "z0 \ s" by blast define z where "z n = (f ^^ n) z0" for n with f z0 have z_in_s: "z n \ s" for n :: nat by (induct n) auto define d where "d = dist (z 0) (z 1)" have fzn: "f (z n) = z (Suc n)" for n by (simp add: z_def) have cf_z: "dist (z n) (z (Suc n)) \ (c ^ n) * d" for n :: nat proof (induct n) case 0 then show ?case by (simp add: d_def) next case (Suc m) with \0 \ c\ have "c * dist (z m) (z (Suc m)) \ c ^ Suc m * d" using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by simp then show ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s] by (simp add: fzn mult_le_cancel_left) qed have cf_z2: "(1 - c) * dist (z m) (z (m + n)) \ (c ^ m) * d * (1 - c ^ n)" for n m :: nat proof (induct n) case 0 show ?case by simp next case (Suc k) from c have "(1 - c) * dist (z m) (z (m + Suc k)) \ (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))" by (simp add: dist_triangle) also from c cf_z[of "m + k"] have "\ \ (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)" by simp also from Suc have "\ \ c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d" by (simp add: field_simps) also have "\ = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)" by (simp add: power_add field_simps) also from c have "\ \ (c ^ m) * d * (1 - c ^ Suc k)" by (simp add: field_simps) finally show ?case by simp qed have "\N. \m n. N \ m \ N \ n \ dist (z m) (z n) < e" if "e > 0" for e proof (cases "d = 0") case True from \1 - c > 0\ have "(1 - c) * x \ 0 \ x \ 0" for x by (simp add: mult_le_0_iff) with c cf_z2[of 0] True have "z n = z0" for n by (simp add: z_def) with \e > 0\ show ?thesis by simp next case False with zero_le_dist[of "z 0" "z 1"] have "d > 0" by (metis d_def less_le) with \1 - c > 0\ \e > 0\ have "0 < e * (1 - c) / d" by simp with c obtain N where N: "c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] by auto have *: "dist (z m) (z n) < e" if "m > n" and as: "m \ N" "n \ N" for m n :: nat proof - from c \n \ N\ have *: "c ^ n \ c ^ N" using power_decreasing[OF \n\N\, of c] by simp from c \m > n\ have "1 - c ^ (m - n) > 0" using power_strict_mono[of c 1 "m - n"] by simp with \d > 0\ \0 < 1 - c\ have **: "d * (1 - c ^ (m - n)) / (1 - c) > 0" by simp from cf_z2[of n "m - n"] \m > n\ have "dist (z m) (z n) \ c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" by (simp add: pos_le_divide_eq[OF \1 - c > 0\] mult.commute dist_commute) also have "\ \ c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" using mult_right_mono[OF * order_less_imp_le[OF **]] by (simp add: mult.assoc) also have "\ < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" using mult_strict_right_mono[OF N **] by (auto simp: mult.assoc) also from c \d > 0\ \1 - c > 0\ have "\ = e * (1 - c ^ (m - n))" by simp also from c \1 - c ^ (m - n) > 0\ \e > 0\ have "\ \ e" using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto finally show ?thesis by simp qed have "dist (z n) (z m) < e" if "N \ m" "N \ n" for m n :: nat proof (cases "n = m") case True with \e > 0\ show ?thesis by simp next case False with *[of n m] *[of m n] and that show ?thesis by (auto simp: dist_commute nat_neq_iff) qed then show ?thesis by auto qed then have "Cauchy z" by (simp add: cauchy_def) then obtain x where "x\s" and x:"(z \ x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto define e where "e = dist (f x) x" have "e = 0" proof (rule ccontr) assume "e \ 0" then have "e > 0" unfolding e_def using zero_le_dist[of "f x" x] by (metis dist_eq_0_iff dist_nz e_def) then obtain N where N:"\n\N. dist (z n) x < e/2" using x[unfolded lim_sequentially, THEN spec[where x="e/2"]] by auto then have N':"dist (z N) x < e/2" by auto have *: "c * dist (z N) x \ dist (z N) x" unfolding mult_le_cancel_right2 using zero_le_dist[of "z N" x] and c by (metis dist_eq_0_iff dist_nz order_less_asym less_le) have "dist (f (z N)) (f x) \ c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]] using z_in_s[of N] \x\s\ using c by auto also have "\ < e/2" using N' and c using * by auto finally show False unfolding fzn using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x] unfolding e_def by auto qed then have "f x = x" by (auto simp: e_def) moreover have "y = x" if "f y = y" "y \ s" for y proof - from \x \ s\ \f x = x\ that have "dist x y \ c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] by simp with c and zero_le_dist[of x y] have "dist x y = 0" by (simp add: mult_le_cancel_right1) then show ?thesis by simp qed ultimately show ?thesis using \x\s\ by blast qed subsection \Edelstein fixed point theorem\ theorem Edelstein_fix: fixes S :: "'a::metric_space set" assumes S: "compact S" "S \ {}" and gs: "(g ` S) \ S" and dist: "\x\S. \y\S. x \ y \ dist (g x) (g y) < dist x y" shows "\!x\S. g x = x" proof - let ?D = "(\x. (x, x)) ` S" have D: "compact ?D" "?D \ {}" by (rule compact_continuous_image) (auto intro!: S continuous_Pair continuous_ident simp: continuous_on_eq_continuous_within) have "\x y e. x \ S \ y \ S \ 0 < e \ dist y x < e \ dist (g y) (g x) < e" using dist by fastforce then have "continuous_on S g" by (auto simp: continuous_on_iff) then have cont: "continuous_on ?D (\x. dist ((g \ fst) x) (snd x))" unfolding continuous_on_eq_continuous_within by (intro continuous_dist ballI continuous_within_compose) (auto intro!: continuous_fst continuous_snd continuous_ident simp: image_image) obtain a where "a \ S" and le: "\x. x \ S \ dist (g a) a \ dist (g x) x" using continuous_attains_inf[OF D cont] by auto have "g a = a" proof (rule ccontr) assume "g a \ a" with \a \ S\ gs have "dist (g (g a)) (g a) < dist (g a) a" by (intro dist[rule_format]) auto moreover have "dist (g a) a \ dist (g (g a)) (g a)" using \a \ S\ gs by (intro le) auto ultimately show False by auto qed moreover have "\x. x \ S \ g x = x \ x = a" using dist[THEN bspec[where x=a]] \g a = a\ and \a\S\ by auto ultimately show "\!x\S. g x = x" using \a \ S\ by blast qed subsection \The diameter of a set\ definition\<^marker>\tag important\ diameter :: "'a::metric_space set \ real" where "diameter S = (if S = {} then 0 else SUP (x,y)\S\S. dist x y)" lemma diameter_empty [simp]: "diameter{} = 0" by (auto simp: diameter_def) lemma diameter_singleton [simp]: "diameter{x} = 0" by (auto simp: diameter_def) lemma diameter_le: assumes "S \ {} \ 0 \ d" and no: "\x y. \x \ S; y \ S\ \ norm(x - y) \ d" shows "diameter S \ d" using assms by (auto simp: dist_norm diameter_def intro: cSUP_least) lemma diameter_bounded_bound: fixes S :: "'a :: metric_space set" assumes S: "bounded S" "x \ S" "y \ S" shows "dist x y \ diameter S" proof - from S obtain z d where z: "\x. x \ S \ dist z x \ d" unfolding bounded_def by auto have "bdd_above (case_prod dist ` (S\S))" proof (intro bdd_aboveI, safe) fix a b assume "a \ S" "b \ S" with z[of a] z[of b] dist_triangle[of a b z] show "dist a b \ 2 * d" by (simp add: dist_commute) qed moreover have "(x,y) \ S\S" using S by auto ultimately have "dist x y \ (SUP (x,y)\S\S. dist x y)" by (rule cSUP_upper2) simp with \x \ S\ show ?thesis by (auto simp: diameter_def) qed lemma diameter_lower_bounded: fixes S :: "'a :: metric_space set" assumes S: "bounded S" and d: "0 < d" "d < diameter S" shows "\x\S. \y\S. d < dist x y" proof (rule ccontr) assume contr: "\ ?thesis" moreover have "S \ {}" using d by (auto simp: diameter_def) ultimately have "diameter S \ d" by (auto simp: not_less diameter_def intro!: cSUP_least) with \d < diameter S\ show False by auto qed lemma diameter_bounded: assumes "bounded S" shows "\x\S. \y\S. dist x y \ diameter S" and "\d>0. d < diameter S \ (\x\S. \y\S. dist x y > d)" using diameter_bounded_bound[of S] diameter_lower_bounded[of S] assms by auto lemma bounded_two_points: "bounded S \ (\e. \x\S. \y\S. dist x y \ e)" by (meson bounded_def diameter_bounded(1)) lemma diameter_compact_attained: assumes "compact S" and "S \ {}" shows "\x\S. \y\S. dist x y = diameter S" proof - have b: "bounded S" using assms(1) by (rule compact_imp_bounded) then obtain x y where xys: "x\S" "y\S" and xy: "\u\S. \v\S. dist u v \ dist x y" using compact_sup_maxdistance[OF assms] by auto then have "diameter S \ dist x y" unfolding diameter_def by (force intro!: cSUP_least) then show ?thesis by (metis b diameter_bounded_bound order_antisym xys) qed lemma diameter_ge_0: assumes "bounded S" shows "0 \ diameter S" by (metis all_not_in_conv assms diameter_bounded_bound diameter_empty dist_self order_refl) lemma diameter_subset: assumes "S \ T" "bounded T" shows "diameter S \ diameter T" proof (cases "S = {} \ T = {}") case True with assms show ?thesis by (force simp: diameter_ge_0) next case False then have "bdd_above ((\x. case x of (x, xa) \ dist x xa) ` (T \ T))" using \bounded T\ diameter_bounded_bound by (force simp: bdd_above_def) with False \S \ T\ show ?thesis apply (simp add: diameter_def) apply (rule cSUP_subset_mono, auto) done qed lemma diameter_closure: assumes "bounded S" shows "diameter(closure S) = diameter S" proof (rule order_antisym) have "False" if d_less_d: "diameter S < diameter (closure S)" proof - define d where "d = diameter(closure S) - diameter(S)" have "d > 0" using that by (simp add: d_def) then have dle: "diameter(closure(S)) - d / 2 < diameter(closure(S))" by simp have dd: "diameter (closure S) - d / 2 = (diameter(closure(S)) + diameter(S)) / 2" by (simp add: d_def field_split_simps) have bocl: "bounded (closure S)" using assms by blast moreover have "0 \ diameter S" using assms diameter_ge_0 by blast ultimately obtain x y where "x \ closure S" "y \ closure S" and xy: "diameter(closure(S)) - d / 2 < dist x y" by (smt (verit) dle d_less_d d_def dd diameter_lower_bounded) then obtain x' y' where x'y': "x' \ S" "dist x' x < d/4" "y' \ S" "dist y' y < d/4" by (metis \0 < d\ zero_less_divide_iff zero_less_numeral closure_approachable) then have "dist x' y' \ diameter S" using assms diameter_bounded_bound by blast with x'y' have "dist x y \ d / 4 + diameter S + d / 4" by (meson add_mono dist_triangle dist_triangle3 less_eq_real_def order_trans) then show ?thesis using xy d_def by linarith qed then show "diameter (closure S) \ diameter S" by fastforce next show "diameter S \ diameter (closure S)" by (simp add: assms bounded_closure closure_subset diameter_subset) qed proposition Lebesgue_number_lemma: assumes "compact S" "\ \ {}" "S \ \\" and ope: "\B. B \ \ \ open B" obtains \ where "0 < \" "\T. \T \ S; diameter T < \\ \ \B \ \. T \ B" proof (cases "S = {}") case True then show ?thesis by (metis \\ \ {}\ zero_less_one empty_subsetI equals0I subset_trans that) next case False { fix x assume "x \ S" then obtain C where C: "x \ C" "C \ \" using \S \ \\\ by blast then obtain r where r: "r>0" "ball x (2*r) \ C" by (metis mult.commute mult_2_right not_le ope openE field_sum_of_halves zero_le_numeral zero_less_mult_iff) then have "\r C. r > 0 \ ball x (2*r) \ C \ C \ \" using C by blast } then obtain r where r: "\x. x \ S \ r x > 0 \ (\C \ \. ball x (2*r x) \ C)" by metis then have "S \ (\x \ S. ball x (r x))" by auto then obtain \ where "finite \" "S \ \\" and \: "\ \ (\x. ball x (r x)) ` S" by (rule compactE [OF \compact S\]) auto then obtain S0 where "S0 \ S" "finite S0" and S0: "\ = (\x. ball x (r x)) ` S0" by (meson finite_subset_image) then have "S0 \ {}" using False \S \ \\\ by auto define \ where "\ = Inf (r ` S0)" have "\ > 0" using \finite S0\ \S0 \ S\ \S0 \ {}\ r by (auto simp: \_def finite_less_Inf_iff) show ?thesis proof show "0 < \" by (simp add: \0 < \\) show "\B \ \. T \ B" if "T \ S" and dia: "diameter T < \" for T proof (cases "T = {}") case True then show ?thesis using \\ \ {}\ by blast next case False then obtain y where "y \ T" by blast then have "y \ S" using \T \ S\ by auto then obtain x where "x \ S0" and x: "y \ ball x (r x)" using \S \ \\\ S0 that by blast have "ball y \ \ ball y (r x)" by (metis \_def \S0 \ {}\ \finite S0\ \x \ S0\ empty_is_image finite_imageI finite_less_Inf_iff imageI less_irrefl not_le subset_ball) also have "... \ ball x (2*r x)" using x by metric finally obtain C where "C \ \" "ball y \ \ C" by (meson r \S0 \ S\ \x \ S0\ dual_order.trans subsetCE) have "bounded T" using \compact S\ bounded_subset compact_imp_bounded \T \ S\ by blast then have "T \ ball y \" using \y \ T\ dia diameter_bounded_bound by fastforce then show ?thesis by (meson \C \ \\ \ball y \ \ C\ subset_eq) qed qed qed subsection \Metric spaces with the Heine-Borel property\ text \ A metric space (or topological vector space) is said to have the Heine-Borel property if every closed and bounded subset is compact. \ class heine_borel = metric_space + assumes bounded_imp_convergent_subsequence: "bounded (range f) \ \l r. strict_mono (r::nat\nat) \ ((f \ r) \ l) sequentially" proposition bounded_closed_imp_seq_compact: fixes S::"'a::heine_borel set" assumes "bounded S" and "closed S" shows "seq_compact S" proof (unfold seq_compact_def, clarify) fix f :: "nat \ 'a" assume f: "\n. f n \ S" with \bounded S\ have "bounded (range f)" by (auto intro: bounded_subset) obtain l r where r: "strict_mono (r :: nat \ nat)" and l: "((f \ r) \ l) sequentially" using bounded_imp_convergent_subsequence [OF \bounded (range f)\] by auto from f have fr: "\n. (f \ r) n \ S" by simp have "l \ S" using \closed S\ fr l by (rule closed_sequentially) show "\l\S. \r. strict_mono r \ ((f \ r) \ l) sequentially" using \l \ S\ r l by blast qed lemma compact_eq_bounded_closed: fixes S :: "'a::heine_borel set" shows "compact S \ bounded S \ closed S" using bounded_closed_imp_seq_compact compact_eq_seq_compact_metric compact_imp_bounded compact_imp_closed by auto lemma bounded_infinite_imp_islimpt: fixes S :: "'a::heine_borel set" assumes "T \ S" "bounded S" "infinite T" obtains x where "x islimpt S" by (meson assms closed_limpt compact_eq_Bolzano_Weierstrass compact_eq_bounded_closed islimpt_subset) lemma compact_Inter: fixes \ :: "'a :: heine_borel set set" assumes com: "\S. S \ \ \ compact S" and "\ \ {}" shows "compact(\ \)" using assms by (meson Inf_lower all_not_in_conv bounded_subset closed_Inter compact_eq_bounded_closed) lemma compact_closure [simp]: fixes S :: "'a::heine_borel set" shows "compact(closure S) \ bounded S" by (meson bounded_closure bounded_subset closed_closure closure_subset compact_eq_bounded_closed) instance\<^marker>\tag important\ real :: heine_borel proof fix f :: "nat \ real" assume f: "bounded (range f)" obtain r :: "nat \ nat" where r: "strict_mono r" "monoseq (f \ r)" unfolding comp_def by (metis seq_monosub) then have "Bseq (f \ r)" unfolding Bseq_eq_bounded by (metis f BseqI' bounded_iff comp_apply rangeI) with r show "\l r. strict_mono r \ (f \ r) \ l" using Bseq_monoseq_convergent[of "f \ r"] by (auto simp: convergent_def) qed lemma compact_lemma_general: fixes f :: "nat \ 'a" fixes proj::"'a \ 'b \ 'c::heine_borel" (infixl "proj" 60) fixes unproj:: "('b \ 'c) \ 'a" assumes finite_basis: "finite basis" assumes bounded_proj: "\k. k \ basis \ bounded ((\x. x proj k) ` range f)" assumes proj_unproj: "\e k. k \ basis \ (unproj e) proj k = e k" assumes unproj_proj: "\x. unproj (\k. x proj k) = x" shows "\d\basis. \l::'a. \ r::nat\nat. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) proj i) (l proj i) < e) sequentially)" proof safe fix d :: "'b set" assume d: "d \ basis" with finite_basis have "finite d" by (blast intro: finite_subset) from this d show "\l::'a. \r::nat\nat. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) proj i) (l proj i) < e) sequentially)" proof (induct d) case empty then show ?case unfolding strict_mono_def by auto next case (insert k d) have k[intro]: "k \ basis" using insert by auto have s': "bounded ((\x. x proj k) ` range f)" using k by (rule bounded_proj) obtain l1::"'a" and r1 where r1: "strict_mono r1" and lr1: "\e > 0. \\<^sub>F n in sequentially. \i\d. dist (f (r1 n) proj i) (l1 proj i) < e" using insert by auto have f': "\n. f (r1 n) proj k \ (\x. x proj k) ` range f" by simp have "bounded (range (\i. f (r1 i) proj k))" by (metis (lifting) bounded_subset f' image_subsetI s') then obtain l2 r2 where r2: "strict_mono r2" and lr2: "(\i. f (r1 (r2 i)) proj k) \ l2" using bounded_imp_convergent_subsequence[of "\i. f (r1 i) proj k"] by (auto simp: o_def) define r where "r = r1 \ r2" have r: "strict_mono r" using r1 and r2 unfolding r_def o_def strict_mono_def by auto moreover define l where "l = unproj (\i. if i = k then l2 else l1 proj i)" { fix e::real assume "e > 0" from lr1 \e > 0\ have N1: "\\<^sub>F n in sequentially. \i\d. dist (f (r1 n) proj i) (l1 proj i) < e" by blast from lr2 \e > 0\ have N2: "\\<^sub>F n in sequentially. dist (f (r1 (r2 n)) proj k) l2 < e" by (rule tendstoD) from r2 N1 have N1': "\\<^sub>F n in sequentially. \i\d. dist (f (r1 (r2 n)) proj i) (l1 proj i) < e" by (rule eventually_subseq) have "\\<^sub>F n in sequentially. \i\insert k d. dist (f (r n) proj i) (l proj i) < e" using N1' N2 by eventually_elim (use insert.prems in \auto simp: l_def r_def o_def proj_unproj\) } ultimately show ?case by auto qed qed lemma bounded_fst: "bounded s \ bounded (fst ` s)" unfolding bounded_def by (metis (erased, opaque_lifting) dist_fst_le image_iff order_trans) lemma bounded_snd: "bounded s \ bounded (snd ` s)" unfolding bounded_def by (metis (no_types, opaque_lifting) dist_snd_le image_iff order.trans) instance\<^marker>\tag important\ prod :: (heine_borel, heine_borel) heine_borel proof fix f :: "nat \ 'a \ 'b" assume f: "bounded (range f)" then have "bounded (fst ` range f)" by (rule bounded_fst) then have s1: "bounded (range (fst \ f))" by (simp add: image_comp) obtain l1 r1 where r1: "strict_mono r1" and l1: "(\n. fst (f (r1 n))) \ l1" using bounded_imp_convergent_subsequence [OF s1] unfolding o_def by fast from f have s2: "bounded (range (snd \ f \ r1))" by (auto simp: image_comp intro: bounded_snd bounded_subset) obtain l2 r2 where r2: "strict_mono r2" and l2: "(\n. snd (f (r1 (r2 n)))) \ l2" using bounded_imp_convergent_subsequence [OF s2] unfolding o_def by fast have l1': "((\n. fst (f (r1 (r2 n)))) \ l1) sequentially" using LIMSEQ_subseq_LIMSEQ [OF l1 r2] unfolding o_def . have l: "((f \ (r1 \ r2)) \ (l1, l2)) sequentially" using tendsto_Pair [OF l1' l2] unfolding o_def by simp have r: "strict_mono (r1 \ r2)" using r1 r2 unfolding strict_mono_def by simp show "\l r. strict_mono r \ ((f \ r) \ l) sequentially" using l r by fast qed subsection \Completeness\ proposition (in metric_space) completeI: assumes "\f. \n. f n \ s \ Cauchy f \ \l\s. f \ l" shows "complete s" using assms unfolding complete_def by fast proposition (in metric_space) completeE: assumes "complete s" and "\n. f n \ s" and "Cauchy f" obtains l where "l \ s" and "f \ l" using assms unfolding complete_def by fast (* TODO: generalize to uniform spaces *) lemma compact_imp_complete: fixes s :: "'a::metric_space set" assumes "compact s" shows "complete s" proof - { fix f assume as: "(\n::nat. f n \ s)" "Cauchy f" from as(1) obtain l r where lr: "l\s" "strict_mono r" "(f \ r) \ l" using assms unfolding compact_def by blast note lr' = seq_suble [OF lr(2)] { fix e :: real assume "e > 0" from as(2) obtain N where N:"\m n. N \ m \ N \ n \ dist (f m) (f n) < e/2" unfolding cauchy_def using \e > 0\ by (meson half_gt_zero) then obtain M where M:"\n\M. dist ((f \ r) n) l < e/2" by (metis dist_self lim_sequentially lr(3)) { fix n :: nat assume n: "n \ max N M" have "dist ((f \ r) n) l < e/2" using n M by auto moreover have "r n \ N" using lr'[of n] n by auto then have "dist (f n) ((f \ r) n) < e/2" using N and n by auto ultimately have "dist (f n) l < e" using n M by metric } then have "\N. \n\N. dist (f n) l < e" by blast } then have "\l\s. (f \ l) sequentially" using \l\s\ unfolding lim_sequentially by auto } then show ?thesis unfolding complete_def by auto qed proposition compact_eq_totally_bounded: "compact s \ complete s \ (\e>0. \k. finite k \ s \ (\x\k. ball x e))" (is "_ \ ?rhs") proof assume assms: "?rhs" then obtain k where k: "\e. 0 < e \ finite (k e)" "\e. 0 < e \ s \ (\x\k e. ball x e)" by (auto simp: choice_iff') show "compact s" proof cases assume "s = {}" then show "compact s" by (simp add: compact_def) next assume "s \ {}" show ?thesis unfolding compact_def proof safe fix f :: "nat \ 'a" assume f: "\n. f n \ s" define e where "e n = 1 / (2 * Suc n)" for n then have [simp]: "\n. 0 < e n" by auto define B where "B n U = (SOME b. infinite {n. f n \ b} \ (\x. b \ ball x (e n) \ U))" for n U { fix n U assume "infinite {n. f n \ U}" then have "\b\k (e n). infinite {i\{n. f n \ U}. f i \ ball b (e n)}" using k f by (intro pigeonhole_infinite_rel) (auto simp: subset_eq) then obtain a where "a \ k (e n)" "infinite {i \ {n. f n \ U}. f i \ ball a (e n)}" .. then have "\b. infinite {i. f i \ b} \ (\x. b \ ball x (e n) \ U)" by (intro exI[of _ "ball a (e n) \ U"] exI[of _ a]) (auto simp: ac_simps) from someI_ex[OF this] have "infinite {i. f i \ B n U}" "\x. B n U \ ball x (e n) \ U" unfolding B_def by auto } note B = this define F where "F = rec_nat (B 0 UNIV) B" { fix n have "infinite {i. f i \ F n}" by (induct n) (auto simp: F_def B) } then have F: "\n. \x. F (Suc n) \ ball x (e n) \ F n" using B by (simp add: F_def) then have F_dec: "\m n. m \ n \ F n \ F m" using decseq_SucI[of F] by (auto simp: decseq_def) obtain sel where sel: "\k i. i < sel k i" "\k i. f (sel k i) \ F k" proof (atomize_elim, unfold all_conj_distrib[symmetric], intro choice allI) fix k i have "infinite ({n. f n \ F k} - {.. i})" using \infinite {n. f n \ F k}\ by auto from infinite_imp_nonempty[OF this] show "\x>i. f x \ F k" by (simp add: set_eq_iff not_le conj_commute) qed define t where "t = rec_nat (sel 0 0) (\n i. sel (Suc n) i)" have "strict_mono t" unfolding strict_mono_Suc_iff by (simp add: t_def sel) moreover have "\i. (f \ t) i \ s" using f by auto moreover have t: "(f \ t) n \ F n" for n by (cases n) (simp_all add: t_def sel) have "Cauchy (f \ t)" proof (safe intro!: metric_CauchyI exI elim!: nat_approx_posE) fix r :: real and N n m assume "1 / Suc N < r" "Suc N \ n" "Suc N \ m" then have "(f \ t) n \ F (Suc N)" "(f \ t) m \ F (Suc N)" "2 * e N < r" using F_dec t by (auto simp: e_def field_simps) with F[of N] obtain x where "dist x ((f \ t) n) < e N" "dist x ((f \ t) m) < e N" by (auto simp: subset_eq) with \2 * e N < r\ show "dist ((f \ t) m) ((f \ t) n) < r" by metric qed ultimately show "\l\s. \r. strict_mono r \ (f \ r) \ l" using assms unfolding complete_def by blast qed qed qed (metis compact_imp_complete compact_imp_seq_compact seq_compact_imp_totally_bounded) lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded (range s)" proof - from assms obtain N :: nat where "\m n. N \ m \ N \ n \ dist (s m) (s n) < 1" unfolding cauchy_def by force then have N:"\n. N \ n \ dist (s N) (s n) < 1" by auto moreover have "bounded (s ` {0..N})" using finite_imp_bounded[of "s ` {1..N}"] by auto then obtain a where a:"\x\s ` {0..N}. dist (s N) x \ a" unfolding bounded_any_center [where a="s N"] by auto ultimately show "?thesis" unfolding bounded_any_center [where a="s N"] apply (rule_tac x="max a 1" in exI, auto) apply (erule_tac x=y in allE) apply (erule_tac x=y in ballE, auto) done qed instance heine_borel < complete_space proof fix f :: "nat \ 'a" assume "Cauchy f" then show "convergent f" unfolding convergent_def using Cauchy_converges_subseq cauchy_imp_bounded bounded_imp_convergent_subsequence by blast qed lemma complete_UNIV: "complete (UNIV :: ('a::complete_space) set)" proof (rule completeI) fix f :: "nat \ 'a" assume "Cauchy f" then show "\l\UNIV. f \ l" using Cauchy_convergent convergent_def by blast qed lemma complete_imp_closed: fixes S :: "'a::metric_space set" shows "complete S \ closed S" by (metis LIMSEQ_imp_Cauchy LIMSEQ_unique closed_sequential_limits completeE) lemma complete_Int_closed: fixes S :: "'a::metric_space set" assumes "complete S" and "closed t" shows "complete (S \ t)" by (metis Int_iff assms closed_sequentially completeE completeI) lemma complete_closed_subset: fixes S :: "'a::metric_space set" assumes "closed S" and "S \ t" and "complete t" shows "complete S" using assms complete_Int_closed [of t S] by (simp add: Int_absorb1) lemma complete_eq_closed: fixes S :: "('a::complete_space) set" shows "complete S \ closed S" proof assume "closed S" then show "complete S" using subset_UNIV complete_UNIV by (rule complete_closed_subset) next assume "complete S" then show "closed S" by (rule complete_imp_closed) qed lemma convergent_eq_Cauchy: fixes S :: "nat \ 'a::complete_space" shows "(\l. (S \ l) sequentially) \ Cauchy S" unfolding Cauchy_convergent_iff convergent_def .. lemma convergent_imp_bounded: fixes S :: "nat \ 'a::metric_space" shows "(S \ l) sequentially \ bounded (range S)" by (intro cauchy_imp_bounded LIMSEQ_imp_Cauchy) lemma frontier_subset_compact: fixes S :: "'a::heine_borel set" shows "compact S \ frontier S \ S" using frontier_subset_closed compact_eq_bounded_closed by blast lemma continuous_closed_imp_Cauchy_continuous: fixes S :: "('a::complete_space) set" shows "\continuous_on S f; closed S; Cauchy \; \n. (\ n) \ S\ \ Cauchy(f \ \)" by (meson LIMSEQ_imp_Cauchy completeE complete_eq_closed continuous_on_sequentially) lemma banach_fix_type: fixes f::"'a::complete_space\'a" assumes c:"0 \ c" "c < 1" and lipschitz:"\x. \y. dist (f x) (f y) \ c * dist x y" shows "\!x. (f x = x)" using assms banach_fix[OF complete_UNIV UNIV_not_empty assms(1,2) subset_UNIV, of f] by auto subsection\<^marker>\tag unimportant\\ Finite intersection property\ text\Also developed in HOL's toplogical spaces theory, but the Heine-Borel type class isn't available there.\ lemma closed_imp_fip: fixes S :: "'a::heine_borel set" assumes "closed S" and T: "T \ \" "bounded T" and clof: "\T. T \ \ \ closed T" and "\\'. \finite \'; \' \ \\ \ S \ \\' \ {}" shows "S \ \\ \ {}" proof - have "compact (S \ T)" using \closed S\ clof compact_eq_bounded_closed T by blast then have "(S \ T) \ \\ \ {}" by (smt (verit, best) Inf_insert Int_assoc assms compact_imp_fip finite_insert insert_subset) then show ?thesis by blast qed lemma closed_imp_fip_compact: fixes S :: "'a::heine_borel set" shows "\closed S; \T. T \ \ \ compact T; \\'. \finite \'; \' \ \\ \ S \ \\' \ {}\ \ S \ \\ \ {}" by (metis closed_imp_fip compact_eq_bounded_closed equals0I finite.emptyI order.refl) lemma closed_fip_Heine_Borel: fixes \ :: "'a::heine_borel set set" assumes "T \ \" "bounded T" and "\T. T \ \ \ closed T" and "\\'. \finite \'; \' \ \\ \ \\' \ {}" shows "\\ \ {}" using closed_imp_fip [OF closed_UNIV] assms by auto lemma compact_fip_Heine_Borel: fixes \ :: "'a::heine_borel set set" assumes clof: "\T. T \ \ \ compact T" and none: "\\'. \finite \'; \' \ \\ \ \\' \ {}" shows "\\ \ {}" by (metis InterI clof closed_fip_Heine_Borel compact_eq_bounded_closed equals0D none) lemma compact_sequence_with_limit: fixes f :: "nat \ 'a::heine_borel" shows "f \ l \ compact (insert l (range f))" by (simp add: closed_limpt compact_eq_bounded_closed convergent_imp_bounded islimpt_insert sequence_unique_limpt) subsection \Properties of Balls and Spheres\ lemma compact_cball[simp]: fixes x :: "'a::heine_borel" shows "compact (cball x e)" using compact_eq_bounded_closed bounded_cball closed_cball by blast lemma compact_frontier_bounded[intro]: fixes S :: "'a::heine_borel set" shows "bounded S \ compact (frontier S)" unfolding frontier_def using compact_eq_bounded_closed by blast lemma compact_frontier[intro]: fixes S :: "'a::heine_borel set" shows "compact S \ compact (frontier S)" using compact_eq_bounded_closed compact_frontier_bounded by blast lemma no_limpt_imp_countable: assumes "\z. \z islimpt (X :: 'a :: {real_normed_vector, heine_borel} set)" shows "countable X" proof - have "countable (\n. cball 0 (real n) \ X)" proof (intro countable_UN[OF _ countable_finite]) fix n :: nat show "finite (cball 0 (real n) \ X)" using assms by (intro finite_not_islimpt_in_compact) auto qed auto also have "(\n. cball 0 (real n)) = (UNIV :: 'a set)" proof safe fix z :: 'a have "norm z \ 0" by simp hence "real (nat \norm z\) \ norm z" by linarith thus "z \ (\n. cball 0 (real n))" by auto qed auto hence "(\n. cball 0 (real n) \ X) = X" by blast finally show "countable X" . qed subsection \Distance from a Set\ lemma distance_attains_sup: assumes "compact s" "s \ {}" shows "\x\s. \y\s. dist a y \ dist a x" proof (rule continuous_attains_sup [OF assms]) { fix x assume "x\s" have "(dist a \ dist a x) (at x within s)" by (intro tendsto_dist tendsto_const tendsto_ident_at) } then show "continuous_on s (dist a)" unfolding continuous_on .. qed text \For \emph{minimal} distance, we only need closure, not compactness.\ lemma distance_attains_inf: fixes a :: "'a::heine_borel" assumes "closed s" and "s \ {}" obtains x where "x\s" "\y. y \ s \ dist a x \ dist a y" proof - from assms obtain b where "b \ s" by auto let ?B = "s \ cball a (dist b a)" have "?B \ {}" using \b \ s\ by (auto simp: dist_commute) moreover have "continuous_on ?B (dist a)" by (auto intro!: continuous_at_imp_continuous_on continuous_dist continuous_ident continuous_const) moreover have "compact ?B" by (intro closed_Int_compact \closed s\ compact_cball) ultimately obtain x where "x \ ?B" "\y\?B. dist a x \ dist a y" by (metis continuous_attains_inf) with that show ?thesis by fastforce qed subsection \Infimum Distance\ definition\<^marker>\tag important\ "infdist x A = (if A = {} then 0 else INF a\A. dist x a)" lemma bdd_below_image_dist[intro, simp]: "bdd_below (dist x ` A)" by (auto intro!: zero_le_dist) lemma infdist_notempty: "A \ {} \ infdist x A = (INF a\A. dist x a)" by (simp add: infdist_def) lemma infdist_nonneg: "0 \ infdist x A" by (auto simp: infdist_def intro: cINF_greatest) lemma infdist_le: "a \ A \ infdist x A \ dist x a" by (auto intro: cINF_lower simp add: infdist_def) lemma infdist_le2: "a \ A \ dist x a \ d \ infdist x A \ d" by (auto intro!: cINF_lower2 simp add: infdist_def) lemma infdist_zero[simp]: "a \ A \ infdist a A = 0" by (auto intro!: antisym infdist_nonneg infdist_le2) lemma infdist_Un_min: assumes "A \ {}" "B \ {}" shows "infdist x (A \ B) = min (infdist x A) (infdist x B)" using assms by (simp add: infdist_def cINF_union inf_real_def) lemma infdist_triangle: "infdist x A \ infdist y A + dist x y" proof (cases "A = {}") case True then show ?thesis by (simp add: infdist_def) next case False then obtain a where "a \ A" by auto have "infdist x A \ Inf {dist x y + dist y a |a. a \ A}" proof (rule cInf_greatest) from \A \ {}\ show "{dist x y + dist y a |a. a \ A} \ {}" by simp fix d assume "d \ {dist x y + dist y a |a. a \ A}" then obtain a where d: "d = dist x y + dist y a" "a \ A" by auto show "infdist x A \ d" unfolding infdist_notempty[OF \A \ {}\] proof (rule cINF_lower2) show "a \ A" by fact show "dist x a \ d" unfolding d by (rule dist_triangle) qed simp qed also have "\ = dist x y + infdist y A" proof (rule cInf_eq, safe) fix a assume "a \ A" then show "dist x y + infdist y A \ dist x y + dist y a" by (auto intro: infdist_le) next fix i assume inf: "\d. d \ {dist x y + dist y a |a. a \ A} \ i \ d" then have "i - dist x y \ infdist y A" unfolding infdist_notempty[OF \A \ {}\] using \a \ A\ by (intro cINF_greatest) (auto simp: field_simps) then show "i \ dist x y + infdist y A" by simp qed finally show ?thesis by simp qed lemma infdist_triangle_abs: "\infdist x A - infdist y A\ \ dist x y" by (metis (full_types) abs_diff_le_iff diff_le_eq dist_commute infdist_triangle) lemma in_closure_iff_infdist_zero: assumes "A \ {}" shows "x \ closure A \ infdist x A = 0" proof assume "x \ closure A" show "infdist x A = 0" proof (rule ccontr) assume "infdist x A \ 0" with infdist_nonneg[of x A] have "infdist x A > 0" by auto then have "ball x (infdist x A) \ closure A = {}" by (smt (verit, best) \x \ closure A\ closure_approachableD infdist_le) then have "x \ closure A" by (metis \0 < infdist x A\ centre_in_ball disjoint_iff_not_equal) then show False using \x \ closure A\ by simp qed next assume x: "infdist x A = 0" then obtain a where "a \ A" by atomize_elim (metis all_not_in_conv assms) have False if "e > 0" "\ (\y\A. dist y x < e)" for e proof - have "infdist x A \ e" using \a \ A\ unfolding infdist_def using that by (force simp: dist_commute intro: cINF_greatest) with x \e > 0\ show False by auto qed then show "x \ closure A" using closure_approachable by blast qed lemma in_closed_iff_infdist_zero: assumes "closed A" "A \ {}" shows "x \ A \ infdist x A = 0" proof - have "x \ closure A \ infdist x A = 0" by (simp add: \A \ {}\ in_closure_iff_infdist_zero) with assms show ?thesis by simp qed lemma infdist_pos_not_in_closed: assumes "closed S" "S \ {}" "x \ S" shows "infdist x S > 0" by (smt (verit, best) assms in_closed_iff_infdist_zero infdist_nonneg) lemma infdist_attains_inf: fixes X::"'a::heine_borel set" assumes "closed X" assumes "X \ {}" obtains x where "x \ X" "infdist y X = dist y x" proof - have "bdd_below (dist y ` X)" by auto from distance_attains_inf[OF assms, of y] obtain x where "x \ X" "\z. z \ X \ dist y x \ dist y z" by auto then have "infdist y X = dist y x" by (metis antisym assms(2) cINF_greatest infdist_def infdist_le) with \x \ X\ show ?thesis .. qed text \Every metric space is a T4 space:\ instance metric_space \ t4_space proof fix S T::"'a set" assume H: "closed S" "closed T" "S \ T = {}" consider "S = {}" | "T = {}" | "S \ {} \ T \ {}" by auto then show "\U V. open U \ open V \ S \ U \ T \ V \ U \ V = {}" proof (cases) case 1 then show ?thesis by blast next case 2 then show ?thesis by blast next case 3 define U where "U = (\x\S. ball x ((infdist x T)/2))" have A: "open U" unfolding U_def by auto have "infdist x T > 0" if "x \ S" for x using H that 3 by (auto intro!: infdist_pos_not_in_closed) then have B: "S \ U" unfolding U_def by auto define V where "V = (\x\T. ball x ((infdist x S)/2))" have C: "open V" unfolding V_def by auto have "infdist x S > 0" if "x \ T" for x using H that 3 by (auto intro!: infdist_pos_not_in_closed) then have D: "T \ V" unfolding V_def by auto have "(ball x ((infdist x T)/2)) \ (ball y ((infdist y S)/2)) = {}" if "x \ S" "y \ T" for x y proof auto fix z assume H: "dist x z * 2 < infdist x T" "dist y z * 2 < infdist y S" have "2 * dist x y \ 2 * dist x z + 2 * dist y z" by metric also have "... < infdist x T + infdist y S" using H by auto finally have "dist x y < infdist x T \ dist x y < infdist y S" by auto then show False using infdist_le[OF \x \ S\, of y] infdist_le[OF \y \ T\, of x] by (auto simp add: dist_commute) qed then have E: "U \ V = {}" unfolding U_def V_def by auto show ?thesis using A B C D E by blast qed qed lemma tendsto_infdist [tendsto_intros]: assumes f: "(f \ l) F" shows "((\x. infdist (f x) A) \ infdist l A) F" proof (rule tendstoI) fix e ::real assume "e > 0" from tendstoD[OF f this] show "eventually (\x. dist (infdist (f x) A) (infdist l A) < e) F" proof (eventually_elim) fix x from infdist_triangle[of l A "f x"] infdist_triangle[of "f x" A l] have "dist (infdist (f x) A) (infdist l A) \ dist (f x) l" by (simp add: dist_commute dist_real_def) also assume "dist (f x) l < e" finally show "dist (infdist (f x) A) (infdist l A) < e" . qed qed lemma continuous_infdist[continuous_intros]: assumes "continuous F f" shows "continuous F (\x. infdist (f x) A)" using assms unfolding continuous_def by (rule tendsto_infdist) lemma continuous_on_infdist [continuous_intros]: assumes "continuous_on S f" shows "continuous_on S (\x. infdist (f x) A)" using assms unfolding continuous_on by (auto intro: tendsto_infdist) lemma compact_infdist_le: fixes A::"'a::heine_borel set" assumes "A \ {}" assumes "compact A" assumes "e > 0" shows "compact {x. infdist x A \ e}" proof - from continuous_closed_vimage[of "{0..e}" "\x. infdist x A"] continuous_infdist[OF continuous_ident, of _ UNIV A] have "closed {x. infdist x A \ e}" by (auto simp: vimage_def infdist_nonneg) moreover from assms obtain x0 b where b: "\x. x \ A \ dist x0 x \ b" "closed A" by (auto simp: compact_eq_bounded_closed bounded_def) { fix y assume "infdist y A \ e" moreover from infdist_attains_inf[OF \closed A\ \A \ {}\, of y] obtain z where "z \ A" "infdist y A = dist y z" by blast ultimately have "dist x0 y \ b + e" using b by metric } then have "bounded {x. infdist x A \ e}" by (auto simp: bounded_any_center[where a=x0] intro!: exI[where x="b + e"]) ultimately show "compact {x. infdist x A \ e}" by (simp add: compact_eq_bounded_closed) qed subsection \Separation between Points and Sets\ proposition separate_point_closed: fixes S :: "'a::heine_borel set" assumes "closed S" and "a \ S" shows "\d>0. \x\S. d \ dist a x" by (metis assms distance_attains_inf empty_iff gt_ex zero_less_dist_iff) proposition separate_compact_closed: fixes S T :: "'a::heine_borel set" assumes "compact S" and T: "closed T" "S \ T = {}" shows "\d>0. \x\S. \y\T. d \ dist x y" proof cases assume "S \ {} \ T \ {}" then have "S \ {}" "T \ {}" by auto let ?inf = "\x. infdist x T" have "continuous_on S ?inf" by (auto intro!: continuous_at_imp_continuous_on continuous_infdist continuous_ident) then obtain x where x: "x \ S" "\y\S. ?inf x \ ?inf y" using continuous_attains_inf[OF \compact S\ \S \ {}\] by auto then have "0 < ?inf x" using T \T \ {}\ in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg) moreover have "\x'\S. \y\T. ?inf x \ dist x' y" using x by (auto intro: order_trans infdist_le) ultimately show ?thesis by auto qed (auto intro!: exI[of _ 1]) proposition separate_closed_compact: fixes S T :: "'a::heine_borel set" assumes S: "closed S" and T: "compact T" and dis: "S \ T = {}" shows "\d>0. \x\S. \y\T. d \ dist x y" by (metis separate_compact_closed[OF T S] dis dist_commute inf_commute) proposition compact_in_open_separated: fixes A::"'a::heine_borel set" assumes A: "A \ {}" "compact A" assumes "open B" assumes "A \ B" obtains e where "e > 0" "{x. infdist x A \ e} \ B" proof atomize_elim have "closed (- B)" "compact A" "- B \ A = {}" using assms by (auto simp: open_Diff compact_eq_bounded_closed) from separate_closed_compact[OF this] obtain d'::real where d': "d'>0" "\x y. x \ B \ y \ A \ d' \ dist x y" by auto define d where "d = d' / 2" hence "d>0" "d < d'" using d' by auto with d' have d: "\x y. x \ B \ y \ A \ d < dist x y" by force show "\e>0. {x. infdist x A \ e} \ B" proof (rule ccontr) assume "\e. 0 < e \ {x. infdist x A \ e} \ B" with \d > 0\ obtain x where x: "infdist x A \ d" "x \ B" by auto then show False by (metis A compact_eq_bounded_closed infdist_attains_inf x d linorder_not_less) qed qed subsection \Uniform Continuity\ lemma uniformly_continuous_onE: assumes "uniformly_continuous_on s f" "0 < e" obtains d where "d>0" "\x x'. \x\s; x'\s; dist x' x < d\ \ dist (f x') (f x) < e" using assms by (auto simp: uniformly_continuous_on_def) lemma uniformly_continuous_on_sequentially: "uniformly_continuous_on s f \ (\x y. (\n. x n \ s) \ (\n. y n \ s) \ (\n. dist (x n) (y n)) \ 0 \ (\n. dist (f(x n)) (f(y n))) \ 0)" (is "?lhs = ?rhs") proof assume ?lhs { fix x y assume x: "\n. x n \ s" and y: "\n. y n \ s" and xy: "((\n. dist (x n) (y n)) \ 0) sequentially" { fix e :: real assume "e > 0" then obtain d where "d > 0" and d: "\x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e" by (metis \?lhs\ uniformly_continuous_onE) obtain N where N: "\n\N. dist (x n) (y n) < d" using xy[unfolded lim_sequentially dist_norm] and \d>0\ by auto then have "\N. \n\N. dist (f (x n)) (f (y n)) < e" using d x y by blast } then have "((\n. dist (f(x n)) (f(y n))) \ 0) sequentially" unfolding lim_sequentially and dist_real_def by auto } then show ?rhs by auto next assume ?rhs { assume "\ ?lhs" then obtain e where "e > 0" "\d>0. \x\s. \x'\s. dist x' x < d \ \ dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto then obtain fa where fa: "\x. 0 < x \ fst (fa x) \ s \ snd (fa x) \ s \ dist (fst (fa x)) (snd (fa x)) < x \ \ dist (f (fst (fa x))) (f (snd (fa x))) < e" using choice[of "\d x. d>0 \ fst x \ s \ snd x \ s \ dist (snd x) (fst x) < d \ \ dist (f (snd x)) (f (fst x)) < e"] by (auto simp: Bex_def dist_commute) define x where "x n = fst (fa (inverse (real n + 1)))" for n define y where "y n = snd (fa (inverse (real n + 1)))" for n have xyn: "\n. x n \ s \ y n \ s" and xy0: "\n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\n. \ dist (f (x n)) (f (y n)) < e" unfolding x_def and y_def using fa by auto { fix e :: real assume "e > 0" then obtain N :: nat where "N \ 0" and N: "0 < inverse (real N) \ inverse (real N) < e" unfolding real_arch_inverse[of e] by auto then have "\N. \n\N. dist (x n) (y n) < e" by (smt (verit, ccfv_SIG) inverse_le_imp_le nat_le_real_less of_nat_le_0_iff xy0) } then have "\e>0. \N. \n\N. dist (f (x n)) (f (y n)) < e" using \?rhs\[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding lim_sequentially dist_real_def by auto then have False using fxy and \e>0\ by auto } then show ?lhs unfolding uniformly_continuous_on_def by blast qed subsection \Continuity on a Compact Domain Implies Uniform Continuity\ text\From the proof of the Heine-Borel theorem: Lemma 2 in section 3.7, page 69 of J. C. Burkill and H. Burkill. A Second Course in Mathematical Analysis (CUP, 2002)\ lemma Heine_Borel_lemma: assumes "compact S" and Ssub: "S \ \\" and opn: "\G. G \ \ \ open G" obtains e where "0 < e" "\x. x \ S \ \G \ \. ball x e \ G" proof - have False if neg: "\e. 0 < e \ \x \ S. \G \ \. \ ball x e \ G" proof - have "\x \ S. \G \ \. \ ball x (1 / Suc n) \ G" for n using neg by simp then obtain f where "\n. f n \ S" and fG: "\G n. G \ \ \ \ ball (f n) (1 / Suc n) \ G" by metis then obtain l r where "l \ S" "strict_mono r" and to_l: "(f \ r) \ l" using \compact S\ compact_def that by metis then obtain G where "l \ G" "G \ \" using Ssub by auto then obtain e where "0 < e" and e: "\z. dist z l < e \ z \ G" using opn open_dist by blast obtain N1 where N1: "\n. n \ N1 \ dist (f (r n)) l < e/2" using to_l apply (simp add: lim_sequentially) using \0 < e\ half_gt_zero that by blast obtain N2 where N2: "of_nat N2 > 2/e" using reals_Archimedean2 by blast obtain x where "x \ ball (f (r (max N1 N2))) (1 / real (Suc (r (max N1 N2))))" and "x \ G" using fG [OF \G \ \\, of "r (max N1 N2)"] by blast then have "dist (f (r (max N1 N2))) x < 1 / real (Suc (r (max N1 N2)))" by simp also have "... \ 1 / real (Suc (max N1 N2))" by (meson Suc_le_mono \strict_mono r\ inverse_of_nat_le nat.discI seq_suble) also have "... \ 1 / real (Suc N2)" by (simp add: field_simps) also have "... < e/2" using N2 \0 < e\ by (simp add: field_simps) finally have "dist (f (r (max N1 N2))) x < e/2" . moreover have "dist (f (r (max N1 N2))) l < e/2" using N1 max.cobounded1 by blast ultimately have "dist x l < e" by metric then show ?thesis using e \x \ G\ by blast qed then show ?thesis by (meson that) qed lemma compact_uniformly_equicontinuous: assumes "compact S" and cont: "\x e. \x \ S; 0 < e\ \ \d. 0 < d \ (\f \ \. \x' \ S. dist x' x < d \ dist (f x') (f x) < e)" and "0 < e" obtains d where "0 < d" "\f x x'. \f \ \; x \ S; x' \ S; dist x' x < d\ \ dist (f x') (f x) < e" proof - obtain d where d_pos: "\x e. \x \ S; 0 < e\ \ 0 < d x e" and d_dist : "\x x' e f. \dist x' x < d x e; x \ S; x' \ S; 0 < e; f \ \\ \ dist (f x') (f x) < e" using cont by metis let ?\ = "((\x. ball x (d x (e/2))) ` S)" have Ssub: "S \ \ ?\" using \0 < e\ d_pos by auto then obtain k where "0 < k" and k: "\x. x \ S \ \G \ ?\. ball x k \ G" by (rule Heine_Borel_lemma [OF \compact S\]) auto moreover have "dist (f v) (f u) < e" if "f \ \" "u \ S" "v \ S" "dist v u < k" for f u v proof - obtain G where "G \ ?\" "u \ G" "v \ G" using k that by (metis \dist v u < k\ \u \ S\ \0 < k\ centre_in_ball subsetD dist_commute mem_ball) then obtain w where w: "dist w u < d w (e/2)" "dist w v < d w (e/2)" "w \ S" by auto with that d_dist have "dist (f w) (f v) < e/2" by (metis \0 < e\ dist_commute half_gt_zero) moreover have "dist (f w) (f u) < e/2" using that d_dist w by (metis \0 < e\ dist_commute divide_pos_pos zero_less_numeral) ultimately show ?thesis using dist_triangle_half_r by blast qed ultimately show ?thesis using that by blast qed corollary compact_uniformly_continuous: fixes f :: "'a :: metric_space \ 'b :: metric_space" assumes f: "continuous_on S f" and S: "compact S" shows "uniformly_continuous_on S f" using f unfolding continuous_on_iff uniformly_continuous_on_def by (force intro: compact_uniformly_equicontinuous [OF S, of "{f}"]) subsection\<^marker>\tag unimportant\\ Theorems relating continuity and uniform continuity to closures\ lemma continuous_on_closure: "continuous_on (closure S) f \ (\x e. x \ closure S \ 0 < e \ (\d. 0 < d \ (\y. y \ S \ dist y x < d \ dist (f y) (f x) < e)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding continuous_on_iff by (metis Un_iff closure_def) next assume R [rule_format]: ?rhs show ?lhs proof fix x and e::real assume "0 < e" and x: "x \ closure S" obtain \::real where "\ > 0" and \: "\y. \y \ S; dist y x < \\ \ dist (f y) (f x) < e/2" using R [of x "e/2"] \0 < e\ x by auto have "dist (f y) (f x) \ e" if y: "y \ closure S" and dyx: "dist y x < \/2" for y proof - obtain \'::real where "\' > 0" and \': "\z. \z \ S; dist z y < \'\ \ dist (f z) (f y) < e/2" using R [of y "e/2"] \0 < e\ y by auto obtain z where "z \ S" and z: "dist z y < min \' \ / 2" using closure_approachable y by (metis \0 < \'\ \0 < \\ divide_pos_pos min_less_iff_conj zero_less_numeral) have "dist (f z) (f y) < e/2" using \' [OF \z \ S\] z \0 < \'\ by metric moreover have "dist (f z) (f x) < e/2" using \[OF \z \ S\] z dyx by metric ultimately show ?thesis by metric qed then show "\d>0. \x'\closure S. dist x' x < d \ dist (f x') (f x) \ e" by (rule_tac x="\/2" in exI) (simp add: \\ > 0\) qed qed lemma continuous_on_closure_sequentially: fixes f :: "'a::metric_space \ 'b :: metric_space" shows "continuous_on (closure S) f \ (\x a. a \ closure S \ (\n. x n \ S) \ x \ a \ (f \ x) \ f a)" (is "?lhs = ?rhs") proof - have "continuous_on (closure S) f \ (\x \ closure S. continuous (at x within S) f)" by (force simp: continuous_on_closure continuous_within_eps_delta) also have "... = ?rhs" by (force simp: continuous_within_sequentially) finally show ?thesis . qed lemma uniformly_continuous_on_closure: fixes f :: "'a::metric_space \ 'b::metric_space" assumes ucont: "uniformly_continuous_on S f" and cont: "continuous_on (closure S) f" shows "uniformly_continuous_on (closure S) f" unfolding uniformly_continuous_on_def proof (intro allI impI) fix e::real assume "0 < e" then obtain d::real where "d>0" and d: "\x x'. \x\S; x'\S; dist x' x < d\ \ dist (f x') (f x) < e/3" using ucont [unfolded uniformly_continuous_on_def, rule_format, of "e/3"] by auto show "\d>0. \x\closure S. \x'\closure S. dist x' x < d \ dist (f x') (f x) < e" proof (rule exI [where x="d/3"], clarsimp simp: \d > 0\) fix x y assume x: "x \ closure S" and y: "y \ closure S" and dyx: "dist y x * 3 < d" obtain d1::real where "d1 > 0" and d1: "\w. \w \ closure S; dist w x < d1\ \ dist (f w) (f x) < e/3" using cont [unfolded continuous_on_iff, rule_format, of "x" "e/3"] \0 < e\ x by auto obtain x' where "x' \ S" and x': "dist x' x < min d1 (d / 3)" using closure_approachable [of x S] by (metis \0 < d1\ \0 < d\ divide_pos_pos min_less_iff_conj x zero_less_numeral) obtain d2::real where "d2 > 0" and d2: "\w \ closure S. dist w y < d2 \ dist (f w) (f y) < e/3" using cont [unfolded continuous_on_iff, rule_format, of "y" "e/3"] \0 < e\ y by auto obtain y' where "y' \ S" and y': "dist y' y < min d2 (d / 3)" using closure_approachable [of y S] by (metis \0 < d2\ \0 < d\ divide_pos_pos min_less_iff_conj y zero_less_numeral) have "dist x' x < d/3" using x' by auto then have "dist x' y' < d" using dyx y' by metric then have "dist (f x') (f y') < e/3" by (rule d [OF \y' \ S\ \x' \ S\]) moreover have "dist (f x') (f x) < e/3" using \x' \ S\ closure_subset x' d1 by (simp add: closure_def) moreover have "dist (f y') (f y) < e/3" using \y' \ S\ closure_subset y' d2 by (simp add: closure_def) ultimately show "dist (f y) (f x) < e" by metric qed qed lemma uniformly_continuous_on_extension_at_closure: fixes f::"'a::metric_space \ 'b::complete_space" assumes uc: "uniformly_continuous_on X f" assumes "x \ closure X" obtains l where "(f \ l) (at x within X)" proof - from assms obtain xs where xs: "xs \ x" "\n. xs n \ X" by (auto simp: closure_sequential) from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF xs] obtain l where l: "(\n. f (xs n)) \ l" by atomize_elim (simp only: convergent_eq_Cauchy) have "(f \ l) (at x within X)" proof (safe intro!: Lim_within_LIMSEQ) fix xs' assume "\n. xs' n \ x \ xs' n \ X" and xs': "xs' \ x" then have "xs' n \ x" "xs' n \ X" for n by auto from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF \xs' \ x\ \xs' _ \ X\] obtain l' where l': "(\n. f (xs' n)) \ l'" by atomize_elim (simp only: convergent_eq_Cauchy) show "(\n. f (xs' n)) \ l" proof (rule tendstoI) fix e::real assume "e > 0" define e' where "e' \ e/2" have "e' > 0" using \e > 0\ by (simp add: e'_def) have "\\<^sub>F n in sequentially. dist (f (xs n)) l < e'" by (simp add: \0 < e'\ l tendstoD) moreover from uc[unfolded uniformly_continuous_on_def, rule_format, OF \e' > 0\] obtain d where d: "d > 0" "\x x'. x \ X \ x' \ X \ dist x x' < d \ dist (f x) (f x') < e'" by auto have "\\<^sub>F n in sequentially. dist (xs n) (xs' n) < d" by (auto intro!: \0 < d\ order_tendstoD tendsto_eq_intros xs xs') ultimately show "\\<^sub>F n in sequentially. dist (f (xs' n)) l < e" proof eventually_elim case (elim n) have "dist (f (xs' n)) l \ dist (f (xs n)) (f (xs' n)) + dist (f (xs n)) l" by metric also have "dist (f (xs n)) (f (xs' n)) < e'" by (auto intro!: d xs \xs' _ \ _\ elim) also note \dist (f (xs n)) l < e'\ also have "e' + e' = e" by (simp add: e'_def) finally show ?case by simp qed qed qed thus ?thesis .. qed lemma uniformly_continuous_on_extension_on_closure: fixes f::"'a::metric_space \ 'b::complete_space" assumes uc: "uniformly_continuous_on X f" obtains g where "uniformly_continuous_on (closure X) g" "\x. x \ X \ f x = g x" "\Y h x. X \ Y \ Y \ closure X \ continuous_on Y h \ (\x. x \ X \ f x = h x) \ x \ Y \ h x = g x" proof - from uc have cont_f: "continuous_on X f" by (simp add: uniformly_continuous_imp_continuous) obtain y where y: "(f \ y x) (at x within X)" if "x \ closure X" for x using uniformly_continuous_on_extension_at_closure[OF assms] by meson let ?g = "\x. if x \ X then f x else y x" have "uniformly_continuous_on (closure X) ?g" unfolding uniformly_continuous_on_def proof safe fix e::real assume "e > 0" define e' where "e' \ e / 3" have "e' > 0" using \e > 0\ by (simp add: e'_def) from uc[unfolded uniformly_continuous_on_def, rule_format, OF \0 < e'\] obtain d where "d > 0" and d: "\x x'. x \ X \ x' \ X \ dist x' x < d \ dist (f x') (f x) < e'" by auto define d' where "d' = d / 3" have "d' > 0" using \d > 0\ by (simp add: d'_def) show "\d>0. \x\closure X. \x'\closure X. dist x' x < d \ dist (?g x') (?g x) < e" proof (safe intro!: exI[where x=d'] \d' > 0\) fix x x' assume x: "x \ closure X" and x': "x' \ closure X" and dist: "dist x' x < d'" then obtain xs xs' where xs: "xs \ x" "\n. xs n \ X" and xs': "xs' \ x'" "\n. xs' n \ X" by (auto simp: closure_sequential) have "\\<^sub>F n in sequentially. dist (xs' n) x' < d'" and "\\<^sub>F n in sequentially. dist (xs n) x < d'" by (auto intro!: \0 < d'\ order_tendstoD tendsto_eq_intros xs xs') moreover have "(\x. f (xs x)) \ y x" if "x \ closure X" "x \ X" "xs \ x" "\n. xs n \ X" for xs x using that not_eventuallyD by (force intro!: filterlim_compose[OF y[OF \x \ closure X\]] simp: filterlim_at) then have "(\x. f (xs' x)) \ ?g x'" "(\x. f (xs x)) \ ?g x" using x x' by (auto intro!: continuous_on_tendsto_compose[OF cont_f] simp: xs' xs) then have "\\<^sub>F n in sequentially. dist (f (xs' n)) (?g x') < e'" "\\<^sub>F n in sequentially. dist (f (xs n)) (?g x) < e'" by (auto intro!: \0 < e'\ order_tendstoD tendsto_eq_intros) ultimately have "\\<^sub>F n in sequentially. dist (?g x') (?g x) < e" proof eventually_elim case (elim n) have "dist (?g x') (?g x) \ dist (f (xs' n)) (?g x') + dist (f (xs' n)) (f (xs n)) + dist (f (xs n)) (?g x)" by (metis add.commute add_le_cancel_left dist_commute dist_triangle dist_triangle_le) also from \dist (xs' n) x' < d'\ \dist x' x < d'\ \dist (xs n) x < d'\ have "dist (xs' n) (xs n) < d" unfolding d'_def by metric with \xs _ \ X\ \xs' _ \ X\ have "dist (f (xs' n)) (f (xs n)) < e'" by (rule d) also note \dist (f (xs' n)) (?g x') < e'\ also note \dist (f (xs n)) (?g x) < e'\ finally show ?case by (simp add: e'_def) qed then show "dist (?g x') (?g x) < e" by simp qed qed moreover have "f x = ?g x" if "x \ X" for x using that by simp moreover { fix Y h x assume Y: "x \ Y" "X \ Y" "Y \ closure X" and cont_h: "continuous_on Y h" and extension: "(\x. x \ X \ f x = h x)" { assume "x \ X" have "x \ closure X" using Y by auto then obtain xs where xs: "xs \ x" "\n. xs n \ X" by (auto simp: closure_sequential) from continuous_on_tendsto_compose[OF cont_h xs(1)] xs(2) Y have hx: "(\x. f (xs x)) \ h x" by (auto simp: subsetD extension) then have "(\x. f (xs x)) \ y x" using \x \ X\ not_eventuallyD xs(2) by (force intro!: filterlim_compose[OF y[OF \x \ closure X\]] simp: filterlim_at xs) with hx have "h x = y x" by (rule LIMSEQ_unique) } then have "h x = ?g x" using extension by auto } ultimately show ?thesis .. qed lemma bounded_uniformly_continuous_image: fixes f :: "'a :: heine_borel \ 'b :: heine_borel" assumes "uniformly_continuous_on S f" "bounded S" shows "bounded(f ` S)" by (metis (no_types, lifting) assms bounded_closure_image compact_closure compact_continuous_image compact_eq_bounded_closed image_cong uniformly_continuous_imp_continuous uniformly_continuous_on_extension_on_closure) subsection \With Abstract Topology (TODO: move and remove dependency?)\ lemma openin_contains_ball: "openin (top_of_set T) S \ S \ T \ (\x \ S. \e. 0 < e \ ball x e \ T \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis IntD2 Int_commute Int_lower1 Int_mono inf.idem openE openin_open) next assume ?rhs then show ?lhs by (smt (verit) open_ball Int_commute Int_iff centre_in_ball in_mono openin_open_Int openin_subopen) qed lemma openin_contains_cball: "openin (top_of_set T) S \ S \ T \ (\x \ S. \e. 0 < e \ cball x e \ T \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (force simp add: openin_contains_ball intro: exI [where x="_/2"]) next assume ?rhs then show ?lhs by (force simp add: openin_contains_ball) qed subsection \Closed Nest\ text \Bounded closed nest property (proof does not use Heine-Borel)\ lemma bounded_closed_nest: fixes S :: "nat \ ('a::heine_borel) set" assumes "\n. closed (S n)" and "\n. S n \ {}" and "\m n. m \ n \ S n \ S m" and "bounded (S 0)" obtains a where "\n. a \ S n" proof - from assms(2) obtain x where x: "\n. x n \ S n" using choice[of "\n x. x \ S n"] by auto from assms(4,1) have "seq_compact (S 0)" by (simp add: bounded_closed_imp_seq_compact) then obtain l r where lr: "l \ S 0" "strict_mono r" "(x \ r) \ l" using x and assms(3) unfolding seq_compact_def by blast have "\n. l \ S n" proof fix n :: nat have "closed (S n)" using assms(1) by simp moreover have "\i. (x \ r) i \ S i" using x and assms(3) and lr(2) [THEN seq_suble] by auto then have "\i. (x \ r) (i + n) \ S n" using assms(3) by (fast intro!: le_add2) moreover have "(\i. (x \ r) (i + n)) \ l" using lr(3) by (rule LIMSEQ_ignore_initial_segment) ultimately show "l \ S n" by (rule closed_sequentially) qed then show ?thesis using that by blast qed text \Decreasing case does not even need compactness, just completeness.\ lemma decreasing_closed_nest: fixes S :: "nat \ ('a::complete_space) set" assumes "\n. closed (S n)" "\n. S n \ {}" "\m n. m \ n \ S n \ S m" "\e. e>0 \ \n. \x\S n. \y\S n. dist x y < e" obtains a where "\n. a \ S n" proof - obtain t where t: "\n. t n \ S n" by (meson assms(2) equals0I) { fix e :: real assume "e > 0" then obtain N where N: "\x\S N. \y\S N. dist x y < e" using assms(4) by blast { fix m n :: nat assume "N \ m \ N \ n" then have "t m \ S N" "t n \ S N" using assms(3) t unfolding subset_eq t by blast+ then have "dist (t m) (t n) < e" using N by auto } then have "\N. \m n. N \ m \ N \ n \ dist (t m) (t n) < e" by auto } then have "Cauchy t" unfolding cauchy_def by auto then obtain l where l:"(t \ l) sequentially" using complete_UNIV unfolding complete_def by auto { fix n :: nat { fix e :: real assume "e > 0" then obtain N :: nat where N: "\n\N. dist (t n) l < e" using l[unfolded lim_sequentially] by auto have "t (max n N) \ S n" by (meson assms(3) contra_subsetD max.cobounded1 t) then have "\y\S n. dist y l < e" using N max.cobounded2 by blast } then have "l \ S n" using closed_approachable[of "S n" l] assms(1) by auto } then show ?thesis using that by blast qed text \Strengthen it to the intersection actually being a singleton.\ lemma decreasing_closed_nest_sing: fixes S :: "nat \ 'a::complete_space set" assumes "\n. closed(S n)" "\n. S n \ {}" "\m n. m \ n \ S n \ S m" "\e. e>0 \ \n. \x \ (S n). \ y\(S n). dist x y < e" shows "\a. \(range S) = {a}" proof - obtain a where a: "\n. a \ S n" using decreasing_closed_nest[of S] using assms by auto { fix b assume b: "b \ \(range S)" { fix e :: real assume "e > 0" then have "dist a b < e" using assms(4) and b and a by blast } then have "dist a b = 0" by (metis dist_eq_0_iff dist_nz less_le) } with a have "\(range S) = {a}" unfolding image_def by auto then show ?thesis .. qed subsection\<^marker>\tag unimportant\ \Making a continuous function avoid some value in a neighbourhood\ lemma continuous_within_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous (at x within s) f" and "f x \ a" shows "\e>0. \y \ s. dist x y < e --> f y \ a" proof - obtain U where "open U" and "f x \ U" and "a \ U" using t1_space [OF \f x \ a\] by fast have "(f \ f x) (at x within s)" using assms(1) by (simp add: continuous_within) then have "eventually (\y. f y \ U) (at x within s)" using \open U\ and \f x \ U\ unfolding tendsto_def by fast then have "eventually (\y. f y \ a) (at x within s)" using \a \ U\ by (fast elim: eventually_mono) then show ?thesis using \f x \ a\ by (auto simp: dist_commute eventually_at) qed lemma continuous_at_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous (at x) f" and "f x \ a" shows "\e>0. \y. dist x y < e \ f y \ a" using assms continuous_within_avoid[of x UNIV f a] by simp lemma continuous_on_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous_on s f" and "x \ s" and "f x \ a" shows "\e>0. \y \ s. dist x y < e \ f y \ a" using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)] continuous_within_avoid[of x s f a] using assms(3) by auto lemma continuous_on_open_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous_on s f" and "open s" and "x \ s" and "f x \ a" shows "\e>0. \y. dist x y < e \ f y \ a" using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)] using continuous_at_avoid[of x f a] assms(4) by auto subsection \Consequences for Real Numbers\ lemma closed_contains_Inf: fixes S :: "real set" shows "S \ {} \ bdd_below S \ closed S \ Inf S \ S" by (metis closure_contains_Inf closure_closed) lemma closed_subset_contains_Inf: fixes A C :: "real set" shows "closed C \ A \ C \ A \ {} \ bdd_below A \ Inf A \ C" by (metis closure_contains_Inf closure_minimal subset_eq) lemma closed_contains_Sup: fixes S :: "real set" shows "S \ {} \ bdd_above S \ closed S \ Sup S \ S" by (subst closure_closed[symmetric], assumption, rule closure_contains_Sup) lemma closed_subset_contains_Sup: fixes A C :: "real set" shows "closed C \ A \ C \ A \ {} \ bdd_above A \ Sup A \ C" by (metis closure_contains_Sup closure_minimal subset_eq) lemma atLeastAtMost_subset_contains_Inf: fixes A :: "real set" and a b :: real shows "A \ {} \ a \ b \ A \ {a..b} \ Inf A \ {a..b}" by (rule closed_subset_contains_Inf) (auto intro: closed_real_atLeastAtMost intro!: bdd_belowI[of A a]) lemma bounded_real: "bounded (S::real set) \ (\a. \x\S. \x\ \ a)" by (simp add: bounded_iff) lemma bounded_imp_bdd_above: "bounded S \ bdd_above (S :: real set)" by (auto simp: bounded_def bdd_above_def dist_real_def) (metis abs_le_D1 abs_minus_commute diff_le_eq) lemma bounded_imp_bdd_below: "bounded S \ bdd_below (S :: real set)" by (auto simp: bounded_def bdd_below_def dist_real_def) (metis abs_le_D1 add.commute diff_le_eq) lemma bounded_has_Sup: fixes S :: "real set" assumes "bounded S" and "S \ {}" shows "\x\S. x \ Sup S" and "\b. (\x\S. x \ b) \ Sup S \ b" proof show "\b. (\x\S. x \ b) \ Sup S \ b" using assms by (metis cSup_least) qed (metis cSup_upper assms(1) bounded_imp_bdd_above) lemma Sup_insert: fixes S :: "real set" shows "bounded S \ Sup (insert x S) = (if S = {} then x else max x (Sup S))" by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If) lemma bounded_has_Inf: fixes S :: "real set" assumes "bounded S" and "S \ {}" shows "\x\S. x \ Inf S" and "\b. (\x\S. x \ b) \ Inf S \ b" proof show "\b. (\x\S. x \ b) \ Inf S \ b" using assms by (metis cInf_greatest) qed (metis cInf_lower assms(1) bounded_imp_bdd_below) lemma Inf_insert: fixes S :: "real set" shows "bounded S \ Inf (insert x S) = (if S = {} then x else min x (Inf S))" by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If) lemma open_real: fixes s :: "real set" shows "open s \ (\x \ s. \e>0. \x'. \x' - x\ < e --> x' \ s)" unfolding open_dist dist_norm by simp lemma islimpt_approachable_real: fixes s :: "real set" shows "x islimpt s \ (\e>0. \x'\ s. x' \ x \ \x' - x\ < e)" unfolding islimpt_approachable dist_norm by simp lemma closed_real: fixes s :: "real set" shows "closed s \ (\x. (\e>0. \x' \ s. x' \ x \ \x' - x\ < e) \ x \ s)" unfolding closed_limpt islimpt_approachable dist_norm by simp lemma continuous_at_real_range: fixes f :: "'a::real_normed_vector \ real" shows "continuous (at x) f \ (\e>0. \d>0. \x'. norm(x' - x) < d --> \f x' - f x\ < e)" by (metis (mono_tags, opaque_lifting) LIM_eq continuous_within norm_eq_zero real_norm_def right_minus_eq) lemma continuous_on_real_range: fixes f :: "'a::real_normed_vector \ real" shows "continuous_on s f \ (\x \ s. \e>0. \d>0. (\x' \ s. norm(x' - x) < d \ \f x' - f x\ < e))" unfolding continuous_on_iff dist_norm by simp lemma continuous_on_closed_Collect_le: fixes f g :: "'a::topological_space \ real" assumes f: "continuous_on s f" and g: "continuous_on s g" and s: "closed s" shows "closed {x \ s. f x \ g x}" proof - have "closed ((\x. g x - f x) -` {0..} \ s)" using closed_real_atLeast continuous_on_diff [OF g f] by (simp add: continuous_on_closed_vimage [OF s]) also have "((\x. g x - f x) -` {0..} \ s) = {x\s. f x \ g x}" by auto finally show ?thesis . qed lemma continuous_le_on_closure: fixes a::real assumes f: "continuous_on (closure s) f" and x: "x \ closure(s)" and xlo: "\x. x \ s ==> f(x) \ a" shows "f(x) \ a" using image_closure_subset [OF f, where T=" {x. x \ a}" ] assms continuous_on_closed_Collect_le[of "UNIV" "\x. x" "\x. a"] by auto lemma continuous_ge_on_closure: fixes a::real assumes f: "continuous_on (closure s) f" and x: "x \ closure(s)" and xlo: "\x. x \ s ==> f(x) \ a" shows "f(x) \ a" using image_closure_subset [OF f, where T=" {x. a \ x}"] assms continuous_on_closed_Collect_le[of "UNIV" "\x. a" "\x. x"] by auto subsection\The infimum of the distance between two sets\ definition\<^marker>\tag important\ setdist :: "'a::metric_space set \ 'a set \ real" where "setdist s t \ (if s = {} \ t = {} then 0 else Inf {dist x y| x y. x \ s \ y \ t})" lemma setdist_empty1 [simp]: "setdist {} t = 0" by (simp add: setdist_def) lemma setdist_empty2 [simp]: "setdist t {} = 0" by (simp add: setdist_def) lemma setdist_pos_le [simp]: "0 \ setdist s t" by (auto simp: setdist_def ex_in_conv [symmetric] intro: cInf_greatest) lemma le_setdistI: assumes "s \ {}" "t \ {}" "\x y. \x \ s; y \ t\ \ d \ dist x y" shows "d \ setdist s t" using assms by (auto simp: setdist_def Set.ex_in_conv [symmetric] intro: cInf_greatest) lemma setdist_le_dist: "\x \ s; y \ t\ \ setdist s t \ dist x y" unfolding setdist_def by (auto intro!: bdd_belowI [where m=0] cInf_lower) lemma le_setdist_iff: "d \ setdist S T \ (\x \ S. \y \ T. d \ dist x y) \ (S = {} \ T = {} \ d \ 0)" by (smt (verit) le_setdistI setdist_def setdist_le_dist) lemma setdist_ltE: assumes "setdist S T < b" "S \ {}" "T \ {}" obtains x y where "x \ S" "y \ T" "dist x y < b" using assms by (auto simp: not_le [symmetric] le_setdist_iff) lemma setdist_refl: "setdist S S = 0" by (metis dist_eq_0_iff ex_in_conv order_antisym setdist_def setdist_le_dist setdist_pos_le) lemma setdist_sym: "setdist S T = setdist T S" by (force simp: setdist_def dist_commute intro!: arg_cong [where f=Inf]) lemma setdist_triangle: "setdist S T \ setdist S {a} + setdist {a} T" proof (cases "S = {} \ T = {}") case True then show ?thesis using setdist_pos_le by fastforce next case False then have "\x. x \ S \ setdist S T - dist x a \ setdist {a} T" using dist_self dist_triangle3 empty_not_insert le_setdist_iff setdist_le_dist singleton_iff by (smt (verit, best) dist_self dist_triangle3 empty_not_insert le_setdist_iff setdist_le_dist singleton_iff) then have "setdist S T - setdist {a} T \ setdist S {a}" using False by (fastforce intro: le_setdistI) then show ?thesis by (simp add: algebra_simps) qed lemma setdist_singletons [simp]: "setdist {x} {y} = dist x y" by (simp add: setdist_def) lemma setdist_Lipschitz: "\setdist {x} S - setdist {y} S\ \ dist x y" apply (subst setdist_singletons [symmetric]) by (metis abs_diff_le_iff diff_le_eq setdist_triangle setdist_sym) lemma continuous_at_setdist [continuous_intros]: "continuous (at x) (\y. (setdist {y} S))" by (force simp: continuous_at_eps_delta dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) lemma continuous_on_setdist [continuous_intros]: "continuous_on T (\y. (setdist {y} S))" by (metis continuous_at_setdist continuous_at_imp_continuous_on) lemma uniformly_continuous_on_setdist: "uniformly_continuous_on T (\y. (setdist {y} S))" by (force simp: uniformly_continuous_on_def dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) lemma setdist_subset_right: "\T \ {}; T \ u\ \ setdist S u \ setdist S T" by (smt (verit, best) in_mono le_setdist_iff) lemma setdist_subset_left: "\S \ {}; S \ T\ \ setdist T u \ setdist S u" by (metis setdist_subset_right setdist_sym) lemma setdist_closure_1 [simp]: "setdist (closure S) T = setdist S T" proof (cases "S = {} \ T = {}") case True then show ?thesis by force next case False { fix y assume "y \ T" have "continuous_on (closure S) (\a. dist a y)" by (auto simp: continuous_intros dist_norm) then have *: "\x. x \ closure S \ setdist S T \ dist x y" by (fast intro: setdist_le_dist \y \ T\ continuous_ge_on_closure) } then show ?thesis by (metis False antisym closure_eq_empty closure_subset le_setdist_iff setdist_subset_left) qed lemma setdist_closure_2 [simp]: "setdist T (closure S) = setdist T S" by (metis setdist_closure_1 setdist_sym) lemma setdist_eq_0I: "\x \ S; x \ T\ \ setdist S T = 0" by (metis antisym dist_self setdist_le_dist setdist_pos_le) lemma setdist_unique: "\a \ S; b \ T; \x y. x \ S \ y \ T ==> dist a b \ dist x y\ \ setdist S T = dist a b" by (force simp add: setdist_le_dist le_setdist_iff intro: antisym) lemma setdist_le_sing: "x \ S ==> setdist S T \ setdist {x} T" using setdist_subset_left by auto lemma infdist_eq_setdist: "infdist x A = setdist {x} A" by (simp add: infdist_def setdist_def Setcompr_eq_image) lemma setdist_eq_infdist: "setdist A B = (if A = {} then 0 else INF a\A. infdist a B)" proof - have "Inf {dist x y |x y. x \ A \ y \ B} = (INF x\A. Inf (dist x ` B))" if "b \ B" "a \ A" for a b proof (rule order_antisym) have "Inf {dist x y |x y. x \ A \ y \ B} \ Inf (dist x ` B)" if "b \ B" "a \ A" "x \ A" for x proof - have "\b'. b' \ B \ Inf {dist x y |x y. x \ A \ y \ B} \ dist x b'" by (metis (mono_tags, lifting) ex_in_conv setdist_def setdist_le_dist that(3)) then show ?thesis by (smt (verit) cINF_greatest ex_in_conv \b \ B\) qed then show "Inf {dist x y |x y. x \ A \ y \ B} \ (INF x\A. Inf (dist x ` B))" by (metis (mono_tags, lifting) cINF_greatest emptyE that) next have *: "\x y. \b \ B; a \ A; x \ A; y \ B\ \ \a\A. Inf (dist a ` B) \ dist x y" by (meson bdd_below_image_dist cINF_lower) show "(INF x\A. Inf (dist x ` B)) \ Inf {dist x y |x y. x \ A \ y \ B}" proof (rule conditionally_complete_lattice_class.cInf_mono) show "bdd_below ((\x. Inf (dist x ` B)) ` A)" by (metis (no_types, lifting) bdd_belowI2 ex_in_conv infdist_def infdist_nonneg that(1)) qed (use that in \auto simp: *\) qed then show ?thesis by (auto simp: setdist_def infdist_def) qed lemma infdist_mono: assumes "A \ B" "A \ {}" shows "infdist x B \ infdist x A" by (simp add: assms infdist_eq_setdist setdist_subset_right) lemma infdist_singleton [simp]: "infdist x {y} = dist x y" by (simp add: infdist_eq_setdist) proposition setdist_attains_inf: assumes "compact B" "B \ {}" obtains y where "y \ B" "setdist A B = infdist y A" proof (cases "A = {}") case True then show thesis by (metis assms diameter_compact_attained infdist_def setdist_def that) next case False obtain y where "y \ B" and min: "\y'. y' \ B \ infdist y A \ infdist y' A" by (metis continuous_attains_inf [OF assms continuous_on_infdist] continuous_on_id) show thesis proof have "setdist A B = (INF y\B. infdist y A)" by (metis \B \ {}\ setdist_eq_infdist setdist_sym) also have "\ = infdist y A" proof (rule order_antisym) show "(INF y\B. infdist y A) \ infdist y A" by (meson \y \ B\ bdd_belowI2 cInf_lower image_eqI infdist_nonneg) next show "infdist y A \ (INF y\B. infdist y A)" by (simp add: \B \ {}\ cINF_greatest min) qed finally show "setdist A B = infdist y A" . qed (fact \y \ B\) qed end diff --git a/src/HOL/Analysis/Locally.thy b/src/HOL/Analysis/Locally.thy --- a/src/HOL/Analysis/Locally.thy +++ b/src/HOL/Analysis/Locally.thy @@ -1,451 +1,972 @@ section \Neighbourhood bases and Locally path-connected spaces\ theory Locally imports - Path_Connected Function_Topology + Path_Connected Function_Topology Sum_Topology begin subsection\Neighbourhood Bases\ text \Useful for "local" properties of various kinds\ definition neighbourhood_base_at where "neighbourhood_base_at x P X \ \W. openin X W \ x \ W \ (\U V. openin X U \ P V \ x \ U \ U \ V \ V \ W)" definition neighbourhood_base_of where "neighbourhood_base_of P X \ \x \ topspace X. neighbourhood_base_at x P X" lemma neighbourhood_base_of: "neighbourhood_base_of P X \ (\W x. openin X W \ x \ W \ (\U V. openin X U \ P V \ x \ U \ U \ V \ V \ W))" unfolding neighbourhood_base_at_def neighbourhood_base_of_def using openin_subset by blast lemma neighbourhood_base_at_mono: "\neighbourhood_base_at x P X; \S. \P S; x \ S\ \ Q S\ \ neighbourhood_base_at x Q X" unfolding neighbourhood_base_at_def by (meson subset_eq) lemma neighbourhood_base_of_mono: "\neighbourhood_base_of P X; \S. P S \ Q S\ \ neighbourhood_base_of Q X" unfolding neighbourhood_base_of_def using neighbourhood_base_at_mono by force lemma open_neighbourhood_base_at: "(\S. \P S; x \ S\ \ openin X S) \ neighbourhood_base_at x P X \ (\W. openin X W \ x \ W \ (\U. P U \ x \ U \ U \ W))" unfolding neighbourhood_base_at_def by (meson subsetD) lemma open_neighbourhood_base_of: "(\S. P S \ openin X S) \ neighbourhood_base_of P X \ (\W x. openin X W \ x \ W \ (\U. P U \ x \ U \ U \ W))" apply (simp add: neighbourhood_base_of, safe, blast) by meson lemma neighbourhood_base_of_open_subset: "\neighbourhood_base_of P X; openin X S\ \ neighbourhood_base_of P (subtopology X S)" apply (clarsimp simp add: neighbourhood_base_of openin_subtopology_alt image_def) apply (rename_tac x V) apply (drule_tac x="S \ V" in spec) apply (simp add: Int_ac) by (metis IntI le_infI1 openin_Int) lemma neighbourhood_base_of_topology_base: "openin X = arbitrary union_of (\W. W \ \) \ neighbourhood_base_of P X \ (\W x. W \ \ \ x \ W \ (\U V. openin X U \ P V \ x \ U \ U \ V \ V \ W))" apply (auto simp: openin_topology_base_unique neighbourhood_base_of) by (meson subset_trans) lemma neighbourhood_base_at_unlocalized: assumes "\S T. \P S; openin X T; x \ T; T \ S\ \ P T" shows "neighbourhood_base_at x P X \ (x \ topspace X \ (\U V. openin X U \ P V \ x \ U \ U \ V \ V \ topspace X))" (is "?lhs = ?rhs") proof assume R: ?rhs show ?lhs unfolding neighbourhood_base_at_def proof clarify fix W assume "openin X W" "x \ W" then have "x \ topspace X" using openin_subset by blast with R obtain U V where "openin X U" "P V" "x \ U" "U \ V" "V \ topspace X" by blast then show "\U V. openin X U \ P V \ x \ U \ U \ V \ V \ W" by (metis IntI \openin X W\ \x \ W\ assms inf_le1 inf_le2 openin_Int) qed qed (simp add: neighbourhood_base_at_def) lemma neighbourhood_base_at_with_subset: "\openin X U; x \ U\ \ (neighbourhood_base_at x P X \ neighbourhood_base_at x (\T. T \ U \ P T) X)" apply (simp add: neighbourhood_base_at_def) apply (metis IntI Int_subset_iff openin_Int) done lemma neighbourhood_base_of_with_subset: "neighbourhood_base_of P X \ neighbourhood_base_of (\t. t \ topspace X \ P t) X" using neighbourhood_base_at_with_subset by (fastforce simp add: neighbourhood_base_of_def) subsection\Locally path-connected spaces\ definition weakly_locally_path_connected_at where "weakly_locally_path_connected_at x X \ neighbourhood_base_at x (path_connectedin X) X" definition locally_path_connected_at where "locally_path_connected_at x X \ neighbourhood_base_at x (\U. openin X U \ path_connectedin X U) X" definition locally_path_connected_space where "locally_path_connected_space X \ neighbourhood_base_of (path_connectedin X) X" lemma locally_path_connected_space_alt: "locally_path_connected_space X \ neighbourhood_base_of (\U. openin X U \ path_connectedin X U) X" (is "?P = ?Q") and locally_path_connected_space_eq_open_path_component_of: "locally_path_connected_space X \ (\U x. openin X U \ x \ U \ openin X (Collect (path_component_of (subtopology X U) x)))" (is "?P = ?R") proof - have ?P if ?Q using locally_path_connected_space_def neighbourhood_base_of_mono that by auto moreover have ?R if P: ?P proof - show ?thesis proof clarify fix U y assume "openin X U" "y \ U" have "\T. openin X T \ x \ T \ T \ Collect (path_component_of (subtopology X U) y)" if "path_component_of (subtopology X U) y x" for x proof - have "x \ U" using path_component_of_equiv that topspace_subtopology by fastforce then have "\Ua. openin X Ua \ (\V. path_connectedin X V \ x \ Ua \ Ua \ V \ V \ U)" using P by (simp add: \openin X U\ locally_path_connected_space_def neighbourhood_base_of) then show ?thesis by (metis dual_order.trans path_component_of_equiv path_component_of_maximal path_connectedin_subtopology subset_iff that) qed then show "openin X (Collect (path_component_of (subtopology X U) y))" using openin_subopen by force qed qed moreover have ?Q if ?R using that apply (simp add: open_neighbourhood_base_of) by (metis mem_Collect_eq openin_subset path_component_of_refl path_connectedin_path_component_of path_connectedin_subtopology that topspace_subtopology_subset) ultimately show "?P = ?Q" "?P = ?R" by blast+ qed lemma locally_path_connected_space: "locally_path_connected_space X \ (\V x. openin X V \ x \ V \ (\U. openin X U \ path_connectedin X U \ x \ U \ U \ V))" by (simp add: locally_path_connected_space_alt open_neighbourhood_base_of) lemma locally_path_connected_space_open_path_components: "locally_path_connected_space X \ (\U c. openin X U \ c \ path_components_of(subtopology X U) \ openin X c)" apply (auto simp: locally_path_connected_space_eq_open_path_component_of path_components_of_def) by (metis imageI inf.absorb_iff2 openin_closedin_eq) lemma openin_path_component_of_locally_path_connected_space: "locally_path_connected_space X \ openin X (Collect (path_component_of X x))" apply (auto simp: locally_path_connected_space_eq_open_path_component_of) by (metis openin_empty openin_topspace path_component_of_eq_empty subtopology_topspace) lemma openin_path_components_of_locally_path_connected_space: "\locally_path_connected_space X; c \ path_components_of X\ \ openin X c" apply (auto simp: locally_path_connected_space_eq_open_path_component_of) by (metis (no_types, lifting) imageE openin_topspace path_components_of_def subtopology_topspace) lemma closedin_path_components_of_locally_path_connected_space: "\locally_path_connected_space X; c \ path_components_of X\ \ closedin X c" by (simp add: closedin_def complement_path_components_of_Union openin_path_components_of_locally_path_connected_space openin_clauses(3) path_components_of_subset) lemma closedin_path_component_of_locally_path_connected_space: assumes "locally_path_connected_space X" shows "closedin X (Collect (path_component_of X x))" proof (cases "x \ topspace X") case True then show ?thesis by (simp add: assms closedin_path_components_of_locally_path_connected_space path_component_in_path_components_of) next case False then show ?thesis by (metis closedin_empty path_component_of_eq_empty) qed lemma weakly_locally_path_connected_at: "weakly_locally_path_connected_at x X \ (\V. openin X V \ x \ V \ (\U. openin X U \ x \ U \ U \ V \ (\y \ U. \C. path_connectedin X C \ C \ V \ x \ C \ y \ C)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: weakly_locally_path_connected_at_def neighbourhood_base_at_def) by (meson order_trans subsetD) next have *: "\V. path_connectedin X V \ U \ V \ V \ W" if "(\y\U. \C. path_connectedin X C \ C \ W \ x \ C \ y \ C)" for W U proof (intro exI conjI) let ?V = "(Collect (path_component_of (subtopology X W) x))" show "path_connectedin X (Collect (path_component_of (subtopology X W) x))" by (meson path_connectedin_path_component_of path_connectedin_subtopology) show "U \ ?V" by (auto simp: path_component_of path_connectedin_subtopology that) show "?V \ W" by (meson path_connectedin_path_component_of path_connectedin_subtopology) qed assume ?rhs then show ?lhs unfolding weakly_locally_path_connected_at_def neighbourhood_base_at_def by (metis "*") qed lemma locally_path_connected_space_im_kleinen: "locally_path_connected_space X \ (\V x. openin X V \ x \ V \ (\U. openin X U \ x \ U \ U \ V \ (\y \ U. \c. path_connectedin X c \ c \ V \ x \ c \ y \ c)))" apply (simp add: locally_path_connected_space_def neighbourhood_base_of_def) apply (simp add: weakly_locally_path_connected_at flip: weakly_locally_path_connected_at_def) using openin_subset apply force done lemma locally_path_connected_space_open_subset: "\locally_path_connected_space X; openin X s\ \ locally_path_connected_space (subtopology X s)" apply (simp add: locally_path_connected_space_def neighbourhood_base_of openin_open_subtopology path_connectedin_subtopology) by (meson order_trans) lemma locally_path_connected_space_quotient_map_image: assumes f: "quotient_map X Y f" and X: "locally_path_connected_space X" shows "locally_path_connected_space Y" unfolding locally_path_connected_space_open_path_components proof clarify fix V C assume V: "openin Y V" and c: "C \ path_components_of (subtopology Y V)" then have sub: "C \ topspace Y" using path_connectedin_path_components_of path_connectedin_subset_topspace path_connectedin_subtopology by blast have "openin X {x \ topspace X. f x \ C}" proof (subst openin_subopen, clarify) fix x assume x: "x \ topspace X" and "f x \ C" let ?T = "Collect (path_component_of (subtopology X {z \ topspace X. f z \ V}) x)" show "\T. openin X T \ x \ T \ T \ {x \ topspace X. f x \ C}" proof (intro exI conjI) have "\U. openin X U \ ?T \ path_components_of (subtopology X U)" proof (intro exI conjI) show "openin X ({z \ topspace X. f z \ V})" using V f openin_subset quotient_map_def by fastforce show "Collect (path_component_of (subtopology X {z \ topspace X. f z \ V}) x) \ path_components_of (subtopology X {z \ topspace X. f z \ V})" by (metis (no_types, lifting) Int_iff \f x \ C\ c mem_Collect_eq path_component_in_path_components_of path_components_of_subset topspace_subtopology topspace_subtopology_subset x) qed with X show "openin X ?T" using locally_path_connected_space_open_path_components by blast show x: "x \ ?T" using V \f x \ C\ c openin_subset path_component_of_equiv path_components_of_subset topspace_subtopology topspace_subtopology_subset x by fastforce have "f ` ?T \ C" proof (rule path_components_of_maximal) show "C \ path_components_of (subtopology Y V)" by (simp add: c) show "path_connectedin (subtopology Y V) (f ` ?T)" proof - have "quotient_map (subtopology X {a \ topspace X. f a \ V}) (subtopology Y V) f" by (simp add: V f quotient_map_restriction) then show ?thesis by (meson path_connectedin_continuous_map_image path_connectedin_path_component_of quotient_imp_continuous_map) qed show "\ disjnt C (f ` ?T)" by (metis (no_types, lifting) \f x \ C\ x disjnt_iff image_eqI) qed then show "?T \ {x \ topspace X. f x \ C}" by (force simp: path_component_of_equiv) qed qed then show "openin Y C" using f sub by (simp add: quotient_map_def) qed lemma homeomorphic_locally_path_connected_space: assumes "X homeomorphic_space Y" shows "locally_path_connected_space X \ locally_path_connected_space Y" proof - obtain f g where "homeomorphic_maps X Y f g" using assms homeomorphic_space_def by fastforce then show ?thesis by (metis (no_types) homeomorphic_map_def homeomorphic_maps_map locally_path_connected_space_quotient_map_image) qed lemma locally_path_connected_space_retraction_map_image: "\retraction_map X Y r; locally_path_connected_space X\ \ locally_path_connected_space Y" using Abstract_Topology.retraction_imp_quotient_map locally_path_connected_space_quotient_map_image by blast lemma locally_path_connected_space_euclideanreal: "locally_path_connected_space euclideanreal" unfolding locally_path_connected_space_def neighbourhood_base_of proof clarsimp fix W and x :: "real" assume "open W" "x \ W" then obtain e where "e > 0" and e: "\x'. \x' - x\ < e \ x' \ W" by (auto simp: open_real) then show "\U. open U \ (\V. path_connected V \ x \ U \ U \ V \ V \ W)" by (force intro!: convex_imp_path_connected exI [where x = "{x-e<.. topspace X") case True then show ?thesis using connectedin_connected_component_of [of X x] apply (clarsimp simp add: connectedin_def connected_space_clopen_in topspace_subtopology_subset cong: conj_cong) apply (drule_tac x="path_component_of_set X x" in spec) by (metis assms closedin_closed_subtopology closedin_connected_component_of closedin_path_component_of_locally_path_connected_space inf.absorb_iff2 inf.orderE openin_path_component_of_locally_path_connected_space openin_subtopology path_component_of_eq_empty path_component_subset_connected_component_of) next case False then show ?thesis using connected_component_of_eq_empty path_component_of_eq_empty by fastforce qed lemma path_components_eq_connected_components_of: "locally_path_connected_space X \ (path_components_of X = connected_components_of X)" by (simp add: path_components_of_def connected_components_of_def image_def path_component_eq_connected_component_of) lemma path_connected_eq_connected_space: "locally_path_connected_space X \ path_connected_space X \ connected_space X" by (metis connected_components_of_subset_sing path_components_eq_connected_components_of path_components_of_subset_singleton) lemma locally_path_connected_space_product_topology: "locally_path_connected_space(product_topology X I) \ topspace(product_topology X I) = {} \ finite {i. i \ I \ ~path_connected_space(X i)} \ (\i \ I. locally_path_connected_space(X i))" (is "?lhs \ ?empty \ ?rhs") proof (cases ?empty) case True then show ?thesis by (simp add: locally_path_connected_space_def neighbourhood_base_of openin_closedin_eq) next case False then obtain z where z: "z \ (\\<^sub>E i\I. topspace (X i))" by auto have ?rhs if L: ?lhs proof - obtain U C where U: "openin (product_topology X I) U" and V: "path_connectedin (product_topology X I) C" and "z \ U" "U \ C" and Csub: "C \ (\\<^sub>E i\I. topspace (X i))" using L apply (clarsimp simp add: locally_path_connected_space_def neighbourhood_base_of) by (metis openin_topspace topspace_product_topology z) then obtain V where finV: "finite {i \ I. V i \ topspace (X i)}" and XV: "\i. i\I \ openin (X i) (V i)" and "z \ Pi\<^sub>E I V" and subU: "Pi\<^sub>E I V \ U" by (force simp: openin_product_topology_alt) show ?thesis proof (intro conjI ballI) have "path_connected_space (X i)" if "i \ I" "V i = topspace (X i)" for i proof - have pc: "path_connectedin (X i) ((\x. x i) ` C)" apply (rule path_connectedin_continuous_map_image [OF _ V]) by (simp add: continuous_map_product_projection \i \ I\) moreover have "((\x. x i) ` C) = topspace (X i)" proof show "(\x. x i) ` C \ topspace (X i)" by (simp add: pc path_connectedin_subset_topspace) have "V i \ (\x. x i) ` (\\<^sub>E i\I. V i)" by (metis \z \ Pi\<^sub>E I V\ empty_iff image_projection_PiE order_refl that(1)) also have "\ \ (\x. x i) ` U" using subU by blast finally show "topspace (X i) \ (\x. x i) ` C" using \U \ C\ that by blast qed ultimately show ?thesis by (simp add: path_connectedin_topspace) qed then have "{i \ I. \ path_connected_space (X i)} \ {i \ I. V i \ topspace (X i)}" by blast with finV show "finite {i \ I. \ path_connected_space (X i)}" using finite_subset by blast next show "locally_path_connected_space (X i)" if "i \ I" for i apply (rule locally_path_connected_space_quotient_map_image [OF _ L, where f = "\x. x i"]) by (metis False Abstract_Topology.retraction_imp_quotient_map retraction_map_product_projection that) qed qed moreover have ?lhs if R: ?rhs proof (clarsimp simp add: locally_path_connected_space_def neighbourhood_base_of) fix F z assume "openin (product_topology X I) F" and "z \ F" then obtain W where finW: "finite {i \ I. W i \ topspace (X i)}" and opeW: "\i. i \ I \ openin (X i) (W i)" and "z \ Pi\<^sub>E I W" "Pi\<^sub>E I W \ F" by (auto simp: openin_product_topology_alt) have "\i \ I. \U C. openin (X i) U \ path_connectedin (X i) C \ z i \ U \ U \ C \ C \ W i \ (W i = topspace (X i) \ path_connected_space (X i) \ U = topspace (X i) \ C = topspace (X i))" (is "\i \ I. ?\ i") proof fix i assume "i \ I" have "locally_path_connected_space (X i)" by (simp add: R \i \ I\) moreover have "openin (X i) (W i) " "z i \ W i" using \z \ Pi\<^sub>E I W\ opeW \i \ I\ by auto ultimately obtain U C where UC: "openin (X i) U" "path_connectedin (X i) C" "z i \ U" "U \ C" "C \ W i" using \i \ I\ by (force simp: locally_path_connected_space_def neighbourhood_base_of) show "?\ i" proof (cases "W i = topspace (X i) \ path_connected_space(X i)") case True then show ?thesis using \z i \ W i\ path_connectedin_topspace by blast next case False then show ?thesis by (meson UC) qed qed then obtain U C where *: "\i. i \ I \ openin (X i) (U i) \ path_connectedin (X i) (C i) \ z i \ (U i) \ (U i) \ (C i) \ (C i) \ W i \ (W i = topspace (X i) \ path_connected_space (X i) \ (U i) = topspace (X i) \ (C i) = topspace (X i))" by metis let ?A = "{i \ I. \ path_connected_space (X i)} \ {i \ I. W i \ topspace (X i)}" have "{i \ I. U i \ topspace (X i)} \ ?A" by (clarsimp simp add: "*") moreover have "finite ?A" by (simp add: that finW) ultimately have "finite {i \ I. U i \ topspace (X i)}" using finite_subset by auto then have "openin (product_topology X I) (Pi\<^sub>E I U)" using * by (simp add: openin_PiE_gen) then show "\U. openin (product_topology X I) U \ (\V. path_connectedin (product_topology X I) V \ z \ U \ U \ V \ V \ F)" apply (rule_tac x="PiE I U" in exI, simp) apply (rule_tac x="PiE I C" in exI) using \z \ Pi\<^sub>E I W\ \Pi\<^sub>E I W \ F\ * apply (simp add: path_connectedin_PiE subset_PiE PiE_iff PiE_mono dual_order.trans) done qed ultimately show ?thesis using False by blast qed +lemma locally_path_connected_is_realinterval: + assumes "is_interval S" + shows "locally_path_connected_space(subtopology euclideanreal S)" + unfolding locally_path_connected_space_def +proof (clarsimp simp add: neighbourhood_base_of openin_subtopology_alt) + fix a U + assume "a \ S" and "a \ U" and "open U" + then obtain r where "r > 0" and r: "ball a r \ U" + by (metis open_contains_ball_eq) + show "\W. open W \ (\V. path_connectedin (top_of_set S) V \ a \ W \ S \ W \ V \ V \ S \ V \ U)" + proof (intro exI conjI) + show "path_connectedin (top_of_set S) (S \ ball a r)" + by (simp add: assms is_interval_Int is_interval_ball_real is_interval_path_connected path_connectedin_subtopology) + show "a \ ball a r" + by (simp add: \0 < r\) + qed (use \0 < r\ r in auto) +qed + +lemma locally_path_connected_real_interval: + "locally_path_connected_space (subtopology euclideanreal{a..b})" + "locally_path_connected_space (subtopology euclideanreal{a<.. + topspace (prod_topology X Y) = {} \ + locally_path_connected_space X \ locally_path_connected_space Y" (is "?lhs=?rhs") +proof (cases "topspace(prod_topology X Y) = {}") + case True + then show ?thesis + by (metis equals0D locally_path_connected_space_def neighbourhood_base_of_def) +next + case False + then have ne: "topspace X \ {}" "topspace Y \ {}" + by simp_all + show ?thesis + proof + assume ?lhs then show ?rhs + by (metis ne locally_path_connected_space_retraction_map_image retraction_map_fst retraction_map_snd) + next + assume ?rhs + with False have X: "locally_path_connected_space X" and Y: "locally_path_connected_space Y" + by auto + show ?lhs + unfolding locally_path_connected_space_def neighbourhood_base_of + proof clarify + fix UV x y + assume UV: "openin (prod_topology X Y) UV" and "(x,y) \ UV" + obtain A B where W12: "openin X A \ openin Y B \ x \ A \ y \ B \ A \ B \ UV" + using X Y by (metis UV \(x,y) \ UV\ openin_prod_topology_alt) + then obtain C D K L + where "openin X C" "path_connectedin X K" "x \ C" "C \ K" "K \ A" + "openin Y D" "path_connectedin Y L" "y \ D" "D \ L" "L \ B" + by (metis X Y locally_path_connected_space) + with W12 \openin X C\ \openin Y D\ + show "\U V. openin (prod_topology X Y) U \ path_connectedin (prod_topology X Y) V \ (x, y) \ U \ U \ V \ V \ UV" + apply (rule_tac x="C \ D" in exI) + apply (rule_tac x="K \ L" in exI) + apply (auto simp: openin_prod_Times_iff path_connectedin_Times) + done + qed + qed +qed + +lemma locally_path_connected_space_sum_topology: + "locally_path_connected_space(sum_topology X I) \ + (\i \ I. locally_path_connected_space (X i))" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + by (smt (verit) homeomorphic_locally_path_connected_space locally_path_connected_space_open_subset topological_property_of_sum_component) +next + assume R: ?rhs + show ?lhs + proof (clarsimp simp add: locally_path_connected_space_def neighbourhood_base_of forall_openin_sum_topology imp_conjL) + fix W i x + assume ope: "\i\I. openin (X i) (W i)" + and "i \ I" and "x \ W i" + then obtain U V where U: "openin (X i) U" and V: "path_connectedin (X i) V" + and "x \ U" "U \ V" "V \ W i" + by (metis R \i \ I\ \x \ W i\ locally_path_connected_space) + show "\U. openin (sum_topology X I) U \ (\V. path_connectedin (sum_topology X I) V \ (i, x) \ U \ U \ V \ V \ Sigma I W)" + proof (intro exI conjI) + show "openin (sum_topology X I) (Pair i ` U)" + by (meson U \i \ I\ open_map_component_injection open_map_def) + show "path_connectedin (sum_topology X I) (Pair i ` V)" + by (meson V \i \ I\ continuous_map_component_injection path_connectedin_continuous_map_image) + show "Pair i ` V \ Sigma I W" + using \V \ W i\ \i \ I\ by force + qed (use \x \ U\ \U \ V\ in auto) + qed +qed + + +subsection\Locally connected spaces\ + +definition weakly_locally_connected_at + where "weakly_locally_connected_at x X \ neighbourhood_base_at x (connectedin X) X" + +definition locally_connected_at + where "locally_connected_at x X \ + neighbourhood_base_at x (\U. openin X U \ connectedin X U ) X" + +definition locally_connected_space + where "locally_connected_space X \ neighbourhood_base_of (connectedin X) X" + + +lemma locally_connected_A: "(\U x. openin X U \ x \ U + \ openin X (connected_component_of_set (subtopology X U) x)) + \ neighbourhood_base_of (\U. openin X U \ connectedin X U) X" + by (smt (verit, best) connected_component_of_refl connectedin_connected_component_of connectedin_subtopology mem_Collect_eq neighbourhood_base_of openin_subset topspace_subtopology_subset) + +lemma locally_connected_B: "locally_connected_space X \ + (\U x. openin X U \ x \ U \ openin X (connected_component_of_set (subtopology X U) x))" + unfolding locally_connected_space_def neighbourhood_base_of + apply (erule all_forward) + apply clarify + apply (subst openin_subopen) + by (smt (verit, ccfv_threshold) Ball_Collect connected_component_of_def connected_component_of_equiv connectedin_subtopology in_mono mem_Collect_eq) + +lemma locally_connected_C: "neighbourhood_base_of (\U. openin X U \ connectedin X U) X \ locally_connected_space X" + using locally_connected_space_def neighbourhood_base_of_mono by auto + + +lemma locally_connected_space_alt: + "locally_connected_space X \ neighbourhood_base_of (\U. openin X U \ connectedin X U) X" + using locally_connected_A locally_connected_B locally_connected_C by blast + +lemma locally_connected_space_eq_open_connected_component_of: + "locally_connected_space X \ + (\U x. openin X U \ x \ U + \ openin X (connected_component_of_set (subtopology X U) x))" + by (meson locally_connected_A locally_connected_B locally_connected_C) + +lemma locally_connected_space: + "locally_connected_space X \ + (\V x. openin X V \ x \ V \ (\U. openin X U \ connectedin X U \ x \ U \ U \ V))" + by (simp add: locally_connected_space_alt open_neighbourhood_base_of) + +lemma locally_path_connected_imp_locally_connected_space: + "locally_path_connected_space X \ locally_connected_space X" + by (simp add: locally_connected_space_def locally_path_connected_space_def neighbourhood_base_of_mono path_connectedin_imp_connectedin) + +lemma locally_connected_space_open_connected_components: + "locally_connected_space X \ + (\U C. openin X U \ C \ connected_components_of(subtopology X U) \ openin X C)" + apply (simp add: locally_connected_space_eq_open_connected_component_of connected_components_of_def) + by (smt (verit) imageE image_eqI inf.orderE inf_commute openin_subset) + +lemma openin_connected_component_of_locally_connected_space: + "locally_connected_space X \ openin X (connected_component_of_set X x)" + by (metis connected_component_of_eq_empty locally_connected_space_eq_open_connected_component_of openin_empty openin_topspace subtopology_topspace) + +lemma openin_connected_components_of_locally_connected_space: + "\locally_connected_space X; C \ connected_components_of X\ \ openin X C" + by (metis locally_connected_space_open_connected_components openin_topspace subtopology_topspace) + +lemma weakly_locally_connected_at: + "weakly_locally_connected_at x X \ + (\V. openin X V \ x \ V + \ (\U. openin X U \ x \ U \ U \ V \ + (\y \ U. \C. connectedin X C \ C \ V \ x \ C \ y \ C)))" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + unfolding neighbourhood_base_at_def weakly_locally_connected_at_def + by (meson subsetD subset_trans) +next + assume R: ?rhs + show ?lhs + unfolding neighbourhood_base_at_def weakly_locally_connected_at_def + proof clarify + fix V + assume "openin X V" and "x \ V" + then obtain U where "openin X U" "x \ U" "U \ V" + and U: "\y\U. \C. connectedin X C \ C \ V \ x \ C \ y \ C" + using R by force + show "\A B. openin X A \ connectedin X B \ x \ A \ A \ B \ B \ V" + proof (intro conjI exI) + show "connectedin X (connected_component_of_set (subtopology X V) x)" + by (meson connectedin_connected_component_of connectedin_subtopology) + show "U \ connected_component_of_set (subtopology X V) x" + using connected_component_of_maximal U + by (simp add: connected_component_of_def connectedin_subtopology subsetI) + show "connected_component_of_set (subtopology X V) x \ V" + using connected_component_of_subset_topspace by fastforce + qed (auto simp: \x \ U\ \openin X U\) + qed +qed + +lemma locally_connected_space_iff_weak: + "locally_connected_space X \ (\x \ topspace X. weakly_locally_connected_at x X)" + by (simp add: locally_connected_space_def neighbourhood_base_of_def weakly_locally_connected_at_def) + +lemma locally_connected_space_im_kleinen: + "locally_connected_space X \ + (\V x. openin X V \ x \ V + \ (\U. openin X U \ x \ U \ U \ V \ + (\y \ U. \C. connectedin X C \ C \ V \ x \ C \ y \ C)))" + unfolding locally_connected_space_iff_weak weakly_locally_connected_at + using openin_subset subsetD by fastforce + +lemma locally_connected_space_open_subset: + "\locally_connected_space X; openin X S\ \ locally_connected_space (subtopology X S)" + apply (simp add: locally_connected_space_def) + by (smt (verit, ccfv_threshold) connectedin_subtopology neighbourhood_base_of openin_open_subtopology subset_trans) + +lemma locally_connected_space_quotient_map_image: + assumes X: "locally_connected_space X" and f: "quotient_map X Y f" + shows "locally_connected_space Y" + unfolding locally_connected_space_open_connected_components +proof clarify + fix V C + assume "openin Y V" and C: "C \ connected_components_of (subtopology Y V)" + then have "C \ topspace Y" + using connected_components_of_subset by force + have ope1: "openin X {a \ topspace X. f a \ V}" + using \openin Y V\ f openin_continuous_map_preimage quotient_imp_continuous_map by blast + define Vf where "Vf \ {z \ topspace X. f z \ V}" + have "openin X {x \ topspace X. f x \ C}" + proof (clarsimp simp: openin_subopen [where S = "{x \ topspace X. f x \ C}"]) + fix x + assume "x \ topspace X" and "f x \ C" + show "\T. openin X T \ x \ T \ T \ {x \ topspace X. f x \ C}" + proof (intro exI conjI) + show "openin X (connected_component_of_set (subtopology X Vf) x)" + by (metis Vf_def X connected_component_of_eq_empty locally_connected_B ope1 openin_empty + openin_subset topspace_subtopology_subset) + show x_in_conn: "x \ connected_component_of_set (subtopology X Vf) x" + using C Vf_def \f x \ C\ \x \ topspace X\ connected_component_of_refl connected_components_of_subset by fastforce + have "connected_component_of_set (subtopology X Vf) x \ topspace X \ Vf" + using connected_component_of_subset_topspace by fastforce + moreover + have "f ` connected_component_of_set (subtopology X Vf) x \ C" + proof (rule connected_components_of_maximal [where X = "subtopology Y V"]) + show "C \ connected_components_of (subtopology Y V)" + by (simp add: C) + have \
: "quotient_map (subtopology X Vf) (subtopology Y V) f" + by (simp add: Vf_def \openin Y V\ f quotient_map_restriction) + then show "connectedin (subtopology Y V) (f ` connected_component_of_set (subtopology X Vf) x)" + by (metis connectedin_connected_component_of connectedin_continuous_map_image quotient_imp_continuous_map) + show "\ disjnt C (f ` connected_component_of_set (subtopology X Vf) x)" + using \f x \ C\ x_in_conn by (auto simp: disjnt_iff) + qed + ultimately + show "connected_component_of_set (subtopology X Vf) x \ {x \ topspace X. f x \ C}" + by blast + qed + qed + then show "openin Y C" + using \C \ topspace Y\ f quotient_map_def by fastforce +qed + + +lemma locally_connected_space_retraction_map_image: + "\retraction_map X Y r; locally_connected_space X\ + \ locally_connected_space Y" + using locally_connected_space_quotient_map_image retraction_imp_quotient_map by blast + +lemma homeomorphic_locally_connected_space: + "X homeomorphic_space Y \ locally_connected_space X \ locally_connected_space Y" + by (meson homeomorphic_map_def homeomorphic_space homeomorphic_space_sym locally_connected_space_quotient_map_image) + +lemma locally_connected_space_euclideanreal: "locally_connected_space euclideanreal" + by (simp add: locally_path_connected_imp_locally_connected_space locally_path_connected_space_euclideanreal) + +lemma locally_connected_is_realinterval: + "is_interval S \ locally_connected_space(subtopology euclideanreal S)" + by (simp add: locally_path_connected_imp_locally_connected_space locally_path_connected_is_realinterval) + +lemma locally_connected_real_interval: + "locally_connected_space (subtopology euclideanreal{a..b})" + "locally_connected_space (subtopology euclideanreal{a<.. locally_connected_at x X" + by (simp add: locally_connected_at_def locally_path_connected_at_def neighbourhood_base_at_mono path_connectedin_imp_connectedin) + +lemma weakly_locally_path_connected_imp_weakly_locally_connected_at: + "weakly_locally_path_connected_at x X + \ weakly_locally_connected_at x X" + by (simp add: neighbourhood_base_at_mono path_connectedin_imp_connectedin weakly_locally_connected_at_def weakly_locally_path_connected_at_def) + + +lemma interior_of_locally_connected_subspace_component: + assumes X: "locally_connected_space X" + and C: "C \ connected_components_of (subtopology X S)" + shows "X interior_of C = C \ X interior_of S" +proof - + obtain Csub: "C \ topspace X" "C \ S" + by (meson C connectedin_connected_components_of connectedin_subset_topspace connectedin_subtopology) + show ?thesis + proof + show "X interior_of C \ C \ X interior_of S" + by (simp add: Csub interior_of_mono interior_of_subset) + have eq: "X interior_of S = \ (connected_components_of (subtopology X (X interior_of S)))" + by (metis Union_connected_components_of interior_of_subset_topspace topspace_subtopology_subset) + moreover have "C \ D \ X interior_of C" + if "D \ connected_components_of (subtopology X (X interior_of S))" for D + proof (cases "C \ D = {}") + case False + have "D \ X interior_of C" + proof (rule interior_of_maximal) + have "connectedin (subtopology X S) D" + by (meson connectedin_connected_components_of connectedin_subtopology interior_of_subset subset_trans that) + then show "D \ C" + by (meson C False connected_components_of_maximal disjnt_def) + show "openin X D" + using X locally_connected_space_open_connected_components openin_interior_of that by blast + qed + then show ?thesis + by blast + qed auto + ultimately show "C \ X interior_of S \ X interior_of C" + by blast + qed +qed + + +lemma frontier_of_locally_connected_subspace_component: + assumes X: "locally_connected_space X" and "closedin X S" + and C: "C \ connected_components_of (subtopology X S)" + shows "X frontier_of C = C \ X frontier_of S" +proof - + obtain Csub: "C \ topspace X" "C \ S" + by (meson C connectedin_connected_components_of connectedin_subset_topspace connectedin_subtopology) + then have "X closure_of C - X interior_of C = C \ X closure_of S - C \ X interior_of S" + using assms + apply (simp add: closure_of_closedin flip: interior_of_locally_connected_subspace_component) + by (metis closedin_connected_components_of closedin_trans_full closure_of_eq inf.orderE) + then show ?thesis + by (simp add: Diff_Int_distrib frontier_of_def) +qed + +(*Similar proof to locally_connected_space_prod_topology*) +lemma locally_connected_space_prod_topology: + "locally_connected_space (prod_topology X Y) \ + topspace (prod_topology X Y) = {} \ + locally_connected_space X \ locally_connected_space Y" (is "?lhs=?rhs") +proof (cases "topspace(prod_topology X Y) = {}") + case True + then show ?thesis + using locally_connected_space_iff_weak by force +next + case False + then have ne: "topspace X \ {}" "topspace Y \ {}" + by simp_all + show ?thesis + proof + assume ?lhs then show ?rhs + by (metis locally_connected_space_quotient_map_image ne quotient_map_fst quotient_map_snd) + next + assume ?rhs + with False have X: "locally_connected_space X" and Y: "locally_connected_space Y" + by auto + show ?lhs + unfolding locally_connected_space_def neighbourhood_base_of + proof clarify + fix UV x y + assume UV: "openin (prod_topology X Y) UV" and "(x,y) \ UV" + + obtain A B where W12: "openin X A \ openin Y B \ x \ A \ y \ B \ A \ B \ UV" + using X Y by (metis UV \(x,y) \ UV\ openin_prod_topology_alt) + then obtain C D K L + where "openin X C" "connectedin X K" "x \ C" "C \ K" "K \ A" + "openin Y D" "connectedin Y L" "y \ D" "D \ L" "L \ B" + by (metis X Y locally_connected_space) + with W12 \openin X C\ \openin Y D\ + show "\U V. openin (prod_topology X Y) U \ connectedin (prod_topology X Y) V \ (x, y) \ U \ U \ V \ V \ UV" + apply (rule_tac x="C \ D" in exI) + apply (rule_tac x="K \ L" in exI) + apply (auto simp: openin_prod_Times_iff connectedin_Times) + done + qed + qed +qed + +(*Same proof as locally_path_connected_space_product_topology*) +lemma locally_connected_space_product_topology: + "locally_connected_space(product_topology X I) \ + topspace(product_topology X I) = {} \ + finite {i. i \ I \ ~connected_space(X i)} \ + (\i \ I. locally_connected_space(X i))" + (is "?lhs \ ?empty \ ?rhs") +proof (cases ?empty) + case True + then show ?thesis + by (simp add: locally_connected_space_def neighbourhood_base_of openin_closedin_eq) +next + case False + then obtain z where z: "z \ (\\<^sub>E i\I. topspace (X i))" + by auto + have ?rhs if L: ?lhs + proof - + obtain U C where U: "openin (product_topology X I) U" + and V: "connectedin (product_topology X I) C" + and "z \ U" "U \ C" and Csub: "C \ (\\<^sub>E i\I. topspace (X i))" + using L apply (clarsimp simp add: locally_connected_space_def neighbourhood_base_of) + by (metis openin_topspace topspace_product_topology z) + then obtain V where finV: "finite {i \ I. V i \ topspace (X i)}" + and XV: "\i. i\I \ openin (X i) (V i)" and "z \ Pi\<^sub>E I V" and subU: "Pi\<^sub>E I V \ U" + by (force simp: openin_product_topology_alt) + show ?thesis + proof (intro conjI ballI) + have "connected_space (X i)" if "i \ I" "V i = topspace (X i)" for i + proof - + have pc: "connectedin (X i) ((\x. x i) ` C)" + apply (rule connectedin_continuous_map_image [OF _ V]) + by (simp add: continuous_map_product_projection \i \ I\) + moreover have "((\x. x i) ` C) = topspace (X i)" + proof + show "(\x. x i) ` C \ topspace (X i)" + by (simp add: pc connectedin_subset_topspace) + have "V i \ (\x. x i) ` (\\<^sub>E i\I. V i)" + by (metis \z \ Pi\<^sub>E I V\ empty_iff image_projection_PiE order_refl that(1)) + also have "\ \ (\x. x i) ` U" + using subU by blast + finally show "topspace (X i) \ (\x. x i) ` C" + using \U \ C\ that by blast + qed + ultimately show ?thesis + by (simp add: connectedin_topspace) + qed + then have "{i \ I. \ connected_space (X i)} \ {i \ I. V i \ topspace (X i)}" + by blast + with finV show "finite {i \ I. \ connected_space (X i)}" + using finite_subset by blast + next + show "locally_connected_space (X i)" if "i \ I" for i + by (meson False L locally_connected_space_quotient_map_image quotient_map_product_projection that) + qed + qed + moreover have ?lhs if R: ?rhs + proof (clarsimp simp add: locally_connected_space_def neighbourhood_base_of) + fix F z + assume "openin (product_topology X I) F" and "z \ F" + then obtain W where finW: "finite {i \ I. W i \ topspace (X i)}" + and opeW: "\i. i \ I \ openin (X i) (W i)" and "z \ Pi\<^sub>E I W" "Pi\<^sub>E I W \ F" + by (auto simp: openin_product_topology_alt) + have "\i \ I. \U C. openin (X i) U \ connectedin (X i) C \ z i \ U \ U \ C \ C \ W i \ + (W i = topspace (X i) \ + connected_space (X i) \ U = topspace (X i) \ C = topspace (X i))" + (is "\i \ I. ?\ i") + proof + fix i assume "i \ I" + have "locally_connected_space (X i)" + by (simp add: R \i \ I\) + moreover have "openin (X i) (W i) " "z i \ W i" + using \z \ Pi\<^sub>E I W\ opeW \i \ I\ by auto + ultimately obtain U C where UC: "openin (X i) U" "connectedin (X i) C" "z i \ U" "U \ C" "C \ W i" + using \i \ I\ by (force simp: locally_connected_space_def neighbourhood_base_of) + show "?\ i" + proof (cases "W i = topspace (X i) \ connected_space(X i)") + case True + then show ?thesis + using \z i \ W i\ connectedin_topspace by blast + next + case False + then show ?thesis + by (meson UC) + qed + qed + then obtain U C where + *: "\i. i \ I \ openin (X i) (U i) \ connectedin (X i) (C i) \ z i \ (U i) \ (U i) \ (C i) \ (C i) \ W i \ + (W i = topspace (X i) \ connected_space (X i) + \ (U i) = topspace (X i) \ (C i) = topspace (X i))" + by metis + let ?A = "{i \ I. \ connected_space (X i)} \ {i \ I. W i \ topspace (X i)}" + have "{i \ I. U i \ topspace (X i)} \ ?A" + by (clarsimp simp add: "*") + moreover have "finite ?A" + by (simp add: that finW) + ultimately have "finite {i \ I. U i \ topspace (X i)}" + using finite_subset by auto + then have "openin (product_topology X I) (Pi\<^sub>E I U)" + using * by (simp add: openin_PiE_gen) + then show "\U. openin (product_topology X I) U \ + (\V. connectedin (product_topology X I) V \ z \ U \ U \ V \ V \ F)" + apply (rule_tac x="PiE I U" in exI, simp) + apply (rule_tac x="PiE I C" in exI) + using \z \ Pi\<^sub>E I W\ \Pi\<^sub>E I W \ F\ * + apply (simp add: connectedin_PiE subset_PiE PiE_iff PiE_mono dual_order.trans) + done + qed + ultimately show ?thesis + using False by blast +qed + +lemma locally_connected_space_sum_topology: + "locally_connected_space(sum_topology X I) \ + (\i \ I. locally_connected_space (X i))" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + by (smt (verit) homeomorphic_locally_connected_space locally_connected_space_open_subset topological_property_of_sum_component) +next + assume R: ?rhs + show ?lhs + proof (clarsimp simp add: locally_connected_space_def neighbourhood_base_of forall_openin_sum_topology imp_conjL) + fix W i x + assume ope: "\i\I. openin (X i) (W i)" + and "i \ I" and "x \ W i" + then obtain U V where U: "openin (X i) U" and V: "connectedin (X i) V" + and "x \ U" "U \ V" "V \ W i" + by (metis R \i \ I\ \x \ W i\ locally_connected_space) + show "\U. openin (sum_topology X I) U \ (\V. connectedin (sum_topology X I) V \ (i,x) \ U \ U \ V \ V \ Sigma I W)" + proof (intro exI conjI) + show "openin (sum_topology X I) (Pair i ` U)" + by (meson U \i \ I\ open_map_component_injection open_map_def) + show "connectedin (sum_topology X I) (Pair i ` V)" + by (meson V \i \ I\ continuous_map_component_injection connectedin_continuous_map_image) + show "Pair i ` V \ Sigma I W" + using \V \ W i\ \i \ I\ by force + qed (use \x \ U\ \U \ V\ in auto) + qed +qed + + end diff --git a/src/HOL/Analysis/Product_Topology.thy b/src/HOL/Analysis/Product_Topology.thy --- a/src/HOL/Analysis/Product_Topology.thy +++ b/src/HOL/Analysis/Product_Topology.thy @@ -1,580 +1,643 @@ section\The binary product topology\ theory Product_Topology imports Function_Topology begin section \Product Topology\ subsection\Definition\ definition prod_topology :: "'a topology \ 'b topology \ ('a \ 'b) topology" where "prod_topology X Y \ topology (arbitrary union_of (\U. U \ {S \ T |S T. openin X S \ openin Y T}))" lemma open_product_open: assumes "open A" shows "\\. \ \ {S \ T |S T. open S \ open T} \ \ \ = A" proof - obtain f g where *: "\u. u \ A \ open (f u) \ open (g u) \ u \ (f u) \ (g u) \ (f u) \ (g u) \ A" using open_prod_def [of A] assms by metis let ?\ = "(\u. f u \ g u) ` A" show ?thesis by (rule_tac x="?\" in exI) (auto simp: dest: *) qed lemma open_product_open_eq: "(arbitrary union_of (\U. \S T. U = S \ T \ open S \ open T)) = open" by (force simp: union_of_def arbitrary_def intro: open_product_open open_Times) lemma openin_prod_topology: "openin (prod_topology X Y) = arbitrary union_of (\U. U \ {S \ T |S T. openin X S \ openin Y T})" unfolding prod_topology_def proof (rule topology_inverse') show "istopology (arbitrary union_of (\U. U \ {S \ T |S T. openin X S \ openin Y T}))" apply (rule istopology_base, simp) by (metis openin_Int Times_Int_Times) qed lemma topspace_prod_topology [simp]: "topspace (prod_topology X Y) = topspace X \ topspace Y" proof - have "topspace(prod_topology X Y) = \ (Collect (openin (prod_topology X Y)))" (is "_ = ?Z") unfolding topspace_def .. also have "\ = topspace X \ topspace Y" proof show "?Z \ topspace X \ topspace Y" apply (auto simp: openin_prod_topology union_of_def arbitrary_def) using openin_subset by force+ next have *: "\A B. topspace X \ topspace Y = A \ B \ openin X A \ openin Y B" by blast show "topspace X \ topspace Y \ ?Z" apply (rule Union_upper) using * by (simp add: openin_prod_topology arbitrary_union_of_inc) qed finally show ?thesis . qed lemma subtopology_Times: shows "subtopology (prod_topology X Y) (S \ T) = prod_topology (subtopology X S) (subtopology Y T)" proof - have "((\U. \S T. U = S \ T \ openin X S \ openin Y T) relative_to S \ T) = (\U. \S' T'. U = S' \ T' \ (openin X relative_to S) S' \ (openin Y relative_to T) T')" by (auto simp: relative_to_def Times_Int_Times fun_eq_iff) metis then show ?thesis by (simp add: topology_eq openin_prod_topology arbitrary_union_of_relative_to flip: openin_relative_to) qed lemma prod_topology_subtopology: "prod_topology (subtopology X S) Y = subtopology (prod_topology X Y) (S \ topspace Y)" "prod_topology X (subtopology Y T) = subtopology (prod_topology X Y) (topspace X \ T)" by (auto simp: subtopology_Times) lemma prod_topology_discrete_topology: "discrete_topology (S \ T) = prod_topology (discrete_topology S) (discrete_topology T)" by (auto simp: discrete_topology_unique openin_prod_topology intro: arbitrary_union_of_inc) lemma prod_topology_euclidean [simp]: "prod_topology euclidean euclidean = euclidean" by (simp add: prod_topology_def open_product_open_eq) lemma prod_topology_subtopology_eu [simp]: "prod_topology (subtopology euclidean S) (subtopology euclidean T) = subtopology euclidean (S \ T)" by (simp add: prod_topology_subtopology subtopology_subtopology Times_Int_Times) lemma openin_prod_topology_alt: "openin (prod_topology X Y) S \ (\x y. (x,y) \ S \ (\U V. openin X U \ openin Y V \ x \ U \ y \ V \ U \ V \ S))" apply (auto simp: openin_prod_topology arbitrary_union_of_alt, fastforce) by (metis mem_Sigma_iff) lemma open_map_fst: "open_map (prod_topology X Y) X fst" unfolding open_map_def openin_prod_topology_alt by (force simp: openin_subopen [of X "fst ` _"] intro: subset_fst_imageI) lemma open_map_snd: "open_map (prod_topology X Y) Y snd" unfolding open_map_def openin_prod_topology_alt by (force simp: openin_subopen [of Y "snd ` _"] intro: subset_snd_imageI) lemma openin_prod_Times_iff: "openin (prod_topology X Y) (S \ T) \ S = {} \ T = {} \ openin X S \ openin Y T" proof (cases "S = {} \ T = {}") case False then show ?thesis apply (simp add: openin_prod_topology_alt openin_subopen [of X S] openin_subopen [of Y T] times_subset_iff, safe) apply (meson|force)+ done qed force lemma closure_of_Times: "(prod_topology X Y) closure_of (S \ T) = (X closure_of S) \ (Y closure_of T)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (clarsimp simp: closure_of_def openin_prod_topology_alt) blast show "?rhs \ ?lhs" by (clarsimp simp: closure_of_def openin_prod_topology_alt) (meson SigmaI subsetD) qed lemma closedin_prod_Times_iff: "closedin (prod_topology X Y) (S \ T) \ S = {} \ T = {} \ closedin X S \ closedin Y T" by (auto simp: closure_of_Times times_eq_iff simp flip: closure_of_eq) lemma interior_of_Times: "(prod_topology X Y) interior_of (S \ T) = (X interior_of S) \ (Y interior_of T)" proof (rule interior_of_unique) show "(X interior_of S) \ Y interior_of T \ S \ T" by (simp add: Sigma_mono interior_of_subset) show "openin (prod_topology X Y) ((X interior_of S) \ Y interior_of T)" by (simp add: openin_prod_Times_iff) next show "T' \ (X interior_of S) \ Y interior_of T" if "T' \ S \ T" "openin (prod_topology X Y) T'" for T' proof (clarsimp; intro conjI) fix a :: "'a" and b :: "'b" assume "(a, b) \ T'" with that obtain U V where UV: "openin X U" "openin Y V" "a \ U" "b \ V" "U \ V \ T'" by (metis openin_prod_topology_alt) then show "a \ X interior_of S" using interior_of_maximal_eq that(1) by fastforce show "b \ Y interior_of T" using UV interior_of_maximal_eq that(1) by (metis SigmaI mem_Sigma_iff subset_eq) qed qed subsection \Continuity\ lemma continuous_map_pairwise: "continuous_map Z (prod_topology X Y) f \ continuous_map Z X (fst \ f) \ continuous_map Z Y (snd \ f)" (is "?lhs = ?rhs") proof - let ?g = "fst \ f" and ?h = "snd \ f" have f: "f x = (?g x, ?h x)" for x by auto show ?thesis proof (cases "(\x \ topspace Z. ?g x \ topspace X) \ (\x \ topspace Z. ?h x \ topspace Y)") case True show ?thesis proof safe assume "continuous_map Z (prod_topology X Y) f" then have "openin Z {x \ topspace Z. fst (f x) \ U}" if "openin X U" for U unfolding continuous_map_def using True that apply clarify apply (drule_tac x="U \ topspace Y" in spec) by (simp add: openin_prod_Times_iff mem_Times_iff cong: conj_cong) with True show "continuous_map Z X (fst \ f)" by (auto simp: continuous_map_def) next assume "continuous_map Z (prod_topology X Y) f" then have "openin Z {x \ topspace Z. snd (f x) \ V}" if "openin Y V" for V unfolding continuous_map_def using True that apply clarify apply (drule_tac x="topspace X \ V" in spec) by (simp add: openin_prod_Times_iff mem_Times_iff cong: conj_cong) with True show "continuous_map Z Y (snd \ f)" by (auto simp: continuous_map_def) next assume Z: "continuous_map Z X (fst \ f)" "continuous_map Z Y (snd \ f)" have *: "openin Z {x \ topspace Z. f x \ W}" if "\w. w \ W \ \U V. openin X U \ openin Y V \ w \ U \ V \ U \ V \ W" for W proof (subst openin_subopen, clarify) fix x :: "'a" assume "x \ topspace Z" and "f x \ W" with that [OF \f x \ W\] obtain U V where UV: "openin X U" "openin Y V" "f x \ U \ V" "U \ V \ W" by auto with Z UV show "\T. openin Z T \ x \ T \ T \ {x \ topspace Z. f x \ W}" apply (rule_tac x="{x \ topspace Z. ?g x \ U} \ {x \ topspace Z. ?h x \ V}" in exI) apply (auto simp: \x \ topspace Z\ continuous_map_def) done qed show "continuous_map Z (prod_topology X Y) f" using True by (simp add: continuous_map_def openin_prod_topology_alt mem_Times_iff *) qed qed (auto simp: continuous_map_def) qed lemma continuous_map_paired: "continuous_map Z (prod_topology X Y) (\x. (f x,g x)) \ continuous_map Z X f \ continuous_map Z Y g" by (simp add: continuous_map_pairwise o_def) lemma continuous_map_pairedI [continuous_intros]: "\continuous_map Z X f; continuous_map Z Y g\ \ continuous_map Z (prod_topology X Y) (\x. (f x,g x))" by (simp add: continuous_map_pairwise o_def) lemma continuous_map_fst [continuous_intros]: "continuous_map (prod_topology X Y) X fst" using continuous_map_pairwise [of "prod_topology X Y" X Y id] by (simp add: continuous_map_pairwise) lemma continuous_map_snd [continuous_intros]: "continuous_map (prod_topology X Y) Y snd" using continuous_map_pairwise [of "prod_topology X Y" X Y id] by (simp add: continuous_map_pairwise) lemma continuous_map_fst_of [continuous_intros]: "continuous_map Z (prod_topology X Y) f \ continuous_map Z X (fst \ f)" by (simp add: continuous_map_pairwise) lemma continuous_map_snd_of [continuous_intros]: "continuous_map Z (prod_topology X Y) f \ continuous_map Z Y (snd \ f)" by (simp add: continuous_map_pairwise) lemma continuous_map_prod_fst: "i \ I \ continuous_map (prod_topology (product_topology (\i. Y) I) X) Y (\x. fst x i)" using continuous_map_componentwise_UNIV continuous_map_fst by fastforce lemma continuous_map_prod_snd: "i \ I \ continuous_map (prod_topology X (product_topology (\i. Y) I)) Y (\x. snd x i)" using continuous_map_componentwise_UNIV continuous_map_snd by fastforce lemma continuous_map_if_iff [simp]: "continuous_map X Y (\x. if P then f x else g x) \ continuous_map X Y (if P then f else g)" by simp lemma continuous_map_if [continuous_intros]: "\P \ continuous_map X Y f; ~P \ continuous_map X Y g\ \ continuous_map X Y (\x. if P then f x else g x)" by simp lemma continuous_map_subtopology_fst [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) X fst" using continuous_map_from_subtopology continuous_map_fst by force lemma continuous_map_subtopology_snd [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) Y snd" using continuous_map_from_subtopology continuous_map_snd by force lemma quotient_map_fst [simp]: "quotient_map(prod_topology X Y) X fst \ (topspace Y = {} \ topspace X = {})" by (auto simp: continuous_open_quotient_map open_map_fst continuous_map_fst) lemma quotient_map_snd [simp]: "quotient_map(prod_topology X Y) Y snd \ (topspace X = {} \ topspace Y = {})" by (auto simp: continuous_open_quotient_map open_map_snd continuous_map_snd) lemma retraction_map_fst: "retraction_map (prod_topology X Y) X fst \ (topspace Y = {} \ topspace X = {})" proof (cases "topspace Y = {}") case True then show ?thesis using continuous_map_image_subset_topspace by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst continuous_map_on_empty) next case False have "\g. continuous_map X (prod_topology X Y) g \ (\x\topspace X. fst (g x) = x)" if y: "y \ topspace Y" for y by (rule_tac x="\x. (x,y)" in exI) (auto simp: y continuous_map_paired) with False have "retraction_map (prod_topology X Y) X fst" by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst) with False show ?thesis by simp qed lemma retraction_map_snd: "retraction_map (prod_topology X Y) Y snd \ (topspace X = {} \ topspace Y = {})" proof (cases "topspace X = {}") case True then show ?thesis using continuous_map_image_subset_topspace by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst continuous_map_on_empty) next case False have "\g. continuous_map Y (prod_topology X Y) g \ (\y\topspace Y. snd (g y) = y)" if x: "x \ topspace X" for x by (rule_tac x="\y. (x,y)" in exI) (auto simp: x continuous_map_paired) with False have "retraction_map (prod_topology X Y) Y snd" by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_snd) with False show ?thesis by simp qed lemma continuous_map_of_fst: "continuous_map (prod_topology X Y) Z (f \ fst) \ topspace Y = {} \ continuous_map X Z f" proof (cases "topspace Y = {}") case True then show ?thesis by (simp add: continuous_map_on_empty) next case False then show ?thesis by (simp add: continuous_compose_quotient_map_eq) qed lemma continuous_map_of_snd: "continuous_map (prod_topology X Y) Z (f \ snd) \ topspace X = {} \ continuous_map Y Z f" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: continuous_map_on_empty) next case False then show ?thesis by (simp add: continuous_compose_quotient_map_eq) qed lemma continuous_map_prod_top: "continuous_map (prod_topology X Y) (prod_topology X' Y') (\(x,y). (f x, g y)) \ topspace (prod_topology X Y) = {} \ continuous_map X X' f \ continuous_map Y Y' g" proof (cases "topspace (prod_topology X Y) = {}") case True then show ?thesis by (simp add: continuous_map_on_empty) next case False then show ?thesis by (simp add: continuous_map_paired case_prod_unfold continuous_map_of_fst [unfolded o_def] continuous_map_of_snd [unfolded o_def]) qed lemma in_prod_topology_closure_of: assumes "z \ (prod_topology X Y) closure_of S" shows "fst z \ X closure_of (fst ` S)" "snd z \ Y closure_of (snd ` S)" using assms continuous_map_eq_image_closure_subset continuous_map_fst apply fastforce using assms continuous_map_eq_image_closure_subset continuous_map_snd apply fastforce done proposition compact_space_prod_topology: "compact_space(prod_topology X Y) \ topspace(prod_topology X Y) = {} \ compact_space X \ compact_space Y" proof (cases "topspace(prod_topology X Y) = {}") case True then show ?thesis using compact_space_topspace_empty by blast next case False then have non_mt: "topspace X \ {}" "topspace Y \ {}" by auto have "compact_space X" "compact_space Y" if "compact_space(prod_topology X Y)" proof - have "compactin X (fst ` (topspace X \ topspace Y))" by (metis compact_space_def continuous_map_fst image_compactin that topspace_prod_topology) moreover have "compactin Y (snd ` (topspace X \ topspace Y))" by (metis compact_space_def continuous_map_snd image_compactin that topspace_prod_topology) ultimately show "compact_space X" "compact_space Y" by (simp_all add: non_mt compact_space_def) qed moreover define \ where "\ \ (\V. topspace X \ V) ` Collect (openin Y)" define \ where "\ \ (\U. U \ topspace Y) ` Collect (openin X)" have "compact_space(prod_topology X Y)" if "compact_space X" "compact_space Y" proof (rule Alexander_subbase_alt) show toptop: "topspace X \ topspace Y \ \(\ \ \)" unfolding \_def \_def by auto fix \ :: "('a \ 'b) set set" assume \: "\ \ \ \ \" "topspace X \ topspace Y \ \\" then obtain \' \' where XY: "\' \ \" "\' \ \" and \eq: "\ = \' \ \'" using subset_UnE by metis then have sub: "topspace X \ topspace Y \ \(\' \ \')" using \ by simp obtain \ \ where \: "\U. U \ \ \ openin X U" "\' = (\U. U \ topspace Y) ` \" and \: "\V. V \ \ \ openin Y V" "\' = (\V. topspace X \ V) ` \" using XY by (clarsimp simp add: \_def \_def subset_image_iff) (force simp add: subset_iff) have "\\. finite \ \ \ \ \' \ \' \ topspace X \ topspace Y \ \ \" proof - have "topspace X \ \\ \ topspace Y \ \\" using \ \ \ \eq by auto then have *: "\\. finite \ \ (\x \ \. x \ (\V. topspace X \ V) ` \ \ x \ (\U. U \ topspace Y) ` \) \ (topspace X \ topspace Y \ \\)" proof assume "topspace X \ \\" with \compact_space X\ \ obtain \ where "finite \" "\ \ \" "topspace X \ \\" by (meson compact_space_alt) with that show ?thesis by (rule_tac x="(\D. D \ topspace Y) ` \" in exI) auto next assume "topspace Y \ \\" with \compact_space Y\ \ obtain \ where "finite \" "\ \ \" "topspace Y \ \\" by (meson compact_space_alt) with that show ?thesis by (rule_tac x="(\C. topspace X \ C) ` \" in exI) auto qed then show ?thesis using that \ \ by blast qed then show "\\. finite \ \ \ \ \ \ topspace X \ topspace Y \ \ \" using \ \eq by blast next have "(finite intersection_of (\x. x \ \ \ x \ \) relative_to topspace X \ topspace Y) = (\U. \S T. U = S \ T \ openin X S \ openin Y T)" (is "?lhs = ?rhs") proof - have "?rhs U" if "?lhs U" for U proof - have "topspace X \ topspace Y \ \ T \ {A \ B |A B. A \ Collect (openin X) \ B \ Collect (openin Y)}" if "finite T" "T \ \ \ \" for T using that proof induction case (insert B \) then show ?case unfolding \_def \_def apply (simp add: Int_ac subset_eq image_def) apply (metis (no_types) openin_Int openin_topspace Times_Int_Times) done qed auto then show ?thesis using that by (auto simp: subset_eq elim!: relative_toE intersection_ofE) qed moreover have "?lhs Z" if Z: "?rhs Z" for Z proof - obtain U V where "Z = U \ V" "openin X U" "openin Y V" using Z by blast then have UV: "U \ V = (topspace X \ topspace Y) \ (U \ V)" by (simp add: Sigma_mono inf_absorb2 openin_subset) moreover have "?lhs ((topspace X \ topspace Y) \ (U \ V))" proof (rule relative_to_inc) show "(finite intersection_of (\x. x \ \ \ x \ \)) (U \ V)" apply (simp add: intersection_of_def \_def \_def) apply (rule_tac x="{(U \ topspace Y),(topspace X \ V)}" in exI) using \openin X U\ \openin Y V\ openin_subset UV apply (fastforce simp add:) done qed ultimately show ?thesis using \Z = U \ V\ by auto qed ultimately show ?thesis by meson qed then show "topology (arbitrary union_of (finite intersection_of (\x. x \ \ \ \) relative_to (topspace X \ topspace Y))) = prod_topology X Y" by (simp add: prod_topology_def) qed ultimately show ?thesis using False by blast qed lemma compactin_Times: "compactin (prod_topology X Y) (S \ T) \ S = {} \ T = {} \ compactin X S \ compactin Y T" by (auto simp: compactin_subspace subtopology_Times compact_space_prod_topology) subsection\Homeomorphic maps\ lemma homeomorphic_maps_prod: "homeomorphic_maps (prod_topology X Y) (prod_topology X' Y') (\(x,y). (f x, g y)) (\(x,y). (f' x, g' y)) \ topspace(prod_topology X Y) = {} \ topspace(prod_topology X' Y') = {} \ homeomorphic_maps X X' f f' \ homeomorphic_maps Y Y' g g'" unfolding homeomorphic_maps_def continuous_map_prod_top by (auto simp: continuous_map_def homeomorphic_maps_def continuous_map_prod_top) lemma homeomorphic_maps_swap: "homeomorphic_maps (prod_topology X Y) (prod_topology Y X) (\(x,y). (y,x)) (\(y,x). (x,y))" by (auto simp: homeomorphic_maps_def case_prod_unfold continuous_map_fst continuous_map_pairedI continuous_map_snd) lemma homeomorphic_map_swap: "homeomorphic_map (prod_topology X Y) (prod_topology Y X) (\(x,y). (y,x))" using homeomorphic_map_maps homeomorphic_maps_swap by metis lemma homeomorphic_space_prod_topology_swap: "(prod_topology X Y) homeomorphic_space (prod_topology Y X)" using homeomorphic_map_swap homeomorphic_space by blast lemma embedding_map_graph: "embedding_map X (prod_topology X Y) (\x. (x, f x)) \ continuous_map X Y f" (is "?lhs = ?rhs") proof assume L: ?lhs have "snd \ (\x. (x, f x)) = f" by force moreover have "continuous_map X Y (snd \ (\x. (x, f x)))" using L unfolding embedding_map_def by (meson continuous_map_in_subtopology continuous_map_snd_of homeomorphic_imp_continuous_map) ultimately show ?rhs by simp next assume R: ?rhs then show ?lhs unfolding homeomorphic_map_maps embedding_map_def homeomorphic_maps_def by (rule_tac x=fst in exI) (auto simp: continuous_map_in_subtopology continuous_map_paired continuous_map_from_subtopology continuous_map_fst) qed lemma homeomorphic_space_prod_topology: "\X homeomorphic_space X''; Y homeomorphic_space Y'\ \ prod_topology X Y homeomorphic_space prod_topology X'' Y'" using homeomorphic_maps_prod unfolding homeomorphic_space_def by blast lemma prod_topology_homeomorphic_space_left: "topspace Y = {b} \ prod_topology X Y homeomorphic_space X" unfolding homeomorphic_space_def by (rule_tac x=fst in exI) (simp add: homeomorphic_map_def inj_on_def flip: homeomorphic_map_maps) lemma prod_topology_homeomorphic_space_right: "topspace X = {a} \ prod_topology X Y homeomorphic_space Y" unfolding homeomorphic_space_def by (rule_tac x=snd in exI) (simp add: homeomorphic_map_def inj_on_def flip: homeomorphic_map_maps) lemma homeomorphic_space_prod_topology_sing1: "b \ topspace Y \ X homeomorphic_space (prod_topology X (subtopology Y {b}))" by (metis empty_subsetI homeomorphic_space_sym inf.absorb_iff2 insert_subset prod_topology_homeomorphic_space_left topspace_subtopology) lemma homeomorphic_space_prod_topology_sing2: "a \ topspace X \ Y homeomorphic_space (prod_topology (subtopology X {a}) Y)" by (metis empty_subsetI homeomorphic_space_sym inf.absorb_iff2 insert_subset prod_topology_homeomorphic_space_right topspace_subtopology) lemma topological_property_of_prod_component: assumes major: "P(prod_topology X Y)" and X: "\x. \x \ topspace X; P(prod_topology X Y)\ \ P(subtopology (prod_topology X Y) ({x} \ topspace Y))" and Y: "\y. \y \ topspace Y; P(prod_topology X Y)\ \ P(subtopology (prod_topology X Y) (topspace X \ {y}))" and PQ: "\X X'. X homeomorphic_space X' \ (P X \ Q X')" and PR: "\X X'. X homeomorphic_space X' \ (P X \ R X')" shows "topspace(prod_topology X Y) = {} \ Q X \ R Y" proof - have "Q X \ R Y" if "topspace(prod_topology X Y) \ {}" proof - from that obtain a b where a: "a \ topspace X" and b: "b \ topspace Y" by force show ?thesis using X [OF a major] and Y [OF b major] homeomorphic_space_prod_topology_sing1 [OF b, of X] homeomorphic_space_prod_topology_sing2 [OF a, of Y] by (simp add: subtopology_Times) (meson PQ PR homeomorphic_space_prod_topology_sing2 homeomorphic_space_sym) qed then show ?thesis by metis qed lemma limitin_pairwise: "limitin (prod_topology X Y) f l F \ limitin X (fst \ f) (fst l) F \ limitin Y (snd \ f) (snd l) F" (is "?lhs = ?rhs") proof assume ?lhs then obtain a b where ev: "\U. \(a,b) \ U; openin (prod_topology X Y) U\ \ \\<^sub>F x in F. f x \ U" and a: "a \ topspace X" and b: "b \ topspace Y" and l: "l = (a,b)" by (auto simp: limitin_def) moreover have "\\<^sub>F x in F. fst (f x) \ U" if "openin X U" "a \ U" for U proof - have "\\<^sub>F c in F. f c \ U \ topspace Y" using b that ev [of "U \ topspace Y"] by (auto simp: openin_prod_topology_alt) then show ?thesis by (rule eventually_mono) (metis (mono_tags, lifting) SigmaE2 prod.collapse) qed moreover have "\\<^sub>F x in F. snd (f x) \ U" if "openin Y U" "b \ U" for U proof - have "\\<^sub>F c in F. f c \ topspace X \ U" using a that ev [of "topspace X \ U"] by (auto simp: openin_prod_topology_alt) then show ?thesis by (rule eventually_mono) (metis (mono_tags, lifting) SigmaE2 prod.collapse) qed ultimately show ?rhs by (simp add: limitin_def) next have "limitin (prod_topology X Y) f (a,b) F" if "limitin X (fst \ f) a F" "limitin Y (snd \ f) b F" for a b using that proof (clarsimp simp: limitin_def) fix Z :: "('a \ 'b) set" assume a: "a \ topspace X" "\U. openin X U \ a \ U \ (\\<^sub>F x in F. fst (f x) \ U)" and b: "b \ topspace Y" "\U. openin Y U \ b \ U \ (\\<^sub>F x in F. snd (f x) \ U)" and Z: "openin (prod_topology X Y) Z" "(a, b) \ Z" then obtain U V where "openin X U" "openin Y V" "a \ U" "b \ V" "U \ V \ Z" using Z by (force simp: openin_prod_topology_alt) then have "\\<^sub>F x in F. fst (f x) \ U" "\\<^sub>F x in F. snd (f x) \ V" by (simp_all add: a b) then show "\\<^sub>F x in F. f x \ Z" by (rule eventually_elim2) (use \U \ V \ Z\ subsetD in auto) qed then show "?rhs \ ?lhs" by (metis prod.collapse) qed +proposition connected_space_prod_topology: + "connected_space(prod_topology X Y) \ + topspace(prod_topology X Y) = {} \ connected_space X \ connected_space Y" (is "?lhs=?rhs") +proof (cases "topspace(prod_topology X Y) = {}") + case True + then show ?thesis + using connected_space_topspace_empty by blast +next + case False + then have nonempty: "topspace X \ {}" "topspace Y \ {}" + by force+ + show ?thesis + proof + assume ?lhs + then show ?rhs + by (meson connected_space_quotient_map_image nonempty quotient_map_fst quotient_map_snd) + next + assume ?rhs + then have conX: "connected_space X" and conY: "connected_space Y" + using False by blast+ + have False + if "openin (prod_topology X Y) U" and "openin (prod_topology X Y) V" + and UV: "topspace X \ topspace Y \ U \ V" "U \ V = {}" + and "U \ {}" and "V \ {}" + for U V + proof - + have Usub: "U \ topspace X \ topspace Y" and Vsub: "V \ topspace X \ topspace Y" + using that by (metis openin_subset topspace_prod_topology)+ + obtain a b where ab: "(a,b) \ U" and a: "a \ topspace X" and b: "b \ topspace Y" + using \U \ {}\ Usub by auto + have "\ topspace X \ topspace Y \ U" + using Usub Vsub \U \ V = {}\ \V \ {}\ by auto + then obtain x y where x: "x \ topspace X" and y: "y \ topspace Y" and "(x,y) \ U" + by blast + have oX: "openin X {x \ topspace X. (x,y) \ U}" "openin X {x \ topspace X. (x,y) \ V}" + and oY: "openin Y {y \ topspace Y. (a,y) \ U}" "openin Y {y \ topspace Y. (a,y) \ V}" + by (force intro: openin_continuous_map_preimage [where Y = "prod_topology X Y"] + simp: that continuous_map_pairwise o_def x y a)+ + have 1: "topspace Y \ {y \ topspace Y. (a,y) \ U} \ {y \ topspace Y. (a,y) \ V}" + using a that(3) by auto + have 2: "{y \ topspace Y. (a,y) \ U} \ {y \ topspace Y. (a,y) \ V} = {}" + using that(4) by auto + have 3: "{y \ topspace Y. (a,y) \ U} \ {}" + using ab b by auto + have 4: "{y \ topspace Y. (a,y) \ V} \ {}" + proof - + show ?thesis + using connected_spaceD [OF conX oX] UV \(x,y) \ U\ a x y + disjoint_iff_not_equal by blast + qed + show ?thesis + using connected_spaceD [OF conY oY 1 2 3 4] by auto + qed + then show ?lhs + unfolding connected_space_def topspace_prod_topology by blast + qed +qed + +lemma connectedin_Times: + "connectedin (prod_topology X Y) (S \ T) \ + S = {} \ T = {} \ connectedin X S \ connectedin Y T" + by (force simp: connectedin_def subtopology_Times connected_space_prod_topology) + end diff --git a/src/HOL/Analysis/Starlike.thy b/src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy +++ b/src/HOL/Analysis/Starlike.thy @@ -1,6348 +1,6353 @@ (* Title: HOL/Analysis/Starlike.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) chapter \Unsorted\ theory Starlike imports Convex_Euclidean_Space Line_Segment begin lemma affine_hull_closed_segment [simp]: "affine hull (closed_segment a b) = affine hull {a,b}" by (simp add: segment_convex_hull) lemma affine_hull_open_segment [simp]: fixes a :: "'a::euclidean_space" shows "affine hull (open_segment a b) = (if a = b then {} else affine hull {a,b})" by (metis affine_hull_convex_hull affine_hull_empty closure_open_segment closure_same_affine_hull segment_convex_hull) lemma rel_interior_closure_convex_segment: fixes S :: "_::euclidean_space set" assumes "convex S" "a \ rel_interior S" "b \ closure S" shows "open_segment a b \ rel_interior S" proof fix x have [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" for u by (simp add: algebra_simps) assume "x \ open_segment a b" then show "x \ rel_interior S" unfolding closed_segment_def open_segment_def using assms by (auto intro: rel_interior_closure_convex_shrink) qed lemma convex_hull_insert_segments: "convex hull (insert a S) = (if S = {} then {a} else \x \ convex hull S. closed_segment a x)" by (force simp add: convex_hull_insert_alt in_segment) lemma Int_convex_hull_insert_rel_exterior: fixes z :: "'a::euclidean_space" assumes "convex C" "T \ C" and z: "z \ rel_interior C" and dis: "disjnt S (rel_interior C)" shows "S \ (convex hull (insert z T)) = S \ (convex hull T)" (is "?lhs = ?rhs") proof have "T = {} \ z \ S" using dis z by (auto simp add: disjnt_def) then show "?lhs \ ?rhs" proof (clarsimp simp add: convex_hull_insert_segments) fix x y assume "x \ S" and y: "y \ convex hull T" and "x \ closed_segment z y" have "y \ closure C" by (metis y \convex C\ \T \ C\ closure_subset contra_subsetD convex_hull_eq hull_mono) moreover have "x \ rel_interior C" by (meson \x \ S\ dis disjnt_iff) moreover have "x \ open_segment z y \ {z, y}" using \x \ closed_segment z y\ closed_segment_eq_open by blast ultimately show "x \ convex hull T" using rel_interior_closure_convex_segment [OF \convex C\ z] using y z by blast qed show "?rhs \ ?lhs" by (meson hull_mono inf_mono subset_insertI subset_refl) qed subsection\<^marker>\tag unimportant\ \Shrinking towards the interior of a convex set\ lemma mem_interior_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto show ?thesis unfolding mem_interior proof (intro exI subsetI conjI) fix y assume "y \ ball (x - e *\<^sub>R (x - c)) (e*d)" then have as: "dist (x - e *\<^sub>R (x - c)) y < e * d" by simp have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using \e > 0\ by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "c - ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = (1 / e) *\<^sub>R (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" using \e > 0\ by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) then have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \1/e\ * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" by (simp add: dist_norm) also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ by (auto simp add:pos_divide_less_eq[OF \e > 0\] mult.commute) finally have "(1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x \ S" using assms(3-5) d by (intro convexD_alt [OF \convex S\]) (auto intro: convexD_alt [OF \convex S\]) with \e > 0\ show "y \ S" by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) qed (use \e>0\ \d>0\ in auto) qed lemma mem_interior_closure_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ closure S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto have "\y\S. norm (y - x) * (1 - e) < e * d" proof (cases "x \ S") case True then show ?thesis using \e > 0\ \d > 0\ by force next case False then have x: "x islimpt S" using assms(3)[unfolded closure_def] by auto show ?thesis proof (cases "e = 1") case True obtain y where "y \ S" "y \ x" "dist y x < 1" using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto then show ?thesis using True \0 < d\ by auto next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using islimpt_approachable x by blast then have "norm (y - x) * (1 - e) < e * d" by (metis "*" dist_norm mult_imp_div_pos_le not_less) then show ?thesis using \y \ S\ by blast qed qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have "(1 - e) * norm (x - y) / e < d" using y \0 < e\ by (simp add: field_simps norm_minus_commute) then have "z \ interior (ball c d)" using \0 < e\ \e \ 1\ by (simp add: interior_open[OF open_ball] z_def dist_norm) then have "z \ interior S" using d interiorI interior_ball by blast then show ?thesis unfolding * using mem_interior_convex_shrink \y \ S\ assms by blast qed lemma in_interior_closure_convex_segment: fixes S :: "'a::euclidean_space set" assumes "convex S" and a: "a \ interior S" and b: "b \ closure S" shows "open_segment a b \ interior S" proof (clarsimp simp: in_segment) fix u::real assume u: "0 < u" "u < 1" have "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" by (simp add: algebra_simps) also have "... \ interior S" using mem_interior_closure_convex_shrink [OF assms] u by simp finally show "(1 - u) *\<^sub>R a + u *\<^sub>R b \ interior S" . qed lemma convex_closure_interior: fixes S :: "'a::euclidean_space set" assumes "convex S" and int: "interior S \ {}" shows "closure(interior S) = closure S" proof - obtain a where a: "a \ interior S" using int by auto have "closure S \ closure(interior S)" proof fix x assume x: "x \ closure S" show "x \ closure (interior S)" proof (cases "x=a") case True then show ?thesis using \a \ interior S\ closure_subset by blast next case False show ?thesis proof (clarsimp simp add: closure_def islimpt_approachable) fix e::real assume xnotS: "x \ interior S" and "0 < e" show "\x'\interior S. x' \ x \ dist x' x < e" proof (intro bexI conjI) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ x" using False \0 < e\ by (auto simp: algebra_simps min_def) show "dist (x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a)) x < e" using \0 < e\ by (auto simp: dist_norm min_def) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ interior S" using \0 < e\ False by (auto simp add: min_def a intro: mem_interior_closure_convex_shrink [OF \convex S\ a x]) qed qed qed qed then show ?thesis by (simp add: closure_mono interior_subset subset_antisym) qed lemma closure_convex_Int_superset: fixes S :: "'a::euclidean_space set" assumes "convex S" "interior S \ {}" "interior S \ closure T" shows "closure(S \ T) = closure S" proof - have "closure S \ closure(interior S)" by (simp add: convex_closure_interior assms) also have "... \ closure (S \ T)" using interior_subset [of S] assms by (metis (no_types, lifting) Int_assoc Int_lower2 closure_mono closure_open_Int_superset inf.orderE open_interior) finally show ?thesis by (simp add: closure_mono dual_order.antisym) qed subsection\<^marker>\tag unimportant\ \Some obvious but surprisingly hard simplex lemmas\ lemma simplex: assumes "finite S" and "0 \ S" shows "convex hull (insert 0 S) = {y. \u. (\x\S. 0 \ u x) \ sum u S \ 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof (simp add: convex_hull_finite set_eq_iff assms, safe) fix x and u :: "'a \ real" assume "0 \ u 0" "\x\S. 0 \ u x" "u 0 + sum u S = 1" then show "\v. (\x\S. 0 \ v x) \ sum v S \ 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by force next fix x and u :: "'a \ real" assume "\x\S. 0 \ u x" "sum u S \ 1" then show "\v. 0 \ v 0 \ (\x\S. 0 \ v x) \ v 0 + sum v S = 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by (rule_tac x="\x. if x = 0 then 1 - sum u S else u x" in exI) (auto simp: sum_delta_notmem assms if_smult) qed lemma substd_simplex: assumes d: "d \ Basis" shows "convex hull (insert 0 d) = {x. (\i\Basis. 0 \ x\i) \ (\i\d. x\i) \ 1 \ (\i\Basis. i \ d \ x\i = 0)}" (is "convex hull (insert 0 ?p) = ?s") proof - let ?D = d have "0 \ ?p" using assms by (auto simp: image_def) from d have "finite d" by (blast intro: finite_subset finite_Basis) show ?thesis unfolding simplex[OF \finite d\ \0 \ ?p\] proof (intro set_eqI; safe) fix u :: "'a \ real" assume as: "\x\?D. 0 \ u x" "sum u ?D \ 1" let ?x = "(\x\?D. u x *\<^sub>R x)" have ind: "\i\Basis. i \ d \ u i = ?x \ i" and notind: "(\i\Basis. i \ d \ ?x \ i = 0)" using substdbasis_expansion_unique[OF assms] by blast+ then have **: "sum u ?D = sum ((\) ?x) ?D" using assms by (auto intro!: sum.cong) show "0 \ ?x \ i" if "i \ Basis" for i using as(1) ind notind that by fastforce show "sum ((\) ?x) ?D \ 1" using "**" as(2) by linarith show "?x \ i = 0" if "i \ Basis" "i \ d" for i using notind that by blast next fix x assume "\i\Basis. 0 \ x \ i" "sum ((\) x) ?D \ 1" "(\i\Basis. i \ d \ x \ i = 0)" with d show "\u. (\x\?D. 0 \ u x) \ sum u ?D \ 1 \ (\x\?D. u x *\<^sub>R x) = x" unfolding substdbasis_expansion_unique[OF assms] by (rule_tac x="inner x" in exI) auto qed qed lemma std_simplex: "convex hull (insert 0 Basis) = {x::'a::euclidean_space. (\i\Basis. 0 \ x\i) \ sum (\i. x\i) Basis \ 1}" using substd_simplex[of Basis] by auto lemma interior_std_simplex: "interior (convex hull (insert 0 Basis)) = {x::'a::euclidean_space. (\i\Basis. 0 < x\i) \ sum (\i. x\i) Basis < 1}" unfolding set_eq_iff mem_interior std_simplex proof (intro allI iffI CollectI; clarify) fix x :: 'a fix e assume "e > 0" and as: "ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" show "(\i\Basis. 0 < x \ i) \ sum ((\) x) Basis < 1" proof safe fix i :: 'a assume i: "i \ Basis" then show "0 < x \ i" using as[THEN subsetD[where c="x - (e/2) *\<^sub>R i"]] and \e > 0\ by (force simp add: inner_simps) next have **: "dist x (x + (e/2) *\<^sub>R (SOME i. i\Basis)) < e" using \e > 0\ unfolding dist_norm by (auto intro!: mult_strict_left_mono simp: SOME_Basis) have "\i. i \ Basis \ (x + (e/2) *\<^sub>R (SOME i. i\Basis)) \ i = x\i + (if i = (SOME i. i\Basis) then e/2 else 0)" by (auto simp: SOME_Basis inner_Basis inner_simps) then have *: "sum ((\) (x + (e/2) *\<^sub>R (SOME i. i\Basis))) Basis = sum (\i. x\i + (if (SOME i. i\Basis) = i then e/2 else 0)) Basis" by (auto simp: intro!: sum.cong) have "sum ((\) x) Basis < sum ((\) (x + (e/2) *\<^sub>R (SOME i. i\Basis))) Basis" using \e > 0\ DIM_positive by (auto simp: SOME_Basis sum.distrib *) also have "\ \ 1" using ** as by force finally show "sum ((\) x) Basis < 1" by auto qed next fix x :: 'a assume as: "\i\Basis. 0 < x \ i" "sum ((\) x) Basis < 1" obtain a :: 'b where "a \ UNIV" using UNIV_witness .. let ?d = "(1 - sum ((\) x) Basis) / real (DIM('a))" show "\e>0. ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" proof (rule_tac x="min (Min (((\) x) ` Basis)) D" for D in exI, intro conjI subsetI CollectI) fix y assume y: "y \ ball x (min (Min ((\) x ` Basis)) ?d)" have "sum ((\) y) Basis \ sum (\i. x\i + ?d) Basis" proof (rule sum_mono) fix i :: 'a assume i: "i \ Basis" have "\y\i - x\i\ \ norm (y - x)" by (metis Basis_le_norm i inner_commute inner_diff_right) also have "... < ?d" using y by (simp add: dist_norm norm_minus_commute) finally have "\y\i - x\i\ < ?d" . then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant by (auto simp add: Suc_le_eq) finally show "sum ((\) y) Basis \ 1" . show "(\i\Basis. 0 \ y \ i)" proof safe fix i :: 'a assume i: "i \ Basis" have "norm (x - y) < Min (((\) x) ` Basis)" using y by (auto simp: dist_norm less_eq_real_def) also have "... \ x\i" using i by auto finally have "norm (x - y) < x\i" . then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i] by (auto simp: inner_simps) qed next have "Min (((\) x) ` Basis) > 0" using as by simp moreover have "?d > 0" using as by (auto simp: Suc_le_eq) ultimately show "0 < min (Min ((\) x ` Basis)) ((1 - sum ((\) x) Basis) / real DIM('a))" by linarith qed qed lemma interior_std_simplex_nonempty: obtains a :: "'a::euclidean_space" where "a \ interior(convex hull (insert 0 Basis))" proof - let ?D = "Basis :: 'a set" let ?a = "sum (\b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis" { fix i :: 'a assume i: "i \ Basis" have "?a \ i = inverse (2 * real DIM('a))" by (rule trans[of _ "sum (\j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) (simp_all add: sum.If_cases i) } note ** = this show ?thesis proof show "?a \ interior(convex hull (insert 0 Basis))" unfolding interior_std_simplex mem_Collect_eq proof safe fix i :: 'a assume i: "i \ Basis" show "0 < ?a \ i" unfolding **[OF i] by (auto simp add: Suc_le_eq) next have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real DIM('a))) ?D" by (auto intro: sum.cong) also have "\ < 1" unfolding sum_constant divide_inverse[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) ?D < 1" by auto qed qed qed lemma rel_interior_substd_simplex: assumes D: "D \ Basis" shows "rel_interior (convex hull (insert 0 D)) = {x::'a::euclidean_space. (\i\D. 0 < x\i) \ (\i\D. x\i) < 1 \ (\i\Basis. i \ D \ x\i = 0)}" (is "_ = ?s") proof - have "finite D" using D finite_Basis finite_subset by blast show ?thesis proof (cases "D = {}") case True then show ?thesis using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto next case False have h0: "affine hull (convex hull (insert 0 D)) = {x::'a::euclidean_space. (\i\Basis. i \ D \ x\i = 0)}" using affine_hull_convex_hull affine_hull_substd_basis assms by auto have aux: "\x::'a. \i\Basis. (\i\D. 0 \ x\i) \ (\i\Basis. i \ D \ x\i = 0) \ 0 \ x\i" by auto { fix x :: "'a::euclidean_space" assume x: "x \ rel_interior (convex hull (insert 0 D))" then obtain e where "e > 0" and "ball x e \ {xa. (\i\Basis. i \ D \ xa\i = 0)} \ convex hull (insert 0 D)" using mem_rel_interior_ball[of x "convex hull (insert 0 D)"] h0 by auto then have as: "\y. \dist x y < e \ (\i\Basis. i \ D \ y\i = 0)\ \ (\i\D. 0 \ y \ i) \ sum ((\) y) D \ 1" using assms by (force simp: substd_simplex) have x0: "(\i\Basis. i \ D \ x\i = 0)" using x rel_interior_subset substd_simplex[OF assms] by auto have "(\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x\i = 0)" proof (intro conjI ballI) fix i :: 'a assume "i \ D" then have "\j\D. 0 \ (x - (e/2) *\<^sub>R i) \ j" using D \e > 0\ x0 by (intro as[THEN conjunct1]) (force simp: dist_norm inner_simps inner_Basis) then show "0 < x \ i" using \e > 0\ \i \ D\ D by (force simp: inner_simps inner_Basis) next obtain a where a: "a \ D" using \D \ {}\ by auto then have **: "dist x (x + (e/2) *\<^sub>R a) < e" using \e > 0\ norm_Basis[of a] D by (auto simp: dist_norm) have "\i. i \ Basis \ (x + (e/2) *\<^sub>R a) \ i = x\i + (if i = a then e/2 else 0)" using a D by (auto simp: inner_simps inner_Basis) then have *: "sum ((\) (x + (e/2) *\<^sub>R a)) D = sum (\i. x\i + (if a = i then e/2 else 0)) D" using D by (intro sum.cong) auto have "a \ Basis" using \a \ D\ D by auto then have h1: "(\i\Basis. i \ D \ (x + (e/2) *\<^sub>R a) \ i = 0)" using x0 D \a\D\ by (auto simp add: inner_add_left inner_Basis) have "sum ((\) x) D < sum ((\) (x + (e/2) *\<^sub>R a)) D" using \e > 0\ \a \ D\ \finite D\ by (auto simp add: * sum.distrib) also have "\ \ 1" using ** h1 as[rule_format, of "x + (e/2) *\<^sub>R a"] by auto finally show "sum ((\) x) D < 1" "\i. i\Basis \ i \ D \ x\i = 0" using x0 by auto qed } moreover { fix x :: "'a::euclidean_space" assume as: "x \ ?s" have "\i. 0 < x\i \ 0 = x\i \ 0 \ x\i" by auto moreover have "\i. i \ D \ i \ D" by auto ultimately have "\i. (\i\D. 0 < x\i) \ (\i. i \ D \ x\i = 0) \ 0 \ x\i" by metis then have h2: "x \ convex hull (insert 0 D)" using as assms by (force simp add: substd_simplex) obtain a where a: "a \ D" using \D \ {}\ by auto define d where "d \ (1 - sum ((\) x) D) / real (card D)" have "\e>0. ball x e \ {x. \i\Basis. i \ D \ x \ i = 0} \ convex hull insert 0 D" unfolding substd_simplex[OF assms] proof (intro exI; safe) have "0 < card D" using \D \ {}\ \finite D\ by (simp add: card_gt_0_iff) have "Min (((\) x) ` D) > 0" using as \D \ {}\ \finite D\ by (simp) moreover have "d > 0" using as \0 < card D\ by (auto simp: d_def) ultimately show "min (Min (((\) x) ` D)) d > 0" by auto fix y :: 'a assume y2: "\i\Basis. i \ D \ y\i = 0" assume "y \ ball x (min (Min ((\) x ` D)) d)" then have y: "dist x y < min (Min ((\) x ` D)) d" by auto have "sum ((\) y) D \ sum (\i. x\i + d) D" proof (rule sum_mono) fix i assume "i \ D" with D have i: "i \ Basis" by auto have "\y\i - x\i\ \ norm (y - x)" by (metis i inner_commute inner_diff_right norm_bound_Basis_le order_refl) also have "... < d" by (metis dist_norm min_less_iff_conj norm_minus_commute y) finally have "\y\i - x\i\ < d" . then show "y \ i \ x \ i + d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant d_def using \0 < card D\ by auto finally show "sum ((\) y) D \ 1" . fix i :: 'a assume i: "i \ Basis" then show "0 \ y\i" proof (cases "i\D") case True have "norm (x - y) < x\i" using y Min_gr_iff[of "(\) x ` D" "norm (x - y)"] \0 < card D\ \i \ D\ by (simp add: dist_norm card_gt_0_iff) then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] by (auto simp: inner_simps) qed (use y2 in auto) qed then have "x \ rel_interior (convex hull (insert 0 D))" using h0 h2 rel_interior_ball by force } ultimately have "\x. x \ rel_interior (convex hull insert 0 D) \ x \ {x. (\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x \ i = 0)}" by blast then show ?thesis by (rule set_eqI) qed qed lemma rel_interior_substd_simplex_nonempty: assumes "D \ {}" and "D \ Basis" obtains a :: "'a::euclidean_space" where "a \ rel_interior (convex hull (insert 0 D))" proof - let ?a = "sum (\b::'a::euclidean_space. inverse (2 * real (card D)) *\<^sub>R b) D" have "finite D" using assms finite_Basis infinite_super by blast then have d1: "0 < real (card D)" using \D \ {}\ by auto { fix i assume "i \ D" have "?a \ i = sum (\j. if i = j then inverse (2 * real (card D)) else 0) D" unfolding inner_sum_left using \i \ D\ by (auto simp: inner_Basis subsetD[OF assms(2)] intro: sum.cong) also have "... = inverse (2 * real (card D))" using \i \ D\ \finite D\ by auto finally have "?a \ i = inverse (2 * real (card D))" . } note ** = this show ?thesis proof show "?a \ rel_interior (convex hull (insert 0 D))" unfolding rel_interior_substd_simplex[OF assms(2)] proof safe fix i assume "i \ D" have "0 < inverse (2 * real (card D))" using d1 by auto also have "\ = ?a \ i" using **[of i] \i \ D\ by auto finally show "0 < ?a \ i" by auto next have "sum ((\) ?a) D = sum (\i. inverse (2 * real (card D))) D" by (rule sum.cong) (rule refl, rule **) also have "\ < 1" unfolding sum_constant divide_real_def[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) D < 1" by auto next fix i assume "i \ Basis" and "i \ D" have "?a \ span D" proof (rule span_sum[of D "(\b. b /\<^sub>R (2 * real (card D)))" D]) { fix x :: "'a::euclidean_space" assume "x \ D" then have "x \ span D" using span_base[of _ "D"] by auto then have "x /\<^sub>R (2 * real (card D)) \ span D" using span_mul[of x "D" "(inverse (real (card D)) / 2)"] by auto } then show "\x. x\D \ x /\<^sub>R (2 * real (card D)) \ span D" by auto qed then show "?a \ i = 0 " using \i \ D\ unfolding span_substd_basis[OF assms(2)] using \i \ Basis\ by auto qed qed qed subsection\<^marker>\tag unimportant\ \Relative interior of convex set\ lemma rel_interior_convex_nonempty_aux: fixes S :: "'n::euclidean_space set" assumes "convex S" and "0 \ S" shows "rel_interior S \ {}" proof (cases "S = {0}") case True then show ?thesis using rel_interior_sing by auto next case False obtain B where B: "independent B \ B \ S \ S \ span B \ card B = dim S" using basis_exists[of S] by metis then have "B \ {}" using B assms \S \ {0}\ span_empty by auto have "insert 0 B \ span B" using subspace_span[of B] subspace_0[of "span B"] span_superset by auto then have "span (insert 0 B) \ span B" using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast then have "convex hull insert 0 B \ span B" using convex_hull_subset_span[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) \ span B" using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast then have *: "span (convex hull insert 0 B) = span B" using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) = span S" using B span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto moreover have "0 \ affine hull (convex hull insert 0 B)" using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S" using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"] assms hull_subset[of S] by auto obtain d and f :: "'n \ 'n" where fd: "card d = card B" "linear f" "f ` B = d" "f ` span B = {x. \i\Basis. i \ d \ x \ i = (0::real)} \ inj_on f (span B)" and d: "d \ Basis" using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B by auto then have "bounded_linear f" using linear_conv_bounded_linear by auto have "d \ {}" using fd B \B \ {}\ by auto have "insert 0 d = f ` (insert 0 B)" using fd linear_0 by auto then have "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))" using convex_hull_linear_image[of f "(insert 0 d)"] convex_hull_linear_image[of f "(insert 0 B)"] \linear f\ by auto moreover have "rel_interior (f ` (convex hull insert 0 B)) = f ` rel_interior (convex hull insert 0 B)" proof (rule rel_interior_injective_on_span_linear_image[OF \bounded_linear f\]) show "inj_on f (span (convex hull insert 0 B))" using fd * by auto qed ultimately have "rel_interior (convex hull insert 0 B) \ {}" using rel_interior_substd_simplex_nonempty[OF \d \ {}\ d] by fastforce moreover have "convex hull (insert 0 B) \ S" using B assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq by auto ultimately show ?thesis using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto qed lemma rel_interior_eq_empty: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior S = {} \ S = {}" proof - { assume "S \ {}" then obtain a where "a \ S" by auto then have "0 \ (+) (-a) ` S" using assms exI[of "(\x. x \ S \ - a + x = 0)" a] by auto then have "rel_interior ((+) (-a) ` S) \ {}" using rel_interior_convex_nonempty_aux[of "(+) (-a) ` S"] convex_translation[of S "-a"] assms by auto then have "rel_interior S \ {}" using rel_interior_translation [of "- a"] by simp } then show ?thesis by auto qed lemma interior_simplex_nonempty: fixes S :: "'N :: euclidean_space set" assumes "independent S" "finite S" "card S = DIM('N)" obtains a where "a \ interior (convex hull (insert 0 S))" proof - have "affine hull (insert 0 S) = UNIV" by (simp add: hull_inc affine_hull_span_0 dim_eq_full[symmetric] assms(1) assms(3) dim_eq_card_independent) moreover have "rel_interior (convex hull insert 0 S) \ {}" using rel_interior_eq_empty [of "convex hull (insert 0 S)"] by auto ultimately have "interior (convex hull insert 0 S) \ {}" by (simp add: rel_interior_interior) with that show ?thesis by auto qed lemma convex_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "convex (rel_interior S)" proof - { fix x y and u :: real assume assm: "x \ rel_interior S" "y \ rel_interior S" "0 \ u" "u \ 1" then have "x \ S" using rel_interior_subset by auto have "x - u *\<^sub>R (x-y) \ rel_interior S" proof (cases "0 = u") case False then have "0 < u" using assm by auto then show ?thesis using assm rel_interior_convex_shrink[of S y x u] assms \x \ S\ by auto next case True then show ?thesis using assm by auto qed then have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ rel_interior S" by (simp add: algebra_simps) } then show ?thesis unfolding convex_alt by auto qed lemma convex_closure_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "closure (rel_interior S) = closure S" proof - have h1: "closure (rel_interior S) \ closure S" using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto show ?thesis proof (cases "S = {}") case False then obtain a where a: "a \ rel_interior S" using rel_interior_eq_empty assms by auto { fix x assume x: "x \ closure S" { assume "x = a" then have "x \ closure (rel_interior S)" using a unfolding closure_def by auto } moreover { assume "x \ a" { fix e :: real assume "e > 0" define e1 where "e1 = min 1 (e/norm (x - a))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (x - a) \ e" using \x \ a\ \e > 0\ le_divide_eq[of e1 e "norm (x - a)"] by simp_all then have *: "x - e1 *\<^sub>R (x - a) \ rel_interior S" using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def by auto have "\y. y \ rel_interior S \ y \ x \ dist y x \ e" using "*" \x \ a\ e1 by force } then have "x islimpt rel_interior S" unfolding islimpt_approachable_le by auto then have "x \ closure(rel_interior S)" unfolding closure_def by auto } ultimately have "x \ closure(rel_interior S)" by auto } then show ?thesis using h1 by auto qed auto qed lemma rel_interior_same_affine_hull: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "affine hull (rel_interior S) = affine hull S" by (metis assms closure_same_affine_hull convex_closure_rel_interior) lemma rel_interior_aff_dim: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "aff_dim (rel_interior S) = aff_dim S" by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull) lemma rel_interior_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (rel_interior S) = rel_interior S" proof - have "openin (top_of_set (affine hull (rel_interior S))) (rel_interior S)" using openin_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto then show ?thesis using rel_interior_def by auto qed lemma rel_interior_rel_open: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_open (rel_interior S)" unfolding rel_open_def using rel_interior_rel_interior assms by auto lemma convex_rel_interior_closure_aux: fixes x y z :: "'n::euclidean_space" assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y" obtains e where "0 < e" "e < 1" "z = y - e *\<^sub>R (y - x)" proof - define e where "e = a / (a + b)" have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)" using assms by (simp add: eq_vector_fraction_iff) also have "\ = (1 / (a + b)) *\<^sub>R (a *\<^sub>R x + b *\<^sub>R y)" using assms scaleR_cancel_left[of "1/(a+b)" "(a + b) *\<^sub>R z" "a *\<^sub>R x + b *\<^sub>R y"] by auto also have "\ = y - e *\<^sub>R (y-x)" using e_def assms by (simp add: divide_simps vector_fraction_eq_iff) (simp add: algebra_simps) finally have "z = y - e *\<^sub>R (y-x)" by auto moreover have "e > 0" "e < 1" using e_def assms by auto ultimately show ?thesis using that[of e] by auto qed lemma convex_rel_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (closure S) = rel_interior S" proof (cases "S = {}") case True then show ?thesis using assms rel_interior_eq_empty by auto next case False have "rel_interior (closure S) \ rel_interior S" using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset by auto moreover { fix z assume z: "z \ rel_interior (closure S)" obtain x where x: "x \ rel_interior S" using \S \ {}\ assms rel_interior_eq_empty by auto have "z \ rel_interior S" proof (cases "x = z") case True then show ?thesis using x by auto next case False obtain e where e: "e > 0" "cball z e \ affine hull closure S \ closure S" using z rel_interior_cball[of "closure S"] by auto hence *: "0 < e/norm(z-x)" using e False by auto define y where "y = z + (e/norm(z-x)) *\<^sub>R (z-x)" have yball: "y \ cball z e" using y_def dist_norm[of z y] e by auto have "x \ affine hull closure S" using x rel_interior_subset_closure hull_inc[of x "closure S"] by blast moreover have "z \ affine hull closure S" using z rel_interior_subset hull_subset[of "closure S"] by blast ultimately have "y \ affine hull closure S" using y_def affine_affine_hull[of "closure S"] mem_affine_3_minus [of "affine hull closure S" z z x "e/norm(z-x)"] by auto then have "y \ closure S" using e yball by auto have "(1 + (e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y" using y_def by (simp add: algebra_simps) then obtain e1 where "0 < e1" "e1 < 1" "z = y - e1 *\<^sub>R (y - x)" using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y] by (auto simp add: algebra_simps) then show ?thesis using rel_interior_closure_convex_shrink assms x \y \ closure S\ by fastforce qed } ultimately show ?thesis by auto qed lemma convex_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "interior (closure S) = interior S" using closure_aff_dim[of S] interior_rel_interior_gen[of S] interior_rel_interior_gen[of "closure S"] convex_rel_interior_closure[of S] assms by auto +lemma open_subset_closure_of_interval: + assumes "open U" "is_interval S" + shows "U \ closure S \ U \ interior S" + by (metis assms convex_interior_closure is_interval_convex open_subset_interior) + lemma closure_eq_rel_interior_eq: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 = rel_interior S2" by (metis convex_rel_interior_closure convex_closure_rel_interior assms) lemma closure_eq_between: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 \ S2 \ S2 \ closure S1" (is "?A \ ?B") proof assume ?A then show ?B by (metis assms closure_subset convex_rel_interior_closure rel_interior_subset) next assume ?B then have "closure S1 \ closure S2" by (metis assms(1) convex_closure_rel_interior closure_mono) moreover from \?B\ have "closure S1 \ closure S2" by (metis closed_closure closure_minimal) ultimately show ?A .. qed lemma open_inter_closure_rel_interior: fixes S A :: "'n::euclidean_space set" assumes "convex S" and "open A" shows "A \ closure S = {} \ A \ rel_interior S = {}" by (metis assms convex_closure_rel_interior open_Int_closure_eq_empty) lemma rel_interior_open_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False then have "open_segment a b = affine hull {a, b} \ ball ((a + b) /\<^sub>R 2) (norm (b - a) / 2)" by (simp add: open_segment_as_ball) then show ?thesis unfolding rel_interior_eq openin_open by (metis Elementary_Metric_Spaces.open_ball False affine_hull_open_segment) qed lemma rel_interior_closed_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(closed_segment a b) = (if a = b then {a} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by auto next case False then show ?thesis by simp (metis closure_open_segment convex_open_segment convex_rel_interior_closure rel_interior_open_segment) qed lemmas rel_interior_segment = rel_interior_closed_segment rel_interior_open_segment subsection\The relative frontier of a set\ definition\<^marker>\tag important\ "rel_frontier S = closure S - rel_interior S" lemma rel_frontier_empty [simp]: "rel_frontier {} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_eq_empty: fixes S :: "'n::euclidean_space set" shows "rel_frontier S = {} \ affine S" unfolding rel_frontier_def using rel_interior_subset_closure by (auto simp add: rel_interior_eq_closure [symmetric]) lemma rel_frontier_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier {a} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_affine_hull: fixes S :: "'a::euclidean_space set" shows "rel_frontier S \ affine hull S" using closure_affine_hull rel_frontier_def by fastforce lemma rel_frontier_cball [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier(cball a r) = (if r = 0 then {} else sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by (force simp: sphere_def) next case equal then show ?thesis by simp next case greater then show ?thesis by simp (metis centre_in_ball empty_iff frontier_cball frontier_def interior_cball interior_rel_interior_gen rel_frontier_def) qed lemma rel_frontier_translation: fixes a :: "'a::euclidean_space" shows "rel_frontier((\x. a + x) ` S) = (\x. a + x) ` (rel_frontier S)" by (simp add: rel_frontier_def translation_diff rel_interior_translation closure_translation) lemma rel_frontier_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ rel_frontier S = frontier S" by (metis frontier_def interior_rel_interior_gen rel_frontier_def) lemma rel_frontier_frontier: fixes S :: "'n::euclidean_space set" shows "affine hull S = UNIV \ rel_frontier S = frontier S" by (simp add: frontier_def rel_frontier_def rel_interior_interior) lemma closest_point_in_rel_frontier: "\closed S; S \ {}; x \ affine hull S - rel_interior S\ \ closest_point S x \ rel_frontier S" by (simp add: closest_point_in_rel_interior closest_point_in_set rel_frontier_def) lemma closed_rel_frontier [iff]: fixes S :: "'n::euclidean_space set" shows "closed (rel_frontier S)" proof - have *: "closedin (top_of_set (affine hull S)) (closure S - rel_interior S)" by (simp add: closed_subset closedin_diff closure_affine_hull openin_rel_interior) show ?thesis proof (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"]) show "closedin (top_of_set (affine hull S)) (rel_frontier S)" by (simp add: "*" rel_frontier_def) qed simp qed lemma closed_rel_boundary: fixes S :: "'n::euclidean_space set" shows "closed S \ closed(S - rel_interior S)" by (metis closed_rel_frontier closure_closed rel_frontier_def) lemma compact_rel_boundary: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(S - rel_interior S)" by (metis bounded_diff closed_rel_boundary closure_eq compact_closure compact_imp_closed) lemma bounded_rel_frontier: fixes S :: "'n::euclidean_space set" shows "bounded S \ bounded(rel_frontier S)" by (simp add: bounded_closure bounded_diff rel_frontier_def) lemma compact_rel_frontier_bounded: fixes S :: "'n::euclidean_space set" shows "bounded S \ compact(rel_frontier S)" using bounded_rel_frontier closed_rel_frontier compact_eq_bounded_closed by blast lemma compact_rel_frontier: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(rel_frontier S)" by (meson compact_eq_bounded_closed compact_rel_frontier_bounded) lemma convex_same_rel_interior_closure: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ closure S = closure T" by (simp add: closure_eq_rel_interior_eq) lemma convex_same_rel_interior_closure_straddle: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ rel_interior S \ T \ T \ closure S" by (simp add: closure_eq_between convex_same_rel_interior_closure) lemma convex_rel_frontier_aff_dim: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" and "S2 \ {}" and "S1 \ rel_frontier S2" shows "aff_dim S1 < aff_dim S2" proof - have "S1 \ closure S2" using assms unfolding rel_frontier_def by auto then have *: "affine hull S1 \ affine hull S2" using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by blast then have "aff_dim S1 \ aff_dim S2" using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] aff_dim_subset[of "affine hull S1" "affine hull S2"] by auto moreover { assume eq: "aff_dim S1 = aff_dim S2" then have "S1 \ {}" using aff_dim_empty[of S1] aff_dim_empty[of S2] \S2 \ {}\ by auto have **: "affine hull S1 = affine hull S2" by (simp_all add: * eq \S1 \ {}\ affine_dim_equal) obtain a where a: "a \ rel_interior S1" using \S1 \ {}\ rel_interior_eq_empty assms by auto obtain T where T: "open T" "a \ T \ S1" "T \ affine hull S1 \ S1" using mem_rel_interior[of a S1] a by auto then have "a \ T \ closure S2" using a assms unfolding rel_frontier_def by auto then obtain b where b: "b \ T \ rel_interior S2" using open_inter_closure_rel_interior[of S2 T] assms T by auto then have "b \ affine hull S1" using rel_interior_subset hull_subset[of S2] ** by auto then have "b \ S1" using T b by auto then have False using b assms unfolding rel_frontier_def by auto } ultimately show ?thesis using less_le by auto qed lemma convex_rel_interior_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "z \ rel_interior S" shows "\x\affine hull S. \m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" proof - obtain e1 where e1: "e1 > 0 \ cball z e1 \ affine hull S \ S" using mem_rel_interior_cball[of z S] assms by auto { fix x assume x: "x \ affine hull S" { assume "x \ z" define m where "m = 1 + e1/norm(x-z)" hence "m > 1" using e1 \x \ z\ by auto { fix e assume e: "e > 1 \ e \ m" have "z \ affine hull S" using assms rel_interior_subset hull_subset[of S] by auto then have *: "(1 - e)*\<^sub>R x + e *\<^sub>R z \ affine hull S" using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x by auto have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x - z))" by (simp add: algebra_simps) also have "\ = (e - 1) * norm (x-z)" using norm_scaleR e by auto also have "\ \ (m - 1) * norm (x - z)" using e mult_right_mono[of _ _ "norm(x-z)"] by auto also have "\ = (e1 / norm (x - z)) * norm (x - z)" using m_def by auto also have "\ = e1" using \x \ z\ e1 by simp finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \ e1" by auto have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \ cball z e1" using m_def ** unfolding cball_def dist_norm by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x+ e *\<^sub>R z \ S" using e * e1 by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" using \m> 1 \ by auto } moreover { assume "x = z" define m where "m = 1 + e1" then have "m > 1" using e1 by auto { fix e assume e: "e > 1 \ e \ m" then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e1 x \x = z\ by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using \m > 1\ by auto } ultimately have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" by blast } then show ?thesis by auto qed lemma convex_rel_interior_if2: fixes S :: "'n::euclidean_space set" assumes "convex S" assumes "z \ rel_interior S" shows "\x\affine hull S. \e. e > 1 \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_if[of S z] assms by auto lemma convex_rel_interior_only_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" assumes "\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" shows "z \ rel_interior S" proof - obtain x where x: "x \ rel_interior S" using rel_interior_eq_empty assms by auto then have "x \ S" using rel_interior_subset by auto then obtain e where e: "e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using assms by auto define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z" then have "y \ S" using e by auto define e1 where "e1 = 1/e" then have "0 < e1 \ e1 < 1" using e by auto then have "z =y - (1 - e1) *\<^sub>R (y - x)" using e1_def y_def by (auto simp add: algebra_simps) then show ?thesis using rel_interior_convex_shrink[of S x y "1-e1"] \0 < e1 \ e1 < 1\ \y \ S\ x assms by auto qed lemma convex_rel_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S "affine"] convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_rel_interior_iff2: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\affine hull S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S] convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "z \ interior S \ (\x. \e. e > 0 \ z + e *\<^sub>R x \ S)" proof (cases "aff_dim S = int DIM('n)") case False { assume "z \ interior S" then have False using False interior_rel_interior_gen[of S] by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0 \ z + e1 *\<^sub>R (x - z) \ S" using r by auto obtain e2 where e2: "e2 > 0 \ z + e2 *\<^sub>R (z - x) \ S" using r by auto define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)" then have x1: "x1 \ affine hull S" using e1 hull_subset[of S] by auto define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)" then have x2: "x2 \ affine hull S" using e2 hull_subset[of S] by auto have *: "e1/(e1+e2) + e2/(e1+e2) = 1" using add_divide_distrib[of e1 e2 "e1+e2"] e1 e2 by simp then have "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2" by (simp add: x1_def x2_def algebra_simps) (simp add: "*" pth_8) then have z: "z \ affine hull S" using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"] x1 x2 affine_affine_hull[of S] * by auto have "x1 - x2 = (e1 + e2) *\<^sub>R (x - z)" using x1_def x2_def by (auto simp add: algebra_simps) then have "x = z+(1/(e1+e2)) *\<^sub>R (x1-x2)" using e1 e2 by simp then have "x \ affine hull S" using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"] x1 x2 z affine_affine_hull[of S] by auto } then have "affine hull S = UNIV" by auto then have "aff_dim S = int DIM('n)" using aff_dim_affine_hull[of S] by (simp) then have False using False by auto } ultimately show ?thesis by auto next case True then have "S \ {}" using aff_dim_empty[of S] by auto have *: "affine hull S = UNIV" using True affine_hull_UNIV by auto { assume "z \ interior S" then have "z \ rel_interior S" using True interior_rel_interior_gen[of S] by auto then have **: "\x. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ * by auto fix x obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \ S" using **[rule_format, of "z-x"] by auto define e where [abs_def]: "e = e1 - 1" then have "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z = z + e *\<^sub>R x" by (simp add: algebra_simps) then have "e > 0" "z + e *\<^sub>R x \ S" using e1 e_def by auto then have "\e. e > 0 \ z + e *\<^sub>R x \ S" by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \ S" using r[rule_format, of "z-x"] by auto define e where "e = e1 + 1" then have "z + e1 *\<^sub>R (z - x) = (1 - e) *\<^sub>R x + e *\<^sub>R z" by (simp add: algebra_simps) then have "e > 1" "(1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using e1 e_def by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" by auto } then have "z \ rel_interior S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ by auto then have "z \ interior S" using True interior_rel_interior_gen[of S] by auto } ultimately show ?thesis by auto qed subsubsection\<^marker>\tag unimportant\ \Relative interior and closure under common operations\ lemma rel_interior_inter_aux: "\{rel_interior S |S. S \ I} \ \I" proof - { fix y assume "y \ \{rel_interior S |S. S \ I}" then have y: "\S \ I. y \ rel_interior S" by auto { fix S assume "S \ I" then have "y \ S" using rel_interior_subset y by auto } then have "y \ \I" by auto } then show ?thesis by auto qed lemma convex_closure_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" proof - obtain x where x: "\S\I. x \ rel_interior S" using assms by auto { fix y assume "y \ \{closure S |S. S \ I}" then have y: "\S \ I. y \ closure S" by auto { assume "y = x" then have "y \ closure (\{rel_interior S |S. S \ I})" using x closure_subset[of "\{rel_interior S |S. S \ I}"] by auto } moreover { assume "y \ x" { fix e :: real assume e: "e > 0" define e1 where "e1 = min 1 (e/norm (y - x))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (y - x) \ e" using \y \ x\ \e > 0\ le_divide_eq[of e1 e "norm (y - x)"] by simp_all define z where "z = y - e1 *\<^sub>R (y - x)" { fix S assume "S \ I" then have "z \ rel_interior S" using rel_interior_closure_convex_shrink[of S x y e1] assms x y e1 z_def by auto } then have *: "z \ \{rel_interior S |S. S \ I}" by auto have "\z. z \ \{rel_interior S |S. S \ I} \ z \ y \ dist z y \ e" using \y \ x\ z_def * e1 e dist_norm[of z y] by (rule_tac x="z" in exI) auto } then have "y islimpt \{rel_interior S |S. S \ I}" unfolding islimpt_approachable_le by blast then have "y \ closure (\{rel_interior S |S. S \ I})" unfolding closure_def by auto } ultimately have "y \ closure (\{rel_interior S |S. S \ I})" by auto } then show ?thesis by auto qed lemma convex_closure_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\I) = \{closure S |S. S \ I}" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_inter_rel_interior_same_closure: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\{rel_interior S |S. S \ I}) = closure (\I)" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "rel_interior (\I) \ \{rel_interior S |S. S \ I}" proof - have "convex (\I)" using assms convex_Inter by auto moreover have "convex (\{rel_interior S |S. S \ I})" using assms convex_rel_interior by (force intro: convex_Inter) ultimately have "rel_interior (\{rel_interior S |S. S \ I}) = rel_interior (\I)" using convex_inter_rel_interior_same_closure assms closure_eq_rel_interior_eq[of "\{rel_interior S |S. S \ I}" "\I"] by blast then show ?thesis using rel_interior_subset[of "\{rel_interior S |S. S \ I}"] by auto qed lemma convex_rel_interior_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" and "finite I" shows "rel_interior (\I) = \{rel_interior S |S. S \ I}" proof - have "\I \ {}" using assms rel_interior_inter_aux[of I] by auto have "convex (\I)" using convex_Inter assms by auto show ?thesis proof (cases "I = {}") case True then show ?thesis using Inter_empty rel_interior_UNIV by auto next case False { fix z assume z: "z \ \{rel_interior S |S. S \ I}" { fix x assume x: "x \ \I" { fix S assume S: "S \ I" then have "z \ rel_interior S" "x \ S" using z x by auto then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S)" using convex_rel_interior_if[of S z] S assms hull_subset[of S] by auto } then obtain mS where mS: "\S\I. mS S > 1 \ (\e. e > 1 \ e \ mS S \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" by metis define e where "e = Min (mS ` I)" then have "e \ mS ` I" using assms \I \ {}\ by simp then have "e > 1" using mS by auto moreover have "\S\I. e \ mS S" using e_def assms by auto ultimately have "\e > 1. (1 - e) *\<^sub>R x + e *\<^sub>R z \ \I" using mS by auto } then have "z \ rel_interior (\I)" using convex_rel_interior_iff[of "\I" z] \\I \ {}\ \convex (\I)\ by auto } then show ?thesis using convex_rel_interior_inter[of I] assms by auto qed qed lemma convex_closure_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" assumes "rel_interior S \ rel_interior T \ {}" shows "closure (S \ T) = closure S \ closure T" using convex_closure_inter[of "{S,T}"] assms by auto lemma convex_rel_interior_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "rel_interior S \ rel_interior T \ {}" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" using convex_rel_interior_finite_inter[of "{S,T}"] assms by auto lemma convex_affine_closure_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "closure (S \ T) = closure S \ T" by (metis affine_imp_convex assms convex_closure_inter_two rel_interior_affine rel_interior_eq_closure) lemma connected_component_1_gen: fixes S :: "'a :: euclidean_space set" assumes "DIM('a) = 1" shows "connected_component S a b \ closed_segment a b \ S" unfolding connected_component_def by (metis (no_types, lifting) assms subsetD subsetI convex_contains_segment convex_segment(1) ends_in_segment connected_convex_1_gen) lemma connected_component_1: fixes S :: "real set" shows "connected_component S a b \ closed_segment a b \ S" by (simp add: connected_component_1_gen) lemma convex_affine_rel_interior_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "rel_interior (S \ T) = rel_interior S \ T" by (simp add: affine_imp_convex assms convex_rel_interior_inter_two rel_interior_affine) lemma convex_affine_rel_frontier_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "interior S \ T \ {}" shows "rel_frontier(S \ T) = frontier S \ T" using assms unfolding rel_frontier_def frontier_def using convex_affine_closure_Int convex_affine_rel_interior_Int rel_interior_nonempty_interior by fastforce lemma rel_interior_convex_Int_affine: fixes S :: "'a::euclidean_space set" assumes "convex S" "affine T" "interior S \ T \ {}" shows "rel_interior(S \ T) = interior S \ T" by (metis Int_emptyI assms convex_affine_rel_interior_Int empty_iff interior_rel_interior_gen) lemma subset_rel_interior_convex: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "S \ closure T" and "\ S \ rel_frontier T" shows "rel_interior S \ rel_interior T" proof - have *: "S \ closure T = S" using assms by auto have "\ rel_interior S \ rel_frontier T" using closure_mono[of "rel_interior S" "rel_frontier T"] closed_rel_frontier[of T] closure_closed[of S] convex_closure_rel_interior[of S] closure_subset[of S] assms by auto then have "rel_interior S \ rel_interior (closure T) \ {}" using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T] by auto then have "rel_interior S \ rel_interior T = rel_interior (S \ closure T)" using assms convex_closure convex_rel_interior_inter_two[of S "closure T"] convex_rel_interior_closure[of T] by auto also have "\ = rel_interior S" using * by auto finally show ?thesis by auto qed lemma rel_interior_convex_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" shows "f ` (rel_interior S) = rel_interior (f ` S)" proof (cases "S = {}") case True then show ?thesis using assms by auto next case False interpret linear f by fact have *: "f ` (rel_interior S) \ f ` S" unfolding image_mono using rel_interior_subset by auto have "f ` S \ f ` (closure S)" unfolding image_mono using closure_subset by auto also have "\ = f ` (closure (rel_interior S))" using convex_closure_rel_interior assms by auto also have "\ \ closure (f ` (rel_interior S))" using closure_linear_image_subset assms by auto finally have "closure (f ` S) = closure (f ` rel_interior S)" using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure closure_mono[of "f ` rel_interior S" "f ` S"] * by auto then have "rel_interior (f ` S) = rel_interior (f ` rel_interior S)" using assms convex_rel_interior linear_conv_bounded_linear[of f] convex_linear_image[of _ S] convex_linear_image[of _ "rel_interior S"] closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"] by auto then have "rel_interior (f ` S) \ f ` rel_interior S" using rel_interior_subset by auto moreover { fix z assume "z \ f ` rel_interior S" then obtain z1 where z1: "z1 \ rel_interior S" "f z1 = z" by auto { fix x assume "x \ f ` S" then obtain x1 where x1: "x1 \ S" "f x1 = x" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 \ S" using convex_rel_interior_iff[of S z1] \convex S\ x1 z1 by auto moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z" using x1 z1 by (simp add: linear_add linear_scale \linear f\) ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using e by auto } then have "z \ rel_interior (f ` S)" using convex_rel_interior_iff[of "f ` S" z] \convex S\ \linear f\ \S \ {}\ convex_linear_image[of f S] linear_conv_bounded_linear[of f] by auto } ultimately show ?thesis by auto qed lemma rel_interior_convex_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "f -` (rel_interior S) \ {}" shows "rel_interior (f -` S) = f -` (rel_interior S)" proof - interpret linear f by fact have "S \ {}" using assms by auto have nonemp: "f -` S \ {}" by (metis assms(3) rel_interior_subset subset_empty vimage_mono) then have "S \ (range f) \ {}" by auto have conv: "convex (f -` S)" using convex_linear_vimage assms by auto then have "convex (S \ range f)" by (simp add: assms(2) convex_Int convex_linear_image linear_axioms) { fix z assume "z \ f -` (rel_interior S)" then have z: "f z \ rel_interior S" by auto { fix x assume "x \ f -` S" then have "f x \ S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \ S" using convex_rel_interior_iff[of S "f z"] z assms \S \ {}\ by auto moreover have "(1 - e) *\<^sub>R f x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R x + e *\<^sub>R z)" using \linear f\ by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f -` S" using e by auto } then have "z \ rel_interior (f -` S)" using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto } moreover { fix z assume z: "z \ rel_interior (f -` S)" { fix x assume "x \ S \ range f" then obtain y where y: "f y = x" "y \ f -` S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \ f -` S" using convex_rel_interior_iff[of "f -` S" z] z conv by auto moreover have "(1 - e) *\<^sub>R x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R y + e *\<^sub>R z)" using \linear f\ y by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R f z \ S \ range f" using e by auto } then have "f z \ rel_interior (S \ range f)" using \convex (S \ (range f))\ \S \ range f \ {}\ convex_rel_interior_iff[of "S \ (range f)" "f z"] by auto moreover have "affine (range f)" by (simp add: linear_axioms linear_subspace_image subspace_imp_affine) ultimately have "f z \ rel_interior S" using convex_affine_rel_interior_Int[of S "range f"] assms by auto then have "z \ f -` (rel_interior S)" by auto } ultimately show ?thesis by auto qed lemma rel_interior_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" proof (cases "S = {} \ T = {}") case True then show ?thesis by auto next case False then have "S \ {}" "T \ {}" by auto then have ri: "rel_interior S \ {}" "rel_interior T \ {}" using rel_interior_eq_empty assms by auto then have "fst -` rel_interior S \ {}" using fst_vimage_eq_Times[of "rel_interior S"] by auto then have "rel_interior ((fst :: 'n * 'm \ 'n) -` S) = fst -` rel_interior S" using linear_fst \convex S\ rel_interior_convex_linear_preimage[of fst S] by auto then have s: "rel_interior (S \ (UNIV :: 'm set)) = rel_interior S \ UNIV" by (simp add: fst_vimage_eq_Times) from ri have "snd -` rel_interior T \ {}" using snd_vimage_eq_Times[of "rel_interior T"] by auto then have "rel_interior ((snd :: 'n * 'm \ 'm) -` T) = snd -` rel_interior T" using linear_snd \convex T\ rel_interior_convex_linear_preimage[of snd T] by auto then have t: "rel_interior ((UNIV :: 'n set) \ T) = UNIV \ rel_interior T" by (simp add: snd_vimage_eq_Times) from s t have *: "rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T) = rel_interior S \ rel_interior T" by auto have "S \ T = S \ (UNIV :: 'm set) \ (UNIV :: 'n set) \ T" by auto then have "rel_interior (S \ T) = rel_interior ((S \ (UNIV :: 'm set)) \ ((UNIV :: 'n set) \ T))" by auto also have "\ = rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T)" using * ri assms convex_Times by (subst convex_rel_interior_inter_two) auto finally show ?thesis using * by auto qed lemma rel_interior_scaleR: fixes S :: "'n::euclidean_space set" assumes "c \ 0" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" using rel_interior_injective_linear_image[of "((*\<^sub>R) c)" S] linear_conv_bounded_linear[of "(*\<^sub>R) c"] linear_scaleR injective_scaleR[of c] assms by auto lemma rel_interior_convex_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" by (metis assms linear_scaleR rel_interior_convex_linear_image) lemma convex_rel_open_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" shows "convex (((*\<^sub>R) c) ` S) \ rel_open (((*\<^sub>R) c) ` S)" by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def) lemma convex_rel_open_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set) \ rel_open S" and "finite I" shows "convex (\I) \ rel_open (\I)" proof (cases "\{rel_interior S |S. S \ I} = {}") case True then have "\I = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def by auto next case False then have "rel_open (\I)" using assms unfolding rel_open_def using convex_rel_interior_finite_inter[of I] by auto then show ?thesis using convex_Inter assms by auto qed lemma convex_rel_open_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f ` S) \ rel_open (f ` S)" by (metis assms convex_linear_image rel_interior_convex_linear_image rel_open_def) lemma convex_rel_open_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f -` S) \ rel_open (f -` S)" proof (cases "f -` (rel_interior S) = {}") case True then have "f -` S = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def by auto next case False then have "rel_open (f -` S)" using assms unfolding rel_open_def using rel_interior_convex_linear_preimage[of f S] by auto then show ?thesis using convex_linear_vimage assms by auto qed lemma rel_interior_projection: fixes S :: "('m::euclidean_space \ 'n::euclidean_space) set" and f :: "'m::euclidean_space \ 'n::euclidean_space set" assumes "convex S" and "f = (\y. {z. (y, z) \ S})" shows "(y, z) \ rel_interior S \ (y \ rel_interior {y. (f y \ {})} \ z \ rel_interior (f y))" proof - { fix y assume "y \ {y. f y \ {}}" then obtain z where "(y, z) \ S" using assms by auto then have "\x. x \ S \ y = fst x" by auto then obtain x where "x \ S" "y = fst x" by blast then have "y \ fst ` S" unfolding image_def by auto } then have "fst ` S = {y. f y \ {}}" unfolding fst_def using assms by auto then have h1: "fst ` rel_interior S = rel_interior {y. f y \ {}}" using rel_interior_convex_linear_image[of fst S] assms linear_fst by auto { fix y assume "y \ rel_interior {y. f y \ {}}" then have "y \ fst ` rel_interior S" using h1 by auto then have *: "rel_interior S \ fst -` {y} \ {}" by auto moreover have aff: "affine (fst -` {y})" unfolding affine_alt by (simp add: algebra_simps) ultimately have **: "rel_interior (S \ fst -` {y}) = rel_interior S \ fst -` {y}" using convex_affine_rel_interior_Int[of S "fst -` {y}"] assms by auto have conv: "convex (S \ fst -` {y})" using convex_Int assms aff affine_imp_convex by auto { fix x assume "x \ f y" then have "(y, x) \ S \ (fst -` {y})" using assms by auto moreover have "x = snd (y, x)" by auto ultimately have "x \ snd ` (S \ fst -` {y})" by blast } then have "snd ` (S \ fst -` {y}) = f y" using assms by auto then have ***: "rel_interior (f y) = snd ` rel_interior (S \ fst -` {y})" using rel_interior_convex_linear_image[of snd "S \ fst -` {y}"] linear_snd conv by auto { fix z assume "z \ rel_interior (f y)" then have "z \ snd ` rel_interior (S \ fst -` {y})" using *** by auto moreover have "{y} = fst ` rel_interior (S \ fst -` {y})" using * ** rel_interior_subset by auto ultimately have "(y, z) \ rel_interior (S \ fst -` {y})" by force then have "(y,z) \ rel_interior S" using ** by auto } moreover { fix z assume "(y, z) \ rel_interior S" then have "(y, z) \ rel_interior (S \ fst -` {y})" using ** by auto then have "z \ snd ` rel_interior (S \ fst -` {y})" by (metis Range_iff snd_eq_Range) then have "z \ rel_interior (f y)" using *** by auto } ultimately have "\z. (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto } then have h2: "\y z. y \ rel_interior {t. f t \ {}} \ (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto { fix y z assume asm: "(y, z) \ rel_interior S" then have "y \ fst ` rel_interior S" by (metis Domain_iff fst_eq_Domain) then have "y \ rel_interior {t. f t \ {}}" using h1 by auto then have "y \ rel_interior {t. f t \ {}}" and "(z \ rel_interior (f y))" using h2 asm by auto } then show ?thesis using h2 by blast qed lemma rel_frontier_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_frontier S \ rel_frontier T \ rel_frontier (S \ T)" by (force simp: rel_frontier_def rel_interior_Times assms closure_Times) subsubsection\<^marker>\tag unimportant\ \Relative interior of convex cone\ lemma cone_rel_interior: fixes S :: "'m::euclidean_space set" assumes "cone S" shows "cone ({0} \ rel_interior S)" proof (cases "S = {}") case True then show ?thesis by (simp add: cone_0) next case False then have *: "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto then have *: "0 \ ({0} \ rel_interior S)" and "\c. c > 0 \ (*\<^sub>R) c ` ({0} \ rel_interior S) = ({0} \ rel_interior S)" by (auto simp add: rel_interior_scaleR) then show ?thesis using cone_iff[of "{0} \ rel_interior S"] by auto qed lemma rel_interior_convex_cone_aux: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) \ c > 0 \ x \ (((*\<^sub>R) c) ` (rel_interior S))" proof (cases "S = {}") case True then show ?thesis by (simp add: cone_hull_empty) next case False then obtain s where "s \ S" by auto have conv: "convex ({(1 :: real)} \ S)" using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"] by auto define f where "f y = {z. (y, z) \ cone hull ({1 :: real} \ S)}" for y then have *: "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) = (c \ rel_interior {y. f y \ {}} \ x \ rel_interior (f c))" using convex_cone_hull[of "{(1 :: real)} \ S"] conv by (subst rel_interior_projection) auto { fix y :: real assume "y \ 0" then have "y *\<^sub>R (1,s) \ cone hull ({1 :: real} \ S)" using cone_hull_expl[of "{(1 :: real)} \ S"] \s \ S\ by auto then have "f y \ {}" using f_def by auto } then have "{y. f y \ {}} = {0..}" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have **: "rel_interior {y. f y \ {}} = {0<..}" using rel_interior_real_semiline by auto { fix c :: real assume "c > 0" then have "f c = ((*\<^sub>R) c ` S)" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have "rel_interior (f c) = (*\<^sub>R) c ` rel_interior S" using rel_interior_convex_scaleR[of S c] assms by auto } then show ?thesis using * ** by auto qed lemma rel_interior_convex_cone: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "rel_interior (cone hull ({1 :: real} \ S)) = {(c, c *\<^sub>R x) | c x. c > 0 \ x \ rel_interior S}" (is "?lhs = ?rhs") proof - { fix z assume "z \ ?lhs" have *: "z = (fst z, snd z)" by auto then have "z \ ?rhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \z \ ?lhs\ by fastforce } moreover { fix z assume "z \ ?rhs" then have "z \ ?lhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms by auto } ultimately show ?thesis by blast qed lemma convex_hull_finite_union: assumes "finite I" assumes "\i\I. convex (S i) \ (S i) \ {}" shows "convex hull (\(S ` I)) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)}" (is "?lhs = ?rhs") proof - have "?lhs \ ?rhs" proof fix x assume "x \ ?rhs" then obtain c s where *: "sum (\i. c i *\<^sub>R s i) I = x" "sum c I = 1" "(\i\I. c i \ 0) \ (\i\I. s i \ S i)" by auto then have "\i\I. s i \ convex hull (\(S ` I))" using hull_subset[of "\(S ` I)" convex] by auto then show "x \ ?lhs" unfolding *(1)[symmetric] using * assms convex_convex_hull by (subst convex_sum) auto qed { fix i assume "i \ I" with assms have "\p. p \ S i" by auto } then obtain p where p: "\i\I. p i \ S i" by metis { fix i assume "i \ I" { fix x assume "x \ S i" define c where "c j = (if j = i then 1::real else 0)" for j then have *: "sum c I = 1" using \finite I\ \i \ I\ sum.delta[of I i "\j::'a. 1::real"] by auto define s where "s j = (if j = i then x else p j)" for j then have "\j. c j *\<^sub>R s j = (if j = i then x else 0)" using c_def by (auto simp add: algebra_simps) then have "x = sum (\i. c i *\<^sub>R s i) I" using s_def c_def \finite I\ \i \ I\ sum.delta[of I i "\j::'a. x"] by auto moreover have "(\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. s i \ S i)" using * c_def s_def p \x \ S i\ by auto ultimately have "x \ ?rhs" by force } then have "?rhs \ S i" by auto } then have *: "?rhs \ \(S ` I)" by auto { fix u v :: real assume uv: "u \ 0 \ v \ 0 \ u + v = 1" fix x y assume xy: "x \ ?rhs \ y \ ?rhs" from xy obtain c s where xc: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)" by auto from xy obtain d t where yc: "y = sum (\i. d i *\<^sub>R t i) I \ (\i\I. d i \ 0) \ sum d I = 1 \ (\i\I. t i \ S i)" by auto define e where "e i = u * c i + v * d i" for i have ge0: "\i\I. e i \ 0" using e_def xc yc uv by simp have "sum (\i. u * c i) I = u * sum c I" by (simp add: sum_distrib_left) moreover have "sum (\i. v * d i) I = v * sum d I" by (simp add: sum_distrib_left) ultimately have sum1: "sum e I = 1" using e_def xc yc uv by (simp add: sum.distrib) define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)" for i { fix i assume i: "i \ I" have "q i \ S i" proof (cases "e i = 0") case True then show ?thesis using i p q_def by auto next case False then show ?thesis using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"] mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"] assms q_def e_def i False xc yc uv by (auto simp del: mult_nonneg_nonneg) qed } then have qs: "\i\I. q i \ S i" by auto { fix i assume i: "i \ I" have "(u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" proof (cases "e i = 0") case True have ge: "u * (c i) \ 0 \ v * d i \ 0" using xc yc uv i by simp moreover from ge have "u * c i \ 0 \ v * d i \ 0" using True e_def i by simp ultimately have "u * c i = 0 \ v * d i = 0" by auto with True show ?thesis by auto next case False then have "(u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i) = q i" using q_def by auto then have "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i)) = (e i) *\<^sub>R (q i)" by auto with False show ?thesis by (simp add: algebra_simps) qed } then have *: "\i\I. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" by auto have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I" using xc yc by (simp add: algebra_simps scaleR_right.sum sum.distrib) also have "\ = sum (\i. e i *\<^sub>R q i) I" using * by auto finally have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (e i) *\<^sub>R (q i)) I" by auto then have "u *\<^sub>R x + v *\<^sub>R y \ ?rhs" using ge0 sum1 qs by auto } then have "convex ?rhs" unfolding convex_def by auto then show ?thesis using \?lhs \ ?rhs\ * hull_minimal[of "\(S ` I)" ?rhs convex] by blast qed lemma convex_hull_union_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "S \ {}" and "convex T" and "T \ {}" shows "convex hull (S \ T) = {u *\<^sub>R s + v *\<^sub>R t | u v s t. u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T}" (is "?lhs = ?rhs") proof define I :: "nat set" where "I = {1, 2}" define s where "s i = (if i = (1::nat) then S else T)" for i have "\(s ` I) = S \ T" using s_def I_def by auto then have "convex hull (\(s ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(s ` I) = {\ i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)}" using assms s_def I_def by (subst convex_hull_finite_union) auto moreover have "{\i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)} \ ?rhs" using s_def I_def by auto ultimately show "?lhs \ ?rhs" by auto { fix x assume "x \ ?rhs" then obtain u v s t where *: "x = u *\<^sub>R s + v *\<^sub>R t \ u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T" by auto then have "x \ convex hull {s, t}" using convex_hull_2[of s t] by auto then have "x \ convex hull (S \ T)" using * hull_mono[of "{s, t}" "S \ T"] by auto } then show "?lhs \ ?rhs" by blast qed proposition ray_to_rel_frontier: fixes a :: "'a::real_inner" assumes "bounded S" and a: "a \ rel_interior S" and aff: "(a + l) \ affine hull S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ rel_frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ rel_interior S" proof - have aaff: "a \ affine hull S" by (meson a hull_subset rel_interior_subset rev_subsetD) let ?D = "{d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" obtain B where "B > 0" and B: "S \ ball a B" using bounded_subset_ballD [OF \bounded S\] by blast have "a + (B / norm l) *\<^sub>R l \ ball a B" by (simp add: dist_norm \l \ 0\) with B have "a + (B / norm l) *\<^sub>R l \ rel_interior S" using rel_interior_subset subsetCE by blast with \B > 0\ \l \ 0\ have nonMT: "?D \ {}" using divide_pos_pos zero_less_norm_iff by fastforce have bdd: "bdd_below ?D" by (metis (no_types, lifting) bdd_belowI le_less mem_Collect_eq) have relin_Ex: "\x. x \ rel_interior S \ \e>0. \x'\affine hull S. dist x' x < e \ x' \ rel_interior S" using openin_rel_interior [of S] by (simp add: openin_euclidean_subtopology_iff) define d where "d = Inf ?D" obtain \ where "0 < \" and \: "\\. \0 \ \; \ < \\ \ (a + \ *\<^sub>R l) \ rel_interior S" proof - obtain e where "e>0" and e: "\x'. x' \ affine hull S \ dist x' a < e \ x' \ rel_interior S" using relin_Ex a by blast show thesis proof (rule_tac \ = "e / norm l" in that) show "0 < e / norm l" by (simp add: \0 < e\ \l \ 0\) next show "a + \ *\<^sub>R l \ rel_interior S" if "0 \ \" "\ < e / norm l" for \ proof (rule e) show "a + \ *\<^sub>R l \ affine hull S" by (metis (no_types) add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show "dist (a + \ *\<^sub>R l) a < e" using that by (simp add: \l \ 0\ dist_norm pos_less_divide_eq) qed qed qed have inint: "\e. \0 \ e; e < d\ \ a + e *\<^sub>R l \ rel_interior S" unfolding d_def using cInf_lower [OF _ bdd] by (metis (no_types, lifting) a add.right_neutral le_less mem_Collect_eq not_less real_vector.scale_zero_left) have "\ \ d" unfolding d_def using \ dual_order.strict_implies_order le_less_linear by (blast intro: cInf_greatest [OF nonMT]) with \0 < \\ have "0 < d" by simp have "a + d *\<^sub>R l \ rel_interior S" proof assume adl: "a + d *\<^sub>R l \ rel_interior S" obtain e where "e > 0" and e: "\x'. x' \ affine hull S \ dist x' (a + d *\<^sub>R l) < e \ x' \ rel_interior S" using relin_Ex adl by blast have "d + e / norm l \ Inf {d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" proof (rule cInf_greatest [OF nonMT], clarsimp) fix x::real assume "0 < x" and nonrel: "a + x *\<^sub>R l \ rel_interior S" show "d + e / norm l \ x" proof (cases "x < d") case True with inint nonrel \0 < x\ show ?thesis by auto next case False then have dle: "x < d + e / norm l \ dist (a + x *\<^sub>R l) (a + d *\<^sub>R l) < e" by (simp add: field_simps \l \ 0\) have ain: "a + x *\<^sub>R l \ affine hull S" by (metis add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show ?thesis using e [OF ain] nonrel dle by force qed qed then show False using \0 < e\ \l \ 0\ by (simp add: d_def [symmetric] field_simps) qed moreover have "a + d *\<^sub>R l \ closure S" proof (clarsimp simp: closure_approachable) fix \::real assume "0 < \" have 1: "a + (d - min d (\ / 2 / norm l)) *\<^sub>R l \ S" proof (rule subsetD [OF rel_interior_subset inint]) show "d - min d (\ / 2 / norm l) < d" using \l \ 0\ \0 < d\ \0 < \\ by auto qed auto have "norm l * min d (\ / (norm l * 2)) \ norm l * (\ / (norm l * 2))" by (metis min_def mult_left_mono norm_ge_zero order_refl) also have "... < \" using \l \ 0\ \0 < \\ by (simp add: field_simps) finally have 2: "norm l * min d (\ / (norm l * 2)) < \" . show "\y\S. dist y (a + d *\<^sub>R l) < \" using 1 2 \0 < d\ \0 < \\ by (rule_tac x="a + (d - min d (\ / 2 / norm l)) *\<^sub>R l" in bexI) (auto simp: algebra_simps) qed ultimately have infront: "a + d *\<^sub>R l \ rel_frontier S" by (simp add: rel_frontier_def) show ?thesis by (rule that [OF \0 < d\ infront inint]) qed corollary ray_to_frontier: fixes a :: "'a::euclidean_space" assumes "bounded S" and a: "a \ interior S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ interior S" proof - have \
: "interior S = rel_interior S" using a rel_interior_nonempty_interior by auto then have "a \ rel_interior S" using a by simp moreover have "a + l \ affine hull S" using a affine_hull_nonempty_interior by blast ultimately show thesis by (metis \
\bounded S\ \l \ 0\ frontier_def ray_to_rel_frontier rel_frontier_def that) qed lemma segment_to_rel_frontier_aux: fixes x :: "'a::euclidean_space" assumes "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "x \ y" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof - have "x + (y - x) \ affine hull S" using hull_inc [OF y] by auto then obtain d where "0 < d" and df: "(x + d *\<^sub>R (y-x)) \ rel_frontier S" and di: "\e. \0 \ e; e < d\ \ (x + e *\<^sub>R (y-x)) \ rel_interior S" by (rule ray_to_rel_frontier [OF \bounded S\ x]) (use xy in auto) show ?thesis proof show "x + d *\<^sub>R (y - x) \ rel_frontier S" by (simp add: df) next have "open_segment x y \ rel_interior S" using rel_interior_closure_convex_segment [OF \convex S\ x] closure_subset y by blast moreover have "x + d *\<^sub>R (y - x) \ open_segment x y" if "d < 1" using xy \0 < d\ that by (force simp: in_segment algebra_simps) ultimately have "1 \ d" using df rel_frontier_def by fastforce moreover have "x = (1 / d) *\<^sub>R x + ((d - 1) / d) *\<^sub>R x" by (metis \0 < d\ add.commute add_divide_distrib diff_add_cancel divide_self_if less_irrefl scaleR_add_left scaleR_one) ultimately show "y \ closed_segment x (x + d *\<^sub>R (y - x))" unfolding in_segment by (rule_tac x="1/d" in exI) (auto simp: algebra_simps) next show "open_segment x (x + d *\<^sub>R (y - x)) \ rel_interior S" proof (rule rel_interior_closure_convex_segment [OF \convex S\ x]) show "x + d *\<^sub>R (y - x) \ closure S" using df rel_frontier_def by auto qed qed qed lemma segment_to_rel_frontier: fixes x :: "'a::euclidean_space" assumes S: "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "\(x = y \ S = {x})" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof (cases "x=y") case True with xy have "S \ {x}" by blast with True show ?thesis by (metis Set.set_insert all_not_in_conv ends_in_segment(1) insert_iff segment_to_rel_frontier_aux[OF S x] that y) next case False then show ?thesis using segment_to_rel_frontier_aux [OF S x y] that by blast qed proposition rel_frontier_not_sing: fixes a :: "'a::euclidean_space" assumes "bounded S" shows "rel_frontier S \ {a}" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain z where "z \ S" by blast then show ?thesis proof (cases "S = {z}") case True then show ?thesis by simp next case False then obtain w where "w \ S" "w \ z" using \z \ S\ by blast show ?thesis proof assume "rel_frontier S = {a}" then consider "w \ rel_frontier S" | "z \ rel_frontier S" using \w \ z\ by auto then show False proof cases case 1 then have w: "w \ rel_interior S" using \w \ S\ closure_subset rel_frontier_def by fastforce have "w + (w - z) \ affine hull S" by (metis \w \ S\ \z \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(w + e *\<^sub>R (w - z)) \ rel_frontier S" using \w \ z\ \z \ S\ by (metis assms ray_to_rel_frontier right_minus_eq w) moreover obtain d where "0 < d" "(w + d *\<^sub>R (z - w)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ w, of "1 *\<^sub>R (z - w)"] \w \ z\ \z \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (z - w) = e *\<^sub>R (w - z)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) next case 2 then have z: "z \ rel_interior S" using \z \ S\ closure_subset rel_frontier_def by fastforce have "z + (z - w) \ affine hull S" by (metis \z \ S\ \w \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(z + e *\<^sub>R (z - w)) \ rel_frontier S" using \w \ z\ \w \ S\ by (metis assms ray_to_rel_frontier right_minus_eq z) moreover obtain d where "0 < d" "(z + d *\<^sub>R (w - z)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ z, of "1 *\<^sub>R (w - z)"] \w \ z\ \w \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (w - z) = e *\<^sub>R (z - w)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) qed qed qed qed subsection\<^marker>\tag unimportant\ \Convexity on direct sums\ lemma closure_sum: fixes S T :: "'a::real_normed_vector set" shows "closure S + closure T \ closure (S + T)" unfolding set_plus_image closure_Times [symmetric] split_def by (intro closure_bounded_linear_image_subset bounded_linear_add bounded_linear_fst bounded_linear_snd) lemma fst_snd_linear: "linear (\(x,y). x + y)" unfolding linear_iff by (simp add: algebra_simps) lemma rel_interior_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S + T) = rel_interior S + rel_interior T" proof - have "rel_interior S + rel_interior T = (\(x,y). x + y) ` (rel_interior S \ rel_interior T)" by (simp add: set_plus_image) also have "\ = (\(x,y). x + y) ` rel_interior (S \ T)" using rel_interior_Times assms by auto also have "\ = rel_interior (S + T)" using fst_snd_linear convex_Times assms rel_interior_convex_linear_image[of "(\(x,y). x + y)" "S \ T"] by (auto simp add: set_plus_image) finally show ?thesis .. qed lemma rel_interior_sum_gen: fixes S :: "'a \ 'n::euclidean_space set" assumes "\i. i\I \ convex (S i)" shows "rel_interior (sum S I) = sum (\i. rel_interior (S i)) I" using rel_interior_sum rel_interior_sing[of "0"] assms by (subst sum_set_cond_linear[of convex], auto simp add: convex_set_plus) lemma convex_rel_open_direct_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S \ T) \ rel_open (S \ T)" by (metis assms convex_Times rel_interior_Times rel_open_def) lemma convex_rel_open_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S + T) \ rel_open (S + T)" by (metis assms convex_set_plus rel_interior_sum rel_open_def) lemma convex_hull_finite_union_cones: assumes "finite I" and "I \ {}" assumes "\i. i\I \ convex (S i) \ cone (S i) \ S i \ {}" shows "convex hull (\(S ` I)) = sum S I" (is "?lhs = ?rhs") proof - { fix x assume "x \ ?lhs" then obtain c xs where x: "x = sum (\i. c i *\<^sub>R xs i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. xs i \ S i)" using convex_hull_finite_union[of I S] assms by auto define s where "s i = c i *\<^sub>R xs i" for i have "\i\I. s i \ S i" using s_def x assms by (simp add: mem_cone) moreover have "x = sum s I" using x s_def by auto ultimately have "x \ ?rhs" using set_sum_alt[of I S] assms by auto } moreover { fix x assume "x \ ?rhs" then obtain s where x: "x = sum s I \ (\i\I. s i \ S i)" using set_sum_alt[of I S] assms by auto define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i then have "x = sum (\i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I" using x assms by auto moreover have "\i\I. xs i \ S i" using x xs_def assms by (simp add: cone_def) moreover have "\i\I. (1 :: real) / of_nat (card I) \ 0" by auto moreover have "sum (\i. (1 :: real) / of_nat (card I)) I = 1" using assms by auto ultimately have "x \ ?lhs" using assms apply (simp add: convex_hull_finite_union[of I S]) by (rule_tac x = "(\i. 1 / (card I))" in exI) auto } ultimately show ?thesis by auto qed lemma convex_hull_union_cones_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "cone S" and "S \ {}" assumes "convex T" and "cone T" and "T \ {}" shows "convex hull (S \ T) = S + T" proof - define I :: "nat set" where "I = {1, 2}" define A where "A i = (if i = (1::nat) then S else T)" for i have "\(A ` I) = S \ T" using A_def I_def by auto then have "convex hull (\(A ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(A ` I) = sum A I" using A_def I_def by (metis assms convex_hull_finite_union_cones empty_iff finite.emptyI finite.insertI insertI1) moreover have "sum A I = S + T" using A_def I_def by (force simp add: set_plus_def) ultimately show ?thesis by auto qed lemma rel_interior_convex_hull_union: fixes S :: "'a \ 'n::euclidean_space set" assumes "finite I" and "\i\I. convex (S i) \ S i \ {}" shows "rel_interior (convex hull (\(S ` I))) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior(S i))}" (is "?lhs = ?rhs") proof (cases "I = {}") case True then show ?thesis using convex_hull_empty by auto next case False define C0 where "C0 = convex hull (\(S ` I))" have "\i\I. C0 \ S i" unfolding C0_def using hull_subset[of "\(S ` I)"] by auto define K0 where "K0 = cone hull ({1 :: real} \ C0)" define K where "K i = cone hull ({1 :: real} \ S i)" for i have "\i\I. K i \ {}" unfolding K_def using assms by (simp add: cone_hull_empty_iff[symmetric]) have convK: "\i\I. convex (K i)" unfolding K_def by (simp add: assms(2) convex_Times convex_cone_hull) have "K0 \ K i" if "i \ I" for i unfolding K0_def K_def by (simp add: Sigma_mono \\i\I. S i \ C0\ hull_mono that) then have "K0 \ \(K ` I)" by auto moreover have "convex K0" unfolding K0_def by (simp add: C0_def convex_Times convex_cone_hull) ultimately have geq: "K0 \ convex hull (\(K ` I))" using hull_minimal[of _ "K0" "convex"] by blast have "\i\I. K i \ {1 :: real} \ S i" using K_def by (simp add: hull_subset) then have "\(K ` I) \ {1 :: real} \ \(S ` I)" by auto then have "convex hull \(K ` I) \ convex hull ({1 :: real} \ \(S ` I))" by (simp add: hull_mono) then have "convex hull \(K ` I) \ {1 :: real} \ C0" unfolding C0_def using convex_hull_Times[of "{(1 :: real)}" "\(S ` I)"] convex_hull_singleton by auto moreover have "cone (convex hull (\(K ` I)))" by (simp add: K_def cone_Union cone_cone_hull cone_convex_hull) ultimately have "convex hull (\(K ` I)) \ K0" unfolding K0_def using hull_minimal[of _ "convex hull (\(K ` I))" "cone"] by blast then have "K0 = convex hull (\(K ` I))" using geq by auto also have "\ = sum K I" using assms False \\i\I. K i \ {}\ cone_hull_eq convK by (intro convex_hull_finite_union_cones; fastforce simp: K_def) finally have "K0 = sum K I" by auto then have *: "rel_interior K0 = sum (\i. (rel_interior (K i))) I" using rel_interior_sum_gen[of I K] convK by auto { fix x assume "x \ ?lhs" then have "(1::real, x) \ rel_interior K0" using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull by auto then obtain k where k: "(1::real, x) = sum k I \ (\i\I. k i \ rel_interior (K i))" using \finite I\ * set_sum_alt[of I "\i. rel_interior (K i)"] by auto { fix i assume "i \ I" then have "convex (S i) \ k i \ rel_interior (cone hull {1} \ S i)" using k K_def assms by auto then have "\ci si. k i = (ci, ci *\<^sub>R si) \ 0 < ci \ si \ rel_interior (S i)" using rel_interior_convex_cone[of "S i"] by auto } then obtain c s where cs: "\i\I. k i = (c i, c i *\<^sub>R s i) \ 0 < c i \ s i \ rel_interior (S i)" by metis then have "x = (\i\I. c i *\<^sub>R s i) \ sum c I = 1" using k by (simp add: sum_prod) then have "x \ ?rhs" using k cs by auto } moreover { fix x assume "x \ ?rhs" then obtain c s where cs: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior (S i))" by auto define k where "k i = (c i, c i *\<^sub>R s i)" for i { fix i assume "i \ I" then have "k i \ rel_interior (K i)" using k_def K_def assms cs rel_interior_convex_cone[of "S i"] by auto } then have "(1, x) \ rel_interior K0" using * set_sum_alt[of I "(\i. rel_interior (K i))"] assms cs by (simp add: k_def) (metis (mono_tags, lifting) sum_prod) then have "x \ ?lhs" using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x] by auto } ultimately show ?thesis by blast qed lemma convex_le_Inf_differential: fixes f :: "real \ real" assumes "convex_on I f" and "x \ interior I" and "y \ I" shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" (is "_ \ _ + Inf (?F x) * (y - x)") proof (cases rule: linorder_cases) assume "x < y" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = min (x + e / 2) ((x + y) / 2)" ultimately have "x < t" "t < y" "t \ ball x e" by (auto simp: dist_real_def field_simps split: split_min) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto define K where "K = x - e / 2" with \0 < e\ have "K \ ball x e" "K < x" by (auto simp: dist_real_def) then have "K \ I" using \interior I \ I\ e(2) by blast have "Inf (?F x) \ (f x - f y) / (x - y)" proof (intro bdd_belowI cInf_lower2) show "(f x - f t) / (x - t) \ ?F x" using \t \ I\ \x < t\ by auto show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" using \convex_on I f\ \x \ I\ \y \ I\ \x < t\ \t < y\ by (rule convex_on_diff) next fix y assume "y \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \K \ I\ _ \K < x\ _]] show "(f K - f x) / (K - x) \ y" by auto qed then show ?thesis using \x < y\ by (simp add: field_simps) next assume "y < x" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = x + e / 2" ultimately have "x < t" "t \ ball x e" by (auto simp: dist_real_def field_simps) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto have "(f x - f y) / (x - y) \ Inf (?F x)" proof (rule cInf_greatest) have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" using \y < x\ by (auto simp: field_simps) also fix z assume "z \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \y \ I\ _ \y < x\]] have "(f y - f x) / (y - x) \ z" by auto finally show "(f x - f y) / (x - y) \ z" . next have "x + e / 2 \ ball x e" using e by (auto simp: dist_real_def) with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto then show "?F x \ {}" by blast qed then show ?thesis using \y < x\ by (simp add: field_simps) qed simp subsection\<^marker>\tag unimportant\\Explicit formulas for interior and relative interior of convex hull\ lemma at_within_cbox_finite: assumes "x \ box a b" "x \ S" "finite S" shows "(at x within cbox a b - S) = at x" proof - have "interior (cbox a b - S) = box a b - S" using \finite S\ by (simp add: interior_diff finite_imp_closed) then show ?thesis using at_within_interior assms by fastforce qed lemma affine_independent_convex_affine_hull: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" "T \ S" shows "convex hull T = affine hull T \ convex hull S" proof - have fin: "finite S" "finite T" using assms aff_independent_finite finite_subset by auto have "convex hull T \ affine hull T" using convex_hull_subset_affine_hull by blast moreover have "convex hull T \ convex hull S" using assms hull_mono by blast moreover have "affine hull T \ convex hull S \ convex hull T" proof - have 0: "\u. sum u S = 0 \ (\v\S. u v = 0) \ (\v\S. u v *\<^sub>R v) \ 0" using affine_dependent_explicit_finite assms(1) fin(1) by auto show ?thesis proof (clarsimp simp add: affine_hull_finite fin) fix u assume S: "(\v\T. u v *\<^sub>R v) \ convex hull S" and T1: "sum u T = 1" then obtain v where v: "\x\S. 0 \ v x" "sum v S = 1" "(\x\S. v x *\<^sub>R x) = (\v\T. u v *\<^sub>R v)" by (auto simp add: convex_hull_finite fin) { fix x assume"x \ T" then have S: "S = (S - T) \ T" \ \split into separate cases\ using assms by auto have [simp]: "(\x\T. v x *\<^sub>R x) + (\x\S - T. v x *\<^sub>R x) = (\x\T. u x *\<^sub>R x)" "sum v T + sum v (S - T) = 1" using v fin S by (auto simp: sum.union_disjoint [symmetric] Un_commute) have "(\x\S. if x \ T then v x - u x else v x) = 0" "(\x\S. (if x \ T then v x - u x else v x) *\<^sub>R x) = 0" using v fin T1 by (subst S, subst sum.union_disjoint, auto simp: algebra_simps sum_subtractf)+ } note [simp] = this have "(\x\T. 0 \ u x)" using 0 [of "\x. if x \ T then v x - u x else v x"] \T \ S\ v(1) by fastforce then show "(\v\T. u v *\<^sub>R v) \ convex hull T" using 0 [of "\x. if x \ T then v x - u x else v x"] \T \ S\ T1 by (fastforce simp add: convex_hull_finite fin) qed qed ultimately show ?thesis by blast qed lemma affine_independent_span_eq: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" "card S = Suc (DIM ('a))" shows "affine hull S = UNIV" proof (cases "S = {}") case True then show ?thesis using assms by simp next case False then obtain a T where T: "a \ T" "S = insert a T" by blast then have fin: "finite T" using assms by (metis finite_insert aff_independent_finite) have "UNIV \ (+) a ` span ((\x. x - a) ` T)" proof (intro card_ge_dim_independent Fun.vimage_subsetD) show "independent ((\x. x - a) ` T)" using T affine_dependent_iff_dependent assms(1) by auto show "dim ((+) a -` UNIV) \ card ((\x. x - a) ` T)" using assms T fin by (auto simp: card_image inj_on_def) qed (use surj_plus in auto) then show ?thesis using T(2) affine_hull_insert_span_gen equalityI by fastforce qed lemma affine_independent_span_gt: fixes S :: "'a::euclidean_space set" assumes ind: "\ affine_dependent S" and dim: "DIM ('a) < card S" shows "affine hull S = UNIV" proof (intro affine_independent_span_eq [OF ind] antisym) show "card S \ Suc DIM('a)" using aff_independent_finite affine_dependent_biggerset ind by fastforce show "Suc DIM('a) \ card S" using Suc_leI dim by blast qed lemma empty_interior_affine_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" and dim: "card S \ DIM ('a)" shows "interior(affine hull S) = {}" using assms proof (induct S rule: finite_induct) case (insert x S) then have "dim (span ((\y. y - x) ` S)) < DIM('a)" by (auto simp: Suc_le_lessD card_image_le dual_order.trans intro!: dim_le_card'[THEN le_less_trans]) then show ?case by (simp add: empty_interior_lowdim affine_hull_insert_span_gen interior_translation) qed auto lemma empty_interior_convex_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" and dim: "card S \ DIM ('a)" shows "interior(convex hull S) = {}" by (metis Diff_empty Diff_eq_empty_iff convex_hull_subset_affine_hull interior_mono empty_interior_affine_hull [OF assms]) lemma explicit_subset_rel_interior_convex_hull: fixes S :: "'a::euclidean_space set" shows "finite S \ {y. \u. (\x \ S. 0 < u x \ u x < 1) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y} \ rel_interior (convex hull S)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=S, simplified]) lemma explicit_subset_rel_interior_convex_hull_minimal: fixes S :: "'a::euclidean_space set" shows "finite S \ {y. \u. (\x \ S. 0 < u x) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y} \ rel_interior (convex hull S)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=S, simplified]) lemma rel_interior_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "rel_interior(convex hull S) = {y. \u. (\x \ S. 0 < u x) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" (is "?lhs = ?rhs") proof show "?rhs \ ?lhs" by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms) next show "?lhs \ ?rhs" proof (cases "\a. S = {a}") case True then show "?lhs \ ?rhs" by force next case False have fs: "finite S" using assms by (simp add: aff_independent_finite) { fix a b and d::real assume ab: "a \ S" "b \ S" "a \ b" then have S: "S = (S - {a,b}) \ {a,b}" \ \split into separate cases\ by auto have "(\x\S. if x = a then - d else if x = b then d else 0) = 0" "(\x\S. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a" using ab fs by (subst S, subst sum.union_disjoint, auto)+ } note [simp] = this { fix y assume y: "y \ convex hull S" "y \ ?rhs" have *: False if ua: "\x\S. 0 \ u x" "sum u S = 1" "\ 0 < u a" "a \ S" and yT: "y = (\x\S. u x *\<^sub>R x)" "y \ T" "open T" and sb: "T \ affine hull S \ {w. \u. (\x\S. 0 \ u x) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = w}" for u T a proof - have ua0: "u a = 0" using ua by auto obtain b where b: "b\S" "a \ b" using ua False by auto obtain e where e: "0 < e" "ball (\x\S. u x *\<^sub>R x) e \ T" using yT by (auto elim: openE) with b obtain d where d: "0 < d" "norm(d *\<^sub>R (a-b)) < e" by (auto intro: that [of "e / 2 / norm(a-b)"]) have "(\x\S. u x *\<^sub>R x) \ affine hull S" using yT y by (metis affine_hull_convex_hull hull_redundant_eq) then have "(\x\S. u x *\<^sub>R x) - d *\<^sub>R (a - b) \ affine hull S" using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2) then have "y - d *\<^sub>R (a - b) \ T \ affine hull S" using d e yT by auto then obtain v where v: "\x\S. 0 \ v x" "sum v S = 1" "(\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x) - d *\<^sub>R (a - b)" using subsetD [OF sb] yT by auto have aff: "\u. sum u S = 0 \ (\v\S. u v = 0) \ (\v\S. u v *\<^sub>R v) \ 0" using assms by (simp add: affine_dependent_explicit_finite fs) show False using ua b d v aff [of "\x. (v x - u x) - (if x = a then -d else if x = b then d else 0)"] by (auto simp: algebra_simps sum_subtractf sum.distrib) qed have "y \ rel_interior (convex hull S)" using y apply (simp add: mem_rel_interior) apply (auto simp: convex_hull_finite [OF fs]) apply (drule_tac x=u in spec) apply (auto intro: *) done } with rel_interior_subset show "?lhs \ ?rhs" by blast qed qed lemma interior_convex_hull_explicit_minimal: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "interior(convex hull S) = (if card(S) \ DIM('a) then {} else {y. \u. (\x \ S. 0 < u x) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = y})" (is "_ = (if _ then _ else ?rhs)") proof (clarsimp simp: aff_independent_finite empty_interior_convex_hull assms) assume S: "\ card S \ DIM('a)" have "interior (convex hull S) = rel_interior(convex hull S)" using assms S by (simp add: affine_independent_span_gt rel_interior_interior) then show "interior(convex hull S) = ?rhs" by (simp add: assms S rel_interior_convex_hull_explicit) qed lemma interior_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "interior(convex hull S) = (if card(S) \ DIM('a) then {} else {y. \u. (\x \ S. 0 < u x \ u x < 1) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = y})" proof - { fix u :: "'a \ real" and a assume "card Basis < card S" and u: "\x. x\S \ 0 < u x" "sum u S = 1" and a: "a \ S" then have cs: "Suc 0 < card S" by (metis DIM_positive less_trans_Suc) obtain b where b: "b \ S" "a \ b" proof (cases "S \ {a}") case True then show thesis using cs subset_singletonD by fastforce qed blast have "u a + u b \ sum u {a,b}" using a b by simp also have "... \ sum u S" using a b u by (intro Groups_Big.sum_mono2) (auto simp: less_imp_le aff_independent_finite assms) finally have "u a < 1" using \b \ S\ u by fastforce } note [simp] = this show ?thesis using assms by (force simp add: not_le interior_convex_hull_explicit_minimal) qed lemma interior_closed_segment_ge2: fixes a :: "'a::euclidean_space" assumes "2 \ DIM('a)" shows "interior(closed_segment a b) = {}" using assms unfolding segment_convex_hull proof - have "card {a, b} \ DIM('a)" using assms by (simp add: card_insert_if linear not_less_eq_eq numeral_2_eq_2) then show "interior (convex hull {a, b}) = {}" by (metis empty_interior_convex_hull finite.insertI finite.emptyI) qed lemma interior_open_segment: fixes a :: "'a::euclidean_space" shows "interior(open_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (simp add: not_le, intro conjI impI) assume "2 \ DIM('a)" then show "interior (open_segment a b) = {}" using interior_closed_segment_ge2 interior_mono segment_open_subset_closed by blast next assume le2: "DIM('a) < 2" show "interior (open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False with le2 have "affine hull (open_segment a b) = UNIV" by (simp add: False affine_independent_span_gt) then show "interior (open_segment a b) = open_segment a b" using rel_interior_interior rel_interior_open_segment by blast qed qed lemma interior_closed_segment: fixes a :: "'a::euclidean_space" shows "interior(closed_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by simp next case False then have "closure (open_segment a b) = closed_segment a b" by simp then show ?thesis by (metis (no_types) convex_interior_closure convex_open_segment interior_open_segment) qed lemmas interior_segment = interior_closed_segment interior_open_segment lemma closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b = closed_segment c d \ {a,b} = {c,d}" proof assume abcd: "closed_segment a b = closed_segment c d" show "{a,b} = {c,d}" proof (cases "a=b \ c=d") case True with abcd show ?thesis by force next case False then have neq: "a \ b \ c \ d" by force have *: "closed_segment c d - {a, b} = rel_interior (closed_segment c d)" using neq abcd by (metis (no_types) open_segment_def rel_interior_closed_segment) have "b \ {c, d}" proof - have "insert b (closed_segment c d) = closed_segment c d" using abcd by blast then show ?thesis by (metis DiffD2 Diff_insert2 False * insertI1 insert_Diff_if open_segment_def rel_interior_closed_segment) qed moreover have "a \ {c, d}" by (metis Diff_iff False * abcd ends_in_segment(1) insertI1 open_segment_def rel_interior_closed_segment) ultimately show "{a, b} = {c, d}" using neq by fastforce qed next assume "{a,b} = {c,d}" then show "closed_segment a b = closed_segment c d" by (simp add: segment_convex_hull) qed lemma closed_open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b \ open_segment c d" by (metis DiffE closed_segment_neq_empty closure_closed_segment closure_open_segment ends_in_segment(1) insertI1 open_segment_def) lemma open_closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b \ closed_segment c d" using closed_open_segment_eq by blast lemma open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b = open_segment c d \ a = b \ c = d \ {a,b} = {c,d}" (is "?lhs = ?rhs") proof assume abcd: ?lhs show ?rhs proof (cases "a=b \ c=d") case True with abcd show ?thesis using finite_open_segment by fastforce next case False then have a2: "a \ b \ c \ d" by force with abcd show ?rhs unfolding open_segment_def by (metis (no_types) abcd closed_segment_eq closure_open_segment) qed next assume ?rhs then show ?lhs by (metis Diff_cancel convex_hull_singleton insert_absorb2 open_segment_def segment_convex_hull) qed subsection\<^marker>\tag unimportant\\Similar results for closure and (relative or absolute) frontier\ lemma closure_convex_hull [simp]: fixes S :: "'a::euclidean_space set" shows "compact S ==> closure(convex hull S) = convex hull S" by (simp add: compact_imp_closed compact_convex_hull) lemma rel_frontier_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "rel_frontier(convex hull S) = {y. \u. (\x \ S. 0 \ u x) \ (\x \ S. u x = 0) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof - have fs: "finite S" using assms by (simp add: aff_independent_finite) have "\u y v. \y \ S; u y = 0; sum u S = 1; \x\S. 0 < v x; sum v S = 1; (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)\ \ \u. sum u S = 0 \ (\v\S. u v \ 0) \ (\v\S. u v *\<^sub>R v) = 0" apply (rule_tac x = "\x. u x - v x" in exI) apply (force simp: sum_subtractf scaleR_diff_left) done then show ?thesis using fs assms apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit) apply (auto simp: convex_hull_finite) apply (metis less_eq_real_def) by (simp add: affine_dependent_explicit_finite) qed lemma frontier_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "frontier(convex hull S) = {y. \u. (\x \ S. 0 \ u x) \ (DIM ('a) < card S \ (\x \ S. u x = 0)) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof - have fs: "finite S" using assms by (simp add: aff_independent_finite) show ?thesis proof (cases "DIM ('a) < card S") case True with assms fs show ?thesis by (simp add: rel_frontier_def frontier_def rel_frontier_convex_hull_explicit [symmetric] interior_convex_hull_explicit_minimal rel_interior_convex_hull_explicit) next case False then have "card S \ DIM ('a)" by linarith then show ?thesis using assms fs apply (simp add: frontier_def interior_convex_hull_explicit finite_imp_compact) apply (simp add: convex_hull_finite) done qed qed lemma rel_frontier_convex_hull_cases: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "rel_frontier(convex hull S) = \{convex hull (S - {x}) |x. x \ S}" proof - have fs: "finite S" using assms by (simp add: aff_independent_finite) { fix u a have "\x\S. 0 \ u x \ a \ S \ u a = 0 \ sum u S = 1 \ \x v. x \ S \ (\x\S - {x}. 0 \ v x) \ sum v (S - {x}) = 1 \ (\x\S - {x}. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" apply (rule_tac x=a in exI) apply (rule_tac x=u in exI) apply (simp add: Groups_Big.sum_diff1 fs) done } moreover { fix a u have "a \ S \ \x\S - {a}. 0 \ u x \ sum u (S - {a}) = 1 \ \v. (\x\S. 0 \ v x) \ (\x\S. v x = 0) \ sum v S = 1 \ (\x\S. v x *\<^sub>R x) = (\x\S - {a}. u x *\<^sub>R x)" apply (rule_tac x="\x. if x = a then 0 else u x" in exI) apply (auto simp: sum.If_cases Diff_eq if_smult fs) done } ultimately show ?thesis using assms apply (simp add: rel_frontier_convex_hull_explicit) apply (auto simp add: convex_hull_finite fs Union_SetCompr_eq) done qed lemma frontier_convex_hull_eq_rel_frontier: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "frontier(convex hull S) = (if card S \ DIM ('a) then convex hull S else rel_frontier(convex hull S))" using assms unfolding rel_frontier_def frontier_def by (simp add: affine_independent_span_gt rel_interior_interior finite_imp_compact empty_interior_convex_hull aff_independent_finite) lemma frontier_convex_hull_cases: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "frontier(convex hull S) = (if card S \ DIM ('a) then convex hull S else \{convex hull (S - {x}) |x. x \ S})" by (simp add: assms frontier_convex_hull_eq_rel_frontier rel_frontier_convex_hull_cases) lemma in_frontier_convex_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" "card S \ Suc (DIM ('a))" "x \ S" shows "x \ frontier(convex hull S)" proof (cases "affine_dependent S") case True with assms obtain y where "y \ S" and y: "y \ affine hull (S - {y})" by (auto simp: affine_dependent_def) moreover have "x \ closure (convex hull S)" by (meson closure_subset hull_inc subset_eq \x \ S\) moreover have "x \ interior (convex hull S)" using assms by (metis Suc_mono affine_hull_convex_hull affine_hull_nonempty_interior \y \ S\ y card.remove empty_iff empty_interior_affine_hull finite_Diff hull_redundant insert_Diff interior_UNIV not_less) ultimately show ?thesis unfolding frontier_def by blast next case False { assume "card S = Suc (card Basis)" then have cs: "Suc 0 < card S" by (simp) with subset_singletonD have "\y \ S. y \ x" by (cases "S \ {x}") fastforce+ } note [dest!] = this show ?thesis using assms unfolding frontier_convex_hull_cases [OF False] Union_SetCompr_eq by (auto simp: le_Suc_eq hull_inc) qed lemma not_in_interior_convex_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" "card S \ Suc (DIM ('a))" "x \ S" shows "x \ interior(convex hull S)" using in_frontier_convex_hull [OF assms] by (metis Diff_iff frontier_def) lemma interior_convex_hull_eq_empty: fixes S :: "'a::euclidean_space set" assumes "card S = Suc (DIM ('a))" shows "interior(convex hull S) = {} \ affine_dependent S" proof show "affine_dependent S \ interior (convex hull S) = {}" proof (clarsimp simp: affine_dependent_def) fix a b assume "b \ S" "b \ affine hull (S - {b})" then have "interior(affine hull S) = {}" using assms by (metis DIM_positive One_nat_def Suc_mono card.remove card.infinite empty_interior_affine_hull eq_iff hull_redundant insert_Diff not_less zero_le_one) then show "interior (convex hull S) = {}" using affine_hull_nonempty_interior by fastforce qed next show "interior (convex hull S) = {} \ affine_dependent S" by (metis affine_hull_convex_hull affine_hull_empty affine_independent_span_eq assms convex_convex_hull empty_not_UNIV rel_interior_eq_empty rel_interior_interior) qed subsection \Coplanarity, and collinearity in terms of affine hull\ definition\<^marker>\tag important\ coplanar where "coplanar S \ \u v w. S \ affine hull {u,v,w}" lemma collinear_affine_hull: "collinear S \ (\u v. S \ affine hull {u,v})" proof (cases "S={}") case True then show ?thesis by simp next case False then obtain x where x: "x \ S" by auto { fix u assume *: "\x y. \x\S; y\S\ \ \c. x - y = c *\<^sub>R u" have "\y c. x - y = c *\<^sub>R u \ \a b. y = a *\<^sub>R x + b *\<^sub>R (x + u) \ a + b = 1" by (rule_tac x="1+c" in exI, rule_tac x="-c" in exI, simp add: algebra_simps) then have "\u v. S \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" using * [OF x] by (rule_tac x=x in exI, rule_tac x="x+u" in exI, force) } moreover { fix u v x y assume *: "S \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" have "\c. x - y = c *\<^sub>R (v-u)" if "x\S" "y\S" proof - obtain a r where "a + r = 1" "x = a *\<^sub>R u + r *\<^sub>R v" using "*" \x \ S\ by blast moreover obtain b s where "b + s = 1" "y = b *\<^sub>R u + s *\<^sub>R v" using "*" \y \ S\ by blast ultimately have "x - y = (r-s) *\<^sub>R (v-u)" by (simp add: algebra_simps) (metis scaleR_left.add) then show ?thesis by blast qed } ultimately show ?thesis unfolding collinear_def affine_hull_2 by blast qed lemma collinear_closed_segment [simp]: "collinear (closed_segment a b)" by (metis affine_hull_convex_hull collinear_affine_hull hull_subset segment_convex_hull) lemma collinear_open_segment [simp]: "collinear (open_segment a b)" unfolding open_segment_def by (metis convex_hull_subset_affine_hull segment_convex_hull dual_order.trans convex_hull_subset_affine_hull Diff_subset collinear_affine_hull) lemma collinear_between_cases: fixes c :: "'a::euclidean_space" shows "collinear {a,b,c} \ between (b,c) a \ between (c,a) b \ between (a,b) c" (is "?lhs = ?rhs") proof assume ?lhs then obtain u v where uv: "\x. x \ {a, b, c} \ \c. x = u + c *\<^sub>R v" by (auto simp: collinear_alt) show ?rhs using uv [of a] uv [of b] uv [of c] by (auto simp: between_1) next assume ?rhs then show ?lhs unfolding between_mem_convex_hull by (metis (no_types, opaque_lifting) collinear_closed_segment collinear_subset hull_redundant hull_subset insert_commute segment_convex_hull) qed lemma subset_continuous_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes "continuous_on (closed_segment a b) f" shows "closed_segment (f a) (f b) \ image f (closed_segment a b)" by (metis connected_segment convex_contains_segment ends_in_segment imageI is_interval_connected_1 is_interval_convex connected_continuous_image [OF assms]) lemma continuous_injective_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (closed_segment a b) = closed_segment (f a) (f b)" proof show "closed_segment (f a) (f b) \ f ` closed_segment a b" by (metis subset_continuous_image_segment_1 contf) show "f ` closed_segment a b \ closed_segment (f a) (f b)" proof (cases "a = b") case True then show ?thesis by auto next case False then have fnot: "f a \ f b" using inj_onD injf by fastforce moreover have "f a \ open_segment (f c) (f b)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def) assume fa: "f a \ closed_segment (f c) (f b)" moreover have "closed_segment (f c) (f b) \ f ` closed_segment c b" by (meson closed_segment_subset contf continuous_on_subset convex_closed_segment ends_in_segment(2) subset_continuous_image_segment_1 that) ultimately have "f a \ f ` closed_segment c b" by blast then have a: "a \ closed_segment c b" by (meson ends_in_segment inj_on_image_mem_iff injf subset_closed_segment that) have cb: "closed_segment c b \ closed_segment a b" by (simp add: closed_segment_subset that) show "f a = f c" proof (rule between_antisym) show "between (f c, f b) (f a)" by (simp add: between_mem_segment fa) show "between (f a, f b) (f c)" by (metis a cb between_antisym between_mem_segment between_triv1 subset_iff) qed qed moreover have "f b \ open_segment (f a) (f c)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def fnot eq_commute) assume fb: "f b \ closed_segment (f a) (f c)" moreover have "closed_segment (f a) (f c) \ f ` closed_segment a c" by (meson contf continuous_on_subset ends_in_segment(1) subset_closed_segment subset_continuous_image_segment_1 that) ultimately have "f b \ f ` closed_segment a c" by blast then have b: "b \ closed_segment a c" by (meson ends_in_segment inj_on_image_mem_iff injf subset_closed_segment that) have ca: "closed_segment a c \ closed_segment a b" by (simp add: closed_segment_subset that) show "f b = f c" proof (rule between_antisym) show "between (f c, f a) (f b)" by (simp add: between_commute between_mem_segment fb) show "between (f b, f a) (f c)" by (metis b between_antisym between_commute between_mem_segment between_triv2 that) qed qed ultimately show ?thesis by (force simp: closed_segment_eq_real_ivl open_segment_eq_real_ivl split: if_split_asm) qed qed lemma continuous_injective_image_open_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (open_segment a b) = open_segment (f a) (f b)" proof - have "f ` (open_segment a b) = f ` (closed_segment a b) - {f a, f b}" by (metis (no_types, opaque_lifting) empty_subsetI ends_in_segment image_insert image_is_empty inj_on_image_set_diff injf insert_subset open_segment_def segment_open_subset_closed) also have "... = open_segment (f a) (f b)" using continuous_injective_image_segment_1 [OF assms] by (simp add: open_segment_def inj_on_image_set_diff [OF injf]) finally show ?thesis . qed lemma collinear_imp_coplanar: "collinear s ==> coplanar s" by (metis collinear_affine_hull coplanar_def insert_absorb2) lemma collinear_small: assumes "finite s" "card s \ 2" shows "collinear s" proof - have "card s = 0 \ card s = 1 \ card s = 2" using assms by linarith then show ?thesis using assms using card_eq_SucD numeral_2_eq_2 by (force simp: card_1_singleton_iff) qed lemma coplanar_small: assumes "finite s" "card s \ 3" shows "coplanar s" proof - consider "card s \ 2" | "card s = Suc (Suc (Suc 0))" using assms by linarith then show ?thesis proof cases case 1 then show ?thesis by (simp add: \finite s\ collinear_imp_coplanar collinear_small) next case 2 then show ?thesis using hull_subset [of "{_,_,_}"] by (fastforce simp: coplanar_def dest!: card_eq_SucD) qed qed lemma coplanar_empty: "coplanar {}" by (simp add: coplanar_small) lemma coplanar_sing: "coplanar {a}" by (simp add: coplanar_small) lemma coplanar_2: "coplanar {a,b}" by (auto simp: card_insert_if coplanar_small) lemma coplanar_3: "coplanar {a,b,c}" by (auto simp: card_insert_if coplanar_small) lemma collinear_affine_hull_collinear: "collinear(affine hull s) \ collinear s" unfolding collinear_affine_hull by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_affine_hull_coplanar: "coplanar(affine hull s) \ coplanar s" unfolding coplanar_def by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "coplanar S" "linear f" shows "coplanar(f ` S)" proof - { fix u v w assume "S \ affine hull {u, v, w}" then have "f ` S \ f ` (affine hull {u, v, w})" by (simp add: image_mono) then have "f ` S \ affine hull (f ` {u, v, w})" by (metis assms(2) linear_conv_bounded_linear affine_hull_linear_image) } then show ?thesis by auto (meson assms(1) coplanar_def) qed lemma coplanar_translation_imp: assumes "coplanar S" shows "coplanar ((\x. a + x) ` S)" proof - obtain u v w where "S \ affine hull {u,v,w}" by (meson assms coplanar_def) then have "(+) a ` S \ affine hull {u + a, v + a, w + a}" using affine_hull_translation [of a "{u,v,w}" for u v w] by (force simp: add.commute) then show ?thesis unfolding coplanar_def by blast qed lemma coplanar_translation_eq: "coplanar((\x. a + x) ` S) \ coplanar S" by (metis (no_types) coplanar_translation_imp translation_galois) lemma coplanar_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "coplanar(f ` S) = coplanar S" proof assume "coplanar S" then show "coplanar (f ` S)" using assms(1) coplanar_linear_image by blast next obtain g where g: "linear g" "g \ f = id" using linear_injective_left_inverse [OF assms] by blast assume "coplanar (f ` S)" then show "coplanar S" by (metis coplanar_linear_image g(1) g(2) id_apply image_comp image_id) qed lemma coplanar_subset: "\coplanar t; S \ t\ \ coplanar S" by (meson coplanar_def order_trans) lemma affine_hull_3_imp_collinear: "c \ affine hull {a,b} \ collinear {a,b,c}" by (metis collinear_2 collinear_affine_hull_collinear hull_redundant insert_commute) lemma collinear_3_imp_in_affine_hull: assumes "collinear {a,b,c}" "a \ b" shows "c \ affine hull {a,b}" proof - obtain u x y where "b - a = y *\<^sub>R u" "c - a = x *\<^sub>R u" using assms unfolding collinear_def by auto with \a \ b\ have "\v. c = (1 - x / y) *\<^sub>R a + v *\<^sub>R b \ 1 - x / y + v = 1" by (simp add: algebra_simps) then show ?thesis by (simp add: hull_inc mem_affine) qed lemma collinear_3_affine_hull: assumes "a \ b" shows "collinear {a,b,c} \ c \ affine hull {a,b}" using affine_hull_3_imp_collinear assms collinear_3_imp_in_affine_hull by blast lemma collinear_3_eq_affine_dependent: "collinear{a,b,c} \ a = b \ a = c \ b = c \ affine_dependent {a,b,c}" proof (cases "a = b \ a = c \ b = c") case True then show ?thesis by (auto simp: insert_commute) next case False then have "collinear{a,b,c}" if "affine_dependent {a,b,c}" using that unfolding affine_dependent_def by (auto simp: insert_Diff_if; metis affine_hull_3_imp_collinear insert_commute) moreover have "affine_dependent {a,b,c}" if "collinear{a,b,c}" using False that by (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if) ultimately show ?thesis using False by blast qed lemma affine_dependent_imp_collinear_3: "affine_dependent {a,b,c} \ collinear{a,b,c}" by (simp add: collinear_3_eq_affine_dependent) lemma collinear_3: "NO_MATCH 0 x \ collinear {x,y,z} \ collinear {0, x-y, z-y}" by (auto simp add: collinear_def) lemma collinear_3_expand: "collinear{a,b,c} \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" proof - have "collinear{a,b,c} = collinear{a,c,b}" by (simp add: insert_commute) also have "... = collinear {0, a - c, b - c}" by (simp add: collinear_3) also have "... \ (a = c \ b = c \ (\ca. b - c = ca *\<^sub>R (a - c)))" by (simp add: collinear_lemma) also have "... \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" by (cases "a = c \ b = c") (auto simp: algebra_simps) finally show ?thesis . qed lemma collinear_aff_dim: "collinear S \ aff_dim S \ 1" proof assume "collinear S" then obtain u and v :: "'a" where "aff_dim S \ aff_dim {u,v}" by (metis \collinear S\ aff_dim_affine_hull aff_dim_subset collinear_affine_hull) then show "aff_dim S \ 1" using order_trans by fastforce next assume "aff_dim S \ 1" then have le1: "aff_dim (affine hull S) \ 1" by simp obtain B where "B \ S" and B: "\ affine_dependent B" "affine hull S = affine hull B" using affine_basis_exists [of S] by auto then have "finite B" "card B \ 2" using B le1 by (auto simp: affine_independent_iff_card) then have "collinear B" by (rule collinear_small) then show "collinear S" by (metis \affine hull S = affine hull B\ collinear_affine_hull_collinear) qed lemma collinear_midpoint: "collinear{a, midpoint a b, b}" proof - have \
: "\a \ midpoint a b; b - midpoint a b \ - 1 *\<^sub>R (a - midpoint a b)\ \ b = midpoint a b" by (simp add: algebra_simps) show ?thesis by (auto simp: collinear_3 collinear_lemma intro: \
) qed lemma midpoint_collinear: fixes a b c :: "'a::real_normed_vector" assumes "a \ c" shows "b = midpoint a c \ collinear{a,b,c} \ dist a b = dist b c" proof - have *: "a - (u *\<^sub>R a + (1 - u) *\<^sub>R c) = (1 - u) *\<^sub>R (a - c)" "u *\<^sub>R a + (1 - u) *\<^sub>R c - c = u *\<^sub>R (a - c)" "\1 - u\ = \u\ \ u = 1/2" for u::real by (auto simp: algebra_simps) have "b = midpoint a c \ collinear{a,b,c}" using collinear_midpoint by blast moreover have "b = midpoint a c \ dist a b = dist b c" if "collinear{a,b,c}" proof - consider "a = c" | u where "b = u *\<^sub>R a + (1 - u) *\<^sub>R c" using \collinear {a,b,c}\ unfolding collinear_3_expand by blast then show ?thesis proof cases case 2 with assms have "dist a b = dist b c \ b = midpoint a c" by (simp add: dist_norm * midpoint_def scaleR_add_right del: divide_const_simps) then show ?thesis by (auto simp: dist_midpoint) qed (use assms in auto) qed ultimately show ?thesis by blast qed lemma between_imp_collinear: fixes x :: "'a :: euclidean_space" assumes "between (a,b) x" shows "collinear {a,x,b}" proof (cases "x = a \ x = b \ a = b") case True with assms show ?thesis by (auto simp: dist_commute) next case False then have False if "\c. b - x \ c *\<^sub>R (a - x)" using that [of "-(norm(b - x) / norm(x - a))"] assms by (simp add: between_norm vector_add_divide_simps flip: real_vector.scale_minus_right) then show ?thesis by (auto simp: collinear_3 collinear_lemma) qed lemma midpoint_between: fixes a b :: "'a::euclidean_space" shows "b = midpoint a c \ between (a,c) b \ dist a b = dist b c" proof (cases "a = c") case False show ?thesis using False between_imp_collinear between_midpoint(1) midpoint_collinear by blast qed (auto simp: dist_commute) lemma collinear_triples: assumes "a \ b" shows "collinear(insert a (insert b S)) \ (\x \ S. collinear{a,b,x})" (is "?lhs = ?rhs") proof safe fix x assume ?lhs and "x \ S" then show "collinear {a, b, x}" using collinear_subset by force next assume ?rhs then have "\x \ S. collinear{a,x,b}" by (simp add: insert_commute) then have *: "\u. x = u *\<^sub>R a + (1 - u) *\<^sub>R b" if "x \ insert a (insert b S)" for x using that assms collinear_3_expand by fastforce+ have "\c. x - y = c *\<^sub>R (b - a)" if x: "x \ insert a (insert b S)" and y: "y \ insert a (insert b S)" for x y proof - obtain u v where "x = u *\<^sub>R a + (1 - u) *\<^sub>R b" "y = v *\<^sub>R a + (1 - v) *\<^sub>R b" using "*" x y by presburger then have "x - y = (v - u) *\<^sub>R (b - a)" by (simp add: scale_left_diff_distrib scale_right_diff_distrib) then show ?thesis .. qed then show ?lhs unfolding collinear_def by metis qed lemma collinear_4_3: assumes "a \ b" shows "collinear {a,b,c,d} \ collinear{a,b,c} \ collinear{a,b,d}" using collinear_triples [OF assms, of "{c,d}"] by (force simp:) lemma collinear_3_trans: assumes "collinear{a,b,c}" "collinear{b,c,d}" "b \ c" shows "collinear{a,b,d}" proof - have "collinear{b,c,a,d}" by (metis (full_types) assms collinear_4_3 insert_commute) then show ?thesis by (simp add: collinear_subset) qed lemma affine_hull_2_alt: fixes a b :: "'a::real_vector" shows "affine hull {a,b} = range (\u. a + u *\<^sub>R (b - a))" proof - have 1: "u *\<^sub>R a + v *\<^sub>R b = a + v *\<^sub>R (b - a)" if "u + v = 1" for u v using that by (simp add: algebra_simps flip: scaleR_add_left) have 2: "a + u *\<^sub>R (b - a) = (1 - u) *\<^sub>R a + u *\<^sub>R b" for u by (auto simp: algebra_simps) show ?thesis by (force simp add: affine_hull_2 dest: 1 intro!: 2) qed lemma interior_convex_hull_3_minimal: fixes a :: "'a::euclidean_space" assumes "\ collinear{a,b,c}" and 2: "DIM('a) = 2" shows "interior(convex hull {a,b,c}) = {v. \x y z. 0 < x \ 0 < y \ 0 < z \ x + y + z = 1 \ x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" (is "?lhs = ?rhs") proof have abc: "a \ b" "a \ c" "b \ c" "\ affine_dependent {a, b, c}" using assms by (auto simp: collinear_3_eq_affine_dependent) with 2 show "?lhs \ ?rhs" by (fastforce simp add: interior_convex_hull_explicit_minimal) show "?rhs \ ?lhs" using abc 2 apply (clarsimp simp add: interior_convex_hull_explicit_minimal) subgoal for x y z by (rule_tac x="\r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI) auto done qed subsection\<^marker>\tag unimportant\\Basic lemmas about hyperplanes and halfspaces\ lemma halfspace_Int_eq: "{x. a \ x \ b} \ {x. b \ a \ x} = {x. a \ x = b}" "{x. b \ a \ x} \ {x. a \ x \ b} = {x. a \ x = b}" by auto lemma hyperplane_eq_Ex: assumes "a \ 0" obtains x where "a \ x = b" by (rule_tac x = "(b / (a \ a)) *\<^sub>R a" in that) (simp add: assms) lemma hyperplane_eq_empty: "{x. a \ x = b} = {} \ a = 0 \ b \ 0" using hyperplane_eq_Ex by (metis (mono_tags, lifting) empty_Collect_eq inner_zero_left) lemma hyperplane_eq_UNIV: "{x. a \ x = b} = UNIV \ a = 0 \ b = 0" proof - have "a = 0 \ b = 0" if "UNIV \ {x. a \ x = b}" using subsetD [OF that, where c = "((b+1) / (a \ a)) *\<^sub>R a"] by (simp add: field_split_simps split: if_split_asm) then show ?thesis by force qed lemma halfspace_eq_empty_lt: "{x. a \ x < b} = {} \ a = 0 \ b \ 0" proof - have "a = 0 \ b \ 0" if "{x. a \ x < b} \ {}" using subsetD [OF that, where c = "((b-1) / (a \ a)) *\<^sub>R a"] by (force simp add: field_split_simps split: if_split_asm) then show ?thesis by force qed lemma halfspace_eq_empty_gt: "{x. a \ x > b} = {} \ a = 0 \ b \ 0" using halfspace_eq_empty_lt [of "-a" "-b"] by simp lemma halfspace_eq_empty_le: "{x. a \ x \ b} = {} \ a = 0 \ b < 0" proof - have "a = 0 \ b < 0" if "{x. a \ x \ b} \ {}" using subsetD [OF that, where c = "((b-1) / (a \ a)) *\<^sub>R a"] by (force simp add: field_split_simps split: if_split_asm) then show ?thesis by force qed lemma halfspace_eq_empty_ge: "{x. a \ x \ b} = {} \ a = 0 \ b > 0" using halfspace_eq_empty_le [of "-a" "-b"] by simp subsection\<^marker>\tag unimportant\\Use set distance for an easy proof of separation properties\ proposition\<^marker>\tag unimportant\ separation_closures: fixes S :: "'a::euclidean_space set" assumes "S \ closure T = {}" "T \ closure S = {}" obtains U V where "U \ V = {}" "open U" "open V" "S \ U" "T \ V" proof (cases "S = {} \ T = {}") case True with that show ?thesis by auto next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on UNIV f" unfolding f_def by (intro continuous_intros continuous_on_setdist) show ?thesis proof (rule_tac U = "{x. f x > 0}" and V = "{x. f x < 0}" in that) show "{x. 0 < f x} \ {x. f x < 0} = {}" by auto show "open {x. 0 < f x}" by (simp add: open_Collect_less contf) show "open {x. f x < 0}" by (simp add: open_Collect_less contf) have "\x. x \ S \ setdist {x} T \ 0" "\x. x \ T \ setdist {x} S \ 0" by (meson False assms disjoint_iff setdist_eq_0_sing_1)+ then show "S \ {x. 0 < f x}" "T \ {x. f x < 0}" using less_eq_real_def by (fastforce simp add: f_def setdist_sing_in_set)+ qed qed lemma separation_normal: fixes S :: "'a::euclidean_space set" assumes "closed S" "closed T" "S \ T = {}" obtains U V where "open U" "open V" "S \ U" "T \ V" "U \ V = {}" using separation_closures [of S T] by (metis assms closure_closed disjnt_def inf_commute) lemma separation_normal_local: fixes S :: "'a::euclidean_space set" assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" obtains S' T' where "openin (top_of_set U) S'" "openin (top_of_set U) T'" "S \ S'" "T \ T'" "S' \ T' = {}" proof (cases "S = {} \ T = {}") case True with that show ?thesis using UT US by (blast dest: closedin_subset) next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on U f" unfolding f_def by (intro continuous_intros) show ?thesis proof (rule_tac S' = "(U \ f -` {0<..})" and T' = "(U \ f -` {..<0})" in that) show "(U \ f -` {0<..}) \ (U \ f -` {..<0}) = {}" by auto show "openin (top_of_set U) (U \ f -` {0<..})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next show "openin (top_of_set U) (U \ f -` {..<0})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next have "S \ U" "T \ U" using closedin_imp_subset assms by blast+ then show "S \ U \ f -` {0<..}" "T \ U \ f -` {..<0}" using assms False by (force simp add: f_def setdist_sing_in_set intro!: setdist_gt_0_closedin)+ qed qed lemma separation_normal_compact: fixes S :: "'a::euclidean_space set" assumes "compact S" "closed T" "S \ T = {}" obtains U V where "open U" "compact(closure U)" "open V" "S \ U" "T \ V" "U \ V = {}" proof - have "closed S" "bounded S" using assms by (auto simp: compact_eq_bounded_closed) then obtain r where "r>0" and r: "S \ ball 0 r" by (auto dest!: bounded_subset_ballD) have **: "closed (T \ - ball 0 r)" "S \ (T \ - ball 0 r) = {}" using assms r by blast+ then obtain U V where UV: "open U" "open V" "S \ U" "T \ - ball 0 r \ V" "U \ V = {}" by (meson \closed S\ separation_normal) then have "compact(closure U)" by (meson bounded_ball bounded_subset compact_closure compl_le_swap2 disjoint_eq_subset_Compl le_sup_iff) with UV show thesis using that by auto qed subsection\Connectedness of the intersection of a chain\ proposition connected_chain: fixes \ :: "'a :: euclidean_space set set" assumes cc: "\S. S \ \ \ compact S \ connected S" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof (cases "\ = {}") case True then show ?thesis by auto next case False then have cf: "compact(\\)" by (simp add: cc compact_Inter) have False if AB: "closed A" "closed B" "A \ B = {}" and ABeq: "A \ B = \\" and "A \ {}" "B \ {}" for A B proof - obtain U V where "open U" "open V" "A \ U" "B \ V" "U \ V = {}" using separation_normal [OF AB] by metis obtain K where "K \ \" "compact K" using cc False by blast then obtain N where "open N" and "K \ N" by blast let ?\ = "insert (U \ V) ((\S. N - S) ` \)" obtain \ where "\ \ ?\" "finite \" "K \ \\" proof (rule compactE [OF \compact K\]) show "K \ \(insert (U \ V) ((-) N ` \))" using \K \ N\ ABeq \A \ U\ \B \ V\ by auto show "\B. B \ insert (U \ V) ((-) N ` \) \ open B" by (auto simp: \open U\ \open V\ open_Un \open N\ cc compact_imp_closed open_Diff) qed then have "finite(\ - {U \ V})" by blast moreover have "\ - {U \ V} \ (\S. N - S) ` \" using \\ \ ?\\ by blast ultimately obtain \ where "\ \ \" "finite \" and Deq: "\ - {U \ V} = (\S. N-S) ` \" using finite_subset_image by metis obtain J where "J \ \" and J: "(\S\\. N - S) \ N - J" proof (cases "\ = {}") case True with \\ \ {}\ that show ?thesis by auto next case False have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (meson \\ \ \\ in_mono local.linear) with \finite \\ \\ \ {}\ have "\J \ \. (\S\\. N - S) \ N - J" proof induction case (insert X \) show ?case proof (cases "\ = {}") case True then show ?thesis by auto next case False then have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (simp add: insert.prems) with insert.IH False obtain J where "J \ \" and J: "(\Y\\. N - Y) \ N - J" by metis have "N - J \ N - X \ N - X \ N - J" by (meson Diff_mono \J \ \\ insert.prems(2) insert_iff order_refl) then show ?thesis proof assume "N - J \ N - X" with J show ?thesis by auto next assume "N - X \ N - J" with J have "N - X \ \ ((-) N ` \) \ N - J" by auto with \J \ \\ show ?thesis by blast qed qed qed simp with \\ \ \\ show ?thesis by (blast intro: that) qed have "K \ \(insert (U \ V) (\ - {U \ V}))" using \K \ \\\ by auto also have "... \ (U \ V) \ (N - J)" by (metis (no_types, opaque_lifting) Deq Un_subset_iff Un_upper2 J Union_insert order_trans sup_ge1) finally have "J \ K \ U \ V" by blast moreover have "connected(J \ K)" by (metis Int_absorb1 \J \ \\ \K \ \\ cc inf.orderE local.linear) moreover have "U \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \A \ {}\ \A \ U\ by blast moreover have "V \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \B \ {}\ \B \ V\ by blast ultimately show False using connectedD [of "J \ K" U V] \open U\ \open V\ \U \ V = {}\ by auto qed with cf show ?thesis by (auto simp: connected_closed_set compact_imp_closed) qed lemma connected_chain_gen: fixes \ :: "'a :: euclidean_space set set" assumes X: "X \ \" "compact X" and cc: "\T. T \ \ \ closed T \ connected T" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof - have "\\ = (\T\\. X \ T)" using X by blast moreover have "connected (\T\\. X \ T)" proof (rule connected_chain) show "\T. T \ (\) X ` \ \ compact T \ connected T" using cc X by auto (metis inf.absorb2 inf.orderE local.linear) show "\S T. S \ (\) X ` \ \ T \ (\) X ` \ \ S \ T \ T \ S" using local.linear by blast qed ultimately show ?thesis by metis qed lemma connected_nest: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. compact(S n)" "\n. connected(S n)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" proof (rule connected_chain) show "\A T. A \ range S \ T \ range S \ A \ T \ T \ A" by (metis image_iff le_cases nest) qed (use S in blast) lemma connected_nest_gen: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. closed(S n)" "\n. connected(S n)" "compact(S k)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" proof (rule connected_chain_gen [of "S k"]) show "\A T. A \ range S \ T \ range S \ A \ T \ T \ A" by (metis imageE le_cases nest) qed (use S in auto) subsection\Proper maps, including projections out of compact sets\ lemma finite_indexed_bound: assumes A: "finite A" "\x. x \ A \ \n::'a::linorder. P x n" shows "\m. \x \ A. \k\m. P x k" using A proof (induction A) case empty then show ?case by force next case (insert a A) then obtain m n where "\x \ A. \k\m. P x k" "P a n" by force then show ?case by (metis dual_order.trans insert_iff le_cases) qed proposition proper_map: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes "closedin (top_of_set S) K" and com: "\U. \U \ T; compact U\ \ compact (S \ f -` U)" and "f ` S \ T" shows "closedin (top_of_set T) (f ` K)" proof - have "K \ S" using assms closedin_imp_subset by metis obtain C where "closed C" and Keq: "K = S \ C" using assms by (auto simp: closedin_closed) have *: "y \ f ` K" if "y \ T" and y: "y islimpt f ` K" for y proof - obtain h where "\n. (\x\K. h n = f x) \ h n \ y" "inj h" and hlim: "(h \ y) sequentially" using \y \ T\ y by (force simp: limpt_sequential_inj) then obtain X where X: "\n. X n \ K \ h n = f (X n) \ h n \ y" by metis then have fX: "\n. f (X n) = h n" by metis define \ where "\ \ \n. {a \ K. f a \ insert y (range (\i. f (X (n + i))))}" have "compact (C \ (S \ f -` insert y (range (\i. f(X(n + i))))))" for n proof (intro closed_Int_compact [OF \closed C\ com] compact_sequence_with_limit) show "insert y (range (\i. f (X (n + i)))) \ T" using X \K \ S\ \f ` S \ T\ \y \ T\ by blast show "(\i. f (X (n + i))) \ y" by (simp add: fX add.commute [of n] LIMSEQ_ignore_initial_segment [OF hlim]) qed then have comf: "compact (\ n)" for n by (simp add: Keq Int_def \_def conj_commute) have ne: "\\ \ {}" if "finite \" and \: "\t. t \ \ \ (\n. t = \ n)" for \ proof - obtain m where m: "\t. t \ \ \ \k\m. t = \ k" by (rule exE [OF finite_indexed_bound [OF \finite \\ \]], force+) have "X m \ \\" using X le_Suc_ex by (fastforce simp: \_def dest: m) then show ?thesis by blast qed have "(\n. \ n) \ {}" proof (rule compact_fip_Heine_Borel) show "\\'. \finite \'; \' \ range \\ \ \ \' \ {}" by (meson ne rangeE subset_eq) qed (use comf in blast) then obtain x where "x \ K" "\n. (f x = y \ (\u. f x = h (n + u)))" by (force simp add: \_def fX) then show ?thesis unfolding image_iff by (metis \inj h\ le_add1 not_less_eq_eq rangeI range_ex1_eq) qed with assms closedin_subset show ?thesis by (force simp: closedin_limpt) qed lemma compact_continuous_image_eq: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes f: "inj_on f S" shows "continuous_on S f \ (\T. compact T \ T \ S \ compact(f ` T))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis continuous_on_subset compact_continuous_image) next assume RHS: ?rhs obtain g where gf: "\x. x \ S \ g (f x) = x" by (metis inv_into_f_f f) then have *: "(S \ f -` U) = g ` U" if "U \ f ` S" for U using that by fastforce have gfim: "g ` f ` S \ S" using gf by auto have **: "compact (f ` S \ g -` C)" if C: "C \ S" "compact C" for C proof - obtain h where "h C \ C \ h C \ S \ compact (f ` C)" by (force simp: C RHS) moreover have "f ` C = (f ` S \ g -` C)" using C gf by auto ultimately show ?thesis using C by auto qed show ?lhs using proper_map [OF _ _ gfim] ** by (simp add: continuous_on_closed * closedin_imp_subset) qed subsection\<^marker>\tag unimportant\\Trivial fact: convexity equals connectedness for collinear sets\ lemma convex_connected_collinear: fixes S :: "'a::euclidean_space set" assumes "collinear S" shows "convex S \ connected S" proof assume "convex S" then show "connected S" using convex_connected by blast next assume S: "connected S" show "convex S" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain a where "a \ S" by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - a = c *\<^sub>R z" by (meson \a \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - a = f x *\<^sub>R z" by metis then have inj_f: "inj_on f (affine hull S)" by (metis diff_add_cancel inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - a" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - a" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" proof (clarsimp simp: dist_norm continuous_on_iff diff) show "\x e. 0 < e \ \d>0. \y \ affine hull S. \f y - f x\ * norm z < d \ \f y - f x\ < e" by (metis \z \ 0\ mult_pos_pos mult_less_iff1 zero_less_norm_iff) qed then have conn_fS: "connected (f ` S)" by (meson S connected_continuous_image continuous_on_subset hull_subset) show ?thesis proof (clarsimp simp: convex_contains_segment) fix x y z assume "x \ S" "y \ S" "z \ closed_segment x y" have False if "z \ S" proof - have "f ` (closed_segment x y) = closed_segment (f x) (f y)" proof (rule continuous_injective_image_segment_1) show "continuous_on (closed_segment x y) f" by (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) show "inj_on f (closed_segment x y)" by (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) qed then have fz: "f z \ closed_segment (f x) (f y)" using \z \ closed_segment x y\ by blast have "z \ affine hull S" by (meson \x \ S\ \y \ S\ \z \ closed_segment x y\ convex_affine_hull convex_contains_segment hull_inc subset_eq) then have fz_notin: "f z \ f ` S" using hull_subset inj_f inj_onD that by fastforce moreover have "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" proof - consider "f x \ f z \ f z \ f y" | "f y \ f z \ f z \ f x" using fz by (auto simp add: closed_segment_eq_real_ivl split: if_split_asm) then have "{.. f ` {x,y} \ {} \ {f z<..} \ f ` {x,y} \ {}" by cases (use fz_notin \x \ S\ \y \ S\ in \auto simp: image_iff\) then show "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" using \x \ S\ \y \ S\ by blast+ qed ultimately show False using connectedD [OF conn_fS, of "{.. S" by meson qed qed qed lemma compact_convex_collinear_segment_alt: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "connected S" "collinear S" obtains a b where "S = closed_segment a b" proof - obtain \ where "\ \ S" using \S \ {}\ by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - \ = c *\<^sub>R z" by (meson \\ \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - \ = f x *\<^sub>R z" by metis let ?g = "\r. r *\<^sub>R z + \" have gf: "?g (f x) = x" if "x \ affine hull S" for x by (metis diff_add_cancel f that) then have inj_f: "inj_on f (affine hull S)" by (metis inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - \" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - \" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" proof (clarsimp simp: dist_norm continuous_on_iff diff) show "\x e. 0 < e \ \d>0. \y \ affine hull S. \f y - f x\ * norm z < d \ \f y - f x\ < e" by (metis \z \ 0\ mult_pos_pos mult_less_iff1 zero_less_norm_iff) qed then have "connected (f ` S)" by (meson \connected S\ connected_continuous_image continuous_on_subset hull_subset) moreover have "compact (f ` S)" by (meson \compact S\ compact_continuous_image_eq cont_f hull_subset inj_f) ultimately obtain x y where "f ` S = {x..y}" by (meson connected_compact_interval_1) then have fS_eq: "f ` S = closed_segment x y" using \S \ {}\ closed_segment_eq_real_ivl by auto obtain a b where "a \ S" "f a = x" "b \ S" "f b = y" by (metis (full_types) ends_in_segment fS_eq imageE) have "f ` (closed_segment a b) = closed_segment (f a) (f b)" proof (rule continuous_injective_image_segment_1) show "continuous_on (closed_segment a b) f" by (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) show "inj_on f (closed_segment a b)" by (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) qed then have "f ` (closed_segment a b) = f ` S" by (simp add: \f a = x\ \f b = y\ fS_eq) then have "?g ` f ` (closed_segment a b) = ?g ` f ` S" by simp moreover have "(\x. f x *\<^sub>R z + \) ` closed_segment a b = closed_segment a b" unfolding image_def using \a \ S\ \b \ S\ by (safe; metis (mono_tags, lifting) convex_affine_hull convex_contains_segment gf hull_subset subsetCE) ultimately have "closed_segment a b = S" using gf by (simp add: image_comp o_def hull_inc cong: image_cong) then show ?thesis using that by blast qed lemma compact_convex_collinear_segment: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "convex S" "collinear S" obtains a b where "S = closed_segment a b" using assms convex_connected_collinear compact_convex_collinear_segment_alt by blast lemma proper_map_from_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and imf: "f ` S \ T" and "compact S" "closedin (top_of_set T) K" shows "compact (S \ f -` K)" by (rule closedin_compact [OF \compact S\] continuous_closedin_preimage_gen assms)+ lemma proper_map_fst: assumes "compact T" "K \ S" "compact K" shows "compact (S \ T \ fst -` K)" proof - have "(S \ T \ fst -` K) = K \ T" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_fst: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact T" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set S) (fst ` c)" proof - have *: "fst ` (S \ T) \ S" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_fst assms) qed lemma proper_map_snd: assumes "compact S" "K \ T" "compact K" shows "compact (S \ T \ snd -` K)" proof - have "(S \ T \ snd -` K) = S \ K" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_snd: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set T) (snd ` c)" proof - have *: "snd ` (S \ T) \ T" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_snd assms) qed lemma closedin_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" and clo: "closedin (top_of_set (S \ T)) U" shows "closedin (top_of_set T) {y. \x. x \ S \ (x, y) \ U}" proof - have "U \ S \ T" by (metis clo closedin_imp_subset) then have "{y. \x. x \ S \ (x, y) \ U} = snd ` U" by force moreover have "closedin (top_of_set T) (snd ` U)" by (rule closed_map_snd [OF assms]) ultimately show ?thesis by simp qed lemma closed_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "('a * 'b::euclidean_space) set" assumes "compact S" and clo: "closed T" shows "closed {y. \x. x \ S \ (x, y) \ T}" proof - have *: "{y. \x. x \ S \ Pair x y \ T} = {y. \x. x \ S \ Pair x y \ ((S \ UNIV) \ T)}" by auto show ?thesis unfolding * by (intro clo closedin_closed_Int closedin_closed_trans [OF _ closed_UNIV] closedin_compact_projection [OF \compact S\]) qed subsubsection\<^marker>\tag unimportant\\Representing affine hull as a finite intersection of hyperplanes\ proposition\<^marker>\tag unimportant\ affine_hull_convex_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "convex S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" proof show "affine hull (S \ T) \ affine hull S" by (simp add: hull_mono) next obtain a where "a \ S" "a \ T" and at: "a \ interior T" using assms interior_subset by blast then obtain e where "e > 0" and e: "cball a e \ T" using mem_interior_cball by blast have *: "x \ (+) a ` span ((\x. x - a) ` (S \ T))" if "x \ S" for x proof (cases "x = a") case True with that span_0 eq_add_iff image_def mem_Collect_eq show ?thesis by blast next case False define k where "k = min (1/2) (e / norm (x-a))" have k: "0 < k" "k < 1" using \e > 0\ False by (auto simp: k_def) then have xa: "(x-a) = inverse k *\<^sub>R k *\<^sub>R (x-a)" by simp have "e / norm (x - a) \ k" using k_def by linarith then have "a + k *\<^sub>R (x - a) \ cball a e" using \0 < k\ False by (simp add: dist_norm) (simp add: field_simps) then have T: "a + k *\<^sub>R (x - a) \ T" using e by blast have S: "a + k *\<^sub>R (x - a) \ S" using k \a \ S\ convexD [OF \convex S\ \a \ S\ \x \ S\, of "1-k" k] by (simp add: algebra_simps) have "inverse k *\<^sub>R k *\<^sub>R (x-a) \ span ((\x. x - a) ` (S \ T))" by (intro span_mul [OF span_base] image_eqI [where x = "a + k *\<^sub>R (x - a)"]) (auto simp: S T) with xa image_iff show ?thesis by fastforce qed have "S \ affine hull (S \ T)" by (force simp: * \a \ S\ \a \ T\ hull_inc affine_hull_span_gen [of a]) then show "affine hull S \ affine hull (S \ T)" by (simp add: subset_hull) qed corollary affine_hull_convex_Int_open: fixes S :: "'a::real_normed_vector set" assumes "convex S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using affine_hull_convex_Int_nonempty_interior assms interior_eq by blast corollary affine_hull_affine_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "affine S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_nonempty_interior affine_imp_convex assms) corollary affine_hull_affine_Int_open: fixes S :: "'a::real_normed_vector set" assumes "affine S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_open affine_imp_convex assms) corollary affine_hull_convex_Int_openin: fixes S :: "'a::real_normed_vector set" assumes "convex S" "openin (top_of_set (affine hull S)) T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using assms unfolding openin_open by (metis affine_hull_convex_Int_open hull_subset inf.orderE inf_assoc) corollary affine_hull_openin: fixes S :: "'a::real_normed_vector set" assumes "openin (top_of_set (affine hull T)) S" "S \ {}" shows "affine hull S = affine hull T" using assms unfolding openin_open by (metis affine_affine_hull affine_hull_affine_Int_open hull_hull) corollary affine_hull_open: fixes S :: "'a::real_normed_vector set" assumes "open S" "S \ {}" shows "affine hull S = UNIV" by (metis affine_hull_convex_Int_nonempty_interior assms convex_UNIV hull_UNIV inf_top.left_neutral interior_open) lemma aff_dim_convex_Int_nonempty_interior: fixes S :: "'a::euclidean_space set" shows "\convex S; S \ interior T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_affine_hull2 affine_hull_convex_Int_nonempty_interior by blast lemma aff_dim_convex_Int_open: fixes S :: "'a::euclidean_space set" shows "\convex S; open T; S \ T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_convex_Int_nonempty_interior interior_eq by blast lemma affine_hull_Diff: fixes S:: "'a::real_normed_vector set" assumes ope: "openin (top_of_set (affine hull S)) S" and "finite F" "F \ S" shows "affine hull (S - F) = affine hull S" proof - have clo: "closedin (top_of_set S) F" using assms finite_imp_closedin by auto moreover have "S - F \ {}" using assms by auto ultimately show ?thesis by (metis ope closedin_def topspace_euclidean_subtopology affine_hull_openin openin_trans) qed lemma affine_hull_halfspace_lt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x < r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_lt [of a r] by (simp add: open_halfspace_lt affine_hull_open) lemma affine_hull_halfspace_le: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r < 0 then {} else UNIV)" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "affine hull closure {x. a \ x < r} = UNIV" using affine_hull_halfspace_lt closure_same_affine_hull by fastforce moreover have "{x. a \ x < r} \ {x. a \ x \ r}" by (simp add: Collect_mono) ultimately show ?thesis using False antisym_conv hull_mono top_greatest by (metis affine_hull_halfspace_lt) qed lemma affine_hull_halfspace_gt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x > r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_gt [of r a] by (simp add: open_halfspace_gt affine_hull_open) lemma affine_hull_halfspace_ge: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r > 0 then {} else UNIV)" using affine_hull_halfspace_le [of "-a" "-r"] by simp lemma aff_dim_halfspace_lt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x < r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_lt open_halfspace_lt) lemma aff_dim_halfspace_le: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r < 0 then -1 else DIM('a))" proof - have "int (DIM('a)) = aff_dim (UNIV::'a set)" by (simp) then have "aff_dim (affine hull {x. a \ x \ r}) = DIM('a)" if "(a = 0 \ r \ 0)" using that by (simp add: affine_hull_halfspace_le not_less) then show ?thesis by (force) qed lemma aff_dim_halfspace_gt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x > r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_gt open_halfspace_gt) lemma aff_dim_halfspace_ge: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r > 0 then -1 else DIM('a))" using aff_dim_halfspace_le [of "-a" "-r"] by simp proposition aff_dim_eq_hyperplane: fixes S :: "'a::euclidean_space set" shows "aff_dim S = DIM('a) - 1 \ (\a b. a \ 0 \ affine hull S = {x. a \ x = b})" (is "?lhs = ?rhs") proof (cases "S = {}") case True then show ?thesis by (auto simp: dest: hyperplane_eq_Ex) next case False then obtain c where "c \ S" by blast show ?thesis proof (cases "c = 0") case True have "?lhs \ (\a. a \ 0 \ span ((\x. x - c) ` S) = {x. a \ x = 0})" by (simp add: aff_dim_eq_dim [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) also have "... \ ?rhs" using span_zero [of S] True \c \ S\ affine_hull_span_0 hull_inc by (fastforce simp add: affine_hull_span_gen [of c] \c = 0\) finally show ?thesis . next case False have xc_im: "x \ (+) c ` {y. a \ y = 0}" if "a \ x = a \ c" for a x proof - have "\y. a \ y = 0 \ c + y = x" by (metis that add.commute diff_add_cancel inner_commute inner_diff_left right_minus_eq) then show "x \ (+) c ` {y. a \ y = 0}" by blast qed have 2: "span ((\x. x - c) ` S) = {x. a \ x = 0}" if "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = b}" for a b proof - have "b = a \ c" using span_0 that by fastforce with that have "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = a \ c}" by simp then have "span ((\x. x - c) ` S) = (\x. x - c) ` {x. a \ x = a \ c}" by (metis (no_types) image_cong translation_galois uminus_add_conv_diff) also have "... = {x. a \ x = 0}" by (force simp: inner_distrib inner_diff_right intro: image_eqI [where x="x+c" for x]) finally show ?thesis . qed have "?lhs = (\a. a \ 0 \ span ((\x. x - c) ` S) = {x. a \ x = 0})" by (simp add: aff_dim_eq_dim [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) also have "... = ?rhs" by (fastforce simp add: affine_hull_span_gen [of c] \c \ S\ hull_inc inner_distrib intro: xc_im intro!: 2) finally show ?thesis . qed qed corollary aff_dim_hyperplane [simp]: fixes a :: "'a::euclidean_space" shows "a \ 0 \ aff_dim {x. a \ x = r} = DIM('a) - 1" by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane) subsection\<^marker>\tag unimportant\\Some stepping theorems\ lemma aff_dim_insert: fixes a :: "'a::euclidean_space" shows "aff_dim (insert a S) = (if a \ affine hull S then aff_dim S else aff_dim S + 1)" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain x s' where S: "S = insert x s'" "x \ s'" by (meson Set.set_insert all_not_in_conv) show ?thesis using S by (force simp add: affine_hull_insert_span_gen span_zero insert_commute [of a] aff_dim_eq_dim [of x] dim_insert) qed lemma affine_dependent_choose: fixes a :: "'a :: euclidean_space" assumes "\(affine_dependent S)" shows "affine_dependent(insert a S) \ a \ S \ a \ affine hull S" (is "?lhs = ?rhs") proof safe assume "affine_dependent (insert a S)" and "a \ S" then show "False" using \a \ S\ assms insert_absorb by fastforce next assume lhs: "affine_dependent (insert a S)" then have "a \ S" by (metis (no_types) assms insert_absorb) moreover have "finite S" using affine_independent_iff_card assms by blast moreover have "aff_dim (insert a S) \ int (card S)" using \finite S\ affine_independent_iff_card \a \ S\ lhs by fastforce ultimately show "a \ affine hull S" by (metis aff_dim_affine_independent aff_dim_insert assms) next assume "a \ S" and "a \ affine hull S" show "affine_dependent (insert a S)" by (simp add: \a \ affine hull S\ \a \ S\ affine_dependent_def) qed lemma affine_independent_insert: fixes a :: "'a :: euclidean_space" shows "\\ affine_dependent S; a \ affine hull S\ \ \ affine_dependent(insert a S)" by (simp add: affine_dependent_choose) lemma subspace_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "subspace S" shows "bounded S \ S = {0}" proof - have "False" if "bounded S" "x \ S" "x \ 0" for x proof - obtain B where B: "\y. y \ S \ norm y < B" "B > 0" using \bounded S\ by (force simp: bounded_pos_less) have "(B / norm x) *\<^sub>R x \ S" using assms subspace_mul \x \ S\ by auto moreover have "norm ((B / norm x) *\<^sub>R x) = B" using that B by (simp add: algebra_simps) ultimately show False using B by force qed then have "bounded S \ S = {0}" using assms subspace_0 by fastforce then show ?thesis by blast qed lemma affine_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "affine S" shows "bounded S \ S = {} \ (\a. S = {a})" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain b where "b \ S" by blast with False assms have "bounded S \ S = {b}" using affine_diffs_subspace [OF assms \b \ S\] by (metis (no_types, lifting) ab_group_add_class.ab_left_minus bounded_translation image_empty image_insert subspace_bounded_eq_trivial translation_invert) then show ?thesis by auto qed lemma affine_bounded_eq_lowdim: fixes S :: "'a::euclidean_space set" assumes "affine S" shows "bounded S \ aff_dim S \ 0" proof show "aff_dim S \ 0 \ bounded S" by (metis aff_dim_sing aff_dim_subset affine_dim_equal affine_sing all_not_in_conv assms bounded_empty bounded_insert dual_order.antisym empty_subsetI insert_subset) qed (use affine_bounded_eq_trivial assms in fastforce) lemma bounded_hyperplane_eq_trivial_0: fixes a :: "'a::euclidean_space" assumes "a \ 0" shows "bounded {x. a \ x = 0} \ DIM('a) = 1" proof assume "bounded {x. a \ x = 0}" then have "aff_dim {x. a \ x = 0} \ 0" by (simp add: affine_bounded_eq_lowdim affine_hyperplane) with assms show "DIM('a) = 1" by (simp add: le_Suc_eq) next assume "DIM('a) = 1" then show "bounded {x. a \ x = 0}" by (simp add: affine_bounded_eq_lowdim affine_hyperplane assms) qed lemma bounded_hyperplane_eq_trivial: fixes a :: "'a::euclidean_space" shows "bounded {x. a \ x = r} \ (if a = 0 then r \ 0 else DIM('a) = 1)" proof (simp add: bounded_hyperplane_eq_trivial_0, clarify) assume "r \ 0" "a \ 0" have "aff_dim {x. y \ x = 0} = aff_dim {x. a \ x = r}" if "y \ 0" for y::'a by (metis that \a \ 0\ aff_dim_hyperplane) then show "bounded {x. a \ x = r} = (DIM('a) = Suc 0)" by (metis One_nat_def \a \ 0\ affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0) qed subsection\<^marker>\tag unimportant\\General case without assuming closure and getting non-strict separation\ proposition\<^marker>\tag unimportant\ separating_hyperplane_closed_point_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "z \ S" obtains a b where "a \ S" "(a - z) \ z < b" "\x. x \ S \ b < (a - z) \ x" proof - obtain y where "y \ S" and y: "\u. u \ S \ dist z y \ dist z u" using distance_attains_inf [of S z] assms by auto then have *: "(y - z) \ z < (y - z) \ z + (norm (y - z))\<^sup>2 / 2" using \y \ S\ \z \ S\ by auto show ?thesis proof (rule that [OF \y \ S\ *]) fix x assume "x \ S" have yz: "0 < (y - z) \ (y - z)" using \y \ S\ \z \ S\ by auto { assume 0: "0 < ((z - y) \ (x - y))" with any_closest_point_dot [OF \convex S\ \closed S\] have False using y \x \ S\ \y \ S\ not_less by blast } then have "0 \ ((y - z) \ (x - y))" by (force simp: not_less inner_diff_left) with yz have "0 < 2 * ((y - z) \ (x - y)) + (y - z) \ (y - z)" by (simp add: algebra_simps) then show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" by (simp add: field_simps inner_diff_left inner_diff_right dot_square_norm [symmetric]) qed qed lemma separating_hyperplane_closed_0_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "0 \ S" obtains a b where "a \ S" "a \ 0" "0 < b" "\x. x \ S \ a \ x > b" using separating_hyperplane_closed_point_inset [OF assms] by simp (metis \0 \ S\) proposition\<^marker>\tag unimportant\ separating_hyperplane_set_0_inspan: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" "0 \ S" obtains a where "a \ span S" "a \ 0" "\x. x \ S \ 0 \ a \ x" proof - define k where [abs_def]: "k c = {x. 0 \ c \ x}" for c :: 'a have "span S \ frontier (cball 0 1) \ \f' \ {}" if f': "finite f'" "f' \ k ` S" for f' proof - obtain C where "C \ S" "finite C" and C: "f' = k ` C" using finite_subset_image [OF f'] by blast obtain a where "a \ S" "a \ 0" using \S \ {}\ \0 \ S\ ex_in_conv by blast then have "norm (a /\<^sub>R (norm a)) = 1" by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp show ?thesis proof (cases "C = {}") case True with C ass show ?thesis by auto next case False have "closed (convex hull C)" using \finite C\ compact_eq_bounded_closed finite_imp_compact_convex_hull by auto moreover have "convex hull C \ {}" by (simp add: False) moreover have "0 \ convex hull C" by (metis \C \ S\ \convex S\ \0 \ S\ convex_hull_subset hull_same insert_absorb insert_subset) ultimately obtain a b where "a \ convex hull C" "a \ 0" "0 < b" and ab: "\x. x \ convex hull C \ a \ x > b" using separating_hyperplane_closed_0_inset by blast then have "a \ S" by (metis \C \ S\ assms(1) subsetCE subset_hull) moreover have "norm (a /\<^sub>R (norm a)) = 1" using \a \ 0\ by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp have "\x. \a \ 0; x \ C\ \ 0 \ x \ a" using ab \0 < b\ by (metis hull_inc inner_commute less_eq_real_def less_trans) then have aa: "a /\<^sub>R (norm a) \ (\c\C. {x. 0 \ c \ x})" by (auto simp add: field_split_simps) show ?thesis unfolding C k_def using ass aa Int_iff empty_iff by force qed qed moreover have "\T. T \ k ` S \ closed T" using closed_halfspace_ge k_def by blast ultimately have "(span S \ frontier(cball 0 1)) \ (\ (k ` S)) \ {}" by (metis compact_imp_fip closed_Int_compact closed_span compact_cball compact_frontier) then show ?thesis unfolding set_eq_iff k_def by simp (metis inner_commute norm_eq_zero that zero_neq_one) qed lemma separating_hyperplane_set_point_inaff: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" and zno: "z \ S" obtains a b where "(z + a) \ affine hull (insert z S)" and "a \ 0" and "a \ z \ b" and "\x. x \ S \ a \ x \ b" proof - from separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] have "convex ((+) (- z) ` S)" using \convex S\ by simp moreover have "(+) (- z) ` S \ {}" by (simp add: \S \ {}\) moreover have "0 \ (+) (- z) ` S" using zno by auto ultimately obtain a where "a \ span ((+) (- z) ` S)" "a \ 0" and a: "\x. x \ ((+) (- z) ` S) \ 0 \ a \ x" using separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] by blast then have szx: "\x. x \ S \ a \ z \ a \ x" by (metis (no_types, lifting) imageI inner_minus_right inner_right_distrib minus_add neg_le_0_iff_le neg_le_iff_le real_add_le_0_iff) moreover have "z + a \ affine hull insert z S" using \a \ span ((+) (- z) ` S)\ affine_hull_insert_span_gen by blast ultimately show ?thesis using \a \ 0\ szx that by auto qed proposition\<^marker>\tag unimportant\ supporting_hyperplane_rel_boundary: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ S" and xno: "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" proof - obtain a b where aff: "(x + a) \ affine hull (insert x (rel_interior S))" and "a \ 0" and "a \ x \ b" and ageb: "\u. u \ (rel_interior S) \ a \ u \ b" using separating_hyperplane_set_point_inaff [of "rel_interior S" x] assms by (auto simp: rel_interior_eq_empty convex_rel_interior) have le_ay: "a \ x \ a \ y" if "y \ S" for y proof - have con: "continuous_on (closure (rel_interior S)) ((\) a)" by (rule continuous_intros continuous_on_subset | blast)+ have y: "y \ closure (rel_interior S)" using \convex S\ closure_def convex_closure_rel_interior \y \ S\ by fastforce show ?thesis using continuous_ge_on_closure [OF con y] ageb \a \ x \ b\ by fastforce qed have 3: "a \ x < a \ y" if "y \ rel_interior S" for y proof - obtain e where "0 < e" "y \ S" and e: "cball y e \ affine hull S \ S" using \y \ rel_interior S\ by (force simp: rel_interior_cball) define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)" have "y' \ cball y e" unfolding y'_def using \0 < e\ by force moreover have "y' \ affine hull S" unfolding y'_def by (metis \x \ S\ \y \ S\ \convex S\ aff affine_affine_hull hull_redundant rel_interior_same_affine_hull hull_inc mem_affine_3_minus2) ultimately have "y' \ S" using e by auto have "a \ x \ a \ y" using le_ay \a \ 0\ \y \ S\ by blast moreover have "a \ x \ a \ y" using le_ay [OF \y' \ S\] \a \ 0\ \0 < e\ not_le by (fastforce simp add: y'_def inner_diff dot_square_norm power2_eq_square) ultimately show ?thesis by force qed show ?thesis by (rule that [OF \a \ 0\ le_ay 3]) qed lemma supporting_hyperplane_relative_frontier: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ closure S" "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ closure S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" using supporting_hyperplane_rel_boundary [of "closure S" x] by (metis assms convex_closure convex_rel_interior_closure) subsection\<^marker>\tag unimportant\\ Some results on decomposing convex hulls: intersections, simplicial subdivision\ lemma fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent(S \ T)" shows convex_hull_Int_subset: "convex hull S \ convex hull T \ convex hull (S \ T)" (is ?C) and affine_hull_Int_subset: "affine hull S \ affine hull T \ affine hull (S \ T)" (is ?A) proof - have [simp]: "finite S" "finite T" using aff_independent_finite assms by blast+ have "sum u (S \ T) = 1 \ (\v\S \ T. u v *\<^sub>R v) = (\v\S. u v *\<^sub>R v)" if [simp]: "sum u S = 1" "sum v T = 1" and eq: "(\x\T. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" for u v proof - define f where "f x = (if x \ S then u x else 0) - (if x \ T then v x else 0)" for x have "sum f (S \ T) = 0" by (simp add: f_def sum_Un sum_subtractf flip: sum.inter_restrict) moreover have "(\x\(S \ T). f x *\<^sub>R x) = 0" by (simp add: eq f_def sum_Un scaleR_left_diff_distrib sum_subtractf if_smult flip: sum.inter_restrict cong: if_cong) ultimately have "\v. v \ S \ T \ f v = 0" using aff_independent_finite assms unfolding affine_dependent_explicit by blast then have u [simp]: "\x. x \ S \ u x = (if x \ T then v x else 0)" by (simp add: f_def) presburger have "sum u (S \ T) = sum u S" by (simp add: sum.inter_restrict) then have "sum u (S \ T) = 1" using that by linarith moreover have "(\v\S \ T. u v *\<^sub>R v) = (\v\S. u v *\<^sub>R v)" by (auto simp: if_smult sum.inter_restrict intro: sum.cong) ultimately show ?thesis by force qed then show ?A ?C by (auto simp: convex_hull_finite affine_hull_finite) qed proposition\<^marker>\tag unimportant\ affine_hull_Int: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent(S \ T)" shows "affine hull (S \ T) = affine hull S \ affine hull T" by (simp add: affine_hull_Int_subset assms hull_mono subset_antisym) proposition\<^marker>\tag unimportant\ convex_hull_Int: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent(S \ T)" shows "convex hull (S \ T) = convex hull S \ convex hull T" by (simp add: convex_hull_Int_subset assms hull_mono subset_antisym) proposition\<^marker>\tag unimportant\ fixes S :: "'a::euclidean_space set set" assumes "\ affine_dependent (\S)" shows affine_hull_Inter: "affine hull (\S) = (\T\S. affine hull T)" (is "?A") and convex_hull_Inter: "convex hull (\S) = (\T\S. convex hull T)" (is "?C") proof - have "finite S" using aff_independent_finite assms finite_UnionD by blast then have "?A \ ?C" using assms proof (induction S rule: finite_induct) case empty then show ?case by auto next case (insert T F) then show ?case proof (cases "F={}") case True then show ?thesis by simp next case False with "insert.prems" have [simp]: "\ affine_dependent (T \ \F)" by (auto intro: affine_dependent_subset) have [simp]: "\ affine_dependent (\F)" using affine_independent_subset insert.prems by fastforce show ?thesis by (simp add: affine_hull_Int convex_hull_Int insert.IH) qed qed then show "?A" "?C" by auto qed proposition\<^marker>\tag unimportant\ in_convex_hull_exchange_unique: fixes S :: "'a::euclidean_space set" assumes naff: "\ affine_dependent S" and a: "a \ convex hull S" and S: "T \ S" "T' \ S" and x: "x \ convex hull (insert a T)" and x': "x \ convex hull (insert a T')" shows "x \ convex hull (insert a (T \ T'))" proof (cases "a \ S") case True then have "\ affine_dependent (insert a T \ insert a T')" using affine_dependent_subset assms by auto then have "x \ convex hull (insert a T \ insert a T')" by (metis IntI convex_hull_Int x x') then show ?thesis by simp next case False then have anot: "a \ T" "a \ T'" using assms by auto have [simp]: "finite S" by (simp add: aff_independent_finite assms) then obtain b where b0: "\s. s \ S \ 0 \ b s" and b1: "sum b S = 1" and aeq: "a = (\s\S. b s *\<^sub>R s)" using a by (auto simp: convex_hull_finite) have fin [simp]: "finite T" "finite T'" using assms infinite_super \finite S\ by blast+ then obtain c c' where c0: "\t. t \ insert a T \ 0 \ c t" and c1: "sum c (insert a T) = 1" and xeq: "x = (\t \ insert a T. c t *\<^sub>R t)" and c'0: "\t. t \ insert a T' \ 0 \ c' t" and c'1: "sum c' (insert a T') = 1" and x'eq: "x = (\t \ insert a T'. c' t *\<^sub>R t)" using x x' by (auto simp: convex_hull_finite) with fin anot have sumTT': "sum c T = 1 - c a" "sum c' T' = 1 - c' a" and wsumT: "(\t \ T. c t *\<^sub>R t) = x - c a *\<^sub>R a" by simp_all have wsumT': "(\t \ T'. c' t *\<^sub>R t) = x - c' a *\<^sub>R a" using x'eq fin anot by simp define cc where "cc \ \x. if x \ T then c x else 0" define cc' where "cc' \ \x. if x \ T' then c' x else 0" define dd where "dd \ \x. cc x - cc' x + (c a - c' a) * b x" have sumSS': "sum cc S = 1 - c a" "sum cc' S = 1 - c' a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 sum_subtractf sum.inter_restrict [symmetric] sumTT') have wsumSS: "(\t \ S. cc t *\<^sub>R t) = x - c a *\<^sub>R a" "(\t \ S. cc' t *\<^sub>R t) = x - c' a *\<^sub>R a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 if_smult sum.inter_restrict [symmetric] wsumT wsumT' cong: if_cong) have sum_dd0: "sum dd S = 0" unfolding dd_def using S by (simp add: sumSS' comm_monoid_add_class.sum.distrib sum_subtractf algebra_simps sum_distrib_right [symmetric] b1) have "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R (\v\S. b v *\<^sub>R v)" for x by (simp add: pth_5 real_vector.scale_sum_right mult.commute) then have *: "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R a" for x using aeq by blast have "(\v \ S. dd v *\<^sub>R v) = 0" unfolding dd_def using S by (simp add: * wsumSS sum.distrib sum_subtractf algebra_simps) then have dd0: "dd v = 0" if "v \ S" for v using naff [unfolded affine_dependent_explicit not_ex, rule_format, of S dd] using that sum_dd0 by force consider "c' a \ c a" | "c a \ c' a" by linarith then show ?thesis proof cases case 1 then have "sum cc S \ sum cc' S" by (simp add: sumSS') then have le: "cc x \ cc' x" if "x \ S" for x using dd0 [OF that] 1 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc(a := c a)) x" by (simp add: c0 cc_def) show "0 \ (cc(a := c a)) a" by (simp add: c0) have "sum (cc(a := c a)) (insert a (T \ T')) = c a + sum (cc(a := c a)) (T \ T')" by (simp add: anot) also have "... = c a + sum (cc(a := c a)) S" using \T \ S\ False cc0 cc_def \a \ S\ by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c a + (1 - c a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS'(1)) finally show "sum (cc(a := c a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = c a *\<^sub>R a + (\x \ T \ T'. (cc(a := c a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c a *\<^sub>R a + (\x \ S. (cc(a := c a)) x *\<^sub>R x)" using \T \ S\ False cc0 by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = x" by simp qed next case 2 then have "sum cc' S \ sum cc S" by (simp add: sumSS') then have le: "cc' x \ cc x" if "x \ S" for x using dd0 [OF that] 2 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc' x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c'0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc'(a := c' a)) x" by (simp add: c'0 cc'_def) show "0 \ (cc'(a := c' a)) a" by (simp add: c'0) have "sum (cc'(a := c' a)) (insert a (T \ T')) = c' a + sum (cc'(a := c' a)) (T \ T')" by (simp add: anot) also have "... = c' a + sum (cc'(a := c' a)) S" using \T \ S\ False cc0 by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c' a + (1 - c' a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS') finally show "sum (cc'(a := c' a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = c' a *\<^sub>R a + (\x \ T \ T'. (cc'(a := c' a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c' a *\<^sub>R a + (\x \ S. (cc'(a := c' a)) x *\<^sub>R x)" using \T \ S\ False cc0 by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = x" by simp qed qed qed corollary\<^marker>\tag unimportant\ convex_hull_exchange_Int: fixes a :: "'a::euclidean_space" assumes "\ affine_dependent S" "a \ convex hull S" "T \ S" "T' \ S" shows "(convex hull (insert a T)) \ (convex hull (insert a T')) = convex hull (insert a (T \ T'))" (is "?lhs = ?rhs") proof (rule subset_antisym) show "?lhs \ ?rhs" using in_convex_hull_exchange_unique assms by blast show "?rhs \ ?lhs" by (metis hull_mono inf_le1 inf_le2 insert_inter_insert le_inf_iff) qed lemma Int_closed_segment: fixes b :: "'a::euclidean_space" assumes "b \ closed_segment a c \ \ collinear{a,b,c}" shows "closed_segment a b \ closed_segment b c = {b}" proof (cases "c = a") case True then show ?thesis using assms collinear_3_eq_affine_dependent by fastforce next case False from assms show ?thesis proof assume "b \ closed_segment a c" moreover have "\ affine_dependent {a, c}" by (simp) ultimately show ?thesis using False convex_hull_exchange_Int [of "{a,c}" b "{a}" "{c}"] by (simp add: segment_convex_hull insert_commute) next assume ncoll: "\ collinear {a, b, c}" have False if "closed_segment a b \ closed_segment b c \ {b}" proof - have "b \ closed_segment a b" and "b \ closed_segment b c" by auto with that obtain d where "b \ d" "d \ closed_segment a b" "d \ closed_segment b c" by force then have d: "collinear {a, d, b}" "collinear {b, d, c}" by (auto simp: between_mem_segment between_imp_collinear) have "collinear {a, b, c}" by (metis (full_types) \b \ d\ collinear_3_trans d insert_commute) with ncoll show False .. qed then show ?thesis by blast qed qed lemma affine_hull_finite_intersection_hyperplanes: fixes S :: "'a::euclidean_space set" obtains \ where "finite \" "of_nat (card \) + aff_dim S = DIM('a)" "affine hull S = \\" "\h. h \ \ \ \a b. a \ 0 \ h = {x. a \ x = b}" proof - obtain b where "b \ S" and indb: "\ affine_dependent b" and eq: "affine hull S = affine hull b" using affine_basis_exists by blast obtain c where indc: "\ affine_dependent c" and "b \ c" and affc: "affine hull c = UNIV" by (metis extend_to_affine_basis affine_UNIV hull_same indb subset_UNIV) then have "finite c" by (simp add: aff_independent_finite) then have fbc: "finite b" "card b \ card c" using \b \ c\ infinite_super by (auto simp: card_mono) have imeq: "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b)) = ((\a. affine hull (c - {a})) ` (c - b))" by blast have card_cb: "(card (c - b)) + aff_dim S = DIM('a)" proof - have aff: "aff_dim (UNIV::'a set) = aff_dim c" by (metis aff_dim_affine_hull affc) have "aff_dim b = aff_dim S" by (metis (no_types) aff_dim_affine_hull eq) then have "int (card b) = 1 + aff_dim S" by (simp add: aff_dim_affine_independent indb) then show ?thesis using fbc aff by (simp add: \\ affine_dependent c\ \b \ c\ aff_dim_affine_independent card_Diff_subset of_nat_diff) qed show ?thesis proof (cases "c = b") case True show ?thesis proof show "int (card {}) + aff_dim S = int DIM('a)" using True card_cb by auto show "affine hull S = \ {}" using True affc eq by blast qed auto next case False have ind: "\ affine_dependent (\a\c - b. c - {a})" by (rule affine_independent_subset [OF indc]) auto have *: "1 + aff_dim (c - {t}) = int (DIM('a))" if t: "t \ c" for t proof - have "insert t c = c" using t by blast then show ?thesis by (metis (full_types) add.commute aff_dim_affine_hull aff_dim_insert aff_dim_UNIV affc affine_dependent_def indc insert_Diff_single t) qed let ?\ = "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b))" show ?thesis proof have "card ((\a. affine hull (c - {a})) ` (c - b)) = card (c - b)" proof (rule card_image) show "inj_on (\a. affine hull (c - {a})) (c - b)" unfolding inj_on_def by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff) qed then show "int (card ?\) + aff_dim S = int DIM('a)" by (simp add: imeq card_cb) show "affine hull S = \ ?\" by (metis Diff_eq_empty_iff INT_simps(4) UN_singleton order_refl \b \ c\ False eq double_diff affine_hull_Inter [OF ind]) have "\a. \a \ c; a \ b\ \ aff_dim (c - {a}) = int (DIM('a) - Suc 0)" by (metis "*" DIM_ge_Suc0 One_nat_def add_diff_cancel_left' int_ops(2) of_nat_diff) then show "\h. h \ ?\ \ \a b. a \ 0 \ h = {x. a \ x = b}" by (auto simp only: One_nat_def aff_dim_eq_hyperplane [symmetric]) qed (use \finite c\ in auto) qed qed lemma affine_hyperplane_sums_eq_UNIV_0: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "0 \ S" and "w \ S" and "a \ w \ 0" shows "{x + y| x y. x \ S \ a \ y = 0} = UNIV" proof - have "subspace S" by (simp add: assms subspace_affine) have span1: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" using \0 \ S\ add.left_neutral by (intro span_mono) force have "w \ span {y. a \ y = 0}" using \a \ w \ 0\ span_induct subspace_hyperplane by auto moreover have "w \ span {x + y |x y. x \ S \ a \ y = 0}" using \w \ S\ by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_base) ultimately have span2: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" by blast have "a \ 0" using assms inner_zero_left by blast then have "DIM('a) - 1 = dim {y. a \ y = 0}" by (simp add: dim_hyperplane) also have "... < dim {x + y |x y. x \ S \ a \ y = 0}" using span1 span2 by (blast intro: dim_psubset) finally have "DIM('a) - 1 < dim {x + y |x y. x \ S \ a \ y = 0}" . then have DD: "dim {x + y |x y. x \ S \ a \ y = 0} = DIM('a)" using antisym dim_subset_UNIV lowdim_subset_hyperplane not_le by fastforce have subs: "subspace {x + y| x y. x \ S \ a \ y = 0}" using subspace_sums [OF \subspace S\ subspace_hyperplane] by simp moreover have "span {x + y| x y. x \ S \ a \ y = 0} = UNIV" using DD dim_eq_full by blast ultimately show ?thesis by (simp add: subs) (metis (lifting) span_eq_iff subs) qed proposition\<^marker>\tag unimportant\ affine_hyperplane_sums_eq_UNIV: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {v. a \ v = b} \ {}" and "S - {v. a \ v = b} \ {}" shows "{x + y| x y. x \ S \ a \ y = b} = UNIV" proof (cases "a = 0") case True with assms show ?thesis by (auto simp: if_splits) next case False obtain c where "c \ S" and c: "a \ c = b" using assms by force with affine_diffs_subspace [OF \affine S\] have "subspace ((+) (- c) ` S)" by blast then have aff: "affine ((+) (- c) ` S)" by (simp add: subspace_imp_affine) have 0: "0 \ (+) (- c) ` S" by (simp add: \c \ S\) obtain d where "d \ S" and "a \ d \ b" and dc: "d-c \ (+) (- c) ` S" using assms by auto then have adc: "a \ (d - c) \ 0" by (simp add: c inner_diff_right) define U where "U \ {x + y |x y. x \ (+) (- c) ` S \ a \ y = 0}" have "u + v \ (+) (c+c) ` U" if "u \ S" "b = a \ v" for u v proof show "u + v = c + c + (u + v - c - c)" by (simp add: algebra_simps) have "\x y. u + v - c - c = x + y \ (\xa\S. x = xa - c) \ a \ y = 0" proof (intro exI conjI) show "u + v - c - c = (u-c) + (v-c)" "a \ (v - c) = 0" by (simp_all add: algebra_simps c that) qed (use that in auto) then show "u + v - c - c \ U" by (auto simp: U_def image_def) qed moreover have "\a \ v = 0; u \ S\ \ \x ya. v + (u + c) = x + ya \ x \ S \ a \ ya = b" for v u by (metis add.left_commute c inner_right_distrib pth_d) ultimately have "{x + y |x y. x \ S \ a \ y = b} = (+) (c+c) ` U" by (fastforce simp: algebra_simps U_def) also have "... = range ((+) (c + c))" by (simp only: U_def affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) also have "... = UNIV" by simp finally show ?thesis . qed lemma aff_dim_sums_Int_0: assumes "affine S" and "affine T" and "0 \ S" "0 \ T" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - have "0 \ {x + y |x y. x \ S \ y \ T}" using assms by force then have 0: "0 \ affine hull {x + y |x y. x \ S \ y \ T}" by (metis (lifting) hull_inc) have sub: "subspace S" "subspace T" using assms by (auto simp: subspace_affine) show ?thesis using dim_sums_Int [OF sub] by (simp add: aff_dim_zero assms 0 hull_inc) qed proposition aff_dim_sums_Int: assumes "affine S" and "affine T" and "S \ T \ {}" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - obtain a where a: "a \ S" "a \ T" using assms by force have aff: "affine ((+) (-a) ` S)" "affine ((+) (-a) ` T)" using affine_translation [symmetric, of "- a"] assms by (simp_all cong: image_cong_simp) have zero: "0 \ ((+) (-a) ` S)" "0 \ ((+) (-a) ` T)" using a assms by auto have "{x + y |x y. x \ (+) (- a) ` S \ y \ (+) (- a) ` T} = (+) (- 2 *\<^sub>R a) ` {x + y| x y. x \ S \ y \ T}" by (force simp: algebra_simps scaleR_2) moreover have "(+) (- a) ` S \ (+) (- a) ` T = (+) (- a) ` (S \ T)" by auto ultimately show ?thesis using aff_dim_sums_Int_0 [OF aff zero] aff_dim_translation_eq by (metis (lifting)) qed lemma aff_dim_affine_Int_hyperplane: fixes a :: "'a::euclidean_space" assumes "affine S" shows "aff_dim(S \ {x. a \ x = b}) = (if S \ {v. a \ v = b} = {} then - 1 else if S \ {v. a \ v = b} then aff_dim S else aff_dim S - 1)" proof (cases "a = 0") case True with assms show ?thesis by auto next case False then have "aff_dim (S \ {x. a \ x = b}) = aff_dim S - 1" if "x \ S" "a \ x \ b" and non: "S \ {v. a \ v = b} \ {}" for x proof - have [simp]: "{x + y| x y. x \ S \ a \ y = b} = UNIV" using affine_hyperplane_sums_eq_UNIV [OF assms non] that by blast show ?thesis using aff_dim_sums_Int [OF assms affine_hyperplane non] by (simp add: of_nat_diff False) qed then show ?thesis by (metis (mono_tags, lifting) inf.orderE aff_dim_empty_eq mem_Collect_eq subsetI) qed lemma aff_dim_lt_full: fixes S :: "'a::euclidean_space set" shows "aff_dim S < DIM('a) \ (affine hull S \ UNIV)" by (metis (no_types) aff_dim_affine_hull aff_dim_le_DIM aff_dim_UNIV affine_hull_UNIV less_le) lemma aff_dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "affine T" "S \ {}" shows "aff_dim S = aff_dim T" proof - show ?thesis proof (rule order_antisym) show "aff_dim S \ aff_dim T" by (blast intro: aff_dim_subset [OF openin_imp_subset] ope) next obtain a where "a \ S" using \S \ {}\ by blast have "S \ T" using ope openin_imp_subset by auto then have "a \ T" using \a \ S\ by auto then have subT': "subspace ((\x. - a + x) ` T)" using affine_diffs_subspace \affine T\ by auto then obtain B where Bsub: "B \ ((\x. - a + x) ` T)" and po: "pairwise orthogonal B" and eq1: "\x. x \ B \ norm x = 1" and "independent B" and cardB: "card B = dim ((\x. - a + x) ` T)" and spanB: "span B = ((\x. - a + x) ` T)" by (rule orthonormal_basis_subspace) auto obtain e where "0 < e" and e: "cball a e \ T \ S" by (meson \a \ S\ openin_contains_cball ope) have "aff_dim T = aff_dim ((\x. - a + x) ` T)" by (metis aff_dim_translation_eq) also have "... = dim ((\x. - a + x) ` T)" using aff_dim_subspace subT' by blast also have "... = card B" by (simp add: cardB) also have "... = card ((\x. e *\<^sub>R x) ` B)" using \0 < e\ by (force simp: inj_on_def card_image) also have "... \ dim ((\x. - a + x) ` S)" proof (simp, rule independent_card_le_dim) have e': "cball 0 e \ (\x. x - a) ` T \ (\x. x - a) ` S" using e by (auto simp: dist_norm norm_minus_commute subset_eq) have "(\x. e *\<^sub>R x) ` B \ cball 0 e \ (\x. x - a) ` T" using Bsub \0 < e\ eq1 subT' \a \ T\ by (auto simp: subspace_def) then show "(\x. e *\<^sub>R x) ` B \ (\x. x - a) ` S" using e' by blast have "inj_on ((*\<^sub>R) e) (span B)" using \0 < e\ inj_on_def by fastforce then show "independent ((\x. e *\<^sub>R x) ` B)" using linear_scale_self \independent B\ linear_dependent_inj_imageD by blast qed also have "... = aff_dim S" using \a \ S\ aff_dim_eq_dim hull_inc by (force cong: image_cong_simp) finally show "aff_dim T \ aff_dim S" . qed qed lemma dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "subspace T" "S \ {}" shows "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis ope dim_subset openin_subset topspace_euclidean_subtopology) next have "dim T = aff_dim S" using aff_dim_openin by (metis aff_dim_subspace \subspace T\ \S \ {}\ ope subspace_affine) also have "... \ dim S" by (metis aff_dim_subset aff_dim_subspace dim_span span_superset subspace_span) finally show "dim T \ dim S" by simp qed subsection\Lower-dimensional affine subsets are nowhere dense\ proposition dense_complement_subspace: fixes S :: "'a :: euclidean_space set" assumes dim_less: "dim T < dim S" and "subspace S" shows "closure(S - T) = S" proof - have "closure(S - U) = S" if "dim U < dim S" "U \ S" for U proof - have "span U \ span S" by (metis neq_iff psubsetI span_eq_dim span_mono that) then obtain a where "a \ 0" "a \ span S" and a: "\y. y \ span U \ orthogonal a y" using orthogonal_to_subspace_exists_gen by metis show ?thesis proof have "closed S" by (simp add: \subspace S\ closed_subspace) then show "closure (S - U) \ S" by (simp add: closure_minimal) show "S \ closure (S - U)" proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" show "\y\S - U. dist y x < e" proof (cases "x \ U") case True let ?y = "x + (e/2 / norm a) *\<^sub>R a" show ?thesis proof show "dist ?y x < e" using \0 < e\ by (simp add: dist_norm) next have "?y \ S" by (metis \a \ span S\ \x \ S\ assms(2) span_eq_iff subspace_add subspace_scale) moreover have "?y \ U" proof - have "e/2 / norm a \ 0" using \0 < e\ \a \ 0\ by auto then show ?thesis by (metis True \a \ 0\ a orthogonal_scaleR orthogonal_self real_vector.scale_eq_0_iff span_add_eq span_base) qed ultimately show "?y \ S - U" by blast qed next case False with \0 < e\ \x \ S\ show ?thesis by force qed qed qed qed moreover have "S - S \ T = S-T" by blast moreover have "dim (S \ T) < dim S" by (metis dim_less dim_subset inf.cobounded2 inf.orderE inf.strict_boundedE not_le) ultimately show ?thesis by force qed corollary\<^marker>\tag unimportant\ dense_complement_affine: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and "affine S" shows "closure(S - T) = S" proof (cases "S \ T = {}") case True then show ?thesis by (metis Diff_triv affine_hull_eq \affine S\ closure_same_affine_hull closure_subset hull_subset subset_antisym) next case False then obtain z where z: "z \ S \ T" by blast then have "subspace ((+) (- z) ` S)" by (meson IntD1 affine_diffs_subspace \affine S\) moreover have "int (dim ((+) (- z) ` T)) < int (dim ((+) (- z) ` S))" thm aff_dim_eq_dim using z less by (simp add: aff_dim_eq_dim_subtract [of z] hull_inc cong: image_cong_simp) ultimately have "closure(((+) (- z) ` S) - ((+) (- z) ` T)) = ((+) (- z) ` S)" by (simp add: dense_complement_subspace) then show ?thesis by (metis closure_translation translation_diff translation_invert) qed corollary\<^marker>\tag unimportant\ dense_complement_openin_affine_hull: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and ope: "openin (top_of_set (affine hull S)) S" shows "closure(S - T) = closure S" proof - have "affine hull S - T \ affine hull S" by blast then have "closure (S \ closure (affine hull S - T)) = closure (S \ (affine hull S - T))" by (rule closure_openin_Int_closure [OF ope]) then show ?thesis by (metis Int_Diff aff_dim_affine_hull affine_affine_hull dense_complement_affine hull_subset inf.orderE less) qed corollary\<^marker>\tag unimportant\ dense_complement_convex: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" shows "closure(S - T) = closure S" proof show "closure (S - T) \ closure S" by (simp add: closure_mono) have "closure (rel_interior S - T) = closure (rel_interior S)" by (simp add: assms dense_complement_openin_affine_hull openin_rel_interior rel_interior_aff_dim rel_interior_same_affine_hull) then show "closure S \ closure (S - T)" by (metis Diff_mono \convex S\ closure_mono convex_closure_rel_interior order_refl rel_interior_subset) qed corollary\<^marker>\tag unimportant\ dense_complement_convex_closed: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" "closed S" shows "closure(S - T) = S" by (simp add: assms dense_complement_convex) subsection\<^marker>\tag unimportant\\Parallel slices, etc\ text\ If we take a slice out of a set, we can do it perpendicularly, with the normal vector to the slice parallel to the affine hull.\ proposition\<^marker>\tag unimportant\ affine_parallel_slice: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {x. a \ x \ b} \ {}" and "\ (S \ {x. a \ x \ b})" obtains a' b' where "a' \ 0" "S \ {x. a' \ x \ b'} = S \ {x. a \ x \ b}" "S \ {x. a' \ x = b'} = S \ {x. a \ x = b}" "\w. w \ S \ (w + a') \ S" proof (cases "S \ {x. a \ x = b} = {}") case True then obtain u v where "u \ S" "v \ S" "a \ u \ b" "a \ v > b" using assms by (auto simp: not_le) define \ where "\ = u + ((b - a \ u) / (a \ v - a \ u)) *\<^sub>R (v - u)" have "\ \ S" by (simp add: \_def \u \ S\ \v \ S\ \affine S\ mem_affine_3_minus) moreover have "a \ \ = b" using \a \ u \ b\ \b < a \ v\ by (simp add: \_def algebra_simps) (simp add: field_simps) ultimately have False using True by force then show ?thesis .. next case False then obtain z where "z \ S" and z: "a \ z = b" using assms by auto with affine_diffs_subspace [OF \affine S\] have sub: "subspace ((+) (- z) ` S)" by blast then have aff: "affine ((+) (- z) ` S)" and span: "span ((+) (- z) ` S) = ((+) (- z) ` S)" by (auto simp: subspace_imp_affine) obtain a' a'' where a': "a' \ span ((+) (- z) ` S)" and a: "a = a' + a''" and "\w. w \ span ((+) (- z) ` S) \ orthogonal a'' w" using orthogonal_subspace_decomp_exists [of "(+) (- z) ` S" "a"] by metis then have "\w. w \ S \ a'' \ (w-z) = 0" by (simp add: span_base orthogonal_def) then have a'': "\w. w \ S \ a'' \ w = (a - a') \ z" by (simp add: a inner_diff_right) then have ba'': "\w. w \ S \ a'' \ w = b - a' \ z" by (simp add: inner_diff_left z) show ?thesis proof (cases "a' = 0") case True with a assms True a'' diff_zero less_irrefl show ?thesis by auto next case False show ?thesis proof show "S \ {x. a' \ x \ a' \ z} = S \ {x. a \ x \ b}" "S \ {x. a' \ x = a' \ z} = S \ {x. a \ x = b}" by (auto simp: a ba'' inner_left_distrib) have "\w. w \ (+) (- z) ` S \ (w + a') \ (+) (- z) ` S" by (metis subspace_add a' span_eq_iff sub) then show "\w. w \ S \ (w + a') \ S" by fastforce qed (use False in auto) qed qed lemma diffs_affine_hull_span: assumes "a \ S" shows "(\x. x - a) ` (affine hull S) = span ((\x. x - a) ` S)" proof - have *: "((\x. x - a) ` (S - {a})) = ((\x. x - a) ` S) - {0}" by (auto simp: algebra_simps) show ?thesis by (auto simp add: algebra_simps affine_hull_span2 [OF assms] *) qed lemma aff_dim_dim_affine_diffs: fixes S :: "'a :: euclidean_space set" assumes "affine S" "a \ S" shows "aff_dim S = dim ((\x. x - a) ` S)" proof - obtain B where aff: "affine hull B = affine hull S" and ind: "\ affine_dependent B" and card: "of_nat (card B) = aff_dim S + 1" using aff_dim_basis_exists by blast then have "B \ {}" using assms by (metis affine_hull_eq_empty ex_in_conv) then obtain c where "c \ B" by auto then have "c \ S" by (metis aff affine_hull_eq \affine S\ hull_inc) have xy: "x - c = y - a \ y = x + 1 *\<^sub>R (a - c)" for x y c and a::'a by (auto simp: algebra_simps) have *: "(\x. x - c) ` S = (\x. x - a) ` S" using assms \c \ S\ by (auto simp: image_iff xy; metis mem_affine_3_minus pth_1) have affS: "affine hull S = S" by (simp add: \affine S\) have "aff_dim S = of_nat (card B) - 1" using card by simp also have "... = dim ((\x. x - c) ` B)" using affine_independent_card_dim_diffs [OF ind \c \ B\] by (simp add: affine_independent_card_dim_diffs [OF ind \c \ B\]) also have "... = dim ((\x. x - c) ` (affine hull B))" by (simp add: diffs_affine_hull_span \c \ B\) also have "... = dim ((\x. x - a) ` S)" by (simp add: affS aff *) finally show ?thesis . qed lemma aff_dim_linear_image_le: assumes "linear f" shows "aff_dim(f ` S) \ aff_dim S" proof - have "aff_dim (f ` T) \ aff_dim T" if "affine T" for T proof (cases "T = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False then obtain a where "a \ T" by auto have 1: "((\x. x - f a) ` f ` T) = {x - f a |x. x \ f ` T}" by auto have 2: "{x - f a| x. x \ f ` T} = f ` ((\x. x - a) ` T)" by (force simp: linear_diff [OF assms]) have "aff_dim (f ` T) = int (dim {x - f a |x. x \ f ` T})" by (simp add: \a \ T\ hull_inc aff_dim_eq_dim [of "f a"] 1 cong: image_cong_simp) also have "... = int (dim (f ` ((\x. x - a) ` T)))" by (force simp: linear_diff [OF assms] 2) also have "... \ int (dim ((\x. x - a) ` T))" by (simp add: dim_image_le [OF assms]) also have "... \ aff_dim T" by (simp add: aff_dim_dim_affine_diffs [symmetric] \a \ T\ \affine T\) finally show ?thesis . qed then have "aff_dim (f ` (affine hull S)) \ aff_dim (affine hull S)" using affine_affine_hull [of S] by blast then show ?thesis using affine_hull_linear_image assms linear_conv_bounded_linear by fastforce qed lemma aff_dim_injective_linear_image [simp]: assumes "linear f" "inj f" shows "aff_dim (f ` S) = aff_dim S" proof (rule antisym) show "aff_dim (f ` S) \ aff_dim S" by (simp add: aff_dim_linear_image_le assms(1)) next obtain g where "linear g" "g \ f = id" using assms(1) assms(2) linear_injective_left_inverse by blast then have "aff_dim S \ aff_dim(g ` f ` S)" by (simp add: image_comp) also have "... \ aff_dim (f ` S)" by (simp add: \linear g\ aff_dim_linear_image_le) finally show "aff_dim S \ aff_dim (f ` S)" . qed lemma choose_affine_subset: assumes "affine S" "-1 \ d" and dle: "d \ aff_dim S" obtains T where "affine T" "T \ S" "aff_dim T = d" proof (cases "d = -1 \ S={}") case True with assms show ?thesis by (metis aff_dim_empty affine_empty bot.extremum that eq_iff) next case False with assms obtain a where "a \ S" "0 \ d" by auto with assms have ss: "subspace ((+) (- a) ` S)" by (simp add: affine_diffs_subspace_subtract cong: image_cong_simp) have "nat d \ dim ((+) (- a) ` S)" by (metis aff_dim_subspace aff_dim_translation_eq dle nat_int nat_mono ss) then obtain T where "subspace T" and Tsb: "T \ span ((+) (- a) ` S)" and Tdim: "dim T = nat d" using choose_subspace_of_subspace [of "nat d" "(+) (- a) ` S"] by blast then have "affine T" using subspace_affine by blast then have "affine ((+) a ` T)" by (metis affine_hull_eq affine_hull_translation) moreover have "(+) a ` T \ S" proof - have "T \ (+) (- a) ` S" by (metis (no_types) span_eq_iff Tsb ss) then show "(+) a ` T \ S" using add_ac by auto qed moreover have "aff_dim ((+) a ` T) = d" by (simp add: aff_dim_subspace Tdim \0 \ d\ \subspace T\ aff_dim_translation_eq) ultimately show ?thesis by (rule that) qed subsection\Paracompactness\ proposition paracompact: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "S \ \\" and opC: "\T. T \ \ \ open T" obtains \' where "S \ \ \'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. x \ S \ \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" proof (cases "S = {}") case True with that show ?thesis by blast next case False have "\T U. x \ U \ open U \ closure U \ T \ T \ \" if "x \ S" for x proof - obtain T where "x \ T" "T \ \" "open T" using assms \x \ S\ by blast then obtain e where "e > 0" "cball x e \ T" by (force simp: open_contains_cball) then show ?thesis by (meson open_ball \T \ \\ ball_subset_cball centre_in_ball closed_cball closure_minimal dual_order.trans) qed then obtain F G where Gin: "x \ G x" and oG: "open (G x)" and clos: "closure (G x) \ F x" and Fin: "F x \ \" if "x \ S" for x by metis then obtain \ where "\ \ G ` S" "countable \" "\\ = \(G ` S)" using Lindelof [of "G ` S"] by (metis image_iff) then obtain K where K: "K \ S" "countable K" and eq: "\(G ` K) = \(G ` S)" by (metis countable_subset_image) with False Gin have "K \ {}" by force then obtain a :: "nat \ 'a" where "range a = K" by (metis range_from_nat_into \countable K\) then have odif: "\n. open (F (a n) - \{closure (G (a m)) |m. m < n})" using \K \ S\ Fin opC by (fastforce simp add:) let ?C = "range (\n. F(a n) - \{closure(G(a m)) |m. m < n})" have enum_S: "\n. x \ F(a n) \ x \ G(a n)" if "x \ S" for x proof - have "\y \ K. x \ G y" using eq that Gin by fastforce then show ?thesis using clos K \range a = K\ closure_subset by blast qed show ?thesis proof show "S \ Union ?C" proof fix x assume "x \ S" define n where "n \ LEAST n. x \ F(a n)" have n: "x \ F(a n)" using enum_S [OF \x \ S\] by (force simp: n_def intro: LeastI) have notn: "x \ F(a m)" if "m < n" for m using that not_less_Least by (force simp: n_def) then have "x \ \{closure (G (a m)) |m. m < n}" using n \K \ S\ \range a = K\ clos notn by fastforce with n show "x \ Union ?C" by blast qed show "\U. U \ ?C \ open U \ (\T. T \ \ \ U \ T)" using Fin \K \ S\ \range a = K\ by (auto simp: odif) show "\V. open V \ x \ V \ finite {U. U \ ?C \ (U \ V \ {})}" if "x \ S" for x proof - obtain n where n: "x \ F(a n)" "x \ G(a n)" using \x \ S\ enum_S by auto have "{U \ ?C. U \ G (a n) \ {}} \ (\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n" proof clarsimp fix k assume "(F (a k) - \{closure (G (a m)) |m. m < k}) \ G (a n) \ {}" then have "k \ n" by auto (metis closure_subset not_le subsetCE) then show "F (a k) - \{closure (G (a m)) |m. m < k} \ (\n. F (a n) - \{closure (G (a m)) |m. m < n}) ` {..n}" by force qed moreover have "finite ((\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n)" by force ultimately have *: "finite {U \ ?C. U \ G (a n) \ {}}" using finite_subset by blast have "a n \ S" using \K \ S\ \range a = K\ by blast then show ?thesis by (blast intro: oG n *) qed qed qed corollary paracompact_closedin: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes cin: "closedin (top_of_set U) S" and oin: "\T. T \ \ \ openin (top_of_set U) T" and "S \ \\" obtains \' where "S \ \ \'" and "\V. V \ \' \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" and "\x. x \ U \ \V. openin (top_of_set U) V \ x \ V \ finite {X. X \ \' \ (X \ V \ {})}" proof - have "\Z. open Z \ (T = U \ Z)" if "T \ \" for T using oin [OF that] by (auto simp: openin_open) then obtain F where opF: "open (F T)" and intF: "U \ F T = T" if "T \ \" for T by metis obtain K where K: "closed K" "U \ K = S" using cin by (auto simp: closedin_closed) have 1: "U \ \(insert (- K) (F ` \))" by clarsimp (metis Int_iff Union_iff \U \ K = S\ \S \ \\\ subsetD intF) have 2: "\T. T \ insert (- K) (F ` \) \ open T" using \closed K\ by (auto simp: opF) obtain \ where "U \ \\" and D1: "\U. U \ \ \ open U \ (\T. T \ insert (- K) (F ` \) \ U \ T)" and D2: "\x. x \ U \ \V. open V \ x \ V \ finite {U \ \. U \ V \ {}}" by (blast intro: paracompact [OF 1 2]) let ?C = "{U \ V |V. V \ \ \ (V \ K \ {})}" show ?thesis proof (rule_tac \' = "{U \ V |V. V \ \ \ (V \ K \ {})}" in that) show "S \ \?C" using \U \ K = S\ \U \ \\\ K by (blast dest!: subsetD) show "\V. V \ ?C \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" using D1 intF by fastforce have *: "{X. (\V. X = U \ V \ V \ \ \ V \ K \ {}) \ X \ (U \ V) \ {}} \ (\x. U \ x) ` {U \ \. U \ V \ {}}" for V by blast show "\V. openin (top_of_set U) V \ x \ V \ finite {X \ ?C. X \ V \ {}}" if "x \ U" for x proof - from D2 [OF that] obtain V where "open V" "x \ V" "finite {U \ \. U \ V \ {}}" by auto with * show ?thesis by (rule_tac x="U \ V" in exI) (auto intro: that finite_subset [OF *]) qed qed qed corollary\<^marker>\tag unimportant\ paracompact_closed: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "closed S" and opC: "\T. T \ \ \ open T" and "S \ \\" obtains \' where "S \ \\'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" by (rule paracompact_closedin [of UNIV S \]) (auto simp: assms) subsection\<^marker>\tag unimportant\\Closed-graph characterization of continuity\ lemma continuous_closed_graph_gen: fixes T :: "'b::real_normed_vector set" assumes contf: "continuous_on S f" and fim: "f ` S \ T" shows "closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" proof - have eq: "((\x. Pair x (f x)) ` S) = (S \ T \ (\z. (f \ fst)z - snd z) -` {0})" using fim by auto show ?thesis unfolding eq by (intro continuous_intros continuous_closedin_preimage continuous_on_subset [OF contf]) auto qed lemma continuous_closed_graph_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" shows "continuous_on S f \ closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" (is "?lhs = ?rhs") proof - have "?lhs" if ?rhs proof (clarsimp simp add: continuous_on_closed_gen [OF fim]) fix U assume U: "closedin (top_of_set T) U" have eq: "(S \ f -` U) = fst ` (((\x. Pair x (f x)) ` S) \ (S \ U))" by (force simp: image_iff) show "closedin (top_of_set S) (S \ f -` U)" by (simp add: U closedin_Int closedin_Times closed_map_fst [OF \compact T\] that eq) qed with continuous_closed_graph_gen assms show ?thesis by blast qed lemma continuous_closed_graph: fixes f :: "'a::topological_space \ 'b::real_normed_vector" assumes "closed S" and contf: "continuous_on S f" shows "closed ((\x. Pair x (f x)) ` S)" proof (rule closedin_closed_trans) show "closedin (top_of_set (S \ UNIV)) ((\x. (x, f x)) ` S)" by (rule continuous_closed_graph_gen [OF contf subset_UNIV]) qed (simp add: \closed S\ closed_Times) lemma continuous_from_closed_graph: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" and clo: "closed ((\x. Pair x (f x)) ` S)" shows "continuous_on S f" using fim clo by (auto intro: closed_subset simp: continuous_closed_graph_eq [OF \compact T\ fim]) lemma continuous_on_Un_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T f" shows "continuous_on (S \ T) f" using pasting_lemma [of "{S,T}" "top_of_set (S \ T)" id euclidean "\i. f" f] contf contg opS opT by (simp add: subtopology_subtopology) (metis inf.absorb2 openin_imp_subset) lemma continuous_on_cases_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T g" and fg: "\x. x \ S \ \P x \ x \ T \ P x \ f x = g x" shows "continuous_on (S \ T) (\x. if P x then f x else g x)" proof - have "\x. x \ S \ (if P x then f x else g x) = f x" "\x. x \ T \ (if P x then f x else g x) = g x" by (simp_all add: fg) then have "continuous_on S (\x. if P x then f x else g x)" "continuous_on T (\x. if P x then f x else g x)" by (simp_all add: contf contg cong: continuous_on_cong) then show ?thesis by (rule continuous_on_Un_local_open [OF opS opT]) qed subsection\<^marker>\tag unimportant\\The union of two collinear segments is another segment\ proposition\<^marker>\tag unimportant\ in_convex_hull_exchange: fixes a :: "'a::euclidean_space" assumes a: "a \ convex hull S" and xS: "x \ convex hull S" obtains b where "b \ S" "x \ convex hull (insert a (S - {b}))" proof (cases "a \ S") case True with xS insert_Diff that show ?thesis by fastforce next case False show ?thesis proof (cases "finite S \ card S \ Suc (DIM('a))") case True then obtain u where u0: "\i. i \ S \ 0 \ u i" and u1: "sum u S = 1" and ua: "(\i\S. u i *\<^sub>R i) = a" using a by (auto simp: convex_hull_finite) obtain v where v0: "\i. i \ S \ 0 \ v i" and v1: "sum v S = 1" and vx: "(\i\S. v i *\<^sub>R i) = x" using True xS by (auto simp: convex_hull_finite) show ?thesis proof (cases "\b. b \ S \ v b = 0") case True then obtain b where b: "b \ S" "v b = 0" by blast show ?thesis proof have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x = a then 0 else v x) = (\x \ S - {b}. if x = a then 0 else v x)" using fin by (force intro: sum.mono_neutral_right) also have "... = (\x \ S - {b}. v x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x)" by (metis \v b = 0\ diff_zero sum.infinite sum_diff1 u1 zero_neq_one) finally show "(\x\insert a (S - {b}). if x = a then 0 else v x) = 1" by (simp add: v1) show "\x. x \ insert a (S - {b}) \ 0 \ (if x = a then 0 else v x)" by (auto simp: v0) have "(\x \ insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = (\x \ S - {b}. (if x = a then 0 else v x) *\<^sub>R x)" using fin by (force intro: sum.mono_neutral_right) also have "... = (\x \ S - {b}. v x *\<^sub>R x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x *\<^sub>R x)" by (metis (no_types, lifting) b(2) diff_zero fin finite.emptyI finite_Diff2 finite_insert scale_eq_0_iff sum_diff1) finally show "(\x\insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) next case False have le_Max: "u i / v i \ Max ((\i. u i / v i) ` S)" if "i \ S" for i by (simp add: True that) have "Max ((\i. u i / v i) ` S) \ (\i. u i / v i) ` S" using True v1 by (auto intro: Max_in) then obtain b where "b \ S" and beq: "Max ((\b. u b / v b) ` S) = u b / v b" by blast then have "0 \ u b / v b" using le_Max beq divide_le_0_iff le_numeral_extra(2) sum_nonpos u1 by (metis False eq_iff v0) then have "0 < u b" "0 < v b" using False \b \ S\ u0 v0 by force+ have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show ?thesis proof show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x=a then v b / u b else v x - (v b / u b) * u x) = v b / u b + (\x \ S - {b}. v x - (v b / u b) * u x)" using \a \ S\ \b \ S\ True by (auto intro!: sum.cong split: if_split_asm) also have "... = v b / u b + (\x \ S - {b}. v x) - (v b / u b) * (\x \ S - {b}. u x)" by (simp add: Groups_Big.sum_subtractf sum_distrib_left) also have "... = (\x\S. v x)" using \0 < u b\ True by (simp add: Groups_Big.sum_diff1 u1 field_simps) finally show "sum (\x. if x=a then v b / u b else v x - (v b / u b) * u x) (insert a (S - {b})) = 1" by (simp add: v1) show "0 \ (if i = a then v b / u b else v i - v b / u b * u i)" if "i \ insert a (S - {b})" for i using \0 < u b\ \0 < v b\ v0 [of i] le_Max [of i] beq that False by (auto simp: field_simps split: if_split_asm) have "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = (v b / u b) *\<^sub>R a + (\x\S - {b}. (v x - v b / u b * u x) *\<^sub>R x)" using \a \ S\ \b \ S\ True by (auto intro!: sum.cong split: if_split_asm) also have "... = (v b / u b) *\<^sub>R a + (\x \ S - {b}. v x *\<^sub>R x) - (v b / u b) *\<^sub>R (\x \ S - {b}. u x *\<^sub>R x)" by (simp add: Groups_Big.sum_subtractf scaleR_left_diff_distrib sum_distrib_left scale_sum_right) also have "... = (\x\S. v x *\<^sub>R x)" using \0 < u b\ True by (simp add: ua vx Groups_Big.sum_diff1 algebra_simps) finally show "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) qed next case False obtain T where "finite T" "T \ S" and caT: "card T \ Suc (DIM('a))" and xT: "x \ convex hull T" using xS by (auto simp: caratheodory [of S]) with False obtain b where b: "b \ S" "b \ T" by (metis antisym subsetI) show ?thesis proof show "x \ convex hull insert a (S - {b})" using \T \ S\ b by (blast intro: subsetD [OF hull_mono xT]) qed (rule \b \ S\) qed qed lemma convex_hull_exchange_Union: fixes a :: "'a::euclidean_space" assumes "a \ convex hull S" shows "convex hull S = (\b \ S. convex hull (insert a (S - {b})))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (blast intro: in_convex_hull_exchange [OF assms]) show "?rhs \ ?lhs" proof clarify fix x b assume"b \ S" "x \ convex hull insert a (S - {b})" then show "x \ convex hull S" if "b \ S" by (metis (no_types) that assms order_refl hull_mono hull_redundant insert_Diff_single insert_subset subsetCE) qed qed lemma Un_closed_segment: fixes a :: "'a::euclidean_space" assumes "b \ closed_segment a c" shows "closed_segment a b \ closed_segment b c = closed_segment a c" proof (cases "c = a") case True with assms show ?thesis by simp next case False with assms have "convex hull {a, b} \ convex hull {b, c} = (\ba\{a, c}. convex hull insert b ({a, c} - {ba}))" by (auto simp: insert_Diff_if insert_commute) then show ?thesis using convex_hull_exchange_Union by (metis assms segment_convex_hull) qed lemma Un_open_segment: fixes a :: "'a::euclidean_space" assumes "b \ open_segment a c" shows "open_segment a b \ {b} \ open_segment b c = open_segment a c" (is "?lhs = ?rhs") proof - have b: "b \ closed_segment a c" by (simp add: assms open_closed_segment) have *: "?rhs \ insert b (open_segment a b \ open_segment b c)" if "{b,c,a} \ open_segment a b \ open_segment b c = {c,a} \ ?rhs" proof - have "insert a (insert c (insert b (open_segment a b \ open_segment b c))) = insert a (insert c (?rhs))" using that by (simp add: insert_commute) then show ?thesis by (metis (no_types) Diff_cancel Diff_eq_empty_iff Diff_insert2 open_segment_def) qed show ?thesis proof show "?lhs \ ?rhs" by (simp add: assms b subset_open_segment) show "?rhs \ ?lhs" using Un_closed_segment [OF b] * by (simp add: closed_segment_eq_open insert_commute) qed qed subsection\Covering an open set by a countable chain of compact sets\ proposition open_Union_compact_subsets: fixes S :: "'a::euclidean_space set" assumes "open S" obtains C where "\n. compact(C n)" "\n. C n \ S" "\n. C n \ interior(C(Suc n))" "\(range C) = S" "\K. \compact K; K \ S\ \ \N. \n\N. K \ (C n)" proof (cases "S = {}") case True then show ?thesis by (rule_tac C = "\n. {}" in that) auto next case False then obtain a where "a \ S" by auto let ?C = "\n. cball a (real n) - (\x \ -S. \e \ ball 0 (1 / real(Suc n)). {x + e})" have "\N. \n\N. K \ (f n)" if "\n. compact(f n)" and sub_int: "\n. f n \ interior (f(Suc n))" and eq: "\(range f) = S" and "compact K" "K \ S" for f K proof - have *: "\n. f n \ (\n. interior (f n))" by (meson Sup_upper2 UNIV_I \\n. f n \ interior (f (Suc n))\ image_iff) have mono: "\m n. m \ n \f m \ f n" by (meson dual_order.trans interior_subset lift_Suc_mono_le sub_int) obtain I where "finite I" and I: "K \ (\i\I. interior (f i))" proof (rule compactE_image [OF \compact K\]) show "K \ (\n. interior (f n))" using \K \ S\ \\(f ` UNIV) = S\ * by blast qed auto { fix n assume n: "Max I \ n" have "(\i\I. interior (f i)) \ f n" by (rule UN_least) (meson dual_order.trans interior_subset mono I Max_ge [OF \finite I\] n) then have "K \ f n" using I by auto } then show ?thesis by blast qed moreover have "\f. (\n. compact(f n)) \ (\n. (f n) \ S) \ (\n. (f n) \ interior(f(Suc n))) \ ((\(range f) = S))" proof (intro exI conjI allI) show "\n. compact (?C n)" by (auto simp: compact_diff open_sums) show "\n. ?C n \ S" by auto show "?C n \ interior (?C (Suc n))" for n proof (simp add: interior_diff, rule Diff_mono) show "cball a (real n) \ ball a (1 + real n)" by (simp add: cball_subset_ball_iff) have cl: "closed (\x\- S. \e\cball 0 (1 / (2 + real n)). {x + e})" using assms by (auto intro: closed_compact_sums) have "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \e \ cball 0 (1 / (2 + real n)). {x + e})" by (intro closure_minimal UN_mono ball_subset_cball order_refl cl) also have "... \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" by (simp add: cball_subset_ball_iff field_split_simps UN_mono) finally show "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" . qed have "S \ \ (range ?C)" proof fix x assume x: "x \ S" then obtain e where "e > 0" and e: "ball x e \ S" using assms open_contains_ball by blast then obtain N1 where "N1 > 0" and N1: "real N1 > 1/e" using reals_Archimedean2 by (metis divide_less_0_iff less_eq_real_def neq0_conv not_le of_nat_0 of_nat_1 of_nat_less_0_iff) obtain N2 where N2: "norm(x - a) \ real N2" by (meson real_arch_simple) have N12: "inverse((N1 + N2) + 1) \ inverse(N1)" using \N1 > 0\ by (auto simp: field_split_simps) have "x \ y + z" if "y \ S" "norm z < 1 / (1 + (real N1 + real N2))" for y z proof - have "e * real N1 < e * (1 + (real N1 + real N2))" by (simp add: \0 < e\) then have "1 / (1 + (real N1 + real N2)) < e" using N1 \e > 0\ by (metis divide_less_eq less_trans mult.commute of_nat_add of_nat_less_0_iff of_nat_Suc) then have "x - z \ ball x e" using that by simp then have "x - z \ S" using e by blast with that show ?thesis by auto qed with N2 show "x \ \ (range ?C)" by (rule_tac a = "N1+N2" in UN_I) (auto simp: dist_norm norm_minus_commute) qed then show "\ (range ?C) = S" by auto qed ultimately show ?thesis using that by metis qed subsection\Orthogonal complement\ definition\<^marker>\tag important\ orthogonal_comp ("_\<^sup>\" [80] 80) where "orthogonal_comp W \ {x. \y \ W. orthogonal y x}" proposition subspace_orthogonal_comp: "subspace (W\<^sup>\)" unfolding subspace_def orthogonal_comp_def orthogonal_def by (auto simp: inner_right_distrib) lemma orthogonal_comp_anti_mono: assumes "A \ B" shows "B\<^sup>\ \ A\<^sup>\" proof fix x assume x: "x \ B\<^sup>\" show "x \ orthogonal_comp A" using x unfolding orthogonal_comp_def by (simp add: orthogonal_def, metis assms in_mono) qed lemma orthogonal_comp_null [simp]: "{0}\<^sup>\ = UNIV" by (auto simp: orthogonal_comp_def orthogonal_def) lemma orthogonal_comp_UNIV [simp]: "UNIV\<^sup>\ = {0}" unfolding orthogonal_comp_def orthogonal_def by auto (use inner_eq_zero_iff in blast) lemma orthogonal_comp_subset: "U \ U\<^sup>\\<^sup>\" by (auto simp: orthogonal_comp_def orthogonal_def inner_commute) lemma subspace_sum_minimal: assumes "S \ U" "T \ U" "subspace U" shows "S + T \ U" proof fix x assume "x \ S + T" then obtain xs xt where "xs \ S" "xt \ T" "x = xs+xt" by (meson set_plus_elim) then show "x \ U" by (meson assms subsetCE subspace_add) qed proposition subspace_sum_orthogonal_comp: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U + U\<^sup>\ = UNIV" proof - obtain B where "B \ U" and ortho: "pairwise orthogonal B" "\x. x \ B \ norm x = 1" and "independent B" "card B = dim U" "span B = U" using orthonormal_basis_subspace [OF assms] by metis then have "finite B" by (simp add: indep_card_eq_dim_span) have *: "\x\B. \y\B. x \ y = (if x=y then 1 else 0)" using ortho norm_eq_1 by (auto simp: orthogonal_def pairwise_def) { fix v let ?u = "\b\B. (v \ b) *\<^sub>R b" have "v = ?u + (v - ?u)" by simp moreover have "?u \ U" by (metis (no_types, lifting) \span B = U\ assms subspace_sum span_base span_mul) moreover have "(v - ?u) \ U\<^sup>\" proof (clarsimp simp: orthogonal_comp_def orthogonal_def) fix y assume "y \ U" with \span B = U\ span_finite [OF \finite B\] obtain u where u: "y = (\b\B. u b *\<^sub>R b)" by auto have "b \ (v - ?u) = 0" if "b \ B" for b using that \finite B\ by (simp add: * algebra_simps inner_sum_right if_distrib [of "(*)v" for v] inner_commute cong: if_cong) then show "y \ (v - ?u) = 0" by (simp add: u inner_sum_left) qed ultimately have "v \ U + U\<^sup>\" using set_plus_intro by fastforce } then show ?thesis by auto qed lemma orthogonal_Int_0: assumes "subspace U" shows "U \ U\<^sup>\ = {0}" using orthogonal_comp_def orthogonal_self by (force simp: assms subspace_0 subspace_orthogonal_comp) lemma orthogonal_comp_self: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U\<^sup>\\<^sup>\ = U" proof have ssU': "subspace (U\<^sup>\)" by (simp add: subspace_orthogonal_comp) have "u \ U" if "u \ U\<^sup>\\<^sup>\" for u proof - obtain v w where "u = v+w" "v \ U" "w \ U\<^sup>\" using subspace_sum_orthogonal_comp [OF assms] set_plus_elim by blast then have "u-v \ U\<^sup>\" by simp moreover have "v \ U\<^sup>\\<^sup>\" using \v \ U\ orthogonal_comp_subset by blast then have "u-v \ U\<^sup>\\<^sup>\" by (simp add: subspace_diff subspace_orthogonal_comp that) ultimately have "u-v = 0" using orthogonal_Int_0 ssU' by blast with \v \ U\ show ?thesis by auto qed then show "U\<^sup>\\<^sup>\ \ U" by auto qed (use orthogonal_comp_subset in auto) lemma ker_orthogonal_comp_adjoint: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "f -` {0} = (range (adjoint f))\<^sup>\" proof - have "\x. \\y. y \ f x = 0\ \ f x = 0" using assms inner_commute all_zero_iff by metis then show ?thesis using assms by (auto simp: orthogonal_comp_def orthogonal_def adjoint_works inner_commute) qed subsection\<^marker>\tag unimportant\ \A non-injective linear function maps into a hyperplane.\ lemma linear_surj_adj_imp_inj: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" "surj (adjoint f)" shows "inj f" proof - have "\x. y = adjoint f x" for y using assms by (simp add: surjD) then show "inj f" using assms unfolding inj_on_def image_def by (metis (no_types) adjoint_works euclidean_eqI) qed \ \\<^url>\https://mathonline.wikidot.com/injectivity-and-surjectivity-of-the-adjoint-of-a-linear-map\\ lemma surj_adjoint_iff_inj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "surj (adjoint f) \ inj f" proof assume "surj (adjoint f)" then show "inj f" by (simp add: assms linear_surj_adj_imp_inj) next assume "inj f" have "f -` {0} = {0}" using assms \inj f\ linear_0 linear_injective_0 by fastforce moreover have "f -` {0} = range (adjoint f)\<^sup>\" by (intro ker_orthogonal_comp_adjoint assms) ultimately have "range (adjoint f)\<^sup>\\<^sup>\ = UNIV" by (metis orthogonal_comp_null) then show "surj (adjoint f)" using adjoint_linear \linear f\ by (subst (asm) orthogonal_comp_self) (simp add: adjoint_linear linear_subspace_image) qed lemma inj_adjoint_iff_surj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "inj (adjoint f) \ surj f" proof assume "inj (adjoint f)" have "(adjoint f) -` {0} = {0}" by (metis \inj (adjoint f)\ adjoint_linear assms surj_adjoint_iff_inj ker_orthogonal_comp_adjoint orthogonal_comp_UNIV) then have "(range(f))\<^sup>\ = {0}" by (metis (no_types, opaque_lifting) adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint set_zero) then show "surj f" by (metis \inj (adjoint f)\ adjoint_adjoint adjoint_linear assms surj_adjoint_iff_inj) next assume "surj f" then have "range f = (adjoint f -` {0})\<^sup>\" by (simp add: adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint) then have "{0} = adjoint f -` {0}" using \surj f\ adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint by force then show "inj (adjoint f)" by (simp add: \surj f\ adjoint_adjoint adjoint_linear assms linear_surj_adj_imp_inj) qed lemma linear_singular_into_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" shows "\ inj f \ (\a. a \ 0 \ (\x. a \ f x = 0))" (is "_ = ?rhs") proof assume "\inj f" then show ?rhs using all_zero_iff by (metis (no_types, opaque_lifting) adjoint_clauses(2) adjoint_linear assms linear_injective_0 linear_injective_imp_surjective linear_surj_adj_imp_inj) next assume ?rhs then show "\inj f" by (metis assms linear_injective_isomorphism all_zero_iff) qed lemma linear_singular_image_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" "\inj f" obtains a where "a \ 0" "\S. f ` S \ {x. a \ x = 0}" using assms by (fastforce simp add: linear_singular_into_hyperplane) end diff --git a/src/HOL/Analysis/Sum_Topology.thy b/src/HOL/Analysis/Sum_Topology.thy --- a/src/HOL/Analysis/Sum_Topology.thy +++ b/src/HOL/Analysis/Sum_Topology.thy @@ -1,146 +1,162 @@ section\Disjoint sum of arbitarily many spaces\ theory Sum_Topology imports Abstract_Topology begin definition sum_topology :: "('a \ 'b topology) \ 'a set \ ('a \ 'b) topology" where "sum_topology X I \ topology (\U. U \ Sigma I (topspace \ X) \ (\i \ I. openin (X i) {x. (i,x) \ U}))" lemma is_sum_topology: "istopology (\U. U \ Sigma I (topspace \ X) \ (\i\I. openin (X i) {x. (i, x) \ U}))" proof - have 1: "{x. (i, x) \ S \ T} = {x. (i, x) \ S} \ {x::'b. (i, x) \ T}" for S T and i::'a by auto have 2: "{x. (i, x) \ \ \} = (\K\\. {x::'b. (i, x) \ K})" for \ and i::'a by auto show ?thesis unfolding istopology_def 1 2 by blast qed lemma openin_sum_topology: "openin (sum_topology X I) U \ U \ Sigma I (topspace \ X) \ (\i \ I. openin (X i) {x. (i,x) \ U})" by (auto simp: sum_topology_def is_sum_topology) lemma openin_disjoint_union: "openin (sum_topology X I) (Sigma I U) \ (\i \ I. openin (X i) (U i))" using openin_subset by (force simp: openin_sum_topology) lemma topspace_sum_topology [simp]: "topspace(sum_topology X I) = Sigma I (topspace \ X)" by (metis comp_apply openin_disjoint_union openin_subset openin_sum_topology openin_topspace subset_antisym) lemma openin_sum_topology_alt: "openin (sum_topology X I) U \ (\T. U = Sigma I T \ (\i \ I. openin (X i) (T i)))" by (bestsimp simp: openin_sum_topology dest: openin_subset) lemma forall_openin_sum_topology: "(\U. openin (sum_topology X I) U \ P U) \ (\T. (\i \ I. openin (X i) (T i)) \ P(Sigma I T))" by (auto simp: openin_sum_topology_alt) lemma exists_openin_sum_topology: "(\U. openin (sum_topology X I) U \ P U) \ (\T. (\i \ I. openin (X i) (T i)) \ P(Sigma I T))" by (auto simp: openin_sum_topology_alt) lemma closedin_sum_topology: "closedin (sum_topology X I) U \ U \ Sigma I (topspace \ X) \ (\i \ I. closedin (X i) {x. (i,x) \ U})" (is "?lhs = ?rhs") proof assume L: ?lhs then have U: "U \ Sigma I (topspace \ X)" using closedin_subset by fastforce then have "\i\I. {x. (i, x) \ U} \ topspace (X i)" by fastforce moreover have "openin (X i) (topspace (X i) - {x. (i, x) \ U})" if "i\I" for i apply (subst openin_subopen) using L that unfolding closedin_def openin_sum_topology by (bestsimp simp: o_def subset_iff) ultimately show ?rhs by (simp add: U closedin_def) next assume R: ?rhs then have "openin (X i) {x. (i, x) \ topspace (sum_topology X I) - U}" if "i\I" for i apply (subst openin_subopen) using that unfolding closedin_def openin_sum_topology by (bestsimp simp: o_def subset_iff) then show ?lhs by (simp add: R closedin_def openin_sum_topology) qed lemma closedin_disjoint_union: "closedin (sum_topology X I) (Sigma I U) \ (\i \ I. closedin (X i) (U i))" using closedin_subset by (force simp: closedin_sum_topology) lemma closedin_sum_topology_alt: "closedin (sum_topology X I) U \ (\T. U = Sigma I T \ (\i \ I. closedin (X i) (T i)))" using closedin_subset unfolding closedin_sum_topology by auto blast lemma forall_closedin_sum_topology: "(\U. closedin (sum_topology X I) U \ P U) \ (\t. (\i \ I. closedin (X i) (t i)) \ P(Sigma I t))" by (metis closedin_sum_topology_alt) lemma exists_closedin_sum_topology: "(\U. closedin (sum_topology X I) U \ P U) \ (\T. (\i \ I. closedin (X i) (T i)) \ P(Sigma I T))" by (simp add: closedin_sum_topology_alt) blast lemma open_map_component_injection: "i \ I \ open_map (X i) (sum_topology X I) (\x. (i,x))" unfolding open_map_def openin_sum_topology using openin_subset openin_subopen by (force simp: image_iff) lemma closed_map_component_injection: assumes "i \ I" shows "closed_map(X i) (sum_topology X I) (\x. (i,x))" proof - have "closedin (X j) {x. j = i \ x \ U}" if "\U S. closedin U S \ S \ topspace U" and "closedin (X i) U" and "j \ I" for U j using that by (cases "j=i") auto then show ?thesis using closedin_subset assms by (force simp: closed_map_def closedin_sum_topology image_iff) qed lemma continuous_map_component_injection: "i \ I \ continuous_map(X i) (sum_topology X I) (\x. (i,x))" apply (clarsimp simp: continuous_map_def openin_sum_topology) by (smt (verit, best) Collect_cong mem_Collect_eq openin_subset subsetD) lemma subtopology_sum_topology: "subtopology (sum_topology X I) (Sigma I S) = sum_topology (\i. subtopology (X i) (S i)) I" unfolding topology_eq proof (intro strip iffI) fix T assume *: "openin (subtopology (sum_topology X I) (Sigma I S)) T" then obtain U where U: "\i\I. openin (X i) (U i) \ T = Sigma I U \ Sigma I S" by (auto simp: openin_subtopology openin_sum_topology_alt) have "T = (SIGMA i:I. U i \ S i)" proof show "T \ (SIGMA i:I. U i \ S i)" by (metis "*" SigmaE Sigma_Int_distrib2 U openin_imp_subset subset_iff) show "(SIGMA i:I. U i \ S i) \ T" using U by fastforce qed then show "openin (sum_topology (\i. subtopology (X i) (S i)) I) T" by (simp add: U openin_disjoint_union openin_subtopology_Int) next fix T assume *: "openin (sum_topology (\i. subtopology (X i) (S i)) I) T" then obtain U where "\i\I. \T. openin (X i) T \ U i = T \ S i" and Teq: "T = Sigma I U" by (auto simp: openin_subtopology openin_sum_topology_alt) then obtain B where "\i. i\I \ openin (X i) (B i) \ U i = B i \ S i" by metis then show "openin (subtopology (sum_topology X I) (Sigma I S)) T" by (auto simp: openin_subtopology openin_sum_topology_alt Teq) qed lemma embedding_map_component_injection: "i \ I \ embedding_map (X i) (sum_topology X I) (\x. (i,x))" by (metis injective_open_imp_embedding_map continuous_map_component_injection open_map_component_injection inj_onI prod.inject) +lemma topological_property_of_sum_component: + assumes major: "P (sum_topology X I)" + and minor: "\X S. \P X; closedin X S; openin X S\ \ P(subtopology X S)" + and PQ: "\X Y. X homeomorphic_space Y \ (P X \ Q Y)" + shows "(\i \ I. Q(X i))" +proof - + have "Q(X i)" if "i \ I" for i + proof - + have "P(subtopology (sum_topology X I) (Pair i ` topspace (X i)))" + by (meson closed_map_component_injection closed_map_def closedin_topspace major minor open_map_component_injection open_map_def openin_topspace that) + then show ?thesis + by (metis PQ embedding_map_component_injection embedding_map_imp_homeomorphic_space homeomorphic_space_sym that) + qed + then show ?thesis by metis +qed + end diff --git a/src/HOL/Statespace/state_space.ML b/src/HOL/Statespace/state_space.ML --- a/src/HOL/Statespace/state_space.ML +++ b/src/HOL/Statespace/state_space.ML @@ -1,758 +1,731 @@ (* Title: HOL/Statespace/state_space.ML Author: Norbert Schirmer, TU Muenchen, 2007 Author: Norbert Schirmer, Apple, 2021 *) signature STATE_SPACE = sig val distinct_compsN : string val getN : string val putN : string val injectN : string val namespaceN : string val projectN : string val valuetypesN : string val namespace_definition : bstring -> typ -> (xstring, string) Expression.expr * (binding * string option * mixfix) list -> string list -> string list -> theory -> theory val define_statespace : string list -> string -> ((string * bool) * (string list * bstring * (string * string) list)) list -> (string * string) list -> theory -> theory val define_statespace_i : string option -> string list -> string -> ((string * bool) * (typ list * bstring * (string * string) list)) list -> (string * typ) list -> theory -> theory val statespace_decl : ((string list * bstring) * (((string * bool) * (string list * xstring * (bstring * bstring) list)) list * (bstring * string) list)) parser - val neq_x_y : Proof.context -> term -> term -> thm option val distinctNameSolver : Simplifier.solver val distinctTree_tac : Proof.context -> int -> tactic val distinct_simproc : Simplifier.simproc val is_statespace : Context.generic -> xstring -> bool val get_comp' : Context.generic -> string -> (typ * string list) option val get_comp : Context.generic -> string -> (typ * string) option (* legacy wrapper, eventually replace by primed version *) val get_comps : Context.generic -> (typ * string list) Termtab.table - val get_silent : Context.generic -> bool - val set_silent : bool -> Context.generic -> Context.generic - + val silent : bool Config.T val gen_lookup_tr : Proof.context -> term -> string -> term val lookup_swap_tr : Proof.context -> term list -> term val lookup_tr : Proof.context -> term list -> term val lookup_tr' : Proof.context -> term list -> term val gen'_update_tr : bool -> bool -> Proof.context -> string -> term -> term -> term val gen_update_tr : (* legacy wrapper, eventually replace by primed version *) bool -> Proof.context -> string -> term -> term -> term val update_tr : Proof.context -> term list -> term val update_tr' : Proof.context -> term list -> term - val trace_name_space_data: Context.generic -> unit - val trace_state_space_data: Context.generic -> unit + val trace_data: Context.generic -> unit end; structure StateSpace : STATE_SPACE = struct (* Names *) val distinct_compsN = "distinct_names" val namespaceN = "_namespace" val valuetypesN = "_valuetypes" val projectN = "project" val injectN = "inject" val getN = "get" val putN = "put" val project_injectL = "StateSpaceLocale.project_inject"; (* Library *) fun fold1 f xs = fold f (tl xs) (hd xs) fun fold1' f [] x = x | fold1' f xs _ = fold1 f xs fun sorted_subset eq [] ys = true | sorted_subset eq (x::xs) [] = false | sorted_subset eq (x::xs) (y::ys) = if eq (x,y) then sorted_subset eq xs ys else sorted_subset eq (x::xs) ys; fun comps_of_distinct_thm thm = Thm.prop_of thm |> (fn (_$(_$t)) => DistinctTreeProver.dest_tree t) |> map (fst o dest_Free) |> sort_strings; fun insert_tagged_distinct_thms tagged_thm tagged_thms = let - fun ins t1 [] = [t1] + fun ins t1 [] = [t1] | ins (t1 as (names1, _)) ((t2 as (names2, _))::thms) = if Ord_List.subset string_ord (names1, names2) then t2::thms else if Ord_List.subset string_ord (names2, names1) then ins t1 thms else t2 :: ins t1 thms - in + in ins tagged_thm tagged_thms end -fun join_tagged_distinct_thms tagged_thms1 tagged_thms2 = +fun join_tagged_distinct_thms tagged_thms1 tagged_thms2 = tagged_thms1 |> fold insert_tagged_distinct_thms tagged_thms2 fun tag_distinct_thm thm = (comps_of_distinct_thm thm, thm) val tag_distinct_thms = map tag_distinct_thm -fun join_distinct_thms thms1 thms2 = - if pointer_eq (thms1, thms2) then thms1 - else join_tagged_distinct_thms (tag_distinct_thms thms1) (tag_distinct_thms thms2) - |> map snd +fun join_distinct_thms (thms1, thms2) = + if pointer_eq (thms1, thms2) then thms1 + else join_tagged_distinct_thms (tag_distinct_thms thms1) (tag_distinct_thms thms2) |> map snd -fun insert_distinct_thm thm thms = join_distinct_thms thms [thm] - -val join_distinctthm_tab = Symtab.join (fn name => fn (thms1, thms2) => join_distinct_thms thms1 thms2) +fun insert_distinct_thm thm thms = join_distinct_thms (thms, [thm]) fun join_declinfo_entry name (T1:typ, names1:string list) (T2, names2) = let fun typ_info T names = @{make_string} T ^ " " ^ Pretty.string_of (Pretty.str_list "(" ")" names); in if T1 = T2 then (T1, distinct (op =) (names1 @ names2)) - else error ("statespace component '" ^ name ^ "' disagrees on type:\n " ^ + else error ("statespace component '" ^ name ^ "' disagrees on type:\n " ^ typ_info T1 names1 ^ " vs. " ^ typ_info T2 names2 - ) + ) end fun guess_name (Free (x,_)) = x | guess_name _ = "unknown" -val join_declinfo = Termtab.join (fn trm => uncurry (join_declinfo_entry (guess_name trm))) +val join_declinfo = Termtab.join (fn t => uncurry (join_declinfo_entry (guess_name t))) (* A component might appear in *different* statespaces within the same context. However, it must be mapped to the same type. Note that this information is currently only properly maintained within contexts where all components are actually "fixes" and not arbitrary terms. Moreover, on - the theory level the info stays empty. + the theory level the info stays empty. This means that on the theory level we do not make use of the syntax and the tree-based distinct simprocs. - N.B: It might still make sense (from a performance perspective) to overcome this limitation + N.B: It might still make sense (from a performance perspective) to overcome this limitation and even use the simprocs when a statespace is interpreted for concrete names like HOL-strings. - Once the distinct-theorem is proven by the interpretation the simproc can use the positions in - the tree to derive distinctness of two strings vs. HOL-string comparison. + Once the distinct-theorem is proven by the interpretation the simproc can use the positions in + the tree to derive distinctness of two strings vs. HOL-string comparison. *) -type namespace_info = - {declinfo: (typ * string list) Termtab.table, (* type, names of statespaces of component *) - distinctthm: thm list Symtab.table, (* minimal list of maximal distinctness-assumptions for a component name *) - silent: bool + +type statespace_info = + {args: (string * sort) list, (* type arguments *) + parents: (typ list * string * string option list) list, + (* type instantiation, state-space name, component renamings *) + components: (string * typ) list, + types: typ list (* range types of state space *) }; -structure NameSpaceData = Generic_Data +structure Data = Generic_Data ( - type T = namespace_info; - val empty = {declinfo = Termtab.empty, distinctthm = Symtab.empty, silent = false}; - fun merge - ({declinfo=declinfo1, distinctthm=distinctthm1, silent=silent1}, - {declinfo=declinfo2, distinctthm=distinctthm2, silent=silent2}) : T = - {declinfo = join_declinfo (declinfo1, declinfo2), - distinctthm = join_distinctthm_tab (distinctthm1, distinctthm2), - silent = silent1 andalso silent2 (* FIXME odd merge *)} + type T = + (typ * string list) Termtab.table * (*declinfo: type, names of statespaces of component*) + thm list Symtab.table * (*distinctthm: minimal list of maximal distinctness-assumptions for a component name*) + statespace_info Symtab.table; + val empty = (Termtab.empty, Symtab.empty, Symtab.empty); + fun merge ((declinfo1, distinctthm1, statespaces1), (declinfo2, distinctthm2, statespaces2)) = + (join_declinfo (declinfo1, declinfo2), + Symtab.join (K join_distinct_thms) (distinctthm1, distinctthm2), + Symtab.merge (K true) (statespaces1, statespaces2)); ); -fun trace_name_space_data context = - tracing ("NameSpaceData: " ^ @{make_string} (NameSpaceData.get context)) - -fun make_namespace_data declinfo distinctthm silent = - {declinfo=declinfo,distinctthm=distinctthm,silent=silent}; - +val get_declinfo = #1 o Data.get +val get_distinctthm = #2 o Data.get +val get_statespaces = #3 o Data.get -fun update_declinfo (n,v) ctxt = - let - val {declinfo,distinctthm,silent} = NameSpaceData.get ctxt; - val v = apsnd single v - in NameSpaceData.put - (make_namespace_data (Termtab.map_default (n,v) (join_declinfo_entry (guess_name n) v) declinfo) distinctthm silent) ctxt - end; +val map_declinfo = Data.map o @{apply 3(1)} +val map_distinctthm = Data.map o @{apply 3(2)} +val map_statespaces = Data.map o @{apply 3(3)} -fun set_silent silent ctxt = - let val {declinfo,distinctthm,...} = NameSpaceData.get ctxt; - in NameSpaceData.put - (make_namespace_data declinfo distinctthm silent) ctxt - end; +fun trace_data context = + tracing ("StateSpace.Data: " ^ @{make_string} + {declinfo = get_declinfo context, + distinctthm = get_distinctthm context, + statespaces = get_statespaces context}) -val get_silent = #silent o NameSpaceData.get; +fun update_declinfo (n,v) = map_declinfo (fn declinfo => + let val vs = apsnd single v + in Termtab.map_default (n, vs) (join_declinfo_entry (guess_name n) vs) declinfo end); fun expression_no_pos (expr, fixes) : Expression.expression = (map (fn (name, inst) => ((name, Position.none), inst)) expr, fixes); fun prove_interpretation_in ctxt_tac (name, expr) thy = thy |> Interpretation.global_sublocale_cmd (name, Position.none) (expression_no_pos expr) [] |> Proof.global_terminal_proof ((Method.Basic (fn ctxt => SIMPLE_METHOD (ctxt_tac ctxt)), Position.no_range), NONE) |> Proof_Context.theory_of fun add_locale name expr elems thy = thy |> Expression.add_locale (Binding.name name) (Binding.name name) [] expr elems |> snd |> Local_Theory.exit; fun add_locale_cmd name expr elems thy = thy |> Expression.add_locale_cmd (Binding.name name) Binding.empty [] (expression_no_pos expr) elems |> snd |> Local_Theory.exit; -type statespace_info = - {args: (string * sort) list, (* type arguments *) - parents: (typ list * string * string option list) list, - (* type instantiation, state-space name, component renamings *) - components: (string * typ) list, - types: typ list (* range types of state space *) - }; +fun is_statespace context name = + Symtab.defined (get_statespaces context) (Locale.intern (Context.theory_of context) name) -structure StateSpaceData = Generic_Data -( - type T = statespace_info Symtab.table; - val empty = Symtab.empty; - fun merge data : T = Symtab.merge (K true) data; -); - -fun is_statespace context name = - Symtab.defined (StateSpaceData.get context) (Locale.intern (Context.theory_of context) name) +fun add_statespace name args parents components types = + map_statespaces (Symtab.update_new (name, {args=args,parents=parents, components=components,types=types})) -fun trace_state_space_data context = - tracing ("StateSpaceData: " ^ @{make_string} (StateSpaceData.get context)) - -fun add_statespace name args parents components types ctxt = - StateSpaceData.put - (Symtab.update_new (name, {args=args,parents=parents, - components=components,types=types}) (StateSpaceData.get ctxt)) - ctxt; - -fun get_statespace ctxt name = - Symtab.lookup (StateSpaceData.get ctxt) name; +val get_statespace = Symtab.lookup o get_statespaces +val the_statespace = the oo get_statespace fun mk_free ctxt name = if Variable.is_fixed ctxt name orelse Variable.is_declared ctxt name then - let + let val n' = Variable.intern_fixed ctxt name |> perhaps Long_Name.dest_hidden; val free = Free (n', Proof_Context.infer_type ctxt (n', dummyT)) in SOME (free) end else (tracing ("variable not fixed or declared: " ^ name); NONE) -fun get_dist_thm ctxt name = Symtab.lookup (#distinctthm (NameSpaceData.get ctxt)) name; +val get_dist_thm = Symtab.lookup o get_distinctthm; -fun get_dist_thm2 ctxt x y = +fun get_dist_thm2 ctxt x y = (let - val dist_thms = [x, y] |> map (#1 o dest_Free) + val dist_thms = [x, y] |> map (#1 o dest_Free) |> map (these o get_dist_thm (Context.Proof ctxt)) |> flat; fun get_paths dist_thm = let val ctree = Thm.cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; val tree = Thm.term_of ctree; val x_path = the (DistinctTreeProver.find_tree x tree); val y_path = the (DistinctTreeProver.find_tree y tree); in SOME (dist_thm, x_path, y_path) end handle Option.Option => NONE - + val result = get_first get_paths dist_thms - in + in result end) - - -fun get_comp' ctxt name = - mk_free (Context.proof_of ctxt) name - |> Option.mapPartial (fn t => + + +fun get_comp' context name = + mk_free (Context.proof_of context) name + |> Option.mapPartial (fn t => let - val declinfo = #declinfo (NameSpaceData.get ctxt) - in + val declinfo = get_declinfo context + in case Termtab.lookup declinfo t of - NONE => (* during syntax phase, types of fixes might not yet be constrained completely *) + NONE => (* during syntax phase, types of fixes might not yet be constrained completely *) AList.lookup (fn (x, Free (n,_)) => n = x | _ => false) (Termtab.dest declinfo) name | some => some end) (* legacy wrapper *) fun get_comp ctxt name = - get_comp' ctxt name |> Option.map (apsnd (fn ns => if null ns then "" else hd ns)) + get_comp' ctxt name |> Option.map (apsnd (fn ns => if null ns then "" else hd ns)) -fun get_comps ctxt = #declinfo (NameSpaceData.get ctxt) +val get_comps = get_declinfo + (*** Tactics ***) fun neq_x_y ctxt x y = (let val (dist_thm, x_path, y_path) = the (get_dist_thm2 ctxt x y); val thm = DistinctTreeProver.distinctTreeProver ctxt dist_thm x_path y_path; in SOME thm end handle Option.Option => NONE) fun distinctTree_tac ctxt = SUBGOAL (fn (goal, i) => (case goal of Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\Not\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ (x as Free _) $ (y as Free _))) => (case neq_x_y ctxt x y of SOME neq => resolve_tac ctxt [neq] i | NONE => no_tac) | _ => no_tac)); val distinctNameSolver = mk_solver "distinctNameSolver" distinctTree_tac; val distinct_simproc = Simplifier.make_simproc \<^context> "StateSpace.distinct_simproc" {lhss = [\<^term>\x = y\], proc = fn _ => fn ctxt => fn ct => (case Thm.term_of ct of Const (\<^const_name>\HOL.eq\,_) $ (x as Free _) $ (y as Free _) => Option.map (fn neq => DistinctTreeProver.neq_to_eq_False OF [neq]) (neq_x_y ctxt x y) | _ => NONE)}; fun interprete_parent name dist_thm_name parent_expr thy = let fun solve_tac ctxt = CSUBGOAL (fn (goal, i) => let val distinct_thm = Proof_Context.get_thm ctxt dist_thm_name; val rule = DistinctTreeProver.distinct_implProver ctxt distinct_thm goal; in resolve_tac ctxt [rule] i end); fun tac ctxt = Locale.intro_locales_tac {strict = true, eager = true} ctxt [] THEN ALLGOALS (solve_tac ctxt); in thy |> prove_interpretation_in tac (name, parent_expr) end; fun namespace_definition name nameT parent_expr parent_comps new_comps thy = let val all_comps = parent_comps @ new_comps; val vars = (map (fn n => (Binding.name n, NONE, NoSyn)) all_comps); val dist_thm_name = distinct_compsN; val dist_thm_full_name = dist_thm_name; fun type_attr phi = Thm.declaration_attribute (fn thm => fn context => (case context of Context.Theory _ => context | Context.Proof ctxt => let - val {declinfo,distinctthm=tt,silent} = NameSpaceData.get context; + val declinfo = get_declinfo context + val tt = get_distinctthm context; val all_names = comps_of_distinct_thm thm; - fun upd name = Symtab.map_default (name, [thm]) (fn thms => insert_distinct_thm thm thms) + fun upd name = Symtab.map_default (name, [thm]) (insert_distinct_thm thm) val tt' = tt |> fold upd all_names; val context' = Context_Position.set_visible false ctxt addsimprocs [distinct_simproc] |> Context_Position.restore_visible ctxt |> Context.Proof - |> NameSpaceData.put {declinfo=declinfo,distinctthm=tt',silent=silent}; + |> map_declinfo (K declinfo) + |> map_distinctthm (K tt'); in context' end)); val attr = Attrib.internal type_attr; val assume = ((Binding.name dist_thm_name, [attr]), [(HOLogic.Trueprop $ (Const (\<^const_name>\all_distinct\, Type (\<^type_name>\tree\, [nameT]) --> HOLogic.boolT) $ DistinctTreeProver.mk_tree (fn n => Free (n, nameT)) nameT (sort fast_string_ord all_comps)), [])]); in thy |> add_locale name ([], vars) [Element.Assumes [assume]] |> Proof_Context.theory_of |> interprete_parent name dist_thm_full_name parent_expr end; fun encode_dot x = if x = #"." then #"_" else x; fun encode_type (TFree (s, _)) = s | encode_type (TVar ((s,i),_)) = "?" ^ s ^ string_of_int i | encode_type (Type (n,Ts)) = let val Ts' = fold1' (fn x => fn y => x ^ "_" ^ y) (map encode_type Ts) ""; val n' = String.map encode_dot n; in if Ts'="" then n' else Ts' ^ "_" ^ n' end; fun project_name T = projectN ^"_"^encode_type T; fun inject_name T = injectN ^"_"^encode_type T; fun add_declaration name decl thy = thy |> Named_Target.init [] name |> (fn lthy => Local_Theory.declaration {syntax = true, pervasive = false} (decl lthy) lthy) |> Local_Theory.exit_global; fun parent_components thy (Ts, pname, renaming) = let - val ctxt = Context.Theory thy; fun rename [] xs = xs | rename (NONE::rs) (x::xs) = x::rename rs xs | rename (SOME r::rs) ((x,T)::xs) = (r,T)::rename rs xs; - val {args, parents, components, ...} = the (Symtab.lookup (StateSpaceData.get ctxt) pname); + val {args, parents, components, ...} = the_statespace (Context.Theory thy) pname; val inst = map fst args ~~ Ts; val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); val parent_comps = maps (fn (Ts',n,rs) => parent_components thy (map subst Ts', n, rs)) parents; val all_comps = rename renaming (parent_comps @ map (apsnd subst) components); in all_comps end; fun statespace_definition state_type args name parents parent_comps components thy = let val full_name = Sign.full_bname thy name; val all_comps = parent_comps @ components; val components' = map (fn (n,T) => (n,(T,full_name))) components; fun parent_expr (prefix, (_, n, rs)) = (suffix namespaceN n, (prefix, (Expression.Positional rs,[]))); val parents_expr = map parent_expr parents; fun distinct_types Ts = let val tab = fold (fn T => fn tab => Typtab.update (T,()) tab) Ts Typtab.empty; in map fst (Typtab.dest tab) end; val Ts = distinct_types (map snd all_comps); val arg_names = map fst args; val valueN = singleton (Name.variant_list arg_names) "'value"; val nameN = singleton (Name.variant_list (valueN :: arg_names)) "'name"; val valueT = TFree (valueN, Sign.defaultS thy); val nameT = TFree (nameN, Sign.defaultS thy); val stateT = nameT --> valueT; fun projectT T = valueT --> T; fun injectT T = T --> valueT; val locinsts = map (fn T => (project_injectL, ((encode_type T,false),(Expression.Positional [SOME (Free (project_name T,projectT T)), SOME (Free ((inject_name T,injectT T)))],[])))) Ts; val locs = maps (fn T => [(Binding.name (project_name T),NONE,NoSyn), (Binding.name (inject_name T),NONE,NoSyn)]) Ts; val constrains = maps (fn T => [(project_name T,projectT T),(inject_name T,injectT T)]) Ts; fun interprete_parent_valuetypes (prefix, (Ts, pname, _)) thy = let - val {args,types,...} = - the (Symtab.lookup (StateSpaceData.get (Context.Theory thy)) pname); + val {args,types,...} = the_statespace (Context.Theory thy) pname; val inst = map fst args ~~ Ts; val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); val pars = maps ((fn T => [project_name T,inject_name T]) o subst) types; val expr = ([(suffix valuetypesN name, (prefix, (Expression.Positional (map SOME pars),[])))],[]); in prove_interpretation_in (fn ctxt => ALLGOALS (solve_tac ctxt (Assumption.all_prems_of ctxt))) (suffix valuetypesN name, expr) thy end; fun interprete_parent (prefix, (_, pname, rs)) = let val expr = ([(pname, (prefix, (Expression.Positional rs,[])))],[]) in prove_interpretation_in (fn ctxt => Locale.intro_locales_tac {strict = true, eager = false} ctxt []) (full_name, expr) end; fun declare_declinfo updates lthy phi ctxt = let fun upd_prf ctxt = let fun upd (n,v) = let val nT = Proof_Context.infer_type (Local_Theory.target_of lthy) (n, dummyT) in Context.proof_map (update_declinfo (Morphism.term phi (Free (n,nT)),v)) end; - val ctxt' = ctxt |> fold upd updates + val ctxt' = ctxt |> fold upd updates in ctxt' end; in Context.mapping I upd_prf ctxt end; fun string_of_typ T = Print_Mode.setmp [] (Syntax.string_of_typ (Config.put show_sorts true (Syntax.init_pretty_global thy))) T; val fixestate = (case state_type of NONE => [] | SOME s => let val fx = Element.Fixes [(Binding.name s,SOME (string_of_typ stateT),NoSyn)]; val cs = Element.Constrains (map (fn (n,T) => (n,string_of_typ T)) ((map (fn (n,_) => (n,nameT)) all_comps) @ constrains)) in [fx,cs] end ) in thy |> namespace_definition (suffix namespaceN name) nameT (parents_expr,[]) (map fst parent_comps) (map fst components) |> Context.theory_map (add_statespace full_name args (map snd parents) components []) |> add_locale (suffix valuetypesN name) (locinsts,locs) [] |> Proof_Context.theory_of |> fold interprete_parent_valuetypes parents |> add_locale_cmd name ([(suffix namespaceN full_name ,(("",false),(Expression.Named [],[]))), (suffix valuetypesN full_name,(("",false),(Expression.Named [],[])))],[]) fixestate |> Proof_Context.theory_of |> fold interprete_parent parents |> add_declaration full_name (declare_declinfo components') end; (* prepare arguments *) fun read_typ ctxt raw_T env = let val ctxt' = fold (Variable.declare_typ o TFree) env ctxt; val T = Syntax.read_typ ctxt' raw_T; val env' = Term.add_tfreesT T env; in (T, env') end; fun cert_typ ctxt raw_T env = let val thy = Proof_Context.theory_of ctxt; val T = Type.no_tvars (Sign.certify_typ thy raw_T) handle TYPE (msg, _, _) => error msg; val env' = Term.add_tfreesT T env; in (T, env') end; fun gen_define_statespace prep_typ state_space args name parents comps thy = let (* - args distinct - only args may occur in comps and parent-instantiations - number of insts must match parent args - no duplicate renamings - renaming should occur in namespace *) val _ = writeln ("Defining statespace " ^ quote name ^ " ..."); val ctxt = Proof_Context.init_global thy; fun add_parent (prefix, (Ts, pname, rs)) env = let val prefix' = (case prefix of ("", mandatory) => (pname, mandatory) | _ => prefix); val full_pname = Sign.full_bname thy pname; val {args,components,...} = (case get_statespace (Context.Theory thy) full_pname of SOME r => r - | NONE => error ("Undefined statespace " ^ quote pname)); + | NONE => error ("Undefined statespace " ^ quote pname)); val (Ts',env') = fold_map (prep_typ ctxt) Ts env handle ERROR msg => cat_error msg ("The error(s) above occurred in parent statespace specification " ^ quote pname); val err_insts = if length args <> length Ts' then ["number of type instantiation(s) does not match arguments of parent statespace " ^ quote pname] else []; val rnames = map fst rs val err_dup_renamings = (case duplicates (op =) rnames of [] => [] | dups => ["Duplicate renaming(s) for " ^ commas dups]) val cnames = map fst components; val err_rename_unknowns = (case subtract (op =) cnames rnames of [] => [] | rs => ["Unknown components " ^ commas rs]); val rs' = map (AList.lookup (op =) rs o fst) components; val errs =err_insts @ err_dup_renamings @ err_rename_unknowns in if null errs then ((prefix', (Ts', full_pname, rs')), env') else error (cat_lines (errs @ ["in parent statespace " ^ quote pname])) end; val (parents',env) = fold_map add_parent parents []; val err_dup_args = (case duplicates (op =) args of [] => [] | dups => ["Duplicate type argument(s) " ^ commas dups]); val err_dup_components = (case duplicates (op =) (map fst comps) of [] => [] | dups => ["Duplicate state-space components " ^ commas dups]); fun prep_comp (n,T) env = let val (T', env') = prep_typ ctxt T env handle ERROR msg => cat_error msg ("The error(s) above occurred in component " ^ quote n) in ((n,T'), env') end; val (comps',env') = fold_map prep_comp comps env; val err_extra_frees = (case subtract (op =) args (map fst env') of [] => [] | extras => ["Extra free type variable(s) " ^ commas extras]); val defaultS = Sign.defaultS thy; val args' = map (fn x => (x, AList.lookup (op =) env x |> the_default defaultS)) args; fun fst_eq ((x:string,_),(y,_)) = x = y; fun snd_eq ((_,t:typ),(_,u)) = t = u; val raw_parent_comps = maps (parent_components thy o snd) parents'; fun check_type (n,T) = (case distinct (snd_eq) (filter (curry fst_eq (n,T)) raw_parent_comps) of [] => [] | [_] => [] | rs => ["Different types for component " ^ quote n ^ ": " ^ commas (map (Syntax.string_of_typ ctxt o snd) rs)]) val err_dup_types = maps check_type (duplicates fst_eq raw_parent_comps) val parent_comps = distinct (fst_eq) raw_parent_comps; val all_comps = parent_comps @ comps'; val err_comp_in_parent = (case duplicates (op =) (map fst all_comps) of [] => [] | xs => ["Components already defined in parents: " ^ commas_quote xs]); val errs = err_dup_args @ err_dup_components @ err_extra_frees @ err_dup_types @ err_comp_in_parent; in if null errs then thy |> statespace_definition state_space args' name parents' parent_comps comps' else error (cat_lines errs) end handle ERROR msg => cat_error msg ("Failed to define statespace " ^ quote name); val define_statespace = gen_define_statespace read_typ NONE; val define_statespace_i = gen_define_statespace cert_typ; (*** parse/print - translations ***) +val silent = Attrib.setup_config_bool \<^binding>\statespace_silent\ (K false); fun gen_lookup_tr ctxt s n = (case get_comp' (Context.Proof ctxt) n of SOME (T, _) => Syntax.const \<^const_name>\StateFun.lookup\ $ Syntax.free (project_name T) $ Syntax.free n $ s | NONE => - if get_silent (Context.Proof ctxt) + if Config.get ctxt silent then Syntax.const \<^const_name>\StateFun.lookup\ $ Syntax.const \<^const_syntax>\undefined\ $ Syntax.free n $ s else raise TERM ("StateSpace.gen_lookup_tr: component " ^ quote n ^ " not defined", [])); fun lookup_tr ctxt [s, x] = (case Term_Position.strip_positions x of Free (n,_) => gen_lookup_tr ctxt s n | _ => raise Match); fun lookup_swap_tr ctxt [Free (n,_),s] = gen_lookup_tr ctxt s n; fun lookup_tr' ctxt [_ $ Free (prj, _), n as (_ $ Free (name, _)), s] = (case get_comp' (Context.Proof ctxt) name of SOME (T, _) => if prj = project_name T then Syntax.const "_statespace_lookup" $ s $ n else raise Match | NONE => raise Match) | lookup_tr' _ _ = raise Match; fun gen'_update_tr const_val id ctxt n v s = let fun pname T = if id then \<^const_name>\Fun.id\ else project_name T; fun iname T = if id then \<^const_name>\Fun.id\ else inject_name T; val v = if const_val then (Syntax.const \<^const_name>\K_statefun\ $ v) else v in (case get_comp' (Context.Proof ctxt) n of SOME (T, _) => Syntax.const \<^const_name>\StateFun.update\ $ Syntax.free (pname T) $ Syntax.free (iname T) $ Syntax.free n $ v $ s | NONE => - if get_silent (Context.Proof ctxt) then + if Config.get ctxt silent then Syntax.const \<^const_name>\StateFun.update\ $ Syntax.const \<^const_syntax>\undefined\ $ Syntax.const \<^const_syntax>\undefined\ $ Syntax.free n $ v $ s else raise TERM ("StateSpace.gen_update_tr: component " ^ n ^ " not defined", [])) end; val gen_update_tr = gen'_update_tr true fun update_tr ctxt [s, x, v] = (case Term_Position.strip_positions x of Free (n, _) => gen'_update_tr true false ctxt n v s | _ => raise Match); fun update_tr' ctxt [_ $ Free (prj, _), _ $ Free (inj, _), n as (_ $ Free (name, _)), (Const (k, _) $ v), s] = if Long_Name.base_name k = Long_Name.base_name \<^const_name>\K_statefun\ then (case get_comp' (Context.Proof ctxt) name of SOME (T, _) => if inj = inject_name T andalso prj = project_name T then Syntax.const "_statespace_update" $ s $ n $ v else raise Match | NONE => raise Match) else raise Match | update_tr' _ _ = raise Match; (*** outer syntax *) local val type_insts = Parse.typ >> single || \<^keyword>\(\ |-- Parse.!!! (Parse.list1 Parse.typ --| \<^keyword>\)\) val comp = Parse.name -- (\<^keyword>\::\ |-- Parse.!!! Parse.typ); fun plus1_unless test scan = scan ::: Scan.repeat (\<^keyword>\+\ |-- Scan.unless test (Parse.!!! scan)); val mapsto = \<^keyword>\=\; val rename = Parse.name -- (mapsto |-- Parse.name); val renames = Scan.optional (\<^keyword>\[\ |-- Parse.!!! (Parse.list1 rename --| \<^keyword>\]\)) []; val parent = Parse_Spec.locale_prefix -- ((type_insts -- Parse.name) || (Parse.name >> pair [])) -- renames >> (fn ((prefix, (insts, name)), renames) => (prefix, (insts, name, renames))); in val statespace_decl = Parse.type_args -- Parse.name -- (\<^keyword>\=\ |-- ((Scan.repeat1 comp >> pair []) || (plus1_unless comp parent -- Scan.optional (\<^keyword>\+\ |-- Parse.!!! (Scan.repeat1 comp)) []))); val _ = Outer_Syntax.command \<^command_keyword>\statespace\ "define state-space as locale context" (statespace_decl >> (fn ((args, name), (parents, comps)) => Toplevel.theory (define_statespace args name parents comps))); end; end; diff --git a/src/HOL/Tools/Lifting/lifting_info.ML b/src/HOL/Tools/Lifting/lifting_info.ML --- a/src/HOL/Tools/Lifting/lifting_info.ML +++ b/src/HOL/Tools/Lifting/lifting_info.ML @@ -1,576 +1,584 @@ (* Title: HOL/Tools/Lifting/lifting_info.ML Author: Ondrej Kuncar Context data for the lifting package. *) signature LIFTING_INFO = sig type quot_map = {rel_quot_thm: thm} val lookup_quot_maps: Proof.context -> string -> quot_map option val print_quot_maps: Proof.context -> unit type pcr = {pcrel_def: thm, pcr_cr_eq: thm} type quotient = {quot_thm: thm, pcr_info: pcr option} val pcr_eq: pcr * pcr -> bool val quotient_eq: quotient * quotient -> bool val transform_quotient: morphism -> quotient -> quotient val lookup_quotients: Proof.context -> string -> quotient option val lookup_quot_thm_quotients: Proof.context -> thm -> quotient option val update_quotients: string -> quotient -> Context.generic -> Context.generic val delete_quotients: thm -> Context.generic -> Context.generic val print_quotients: Proof.context -> unit type restore_data = {quotient : quotient, transfer_rules: thm Item_Net.T} val lookup_restore_data: Proof.context -> string -> restore_data option val init_restore_data: string -> quotient -> Context.generic -> Context.generic val add_transfer_rules_in_restore_data: string -> thm Item_Net.T -> Context.generic -> Context.generic val get_relator_eq_onp_rules: Proof.context -> thm list val get_reflexivity_rules: Proof.context -> thm list val add_reflexivity_rule_attribute: attribute type relator_distr_data = {pos_mono_rule: thm, neg_mono_rule: thm, pos_distr_rules: thm list, neg_distr_rules: thm list} val lookup_relator_distr_data: Proof.context -> string -> relator_distr_data option val add_no_code_type: string -> Context.generic -> Context.generic val is_no_code_type: Proof.context -> string -> bool val get_quot_maps : Proof.context -> quot_map Symtab.table val get_quotients : Proof.context -> quotient Symtab.table val get_relator_distr_data : Proof.context -> relator_distr_data Symtab.table val get_restore_data : Proof.context -> restore_data Symtab.table val get_no_code_types : Proof.context -> Symset.T end structure Lifting_Info: LIFTING_INFO = struct open Lifting_Util (* context data *) type quot_map = {rel_quot_thm: thm} type pcr = {pcrel_def: thm, pcr_cr_eq: thm} type quotient = {quot_thm: thm, pcr_info: pcr option} type relator_distr_data = {pos_mono_rule: thm, neg_mono_rule: thm, pos_distr_rules: thm list, neg_distr_rules: thm list} type restore_data = {quotient : quotient, transfer_rules: thm Item_Net.T} fun pcr_eq ({pcrel_def = pcrel_def1, pcr_cr_eq = pcr_cr_eq1}, {pcrel_def = pcrel_def2, pcr_cr_eq = pcr_cr_eq2}) = Thm.eq_thm (pcrel_def1, pcrel_def2) andalso Thm.eq_thm (pcr_cr_eq1, pcr_cr_eq2) fun quotient_eq ({quot_thm = quot_thm1, pcr_info = pcr_info1}, {quot_thm = quot_thm2, pcr_info = pcr_info2}) = Thm.eq_thm (quot_thm1, quot_thm2) andalso eq_option pcr_eq (pcr_info1, pcr_info2) fun join_restore_data key (rd1:restore_data, rd2) = if pointer_eq (rd1, rd2) then raise Symtab.SAME else if not (quotient_eq (#quotient rd1, #quotient rd2)) then raise Symtab.DUP key else { quotient = #quotient rd1, transfer_rules = Item_Net.merge (#transfer_rules rd1, #transfer_rules rd2)} structure Data = Generic_Data ( type T = { quot_maps : quot_map Symtab.table, quotients : quotient Symtab.table, reflexivity_rules : thm Item_Net.T, relator_distr_data : relator_distr_data Symtab.table, restore_data : restore_data Symtab.table, no_code_types : Symset.T } val empty = { quot_maps = Symtab.empty, quotients = Symtab.empty, reflexivity_rules = Thm.item_net, relator_distr_data = Symtab.empty, restore_data = Symtab.empty, no_code_types = Symset.empty } fun merge ( { quot_maps = qm1, quotients = q1, reflexivity_rules = rr1, relator_distr_data = rdd1, restore_data = rd1, no_code_types = nct1 }, { quot_maps = qm2, quotients = q2, reflexivity_rules = rr2, relator_distr_data = rdd2, restore_data = rd2, no_code_types = nct2 } ) = { quot_maps = Symtab.merge (K true) (qm1, qm2), quotients = Symtab.merge (K true) (q1, q2), reflexivity_rules = Item_Net.merge (rr1, rr2), relator_distr_data = Symtab.merge (K true) (rdd1, rdd2), restore_data = Symtab.join join_restore_data (rd1, rd2), no_code_types = Symset.merge (nct1, nct2) } ) fun map_data f1 f2 f3 f4 f5 f6 { quot_maps, quotients, reflexivity_rules, relator_distr_data, restore_data, no_code_types } = { quot_maps = f1 quot_maps, quotients = f2 quotients, reflexivity_rules = f3 reflexivity_rules, relator_distr_data = f4 relator_distr_data, restore_data = f5 restore_data, no_code_types = f6 no_code_types } fun map_quot_maps f = map_data f I I I I I fun map_quotients f = map_data I f I I I I fun map_reflexivity_rules f = map_data I I f I I I fun map_relator_distr_data f = map_data I I I f I I fun map_restore_data f = map_data I I I I f I fun map_no_code_types f = map_data I I I I I f val get_quot_maps' = #quot_maps o Data.get val get_quotients' = #quotients o Data.get val get_reflexivity_rules' = #reflexivity_rules o Data.get val get_relator_distr_data' = #relator_distr_data o Data.get val get_restore_data' = #restore_data o Data.get val get_no_code_types' = #no_code_types o Data.get val get_quot_maps = get_quot_maps' o Context.Proof val get_quotients = get_quotients' o Context.Proof val get_relator_distr_data = get_relator_distr_data' o Context.Proof val get_restore_data = get_restore_data' o Context.Proof val get_no_code_types = get_no_code_types' o Context.Proof (* info about Quotient map theorems *) val lookup_quot_maps = Symtab.lookup o get_quot_maps fun quot_map_thm_sanity_check rel_quot_thm ctxt = let fun quot_term_absT ctxt quot_term = let - val (_, abs, _, _) = (dest_Quotient o HOLogic.dest_Trueprop) quot_term + val (_, abs, _, _) = dest_Quotient (HOLogic.dest_Trueprop quot_term) handle TERM (_, [t]) => error (Pretty.string_of (Pretty.block [Pretty.str "The Quotient map theorem is not in the right form.", Pretty.brk 1, Pretty.str "The following term is not the Quotient predicate:", Pretty.brk 1, Syntax.pretty_term ctxt t])) in fastype_of abs end val ((_, [rel_quot_thm_fixed]), ctxt') = Variable.importT [rel_quot_thm] ctxt val rel_quot_thm_prop = Thm.prop_of rel_quot_thm_fixed val rel_quot_thm_concl = Logic.strip_imp_concl rel_quot_thm_prop val rel_quot_thm_prems = Logic.strip_imp_prems rel_quot_thm_prop; val concl_absT = quot_term_absT ctxt' rel_quot_thm_concl val concl_tfrees = Term.add_tfree_namesT (concl_absT) [] val prems_tfrees = fold (fn typ => fn list => Term.add_tfree_namesT (quot_term_absT ctxt' typ) list) rel_quot_thm_prems [] val extra_prem_tfrees = case subtract (op =) concl_tfrees prems_tfrees of [] => [] - | extras => [Pretty.block ([Pretty.str "Extra type variables in the premises:", - Pretty.brk 1] @ - ((Pretty.commas o map (Pretty.str o quote)) extras) @ - [Pretty.str "."])] + | extras => + [Pretty.block ([Pretty.str "Extra type variables in the premises:", + Pretty.brk 1] @ + Pretty.commas (map (Pretty.str o quote) extras) @ + [Pretty.str "."])] val errs = extra_prem_tfrees in if null errs then () else error (cat_lines (["Sanity check of the quotient map theorem failed:",""] @ (map Pretty.string_of errs))) end fun add_quot_map rel_quot_thm context = let val _ = Context.cases (K ()) (quot_map_thm_sanity_check rel_quot_thm) context - val rel_quot_thm_concl = (Logic.strip_imp_concl o Thm.prop_of) rel_quot_thm - val (_, abs, _, _) = (dest_Quotient o HOLogic.dest_Trueprop) rel_quot_thm_concl - val relatorT_name = (fst o dest_Type o fst o dest_funT o fastype_of) abs + val rel_quot_thm_concl = Logic.strip_imp_concl (Thm.prop_of rel_quot_thm) + val (_, abs, _, _) = dest_Quotient (HOLogic.dest_Trueprop rel_quot_thm_concl) + val relatorT_name = fst (dest_Type (fst (dest_funT (fastype_of abs)))) val minfo = {rel_quot_thm = Thm.trim_context rel_quot_thm} - in - Data.map (map_quot_maps (Symtab.update (relatorT_name, minfo))) context - end + in (Data.map o map_quot_maps) (Symtab.update (relatorT_name, minfo)) context end val _ = Theory.setup (Attrib.setup \<^binding>\quot_map\ (Scan.succeed (Thm.declaration_attribute add_quot_map)) "declaration of the Quotient map theorem") fun print_quot_maps ctxt = let fun prt_map (ty_name, {rel_quot_thm}) = Pretty.block (separate (Pretty.brk 2) [Pretty.str "type:", Pretty.str ty_name, Pretty.str "quot. theorem:", Syntax.pretty_term ctxt (Thm.prop_of rel_quot_thm)]) in map prt_map (Symtab.dest (get_quot_maps ctxt)) |> Pretty.big_list "maps for type constructors:" |> Pretty.writeln end (* info about quotient types *) fun transform_pcr_info phi {pcrel_def, pcr_cr_eq} = {pcrel_def = Morphism.thm phi pcrel_def, pcr_cr_eq = Morphism.thm phi pcr_cr_eq} fun transform_quotient phi {quot_thm, pcr_info} = {quot_thm = Morphism.thm phi quot_thm, pcr_info = Option.map (transform_pcr_info phi) pcr_info} fun lookup_quotients ctxt type_name = Symtab.lookup (get_quotients ctxt) type_name |> Option.map (transform_quotient (Morphism.transfer_morphism' ctxt)) fun lookup_quot_thm_quotients ctxt quot_thm = let val (_, qtyp) = quot_thm_rty_qty quot_thm - val qty_full_name = (fst o dest_Type) qtyp + val qty_full_name = fst (dest_Type qtyp) fun compare_data (data:quotient) = Thm.eq_thm_prop (#quot_thm data, quot_thm) in case lookup_quotients ctxt qty_full_name of SOME quotient => if compare_data quotient then SOME quotient else NONE | NONE => NONE end fun update_quotients type_name qinfo context = let val qinfo' = transform_quotient Morphism.trim_context_morphism qinfo - in Data.map (map_quotients (Symtab.update (type_name, qinfo'))) context end + in (Data.map o map_quotients) (Symtab.update (type_name, qinfo')) context end fun delete_quotients quot_thm context = let val (_, qtyp) = quot_thm_rty_qty quot_thm - val qty_full_name = (fst o dest_Type) qtyp + val qty_full_name = fst (dest_Type qtyp) in if is_some (lookup_quot_thm_quotients (Context.proof_of context) quot_thm) - then Data.map (map_quotients (Symtab.delete qty_full_name)) context + then (Data.map o map_quotients) (Symtab.delete qty_full_name) context else context end fun print_quotients ctxt = let fun prt_quot (qty_name, {quot_thm, pcr_info}: quotient) = Pretty.block (separate (Pretty.brk 2) ([Pretty.str "type:", Pretty.str qty_name, Pretty.str "quot thm:", Thm.pretty_thm ctxt quot_thm] @ (case pcr_info of NONE => [] | SOME {pcrel_def, pcr_cr_eq, ...} => [Pretty.str "pcrel_def thm:", Thm.pretty_thm ctxt pcrel_def, Pretty.str "pcr_cr_eq thm:", Thm.pretty_thm ctxt pcr_cr_eq]))) in map prt_quot (Symtab.dest (get_quotients ctxt)) |> Pretty.big_list "quotients:" |> Pretty.writeln end val _ = Theory.setup (Attrib.setup \<^binding>\quot_del\ (Scan.succeed (Thm.declaration_attribute delete_quotients)) "deletes the Quotient theorem") (* data for restoring Transfer/Lifting context *) fun lookup_restore_data ctxt bundle_name = Symtab.lookup (get_restore_data ctxt) bundle_name fun update_restore_data bundle_name restore_data context = - Data.map (map_restore_data (Symtab.update (bundle_name, restore_data))) context + (Data.map o map_restore_data) (Symtab.update (bundle_name, restore_data)) context fun init_restore_data bundle_name qinfo context = update_restore_data bundle_name { quotient = qinfo, transfer_rules = Thm.item_net } context fun add_transfer_rules_in_restore_data bundle_name transfer_rules context = (case Symtab.lookup (get_restore_data' context) bundle_name of SOME restore_data => update_restore_data bundle_name { quotient = #quotient restore_data, transfer_rules = Item_Net.merge ((#transfer_rules restore_data), transfer_rules) } context | NONE => error ("The restore data " ^ quote bundle_name ^ " is not defined.")) (* theorems that a relator of an eq_onp is an eq_onp of the corresponding predicate *) fun get_relator_eq_onp_rules ctxt = map safe_mk_meta_eq (rev (Named_Theorems.get ctxt \<^named_theorems>\relator_eq_onp\)) (* info about reflexivity rules *) fun get_reflexivity_rules ctxt = Item_Net.content (get_reflexivity_rules' (Context.Proof ctxt)) |> map (Thm.transfer' ctxt) fun add_reflexivity_rule thm = - Data.map (map_reflexivity_rules (Item_Net.update (Thm.trim_context thm))) + (Data.map o map_reflexivity_rules) (Item_Net.update (Thm.trim_context thm)) val add_reflexivity_rule_attribute = Thm.declaration_attribute add_reflexivity_rule (* info about relator distributivity theorems *) fun map_relator_distr_data' f1 f2 f3 f4 {pos_mono_rule, neg_mono_rule, pos_distr_rules, neg_distr_rules} = {pos_mono_rule = f1 pos_mono_rule, neg_mono_rule = f2 neg_mono_rule, pos_distr_rules = f3 pos_distr_rules, neg_distr_rules = f4 neg_distr_rules} fun map_pos_mono_rule f = map_relator_distr_data' f I I I fun map_neg_mono_rule f = map_relator_distr_data' I f I I fun map_pos_distr_rules f = map_relator_distr_data' I I f I fun map_neg_distr_rules f = map_relator_distr_data' I I I f fun introduce_polarities rule = let val dest_less_eq = HOLogic.dest_bin \<^const_name>\less_eq\ dummyT val prems_pairs = map (dest_less_eq o HOLogic.dest_Trueprop) (Thm.prems_of rule) val equal_prems = filter op= prems_pairs val _ = if null equal_prems then () else error "The rule contains reflexive assumptions." val concl_pairs = rule |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_less_eq |> apply2 (snd o strip_comb) |> op ~~ |> filter_out op = val _ = if has_duplicates op= concl_pairs then error "The rule contains duplicated variables in the conlusion." else () fun rewrite_prem prem_pair = if member op= concl_pairs prem_pair then HOLogic.Trueprop_conv (Conv.rewr_conv (Thm.symmetric @{thm POS_def})) else if member op= concl_pairs (swap prem_pair) then HOLogic.Trueprop_conv (Conv.rewr_conv (Thm.symmetric @{thm NEG_def})) else error "The rule contains a non-relevant assumption." fun rewrite_prems [] = Conv.all_conv | rewrite_prems (x::xs) = Conv.implies_conv (rewrite_prem x) (rewrite_prems xs) val rewrite_prems_conv = rewrite_prems prems_pairs val rewrite_concl_conv = Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.rewr_conv (Thm.symmetric @{thm POS_def}))) in (Conv.fconv_rule (rewrite_prems_conv then_conv rewrite_concl_conv)) rule end handle TERM _ => error "The rule has a wrong format." | CTERM _ => error "The rule has a wrong format." fun negate_mono_rule mono_rule = let val rewr_conv = HOLogic.Trueprop_conv (Conv.rewrs_conv [@{thm POS_NEG}, @{thm NEG_POS}]) in Conv.fconv_rule (Conv.prems_conv ~1 rewr_conv then_conv Conv.concl_conv ~1 rewr_conv) mono_rule end; fun add_reflexivity_rules mono_rule context = let val ctxt = Context.proof_of context val thy = Context.theory_of context fun find_eq_rule thm = let - val concl_rhs = (hd o get_args 1 o HOLogic.dest_Trueprop o Thm.concl_of) thm + val concl_rhs = hd (get_args 1 (HOLogic.dest_Trueprop (Thm.concl_of thm))) val rules = Transfer.retrieve_relator_eq ctxt concl_rhs in find_first (fn th => Pattern.matches thy (concl_rhs, - (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.concl_of) th)) rules + fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.concl_of th))))) rules end val eq_rule = find_eq_rule mono_rule; val eq_rule = if is_some eq_rule then the eq_rule else error "No corresponding rule that the relator preserves equality was found." in context |> add_reflexivity_rule (Drule.zero_var_indexes (@{thm ord_le_eq_trans} OF [mono_rule, eq_rule])) |> add_reflexivity_rule (Drule.zero_var_indexes (@{thm ord_eq_le_trans} OF [sym OF [eq_rule], mono_rule])) end fun add_mono_rule mono_rule context = let val pol_mono_rule = introduce_polarities mono_rule - val mono_ruleT_name = (fst o dest_Type o fst o relation_types o fst o relation_types o snd o - dest_Const o head_of o HOLogic.dest_Trueprop o Thm.concl_of) pol_mono_rule + val mono_ruleT_name = + fst (dest_Type (fst (relation_types (fst (relation_types + (snd (dest_Const (head_of (HOLogic.dest_Trueprop (Thm.concl_of pol_mono_rule)))))))))) in if Symtab.defined (get_relator_distr_data' context) mono_ruleT_name then (if Context_Position.is_visible_generic context then warning ("Monotonicity rule for type " ^ quote mono_ruleT_name ^ " is already_defined.") else (); context) else let val neg_mono_rule = negate_mono_rule pol_mono_rule - val relator_distr_data = {pos_mono_rule = pol_mono_rule, neg_mono_rule = neg_mono_rule, - pos_distr_rules = [], neg_distr_rules = []} + val relator_distr_data = + {pos_mono_rule = Thm.trim_context pol_mono_rule, + neg_mono_rule = Thm.trim_context neg_mono_rule, + pos_distr_rules = [], + neg_distr_rules = []} in context - |> Data.map (map_relator_distr_data (Symtab.update (mono_ruleT_name, relator_distr_data))) + |> (Data.map o map_relator_distr_data) (Symtab.update (mono_ruleT_name, relator_distr_data)) |> add_reflexivity_rules mono_rule end end; local fun add_distr_rule update_entry distr_rule context = let - val distr_ruleT_name = (fst o dest_Type o fst o relation_types o fst o relation_types o snd o - dest_Const o head_of o HOLogic.dest_Trueprop o Thm.concl_of) distr_rule + val distr_ruleT_name = + fst (dest_Type (fst (relation_types (fst (relation_types + (snd (dest_Const (head_of (HOLogic.dest_Trueprop (Thm.concl_of distr_rule)))))))))) in if Symtab.defined (get_relator_distr_data' context) distr_ruleT_name then - Data.map (map_relator_distr_data (Symtab.map_entry distr_ruleT_name (update_entry distr_rule))) - context + (Data.map o map_relator_distr_data) + (Symtab.map_entry distr_ruleT_name (update_entry distr_rule)) context else error "The monotonicity rule is not defined." end fun rewrite_concl_conv thm ctm = Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.rewr_conv (Thm.symmetric thm))) ctm handle CTERM _ => error "The rule has a wrong format." in fun add_pos_distr_rule distr_rule context = let - val distr_rule = Conv.fconv_rule (rewrite_concl_conv @{thm POS_def}) distr_rule + val distr_rule' = Conv.fconv_rule (rewrite_concl_conv @{thm POS_def}) distr_rule fun update_entry distr_rule data = - map_pos_distr_rules (cons (@{thm POS_trans} OF [distr_rule, #pos_mono_rule data])) data + data |> (map_pos_distr_rules o cons) + (Thm.trim_context (@{thm POS_trans} OF + [distr_rule, Thm.transfer'' context (#pos_mono_rule data)])) in - add_distr_rule update_entry distr_rule context + add_distr_rule update_entry distr_rule' context end handle THM _ => error "Combining of the distr. rule and the monotonicity rule together has failed." fun add_neg_distr_rule distr_rule context = let - val distr_rule = Conv.fconv_rule (rewrite_concl_conv @{thm NEG_def}) distr_rule + val distr_rule' = Conv.fconv_rule (rewrite_concl_conv @{thm NEG_def}) distr_rule fun update_entry distr_rule data = - map_neg_distr_rules (cons (@{thm NEG_trans} OF [distr_rule, #neg_mono_rule data])) data + data |> (map_neg_distr_rules o cons) + (Thm.trim_context (@{thm NEG_trans} OF + [distr_rule, Thm.transfer'' context (#neg_mono_rule data)])) in - add_distr_rule update_entry distr_rule context + add_distr_rule update_entry distr_rule' context end handle THM _ => error "Combining of the distr. rule and the monotonicity rule together has failed." end local val eq_refl2 = sym RS @{thm eq_refl} in fun add_eq_distr_rule distr_rule context = let val pos_distr_rule = @{thm eq_refl} OF [distr_rule] val neg_distr_rule = eq_refl2 OF [distr_rule] in context |> add_pos_distr_rule pos_distr_rule |> add_neg_distr_rule neg_distr_rule end end; local fun sanity_check rule = let val assms = map (perhaps (try HOLogic.dest_Trueprop)) (Thm.prems_of rule) - val concl = (perhaps (try HOLogic.dest_Trueprop)) (Thm.concl_of rule); + val concl = perhaps (try HOLogic.dest_Trueprop) (Thm.concl_of rule); val (lhs, rhs) = (case concl of Const (\<^const_name>\less_eq\, _) $ (lhs as Const (\<^const_name>\relcompp\,_) $ _ $ _) $ rhs => (lhs, rhs) | Const (\<^const_name>\less_eq\, _) $ rhs $ (lhs as Const (\<^const_name>\relcompp\,_) $ _ $ _) => (lhs, rhs) | Const (\<^const_name>\HOL.eq\, _) $ (lhs as Const (\<^const_name>\relcompp\,_) $ _ $ _) $ rhs => (lhs, rhs) | _ => error "The rule has a wrong format.") val lhs_vars = Term.add_vars lhs [] val rhs_vars = Term.add_vars rhs [] val assms_vars = fold Term.add_vars assms []; val _ = if has_duplicates op= lhs_vars then error "Left-hand side has variable duplicates" else () val _ = if subset op= (rhs_vars, lhs_vars) then () else error "Extra variables in the right-hand side of the rule" val _ = if subset op= (assms_vars, lhs_vars) then () else error "Extra variables in the assumptions of the rule" val rhs_args = (snd o strip_comb) rhs; fun check_comp t = (case t of Const (\<^const_name>\relcompp\, _) $ Var _ $ Var _ => () | _ => error "There is an argument on the rhs that is not a composition.") val _ = map check_comp rhs_args in () end in fun add_distr_rule distr_rule context = let val _ = sanity_check distr_rule - val concl = (perhaps (try HOLogic.dest_Trueprop)) (Thm.concl_of distr_rule) + val concl = perhaps (try HOLogic.dest_Trueprop) (Thm.concl_of distr_rule) in (case concl of Const (\<^const_name>\less_eq\, _) $ (Const (\<^const_name>\relcompp\,_) $ _ $ _) $ _ => add_pos_distr_rule distr_rule context | Const (\<^const_name>\less_eq\, _) $ _ $ (Const (\<^const_name>\relcompp\,_) $ _ $ _) => add_neg_distr_rule distr_rule context | Const (\<^const_name>\HOL.eq\, _) $ (Const (\<^const_name>\relcompp\,_) $ _ $ _) $ _ => add_eq_distr_rule distr_rule context) end end fun get_distr_rules_raw context = Symtab.fold (fn (_, {pos_distr_rules, neg_distr_rules, ...}) => fn rules => pos_distr_rules @ neg_distr_rules @ rules) (get_relator_distr_data' context) [] |> map (Thm.transfer'' context) fun get_mono_rules_raw context = Symtab.fold (fn (_, {pos_mono_rule, neg_mono_rule, ...}) => fn rules => [pos_mono_rule, neg_mono_rule] @ rules) (get_relator_distr_data' context) [] |> map (Thm.transfer'' context) val lookup_relator_distr_data = Symtab.lookup o get_relator_distr_data val _ = Theory.setup (Attrib.setup \<^binding>\relator_mono\ (Scan.succeed (Thm.declaration_attribute add_mono_rule)) "declaration of relator's monotonicity" #> Attrib.setup \<^binding>\relator_distr\ (Scan.succeed (Thm.declaration_attribute add_distr_rule)) "declaration of relator's distributivity over OO" #> Global_Theory.add_thms_dynamic (\<^binding>\relator_distr_raw\, get_distr_rules_raw) #> Global_Theory.add_thms_dynamic (\<^binding>\relator_mono_raw\, get_mono_rules_raw)) (* no_code types *) fun add_no_code_type type_name context = Data.map (map_no_code_types (Symset.insert type_name)) context; -fun is_no_code_type context type_name = (Symset.member o get_no_code_types) context type_name +val is_no_code_type = Symset.member o get_no_code_types; (* setup fixed eq_onp rules *) val _ = Context.>> (fold (Named_Theorems.add_thm \<^named_theorems>\relator_eq_onp\ o Transfer.prep_transfer_domain_thm \<^context>) @{thms composed_equiv_rel_eq_onp composed_equiv_rel_eq_eq_onp}) (* setup fixed reflexivity rules *) val _ = Context.>> (fold add_reflexivity_rule @{thms order_refl[of "(=)"] eq_onp_le_eq Quotient_composition_le_eq Quotient_composition_ge_eq bi_unique_OO bi_total_OO right_unique_OO right_total_OO left_unique_OO left_total_OO}) (* outer syntax commands *) val _ = Outer_Syntax.command \<^command_keyword>\print_quot_maps\ "print quotient map functions" (Scan.succeed (Toplevel.keep (print_quot_maps o Toplevel.context_of))) val _ = Outer_Syntax.command \<^command_keyword>\print_quotients\ "print quotients" (Scan.succeed (Toplevel.keep (print_quotients o Toplevel.context_of))) end diff --git a/src/HOL/Tools/SMT/z3_replay_rules.ML b/src/HOL/Tools/SMT/z3_replay_rules.ML --- a/src/HOL/Tools/SMT/z3_replay_rules.ML +++ b/src/HOL/Tools/SMT/z3_replay_rules.ML @@ -1,53 +1,51 @@ (* Title: HOL/Tools/SMT/z3_replay_rules.ML Author: Sascha Boehme, TU Muenchen Custom rules for Z3 proof replay. *) signature Z3_REPLAY_RULES = sig val apply: Proof.context -> cterm -> thm option end; structure Z3_Replay_Rules: Z3_REPLAY_RULES = struct structure Data = Generic_Data ( type T = thm Net.net val empty = Net.empty val merge = Net.merge Thm.eq_thm ) fun maybe_instantiate ct thm = try Thm.first_order_match (Thm.cprop_of thm, ct) |> Option.map (fn inst => Thm.instantiate inst thm) fun apply ctxt ct = let val net = Data.get (Context.Proof ctxt) - val xthms = Net.match_term net (Thm.term_of ct) + val xthms = map (Thm.transfer' ctxt) (Net.match_term net (Thm.term_of ct)) fun select ct = map_filter (maybe_instantiate ct) xthms fun select' ct = let val thm = Thm.trivial ct in map_filter (try (fn rule => rule COMP thm)) xthms end in try hd (case select ct of [] => select' ct | xthms' => xthms') end val prep = `Thm.prop_of fun ins thm net = Net.insert_term Thm.eq_thm (prep thm) net handle Net.INSERT => net fun del thm net = Net.delete_term Thm.eq_thm (prep thm) net handle Net.DELETE => net -val add = Thm.declaration_attribute (Data.map o ins) +val add = Thm.declaration_attribute (Data.map o ins o Thm.trim_context) val del = Thm.declaration_attribute (Data.map o del) -val name = Binding.name "z3_rule" - -val description = "declaration of Z3 proof rules" - -val _ = Theory.setup (Attrib.setup name (Attrib.add_del add del) description #> - Global_Theory.add_thms_dynamic (name, Net.content o Data.get)) +val _ = Theory.setup + (Attrib.setup \<^binding>\z3_rule\ (Attrib.add_del add del) "declaration of Z3 proof rules" #> + Global_Theory.add_thms_dynamic (\<^binding>\z3_rule\, + fn context => map (Thm.transfer'' context) (Net.content (Data.get context)))) end; diff --git a/src/Pure/Concurrent/thread_position.ML b/src/Pure/Concurrent/thread_position.ML --- a/src/Pure/Concurrent/thread_position.ML +++ b/src/Pure/Concurrent/thread_position.ML @@ -1,27 +1,27 @@ (* Title: Pure/Concurrent/thread_position.ML Author: Makarius Thread-local position information. *) signature THREAD_POSITION = sig - type T = {line: int, offset: int, end_offset: int, props: {file: string, id: string}} + type T = {line: int, offset: int, end_offset: int, props: {label: string, file: string, id: string}} val none: T val get: unit -> T val setmp: T -> ('a -> 'b) -> 'a -> 'b end; structure Thread_Position: THREAD_POSITION = struct -type T = {line: int, offset: int, end_offset: int, props: {file: string, id: string}}; +type T = {line: int, offset: int, end_offset: int, props: {label: string, file: string, id: string}}; val var = Thread_Data.var () : T Thread_Data.var; -val none: T = {line = 0, offset = 0, end_offset = 0, props = {file = "", id = ""}}; +val none: T = {line = 0, offset = 0, end_offset = 0, props = {label = "", file = "", id = ""}}; fun get () = (case Thread_Data.get var of NONE => none | SOME pos => pos); fun setmp pos f x = Thread_Data.setmp var (if pos = none then NONE else SOME pos) f x; end; diff --git a/src/Pure/Concurrent/unsynchronized.ML b/src/Pure/Concurrent/unsynchronized.ML --- a/src/Pure/Concurrent/unsynchronized.ML +++ b/src/Pure/Concurrent/unsynchronized.ML @@ -1,64 +1,68 @@ (* Title: Pure/Concurrent/unsynchronized.ML Author: Makarius Raw ML references as unsynchronized state variables. *) signature UNSYNCHRONIZED = sig datatype ref = datatype ref val := : 'a ref * 'a -> unit val ! : 'a ref -> 'a val change: 'a ref -> ('a -> 'a) -> unit val change_result: 'a ref -> ('a -> 'b * 'a) -> 'b val inc: int ref -> int val dec: int ref -> int val add: int ref -> int -> int val setmp: 'a ref -> 'a -> ('b -> 'c) -> 'b -> 'c type 'a weak_ref = 'a ref option ref val weak_ref: 'a -> 'a weak_ref val weak_peek: 'a weak_ref -> 'a option + val weak_active: 'a weak_ref -> bool end; structure Unsynchronized: UNSYNCHRONIZED = struct (* regular references *) datatype ref = datatype ref; val op := = op :=; val ! = !; fun change r f = r := f (! r); fun change_result r f = let val (x, y) = f (! r) in r := y; x end; fun inc i = (i := ! i + (1: int); ! i); fun dec i = (i := ! i - (1: int); ! i); fun add i n = (i := ! i + (n: int); ! i); fun setmp flag value f x = Thread_Attributes.uninterruptible (fn restore_attributes => fn () => let val orig_value = ! flag; val _ = flag := value; val result = Exn.capture (restore_attributes f) x; val _ = flag := orig_value; in Exn.release result end) (); (* weak references *) (*see also $ML_SOURCES/basis/Weak.sml*) type 'a weak_ref = 'a ref option ref; fun weak_ref a = Weak.weak (SOME (ref a)); fun weak_peek (ref (SOME (ref a))) = SOME a | weak_peek _ = NONE; +fun weak_active (ref (SOME (ref _))) = true + | weak_active _ = false; + end; ML_Name_Space.forget_val "ref"; ML_Name_Space.forget_type "ref"; diff --git a/src/Pure/General/position.ML b/src/Pure/General/position.ML --- a/src/Pure/General/position.ML +++ b/src/Pure/General/position.ML @@ -1,316 +1,331 @@ (* Title: Pure/General/position.ML Author: Makarius Source positions starting from 1; values <= 0 mean "absent". Count Isabelle symbols, not UTF8 bytes nor UTF16 characters. Position range specifies a right-open interval offset .. end_offset (exclusive). *) signature POSITION = sig eqtype T val ord: T ord val make: Thread_Position.T -> T val dest: T -> Thread_Position.T val line_of: T -> int option val offset_of: T -> int option val end_offset_of: T -> int option + val label_of: T -> string option val file_of: T -> string option val id_of: T -> string option val symbol: Symbol.symbol -> T -> T val symbol_explode: string -> T -> T val distance_of: T * T -> int option val none: T val start: T + val label: string -> T -> T val file_only: string -> T val file: string -> T val line_file_only: int -> string -> T val line_file: int -> string -> T val line: int -> T val id: string -> T val id_only: string -> T val put_id: string -> T -> T val copy_id: T -> T -> T val parse_id: T -> int option val shift_offsets: {remove_id: bool} -> int -> T -> T val adjust_offsets: (int -> int option) -> T -> T val of_props: {line: int, offset: int, end_offset: int, props: Properties.T} -> T val of_properties: Properties.T -> T val properties_of: T -> Properties.T val def_properties_of: T -> Properties.T val entity_markup: string -> string * T -> Markup.T val make_entity_markup: {def: bool} -> serial -> string -> string * T -> Markup.T val markup: T -> Markup.T -> Markup.T val is_reported: T -> bool val is_reported_range: T -> bool val reported_text: T -> Markup.T -> string -> string val report_text: T -> Markup.T -> string -> unit val report: T -> Markup.T -> unit type report = T * Markup.T type report_text = report * string val reports_text: report_text list -> unit val reports: report list -> unit val store_reports: report_text list Unsynchronized.ref -> T list -> ('a -> Markup.T list) -> 'a -> unit val append_reports: report_text list Unsynchronized.ref -> report list -> unit val here_strs: T -> string * string val here: T -> string val here_list: T list -> string type range = T * T val no_range: range val no_range_position: T -> T val range_position: range -> T val range: T * T -> range val range_of_properties: Properties.T -> range val properties_of_range: range -> Properties.T val thread_data: unit -> T val setmp_thread_data: T -> ('a -> 'b) -> 'a -> 'b + val setmp_thread_data_label: string -> ('a -> 'b) -> 'a -> 'b val default: T -> bool * T end; structure Position: POSITION = struct (* datatype position *) datatype T = Pos of Thread_Position.T; val make = Pos; fun dest (Pos args) = args; val ord = pointer_eq_ord (int_ord o apply2 (#line o dest) ||| int_ord o apply2 (#offset o dest) ||| int_ord o apply2 (#end_offset o dest) ||| + fast_string_ord o apply2 (#label o #props o dest) ||| fast_string_ord o apply2 (#file o #props o dest) ||| fast_string_ord o apply2 (#id o #props o dest)); (* fields *) fun valid (i: int) = i > 0; fun maybe_valid i = if valid i then SOME i else NONE; val invalid = not o valid; fun invalid_pos (Pos {line, offset, ...}) = invalid line andalso invalid offset; val line_of = maybe_valid o #line o dest; val offset_of = maybe_valid o #offset o dest; val end_offset_of = maybe_valid o #end_offset o dest; +fun label_of (Pos {props = {label, ...}, ...}) = if label = "" then NONE else SOME label; fun file_of (Pos {props = {file, ...}, ...}) = if file = "" then NONE else SOME file; fun id_of (Pos {props = {id, ...}, ...}) = if id = "" then NONE else SOME id; (* symbol positions *) fun symbols [] pos = pos | symbols ss (pos as Pos {line, offset, end_offset, props}) = let val line' = if valid line then fold (fn s => s = "\n" ? Integer.add 1) ss line else line; val offset' = if valid offset then fold (fn s => Symbol.not_eof s ? Integer.add 1) ss offset else offset; in if line = line' andalso offset = offset' then pos else Pos {line = line', offset = offset', end_offset = end_offset, props = props} end; val symbol = symbols o single; fun symbol_explode str pos = if str = "" orelse invalid_pos pos then pos else fold symbol (Symbol.explode str) pos; (* distance of adjacent positions *) fun distance_of (Pos {offset, ...}, Pos {offset = offset', ...}) = if valid offset andalso valid offset' then SOME (offset' - offset) else NONE; (* make position *) -type props = {file: string, id: string}; +type props = {label: string, file: string, id: string}; -val no_props: props = {file = "", id = ""}; +val no_props: props = {label = "", file = "", id = ""}; fun file_props name : props = - if name = "" then no_props else {file = name, id = ""}; + if name = "" then no_props else {label = "", file = name, id = ""}; fun id_props id : props = - if id = "" then no_props else {file = "", id = id}; + if id = "" then no_props else {label = "", file = "", id = id}; val none = Pos {line = 0, offset = 0, end_offset = 0, props = no_props}; val start = Pos {line = 1, offset = 1, end_offset = 0, props = no_props}; +fun label label (Pos {line, offset, end_offset, props = {label = _, file, id}}) = + Pos {line = line, offset = offset, end_offset = end_offset, + props = {label = label, file = file, id = id}}; fun file_only name = Pos {line = 0, offset = 0, end_offset = 0, props = file_props name}; fun file name = Pos {line = 1, offset = 1, end_offset = 0, props = file_props name}; fun line_file_only line name = Pos {line = line, offset = 0, end_offset = 0, props = file_props name}; fun line_file line name = Pos {line = line, offset = 1, end_offset = 0, props = file_props name}; fun line line = line_file line ""; fun id id = Pos {line = 0, offset = 1, end_offset = 0, props = id_props id}; fun id_only id = Pos {line = 0, offset = 0, end_offset = 0, props = id_props id}; -fun put_id id' (pos as Pos {line, offset, end_offset, props = {file, id}}) = +fun put_id id' (pos as Pos {line, offset, end_offset, props = {label, file, id}}) = if id = id' then pos - else Pos {line = line, offset = offset, end_offset = end_offset, props = {file = file, id = id'}}; + else Pos {line = line, offset = offset, end_offset = end_offset, + props = {label = label, file = file, id = id'}}; fun copy_id pos = (case id_of pos of NONE => I | SOME id => put_id id); fun parse_id pos = Option.map Value.parse_int (id_of pos); (* adjust offsets *) fun shift_offsets {remove_id} shift (pos as Pos {line, offset, end_offset, props}) = if shift < 0 then raise Fail "Illegal offset shift" else if shift > 0 andalso valid line then raise Fail "Illegal line position" else let val offset' = if valid offset then offset + shift else offset; val end_offset' = if valid end_offset then end_offset + shift else end_offset; val props' = if remove_id then file_props (#file props) else props; in if offset = offset' andalso end_offset = end_offset' andalso props = props' then pos else Pos {line = line, offset = offset', end_offset = end_offset', props = props'} end; fun adjust_offsets adjust pos = if is_none (file_of pos) then (case parse_id pos of SOME id => (case adjust id of SOME shift => shift_offsets {remove_id = true} shift pos | NONE => pos) | NONE => pos) else pos; (* markup properties *) fun of_props {line, offset, end_offset, props} = Pos {line = line, offset = offset, end_offset = end_offset, props = - {file = Properties.get_string props Markup.fileN, + {label = Properties.get_string props Markup.labelN, + file = Properties.get_string props Markup.fileN, id = Properties.get_string props Markup.idN}}; fun of_properties props = of_props { line = Properties.get_int props Markup.lineN, offset = Properties.get_int props Markup.offsetN, end_offset = Properties.get_int props Markup.end_offsetN, props = props}; -fun properties_of (Pos {line, offset, end_offset, props = {file, id}}) = +fun properties_of (Pos {line, offset, end_offset, props = {label, file, id}}) = Properties.make_int Markup.lineN line @ Properties.make_int Markup.offsetN offset @ Properties.make_int Markup.end_offsetN end_offset @ + Properties.make_string Markup.labelN label @ Properties.make_string Markup.fileN file @ Properties.make_string Markup.idN id; val def_properties_of = properties_of #> map (apfst Markup.def_name); fun entity_markup kind (name, pos) = Markup.entity kind name |> Markup.properties (def_properties_of pos); fun make_entity_markup {def} serial kind (name, pos) = let val props = if def then (Markup.defN, Value.print_int serial) :: properties_of pos else (Markup.refN, Value.print_int serial) :: def_properties_of pos; in Markup.entity kind name |> Markup.properties props end; val markup = Markup.properties o properties_of; (* reports *) fun is_reported pos = is_some (offset_of pos) andalso is_some (id_of pos); fun is_reported_range pos = is_reported pos andalso is_some (end_offset_of pos); fun reported_text pos m txt = if is_reported pos then Markup.markup (markup pos m) txt else ""; fun report_text pos markup txt = Output.report [reported_text pos markup txt]; fun report pos markup = report_text pos markup ""; type report = T * Markup.T; type report_text = report * string; val reports_text = map (fn ((pos, m), txt) => if is_reported pos then Markup.markup (markup pos m) txt else "") #> Output.report; val reports = map (rpair "") #> reports_text; fun store_reports _ [] _ _ = () | store_reports (r: report_text list Unsynchronized.ref) ps markup x = let val ms = markup x in Unsynchronized.change r (fold (fn p => fold (fn m => cons ((p, m), "")) ms) ps) end; fun append_reports (r: report_text list Unsynchronized.ref) reports = Unsynchronized.change r (append (map (rpair "") reports)); (* here: user output *) fun here_strs pos = (case (line_of pos, file_of pos) of (SOME i, NONE) => (" ", "(line " ^ Value.print_int i ^ ")") | (SOME i, SOME name) => (" ", "(line " ^ Value.print_int i ^ " of " ^ quote name ^ ")") | (NONE, SOME name) => (" ", "(file " ^ quote name ^ ")") | _ => if is_reported pos then ("", "\092<^here>") else ("", "")); fun here pos = let val props = properties_of pos; val (s1, s2) = here_strs pos; in if s2 = "" then "" else s1 ^ Markup.markup (Markup.properties props Markup.position) s2 end; val here_list = map here #> distinct (op =) #> implode; (* range *) type range = T * T; val no_range = (none, none); fun no_range_position (Pos {line, offset, end_offset = _, props}) = Pos {line = line, offset = offset, end_offset = 0, props = props}; fun range_position (Pos {line, offset, end_offset = _, props}, Pos {offset = offset', ...}) = Pos {line = line, offset = offset, end_offset = offset', props = props}; fun range (pos, pos') = (range_position (pos, pos'), no_range_position pos'); fun range_of_properties props = let val pos = of_properties props; val pos' = of_props {line = Properties.get_int props Markup.end_lineN, offset = Properties.get_int props Markup.end_offsetN, end_offset = 0, props = props}; in (pos, pos') end; fun properties_of_range (pos, pos') = properties_of pos @ Properties.make_int Markup.end_lineN (the_default 0 (line_of pos')); (* thread data *) val thread_data = Pos o Thread_Position.get; fun setmp_thread_data pos = Thread_Position.setmp (dest pos); +fun setmp_thread_data_label a f x = + if a = "" then f x + else setmp_thread_data (label a (thread_data ())) f x; + fun default pos = if pos = none then (false, thread_data ()) else (true, pos); end; diff --git a/src/Pure/General/position.scala b/src/Pure/General/position.scala --- a/src/Pure/General/position.scala +++ b/src/Pure/General/position.scala @@ -1,162 +1,164 @@ /* Title: Pure/General/position.scala Author: Makarius Position properties. */ package isabelle object Position { type T = Properties.T val none: T = Nil val Line = new Properties.Int(Markup.LINE) val Offset = new Properties.Int(Markup.OFFSET) val End_Offset = new Properties.Int(Markup.END_OFFSET) + val Label = new Properties.String(Markup.LABEL) val File = new Properties.String(Markup.FILE) val Id = new Properties.Long(Markup.ID) val Id_String = new Properties.String(Markup.ID) val Def_Line = new Properties.Int(Markup.DEF_LINE) val Def_Offset = new Properties.Int(Markup.DEF_OFFSET) val Def_End_Offset = new Properties.Int(Markup.DEF_END_OFFSET) + val Def_Label = new Properties.String(Markup.DEF_LABEL) val Def_File = new Properties.String(Markup.DEF_FILE) val Def_Id = new Properties.Long(Markup.DEF_ID) object Line_File { def apply(line: Int, file: String): T = (if (line > 0) Line(line) else Nil) ::: (if (file != "") File(file) else Nil) def unapply(pos: T): Option[(Int, String)] = (pos, pos) match { case (Line(i), File(name)) => Some((i, name)) case (_, File(name)) => Some((1, name)) case _ => None } } object Def_Line_File { def unapply(pos: T): Option[(Int, String)] = (pos, pos) match { case (Def_Line(i), Def_File(name)) => Some((i, name)) case (_, Def_File(name)) => Some((1, name)) case _ => None } } object Range { def apply(range: Symbol.Range): T = Offset(range.start) ::: End_Offset(range.stop) def unapply(pos: T): Option[Symbol.Range] = (pos, pos) match { case (Offset(start), End_Offset(stop)) if start <= stop => Some(Text.Range(start, stop)) case (Offset(start), _) => Some(Text.Range(start, start + 1)) case _ => None } } object Def_Range { def apply(range: Symbol.Range): T = Def_Offset(range.start) ::: Def_End_Offset(range.stop) def unapply(pos: T): Option[Symbol.Range] = (pos, pos) match { case (Def_Offset(start), Def_End_Offset(stop)) if start <= stop => Some(Text.Range(start, stop)) case (Def_Offset(start), _) => Some(Text.Range(start, start + 1)) case _ => None } } object Item_Id { def unapply(pos: T): Option[(Long, Symbol.Range)] = pos match { case Id(id) => val offset = Offset.unapply(pos) getOrElse 0 val end_offset = End_Offset.unapply(pos) getOrElse (offset + 1) Some((id, Text.Range(offset, end_offset))) case _ => None } } object Item_Def_Id { def unapply(pos: T): Option[(Long, Symbol.Range)] = pos match { case Def_Id(id) => val offset = Def_Offset.unapply(pos) getOrElse 0 val end_offset = Def_End_Offset.unapply(pos) getOrElse (offset + 1) Some((id, Text.Range(offset, end_offset))) case _ => None } } object Item_File { def unapply(pos: T): Option[(String, Int, Symbol.Range)] = pos match { case Line_File(line, name) => val offset = Offset.unapply(pos) getOrElse 0 val end_offset = End_Offset.unapply(pos) getOrElse (offset + 1) Some((name, line, Text.Range(offset, end_offset))) case _ => None } } object Item_Def_File { def unapply(pos: T): Option[(String, Int, Symbol.Range)] = pos match { case Def_Line_File(line, name) => val offset = Def_Offset.unapply(pos) getOrElse 0 val end_offset = Def_End_Offset.unapply(pos) getOrElse (offset + 1) Some((name, line, Text.Range(offset, end_offset))) case _ => None } } /* here: user output */ def here(props: T, delimited: Boolean = true): String = { val pos = props.filter(Markup.position_property) if_proper(pos, { val s0 = (Line.unapply(pos), File.unapply(pos)) match { case (Some(i), None) => "line " + i.toString case (Some(i), Some(name)) => "line " + i.toString + " of " + quote(name) case (None, Some(name)) => "file " + quote(name) case _ => "" } val s = if (s0 == "") s0 else if (delimited) " (" + s0 + ")" else s0 Markup(Markup.POSITION, pos).markup(s) }) } /* JSON representation */ object JSON { def apply(pos: T): isabelle.JSON.Object.T = isabelle.JSON.Object.empty ++ Line.unapply(pos).map(Line.name -> _) ++ Offset.unapply(pos).map(Offset.name -> _) ++ End_Offset.unapply(pos).map(End_Offset.name -> _) ++ File.unapply(pos).map(File.name -> _) ++ Id.unapply(pos).map(Id.name -> _) def unapply(json: isabelle.JSON.T): Option[T] = for { line <- isabelle.JSON.int_default(json, Line.name) offset <- isabelle.JSON.int_default(json, Offset.name) end_offset <- isabelle.JSON.int_default(json, End_Offset.name) file <- isabelle.JSON.string_default(json, File.name) id <- isabelle.JSON.long_default(json, Id.name) } yield { def defined(name: String): Boolean = isabelle.JSON.value(json, name).isDefined (if (defined(Line.name)) Line(line) else Nil) ::: (if (defined(Offset.name)) Offset(offset) else Nil) ::: (if (defined(End_Offset.name)) End_Offset(end_offset) else Nil) ::: (if (defined(File.name)) File(file) else Nil) ::: (if (defined(Id.name)) Id(id) else Nil) } } } diff --git a/src/Pure/Isar/attrib.ML b/src/Pure/Isar/attrib.ML --- a/src/Pure/Isar/attrib.ML +++ b/src/Pure/Isar/attrib.ML @@ -1,640 +1,646 @@ (* Title: Pure/Isar/attrib.ML Author: Markus Wenzel, TU Muenchen Symbolic representation of attributes -- with name and syntax. *) signature ATTRIB = sig type thms = Attrib.thms type fact = Attrib.fact val print_attributes: bool -> Proof.context -> unit val attribute_space: Context.generic -> Name_Space.T val define_global: binding -> (Token.src -> attribute) -> string -> theory -> string * theory val define: binding -> (Token.src -> attribute) -> string -> local_theory -> string * local_theory val check_name_generic: Context.generic -> xstring * Position.T -> string val check_name: Proof.context -> xstring * Position.T -> string val check_src: Proof.context -> Token.src -> Token.src val attribs: Token.src list context_parser val opt_attribs: Token.src list context_parser val pretty_attribs: Proof.context -> Token.src list -> Pretty.T list val pretty_binding: Proof.context -> Attrib.binding -> string -> Pretty.T list val attribute: Proof.context -> Token.src -> attribute val attribute_global: theory -> Token.src -> attribute val attribute_cmd: Proof.context -> Token.src -> attribute val attribute_cmd_global: theory -> Token.src -> attribute val map_specs: ('a list -> 'att list) -> (('c * 'a list) * 'b) list -> (('c * 'att list) * 'b) list val map_facts: ('a list -> 'att list) -> (('c * 'a list) * ('d * 'a list) list) list -> (('c * 'att list) * ('d * 'att list) list) list val map_facts_refs: ('a list -> 'att list) -> ('b -> 'fact) -> (('c * 'a list) * ('b * 'a list) list) list -> (('c * 'att list) * ('fact * 'att list) list) list val trim_context_binding: Attrib.binding -> Attrib.binding val trim_context_thms: thms -> thms val trim_context_fact: fact -> fact val global_notes: string -> fact list -> theory -> (string * thm list) list * theory val local_notes: string -> fact list -> Proof.context -> (string * thm list) list * Proof.context val generic_notes: string -> fact list -> Context.generic -> (string * thm list) list * Context.generic val lazy_notes: string -> binding * thm list lazy -> Context.generic -> Context.generic val eval_thms: Proof.context -> (Facts.ref * Token.src list) list -> thm list val attribute_syntax: attribute context_parser -> Token.src -> attribute val setup: binding -> attribute context_parser -> string -> theory -> theory val local_setup: binding -> attribute context_parser -> string -> local_theory -> local_theory val attribute_setup: bstring * Position.T -> Input.source -> string -> local_theory -> local_theory val internal: (morphism -> attribute) -> Token.src val internal_declaration: declaration -> thms val add_del: attribute -> attribute -> attribute context_parser val thm: thm context_parser val thms: thm list context_parser val multi_thm: thm list context_parser val transform_facts: morphism -> fact list -> fact list val partial_evaluation: Proof.context -> fact list -> fact list val print_options: bool -> Proof.context -> unit val config_bool: binding -> (Context.generic -> bool) -> bool Config.T * (theory -> theory) val config_int: binding -> (Context.generic -> int) -> int Config.T * (theory -> theory) val config_real: binding -> (Context.generic -> real) -> real Config.T * (theory -> theory) val config_string: binding -> (Context.generic -> string) -> string Config.T * (theory -> theory) val setup_config_bool: binding -> (Context.generic -> bool) -> bool Config.T val setup_config_int: binding -> (Context.generic -> int) -> int Config.T val setup_config_real: binding -> (Context.generic -> real) -> real Config.T val setup_config_string: binding -> (Context.generic -> string) -> string Config.T val option_bool: string * Position.T -> bool Config.T * (theory -> theory) val option_int: string * Position.T -> int Config.T * (theory -> theory) val option_real: string * Position.T -> real Config.T * (theory -> theory) val option_string: string * Position.T -> string Config.T * (theory -> theory) val setup_option_bool: string * Position.T -> bool Config.T val setup_option_int: string * Position.T -> int Config.T val setup_option_real: string * Position.T -> real Config.T val setup_option_string: string * Position.T -> string Config.T val consumes: int -> Token.src val constraints: int -> Token.src val cases_open: Token.src val case_names: string list -> Token.src val case_conclusion: string * string list -> Token.src end; structure Attrib: sig type binding = Attrib.binding include ATTRIB end = struct open Attrib; (** named attributes **) (* theory data *) structure Attributes = Generic_Data ( type T = ((Token.src -> attribute) * string) Name_Space.table; val empty : T = Name_Space.empty_table Markup.attributeN; fun merge data : T = Name_Space.merge_tables data; ); val ops_attributes = {get_data = Attributes.get, put_data = Attributes.put}; val get_attributes = Attributes.get o Context.Proof; fun print_attributes verbose ctxt = let val attribs = get_attributes ctxt; fun prt_attr (name, (_, "")) = Pretty.mark_str name | prt_attr (name, (_, comment)) = Pretty.block (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); in [Pretty.big_list "attributes:" (map prt_attr (Name_Space.markup_table verbose ctxt attribs))] |> Pretty.writeln_chunks end; val attribute_space = Name_Space.space_of_table o Attributes.get; (* define *) fun define_global binding att comment = Entity.define_global ops_attributes binding (att, comment); fun define binding att comment = Entity.define ops_attributes binding (att, comment); (* check *) fun check_name_generic context = #1 o Name_Space.check context (Attributes.get context); val check_name = check_name_generic o Context.Proof; fun check_src ctxt src = let val _ = if Token.checked_src src then () else Context_Position.report ctxt (#1 (Token.range_of src)) Markup.language_attribute; in #1 (Token.check_src ctxt get_attributes src) end; val attribs = Args.context -- Scan.lift Parse.attribs >> (fn (ctxt, srcs) => map (check_src ctxt) srcs); val opt_attribs = Scan.optional attribs []; (* pretty printing *) fun pretty_attribs _ [] = [] | pretty_attribs ctxt srcs = [Pretty.enum "," "[" "]" (map (Token.pretty_src ctxt) srcs)]; fun pretty_binding ctxt (b, atts) sep = (case (Binding.is_empty b, null atts) of (true, true) => [] | (false, true) => [Pretty.block [Binding.pretty b, Pretty.str sep]] | (true, false) => [Pretty.block (pretty_attribs ctxt atts @ [Pretty.str sep])] | (false, false) => [Pretty.block (Binding.pretty b :: Pretty.brk 1 :: pretty_attribs ctxt atts @ [Pretty.str sep])]); (* get attributes *) fun attribute_generic context = - let val table = Attributes.get context - in fn src => #1 (Name_Space.get table (#1 (Token.name_of_src src))) src end; + let val table = Attributes.get context in + fn src => + let + val name = #1 (Token.name_of_src src); + val label = Long_Name.qualify Markup.attributeN name; + val att = #1 (Name_Space.get table name) src; + in Position.setmp_thread_data_label label att end + end; val attribute = attribute_generic o Context.Proof; val attribute_global = attribute_generic o Context.Theory; fun attribute_cmd ctxt = attribute ctxt o check_src ctxt; fun attribute_cmd_global thy = attribute_global thy o check_src (Proof_Context.init_global thy); (* attributed declarations *) fun map_specs f = map (apfst (apsnd f)); fun map_facts f = map (apfst (apsnd f) o apsnd (map (apsnd f))); fun map_facts_refs f g = map_facts f #> map (apsnd (map (apfst g))); (* fact expressions *) val trim_context_binding: Attrib.binding -> Attrib.binding = apsnd (map Token.trim_context_src); val trim_context_thms: thms -> thms = (map o apfst o map) Thm.trim_context; fun trim_context_fact (binding, thms) = (trim_context_binding binding, trim_context_thms thms); fun global_notes kind facts thy = thy |> Global_Theory.note_thmss kind (map_facts (map (attribute_global thy)) facts); fun local_notes kind facts ctxt = ctxt |> Proof_Context.note_thmss kind (map_facts (map (attribute ctxt)) facts); fun generic_notes kind facts context = context |> Context.mapping_result (global_notes kind facts) (local_notes kind facts); fun lazy_notes kind arg = Context.mapping (Global_Theory.add_thms_lazy kind arg) (Proof_Context.add_thms_lazy kind arg); fun eval_thms ctxt srcs = ctxt |> Proof_Context.note_thmss "" (map_facts_refs (map (attribute_cmd ctxt)) (Proof_Context.get_fact ctxt) [(Binding.empty_atts, srcs)]) |> fst |> maps snd; (* attribute setup *) fun attribute_syntax scan src (context, th) = let val (a, context') = Token.syntax_generic scan src context in a (context', th) end; fun setup binding scan comment = define_global binding (attribute_syntax scan) comment #> snd; fun local_setup binding scan comment = define binding (attribute_syntax scan) comment #> snd; fun attribute_setup binding source comment = ML_Context.expression (Input.pos_of source) (ML_Lex.read ("Theory.local_setup (Attrib.local_setup (" ^ ML_Syntax.make_binding binding ^ ") (") @ ML_Lex.read_source source @ ML_Lex.read (") " ^ ML_Syntax.print_string comment ^ ")")) |> Context.proof_map; (* internal attribute *) val _ = Theory.setup (setup (Binding.make ("attribute", \<^here>)) (Scan.lift Args.internal_attribute >> Morphism.form) "internal attribute"); fun internal_name ctxt name = Token.make_src (name, Position.none) [] |> check_src ctxt |> hd; val internal_attribute_name = internal_name (Context.the_local_context ()) "attribute"; fun internal att = internal_attribute_name :: [Token.make_string ("", Position.none) |> Token.assign (SOME (Token.Attribute att))]; fun internal_declaration decl = [([Drule.dummy_thm], [internal (fn phi => Thm.declaration_attribute (K (decl phi)))])]; (* add/del syntax *) fun add_del add del = Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add); (** parsing attributed theorems **) local val fact_name = Parse.position Args.internal_fact >> (fn (_, pos) => ("", pos)) || Args.name_position; fun gen_thm pick = Scan.depend (fn context => let val get = Proof_Context.get_fact_generic context; val get_fact = get o Facts.Fact; fun get_named is_sel pos name = let val (a, ths) = get (Facts.Named ((name, pos), NONE)) in (if is_sel then NONE else a, ths) end; in Parse.$$$ "[" |-- Scan.pass context attribs --| Parse.$$$ "]" >> (fn srcs => let val atts = map (attribute_generic context) srcs; val (th', context') = fold (uncurry o Thm.apply_attribute) atts (Drule.dummy_thm, context); in (context', pick ("", Position.none) [th']) end) || (Scan.ahead Args.alt_name -- Args.named_fact get_fact >> (fn (s, fact) => ("", Facts.Fact s, fact)) || Scan.ahead (fact_name -- Scan.option Parse.thm_sel) :|-- (fn ((name, pos), sel) => Args.named_fact (get_named (is_some sel) pos) --| Scan.option Parse.thm_sel >> (fn fact => (name, Facts.Named ((name, pos), sel), fact)))) -- Scan.pass context opt_attribs >> (fn ((name, thmref, fact), srcs) => let val ths = Facts.select thmref fact; val atts = map (attribute_generic context) srcs; val (ths', context') = fold_map (curry (fold (uncurry o Thm.apply_attribute) atts)) ths context; in (context', pick (name, Facts.ref_pos thmref) ths') end) end); in val thm = gen_thm Facts.the_single; val multi_thm = gen_thm (K I); val thms = Scan.repeats multi_thm; end; (* transform fact expressions *) fun transform_facts phi = map (fn ((a, atts), bs) => ((Morphism.binding phi a, (map o map) (Token.transform phi) atts), bs |> map (fn (ths, btts) => (Morphism.fact phi ths, (map o map) (Token.transform phi) btts)))); (** partial evaluation -- observing rule/declaration/mixed attributes **) (*NB: result length may change due to rearrangement of symbolic expression*) local fun apply_att src (context, th) = let val src1 = map Token.init_assignable src; val result = attribute_generic context src1 (context, th); val src2 = map Token.closure src1; in (src2, result) end; fun err msg src = let val (name, pos) = Token.name_of_src src in error (msg ^ " " ^ quote name ^ Position.here pos) end; fun eval src ((th, dyn), (decls, context)) = (case (apply_att src (context, th), dyn) of ((_, (NONE, SOME th')), NONE) => ((th', NONE), (decls, context)) | ((_, (NONE, SOME _)), SOME _) => err "Mixed dynamic attribute followed by static rule" src | ((src', (SOME context', NONE)), NONE) => let val decls' = (case decls of [] => [(th, [src'])] | (th2, srcs2) :: rest => if Thm.eq_thm_strict (th, th2) then ((th2, src' :: srcs2) :: rest) else (th, [src']) :: (th2, srcs2) :: rest); in ((th, NONE), (decls', context')) end | ((src', (opt_context', opt_th')), _) => let val context' = the_default context opt_context'; val th' = the_default th opt_th'; val dyn' = (case dyn of NONE => SOME (th, [src']) | SOME (dyn_th, srcs) => SOME (dyn_th, src' :: srcs)); in ((th', dyn'), (decls, context')) end); in fun partial_evaluation ctxt facts = (facts, Context.Proof (Context_Position.not_really ctxt)) |-> fold_map (fn ((b, more_atts), fact) => fn context => let val (fact', (decls, context')) = (fact, ([], context)) |-> fold_map (fn (ths, atts) => fn res1 => (ths, res1) |-> fold_map (fn th => fn res2 => let val ((th', dyn'), res3) = fold eval (atts @ more_atts) ((th, NONE), res2); val th_atts' = (case dyn' of NONE => (th', []) | SOME (dyn_th', atts') => (dyn_th', rev atts')); in (th_atts', res3) end)) |>> flat; val decls' = rev (map (apsnd rev) decls); val facts' = if eq_list (eq_fst Thm.eq_thm_strict) (decls', fact') then [((b, []), map2 (fn (th, atts1) => fn (_, atts2) => (th, atts1 @ atts2)) decls' fact')] else if null decls' then [((b, []), fact')] else [(Binding.empty_atts, decls'), ((b, []), fact')]; in (facts', context') end) |> fst |> flat |> map (apsnd (map (apfst single))) |> filter_out (fn (b, fact) => Binding.is_empty_atts b andalso forall (null o #2) fact); end; (** configuration options **) (* naming *) structure Configs = Theory_Data ( type T = Config.value Config.T Symtab.table; val empty = Symtab.empty; fun merge data = Symtab.merge (K true) data; ); fun print_options verbose ctxt = let fun prt (name, config) = let val value = Config.get ctxt config in Pretty.block [Pretty.mark_str name, Pretty.str (": " ^ Config.print_type value ^ " ="), Pretty.brk 1, Pretty.str (Config.print_value value)] end; val space = attribute_space (Context.Proof ctxt); val configs = Name_Space.markup_entries verbose ctxt space (Symtab.dest (Configs.get (Proof_Context.theory_of ctxt))); in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end; (* concrete syntax *) local val equals = Parse.$$$ "="; fun scan_value (Config.Bool _) = equals -- Args.$$$ "false" >> K (Config.Bool false) || equals -- Args.$$$ "true" >> K (Config.Bool true) || Scan.succeed (Config.Bool true) | scan_value (Config.Int _) = equals |-- Parse.int >> Config.Int | scan_value (Config.Real _) = equals |-- Parse.real >> Config.Real | scan_value (Config.String _) = equals |-- Args.name >> Config.String; fun scan_config thy config = let val config_type = Config.get_global thy config in scan_value config_type >> (K o Thm.declaration_attribute o K o Config.put_generic config) end; fun register binding config thy = let val name = Sign.full_name thy binding in thy |> setup binding (Scan.lift (scan_config thy config) >> Morphism.form) "configuration option" |> Configs.map (Symtab.update (name, config)) end; fun declare make coerce binding default = let val name = Binding.name_of binding; val pos = Binding.pos_of binding; val config_value = Config.declare (name, pos) (make o default); val config = coerce config_value; in (config, register binding config_value) end; in fun register_config config = register (Binding.make (Config.name_of config, Config.pos_of config)) config; val register_config_bool = register_config o Config.bool_value; val register_config_int = register_config o Config.int_value; val register_config_real = register_config o Config.real_value; val register_config_string = register_config o Config.string_value; val config_bool = declare Config.Bool Config.bool; val config_int = declare Config.Int Config.int; val config_real = declare Config.Real Config.real; val config_string = declare Config.String Config.string; end; (* implicit setup *) local fun setup_config declare_config binding default = let val (config, setup) = declare_config binding default; val _ = Theory.setup setup; in config end; in val setup_config_bool = setup_config config_bool; val setup_config_int = setup_config config_int; val setup_config_string = setup_config config_string; val setup_config_real = setup_config config_real; end; (* system options *) local fun declare_option coerce (name, pos) = let val config = Config.declare_option (name, pos); in (coerce config, register_config config) end; fun setup_option coerce (name, pos) = let val config = Config.declare_option (name, pos); val _ = Theory.setup (register_config config); in coerce config end; in val option_bool = declare_option Config.bool; val option_int = declare_option Config.int; val option_real = declare_option Config.real; val option_string = declare_option Config.string; val setup_option_bool = setup_option Config.bool; val setup_option_int = setup_option Config.int; val setup_option_real = setup_option Config.real; val setup_option_string = setup_option Config.string; end; (* theory setup *) val _ = Theory.setup (setup \<^binding>\tagged\ (Scan.lift (Args.name -- Args.name) >> Thm.tag) "tagged theorem" #> setup \<^binding>\untagged\ (Scan.lift Args.name >> Thm.untag) "untagged theorem" #> setup \<^binding>\kind\ (Scan.lift Args.name >> Thm.kind) "theorem kind" #> setup \<^binding>\THEN\ (Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- thm >> (fn (i, B) => Thm.rule_attribute [B] (fn _ => fn A => A RSN (i, B)))) "resolution with rule" #> setup \<^binding>\OF\ (thms >> (fn Bs => Thm.rule_attribute Bs (fn _ => fn A => A OF Bs))) "rule resolved with facts" #> setup \<^binding>\rename_abs\ (Scan.lift (Scan.repeat (Args.maybe Args.name)) >> (fn vs => Thm.rule_attribute [] (K (Drule.rename_bvars' vs)))) "rename bound variables in abstractions" #> setup \<^binding>\unfolded\ (thms >> (fn ths => Thm.rule_attribute ths (fn context => Local_Defs.unfold (Context.proof_of context) ths))) "unfolded definitions" #> setup \<^binding>\folded\ (thms >> (fn ths => Thm.rule_attribute ths (fn context => Local_Defs.fold (Context.proof_of context) ths))) "folded definitions" #> setup \<^binding>\consumes\ (Scan.lift (Scan.optional Parse.int 1) >> Rule_Cases.consumes) "number of consumed facts" #> setup \<^binding>\constraints\ (Scan.lift Parse.nat >> Rule_Cases.constraints) "number of equality constraints" #> setup \<^binding>\cases_open\ (Scan.succeed Rule_Cases.cases_open) "rule with open parameters" #> setup \<^binding>\case_names\ (Scan.lift (Scan.repeat (Args.name -- Scan.optional (Parse.$$$ "[" |-- Scan.repeat1 (Args.maybe Args.name) --| Parse.$$$ "]") [])) >> (fn cs => Rule_Cases.cases_hyp_names (map #1 cs) (map (map (the_default Rule_Cases.case_hypsN) o #2) cs))) "named rule cases" #> setup \<^binding>\case_conclusion\ (Scan.lift (Args.name -- Scan.repeat Args.name) >> Rule_Cases.case_conclusion) "named conclusion of rule cases" #> setup \<^binding>\params\ (Scan.lift (Parse.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params) "named rule parameters" #> setup \<^binding>\rule_format\ (Scan.lift (Args.mode "no_asm") >> (fn true => Object_Logic.rule_format_no_asm | false => Object_Logic.rule_format)) "result put into canonical rule format" #> setup \<^binding>\elim_format\ (Scan.succeed (Thm.rule_attribute [] (K Tactic.make_elim))) "destruct rule turned into elimination rule format" #> setup \<^binding>\no_vars\ (Scan.succeed (Thm.rule_attribute [] (Variable.import_vars o Context.proof_of))) "imported schematic variables" #> setup \<^binding>\atomize\ (Scan.succeed Object_Logic.declare_atomize) "declaration of atomize rule" #> setup \<^binding>\rulify\ (Scan.succeed Object_Logic.declare_rulify) "declaration of rulify rule" #> setup \<^binding>\rotated\ (Scan.lift (Scan.optional Parse.int 1 >> (fn n => Thm.rule_attribute [] (fn _ => rotate_prems n)))) "rotated theorem premises" #> setup \<^binding>\defn\ (add_del Local_Defs.defn_add Local_Defs.defn_del) "declaration of definitional transformations" #> setup \<^binding>\abs_def\ (Scan.succeed (Thm.rule_attribute [] (Local_Defs.abs_def_rule o Context.proof_of))) "abstract over free variables of definitional theorem" #> register_config_bool Goal.quick_and_dirty #> register_config_bool Ast.trace #> register_config_bool Ast.stats #> register_config_bool Printer.show_brackets #> register_config_bool Printer.show_sorts #> register_config_bool Printer.show_types #> register_config_bool Printer.show_markup #> register_config_bool Printer.show_structs #> register_config_bool Printer.show_question_marks #> register_config_bool Syntax.ambiguity_warning #> register_config_int Syntax.ambiguity_limit #> register_config_bool Syntax_Trans.eta_contract #> register_config_bool Name_Space.names_long #> register_config_bool Name_Space.names_short #> register_config_bool Name_Space.names_unique #> register_config_int ML_Print_Depth.print_depth #> register_config_string ML_Env.ML_environment #> register_config_bool ML_Env.ML_read_global #> register_config_bool ML_Env.ML_write_global #> register_config_bool ML_Options.source_trace #> register_config_bool ML_Options.exception_trace #> register_config_bool ML_Options.exception_debugger #> register_config_bool ML_Options.debugger #> register_config_bool Proof_Context.show_abbrevs #> register_config_int Goal_Display.goals_limit #> register_config_bool Goal_Display.show_main_goal #> register_config_bool Thm.show_consts #> register_config_bool Thm.show_hyps #> register_config_bool Thm.show_tags #> register_config_bool Pattern.unify_trace_failure #> register_config_int Unify.trace_bound #> register_config_int Unify.search_bound #> register_config_bool Unify.trace_simp #> register_config_bool Unify.trace_types #> register_config_int Raw_Simplifier.simp_depth_limit #> register_config_int Raw_Simplifier.simp_trace_depth_limit #> register_config_bool Raw_Simplifier.simp_debug #> register_config_bool Raw_Simplifier.simp_trace #> register_config_bool Local_Defs.unfold_abs_def); (* internal source *) local val internal = internal_name (Context.the_local_context ()); val consumes_name = internal "consumes"; val constraints_name = internal "constraints"; val cases_open_name = internal "cases_open"; val case_names_name = internal "case_names"; val case_conclusion_name = internal "case_conclusion"; fun make_string s = Token.make_string (s, Position.none); in fun consumes i = consumes_name :: Token.make_int i; fun constraints i = constraints_name :: Token.make_int i; val cases_open = [cases_open_name]; fun case_names names = case_names_name :: map make_string names; fun case_conclusion (name, names) = case_conclusion_name :: map make_string (name :: names); end; end; \ No newline at end of file diff --git a/src/Pure/Isar/local_theory.ML b/src/Pure/Isar/local_theory.ML --- a/src/Pure/Isar/local_theory.ML +++ b/src/Pure/Isar/local_theory.ML @@ -1,428 +1,436 @@ (* Title: Pure/Isar/local_theory.ML Author: Makarius Local theory operations, with abstract target context. *) type local_theory = Proof.context; type generic_theory = Context.generic; structure Attrib = struct type binding = binding * Token.src list; type thms = (thm list * Token.src list) list; type fact = binding * thms; end; structure Locale = struct type registration = {inst: string * morphism, mixin: (morphism * bool) option, export: morphism}; end; signature LOCAL_THEORY = sig type operations val assert: local_theory -> local_theory val level: Proof.context -> int val map_contexts: (int -> Proof.context -> Proof.context) -> local_theory -> local_theory val background_naming_of: local_theory -> Name_Space.naming val map_background_naming: (Name_Space.naming -> Name_Space.naming) -> local_theory -> local_theory val restore_background_naming: local_theory -> local_theory -> local_theory val full_name: local_theory -> binding -> string val new_group: local_theory -> local_theory val reset_group: local_theory -> local_theory val standard_morphism: local_theory -> Proof.context -> morphism val standard_morphism_theory: local_theory -> morphism val standard_form: local_theory -> Proof.context -> (morphism -> 'a) -> 'a val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory val raw_theory: (theory -> theory) -> local_theory -> local_theory val background_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory val background_theory: (theory -> theory) -> local_theory -> local_theory val target_of: local_theory -> Proof.context val target: (Proof.context -> Proof.context) -> local_theory -> local_theory val target_morphism: local_theory -> morphism val propagate_ml_env: generic_theory -> generic_theory + val touch_ml_env: generic_theory -> generic_theory val operations_of: local_theory -> operations val define: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val define_internal: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val note: Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory val notes: Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val notes_kind: string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory val declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory val theory_registration: Locale.registration -> local_theory -> local_theory val locale_dependency: Locale.registration -> local_theory -> local_theory val pretty: local_theory -> Pretty.T list val add_thms_dynamic: binding * (Context.generic -> thm list) -> local_theory -> local_theory val set_defsort: sort -> local_theory -> local_theory val syntax: bool -> Syntax.mode -> (string * typ * mixfix) list -> local_theory -> local_theory val syntax_cmd: bool -> Syntax.mode -> (string * string * mixfix) list -> local_theory -> local_theory val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory val type_notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val type_alias: binding -> string -> local_theory -> local_theory val const_alias: binding -> string -> local_theory -> local_theory val init: {background_naming: Name_Space.naming, setup: theory -> Proof.context, conclude: local_theory -> Proof.context} -> operations -> theory -> local_theory val exit: local_theory -> Proof.context val exit_global: local_theory -> theory val exit_result: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * Proof.context val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory val begin_nested: local_theory -> Binding.scope * local_theory val end_nested: local_theory -> local_theory val end_nested_result: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * local_theory end; signature PRIVATE_LOCAL_THEORY = sig include LOCAL_THEORY val reset: local_theory -> local_theory end structure Local_Theory: PRIVATE_LOCAL_THEORY = struct (** local theory data **) (* type lthy *) type operations = {define: bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory, notes: string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory, abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory, declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory, theory_registration: Locale.registration -> local_theory -> local_theory, locale_dependency: Locale.registration -> local_theory -> local_theory, pretty: local_theory -> Pretty.T list}; type lthy = {background_naming: Name_Space.naming, operations: operations, conclude: Proof.context -> Proof.context, target: Proof.context}; fun make_lthy (background_naming, operations, conclude, target) : lthy = {background_naming = background_naming, operations = operations, conclude = conclude, target = target}; (* context data *) structure Data = Proof_Data ( type T = lthy list; fun init _ = []; ); (* nested structure *) val level = length o Data.get; (*1: main target at bottom, >= 2: nested target context*) fun assert lthy = if level lthy = 0 then error "Missing local theory context" else lthy; fun assert_bottom lthy = let val _ = assert lthy; in if level lthy > 1 then error "Not at bottom of local theory nesting" else lthy end; fun assert_not_bottom lthy = let val _ = assert lthy; in if level lthy = 1 then error "Already at bottom of local theory nesting" else lthy end; val bottom_of = List.last o Data.get o assert; val top_of = hd o Data.get o assert; fun map_top f = assert #> Data.map (fn {background_naming, operations, conclude, target} :: parents => make_lthy (f (background_naming, operations, conclude, target)) :: parents); fun reset lthy = #target (top_of lthy) |> Data.put (Data.get lthy); fun map_contexts f lthy = let val n = level lthy in lthy |> (Data.map o map_index) (fn (i, {background_naming, operations, conclude, target}) => make_lthy (background_naming, operations, conclude, target |> Context_Position.set_visible false |> f (n - i - 1) |> Context_Position.restore_visible target)) |> f n end; (* naming for background theory *) val background_naming_of = #background_naming o top_of; fun map_background_naming f = map_top (fn (background_naming, operations, conclude, target) => (f background_naming, operations, conclude, target)); val restore_background_naming = map_background_naming o K o background_naming_of; val full_name = Name_Space.full_name o background_naming_of; val new_group = map_background_naming Name_Space.new_group; val reset_group = map_background_naming Name_Space.reset_group; (* standard morphisms *) fun standard_morphism lthy ctxt = Proof_Context.norm_export_morphism lthy ctxt $> Morphism.binding_morphism "Local_Theory.standard_binding" (Name_Space.transform_binding (Proof_Context.naming_of lthy)); fun standard_morphism_theory lthy = standard_morphism lthy (Proof_Context.init_global (Proof_Context.theory_of lthy)); fun standard_form lthy ctxt x = Morphism.form (Morphism.transform (standard_morphism lthy ctxt) x); (* background theory *) fun raw_theory_result f lthy = let val (res, thy') = f (Proof_Context.theory_of lthy); val lthy' = map_contexts (K (Proof_Context.transfer thy')) lthy; in (res, lthy') end; fun raw_theory f = #2 o raw_theory_result (f #> pair ()); fun background_theory_result f lthy = let val naming = background_naming_of lthy |> Name_Space.transform_naming (Proof_Context.naming_of lthy); in lthy |> raw_theory_result (fn thy => thy |> Sign.map_naming (K naming) |> f ||> Sign.restore_naming thy) end; fun background_theory f = #2 o background_theory_result (f #> pair ()); (* target contexts *) val target_of = #target o bottom_of; fun target f lthy = let val ctxt = target_of lthy; val ctxt' = ctxt |> Context_Position.set_visible false |> f |> Context_Position.restore_visible ctxt; val thy' = Proof_Context.theory_of ctxt'; in map_contexts (fn 0 => K ctxt' | _ => Proof_Context.transfer thy') lthy end; fun target_morphism lthy = standard_morphism lthy (target_of lthy); fun propagate_ml_env (context as Context.Proof lthy) = let val inherit = ML_Env.inherit [context] in lthy |> background_theory (Context.theory_map inherit) |> map_contexts (K (Context.proof_map inherit)) |> Context.Proof end | propagate_ml_env context = context; +fun touch_ml_env context = + if Context.enabled_tracing () then + (case context of + Context.Theory _ => ML_Env.touch context + | Context.Proof _ => context) + else context; + (** operations **) val operations_of = #operations o top_of; fun operation f lthy = f (operations_of lthy) lthy; fun operation1 f x = operation (fn ops => f ops x); fun operation2 f x y = operation (fn ops => f ops x y); (* primitives *) val pretty = operation #pretty; val abbrev = operation2 #abbrev; val define = operation2 #define false; val define_internal = operation2 #define true; val notes_kind = operation2 #notes; val declaration = operation2 #declaration; val theory_registration = operation1 #theory_registration; fun locale_dependency registration = assert_bottom #> operation1 #locale_dependency registration; (* theorems *) val notes = notes_kind ""; fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single; fun add_thms_dynamic (binding, f) lthy = lthy |> background_theory_result (fn thy => thy |> Global_Theory.add_thms_dynamic' (Sign.inherit_naming thy lthy) (binding, f)) |-> (fn name => map_contexts (fn _ => fn ctxt => Proof_Context.transfer_facts (Proof_Context.theory_of ctxt) ctxt) #> declaration {syntax = false, pervasive = false} (fn phi => let val binding' = Morphism.binding phi binding in Context.mapping (Global_Theory.alias_fact binding' name) (Proof_Context.alias_fact binding' name) end)); (* default sort *) fun set_defsort S = declaration {syntax = true, pervasive = false} (K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S))); (* syntax *) fun gen_syntax prep_type add mode raw_args lthy = let val args = map (fn (c, T, mx) => (c, prep_type lthy T, mx)) raw_args; val args' = map (fn (c, T, mx) => (c, T, Mixfix.reset_pos mx)) args; val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.syntax add mode args; in declaration {syntax = true, pervasive = false} (fn _ => Proof_Context.generic_syntax add mode args') lthy end; val syntax = gen_syntax (K I); val syntax_cmd = gen_syntax Proof_Context.read_typ_syntax; (* notation *) local fun gen_type_notation prep_type add mode raw_args lthy = let val prepare = prep_type lthy #> Logic.type_map (Assumption.export_term lthy (target_of lthy)); val args = map (apfst prepare) raw_args; val args' = map (apsnd Mixfix.reset_pos) args; val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.type_notation add mode args; in declaration {syntax = true, pervasive = false} (Proof_Context.generic_type_notation add mode args') lthy end; fun gen_notation prep_const add mode raw_args lthy = let val prepare = prep_const lthy #> Assumption.export_term lthy (target_of lthy); val args = map (apfst prepare) raw_args; val args' = map (apsnd Mixfix.reset_pos) args; val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.notation add mode args; in declaration {syntax = true, pervasive = false} (Proof_Context.generic_notation add mode args') lthy end; in val type_notation = gen_type_notation (K I); val type_notation_cmd = gen_type_notation (Proof_Context.read_type_name {proper = true, strict = false}); val notation = gen_notation (K I); val notation_cmd = gen_notation (Proof_Context.read_const {proper = false, strict = false}); end; (* name space aliases *) fun syntax_alias global_alias local_alias b name = declaration {syntax = true, pervasive = false} (fn phi => let val b' = Morphism.binding phi b in Context.mapping (global_alias b' name) (local_alias b' name) end); val type_alias = syntax_alias Sign.type_alias Proof_Context.type_alias; val const_alias = syntax_alias Sign.const_alias Proof_Context.const_alias; (** manage targets **) (* main target *) fun init_target background_naming conclude operations target = Data.map (K [make_lthy (background_naming, operations, conclude, target)]) target fun init {background_naming, setup, conclude} operations thy = thy |> Sign.change_begin |> setup |> init_target background_naming (conclude #> target_of #> Sign.change_end_local) operations; val exit_of = #conclude o bottom_of; fun exit lthy = exit_of lthy (assert_bottom lthy); val exit_global = Proof_Context.theory_of o exit; fun exit_result decl (x, lthy) = let val ctxt = exit lthy; val phi = standard_morphism lthy ctxt; in (decl phi x, ctxt) end; fun exit_result_global decl (x, lthy) = let val thy = exit_global lthy; val thy_ctxt = Proof_Context.init_global thy; val phi = standard_morphism lthy thy_ctxt; in (decl phi x, thy) end; (* nested targets *) fun begin_nested lthy = let val _ = assert lthy; val (scope, target) = Proof_Context.new_scope lthy; val entry = make_lthy (background_naming_of lthy, operations_of lthy, Proof_Context.restore_naming lthy, target); val lthy' = Data.map (cons entry) target; in (scope, lthy') end; fun end_nested lthy = let val _ = assert_not_bottom lthy; val ({conclude, ...} :: rest) = Data.get lthy; in lthy |> Data.put rest |> reset |> conclude end; fun end_nested_result decl (x, lthy) = let val outer_lthy = end_nested lthy; val phi = Proof_Context.export_morphism lthy outer_lthy; in (decl phi x, outer_lthy) end; end; diff --git a/src/Pure/Isar/outer_syntax.ML b/src/Pure/Isar/outer_syntax.ML --- a/src/Pure/Isar/outer_syntax.ML +++ b/src/Pure/Isar/outer_syntax.ML @@ -1,332 +1,333 @@ (* Title: Pure/Isar/outer_syntax.ML Author: Markus Wenzel, TU Muenchen Isabelle/Isar outer syntax. *) signature OUTER_SYNTAX = sig val help: theory -> string list -> unit val print_commands: theory -> unit type command_keyword = string * Position.T val command: command_keyword -> string -> (Toplevel.transition -> Toplevel.transition) parser -> unit val maybe_begin_local_theory: command_keyword -> string -> (local_theory -> local_theory) parser -> (theory -> local_theory) parser -> unit val local_theory': command_keyword -> string -> (bool -> local_theory -> local_theory) parser -> unit val local_theory: command_keyword -> string -> (local_theory -> local_theory) parser -> unit val local_theory_to_proof': command_keyword -> string -> (bool -> local_theory -> Proof.state) parser -> unit val local_theory_to_proof: command_keyword -> string -> (local_theory -> Proof.state) parser -> unit val bootstrap: bool Config.T val parse_spans: Token.T list -> Command_Span.span list val make_span: Token.T list -> Command_Span.span val parse_span: theory -> (unit -> theory) -> Token.T list -> Toplevel.transition val parse_text: theory -> (unit -> theory) -> Position.T -> string -> Toplevel.transition list val command_reports: theory -> Token.T -> Position.report_text list val check_command: Proof.context -> command_keyword -> string end; structure Outer_Syntax: OUTER_SYNTAX = struct (** outer syntax **) (* errors *) fun err_command msg name ps = error (msg ^ quote (Markup.markup Markup.keyword1 name) ^ Position.here_list ps); fun err_dup_command name ps = err_command "Duplicate outer syntax command " name ps; (* command parsers *) datatype command_parser = Parser of (Toplevel.transition -> Toplevel.transition) parser | Restricted_Parser of (bool * Position.T) option -> (Toplevel.transition -> Toplevel.transition) parser; datatype command = Command of {comment: string, command_parser: command_parser, pos: Position.T, id: serial}; fun eq_command (Command {id = id1, ...}, Command {id = id2, ...}) = id1 = id2; fun new_command comment command_parser pos = Command {comment = comment, command_parser = command_parser, pos = pos, id = serial ()}; fun command_pos (Command {pos, ...}) = pos; fun command_markup def (name, Command {pos, id, ...}) = Position.make_entity_markup def id Markup.commandN (name, pos); fun pretty_command (cmd as (name, Command {comment, ...})) = Pretty.block (Pretty.marks_str ([Active.make_markup Markup.sendbackN {implicit = true, properties = [Markup.padding_line]}, command_markup {def = false} cmd], name) :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); (* theory data *) structure Data = Theory_Data ( type T = command Symtab.table; val empty = Symtab.empty; fun merge data : T = data |> Symtab.join (fn name => fn (cmd1, cmd2) => if eq_command (cmd1, cmd2) then raise Symtab.SAME else err_dup_command name [command_pos cmd1, command_pos cmd2]); ); val get_commands = Data.get; val dest_commands = get_commands #> Symtab.dest #> sort_by #1; val lookup_commands = Symtab.lookup o get_commands; fun help thy pats = dest_commands thy |> filter (fn (name, _) => forall (fn pat => match_string pat name) pats) |> map pretty_command |> Pretty.writeln_chunks; fun print_commands thy = let val keywords = Thy_Header.get_keywords thy; val minor = Scan.dest_lexicon (Keyword.minor_keywords keywords); val commands = dest_commands thy; in [Pretty.strs ("keywords:" :: map quote minor), Pretty.big_list "commands:" (map pretty_command commands)] |> Pretty.writeln_chunks end; (* maintain commands *) fun add_command name cmd thy = if member (op =) Thy_Header.bootstrap_thys (Context.theory_base_name thy) then thy else let val _ = Keyword.is_command (Thy_Header.get_keywords thy) name orelse err_command "Undeclared outer syntax command " name [command_pos cmd]; val _ = (case lookup_commands thy name of NONE => () | SOME cmd' => err_dup_command name [command_pos cmd, command_pos cmd']); val _ = Context_Position.report_generic (Context.the_generic_context ()) (command_pos cmd) (command_markup {def = true} (name, cmd)); in Data.map (Symtab.update (name, cmd)) thy end; val _ = Theory.setup (Theory.at_end (fn thy => let val command_keywords = Scan.dest_lexicon (Keyword.major_keywords (Thy_Header.get_keywords thy)); val _ = (case subtract (op =) (map #1 (dest_commands thy)) command_keywords of [] => () | missing => error ("Missing outer syntax command(s) " ^ commas_quote missing)) in NONE end)); (* implicit theory setup *) type command_keyword = string * Position.T; fun raw_command (name, pos) comment command_parser = let val setup = add_command name (new_command comment command_parser pos) in Context.>> (Context.mapping setup (Local_Theory.background_theory setup)) end; fun command (name, pos) comment parse = raw_command (name, pos) comment (Parser parse); fun maybe_begin_local_theory command_keyword comment parse_local parse_global = raw_command command_keyword comment (Restricted_Parser (fn restricted => Parse.opt_target -- parse_local >> (fn (target, f) => Toplevel.local_theory restricted target f) || (if is_some restricted then Scan.fail else parse_global >> Toplevel.begin_main_target true))); fun local_theory_command trans command_keyword comment parse = raw_command command_keyword comment (Restricted_Parser (fn restricted => Parse.opt_target -- parse >> (fn (target, f) => trans restricted target f))); val local_theory' = local_theory_command (fn a => fn b => fn c => Toplevel.local_theory' a b c NONE); val local_theory = local_theory_command Toplevel.local_theory; val local_theory_to_proof' = local_theory_command Toplevel.local_theory_to_proof'; val local_theory_to_proof = local_theory_command Toplevel.local_theory_to_proof; (** toplevel parsing **) (* parse spans *) local fun ship span = let val kind = if forall Token.is_ignored span then Command_Span.Ignored_Span else if exists Token.is_error span then Command_Span.Malformed_Span else (case find_first Token.is_command span of NONE => Command_Span.Malformed_Span | SOME cmd => Command_Span.Command_Span (Token.content_of cmd, Token.pos_of cmd)); in cons (Command_Span.Span (kind, span)) end; fun flush (result, content, ignored) = result |> not (null content) ? ship (rev content) |> not (null ignored) ? ship (rev ignored); fun parse tok (result, content, ignored) = if Token.is_ignored tok then (result, content, tok :: ignored) else if Token.is_command_modifier tok orelse Token.is_command tok andalso (not (exists Token.is_command_modifier content) orelse exists Token.is_command content) then (flush (result, content, ignored), [tok], []) else (result, tok :: (ignored @ content), []); in fun parse_spans toks = fold parse toks ([], [], []) |> flush |> rev; end; fun make_span toks = (case parse_spans toks of [span] => span | _ => Command_Span.Span (Command_Span.Malformed_Span, toks)); (* parse commands *) val bootstrap = Config.declare_bool ("Outer_Syntax.bootstrap", \<^here>) (K true); local val before_command = Scan.option (Parse.position (Parse.private >> K true || Parse.qualified >> K false)); fun parse_command thy markers = Scan.ahead (before_command |-- Parse.position Parse.command) :|-- (fn (name, pos) => let val keywords = Thy_Header.get_keywords thy; val tr0 = Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |> Keyword.is_proof_open keywords name ? Toplevel.skip_proof_open |> Keyword.is_proof_close keywords name ? Toplevel.skip_proof_close; val command_markers = Parse.command |-- Document_Source.old_tags >> (fn tags => Toplevel.markers (map Document_Marker.legacy_tag tags @ markers) tr0); in (case lookup_commands thy name of SOME (Command {command_parser = Parser parse, ...}) => Parse.!!! (command_markers -- parse) >> (op |>) | SOME (Command {command_parser = Restricted_Parser parse, ...}) => before_command :|-- (fn restricted => Parse.!!! (command_markers -- parse restricted)) >> (op |>) | NONE => Scan.fail_with (fn _ => fn _ => let val msg = if Config.get_global thy bootstrap then "missing theory context for command " else "undefined command "; in msg ^ quote (Markup.markup Markup.keyword1 name) end)) end); in fun parse_span thy init span = let val full_range = Token.range_of span; val core_range = Token.core_range_of span; val markers = map Token.input_of (filter Token.is_document_marker span); fun parse () = filter Token.is_proper span |> Source.of_list |> Source.source Token.stopper (Scan.bulk (fn xs => Parse.!!! (parse_command thy markers) xs)) |> Source.exhaust; in (case parse () of [tr] => Toplevel.modify_init init tr | [] => Toplevel.ignored (#1 full_range) | _ => Toplevel.malformed (#1 core_range) "Exactly one command expected") handle ERROR msg => Toplevel.malformed (#1 core_range) msg end; fun parse_text thy init pos text = Symbol_Pos.explode (text, pos) |> Token.tokenize (Thy_Header.get_keywords thy) {strict = false} |> parse_spans |> map (Command_Span.content #> parse_span thy init); end; (* check commands *) fun command_reports thy tok = if Token.is_command tok then let val name = Token.content_of tok in (case lookup_commands thy name of NONE => [] | SOME cmd => [((Token.pos_of tok, command_markup {def = false} (name, cmd)), "")]) end else []; fun check_command ctxt (name, pos) = let val thy = Proof_Context.theory_of ctxt; val keywords = Thy_Header.get_keywords thy; in if Keyword.is_command keywords name then let val markup = Token.explode0 keywords name |> maps (command_reports thy) |> map (#2 o #1); val _ = Context_Position.reports ctxt (map (pair pos) markup); in name end else let val completion_report = Completion.make_report (name, pos) (fn completed => Keyword.dest_commands keywords |> filter completed |> sort_strings |> map (fn a => (a, (Markup.commandN, a)))); in error ("Bad command " ^ quote name ^ Position.here pos ^ completion_report) end end; (* 'ML' command -- required for bootstrapping Isar *) val _ = command ("ML", \<^here>) "ML text within theory or local theory" (Parse.ML_source >> (fn source => Toplevel.generic_theory - (ML_Context.exec (fn () => + (Local_Theory.touch_ml_env #> + ML_Context.exec (fn () => ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags) source) #> Local_Theory.propagate_ml_env))); end; diff --git a/src/Pure/Isar/proof_display.ML b/src/Pure/Isar/proof_display.ML --- a/src/Pure/Isar/proof_display.ML +++ b/src/Pure/Isar/proof_display.ML @@ -1,393 +1,422 @@ (* Title: Pure/Isar/proof_display.ML Author: Makarius Printing of theorems, goals, results etc. *) signature PROOF_DISPLAY = sig val pp_context: Proof.context -> Pretty.T val pp_thm: (unit -> theory) -> thm -> Pretty.T val pp_typ: (unit -> theory) -> typ -> Pretty.T val pp_term: (unit -> theory) -> term -> Pretty.T val pp_ctyp: (unit -> theory) -> ctyp -> Pretty.T val pp_cterm: (unit -> theory) -> cterm -> Pretty.T val pretty_theory: bool -> Proof.context -> Pretty.T val pretty_definitions: bool -> Proof.context -> Pretty.T val pretty_theorems_diff: bool -> theory list -> Proof.context -> Pretty.T list val pretty_theorems: bool -> Proof.context -> Pretty.T list val string_of_rule: Proof.context -> string -> thm -> string val pretty_goal_header: thm -> Pretty.T val string_of_goal: Proof.context -> thm -> string val pretty_goal_facts: Proof.context -> string -> thm list option -> Pretty.T list val pretty_goal_inst: Proof.context -> term list list -> thm -> Pretty.T list val method_error: string -> Position.T -> {context: Proof.context, facts: thm list, goal: thm} -> 'a Seq.result val show_results: bool Config.T val print_results: bool -> Position.T -> Proof.context -> ((string * string) * (string * thm list) list) -> unit val print_consts: bool -> Position.T -> Proof.context -> (string * typ -> bool) -> (string * typ) list -> unit + val markup_extern_label: Position.T -> (Markup.T * xstring) option + val print_label: Position.T -> string + val print_context_tracing: Context.generic * Position.T -> string end structure Proof_Display: PROOF_DISPLAY = struct (** toplevel pretty printing **) fun pp_context ctxt = (if Config.get ctxt Proof_Context.debug then Pretty.quote (Pretty.big_list "proof context:" (Proof_Context.pretty_context ctxt)) else Pretty.str ""); fun default_context mk_thy = (case Context.get_generic_context () of SOME (Context.Proof ctxt) => ctxt | SOME (Context.Theory thy) => try Syntax.init_pretty_global thy |> \<^if_none>\Syntax.init_pretty_global (mk_thy ())\ | NONE => Syntax.init_pretty_global (mk_thy ())); fun pp_thm mk_thy th = Thm.pretty_thm_raw (default_context mk_thy) {quote = true, show_hyps = false} th; fun pp_typ mk_thy T = Pretty.quote (Syntax.pretty_typ (default_context mk_thy) T); fun pp_term mk_thy t = Pretty.quote (Syntax.pretty_term (default_context mk_thy) t); fun pp_ctyp mk_thy cT = pp_typ mk_thy (Thm.typ_of cT); fun pp_cterm mk_thy ct = pp_term mk_thy (Thm.term_of ct); (** theory content **) fun pretty_theory verbose ctxt = let val thy = Proof_Context.theory_of ctxt; fun prt_cls c = Syntax.pretty_sort ctxt [c]; fun prt_sort S = Syntax.pretty_sort ctxt S; fun prt_arity t (c, Ss) = Syntax.pretty_arity ctxt (t, Ss, [c]); fun prt_typ ty = Pretty.quote (Syntax.pretty_typ ctxt ty); val prt_typ_no_tvars = prt_typ o Logic.unvarifyT_global; fun prt_term t = Pretty.quote (Syntax.pretty_term ctxt t); val prt_term_no_vars = prt_term o Logic.unvarify_global; fun prt_const (c, ty) = [Pretty.mark_str c, Pretty.str " ::", Pretty.brk 1, prt_typ_no_tvars ty]; fun pretty_classrel (c, []) = prt_cls c | pretty_classrel (c, cs) = Pretty.block (prt_cls c :: Pretty.str " <" :: Pretty.brk 1 :: Pretty.commas (map prt_cls cs)); fun pretty_default S = Pretty.block [Pretty.str "default sort:", Pretty.brk 1, prt_sort S]; val tfrees = map (fn v => TFree (v, [])); fun pretty_type syn (t, Type.LogicalType n) = if syn then NONE else SOME (prt_typ (Type (t, tfrees (Name.invent Name.context Name.aT n)))) | pretty_type syn (t, Type.Abbreviation (vs, U, syn')) = if syn <> syn' then NONE else SOME (Pretty.block [prt_typ (Type (t, tfrees vs)), Pretty.str " =", Pretty.brk 1, prt_typ U]) | pretty_type syn (t, Type.Nonterminal) = if not syn then NONE else SOME (prt_typ (Type (t, []))); val pretty_arities = maps (fn (t, ars) => map (prt_arity t) ars); fun pretty_abbrev (c, (ty, t)) = Pretty.block (prt_const (c, ty) @ [Pretty.str " \", Pretty.brk 1, prt_term_no_vars t]); val tsig = Sign.tsig_of thy; val consts = Sign.consts_of thy; val {const_space, constants, constraints} = Consts.dest consts; val {classes, default, types, ...} = Type.rep_tsig tsig; val type_space = Type.type_space tsig; val (class_space, class_algebra) = classes; val classes = Sorts.classes_of class_algebra; val arities = Sorts.arities_of class_algebra; val clsses = Name_Space.extern_entries verbose ctxt class_space (map (fn ((c, _), cs) => (c, Sign.minimize_sort thy cs)) (Graph.dest classes)) |> map (apfst #1); val tdecls = Name_Space.extern_table verbose ctxt types |> map (apfst #1); val arties = Name_Space.extern_entries verbose ctxt type_space (Symtab.dest arities) |> map (apfst #1); val cnsts = Name_Space.markup_entries verbose ctxt const_space constants; val log_cnsts = map_filter (fn (c, (ty, NONE)) => SOME (c, ty) | _ => NONE) cnsts; val abbrevs = map_filter (fn (c, (ty, SOME t)) => SOME (c, (ty, t)) | _ => NONE) cnsts; val cnstrs = Name_Space.markup_entries verbose ctxt const_space constraints; in Pretty.chunks [Pretty.big_list "classes:" (map pretty_classrel clsses), pretty_default default, Pretty.big_list "syntactic types:" (map_filter (pretty_type true) tdecls), Pretty.big_list "logical types:" (map_filter (pretty_type false) tdecls), Pretty.big_list "type arities:" (pretty_arities arties), Pretty.big_list "logical consts:" (map (Pretty.block o prt_const) log_cnsts), Pretty.big_list "abbreviations:" (map pretty_abbrev abbrevs), Pretty.big_list "constraints:" (map (Pretty.block o prt_const) cnstrs), Pretty.block (Pretty.breaks (Pretty.str "oracles:" :: map Pretty.mark_str (Thm.extern_oracles verbose ctxt)))] end; (* theorems *) fun pretty_theorems_diff verbose prev_thys ctxt = let val pretty_fact = Proof_Context.pretty_fact ctxt; val facts = Global_Theory.facts_of (Proof_Context.theory_of ctxt); val thmss = Facts.dest_static verbose (map Global_Theory.facts_of prev_thys) facts; val prts = map #1 (sort_by (#1 o #2) (map (`pretty_fact) thmss)); in if null prts then [] else [Pretty.big_list "theorems:" prts] end; fun pretty_theorems verbose ctxt = pretty_theorems_diff verbose (Theory.parents_of (Proof_Context.theory_of ctxt)) ctxt; (* definitions *) fun pretty_definitions verbose ctxt = let val thy = Proof_Context.theory_of ctxt; val context = Proof_Context.defs_context ctxt; val const_space = Proof_Context.const_space ctxt; val type_space = Proof_Context.type_space ctxt; val item_space = fn Defs.Const => const_space | Defs.Type => type_space; fun prune_item (k, c) = not verbose andalso Name_Space.is_concealed (item_space k) c; fun extern_item (kind, name) = let val xname = Name_Space.extern ctxt (item_space kind) name in (xname, (kind, name)) end; fun extern_item_ord ((xname1, (kind1, _)), (xname2, (kind2, _))) = (case Defs.item_kind_ord (kind1, kind2) of EQUAL => string_ord (xname1, xname2) | ord => ord); fun sort_items f = sort (extern_item_ord o apply2 f); fun pretty_entry ((_: string, item), args) = Defs.pretty_entry context (item, args); fun pretty_reduct (lhs, rhs) = Pretty.block ([pretty_entry lhs, Pretty.str " ->", Pretty.brk 2] @ Pretty.commas (map pretty_entry (sort_items fst rhs))); fun pretty_restrict (entry, name) = Pretty.block ([pretty_entry entry, Pretty.brk 2, Pretty.str ("(from " ^ quote name ^ ")")]); val defs = Theory.defs_of thy; val {restricts, reducts} = Defs.dest defs; val (reds1, reds2) = filter_out (prune_item o #1 o #1) reducts |> map (fn (lhs, rhs) => (apfst extern_item lhs, map (apfst extern_item) (filter_out (prune_item o fst) rhs))) |> sort_items (#1 o #1) |> filter_out (null o #2) |> List.partition (Defs.plain_args o #2 o #1); val rests = restricts |> map (apfst (apfst extern_item)) |> sort_items (#1 o #1); in Pretty.big_list "definitions:" (map (Pretty.text_fold o single) [Pretty.big_list "non-overloaded LHS:" (map pretty_reduct reds1), Pretty.big_list "overloaded LHS:" (map pretty_reduct reds2), Pretty.big_list "pattern restrictions due to incomplete LHS:" (map pretty_restrict rests)]) end; (** proof items **) (* refinement rule *) fun pretty_rule ctxt s thm = Pretty.block [Pretty.str (s ^ " attempt to solve goal by exported rule:"), Pretty.fbrk, Thm.pretty_thm ctxt thm]; val string_of_rule = Pretty.string_of ooo pretty_rule; (* goals *) local fun subgoals 0 = [] | subgoals 1 = [Pretty.brk 1, Pretty.str "(1 subgoal)"] | subgoals n = [Pretty.brk 1, Pretty.str ("(" ^ string_of_int n ^ " subgoals)")]; in fun pretty_goal_header goal = Pretty.block ([Pretty.keyword1 "goal"] @ subgoals (Thm.nprems_of goal) @ [Pretty.str ":"]); end; fun string_of_goal ctxt goal = Pretty.string_of (Pretty.chunks [pretty_goal_header goal, Goal_Display.pretty_goal ctxt goal]); (* goal facts *) fun pretty_goal_facts _ _ NONE = [] | pretty_goal_facts ctxt s (SOME ths) = (single o Pretty.block o Pretty.fbreaks) [if s = "" then Pretty.str "this:" else Pretty.block [Pretty.keyword1 s, Pretty.brk 1, Pretty.str "this:"], Proof_Context.pretty_fact ctxt ("", ths)]; (* goal instantiation *) fun pretty_goal_inst ctxt propss goal = let val title = "goal instantiation:"; fun prt_inst env = if Envir.is_empty env then [] else let val Envir.Envir {tyenv, tenv, ...} = env; val prt_type = Syntax.pretty_typ ctxt; val prt_term = Syntax.pretty_term ctxt; fun instT v = let val T = TVar v; val T' = Envir.subst_type tyenv T; in if T = T' then NONE else SOME (prt_type T, prt_type T') end; fun inst v = let val t = Var v; val t' = Envir.subst_term (tyenv, tenv) (Envir.subst_term_types tyenv t); in if t aconv t' then NONE else SOME (prt_term t, prt_term t') end; fun inst_pair (x, y) = Pretty.item [x, Pretty.str " \", Pretty.brk 1, y]; val prts = ((fold o fold) Term.add_tvars propss [] |> sort Term_Ord.tvar_ord |> map_filter instT) @ ((fold o fold) Term.add_vars propss [] |> sort Term_Ord.var_ord |> map_filter inst); in if null prts then [] else [Pretty.big_list title (map inst_pair prts)] end; exception LOST_STRUCTURE; fun goal_matcher () = let val concl = Logic.unprotect (Thm.concl_of goal) handle TERM _ => raise LOST_STRUCTURE; val goal_propss = filter_out null propss; val results = Logic.dest_conjunction_balanced (length goal_propss) concl |> map2 Logic.dest_conjunction_balanced (map length goal_propss) handle TERM _ => raise LOST_STRUCTURE; in Unify.matcher (Context.Proof ctxt) (map (Soft_Type_System.purge ctxt) (flat goal_propss)) (flat results) end; fun failure msg = (warning (title ^ " " ^ msg); []); in (case goal_matcher () of SOME env => prt_inst env | NONE => failure "match failed") handle LOST_STRUCTURE => failure "lost goal structure" end; (* method_error *) fun method_error kind pos {context = ctxt, facts, goal} = Seq.Error (fn () => let val pr_header = "Failed to apply " ^ (if kind = "" then "" else kind ^ " ") ^ "proof method" ^ Position.here pos ^ ":\n"; val pr_facts = if null facts then "" else Pretty.string_of (Pretty.chunks (pretty_goal_facts ctxt "using" (SOME facts))) ^ "\n"; val pr_goal = string_of_goal ctxt goal; in pr_header ^ pr_facts ^ pr_goal end); (* results *) fun position_markup pos = Pretty.mark (Position.markup pos Markup.position); val interactive = Config.declare_bool ("interactive", \<^here>) (K false); val show_results = Attrib.setup_config_bool \<^binding>\show_results\ (fn context => Config.get_generic context interactive orelse Options.default_bool \<^system_option>\show_states\); fun no_print int ctxt = not (Config.get (Config.put interactive int ctxt) show_results); local fun pretty_fact_name pos (kind, "") = position_markup pos (Pretty.keyword1 kind) | pretty_fact_name pos (kind, name) = Pretty.block [position_markup pos (Pretty.keyword1 kind), Pretty.brk 1, Pretty.str (Long_Name.base_name name), Pretty.str ":"]; fun pretty_facts ctxt = flat o (separate [Pretty.fbrk, Pretty.keyword2 "and", Pretty.str " "]) o map (single o Proof_Context.pretty_fact ctxt); in fun print_results int pos ctxt ((kind, name), facts) = if kind = "" orelse no_print int ctxt then () else if name = "" then (Output.state o Pretty.string_of) (Pretty.block (position_markup pos (Pretty.keyword1 kind) :: Pretty.brk 1 :: pretty_facts ctxt facts)) else (Output.state o Pretty.string_of) (case facts of [fact] => Pretty.blk (1, [pretty_fact_name pos (kind, name), Pretty.fbrk, Proof_Context.pretty_fact ctxt fact]) | _ => Pretty.blk (1, [pretty_fact_name pos (kind, name), Pretty.fbrk, Pretty.blk (1, Pretty.str "(" :: pretty_facts ctxt facts @ [Pretty.str ")"])])); end; (* consts *) local fun pretty_var ctxt (x, T) = Pretty.block [Pretty.str x, Pretty.str " ::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt T)]; fun pretty_vars pos ctxt kind vs = Pretty.block (Pretty.fbreaks (position_markup pos (Pretty.mark_str kind) :: map (pretty_var ctxt) vs)); val fixes = (Markup.keyword2, "fixes"); val consts = (Markup.keyword1, "consts"); fun pretty_consts pos ctxt pred cs = (case filter pred (#1 (Proof_Context.inferred_fixes ctxt)) of [] => pretty_vars pos ctxt consts cs | ps => Pretty.chunks [pretty_vars pos ctxt fixes ps, pretty_vars pos ctxt consts cs]); in fun print_consts int pos ctxt pred cs = if null cs orelse no_print int ctxt then () else Output.state (Pretty.string_of (pretty_consts pos ctxt pred cs)); end; + +(* position label *) + +val command_prefix = Markup.commandN ^ "."; + +fun markup_extern_label pos = + Position.label_of pos |> Option.map (fn label => + (case try (unprefix command_prefix) label of + SOME cmd => (Markup.keyword1, Long_Name.base_name cmd) + | NONE => (Markup.empty, label))); + +fun print_label pos = + (case markup_extern_label pos of + NONE => "" + | SOME (m, s) => Markup.markup m s); + + +(* context tracing *) + +fun context_type (Context.Theory _) = "theory" + | context_type (Context.Proof ctxt) = + if can Local_Theory.assert ctxt then "local_theory" else "Proof.context"; + +fun print_context_tracing (context, pos) = + context_type context ^ (case print_label pos of "" => "" | s => " " ^ s) ^ Position.here pos; + end; diff --git a/src/Pure/Isar/toplevel.ML b/src/Pure/Isar/toplevel.ML --- a/src/Pure/Isar/toplevel.ML +++ b/src/Pure/Isar/toplevel.ML @@ -1,810 +1,810 @@ (* Title: Pure/Isar/toplevel.ML Author: Markus Wenzel, TU Muenchen Isabelle/Isar toplevel transactions. *) signature TOPLEVEL = sig exception UNDEF type state val make_state: theory option -> state val is_toplevel: state -> bool val is_theory: state -> bool val is_proof: state -> bool val is_skipped_proof: state -> bool val level: state -> int val previous_theory_of: state -> theory option val output_of: state -> Latex.text option val context_of: state -> Proof.context val generic_theory_of: state -> generic_theory val theory_of: state -> theory val proof_of: state -> Proof.state val proof_position_of: state -> int val is_end_theory: state -> bool val end_theory: Position.T -> state -> theory val presentation_context: state -> Proof.context val presentation_state: Proof.context -> state val pretty_context: state -> Pretty.T list val pretty_state: state -> Pretty.T list val string_of_state: state -> string val pretty_abstract: state -> Pretty.T type presentation = state -> Latex.text type transition val empty: transition val name_of: transition -> string val pos_of: transition -> Position.T val timing_of: transition -> Time.time val type_error: transition -> string val name: string -> transition -> transition val position: Position.T -> transition -> transition val markers: Input.source list -> transition -> transition val timing: Time.time -> transition -> transition val init_theory: (unit -> theory) -> transition -> transition val is_init: transition -> bool val modify_init: (unit -> theory) -> transition -> transition val exit: transition -> transition val present: presentation -> transition -> transition val keep: (state -> unit) -> transition -> transition val keep': (bool -> state -> unit) -> transition -> transition val keep_proof: (state -> unit) -> transition -> transition val is_ignored: transition -> bool val is_malformed: transition -> bool val ignored: Position.T -> transition val malformed: Position.T -> string -> transition val generic_theory: (generic_theory -> generic_theory) -> transition -> transition val theory': (bool -> theory -> theory) -> presentation option -> transition -> transition val theory: (theory -> theory) -> transition -> transition val begin_main_target: bool -> (theory -> local_theory) -> transition -> transition val end_main_target: transition -> transition val begin_nested_target: (Context.generic -> local_theory) -> transition -> transition val end_nested_target: transition -> transition val local_theory': (bool * Position.T) option -> (xstring * Position.T) option -> (bool -> local_theory -> local_theory) -> presentation option -> transition -> transition val local_theory: (bool * Position.T) option -> (xstring * Position.T) option -> (local_theory -> local_theory) -> transition -> transition val present_local_theory: (xstring * Position.T) option -> presentation -> transition -> transition val local_theory_to_proof': (bool * Position.T) option -> (xstring * Position.T) option -> (bool -> local_theory -> Proof.state) -> transition -> transition val local_theory_to_proof: (bool * Position.T) option -> (xstring * Position.T) option -> (local_theory -> Proof.state) -> transition -> transition val theory_to_proof: (theory -> Proof.state) -> transition -> transition val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition val forget_proof: transition -> transition val proofs': (bool -> Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition val proofs: (Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition val proof: (Proof.state -> Proof.state) -> transition -> transition val actual_proof: (Proof_Node.T -> Proof_Node.T) -> transition -> transition val skip_proof: (unit -> unit) -> transition -> transition val skip_proof_open: transition -> transition val skip_proof_close: transition -> transition val exec_id: Document_ID.exec -> transition -> transition val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b val transition: bool -> transition -> state -> state * (exn * string) option val command_errors: bool -> transition -> state -> Runtime.error list * state option val command_exception: bool -> transition -> state -> state val reset_theory: state -> state option val reset_proof: state -> state option val reset_notepad: state -> state option val fork_presentation: transition -> transition * transition type result val join_results: result -> (state * transition * state) list val element_result: Keyword.keywords -> transition Thy_Element.element -> state -> result * state end; structure Toplevel: TOPLEVEL = struct (** toplevel state **) exception UNDEF = Runtime.UNDEF; (* datatype node *) datatype node = Toplevel (*toplevel outside of theory body*) | Theory of generic_theory (*global or local theory*) | Proof of Proof_Node.T * ((Proof.context -> generic_theory) * generic_theory) (*proof node, finish, original theory*) | Skipped_Proof of int * (generic_theory * generic_theory); (*proof depth, resulting theory, original theory*) val theory_node = fn Theory gthy => SOME gthy | _ => NONE; val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE; val skipped_proof_node = fn Skipped_Proof _ => true | _ => false; fun cases_node f _ _ Toplevel = f () | cases_node _ g _ (Theory gthy) = g gthy | cases_node _ _ h (Proof (prf, _)) = h (Proof_Node.current prf) | cases_node _ g _ (Skipped_Proof (_, (gthy, _))) = g gthy; fun cases_proper_node g h = cases_node (fn () => raise UNDEF) g h; val get_theory = cases_node (K NONE) (SOME o Context.theory_of) (SOME o Proof.theory_of); (* datatype state *) type node_presentation = node * Proof.context; fun init_presentation () = Proof_Context.init_global (Theory.get_pure_bootstrap ()); fun node_presentation node = (node, cases_node init_presentation Context.proof_of Proof.context_of node); datatype state = State of node_presentation * (theory option * Latex.text future option); (*current node with presentation context, previous theory, document output*) fun node_of (State ((node, _), _)) = node; fun previous_theory_of (State (_, (prev_thy, _))) = prev_thy; fun output_of (State (_, (_, output))) = Option.map Future.join output; fun make_state opt_thy = let val node = (case opt_thy of NONE => Toplevel | SOME thy => Theory (Context.Theory thy)) in State (node_presentation node, (NONE, NONE)) end; fun level state = (case node_of state of Toplevel => 0 | Theory _ => 0 | Proof (prf, _) => Proof.level (Proof_Node.current prf) | Skipped_Proof (d, _) => d + 1); (*different notion of proof depth!*) fun str_of_state state = (case node_of state of Toplevel => (case previous_theory_of state of NONE => "at top level" | SOME thy => "at top level, result theory " ^ quote (Context.theory_base_name thy)) | Theory (Context.Theory _) => "in theory mode" | Theory (Context.Proof _) => "in local theory mode" | Proof _ => "in proof mode" | Skipped_Proof _ => "in skipped proof mode"); (* current node *) fun is_toplevel state = (case node_of state of Toplevel => true | _ => false); fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state)); fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state)); fun is_skipped_proof state = not (is_toplevel state) andalso skipped_proof_node (node_of state); fun proper_node_of state = if is_toplevel state then raise UNDEF else node_of state; fun proper_node_case f g state = cases_proper_node f g (proper_node_of state); val context_of = proper_node_case Context.proof_of Proof.context_of; val generic_theory_of = proper_node_case I (Context.Proof o Proof.context_of); val theory_of = proper_node_case Context.theory_of Proof.theory_of; val proof_of = proper_node_case (fn _ => error "No proof state") I; fun proof_position_of state = (case proper_node_of state of Proof (prf, _) => Proof_Node.position prf | _ => ~1); fun is_end_theory (State ((Toplevel, _), (SOME _, _))) = true | is_end_theory _ = false; fun end_theory _ (State ((Toplevel, _), (SOME thy, _))) = thy | end_theory pos _ = error ("Malformed theory" ^ Position.here pos); (* presentation context *) structure Presentation_State = Proof_Data ( type T = state option; fun init _ = NONE; ); fun presentation_context0 (State ((_, pr_ctxt), _)) = pr_ctxt; fun presentation_context (state as State (current, _)) = presentation_context0 state |> Presentation_State.put (SOME (State (current, (NONE, NONE)))); fun presentation_state ctxt = Presentation_State.get ctxt |> \<^if_none>\State (node_presentation (Theory (Context.Proof ctxt)), (NONE, NONE))\; (* print state *) fun pretty_context state = if is_toplevel state then [] else let val gthy = (case node_of state of Toplevel => raise Match | Theory gthy => gthy | Proof (_, (_, gthy)) => gthy | Skipped_Proof (_, (_, gthy)) => gthy); val lthy = Context.cases Named_Target.theory_init I gthy; in Local_Theory.pretty lthy end; fun pretty_state state = (case node_of state of Toplevel => [] | Theory _ => [] | Proof (prf, _) => Proof.pretty_state (Proof_Node.current prf) | Skipped_Proof (d, _) => [Pretty.str ("skipped proof: depth " ^ string_of_int d)]); val string_of_state = pretty_state #> Pretty.chunks #> Pretty.string_of; fun pretty_abstract state = Pretty.str (""); val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty_abstract); (** toplevel transitions **) (* primitive transitions *) type presentation = state -> Latex.text; datatype trans = (*init theory*) Init of unit -> theory | (*formal exit of theory*) Exit | (*keep state unchanged*) Keep of (bool -> state -> unit) * presentation option | (*node transaction and presentation*) Transaction of (bool -> node -> node_presentation) * presentation option; local fun present_state fork g node_pr prev_thy = let val state = State (node_pr, (prev_thy, NONE)); fun present pr = if fork andalso Future.proofs_enabled 1 then Execution.fork {name = "Toplevel.present_state", pos = Position.thread_data (), pri = ~1} (fn () => pr state) else Future.value (pr state); in State (node_pr, (prev_thy, Option.map present g)) end; fun no_update f g state = Runtime.controlled_execution (try generic_theory_of state) (fn () => let val prev_thy = previous_theory_of state; val () = f state; val node_pr = node_presentation (node_of state); in present_state false g node_pr prev_thy end) () fun update f g state = Runtime.controlled_execution (try generic_theory_of state) (fn () => let val prev_thy = previous_theory_of state; val node_pr' = f (node_of state); in present_state true g node_pr' prev_thy end) (); fun apply_tr int trans state = (case (trans, node_of state) of (Init f, Toplevel) => Runtime.controlled_execution NONE (fn () => State (node_presentation (Theory (Context.Theory (f ()))), (NONE, NONE))) () | (Exit, Theory (Context.Theory thy)) => let val State ((node', pr_ctxt), _) = state |> update (fn _ => node_presentation (Theory (Context.Theory (tap Thm.expose_theory (Theory.end_theory thy))))) NONE; in State ((Toplevel, pr_ctxt), (get_theory node', NONE)) end | (Keep (f, g), _) => no_update (fn x => f int x) g state | (Transaction _, Toplevel) => raise UNDEF | (Transaction (f, g), _) => update (fn x => f int x) g state | _ => raise UNDEF); fun apply_body _ [] _ = raise UNDEF | apply_body int [tr] state = apply_tr int tr state | apply_body int (tr :: trs) state = apply_body int trs state handle Runtime.UNDEF => apply_tr int tr state; fun apply_markers name markers (state as State ((node, pr_ctxt), prev_thy)) = let val state' = Runtime.controlled_execution (try generic_theory_of state) (fn () => State ((node, fold (Document_Marker.evaluate name) markers pr_ctxt), prev_thy)) (); in (state', NONE) end; in fun apply_capture int name markers trans state = (apply_body int trans state |> apply_markers name markers) handle exn => (state, SOME exn); end; (* datatype transition *) datatype transition = Transition of {name: string, (*command name*) pos: Position.T, (*source position*) markers: Input.source list, (*semantic document markers*) timing: Time.time, (*prescient timing information*) trans: trans list}; (*primitive transitions (union)*) fun make_transition (name, pos, markers, timing, trans) = Transition {name = name, pos = pos, markers = markers, timing = timing, trans = trans}; fun map_transition f (Transition {name, pos, markers, timing, trans}) = make_transition (f (name, pos, markers, timing, trans)); val empty = make_transition ("", Position.none, [], Time.zeroTime, []); (* diagnostics *) fun name_of (Transition {name, ...}) = name; fun pos_of (Transition {pos, ...}) = pos; fun timing_of (Transition {timing, ...}) = timing; fun command_msg msg tr = msg ^ "command " ^ quote (Markup.markup Markup.keyword1 (name_of tr)) ^ Position.here (pos_of tr); fun at_command tr = command_msg "At " tr; fun type_error tr = command_msg "Bad context for " tr; (* modify transitions *) fun name name = map_transition (fn (_, pos, markers, timing, trans) => (name, pos, markers, timing, trans)); fun position pos = map_transition (fn (name, _, markers, timing, trans) => (name, pos, markers, timing, trans)); fun markers markers = map_transition (fn (name, pos, _, timing, trans) => (name, pos, markers, timing, trans)); fun timing timing = map_transition (fn (name, pos, markers, _, trans) => (name, pos, markers, timing, trans)); fun add_trans tr = map_transition (fn (name, pos, markers, timing, trans) => (name, pos, markers, timing, tr :: trans)); val reset_trans = map_transition (fn (name, pos, markers, timing, _) => (name, pos, markers, timing, [])); (* basic transitions *) fun init_theory f = add_trans (Init f); fun is_init (Transition {trans = [Init _], ...}) = true | is_init _ = false; fun modify_init f tr = if is_init tr then init_theory f (reset_trans tr) else tr; val exit = add_trans Exit; fun present_transaction f g = add_trans (Transaction (f, g)); fun transaction f = present_transaction f NONE; fun transaction0 f = present_transaction (node_presentation oo f) NONE; fun present g = add_trans (Keep (fn _ => fn _ => (), SOME g)); fun keep f = add_trans (Keep (K f, NONE)); fun keep' f = add_trans (Keep (f, NONE)); fun keep_proof f = keep (fn st => if is_proof st then f st else if is_skipped_proof st then () else warning "No proof state"); val ignoredN = ""; val malformedN = ""; fun is_ignored tr = name_of tr = ignoredN; fun is_malformed tr = name_of tr = malformedN; fun ignored pos = empty |> name ignoredN |> position pos |> keep (fn _ => ()); fun malformed pos msg = empty |> name malformedN |> position pos |> keep (fn _ => error msg); (* theory transitions *) fun generic_theory f = transaction (fn _ => (fn Theory gthy => node_presentation (Theory (f gthy)) | _ => raise UNDEF)); fun theory' f = present_transaction (fn int => (fn Theory (Context.Theory thy) => let val thy' = thy |> Sign.new_group |> f int |> Sign.reset_group; in node_presentation (Theory (Context.Theory thy')) end | _ => raise UNDEF)); fun theory f = theory' (K f) NONE; fun begin_main_target begin f = transaction (fn _ => (fn Theory (Context.Theory thy) => let val lthy = f thy; val gthy = if begin then Context.Proof lthy else Target_Context.end_named_cmd lthy; val _ = (case Local_Theory.pretty lthy of [] => () | prts => Output.state (Pretty.string_of (Pretty.chunks prts))); in (Theory gthy, lthy) end | _ => raise UNDEF)); val end_main_target = transaction (fn _ => (fn Theory (Context.Proof lthy) => (Theory (Target_Context.end_named_cmd lthy), lthy) | _ => raise UNDEF)); fun begin_nested_target f = transaction0 (fn _ => (fn Theory gthy => let val lthy' = f gthy; in Theory (Context.Proof lthy') end | _ => raise UNDEF)); val end_nested_target = transaction (fn _ => (fn Theory (Context.Proof lthy) => (case try Target_Context.end_nested_cmd lthy of SOME gthy' => (Theory gthy', lthy) | NONE => raise UNDEF) | _ => raise UNDEF)); fun restricted_context (SOME (strict, scope)) = Proof_Context.map_naming (Name_Space.restricted strict scope) | restricted_context NONE = I; fun local_theory' restricted target f = present_transaction (fn int => (fn Theory gthy => let val (finish, lthy) = Target_Context.switch_named_cmd target gthy; val lthy' = lthy |> restricted_context restricted |> Local_Theory.new_group |> f int |> Local_Theory.reset_group; in (Theory (finish lthy'), lthy') end | _ => raise UNDEF)); fun local_theory restricted target f = local_theory' restricted target (K f) NONE; fun present_local_theory target g = present_transaction (fn _ => (fn Theory gthy => let val (finish, lthy) = Target_Context.switch_named_cmd target gthy; in (Theory (finish lthy), lthy) end | _ => raise UNDEF)) (SOME g); (* proof transitions *) fun end_proof f = transaction (fn int => (fn Proof (prf, (finish, _)) => let val state = Proof_Node.current prf in if can (Proof.assert_bottom true) state then let val ctxt' = f int state; val gthy' = finish ctxt'; in (Theory gthy', ctxt') end else raise UNDEF end | Skipped_Proof (0, (gthy, _)) => node_presentation (Theory gthy) | _ => raise UNDEF)); local fun begin_proof init_proof = transaction0 (fn int => (fn Theory gthy => let val (finish, prf) = init_proof int gthy; val document = Options.default_string "document"; val skip = (document = "" orelse document = "false") andalso Goal.skip_proofs_enabled (); val schematic_goal = try Proof.schematic_goal prf; val _ = if skip andalso schematic_goal = SOME true then warning "Cannot skip proof of schematic goal statement" else (); in if skip andalso schematic_goal = SOME false then Skipped_Proof (0, (finish (Proof.global_skip_proof true prf), gthy)) else Proof (Proof_Node.init prf, (finish, gthy)) end | _ => raise UNDEF)); in fun local_theory_to_proof' restricted target f = begin_proof (fn int => fn gthy => let val (finish, lthy) = Target_Context.switch_named_cmd target gthy; val prf = lthy |> restricted_context restricted |> Local_Theory.new_group |> f int; in (finish o Local_Theory.reset_group, prf) end); fun local_theory_to_proof restricted target f = local_theory_to_proof' restricted target (K f); fun theory_to_proof f = begin_proof (fn _ => fn gthy => (Context.Theory o Sign.reset_group o Sign.change_check o Proof_Context.theory_of, (case gthy of Context.Theory thy => f (Sign.new_group thy) | _ => raise UNDEF))); end; val forget_proof = transaction0 (fn _ => (fn Proof (prf, (_, orig_gthy)) => if Proof.is_notepad (Proof_Node.current prf) then raise UNDEF else Theory orig_gthy | Skipped_Proof (_, (_, orig_gthy)) => Theory orig_gthy | _ => raise UNDEF)); fun proofs' f = transaction0 (fn int => (fn Proof (prf, x) => Proof (Proof_Node.applys (f int) prf, x) | skip as Skipped_Proof _ => skip | _ => raise UNDEF)); fun proof' f = proofs' ((Seq.single o Seq.Result) oo f); val proofs = proofs' o K; val proof = proof' o K; (* skipped proofs *) fun actual_proof f = transaction0 (fn _ => (fn Proof (prf, x) => Proof (f prf, x) | _ => raise UNDEF)); fun skip_proof f = transaction0 (fn _ => (fn skip as Skipped_Proof _ => (f (); skip) | _ => raise UNDEF)); val skip_proof_open = transaction0 (fn _ => (fn Skipped_Proof (d, x) => Skipped_Proof (d + 1, x) | _ => raise UNDEF)); val skip_proof_close = transaction0 (fn _ => (fn Skipped_Proof (0, (gthy, _)) => Theory gthy | Skipped_Proof (d, x) => Skipped_Proof (d - 1, x) | _ => raise UNDEF)); (** toplevel transactions **) (* runtime position *) fun exec_id id (tr as Transition {pos, ...}) = let val put_id = Position.put_id (Document_ID.print id) in position (put_id pos) tr end; -fun setmp_thread_position (Transition {pos, ...}) f x = - Position.setmp_thread_data pos f x; +fun setmp_thread_position (Transition {pos, name, ...}) f x = + Position.setmp_thread_data (Position.label (Long_Name.qualify Markup.commandN name) pos) f x; (* apply transitions *) local fun show_state (st', opt_err) = let val enabled = is_none opt_err andalso Options.default_bool \<^system_option>\show_states\; val opt_err' = if enabled then (case Exn.capture (Output.state o string_of_state) st' of Exn.Exn exn => SOME exn | Exn.Res _ => opt_err) else opt_err; in (st', opt_err') end; fun app int (tr as Transition {name, markers, trans, ...}) = setmp_thread_position tr (Timing.protocol (name_of tr) (pos_of tr) (apply_capture int name markers trans) ##> Option.map (fn Runtime.UNDEF => ERROR (type_error tr) | exn => exn) #> show_state); in fun transition int tr st = let val (st', opt_err) = Context.setmp_generic_context (try (Context.Proof o presentation_context0) st) (fn () => app int tr st) (); val opt_err' = opt_err |> Option.map (fn Runtime.EXCURSION_FAIL exn_info => exn_info | exn => (Runtime.exn_context (try context_of st) exn, at_command tr)); in (st', opt_err') end; end; (* managed commands *) fun command_errors int tr st = (case transition int tr st of (st', NONE) => ([], SOME st') | (_, SOME (exn, _)) => (Runtime.exn_messages exn, NONE)); fun command_exception int tr st = (case transition int tr st of (st', NONE) => st' | (_, SOME (exn, info)) => if Exn.is_interrupt exn then Exn.reraise exn else raise Runtime.EXCURSION_FAIL (exn, info)); val command = command_exception false; (* reset state *) local fun reset_state check trans st = if check st then NONE else #2 (command_errors false (trans empty) st); in val reset_theory = reset_state is_theory forget_proof; val reset_proof = reset_state is_proof (transaction0 (fn _ => (fn Theory gthy => Skipped_Proof (0, (gthy, gthy)) | _ => raise UNDEF))); val reset_notepad = reset_state (fn st => (case try proof_of st of SOME state => not (Proof.is_notepad state) orelse can Proof.end_notepad state | NONE => true)) (proof Proof.reset_notepad); end; (* scheduled proof result *) datatype result = Result of transition * state | Result_List of result list | Result_Future of result future; fun join_results result = let fun add (tr, st') res = (case res of [] => [(make_state NONE, tr, st')] | (_, _, st) :: _ => (st, tr, st') :: res); fun acc (Result r) = add r | acc (Result_List rs) = fold acc rs | acc (Result_Future x) = acc (Future.join x); in rev (acc result []) end; local structure Result = Proof_Data ( type T = result; fun init _ = Result_List []; ); val get_result = Result.get o Proof.context_of; val put_result = Proof.map_context o Result.put; fun timing_estimate elem = let val trs = tl (Thy_Element.flat_element elem) in fold (fn tr => fn t => timing_of tr + t) trs Time.zeroTime end; fun future_proofs_enabled estimate st = (case try proof_of st of NONE => false | SOME state => not (Proofterm.proofs_enabled ()) andalso not (Proof.is_relevant state) andalso (if can (Proof.assert_bottom true) state then Future.proofs_enabled 1 else Future.proofs_enabled 2 orelse Future.proofs_enabled_timing estimate)); val empty_markers = markers []; val empty_trans = reset_trans #> keep (K ()); in fun fork_presentation tr = (tr |> empty_markers, tr |> empty_trans); fun atom_result keywords tr st = let val st' = if Future.proofs_enabled 1 andalso Keyword.is_diag keywords (name_of tr) then let val (tr1, tr2) = fork_presentation tr; val _ = Execution.fork {name = "Toplevel.diag", pos = pos_of tr, pri = ~1} (fn () => command tr1 st); in command tr2 st end else command tr st; in (Result (tr, st'), st') end; fun element_result keywords (Thy_Element.Element (tr, NONE)) st = atom_result keywords tr st | element_result keywords (elem as Thy_Element.Element (head_tr, SOME element_rest)) st = let val (head_result, st') = atom_result keywords head_tr st; val (body_elems, end_tr) = element_rest; val estimate = timing_estimate elem; in if not (future_proofs_enabled estimate st') then let val proof_trs = maps Thy_Element.flat_element body_elems @ [end_tr]; val (proof_results, st'') = fold_map (atom_result keywords) proof_trs st'; in (Result_List (head_result :: proof_results), st'') end else let val (end_tr1, end_tr2) = fork_presentation end_tr; val finish = Context.Theory o Proof_Context.theory_of; val future_proof = Proof.future_proof (fn state => Execution.fork {name = "Toplevel.future_proof", pos = pos_of head_tr, pri = ~1} (fn () => let val State ((Proof (prf, (_, orig_gthy)), _), prev_thy) = st'; val node' = Proof (Proof_Node.apply (K state) prf, (finish, orig_gthy)); val (results, result_state) = State (node_presentation node', prev_thy) |> fold_map (element_result keywords) body_elems ||> command end_tr1; in (Result_List results, presentation_context0 result_state) end)) #> (fn (res, state') => state' |> put_result (Result_Future res)); val forked_proof = proof (future_proof #> (fn state => state |> Proof.local_done_proof |> put_result (get_result state))) o end_proof (fn _ => future_proof #> (fn state => state |> Proof.global_done_proof |> Result.put (get_result state))); val st'' = st' |> command (head_tr |> reset_trans |> forked_proof); val end_st = st'' |> command end_tr2; val end_result = Result (end_tr, end_st); val result = Result_List [head_result, Result.get (presentation_context0 st''), end_result]; in (result, end_st) end end; end; end; diff --git a/src/Pure/ML/ml_compiler.ML b/src/Pure/ML/ml_compiler.ML --- a/src/Pure/ML/ml_compiler.ML +++ b/src/Pure/ML/ml_compiler.ML @@ -1,304 +1,305 @@ (* Title: Pure/ML/ml_compiler.ML Author: Makarius Runtime compilation and evaluation. *) signature ML_COMPILER = sig type flags = {environment: string, redirect: bool, verbose: bool, debug: bool option, writeln: string -> unit, warning: string -> unit} val debug_flags: bool option -> flags val flags: flags val verbose: bool -> flags -> flags val eval: flags -> Position.T -> ML_Lex.token list -> unit end; structure ML_Compiler: ML_COMPILER = struct (* flags *) type flags = {environment: string, redirect: bool, verbose: bool, debug: bool option, writeln: string -> unit, warning: string -> unit}; fun debug_flags opt_debug : flags = {environment = "", redirect = false, verbose = false, debug = opt_debug, writeln = writeln, warning = warning}; val flags = debug_flags NONE; fun verbose b (flags: flags) = {environment = #environment flags, redirect = #redirect flags, verbose = b, debug = #debug flags, writeln = #writeln flags, warning = #warning flags}; (* parse trees *) fun breakpoint_position loc = let val pos = Position.no_range_position (Exn_Properties.position_of_polyml_location loc) in (case Position.offset_of pos of NONE => pos | SOME 1 => pos | SOME j => Position.properties_of pos |> Properties.put (Markup.offsetN, Value.print_int (j - 1)) |> Position.of_properties) end; fun report_parse_tree redirect depth name_space parse_tree = let val reports_enabled = (case Context.get_generic_context () of SOME context => Context_Position.reports_enabled_generic context | NONE => Context_Position.reports_enabled0 ()); fun is_reported pos = reports_enabled andalso Position.is_reported pos; (* syntax reports *) fun reported_types loc types = let val pos = Exn_Properties.position_of_polyml_location loc in is_reported pos ? let val xml = PolyML.NameSpace.Values.printType (types, depth, SOME name_space) |> Pretty.from_polyml |> Pretty.string_of |> Output.output |> YXML.parse_body; in cons (pos, fn () => Markup.ML_typing, fn () => YXML.string_of_body xml) end end; fun reported_entity kind loc decl = let val pos = Exn_Properties.position_of_polyml_location loc; val def_pos = Exn_Properties.position_of_polyml_location decl; in (is_reported pos andalso pos <> def_pos) ? cons (pos, fn () => Position.entity_markup kind ("", def_pos), fn () => "") end; fun reported_entity_id def id loc = let val pos = Exn_Properties.position_of_polyml_location loc; in (is_reported pos andalso id <> 0) ? let fun markup () = (Markup.entityN, [(if def then Markup.defN else Markup.refN, Value.print_int id)]); in cons (pos, markup, fn () => "") end end; fun reported_completions loc names = let val pos = Exn_Properties.position_of_polyml_location loc in if is_reported pos andalso not (null names) then let val completion = Completion.names pos (map (fn a => (a, ("ML", a))) names); val xml = Completion.encode completion; in cons (pos, fn () => Markup.completion, fn () => YXML.string_of_body xml) end else I end; fun reported _ (PolyML.PTnextSibling tree) = reported_tree (tree ()) | reported _ (PolyML.PTfirstChild tree) = reported_tree (tree ()) | reported loc (PolyML.PTdefId id) = reported_entity_id true (FixedInt.toLarge id) loc | reported loc (PolyML.PTrefId id) = reported_entity_id false (FixedInt.toLarge id) loc | reported loc (PolyML.PTtype types) = reported_types loc types | reported loc (PolyML.PTdeclaredAt decl) = reported_entity Markup.ML_defN loc decl | reported loc (PolyML.PTcompletions names) = reported_completions loc names | reported _ _ = I and reported_tree (loc, props) = fold (reported loc) props; val persistent_reports = reported_tree parse_tree []; fun output () = persistent_reports |> map (fn (pos, markup, text) => Position.reported_text pos (markup ()) (text ())) |> Output.report; val _ = if not (null persistent_reports) andalso redirect andalso not (Options.default_bool "pide_reports") andalso Future.enabled () then Execution.print {name = "ML_Compiler.report", pos = Position.thread_data (), pri = Task_Queue.urgent_pri} output else output (); (* breakpoints *) fun breakpoints _ (PolyML.PTnextSibling tree) = breakpoints_tree (tree ()) | breakpoints _ (PolyML.PTfirstChild tree) = breakpoints_tree (tree ()) | breakpoints loc (PolyML.PTbreakPoint b) = let val pos = breakpoint_position loc in if is_reported pos then let val id = serial (); in cons ((pos, Markup.ML_breakpoint id), (id, (b, pos))) end else I end | breakpoints _ _ = I and breakpoints_tree (loc, props) = fold (breakpoints loc) props; val all_breakpoints = rev (breakpoints_tree parse_tree []); val _ = Position.reports (map #1 all_breakpoints); in map (fn (_, (id, (b, pos))) => (id, (b, Position.dest pos))) all_breakpoints end; (* eval ML source tokens *) fun eval (flags: flags) pos toks = let val opt_context = Context.get_generic_context (); val env as {debug, name_space, add_breakpoints} = (case (ML_Recursive.get (), #environment flags <> "") of (SOME env, false) => env | _ => {debug = (case #debug flags of SOME debug => debug | NONE => ML_Options.debugger_enabled opt_context), name_space = ML_Env.make_name_space (#environment flags), add_breakpoints = ML_Env.add_breakpoints}); (* input *) val position_props = - filter (fn (a, _) => a = Markup.idN orelse a = Markup.fileN) (Position.properties_of pos); + filter (fn (a, _) => a = Markup.labelN orelse a = Markup.fileN orelse a = Markup.idN) + (Position.properties_of pos); val location_props = op ^ (YXML.output_markup (":", position_props)); val {explode_token, ...} = ML_Env.operations opt_context (#environment flags); fun token_content tok = if ML_Lex.is_comment tok then NONE else SOME (`explode_token tok); val input_buffer = Unsynchronized.ref (map_filter token_content toks); fun get () = (case ! input_buffer of (c :: cs, tok) :: rest => (input_buffer := (cs, tok) :: rest; SOME c) | ([], _) :: rest => (input_buffer := rest; SOME #" ") | [] => NONE); fun get_pos () = (case ! input_buffer of (_ :: _, tok) :: _ => ML_Lex.pos_of tok | ([], tok) :: _ => ML_Lex.end_pos_of tok | [] => Position.none); (* output *) val writeln_buffer = Unsynchronized.ref Buffer.empty; fun write s = Unsynchronized.change writeln_buffer (Buffer.add s); fun output_writeln () = #writeln flags (trim_line (Buffer.content (! writeln_buffer))); val warnings = Unsynchronized.ref ([]: string list); fun warn msg = Unsynchronized.change warnings (cons msg); fun output_warnings () = List.app (#warning flags) (rev (! warnings)); val error_buffer = Unsynchronized.ref Buffer.empty; fun err msg = Unsynchronized.change error_buffer (Buffer.add msg #> Buffer.add "\n"); fun flush_error () = #writeln flags (trim_line (Buffer.content (! error_buffer))); fun raise_error msg = error (trim_line (Buffer.content (Buffer.add msg (! error_buffer)))); fun message {message = msg, hard, location = loc, context = _} = let val pos = Exn_Properties.position_of_polyml_location loc; val txt = (if hard then "ML error" else "ML warning") ^ Position.here pos ^ ":\n" ^ Pretty.string_of (Pretty.from_polyml msg); in if hard then err txt else warn txt end; (* results *) val depth = FixedInt.fromInt (ML_Print_Depth.get_print_depth ()); fun apply_result {fixes, types, signatures, structures, functors, values} = let fun display disp x = if depth > 0 then (write (disp x |> Pretty.from_polyml |> Pretty.string_of); write "\n") else (); fun apply_fix (a, b) = (#enterFix name_space (a, b); display PolyML.NameSpace.Infixes.print b); fun apply_type (a, b) = (#enterType name_space (a, b); display PolyML.NameSpace.TypeConstrs.print (b, depth, SOME name_space)); fun apply_sig (a, b) = (#enterSig name_space (a, b); display PolyML.NameSpace.Signatures.print (b, depth, SOME name_space)); fun apply_struct (a, b) = (#enterStruct name_space (a, b); display PolyML.NameSpace.Structures.print (b, depth, SOME name_space)); fun apply_funct (a, b) = (#enterFunct name_space (a, b); display PolyML.NameSpace.Functors.print (b, depth, SOME name_space)); fun apply_val (a, b) = (#enterVal name_space (a, b); display PolyML.NameSpace.Values.printWithType (b, depth, SOME name_space)); in List.app apply_fix fixes; List.app apply_type types; List.app apply_sig signatures; List.app apply_struct structures; List.app apply_funct functors; List.app apply_val values end; exception STATIC_ERRORS of unit; fun result_fun (phase1, phase2) () = ((case phase1 of NONE => () | SOME parse_tree => add_breakpoints (report_parse_tree (#redirect flags) depth name_space parse_tree)); (case phase2 of NONE => raise STATIC_ERRORS () | SOME code => apply_result ((code |> Runtime.debugging opt_context |> Runtime.toplevel_error (err o Runtime.exn_message)) ()))); (* compiler invocation *) val parameters = [PolyML.Compiler.CPOutStream write, PolyML.Compiler.CPNameSpace name_space, PolyML.Compiler.CPErrorMessageProc message, PolyML.Compiler.CPLineNo (the_default 0 o Position.line_of o get_pos), PolyML.Compiler.CPLineOffset (the_default 0 o Position.offset_of o get_pos), PolyML.Compiler.CPFileName location_props, PolyML.Compiler.CPPrintDepth ML_Print_Depth.get_print_depth, PolyML.Compiler.CPCompilerResultFun result_fun, PolyML.Compiler.CPPrintInAlphabeticalOrder false, PolyML.Compiler.CPDebug debug, PolyML.Compiler.CPBindingSeq serial]; val _ = (while not (List.null (! input_buffer)) do ML_Recursive.recursive env (fn () => PolyML.compiler (get, parameters) ())) handle exn => if Exn.is_interrupt exn then Exn.reraise exn else let val exn_msg = (case exn of STATIC_ERRORS () => "" | Runtime.TOPLEVEL_ERROR => "" | _ => "Exception- " ^ Pretty.string_of (Runtime.pretty_exn exn) ^ " raised"); val _ = output_warnings (); val _ = output_writeln (); in raise_error exn_msg end; in if #verbose flags then (output_warnings (); flush_error (); output_writeln ()) else () end; end; diff --git a/src/Pure/ML/ml_context.ML b/src/Pure/ML/ml_context.ML --- a/src/Pure/ML/ml_context.ML +++ b/src/Pure/ML/ml_context.ML @@ -1,235 +1,237 @@ (* Title: Pure/ML/ml_context.ML Author: Makarius ML context and antiquotations. *) signature ML_CONTEXT = sig val check_antiquotation: Proof.context -> xstring * Position.T -> string val struct_name: Proof.context -> string val variant: string -> Proof.context -> string * Proof.context type decl = Proof.context -> ML_Lex.token list * ML_Lex.token list type decls = Proof.context -> (ML_Lex.token list * ML_Lex.token list) list type antiquotation = Position.range -> Token.src -> Proof.context -> decl * Proof.context val add_antiquotation: binding -> bool -> antiquotation -> theory -> theory val print_antiquotations: bool -> Proof.context -> unit val expand_antiquotes: ML_Lex.token Antiquote.antiquote list -> Proof.context -> decl * Proof.context val expand_antiquotes_list: ML_Lex.token Antiquote.antiquote list list -> Proof.context -> decls * Proof.context val eval: ML_Compiler.flags -> Position.T -> ML_Lex.token Antiquote.antiquote list -> unit val eval_file: ML_Compiler.flags -> Path.T -> unit val eval_source: ML_Compiler.flags -> Input.source -> unit val eval_in: Proof.context option -> ML_Compiler.flags -> Position.T -> ML_Lex.token Antiquote.antiquote list -> unit val eval_source_in: Proof.context option -> ML_Compiler.flags -> Input.source -> unit val exec: (unit -> unit) -> Context.generic -> Context.generic val expression: Position.T -> ML_Lex.token Antiquote.antiquote list -> Context.generic -> Context.generic end structure ML_Context: ML_CONTEXT = struct (** ML antiquotations **) (* names for generated environment *) structure Names = Proof_Data ( type T = string * Name.context; val init_names = ML_Syntax.reserved |> Name.declare "ML_context"; fun init _ = ("Isabelle0", init_names); ); fun struct_name ctxt = #1 (Names.get ctxt); val struct_begin = (Names.map o apfst) (fn _ => "Isabelle" ^ serial_string ()); fun variant a ctxt = let val names = #2 (Names.get ctxt); val (b, names') = Name.variant (Name.desymbolize (SOME false) a) names; val ctxt' = (Names.map o apsnd) (K names') ctxt; in (b, ctxt') end; (* theory data *) type decl = Proof.context -> ML_Lex.token list * ML_Lex.token list; (*final context -> ML env, ML body*) type decls = Proof.context -> (ML_Lex.token list * ML_Lex.token list) list; type antiquotation = Position.range -> Token.src -> Proof.context -> decl * Proof.context; structure Antiquotations = Theory_Data ( type T = (bool * antiquotation) Name_Space.table; val empty : T = Name_Space.empty_table Markup.ML_antiquotationN; fun merge data : T = Name_Space.merge_tables data; ); val get_antiquotations = Antiquotations.get o Proof_Context.theory_of; fun check_antiquotation ctxt = #1 o Name_Space.check (Context.Proof ctxt) (get_antiquotations ctxt); fun add_antiquotation name embedded antiquotation thy = thy |> Antiquotations.map (Name_Space.define (Context.Theory thy) true (name, (embedded, antiquotation)) #> #2); fun print_antiquotations verbose ctxt = Pretty.big_list "ML antiquotations:" (map (Pretty.mark_str o #1) (Name_Space.markup_table verbose ctxt (get_antiquotations ctxt))) |> Pretty.writeln; fun apply_antiquotation range src ctxt = let val (src', (embedded, antiquotation)) = Token.check_src ctxt get_antiquotations src; val _ = if Options.default_bool "update_control_cartouches" then Context_Position.reports_text ctxt (Antiquote.update_reports embedded (#1 range) (map Token.content_of src')) else (); in antiquotation range src' ctxt end; (* parsing *) local val antiq = Parse.!!! ((Parse.token Parse.liberal_name ::: Parse.args) --| Scan.ahead Parse.eof); fun expand_antiquote ant ctxt = (case ant of Antiquote.Text tok => (K ([], [tok]), ctxt) | Antiquote.Control {name, range, body} => ctxt |> apply_antiquotation range (Token.make_src name (if null body then [] else [Token.read_cartouche body])) | Antiquote.Antiq {range, body, ...} => ctxt |> apply_antiquotation range (Parse.read_antiq (Thy_Header.get_keywords' ctxt) antiq (body, #1 range))); in fun expand_antiquotes ants ctxt = let val (decls, ctxt') = fold_map expand_antiquote ants ctxt; fun decl ctxt'' = decls |> map (fn d => d ctxt'') |> split_list |> apply2 flat; in (decl, ctxt') end; fun expand_antiquotes_list antss ctxt = let val (decls, ctxt') = fold_map expand_antiquotes antss ctxt; fun decl' ctxt'' = map (fn decl => decl ctxt'') decls; in (decl', ctxt') end end; (* evaluation *) local fun make_env name visible = (ML_Lex.tokenize ("structure " ^ name ^ " =\nstruct\n\ \val ML_context = Context_Position.set_visible " ^ Bool.toString visible ^ " (Context.the_local_context ());\n"), ML_Lex.tokenize "end;"); fun reset_env name = ML_Lex.tokenize ("structure " ^ name ^ " = struct end"); fun eval_antiquotes ants opt_context = if forall Antiquote.is_text ants orelse is_none opt_context then (([], map (Antiquote.the_text "No context -- cannot expand ML antiquotations") ants), Option.map Context.proof_of opt_context) else let val context = the opt_context; val visible = (case context of Context.Proof ctxt => Context_Position.is_visible ctxt | _ => true); val ctxt = struct_begin (Context.proof_of context); val (begin_env, end_env) = make_env (struct_name ctxt) visible; val (decl, ctxt') = expand_antiquotes ants ctxt; val (ml_env, ml_body) = decl ctxt'; in ((begin_env @ ml_env @ end_env, ml_body), SOME ctxt') end; in fun eval flags pos ants = let val non_verbose = ML_Compiler.verbose false flags; (*prepare source text*) val ((env, body), env_ctxt) = eval_antiquotes ants (Context.get_generic_context ()); val _ = (case env_ctxt of SOME ctxt => if Config.get ctxt ML_Options.source_trace andalso Context_Position.is_visible ctxt then tracing (cat_lines [ML_Lex.flatten env, ML_Lex.flatten body]) else () | NONE => ()); (*prepare environment*) val _ = Context.setmp_generic_context (Option.map (Context.Proof o Context_Position.set_visible false) env_ctxt) (fn () => (ML_Compiler.eval non_verbose Position.none env; Context.get_generic_context ())) () |> (fn NONE => () | SOME context' => Context.>> (ML_Env.inherit [context'])); (*eval body*) val _ = ML_Compiler.eval flags pos body; (*clear environment*) val _ = (case (env_ctxt, is_some (Context.get_generic_context ())) of (SOME ctxt, true) => let val name = struct_name ctxt; val _ = ML_Compiler.eval non_verbose Position.none (reset_env name); val _ = Context.>> (ML_Env.forget_structure name); in () end | _ => ()); in () end; end; (* derived versions *) fun eval_file flags path = let val pos = Position.file (File.symbolic_path path) in eval flags pos (ML_Lex.read_text (File.read path, pos)) end; fun eval_source flags source = let val opt_context = Context.get_generic_context (); val {read_source, ...} = ML_Env.operations opt_context (#environment flags); in eval flags (Input.pos_of source) (read_source source) end; fun eval_in ctxt flags pos ants = Context.setmp_generic_context (Option.map Context.Proof ctxt) (fn () => eval flags pos ants) (); fun eval_source_in ctxt flags source = Context.setmp_generic_context (Option.map Context.Proof ctxt) (fn () => eval_source flags source) (); fun exec (e: unit -> unit) context = (case Context.setmp_generic_context (SOME context) (fn () => (e (); Context.get_generic_context ())) () of SOME context' => context' | NONE => error "Missing context after execution"); -fun expression pos ants = exec (fn () => eval ML_Compiler.flags pos ants); +fun expression pos ants = + Local_Theory.touch_ml_env #> + exec (fn () => eval ML_Compiler.flags pos ants); end; val ML = ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags); diff --git a/src/Pure/ML/ml_env.ML b/src/Pure/ML/ml_env.ML --- a/src/Pure/ML/ml_env.ML +++ b/src/Pure/ML/ml_env.ML @@ -1,328 +1,331 @@ (* Title: Pure/ML/ml_env.ML Author: Makarius Toplevel environment for Standard ML and Isabelle/ML within the implicit context. *) signature ML_ENV = sig val Isabelle: string val SML: string val ML_environment: string Config.T val ML_read_global: bool Config.T val ML_write_global: bool Config.T val inherit: Context.generic list -> Context.generic -> Context.generic type operations = {read_source: Input.source -> ML_Lex.token Antiquote.antiquote list, explode_token: ML_Lex.token -> char list, ML_system: bool} type environment = {read: string, write: string} val parse_environment: Context.generic option -> string -> environment val print_environment: environment -> string val SML_export: string val SML_import: string val setup: string -> operations -> theory -> theory val Isabelle_operations: operations val SML_operations: operations val operations: Context.generic option -> string -> operations + val touch: Context.generic -> Context.generic val forget_structure: string -> Context.generic -> Context.generic val bootstrap_name_space: Context.generic -> Context.generic val add_name_space: string -> ML_Name_Space.T -> Context.generic -> Context.generic val make_name_space: string -> ML_Name_Space.T val context: ML_Compiler0.context val name_space: ML_Name_Space.T val check_functor: string -> unit val get_breakpoint: Context.generic -> serial -> (bool Unsynchronized.ref * Position.T) option val add_breakpoints: (serial * (bool Unsynchronized.ref * Thread_Position.T)) list -> unit end structure ML_Env: ML_ENV = struct (* named environments *) val Isabelle = "Isabelle"; val SML = "SML"; val ML_environment = Config.declare_string ("ML_environment", \<^here>) (K Isabelle); (* global read/write *) val ML_read_global = Config.declare_bool ("ML_read_global", \<^here>) (K true); val ML_write_global = Config.declare_bool ("ML_write_global", \<^here>) (K true); val read_global = Config.apply_generic ML_read_global; val write_global = Config.apply_generic ML_write_global; (* name space tables *) type tables = PolyML.NameSpace.Values.value Symtab.table * PolyML.NameSpace.TypeConstrs.typeConstr Symtab.table * PolyML.NameSpace.Infixes.fixity Symtab.table * PolyML.NameSpace.Structures.structureVal Symtab.table * PolyML.NameSpace.Signatures.signatureVal Symtab.table * PolyML.NameSpace.Functors.functorVal Symtab.table; val empty_tables: tables = (Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty); fun merge_tables ((val1, type1, fixity1, structure1, signature1, functor1), (val2, type2, fixity2, structure2, signature2, functor2)) : tables = (Symtab.merge (K true) (val1, val2), Symtab.merge (K true) (type1, type2), Symtab.merge (K true) (fixity1, fixity2), Symtab.merge (K true) (structure1, structure2), Symtab.merge (K true) (signature1, signature2), Symtab.merge (K true) (functor1, functor2)); fun update_tables_values vals (val1, type1, fixity1, structure1, signature1, functor1) : tables = (fold Symtab.update vals val1, type1, fixity1, structure1, signature1, functor1); val sml_tables: tables = (Symtab.make ML_Name_Space.sml_val, Symtab.make ML_Name_Space.sml_type, Symtab.make ML_Name_Space.sml_fixity, Symtab.make ML_Name_Space.sml_structure, Symtab.make ML_Name_Space.sml_signature, Symtab.make ML_Name_Space.sml_functor); (* context data *) type operations = {read_source: Input.source -> ML_Lex.token Antiquote.antiquote list, explode_token: ML_Lex.token -> char list, ML_system: bool}; local type env = tables * operations; type data = env Name_Space.table * (bool Unsynchronized.ref * Position.T) Inttab.table; val empty_data: data = (Name_Space.empty_table "ML_environment", Inttab.empty); fun merge_data ((envs1, breakpoints1), (envs2, breakpoints2)) : data = ((envs1, envs2) |> Name_Space.join_tables (K (fn ((tables1, ops), (tables2, _)) => (merge_tables (tables1, tables2), ops))), Inttab.merge (K true) (breakpoints1, breakpoints2)); in structure Data = Generic_Data ( type T = data; val empty = empty_data; val merge = merge_data; ); fun inherit contexts = Data.put (Library.foldl1 merge_data (map Data.get contexts)); end; val get_env = Name_Space.get o #1 o Data.get; val get_tables = #1 oo get_env; +val touch = Data.map I; + fun update_tables env_name f context = context |> (Data.map o apfst) (fn envs => let val _ = Name_Space.get envs env_name; in Name_Space.map_table_entry env_name (apfst f) envs end); fun forget_structure name context = (if write_global context then ML_Name_Space.forget_structure name else (); context |> update_tables Isabelle (fn tables => (#1 tables, #2 tables, #3 tables, Symtab.delete_safe name (#4 tables), #5 tables, #6 tables))); (* environment name *) type environment = {read: string, write: string}; val separator = ">"; fun parse_environment opt_context environment = let fun check name = (case opt_context of NONE => if name = Isabelle then name else error ("Bad ML environment name " ^ quote name ^ " -- no context") | SOME context => if name = Isabelle then name else (get_env context name; name)); val spec = if environment = "" then (case opt_context of NONE => Isabelle | SOME context => Config.get_generic context ML_environment) else environment; val (read, write) = (case space_explode separator spec of [a] => (a, a) | [a, b] => (a, b) | _ => error ("Malformed ML environment specification: " ^ quote spec)); in {read = check read, write = check write} end; fun print_environment {read, write} = read ^ separator ^ write; val SML_export = print_environment {read = SML, write = Isabelle}; val SML_import = print_environment {read = Isabelle, write = SML}; (* setup operations *) val ML_system_values = #allVal (ML_Name_Space.global) () |> filter (member (op =) ["ML_system_pretty", "ML_system_pp", "ML_system_overload"] o #1); fun setup env_name ops thy = thy |> (Context.theory_map o Data.map o apfst) (fn envs => let val thy' = Sign.map_naming (K Name_Space.global_naming) thy; val tables = (if env_name = Isabelle then empty_tables else sml_tables) |> #ML_system ops ? update_tables_values ML_system_values; val (_, envs') = Name_Space.define (Context.Theory thy') true (Binding.name env_name, (tables, ops)) envs; in envs' end); val Isabelle_operations: operations = {read_source = ML_Lex.read_source, explode_token = ML_Lex.check_content_of #> Symbol.explode #> maps (Symbol.esc #> String.explode), ML_system = false}; val SML_operations: operations = {read_source = ML_Lex.read_source_sml, explode_token = ML_Lex.check_content_of #> String.explode, ML_system = false}; val _ = Theory.setup (setup Isabelle Isabelle_operations #> setup SML SML_operations); fun operations opt_context environment = let val env = #read (parse_environment opt_context environment) in (case opt_context of NONE => Isabelle_operations | SOME context => #2 (get_env context env)) end; (* name space *) val bootstrap_name_space = update_tables Isabelle (fn (tables: tables) => let fun update entries (x, y) = member (op =) entries x ? Symtab.update (x, y); val bootstrap_val = update ML_Name_Space.bootstrap_values; val bootstrap_structure = update ML_Name_Space.bootstrap_structures; val bootstrap_signature = update ML_Name_Space.bootstrap_signatures; val (val1, type1, fixity1, structure1, signature1, functor1) = sml_tables; val val2 = val1 |> fold bootstrap_val (#allVal ML_Name_Space.global ()) |> Symtab.fold bootstrap_val (#1 tables); val structure2 = structure1 |> fold bootstrap_structure (#allStruct ML_Name_Space.global ()) |> Symtab.fold bootstrap_structure (#4 tables); val signature2 = signature1 |> fold bootstrap_signature (#allSig ML_Name_Space.global ()) |> Symtab.fold bootstrap_signature (#5 tables); in (val2, type1, fixity1, structure2, signature2, functor1) end); fun add_name_space env (space: ML_Name_Space.T) = update_tables env (fn (val1, type1, fixity1, structure1, signature1, functor1) => let val val2 = fold Symtab.update (#allVal space ()) val1; val type2 = fold Symtab.update (#allType space ()) type1; val fixity2 = fold Symtab.update (#allFix space ()) fixity1; val structure2 = fold Symtab.update (#allStruct space ()) structure1; val signature2 = fold Symtab.update (#allSig space ()) signature1; val functor2 = fold Symtab.update (#allFunct space ()) functor1; in (val2, type2, fixity2, structure2, signature2, functor2) end); fun make_name_space environment : ML_Name_Space.T = let val {read, write} = parse_environment (Context.get_generic_context ()) environment; fun lookup_env sel1 context name = Symtab.lookup (sel1 (get_tables context read)) name; fun dest_env sel1 context = Symtab.dest (sel1 (get_tables context read)); fun enter_env ap1 entry = update_tables write (ap1 (Symtab.update entry)); fun lookup sel1 sel2 name = if read = Isabelle then (case Context.get_generic_context () of NONE => sel2 ML_Name_Space.global name | SOME context => (case lookup_env sel1 context name of NONE => if read_global context then sel2 ML_Name_Space.global name else NONE | some => some)) else lookup_env sel1 (Context.the_generic_context ()) name; fun all sel1 sel2 () = sort_distinct (string_ord o apply2 #1) (if read = Isabelle then (case Context.get_generic_context () of NONE => sel2 ML_Name_Space.global () | SOME context => dest_env sel1 context @ (if read_global context then sel2 ML_Name_Space.global () else [])) else dest_env sel1 (Context.the_generic_context ())); fun enter ap1 sel2 entry = if write = Isabelle then (case Context.get_generic_context () of NONE => sel2 ML_Name_Space.global entry | SOME context => (if write_global context then sel2 ML_Name_Space.global entry else (); Context.>> (enter_env ap1 entry))) else Context.>> (enter_env ap1 entry); in {lookupVal = lookup #1 #lookupVal, lookupType = lookup #2 #lookupType, lookupFix = lookup #3 #lookupFix, lookupStruct = lookup #4 #lookupStruct, lookupSig = lookup #5 #lookupSig, lookupFunct = lookup #6 #lookupFunct, enterVal = enter (fn h => fn (a, b, c, d, e, f) => (h a, b, c, d, e, f)) #enterVal, enterType = enter (fn h => fn (a, b, c, d, e, f) => (a, h b, c, d, e, f)) #enterType, enterFix = enter (fn h => fn (a, b, c, d, e, f) => (a, b, h c, d, e, f)) #enterFix, enterStruct = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, h d, e, f)) #enterStruct, enterSig = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, h e, f)) #enterSig, enterFunct = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, e, h f)) #enterFunct, allVal = all #1 #allVal, allType = all #2 #allType, allFix = all #3 #allFix, allStruct = all #4 #allStruct, allSig = all #5 #allSig, allFunct = all #6 #allFunct} end; val context: ML_Compiler0.context = {name_space = make_name_space "", print_depth = NONE, here = Position.here oo Position.line_file, print = writeln, error = error}; val name_space = #name_space context; val is_functor = is_some o #lookupFunct name_space; fun check_functor name = if not (is_functor "Table") (*mask dummy version of name_space*) orelse is_functor name then () else error ("Unknown ML functor: " ^ quote name); (* breakpoints *) val get_breakpoint = Inttab.lookup o #2 o Data.get; fun add_breakpoints more_breakpoints = if is_some (Context.get_generic_context ()) then Context.>> ((Data.map o apsnd) (fold (Inttab.update_new o (apsnd o apsnd) Position.make) more_breakpoints)) else (); end; diff --git a/src/Pure/ML/ml_file.ML b/src/Pure/ML/ml_file.ML --- a/src/Pure/ML/ml_file.ML +++ b/src/Pure/ML/ml_file.ML @@ -1,39 +1,40 @@ (* Title: Pure/ML/ml_file.ML Author: Makarius Commands to load ML files. *) signature ML_FILE = sig val command: string -> bool option -> (theory -> Token.file) -> Toplevel.transition -> Toplevel.transition val ML: bool option -> (theory -> Token.file) -> Toplevel.transition -> Toplevel.transition val SML: bool option -> (theory -> Token.file) -> Toplevel.transition -> Toplevel.transition end; structure ML_File: ML_FILE = struct fun command environment debug get_file = Toplevel.generic_theory (fn gthy => let val file = get_file (Context.theory_of gthy); val provide = Resources.provide_file file; val source = Token.file_source file; val _ = Document_Output.check_comments (Context.proof_of gthy) (Input.source_explode source); val flags: ML_Compiler.flags = {environment = environment, redirect = true, verbose = true, debug = debug, writeln = writeln, warning = warning}; in gthy + |> Local_Theory.touch_ml_env |> ML_Context.exec (fn () => ML_Context.eval_source flags source) |> Local_Theory.propagate_ml_env |> Context.mapping provide (Local_Theory.background_theory provide) end); val ML = command ""; val SML = command ML_Env.SML; end; diff --git a/src/Pure/PIDE/markup.ML b/src/Pure/PIDE/markup.ML --- a/src/Pure/PIDE/markup.ML +++ b/src/Pure/PIDE/markup.ML @@ -1,872 +1,874 @@ (* Title: Pure/PIDE/markup.ML Author: Makarius Quasi-abstract markup elements. *) signature MARKUP = sig type T = string * Properties.T val empty: T val is_empty: T -> bool val properties: Properties.T -> T -> T val nameN: string val name: string -> T -> T val xnameN: string val xname: string -> T -> T val kindN: string val serialN: string val serial_properties: int -> Properties.T val instanceN: string val meta_titleN: string val meta_title: T val meta_creatorN: string val meta_creator: T val meta_contributorN: string val meta_contributor: T val meta_dateN: string val meta_date: T val meta_licenseN: string val meta_license: T val meta_descriptionN: string val meta_description: T val languageN: string val symbolsN: string val delimitedN: string val is_delimited: Properties.T -> bool val language: {name: string, symbols: bool, antiquotes: bool, delimited: bool} -> T val language': {name: string, symbols: bool, antiquotes: bool} -> bool -> T val language_Isar: bool -> T val language_method: T val language_attribute: T val language_sort: bool -> T val language_type: bool -> T val language_term: bool -> T val language_prop: bool -> T val language_ML: bool -> T val language_SML: bool -> T val language_document: bool -> T val language_document_marker: T val language_antiquotation: T val language_text: bool -> T val language_verbatim: bool -> T val language_latex: bool -> T val language_rail: T val language_path: bool -> T val language_url: bool -> T val language_mixfix: T val bindingN: string val binding: T val entityN: string val entity: string -> string -> T val defN: string val refN: string val completionN: string val completion: T val no_completionN: string val no_completion: T val updateN: string val update: T val lineN: string val end_lineN: string val offsetN: string val end_offsetN: string + val labelN: string val fileN: string val idN: string val positionN: string val position: T val position_properties: string list val position_property: Properties.entry -> bool val def_name: string -> string val expressionN: string val expression: string -> T val citationN: string val citation: string -> T val pathN: string val path: string -> T val export_pathN: string val export_path: string -> T val urlN: string val url: string -> T val docN: string val doc: string -> T val toolN: string val tool: string -> T val markupN: string val consistentN: string val unbreakableN: string val block_properties: string list val indentN: string val widthN: string val blockN: string val block: bool -> int -> T val breakN: string val break: int -> int -> T val fbreakN: string val fbreak: T val itemN: string val item: T val wordsN: string val words: T val hiddenN: string val hidden: T val deleteN: string val delete: T val load_commandN: string val bash_functionN: string val bibtex_entryN: string val scala_functionN: string val system_optionN: string val sessionN: string val theoryN: string val classN: string val localeN: string val bundleN: string val type_nameN: string val constantN: string val axiomN: string val factN: string val oracleN: string val fixedN: string val fixed: string -> T val caseN: string val case_: string -> T val dynamic_factN: string val dynamic_fact: string -> T val literal_factN: string val literal_fact: string -> T val attributeN: string val methodN: string val method_modifierN: string val tfreeN: string val tfree: T val tvarN: string val tvar: T val freeN: string val free: T val skolemN: string val skolem: T val boundN: string val bound: T val varN: string val var: T val numeralN: string val numeral: T val literalN: string val literal: T val delimiterN: string val delimiter: T val inner_stringN: string val inner_string: T val inner_cartoucheN: string val inner_cartouche: T val token_rangeN: string val token_range: T val sortingN: string val sorting: T val typingN: string val typing: T val class_parameterN: string val class_parameter: T val ML_keyword1N: string val ML_keyword1: T val ML_keyword2N: string val ML_keyword2: T val ML_keyword3N: string val ML_keyword3: T val ML_delimiterN: string val ML_delimiter: T val ML_tvarN: string val ML_tvar: T val ML_numeralN: string val ML_numeral: T val ML_charN: string val ML_char: T val ML_stringN: string val ML_string: T val ML_commentN: string val ML_comment: T val ML_defN: string val ML_openN: string val ML_structureN: string val ML_typingN: string val ML_typing: T val ML_breakpointN: string val ML_breakpoint: int -> T val antiquotedN: string val antiquoted: T val antiquoteN: string val antiquote: T val file_typeN: string val antiquotationN: string val ML_antiquotationN: string val document_antiquotationN: string val document_antiquotation_optionN: string val prismjs_languageN: string val raw_textN: string val raw_text: T val plain_textN: string val plain_text: T val paragraphN: string val paragraph: T val text_foldN: string val text_fold: T val document_markerN: string val document_marker: T val document_tagN: string val document_tag: string -> T val document_latexN: string val document_latex: T val latex_outputN: string val latex_output: T val latex_macro0N: string val latex_macro0: string -> T val latex_macroN: string val latex_macro: string -> T val latex_environmentN: string val latex_environment: string -> T val latex_headingN: string val latex_heading: string -> T val latex_bodyN: string val latex_body: string -> T val latex_citeN: string val latex_cite: {kind: string, citations: string} -> T val latex_index_itemN: string val latex_index_item: T val latex_index_entryN: string val latex_index_entry: string -> T val latex_delimN: string val latex_delim: string -> T val latex_tagN: string val latex_tag: string -> T val optional_argumentN: string val optional_argument: string -> T -> T val markdown_paragraphN: string val markdown_paragraph: T val markdown_itemN: string val markdown_item: T val markdown_bulletN: string val markdown_bullet: int -> T val markdown_listN: string val markdown_list: string -> T val itemizeN: string val enumerateN: string val descriptionN: string val inputN: string val input: bool -> Properties.T -> T val command_keywordN: string val command_keyword: T val command_spanN: string val command_span: string -> T val commandN: string val command_properties: T -> T val keywordN: string val keyword_properties: T -> T val stringN: string val string: T val alt_stringN: string val alt_string: T val verbatimN: string val verbatim: T val cartoucheN: string val cartouche: T val commentN: string val comment: T val keyword1N: string val keyword1: T val keyword2N: string val keyword2: T val keyword3N: string val keyword3: T val quasi_keywordN: string val quasi_keyword: T val improperN: string val improper: T val operatorN: string val operator: T val comment1N: string val comment1: T val comment2N: string val comment2: T val comment3N: string val comment3: T val elapsedN: string val cpuN: string val gcN: string val timing_properties: {elapsed: Time.time, cpu: Time.time, gc: Time.time} -> Properties.T val parse_command_timing_properties: Properties.T -> ({file: string, offset: int, name: string} * Time.time) option val timingN: string val timing: {elapsed: Time.time, cpu: Time.time, gc: Time.time} -> T val command_indentN: string val command_indent: int -> T val goalN: string val goal: T val subgoalN: string val subgoal: string -> T val taskN: string val forkedN: string val forked: T val joinedN: string val joined: T val runningN: string val running: T val finishedN: string val finished: T val failedN: string val failed: T val canceledN: string val canceled: T val initializedN: string val initialized: T val finalizedN: string val finalized: T val consolidatingN: string val consolidating: T val consolidatedN: string val consolidated: T val exec_idN: string val initN: string val statusN: string val status: T val resultN: string val result: T val writelnN: string val writeln: T val stateN: string val state: T val informationN: string val information: T val tracingN: string val tracing: T val warningN: string val warning: T val legacyN: string val legacy: T val errorN: string val error: T val systemN: string val system: T val protocolN: string val reportN: string val report: T val no_reportN: string val no_report: T val badN: string val bad: unit -> T val intensifyN: string val intensify: T val countN: string val ML_profiling_entryN: string val ML_profiling_entry: {name: string, count: int} -> T val ML_profilingN: string val ML_profiling: string -> T val browserN: string val graphviewN: string val theory_exportsN: string val sendbackN: string val paddingN: string val padding_line: Properties.entry val padding_command: Properties.entry val dialogN: string val dialog: serial -> string -> T val jedit_actionN: string val function: string -> Properties.entry val ML_statistics: {pid: int, stats_dir: string} -> Properties.T val commands_accepted: Properties.T val assign_update: Properties.T val removed_versions: Properties.T val invoke_scala: string -> string -> Properties.T val cancel_scala: string -> Properties.T val task_statistics: Properties.entry val command_timing: Properties.entry val theory_timing: Properties.entry val session_timing: Properties.entry val loading_theory: string -> Properties.T val build_session_finished: Properties.T val print_operations: Properties.T val exportN: string type export_args = {id: string option, serial: serial, theory_name: string, name: string, executable: bool, compress: bool, strict: bool} val export: export_args -> Properties.T val debugger_state: string -> Properties.T val debugger_output: string -> Properties.T val simp_trace_panelN: string val simp_trace_logN: string val simp_trace_stepN: string val simp_trace_recurseN: string val simp_trace_hintN: string val simp_trace_ignoreN: string val simp_trace_cancel: serial -> Properties.T type output = Output.output * Output.output val no_output: output val add_mode: string -> (T -> output) -> unit val output: T -> output val enclose: T -> Output.output -> Output.output val markup: T -> string -> string val markups: T list -> string -> string val markup_only: T -> string val markup_report: string -> string end; structure Markup: MARKUP = struct (** markup elements **) (* basic markup *) type T = string * Properties.T; val empty = ("", []); fun is_empty ("", _) = true | is_empty _ = false; fun properties more_props ((elem, props): T) = (elem, fold_rev Properties.put more_props props); fun markup_elem name = (name, (name, []): T); fun markup_string name prop = (name, fn s => (name, [(prop, s)]): T); fun markup_int name prop = (name, fn i => (name, [(prop, Value.print_int i)]): T); (* misc properties *) val nameN = "name"; fun name a = properties [(nameN, a)]; val xnameN = "xname"; fun xname a = properties [(xnameN, a)]; val kindN = "kind"; val serialN = "serial"; fun serial_properties i = [(serialN, Value.print_int i)]; val instanceN = "instance"; (* meta data -- see https://www.dublincore.org/specifications/dublin-core/dcmi-terms *) val (meta_titleN, meta_title) = markup_elem "meta_title"; val (meta_creatorN, meta_creator) = markup_elem "meta_creator"; val (meta_contributorN, meta_contributor) = markup_elem "meta_contributor"; val (meta_dateN, meta_date) = markup_elem "meta_date"; val (meta_licenseN, meta_license) = markup_elem "meta_license"; val (meta_descriptionN, meta_description) = markup_elem "meta_description"; (* embedded languages *) val languageN = "language"; val symbolsN = "symbols"; val antiquotesN = "antiquotes"; val delimitedN = "delimited" fun is_delimited props = Properties.get props delimitedN = SOME "true"; fun language {name, symbols, antiquotes, delimited} = (languageN, [(nameN, name), (symbolsN, Value.print_bool symbols), (antiquotesN, Value.print_bool antiquotes), (delimitedN, Value.print_bool delimited)]); fun language' {name, symbols, antiquotes} delimited = language {name = name, symbols = symbols, antiquotes = antiquotes, delimited = delimited}; val language_Isar = language' {name = "Isar", symbols = true, antiquotes = false}; val language_method = language {name = "method", symbols = true, antiquotes = false, delimited = false}; val language_attribute = language {name = "attribute", symbols = true, antiquotes = false, delimited = false}; val language_sort = language' {name = "sort", symbols = true, antiquotes = false}; val language_type = language' {name = "type", symbols = true, antiquotes = false}; val language_term = language' {name = "term", symbols = true, antiquotes = false}; val language_prop = language' {name = "prop", symbols = true, antiquotes = false}; val language_ML = language' {name = "ML", symbols = false, antiquotes = true}; val language_SML = language' {name = "SML", symbols = false, antiquotes = false}; val language_document = language' {name = "document", symbols = false, antiquotes = true}; val language_document_marker = language {name = "document_marker", symbols = true, antiquotes = true, delimited = true}; val language_antiquotation = language {name = "antiquotation", symbols = true, antiquotes = false, delimited = true}; val language_text = language' {name = "text", symbols = true, antiquotes = false}; val language_verbatim = language' {name = "verbatim_text", symbols = true, antiquotes = false}; val language_latex = language' {name = "latex", symbols = false, antiquotes = false}; val language_rail = language {name = "rail", symbols = true, antiquotes = true, delimited = true}; val language_path = language' {name = "path", symbols = false, antiquotes = false}; val language_url = language' {name = "url", symbols = false, antiquotes = false}; val language_mixfix = language {name = "mixfix_annotation", symbols = true, antiquotes = false, delimited = true}; (* formal entities *) val (bindingN, binding) = markup_elem "binding"; val entityN = "entity"; fun entity kind name = (entityN, (if name = "" then [] else [(nameN, name)]) @ (if kind = "" then [] else [(kindN, kind)])); val defN = "def"; val refN = "ref"; (* completion *) val (completionN, completion) = markup_elem "completion"; val (no_completionN, no_completion) = markup_elem "no_completion"; val (updateN, update) = markup_elem "update"; (* position *) val lineN = "line"; val end_lineN = "end_line"; val offsetN = "offset"; val end_offsetN = "end_offset"; +val labelN = "label"; val fileN = "file"; val idN = "id"; val (positionN, position) = markup_elem "position"; -val position_properties = [lineN, offsetN, end_offsetN, fileN, idN]; +val position_properties = [lineN, offsetN, end_offsetN, labelN, fileN, idN]; fun position_property (entry: Properties.entry) = member (op =) position_properties (#1 entry); (* position "def" names *) fun make_def a = "def_" ^ a; val def_names = Symtab.make (map (fn a => (a, make_def a)) position_properties); fun def_name a = (case Symtab.lookup def_names a of SOME b => b | NONE => make_def a); (* expression *) val expressionN = "expression"; fun expression kind = (expressionN, if kind = "" then [] else [(kindN, kind)]); (* citation *) val (citationN, citation) = markup_string "citation" nameN; (* external resources *) val (pathN, path) = markup_string "path" nameN; val (export_pathN, export_path) = markup_string "export_path" nameN; val (urlN, url) = markup_string "url" nameN; val (docN, doc) = markup_string "doc" nameN; val (toolN, tool) = markup_string "tool" nameN; (* pretty printing *) val markupN = "markup"; val consistentN = "consistent"; val unbreakableN = "unbreakable"; val indentN = "indent"; val block_properties = [markupN, consistentN, unbreakableN, indentN]; val widthN = "width"; val blockN = "block"; fun block c i = (blockN, (if c then [(consistentN, Value.print_bool c)] else []) @ (if i <> 0 then [(indentN, Value.print_int i)] else [])); val breakN = "break"; fun break w i = (breakN, (if w <> 0 then [(widthN, Value.print_int w)] else []) @ (if i <> 0 then [(indentN, Value.print_int i)] else [])); val (fbreakN, fbreak) = markup_elem "fbreak"; val (itemN, item) = markup_elem "item"; (* text properties *) val (wordsN, words) = markup_elem "words"; val (hiddenN, hidden) = markup_elem "hidden"; val (deleteN, delete) = markup_elem "delete"; (* misc entities *) val load_commandN = "load_command"; val bash_functionN = "bash_function"; val bibtex_entryN = "bibtex_entry"; val scala_functionN = "scala_function"; val system_optionN = "system_option"; val sessionN = "session"; val theoryN = "theory"; val classN = "class"; val localeN = "locale"; val bundleN = "bundle"; val type_nameN = "type_name"; val constantN = "constant"; val axiomN = "axiom"; val factN = "fact"; val oracleN = "oracle"; val (fixedN, fixed) = markup_string "fixed" nameN; val (caseN, case_) = markup_string "case" nameN; val (dynamic_factN, dynamic_fact) = markup_string "dynamic_fact" nameN; val (literal_factN, literal_fact) = markup_string "literal_fact" nameN; val attributeN = "attribute"; val methodN = "method"; val method_modifierN = "method_modifier"; (* inner syntax *) val (tfreeN, tfree) = markup_elem "tfree"; val (tvarN, tvar) = markup_elem "tvar"; val (freeN, free) = markup_elem "free"; val (skolemN, skolem) = markup_elem "skolem"; val (boundN, bound) = markup_elem "bound"; val (varN, var) = markup_elem "var"; val (numeralN, numeral) = markup_elem "numeral"; val (literalN, literal) = markup_elem "literal"; val (delimiterN, delimiter) = markup_elem "delimiter"; val (inner_stringN, inner_string) = markup_elem "inner_string"; val (inner_cartoucheN, inner_cartouche) = markup_elem "inner_cartouche"; val (token_rangeN, token_range) = markup_elem "token_range"; val (sortingN, sorting) = markup_elem "sorting"; val (typingN, typing) = markup_elem "typing"; val (class_parameterN, class_parameter) = markup_elem "class_parameter"; (* ML *) val (ML_keyword1N, ML_keyword1) = markup_elem "ML_keyword1"; val (ML_keyword2N, ML_keyword2) = markup_elem "ML_keyword2"; val (ML_keyword3N, ML_keyword3) = markup_elem "ML_keyword3"; val (ML_delimiterN, ML_delimiter) = markup_elem "ML_delimiter"; val (ML_tvarN, ML_tvar) = markup_elem "ML_tvar"; val (ML_numeralN, ML_numeral) = markup_elem "ML_numeral"; val (ML_charN, ML_char) = markup_elem "ML_char"; val (ML_stringN, ML_string) = markup_elem "ML_string"; val (ML_commentN, ML_comment) = markup_elem "ML_comment"; val ML_defN = "ML_def"; val ML_openN = "ML_open"; val ML_structureN = "ML_structure"; val (ML_typingN, ML_typing) = markup_elem "ML_typing"; val (ML_breakpointN, ML_breakpoint) = markup_int "ML_breakpoint" serialN; (* antiquotations *) val (antiquotedN, antiquoted) = markup_elem "antiquoted"; val (antiquoteN, antiquote) = markup_elem "antiquote"; val file_typeN = "file_type"; val antiquotationN = "antiquotation"; val ML_antiquotationN = "ML_antiquotation"; val document_antiquotationN = "document_antiquotation"; val document_antiquotation_optionN = "document_antiquotation_option"; (* document text *) val prismjs_languageN = "prismjs_language"; val (raw_textN, raw_text) = markup_elem "raw_text"; val (plain_textN, plain_text) = markup_elem "plain_text"; val (paragraphN, paragraph) = markup_elem "paragraph"; val (text_foldN, text_fold) = markup_elem "text_fold"; val (document_markerN, document_marker) = markup_elem "document_marker"; val (document_tagN, document_tag) = markup_string "document_tag" nameN; (* LaTeX *) val (document_latexN, document_latex) = markup_elem "document_latex"; val (latex_outputN, latex_output) = markup_elem "latex_output"; val (latex_macro0N, latex_macro0) = markup_string "latex_macro0" nameN; val (latex_macroN, latex_macro) = markup_string "latex_macro" nameN; val (latex_environmentN, latex_environment) = markup_string "latex_environment" nameN; val (latex_headingN, latex_heading) = markup_string "latex_heading" kindN; val (latex_bodyN, latex_body) = markup_string "latex_body" kindN; val (latex_citeN, _) = markup_elem "latex_cite"; fun latex_cite {kind, citations} = (latex_citeN, (if kind = "" then [] else [(kindN, kind)]) @ (if citations = "" then [] else [("citations", citations)])); val (latex_index_itemN, latex_index_item) = markup_elem "latex_index_item"; val (latex_index_entryN, latex_index_entry) = markup_string "latex_index_entry" kindN; val (latex_delimN, latex_delim) = markup_string "latex_delim" nameN; val (latex_tagN, latex_tag) = markup_string "latex_tag" nameN; val optional_argumentN = "optional_argument"; fun optional_argument arg = properties [(optional_argumentN, arg)]; (* Markdown document structure *) val (markdown_paragraphN, markdown_paragraph) = markup_elem "markdown_paragraph"; val (markdown_itemN, markdown_item) = markup_elem "markdown_item"; val (markdown_bulletN, markdown_bullet) = markup_int "markdown_bullet" "depth"; val (markdown_listN, markdown_list) = markup_string "markdown_list" kindN; val itemizeN = "itemize"; val enumerateN = "enumerate"; val descriptionN = "description"; (* formal input *) val inputN = "input"; fun input delimited props = (inputN, (delimitedN, Value.print_bool delimited) :: props); (* outer syntax *) val (command_keywordN, command_keyword) = markup_elem "command_keyword"; val (command_spanN, command_span) = markup_string "command_span" nameN; val commandN = "command"; val command_properties = properties [(kindN, commandN)]; val keywordN = "keyword"; val keyword_properties = properties [(kindN, keywordN)]; val (keyword1N, keyword1) = markup_elem "keyword1"; val (keyword2N, keyword2) = markup_elem "keyword2"; val (keyword3N, keyword3) = markup_elem "keyword3"; val (quasi_keywordN, quasi_keyword) = markup_elem "quasi_keyword"; val (improperN, improper) = markup_elem "improper"; val (operatorN, operator) = markup_elem "operator"; val (stringN, string) = markup_elem "string"; val (alt_stringN, alt_string) = markup_elem "alt_string"; val (verbatimN, verbatim) = markup_elem "verbatim"; val (cartoucheN, cartouche) = markup_elem "cartouche"; val (commentN, comment) = markup_elem "comment"; (* comments *) val (comment1N, comment1) = markup_elem "comment1"; val (comment2N, comment2) = markup_elem "comment2"; val (comment3N, comment3) = markup_elem "comment3"; (* timing *) val elapsedN = "elapsed"; val cpuN = "cpu"; val gcN = "gc"; fun timing_properties {elapsed, cpu, gc} = [(elapsedN, Value.print_time elapsed), (cpuN, Value.print_time cpu), (gcN, Value.print_time gc)]; val timingN = "timing"; fun timing t = (timingN, timing_properties t); (* command timing *) fun parse_command_timing_properties props = (case (Properties.get props fileN, Properties.get props offsetN, Properties.get props nameN) of (SOME file, SOME offset, SOME name) => SOME ({file = file, offset = Value.parse_int offset, name = name}, Properties.get_seconds props elapsedN) | _ => NONE); (* indentation *) val (command_indentN, command_indent) = markup_int "command_indent" indentN; (* goals *) val (goalN, goal) = markup_elem "goal"; val (subgoalN, subgoal) = markup_string "subgoal" nameN; (* command status *) val taskN = "task"; val (forkedN, forked) = markup_elem "forked"; val (joinedN, joined) = markup_elem "joined"; val (runningN, running) = markup_elem "running"; val (finishedN, finished) = markup_elem "finished"; val (failedN, failed) = markup_elem "failed"; val (canceledN, canceled) = markup_elem "canceled"; val (initializedN, initialized) = markup_elem "initialized"; val (finalizedN, finalized) = markup_elem "finalized"; val (consolidatingN, consolidating) = markup_elem "consolidating"; val (consolidatedN, consolidated) = markup_elem "consolidated"; (* messages *) val exec_idN = "exec_id"; val initN = "init"; val (statusN, status) = markup_elem "status"; val (resultN, result) = markup_elem "result"; val (writelnN, writeln) = markup_elem "writeln"; val (stateN, state) = markup_elem "state" val (informationN, information) = markup_elem "information"; val (tracingN, tracing) = markup_elem "tracing"; val (warningN, warning) = markup_elem "warning"; val (legacyN, legacy) = markup_elem "legacy"; val (errorN, error) = markup_elem "error"; val (systemN, system) = markup_elem "system"; val protocolN = "protocol"; val (reportN, report) = markup_elem "report"; val (no_reportN, no_report) = markup_elem "no_report"; val badN = "bad"; fun bad () = (badN, serial_properties (serial ())); val (intensifyN, intensify) = markup_elem "intensify"; (* ML profiling *) val countN = "count"; val ML_profiling_entryN = "ML_profiling_entry"; fun ML_profiling_entry {name, count} = (ML_profiling_entryN, [(nameN, name), (countN, Value.print_int count)]); val (ML_profilingN, ML_profiling) = markup_string "ML_profiling" kindN; (* active areas *) val browserN = "browser" val graphviewN = "graphview"; val theory_exportsN = "theory_exports"; val sendbackN = "sendback"; val paddingN = "padding"; val padding_line = (paddingN, "line"); val padding_command = (paddingN, "command"); val dialogN = "dialog"; fun dialog i result = (dialogN, [(serialN, Value.print_int i), (resultN, result)]); val jedit_actionN = "jedit_action"; (* protocol message functions *) fun function name = ("function", name); fun ML_statistics {pid, stats_dir} = [function "ML_statistics", ("pid", Value.print_int pid), ("stats_dir", stats_dir)]; val commands_accepted = [function "commands_accepted"]; val assign_update = [function "assign_update"]; val removed_versions = [function "removed_versions"]; fun invoke_scala name id = [function "invoke_scala", (nameN, name), (idN, id)]; fun cancel_scala id = [function "cancel_scala", (idN, id)]; val task_statistics = function "task_statistics"; val command_timing = function "command_timing"; val theory_timing = function "theory_timing"; val session_timing = function "session_timing"; fun loading_theory name = [function "loading_theory", (nameN, name)]; val build_session_finished = [function "build_session_finished"]; val print_operations = [function "print_operations"]; (* export *) val exportN = "export"; type export_args = {id: string option, serial: serial, theory_name: string, name: string, executable: bool, compress: bool, strict: bool}; fun export ({id, serial, theory_name, name, executable, compress, strict}: export_args) = [function exportN, (idN, the_default "" id), (serialN, Value.print_int serial), ("theory_name", theory_name), (nameN, name), ("executable", Value.print_bool executable), ("compress", Value.print_bool compress), ("strict", Value.print_bool strict)]; (* debugger *) fun debugger_state name = [function "debugger_state", (nameN, name)]; fun debugger_output name = [function "debugger_output", (nameN, name)]; (* simplifier trace *) val simp_trace_panelN = "simp_trace_panel"; val simp_trace_logN = "simp_trace_log"; val simp_trace_stepN = "simp_trace_step"; val simp_trace_recurseN = "simp_trace_recurse"; val simp_trace_hintN = "simp_trace_hint"; val simp_trace_ignoreN = "simp_trace_ignore"; fun simp_trace_cancel i = [function "simp_trace_cancel", (serialN, Value.print_int i)]; (** print mode operations **) type output = Output.output * Output.output; val no_output = ("", ""); local val default = {output = Output_Primitives.markup_fn}; val modes = Synchronized.var "Markup.modes" (Symtab.make [("", default)]); in fun add_mode name output = Synchronized.change modes (fn tab => (if not (Symtab.defined tab name) then () else Output.warning ("Redefining markup mode " ^ quote name); Symtab.update (name, {output = output}) tab)); fun get_mode () = the_default default (Library.get_first (Symtab.lookup (Synchronized.value modes)) (print_mode_value ())); end; fun output m = if is_empty m then no_output else #output (get_mode ()) m; val enclose = output #-> Library.enclose; fun markup m = let val (bg, en) = output m in Library.enclose (Output.escape bg) (Output.escape en) end; val markups = fold_rev markup; fun markup_only m = markup m ""; fun markup_report "" = "" | markup_report txt = markup report txt; end; diff --git a/src/Pure/PIDE/markup.scala b/src/Pure/PIDE/markup.scala --- a/src/Pure/PIDE/markup.scala +++ b/src/Pure/PIDE/markup.scala @@ -1,773 +1,775 @@ /* Title: Pure/PIDE/markup.scala Author: Makarius Quasi-abstract markup elements. */ package isabelle object Markup { /* elements */ object Elements { def apply(elems: Set[String]): Elements = new Elements(elems) def apply(elems: String*): Elements = apply(Set(elems: _*)) val empty: Elements = apply() val full: Elements = new Elements(Set.empty) { override def apply(elem: String): Boolean = true override def toString: String = "Elements.full" } } sealed class Elements private[Markup](private val rep: Set[String]) { def apply(elem: String): Boolean = rep.contains(elem) def + (elem: String): Elements = new Elements(rep + elem) def ++ (elems: Elements): Elements = new Elements(rep ++ elems.rep) def - (elem: String): Elements = new Elements(rep - elem) def -- (elems: Elements): Elements = new Elements(rep -- elems.rep) override def toString: String = rep.mkString("Elements(", ",", ")") } /* properties */ val NAME = "name" val Name = new Properties.String(NAME) val XNAME = "xname" val XName = new Properties.String(XNAME) val KIND = "kind" val Kind = new Properties.String(KIND) val CONTENT = "content" val Content = new Properties.String(CONTENT) val SERIAL = "serial" val Serial = new Properties.Long(SERIAL) val INSTANCE = "instance" val Instance = new Properties.String(INSTANCE) /* basic markup */ val Empty: Markup = Markup("", Nil) val Broken: Markup = Markup("broken", Nil) class Markup_Elem(val name: String) { def apply(props: Properties.T = Nil): Markup = Markup(name, props) def unapply(markup: Markup): Option[Properties.T] = if (markup.name == name) Some(markup.properties) else None } class Markup_String(val name: String, prop: String) { val Prop: Properties.String = new Properties.String(prop) def apply(s: String): Markup = Markup(name, Prop(s)) def unapply(markup: Markup): Option[String] = if (markup.name == name) Prop.unapply(markup.properties) else None def get(markup: Markup): String = unapply(markup).getOrElse("") } class Markup_Int(val name: String, prop: String) { val Prop: Properties.Int = new Properties.Int(prop) def apply(i: Int): Markup = Markup(name, Prop(i)) def unapply(markup: Markup): Option[Int] = if (markup.name == name) Prop.unapply(markup.properties) else None def get(markup: Markup): Int = unapply(markup).getOrElse(0) } class Markup_Long(val name: String, prop: String) { val Prop: Properties.Long = new Properties.Long(prop) def apply(i: Long): Markup = Markup(name, Prop(i)) def unapply(markup: Markup): Option[Long] = if (markup.name == name) Prop.unapply(markup.properties) else None def get(markup: Markup): Long = unapply(markup).getOrElse(0) } /* meta data */ val META_TITLE = "meta_title" val META_CREATOR = "meta_creator" val META_CONTRIBUTOR = "meta_contributor" val META_DATE = "meta_date" val META_LICENSE = "meta_license" val META_DESCRIPTION = "meta_description" /* formal entities */ val BINDING = "binding" val ENTITY = "entity" object Entity { val Def = new Markup_Long(ENTITY, "def") val Ref = new Markup_Long(ENTITY, "ref") object Occ { def unapply(markup: Markup): Option[Long] = Def.unapply(markup) orElse Ref.unapply(markup) } def unapply(markup: Markup): Option[(String, String)] = markup match { case Markup(ENTITY, props) => Some((Kind.get(props), Name.get(props))) case _ => None } } /* completion */ val COMPLETION = "completion" val NO_COMPLETION = "no_completion" val UPDATE = "update" /* position */ val LINE = "line" val END_LINE = "line" val OFFSET = "offset" val END_OFFSET = "end_offset" + val LABEL = "label" val FILE = "file" val ID = "id" val DEF_LINE = "def_line" val DEF_OFFSET = "def_offset" val DEF_END_OFFSET = "def_end_offset" + val DEF_LABEL = "def_label" val DEF_FILE = "def_file" val DEF_ID = "def_id" val POSITION = "position" - val POSITION_PROPERTIES = Set(LINE, OFFSET, END_OFFSET, FILE, ID) + val POSITION_PROPERTIES = Set(LINE, OFFSET, END_OFFSET, LABEL, FILE, ID) def position_property(entry: Properties.Entry): Boolean = POSITION_PROPERTIES(entry._1) /* position "def" name */ private val def_names: Map[String, String] = Map(LINE -> DEF_LINE, OFFSET -> DEF_OFFSET, END_OFFSET -> DEF_END_OFFSET, - FILE -> DEF_FILE, ID -> DEF_ID) + LABEL -> DEF_LABEL, FILE -> DEF_FILE, ID -> DEF_ID) def def_name(a: String): String = def_names.getOrElse(a, a + "def_") /* expression */ val EXPRESSION = "expression" object Expression { def unapply(markup: Markup): Option[String] = markup match { case Markup(EXPRESSION, props) => Some(Kind.get(props)) case _ => None } } /* citation */ val CITATION = "citation" val Citation = new Markup_String(CITATION, NAME) /* embedded languages */ val Symbols = new Properties.Boolean("symbols") val Antiquotes = new Properties.Boolean("antiquotes") val Delimited = new Properties.Boolean("delimited") val LANGUAGE = "language" object Language { val DOCUMENT = "document" val ANTIQUOTATION = "antiquotation" val ML = "ML" val SML = "SML" val PATH = "path" val UNKNOWN = "unknown" def apply(name: String, symbols: Boolean, antiquotes: Boolean, delimited: Boolean): Language = new Language(name, symbols, antiquotes, delimited) val outer: Language = apply("", true, false, false) def unapply(markup: Markup): Option[Language] = markup match { case Markup(LANGUAGE, props) => (props, props, props, props) match { case (Name(name), Symbols(symbols), Antiquotes(antiquotes), Delimited(delimited)) => Some(apply(name, symbols, antiquotes, delimited)) case _ => None } case _ => None } } class Language private( val name: String, val symbols: Boolean, val antiquotes: Boolean, val delimited: Boolean ) { override def toString: String = name def is_document: Boolean = name == Language.DOCUMENT def is_antiquotation: Boolean = name == Language.ANTIQUOTATION def is_path: Boolean = name == Language.PATH def description: String = Word.implode(Word.explode('_', name)) } /* external resources */ val PATH = "path" val Path = new Markup_String(PATH, NAME) val EXPORT_PATH = "export_path" val Export_Path = new Markup_String(EXPORT_PATH, NAME) val URL = "url" val Url = new Markup_String(URL, NAME) val DOC = "doc" val Doc = new Markup_String(DOC, NAME) /* pretty printing */ val Consistent = new Properties.Boolean("consistent") val Indent = new Properties.Int("indent") val Width = new Properties.Int("width") object Block { val name = "block" def apply(c: Boolean, i: Int): Markup = Markup(name, (if (c) Consistent(c) else Nil) ::: (if (i != 0) Indent(i) else Nil)) def unapply(markup: Markup): Option[(Boolean, Int)] = if (markup.name == name) { val c = Consistent.get(markup.properties) val i = Indent.get(markup.properties) Some((c, i)) } else None } object Break { val name = "break" def apply(w: Int, i: Int): Markup = Markup(name, (if (w != 0) Width(w) else Nil) ::: (if (i != 0) Indent(i) else Nil)) def unapply(markup: Markup): Option[(Int, Int)] = if (markup.name == name) { val w = Width.get(markup.properties) val i = Indent.get(markup.properties) Some((w, i)) } else None } val ITEM = "item" val BULLET = "bullet" val SEPARATOR = "separator" /* text properties */ val WORDS = "words" val HIDDEN = "hidden" val DELETE = "delete" /* misc entities */ val SESSION = "session" val THEORY = "theory" val CLASS = "class" val LOCALE = "locale" val BUNDLE = "bundle" val TYPE_NAME = "type_name" val CONSTANT = "constant" val AXIOM = "axiom" val FACT = "fact" val ORACLE = "oracle" val FIXED = "fixed" val CASE = "case" val DYNAMIC_FACT = "dynamic_fact" val LITERAL_FACT = "literal_fact" val ATTRIBUTE = "attribute" val METHOD = "method" val METHOD_MODIFIER = "method_modifier" /* inner syntax */ val TFREE = "tfree" val TVAR = "tvar" val FREE = "free" val SKOLEM = "skolem" val BOUND = "bound" val VAR = "var" val NUMERAL = "numeral" val LITERAL = "literal" val DELIMITER = "delimiter" val INNER_STRING = "inner_string" val INNER_CARTOUCHE = "inner_cartouche" val TOKEN_RANGE = "token_range" val SORTING = "sorting" val TYPING = "typing" val CLASS_PARAMETER = "class_parameter" /* antiquotations */ val ANTIQUOTED = "antiquoted" val ANTIQUOTE = "antiquote" val ML_ANTIQUOTATION = "ML_antiquotation" val DOCUMENT_ANTIQUOTATION = "document_antiquotation" val DOCUMENT_ANTIQUOTATION_OPTION = "document_antiquotation_option" /* document text */ val RAW_TEXT = "raw_text" val PLAIN_TEXT = "plain_text" val PARAGRAPH = "paragraph" val TEXT_FOLD = "text_fold" object Document_Tag extends Markup_String("document_tag", NAME) { val IMPORTANT = "important" val UNIMPORTANT = "unimportant" } /* LaTeX */ val Document_Latex = new Markup_Elem("document_latex") val Latex_Output = new Markup_Elem("latex_output") val Latex_Macro0 = new Markup_String("latex_macro0", NAME) val Latex_Macro = new Markup_String("latex_macro", NAME) val Latex_Environment = new Markup_String("latex_environment", NAME) val Latex_Heading = new Markup_String("latex_heading", KIND) val Latex_Body = new Markup_String("latex_body", KIND) val Latex_Delim = new Markup_String("latex_delim", NAME) val Latex_Tag = new Markup_String("latex_tag", NAME) val Latex_Cite = new Markup_Elem("latex_cite") val Citations = new Properties.String("citations") val Latex_Index_Item = new Markup_Elem("latex_index_item") val Latex_Index_Entry = new Markup_String("latex_index_entry", KIND) val Optional_Argument = new Properties.String("optional_argument") /* Markdown document structure */ val MARKDOWN_PARAGRAPH = "markdown_paragraph" val MARKDOWN_ITEM = "markdown_item" val Markdown_Bullet = new Markup_Int("markdown_bullet", "depth") val Markdown_List = new Markup_String("markdown_list", "kind") val ITEMIZE = "itemize" val ENUMERATE = "enumerate" val DESCRIPTION = "description" /* ML */ val ML_KEYWORD1 = "ML_keyword1" val ML_KEYWORD2 = "ML_keyword2" val ML_KEYWORD3 = "ML_keyword3" val ML_DELIMITER = "ML_delimiter" val ML_TVAR = "ML_tvar" val ML_NUMERAL = "ML_numeral" val ML_CHAR = "ML_char" val ML_STRING = "ML_string" val ML_COMMENT = "ML_comment" val ML_DEF = "ML_def" val ML_OPEN = "ML_open" val ML_STRUCTURE = "ML_structure" val ML_TYPING = "ML_typing" val ML_BREAKPOINT = "ML_breakpoint" /* outer syntax */ val COMMAND_SPAN = "command_span" val Command_Span = new Markup_String(COMMAND_SPAN, NAME) val COMMAND = "command" val KEYWORD = "keyword" val KEYWORD1 = "keyword1" val KEYWORD2 = "keyword2" val KEYWORD3 = "keyword3" val QUASI_KEYWORD = "quasi_keyword" val IMPROPER = "improper" val OPERATOR = "operator" val STRING = "string" val ALT_STRING = "alt_string" val CARTOUCHE = "cartouche" val COMMENT = "comment" val LOAD_COMMAND = "load_command" /* comments */ val COMMENT1 = "comment1" val COMMENT2 = "comment2" val COMMENT3 = "comment3" /* timing */ val Elapsed = new Properties.Double("elapsed") val CPU = new Properties.Double("cpu") val GC = new Properties.Double("gc") object Timing_Properties { def apply(timing: isabelle.Timing): Properties.T = Elapsed(timing.elapsed.seconds) ::: CPU(timing.cpu.seconds) ::: GC(timing.gc.seconds) def unapply(props: Properties.T): Option[isabelle.Timing] = (props, props, props) match { case (Elapsed(elapsed), CPU(cpu), GC(gc)) => Some(new isabelle.Timing(Time.seconds(elapsed), Time.seconds(cpu), Time.seconds(gc))) case _ => None } def get(props: Properties.T): isabelle.Timing = unapply(props).getOrElse(isabelle.Timing.zero) } val TIMING = "timing" object Timing { def apply(timing: isabelle.Timing): Markup = Markup(TIMING, Timing_Properties(timing)) def unapply(markup: Markup): Option[isabelle.Timing] = markup match { case Markup(TIMING, Timing_Properties(timing)) => Some(timing) case _ => None } } /* process result */ val Return_Code = new Properties.Int("return_code") object Process_Result { def apply(result: Process_Result): Properties.T = Return_Code(result.rc) ::: (if (result.timing.is_zero) Nil else Timing_Properties(result.timing)) def unapply(props: Properties.T): Option[Process_Result] = props match { case Return_Code(rc) => val timing = Timing_Properties.unapply(props).getOrElse(isabelle.Timing.zero) Some(isabelle.Process_Result(rc, timing = timing)) case _ => None } } /* command indentation */ val Command_Indent = new Markup_Int("command_indent", Indent.name) /* goals */ val GOAL = "goal" val SUBGOAL = "subgoal" /* command status */ val TASK = "task" val ACCEPTED = "accepted" val FORKED = "forked" val JOINED = "joined" val RUNNING = "running" val FINISHED = "finished" val FAILED = "failed" val CANCELED = "canceled" val INITIALIZED = "initialized" val FINALIZED = "finalized" val CONSOLIDATING = "consolidating" val CONSOLIDATED = "consolidated" /* interactive documents */ val VERSION = "version" val ASSIGN = "assign" /* prover process */ val PROVER_COMMAND = "prover_command" val PROVER_ARG = "prover_arg" /* messages */ val INIT = "init" val STATUS = "status" val REPORT = "report" val RESULT = "result" val WRITELN = "writeln" val STATE = "state" val INFORMATION = "information" val TRACING = "tracing" val WARNING = "warning" val LEGACY = "legacy" val ERROR = "error" val NODES_STATUS = "nodes_status" val PROTOCOL = "protocol" val SYSTEM = "system" val STDOUT = "stdout" val STDERR = "stderr" val EXIT = "exit" val WRITELN_MESSAGE = "writeln_message" val STATE_MESSAGE = "state_message" val INFORMATION_MESSAGE = "information_message" val TRACING_MESSAGE = "tracing_message" val WARNING_MESSAGE = "warning_message" val LEGACY_MESSAGE = "legacy_message" val ERROR_MESSAGE = "error_message" val messages = Map( WRITELN -> WRITELN_MESSAGE, STATE -> STATE_MESSAGE, INFORMATION -> INFORMATION_MESSAGE, TRACING -> TRACING_MESSAGE, WARNING -> WARNING_MESSAGE, LEGACY -> LEGACY_MESSAGE, ERROR -> ERROR_MESSAGE) val message: String => String = messages.withDefault((s: String) => s) val NO_REPORT = "no_report" val BAD = "bad" val INTENSIFY = "intensify" /* ML profiling */ val COUNT = "count" val ML_PROFILING_ENTRY = "ML_profiling_entry" val ML_PROFILING = "ML_profiling" object ML_Profiling { def unapply_entry(tree: XML.Tree): Option[isabelle.ML_Profiling.Entry] = tree match { case XML.Elem(Markup(ML_PROFILING_ENTRY, List((NAME, name), (COUNT, Value.Long(count)))), _) => Some(isabelle.ML_Profiling.Entry(name, count)) case _ => None } def unapply_report(tree: XML.Tree): Option[isabelle.ML_Profiling.Report] = tree match { case XML.Elem(Markup(ML_PROFILING, List((KIND, kind))), body) => Some(isabelle.ML_Profiling.Report(kind, body.flatMap(unapply_entry))) case _ => None } } /* active areas */ val BROWSER = "browser" val GRAPHVIEW = "graphview" val THEORY_EXPORTS = "theory_exports" val SENDBACK = "sendback" val PADDING = "padding" val PADDING_LINE = (PADDING, "line") val PADDING_COMMAND = (PADDING, "command") val DIALOG = "dialog" val Result = new Properties.String(RESULT) val JEDIT_ACTION = "jedit_action" /* protocol message functions */ val FUNCTION = "function" class Function(val name: String) { val THIS: Properties.Entry = (FUNCTION, name) } class Properties_Function(name: String) extends Function(name) { def unapply(props: Properties.T): Option[Properties.T] = props match { case THIS :: args => Some(args) case _ => None } } class Name_Function(name: String) extends Function(name) { def unapply(props: Properties.T): Option[String] = props match { case List(THIS, (NAME, a)) => Some(a) case _ => None } } object ML_Statistics extends Function("ML_statistics") { def unapply(props: Properties.T): Option[(Long, String)] = props match { case List(THIS, ("pid", Value.Long(pid)), ("stats_dir", stats_dir)) => Some((pid, stats_dir)) case _ => None } } val command_timing_properties: Set[String] = Set(FILE, OFFSET, NAME, Elapsed.name) def command_timing_property(entry: Properties.Entry): Boolean = command_timing_properties(entry._1) object Command_Timing extends Properties_Function("command_timing") object Theory_Timing extends Properties_Function("theory_timing") object Session_Timing extends Properties_Function("session_timing") { val Threads = new Properties.Int("threads") } object Task_Statistics extends Properties_Function("task_statistics") object Loading_Theory extends Properties_Function("loading_theory") object Build_Session_Finished extends Function("build_session_finished") object Commands_Accepted extends Function("commands_accepted") object Assign_Update extends Function("assign_update") object Removed_Versions extends Function("removed_versions") object Invoke_Scala extends Function("invoke_scala") { def unapply(props: Properties.T): Option[(String, String)] = props match { case List(THIS, (NAME, name), (ID, id)) => Some((name, id)) case _ => None } } object Cancel_Scala extends Function("cancel_scala") { def unapply(props: Properties.T): Option[String] = props match { case List(THIS, (ID, id)) => Some(id) case _ => None } } val PRINT_OPERATIONS = "print_operations" /* export */ val EXPORT = "export" val THEORY_NAME = "theory_name" val EXECUTABLE = "executable" val COMPRESS = "compress" val STRICT = "strict" /* debugger output */ val DEBUGGER_STATE = "debugger_state" object Debugger_State { def unapply(props: Properties.T): Option[String] = props match { case List((FUNCTION, DEBUGGER_STATE), (NAME, name)) => Some(name) case _ => None } } val DEBUGGER_OUTPUT = "debugger_output" object Debugger_Output { def unapply(props: Properties.T): Option[String] = props match { case List((FUNCTION, DEBUGGER_OUTPUT), (NAME, name)) => Some(name) case _ => None } } /* simplifier trace */ val SIMP_TRACE_PANEL = "simp_trace_panel" val SIMP_TRACE_LOG = "simp_trace_log" val SIMP_TRACE_STEP = "simp_trace_step" val SIMP_TRACE_RECURSE = "simp_trace_recurse" val SIMP_TRACE_HINT = "simp_trace_hint" val SIMP_TRACE_IGNORE = "simp_trace_ignore" val SIMP_TRACE_CANCEL = "simp_trace_cancel" object Simp_Trace_Cancel { def unapply(props: Properties.T): Option[Long] = props match { case (FUNCTION, SIMP_TRACE_CANCEL) :: Serial(i) => Some(i) case _ => None } } /* XML data representation */ def encode: XML.Encode.T[Markup] = { (markup: Markup) => import XML.Encode._ pair(string, properties)((markup.name, markup.properties)) } def decode: XML.Decode.T[Markup] = { (body: XML.Body) => import XML.Decode._ val (name, props) = pair(string, properties)(body) Markup(name, props) } } sealed case class Markup(name: String, properties: Properties.T) { def is_empty: Boolean = name.isEmpty def position_properties: Position.T = properties.filter(Markup.position_property) def markup(s: String): String = YXML.string_of_tree(XML.Elem(this, List(XML.Text(s)))) def update_properties(more_props: Properties.T): Markup = if (more_props.isEmpty) this else Markup(name, more_props.foldRight(properties) { case (p, ps) => Properties.put(ps, p) }) def + (entry: Properties.Entry): Markup = Markup(name, Properties.put(properties, entry)) } diff --git a/src/Pure/PIDE/session.ML b/src/Pure/PIDE/session.ML --- a/src/Pure/PIDE/session.ML +++ b/src/Pure/PIDE/session.ML @@ -1,45 +1,51 @@ (* Title: Pure/PIDE/session.ML Author: Makarius Prover session: persistent state of logic image. *) signature SESSION = sig val init: string -> unit val get_name: unit -> string val welcome: unit -> string val shutdown: unit -> unit val finish: unit -> unit + val print_context_tracing: (Context.generic -> bool) -> unit end; structure Session: SESSION = struct (* session name *) val session = Synchronized.var "Session.session" ""; fun init name = Synchronized.change session (K name); fun get_name () = Synchronized.value session; fun description () = (case get_name () of "" => "Isabelle" | name => "Isabelle/" ^ name); fun welcome () = "Welcome to " ^ description () ^ Isabelle_System.isabelle_heading (); (* finish *) fun shutdown () = (Execution.shutdown (); Event_Timer.shutdown (); Future.shutdown ()); fun finish () = (shutdown (); Par_List.map (Global_Theory.get_thm_names o Thy_Info.get_theory) (Thy_Info.get_names ()); Thy_Info.finish (); - shutdown ()); + shutdown (); + ignore (Context.finish_tracing ())); + +fun print_context_tracing pred = + List.app (writeln o Proof_Display.print_context_tracing) + (filter (pred o #1) (#contexts (Context.finish_tracing ()))); end; diff --git a/src/Pure/System/isabelle_process.ML b/src/Pure/System/isabelle_process.ML --- a/src/Pure/System/isabelle_process.ML +++ b/src/Pure/System/isabelle_process.ML @@ -1,226 +1,229 @@ (* Title: Pure/System/isabelle_process.ML Author: Makarius Isabelle process wrapper. *) signature ISABELLE_PROCESS = sig val is_active: unit -> bool val reset_tracing: Document_ID.exec -> unit val crashes: exn list Synchronized.var val init_options: unit -> unit val init_options_interactive: unit -> unit val init: unit -> unit val init_build: unit -> unit end; structure Isabelle_Process: ISABELLE_PROCESS = struct (* print mode *) val isabelle_processN = "isabelle_process"; fun is_active () = Print_Mode.print_mode_active isabelle_processN; val _ = Output.add_mode isabelle_processN Output.default_output Output.default_escape; val _ = Markup.add_mode isabelle_processN YXML.output_markup; val protocol_modes1 = [Syntax_Trans.no_bracketsN, Syntax_Trans.no_type_bracketsN]; val protocol_modes2 = [isabelle_processN, Pretty.symbolicN]; (* restricted tracing messages *) val tracing_messages = Synchronized.var "tracing_messages" (Inttab.empty: int Inttab.table); fun reset_tracing exec_id = Synchronized.change tracing_messages (Inttab.delete_safe exec_id); fun update_tracing () = (case Position.parse_id (Position.thread_data ()) of NONE => () | SOME exec_id => let val ok = Synchronized.change_result tracing_messages (fn tab => let val n = the_default 0 (Inttab.lookup tab exec_id) + 1; val limit = Options.default_int "editor_tracing_messages"; val ok = limit <= 0 orelse n <= limit; in (ok, Inttab.update (exec_id, n) tab) end); in if ok then () else let val (text, promise) = Active.dialog_text (); val _ = writeln ("Tracing paused. " ^ text "Stop" ^ ", or continue with next " ^ text "100" ^ ", " ^ text "1000" ^ ", " ^ text "10000" ^ " messages?") val m = Value.parse_int (Future.join promise) handle Fail _ => error "Stopped"; in Synchronized.change tracing_messages (Inttab.map_default (exec_id, 0) (fn k => k - m)) end end); (* init protocol -- uninterruptible *) val crashes = Synchronized.var "Isabelle_Process.crashes" ([]: exn list); local fun recover crash = (Synchronized.change crashes (cons crash); Output.physical_stderr "Recovered from Isabelle process crash -- see also Isabelle_Process.crashes\n"); fun ml_statistics () = Output.protocol_message (Markup.ML_statistics {pid = ML_Pid.get (), stats_dir = getenv "POLYSTATSDIR"}) []; in fun init_protocol modes = Thread_Attributes.uninterruptible (fn _ => fn (address, password) => let val _ = SHA1.test_samples () handle exn as Fail msg => (Output.physical_stderr (msg ^ "\n"); Exn.reraise exn); val _ = Output.physical_stderr Symbol.STX; (* streams *) val (in_stream, out_stream) = Socket_IO.open_streams address; val _ = Byte_Message.write_line out_stream (Bytes.string password); val _ = TextIO.StreamIO.setBufferMode (TextIO.getOutstream TextIO.stdOut, IO.LINE_BUF); val _ = TextIO.StreamIO.setBufferMode (TextIO.getOutstream TextIO.stdErr, IO.LINE_BUF); val _ = BinIO.StreamIO.setBufferMode (BinIO.getOutstream out_stream, IO.BLOCK_BUF); (* messages *) val message_channel = Message_Channel.make out_stream; val message = Message_Channel.message message_channel; fun standard_message props name ss = if forall (fn s => s = "") ss then () else let val pos_props = if exists Markup.position_property props then props else props @ Position.properties_of (Position.thread_data ()); in message name pos_props [XML.blob ss] end; fun report_message ss = if Context_Position.reports_enabled0 () then standard_message [] Markup.reportN ss else (); val serial_props = Markup.serial_properties o serial; val message_context = Unsynchronized.setmp Private_Output.status_fn (standard_message [] Markup.statusN) #> Unsynchronized.setmp Private_Output.report_fn report_message #> Unsynchronized.setmp Private_Output.result_fn (fn props => fn s => standard_message (props @ serial_props ()) Markup.resultN s) #> Unsynchronized.setmp Private_Output.writeln_fn (fn s => standard_message (serial_props ()) Markup.writelnN s) #> Unsynchronized.setmp Private_Output.state_fn (fn s => standard_message (serial_props ()) Markup.stateN s) #> Unsynchronized.setmp Private_Output.information_fn (fn s => standard_message (serial_props ()) Markup.informationN s) #> Unsynchronized.setmp Private_Output.tracing_fn (fn s => (update_tracing (); standard_message (serial_props ()) Markup.tracingN s)) #> Unsynchronized.setmp Private_Output.warning_fn (fn s => standard_message (serial_props ()) Markup.warningN s) #> Unsynchronized.setmp Private_Output.legacy_fn (fn s => standard_message (serial_props ()) Markup.legacyN s) #> Unsynchronized.setmp Private_Output.error_message_fn (fn (i, s) => standard_message (Markup.serial_properties i) Markup.errorN s) #> Unsynchronized.setmp Private_Output.system_message_fn (fn ss => message Markup.systemN [] [XML.blob ss]) #> Unsynchronized.setmp Private_Output.protocol_message_fn (fn props => fn chunks => message Markup.protocolN props chunks) #> Unsynchronized.setmp print_mode ((! print_mode @ #1 modes) |> fold (update op =) (#2 modes)); (* protocol *) fun protocol_loop () = let val _ = (case Byte_Message.read_message in_stream of NONE => raise Protocol_Command.STOP 0 | SOME [] => Output.system_message "Isabelle process: no input" | SOME (name :: args) => Protocol_Command.run (Bytes.content name) args) handle exn => if Protocol_Command.is_protocol_exn exn then Exn.reraise exn else (Runtime.exn_system_message exn handle crash => recover crash); in protocol_loop () end; fun protocol () = (message Markup.initN [] [[XML.Text (Session.welcome ())]]; ml_statistics (); protocol_loop ()); val result = Exn.capture (message_context protocol) (); (* shutdown *) val _ = Future.shutdown (); val _ = Execution.reset (); val _ = Message_Channel.shutdown message_channel; val _ = BinIO.closeIn in_stream; val _ = BinIO.closeOut out_stream; val _ = Options.reset_default (); in (case result of Exn.Exn (Protocol_Command.STOP rc) => if rc = 0 then () else exit rc | _ => Exn.release result) end); end; (* init options *) fun init_options () = (ML_Print_Depth.set_print_depth (Options.default_int "ML_print_depth"); Multithreading.trace := Options.default_int "threads_trace"; Multithreading.max_threads_update (Options.default_int "threads"); Multithreading.parallel_proofs := Options.default_int "parallel_proofs"; if Options.default_bool "export_standard_proofs" then Proofterm.proofs := 2 else (); let val proofs = Options.default_int "record_proofs" in if proofs < 0 then () else Proofterm.proofs := proofs end; - Printer.show_markup_default := false); + Printer.show_markup_default := false; + Context.theory_tracing := Options.default_bool "context_theory_tracing"; + Context.proof_tracing := Options.default_bool "context_proof_tracing"; + Context.data_timing := Options.default_bool "context_data_timing"); fun init_options_interactive () = (init_options (); Multithreading.parallel_proofs := (if Options.default_int "parallel_proofs" > 0 then 3 else 0); Printer.show_markup_default := true); (* generic init *) fun init_modes modes = let val address = Options.default_string \<^system_option>\system_channel_address\; val password = Options.default_string \<^system_option>\system_channel_password\; in if address <> "" andalso password <> "" then init_protocol modes (address, password) else init_options () end; fun init () = init_modes (protocol_modes1, protocol_modes2); fun init_build () = init_modes ([], protocol_modes2); end; diff --git a/src/Pure/Thy/thy_info.ML b/src/Pure/Thy/thy_info.ML --- a/src/Pure/Thy/thy_info.ML +++ b/src/Pure/Thy/thy_info.ML @@ -1,466 +1,466 @@ (* Title: Pure/Thy/thy_info.ML Author: Makarius Global theory info database, with auto-loading according to theory and file dependencies, and presentation of theory documents. *) signature THY_INFO = sig type presentation_context = {options: Options.T, file_pos: Position.T, adjust_pos: Position.T -> Position.T, segments: Document_Output.segment list} val adjust_pos_properties: presentation_context -> Position.T -> Properties.T val apply_presentation: presentation_context -> theory -> unit val add_presentation: (presentation_context -> theory -> unit) -> theory -> theory val get_names: unit -> string list val lookup_theory: string -> theory option val defined_theory: string -> bool val get_theory: string -> theory val master_directory: string -> Path.T val remove_thy: string -> unit val use_theories: Options.T -> string -> (string * Position.T) list -> (theory * Document_Output.segment list) list val use_thy: string -> unit val script_thy: Position.T -> string -> theory -> theory val register_thy: theory -> unit val finish: unit -> unit end; structure Thy_Info: THY_INFO = struct (** theory presentation **) (* hook for consolidated theory *) type presentation_context = {options: Options.T, file_pos: Position.T, adjust_pos: Position.T -> Position.T, segments: Document_Output.segment list}; fun adjust_pos_properties (context: presentation_context) pos = let val props = Position.properties_of pos; val props' = Position.properties_of (#adjust_pos context pos); in filter (fn (a, _) => a = Markup.offsetN orelse a = Markup.end_offsetN) props' @ - filter (fn (a, _) => a = Markup.idN orelse a = Markup.fileN) props + filter (fn (a, _) => a = Markup.labelN orelse a = Markup.fileN orelse a = Markup.idN) props end; structure Presentation = Theory_Data ( type T = ((presentation_context -> theory -> unit) * stamp) list; val empty = []; fun merge data : T = Library.merge (eq_snd op =) data; ); fun apply_presentation (context: presentation_context) thy = ignore (Par_List.map (fn (f, _) => f context thy) (Presentation.get thy)); fun add_presentation f = Presentation.map (cons (f, stamp ())); val _ = Theory.setup (add_presentation (fn {options, segments, ...} => fn thy => if exists (Toplevel.is_skipped_proof o #state) segments then () else let val keywords = Thy_Header.get_keywords thy; val thy_name = Context.theory_base_name thy; val latex = Document_Output.present_thy options keywords thy_name segments; in if Options.string options "document" = "false" then () else Export.export thy \<^path_binding>\document/latex\ latex end)); (** thy database **) (* messages *) val show_path = space_implode " via " o map quote; fun cycle_msg names = "Cyclic dependency of " ^ show_path names; (* derived graph operations *) fun add_deps name parents G = String_Graph.add_deps_acyclic (name, parents) G handle String_Graph.CYCLES namess => error (cat_lines (map cycle_msg namess)); fun new_entry name parents entry = String_Graph.new_node (name, entry) #> add_deps name parents; (* global thys *) type deps = {master: (Path.T * SHA1.digest), (*master dependencies for thy file*) imports: (string * Position.T) list}; (*source specification of imports (partially qualified)*) fun make_deps master imports : deps = {master = master, imports = imports}; fun master_dir_deps (d: deps option) = the_default Path.current (Option.map (Path.dir o #1 o #master) d); local val global_thys = Synchronized.var "Thy_Info.thys" (String_Graph.empty: (deps option * theory option) String_Graph.T); in fun get_thys () = Synchronized.value global_thys; fun change_thys f = Synchronized.change global_thys f; end; fun get_names () = String_Graph.topological_order (get_thys ()); (* access thy *) fun lookup thys name = try (String_Graph.get_node thys) name; fun lookup_thy name = lookup (get_thys ()) name; fun get thys name = (case lookup thys name of SOME thy => thy | NONE => error ("Theory loader: nothing known about theory " ^ quote name)); fun get_thy name = get (get_thys ()) name; (* access deps *) val lookup_deps = Option.map #1 o lookup_thy; val master_directory = master_dir_deps o #1 o get_thy; (* access theory *) fun lookup_theory name = (case lookup_thy name of SOME (_, SOME theory) => SOME theory | _ => NONE); val defined_theory = is_some o lookup_theory; fun get_theory name = (case lookup_theory name of SOME theory => theory | _ => error ("Theory loader: undefined entry for theory " ^ quote name)); val get_imports = Resources.imports_of o get_theory; (** thy operations **) (* remove *) fun remove name thys = (case lookup thys name of NONE => thys | SOME (NONE, _) => error ("Cannot update finished theory " ^ quote name) | SOME _ => let val succs = String_Graph.all_succs thys [name]; val _ = writeln ("Theory loader: removing " ^ commas_quote succs); in fold String_Graph.del_node succs thys end); val remove_thy = change_thys o remove; (* update *) fun update deps theory thys = let val name = Context.theory_long_name theory; val parents = map Context.theory_long_name (Theory.parents_of theory); val thys' = remove name thys; val _ = map (get thys') parents; in new_entry name parents (SOME deps, SOME theory) thys' end; fun update_thy deps theory = change_thys (update deps theory); (* scheduling loader tasks *) datatype result = Result of {theory: theory, exec_id: Document_ID.exec, present: unit -> presentation_context option, commit: unit -> unit}; fun theory_result theory = Result {theory = theory, exec_id = Document_ID.none, present = K NONE, commit = I}; fun result_theory (Result {theory, ...}) = theory; fun result_commit (Result {commit, ...}) = commit; datatype task = Task of string list * (theory list -> result) | Finished of theory; local fun consolidate_theory (Exn.Exn exn) = [Exn.Exn exn] | consolidate_theory (Exn.Res (Result {theory, exec_id, ...})) = let val _ = Execution.join [exec_id]; val res = Exn.capture Thm.consolidate_theory theory; val exns = maps Task_Queue.group_status (Execution.peek exec_id); in res :: map Exn.Exn exns end; fun present_theory (Exn.Exn exn) = [Exn.Exn exn] | present_theory (Exn.Res (Result {theory, present, ...})) = (case present () of NONE => [] | SOME (context as {segments, ...}) => [Exn.capture (fn () => (apply_presentation context theory; (theory, segments))) ()]); in val schedule_theories = Thread_Attributes.uninterruptible (fn _ => fn tasks => let fun join_parents deps name parents = (case map #1 (filter (not o can Future.join o #2) deps) of [] => map (result_theory o Future.join o the o AList.lookup (op =) deps) parents | bad => error ("Failed to load theory " ^ quote name ^ " (unresolved " ^ commas_quote bad ^ ")")); val futures = tasks |> String_Graph.schedule (fn deps => fn (name, task) => (case task of Task (parents, body) => if Multithreading.max_threads () > 1 then (singleton o Future.forks) {name = "theory:" ^ name, group = NONE, deps = map (Future.task_of o #2) deps, pri = 0, interrupts = true} (fn () => body (join_parents deps name parents)) else Future.value_result (Exn.capture (fn () => body (join_parents deps name parents)) ()) | Finished theory => Future.value (theory_result theory))); val results1 = futures |> maps (consolidate_theory o Future.join_result); val present_results = futures |> maps (present_theory o Future.join_result); val results2 = (map o Exn.map_res) (K ()) present_results; val results3 = futures |> map (fn future => Exn.capture (fn () => result_commit (Future.join future) ()) ()); val results4 = map Exn.Exn (maps Task_Queue.group_status (Execution.reset ())); val _ = Par_Exn.release_all (results1 @ results2 @ results3 @ results4); in Par_Exn.release_all present_results end); end; (* eval theory *) fun eval_thy options master_dir header text_pos text parents = let val (name, _) = #name header; val keywords = fold (curry Keyword.merge_keywords o Thy_Header.get_keywords) parents (Keyword.add_keywords (#keywords header) Keyword.empty_keywords); val spans = Outer_Syntax.parse_spans (Token.explode keywords text_pos text); val elements = Thy_Element.parse_elements keywords spans; val text_id = Position.copy_id text_pos Position.none; fun init () = Resources.begin_theory master_dir header parents; fun excursion () = let fun element_result span_elem (st, _) = let fun prepare span = let val tr = Position.setmp_thread_data text_id (fn () => Command.read_span keywords st master_dir init span) (); in Toplevel.timing (Resources.last_timing tr) tr end; val elem = Thy_Element.map_element prepare span_elem; val (results, st') = Toplevel.element_result keywords elem st; val pos' = Toplevel.pos_of (Thy_Element.last_element elem); in (results, (st', pos')) end; val (results, (end_state, end_pos)) = fold_map element_result elements (Toplevel.make_state NONE, Position.none); val thy = Toplevel.end_theory end_pos end_state; in (results, thy) end; val (results, thy) = cond_timeit true ("theory " ^ quote name) excursion; fun present () : presentation_context = let val segments = (spans ~~ maps Toplevel.join_results results) |> map (fn (span, (st, tr, st')) => {span = span, prev_state = st, command = tr, state = st'}); in {options = options, file_pos = text_pos, adjust_pos = I, segments = segments} end; in (thy, present) end; (* require_thy -- checking database entries wrt. the file-system *) local fun required_by _ [] = "" | required_by s initiators = s ^ "(required by " ^ show_path (rev initiators) ^ ")"; fun load_thy options initiators deps text (name, header_pos) keywords parents = let val {master = (thy_path, _), imports} = deps; val dir = Path.dir thy_path; val exec_id = Document_ID.make (); val id = Document_ID.print exec_id; val put_id = Position.put_id id; val _ = Execution.running Document_ID.none exec_id [] orelse raise Fail ("Failed to register execution: " ^ id); val text_pos = put_id (Position.file (File.symbolic_path thy_path)); val text_props = Position.properties_of text_pos; val _ = remove_thy name; val _ = writeln ("Loading theory " ^ quote name ^ required_by " " initiators); val _ = Output.try_protocol_message (Markup.loading_theory name @ text_props) [XML.blob [text]]; val _ = Position.setmp_thread_data (Position.id_only id) (fn () => (parents, map #2 imports) |> ListPair.app (fn (thy, pos) => Context_Position.reports_global thy [(put_id pos, Theory.get_markup thy)])) (); val timing_start = Timing.start (); val header = Thy_Header.make (name, put_id header_pos) imports keywords; val (theory, present) = eval_thy options dir header text_pos text (if name = Context.PureN then [Context.the_global_context ()] else parents); val timing_result = Timing.result timing_start; val timing_props = [Markup.theory_timing, (Markup.nameN, name)]; val _ = Output.try_protocol_message (timing_props @ Markup.timing_properties timing_result) [] fun commit () = update_thy deps theory; in Result {theory = theory, exec_id = exec_id, present = SOME o present, commit = commit} end; fun check_thy_deps dir name = (case lookup_deps name of SOME NONE => (true, NONE, Position.none, get_imports name, []) | NONE => let val {master, text, theory_pos, imports, keywords} = Resources.check_thy dir name in (false, SOME (make_deps master imports, text), theory_pos, imports, keywords) end | SOME (SOME {master, ...}) => let val {master = master', text = text', theory_pos = theory_pos', imports = imports', keywords = keywords'} = Resources.check_thy dir name; val deps' = SOME (make_deps master' imports', text'); val current = #2 master = #2 master' andalso (case lookup_theory name of NONE => false | SOME theory => Resources.loaded_files_current theory); in (current, deps', theory_pos', imports', keywords') end); fun task_finished (Task _) = false | task_finished (Finished _) = true; in fun require_thys options initiators qualifier dir strs tasks = fold_map (require_thy options initiators qualifier dir) strs tasks |>> forall I and require_thy options initiators qualifier dir (s, require_pos) tasks = let val {master_dir, theory_name, ...} = Resources.import_name qualifier dir s; in (case try (String_Graph.get_node tasks) theory_name of SOME task => (task_finished task, tasks) | NONE => let val _ = member (op =) initiators theory_name andalso error (cycle_msg initiators); val (current, deps, theory_pos, imports, keywords) = check_thy_deps master_dir theory_name handle ERROR msg => cat_error msg ("The error(s) above occurred for theory " ^ quote theory_name ^ Position.here require_pos ^ required_by "\n" initiators); val qualifier' = Resources.theory_qualifier theory_name; val dir' = dir + master_dir_deps (Option.map #1 deps); val parents = map (#theory_name o Resources.import_name qualifier' dir' o #1) imports; val (parents_current, tasks') = require_thys options (theory_name :: initiators) qualifier' dir' imports tasks; val all_current = current andalso parents_current; val task = if all_current then Finished (get_theory theory_name) else (case deps of NONE => raise Fail "Malformed deps" | SOME (dep, text) => Task (parents, load_thy options initiators dep text (theory_name, theory_pos) keywords)); val tasks'' = new_entry theory_name parents task tasks'; in (all_current, tasks'') end) end; end; (* use theories *) fun use_theories options qualifier imports = schedule_theories (#2 (require_thys options [] qualifier Path.current imports String_Graph.empty)); fun use_thy name = ignore (use_theories (Options.default ()) Resources.default_qualifier [(name, Position.none)]); (* toplevel scripting -- without maintaining database *) fun script_thy pos txt thy = let val trs = Outer_Syntax.parse_text thy (K thy) pos txt; val end_pos = if null trs then pos else Toplevel.pos_of (List.last trs); val end_state = fold (Toplevel.command_exception true) trs (Toplevel.make_state NONE); in Toplevel.end_theory end_pos end_state end; (* register theory *) fun register_thy theory = let val name = Context.theory_long_name theory; val {master, ...} = Resources.check_thy (Resources.master_directory theory) name; val imports = Resources.imports_of theory; in change_thys (fn thys => let val thys' = remove name thys; val _ = writeln ("Registering theory " ^ quote name); in update (make_deps master imports) theory thys' end) end; (* finish all theories *) fun finish () = change_thys (String_Graph.map (fn _ => fn (_, entry) => (NONE, entry))); end; fun use_thy name = Runtime.toplevel_program (fn () => Thy_Info.use_thy name); diff --git a/src/Pure/context.ML b/src/Pure/context.ML --- a/src/Pure/context.ML +++ b/src/Pure/context.ML @@ -1,824 +1,827 @@ (* Title: Pure/context.ML Author: Markus Wenzel, TU Muenchen Generic theory contexts with unique identity, arbitrarily typed data, monotonic development graph and history support. Generic proof contexts with arbitrarily typed data. Firm naming conventions: thy, thy', thy1, thy2: theory ctxt, ctxt', ctxt1, ctxt2: Proof.context context: Context.generic *) signature BASIC_CONTEXT = sig type theory exception THEORY of string * theory list structure Proof: sig type context end structure Proof_Context: sig val theory_of: Proof.context -> theory val init_global: theory -> Proof.context val get_global: {long: bool} -> theory -> string -> Proof.context end end; signature CONTEXT = sig include BASIC_CONTEXT (*theory data*) type data_kind = int val data_kinds: unit -> (data_kind * Position.T) list (*theory context*) type id = int type theory_id val theory_id: theory -> theory_id - val timing: bool Unsynchronized.ref + val data_timing: bool Unsynchronized.ref val parents_of: theory -> theory list val ancestors_of: theory -> theory list val theory_id_ord: theory_id ord val theory_id_name: {long: bool} -> theory_id -> string val theory_long_name: theory -> string val theory_base_name: theory -> string val theory_name: {long: bool} -> theory -> string val theory_identifier: theory -> id val PureN: string val pretty_thy: theory -> Pretty.T val pretty_abbrev_thy: theory -> Pretty.T val get_theory: {long: bool} -> theory -> string -> theory val eq_thy_id: theory_id * theory_id -> bool val eq_thy: theory * theory -> bool val proper_subthy_id: theory_id * theory_id -> bool val proper_subthy: theory * theory -> bool val subthy_id: theory_id * theory_id -> bool val subthy: theory * theory -> bool val join_thys: theory list -> theory val begin_thy: string -> theory list -> theory val finish_thy: theory -> theory val theory_data_sizeof1: theory -> (Position.T * int) list (*proof context*) val raw_transfer: theory -> Proof.context -> Proof.context (*certificate*) datatype certificate = Certificate of theory | Certificate_Id of theory_id val certificate_theory: certificate -> theory val certificate_theory_id: certificate -> theory_id val eq_certificate: certificate * certificate -> bool val join_certificate: certificate * certificate -> certificate (*generic context*) datatype generic = Theory of theory | Proof of Proof.context - val trace_theories: bool Unsynchronized.ref - val trace_proofs: bool Unsynchronized.ref - val allocations_trace: unit -> + val theory_tracing: bool Unsynchronized.ref + val proof_tracing: bool Unsynchronized.ref + val enabled_tracing: unit -> bool + val finish_tracing: unit -> {contexts: (generic * Position.T) list, active_contexts: int, active_theories: int, active_proofs: int, total_contexts: int, total_theories: int, total_proofs: int} val cases: (theory -> 'a) -> (Proof.context -> 'a) -> generic -> 'a val mapping: (theory -> theory) -> (Proof.context -> Proof.context) -> generic -> generic val mapping_result: (theory -> 'a * theory) -> (Proof.context -> 'a * Proof.context) -> generic -> 'a * generic val the_theory: generic -> theory val the_proof: generic -> Proof.context val map_theory: (theory -> theory) -> generic -> generic val map_proof: (Proof.context -> Proof.context) -> generic -> generic val map_theory_result: (theory -> 'a * theory) -> generic -> 'a * generic val map_proof_result: (Proof.context -> 'a * Proof.context) -> generic -> 'a * generic val theory_map: (generic -> generic) -> theory -> theory val proof_map: (generic -> generic) -> Proof.context -> Proof.context val theory_of: generic -> theory (*total*) val proof_of: generic -> Proof.context (*total*) (*thread data*) val get_generic_context: unit -> generic option val put_generic_context: generic option -> unit val setmp_generic_context: generic option -> ('a -> 'b) -> 'a -> 'b val the_generic_context: unit -> generic val the_global_context: unit -> theory val the_local_context: unit -> Proof.context val >> : (generic -> generic) -> unit val >>> : (generic -> 'a * generic) -> 'a end; signature PRIVATE_CONTEXT = sig include CONTEXT structure Theory_Data: sig val declare: Position.T -> Any.T -> ((theory * Any.T) list -> Any.T) -> data_kind val get: data_kind -> (Any.T -> 'a) -> theory -> 'a val put: data_kind -> ('a -> Any.T) -> 'a -> theory -> theory end structure Proof_Data: sig val declare: (theory -> Any.T) -> data_kind val get: data_kind -> (Any.T -> 'a) -> Proof.context -> 'a val put: data_kind -> ('a -> Any.T) -> 'a -> Proof.context -> Proof.context end end; structure Context: PRIVATE_CONTEXT = struct (*** type definitions ***) (* context data *) (*private copy avoids potential conflict of table exceptions*) structure Datatab = Table(type key = int val ord = int_ord); type data_kind = int; val data_kind = Counter.make (); (* theory identity *) type id = int; val new_id = Counter.make (); abstype theory_id = Thy_Id of {id: id, (*identifier*) ids: Intset.T, (*cumulative identifiers -- symbolic body content*) name: string, (*official theory name*) stage: int} (*index for anonymous updates*) with fun rep_theory_id (Thy_Id args) = args; val make_theory_id = Thy_Id; end; (* theory allocation state *) type state = {stage: int} Synchronized.var; fun make_state () : state = Synchronized.var "Context.state" {stage = 0}; fun next_stage (state: state) = Synchronized.change_result state (fn {stage} => (stage + 1, {stage = stage + 1})); (* theory and proof context *) datatype theory = Thy_Undef | Thy of (*allocation state*) state * (*identity*) {theory_id: theory_id, theory_token: theory Unsynchronized.ref, theory_token_pos: Position.T} * (*ancestry*) {parents: theory list, (*immediate predecessors*) ancestors: theory list} * (*all predecessors -- canonical reverse order*) (*data*) Any.T Datatab.table; (*body content*) datatype proof = Prf_Undef | Prf of (*identity*) proof Unsynchronized.ref * (*token*) Position.T * (*token_pos*) theory * (*data*) Any.T Datatab.table; structure Proof = struct type context = proof end; datatype generic = Theory of theory | Proof of Proof.context; (* heap allocations *) -val trace_theories = Unsynchronized.ref false; -val trace_proofs = Unsynchronized.ref false; +val theory_tracing = Unsynchronized.ref false; +val proof_tracing = Unsynchronized.ref false; + +fun enabled_tracing () = ! theory_tracing orelse ! proof_tracing; local +fun cons_tokens var token = + Synchronized.change var (fn (n, tokens) => (n + 1, Weak.weak (SOME token) :: tokens)); + +fun finish_tokens var = + Synchronized.change_result var (fn (n, tokens) => + let + val tokens' = filter Unsynchronized.weak_active tokens; + val results = map_filter Unsynchronized.weak_peek tokens'; + in ((n, results), (n, tokens')) end); + fun make_token guard var token0 = if ! guard then let val token = Unsynchronized.ref (! token0); val pos = Position.thread_data (); - fun assign res = - (token := res; Synchronized.change var (cons (Weak.weak (SOME token))); res); + fun assign res = (token := res; cons_tokens var token; res); in (token, pos, assign) end else (token0, Position.none, I); -val theory_tokens = Synchronized.var "theory_tokens" ([]: theory Unsynchronized.weak_ref list); -val proof_tokens = Synchronized.var "proof_tokens" ([]: Proof.context Unsynchronized.weak_ref list); +val theory_tokens = Synchronized.var "theory_tokens" (0, []: theory Unsynchronized.weak_ref list); +val proof_tokens = Synchronized.var "proof_tokens" (0, []: proof Unsynchronized.weak_ref list); val theory_token0 = Unsynchronized.ref Thy_Undef; val proof_token0 = Unsynchronized.ref Prf_Undef; in -fun theory_token () = make_token trace_theories theory_tokens theory_token0; -fun proof_token () = make_token trace_proofs proof_tokens proof_token0; +fun theory_token () = make_token theory_tracing theory_tokens theory_token0; +fun proof_token () = make_token proof_tracing proof_tokens proof_token0; -fun allocations_trace () = +fun finish_tracing () = let val _ = ML_Heap.full_gc (); - val trace1 = Synchronized.value theory_tokens; - val trace2 = Synchronized.value proof_tokens; + val (total_theories, token_theories) = finish_tokens theory_tokens; + val (total_proofs, token_proofs) = finish_tokens proof_tokens; - fun cons1 r = - (case Unsynchronized.weak_peek r of - SOME (thy as Thy (_, {theory_token_pos, ...}, _, _)) => - cons (Theory thy, theory_token_pos) - | _ => I); - fun cons2 r = - (case Unsynchronized.weak_peek r of - SOME (ctxt as Prf (_, proof_token_pos, _, _)) => - cons (Proof ctxt, proof_token_pos) - | _ => I); + fun cons1 (thy as Thy (_, {theory_token_pos, ...}, _, _)) = cons (Theory thy, theory_token_pos) + | cons1 _ = I; + fun cons2 (ctxt as Prf (_, proof_token_pos, _, _)) = cons (Proof ctxt, proof_token_pos) + | cons2 _ = I; - val contexts = build (fold cons1 trace1 #> fold cons2 trace2); + val contexts = build (fold cons1 token_theories #> fold cons2 token_proofs); val active_theories = fold (fn (Theory _, _) => Integer.add 1 | _ => I) contexts 0; val active_proofs = fold (fn (Proof _, _) => Integer.add 1 | _ => I) contexts 0; - - val total_theories = length trace1; - val total_proofs = length trace2; in {contexts = contexts, active_contexts = active_theories + active_proofs, active_theories = active_theories, active_proofs = active_proofs, total_contexts = total_theories + total_proofs, total_theories = total_theories, total_proofs = total_proofs} end; end; (*** theory operations ***) fun rep_theory (Thy args) = args; exception THEORY of string * theory list; val state_of = #1 o rep_theory; val theory_identity = #2 o rep_theory; val theory_id = #theory_id o theory_identity; val identity_of = rep_theory_id o theory_id; val ancestry_of = #3 o rep_theory; val data_of = #4 o rep_theory; fun make_ancestry parents ancestors = {parents = parents, ancestors = ancestors}; fun stage_final stage = stage = 0; val theory_id_stage = #stage o rep_theory_id; val theory_id_final = stage_final o theory_id_stage; val theory_id_ord = int_ord o apply2 (#id o rep_theory_id); fun theory_id_name {long} thy_id = let val name = #name (rep_theory_id thy_id) in if long then name else Long_Name.base_name name end; val theory_long_name = #name o identity_of; val theory_base_name = Long_Name.base_name o theory_long_name; fun theory_name {long} = if long then theory_long_name else theory_base_name; val theory_identifier = #id o identity_of; val parents_of = #parents o ancestry_of; val ancestors_of = #ancestors o ancestry_of; (* names *) val PureN = "Pure"; fun display_name thy_id = let val name = theory_id_name {long = false} thy_id; val final = theory_id_final thy_id; in if final then name else name ^ ":" ^ string_of_int (theory_id_stage thy_id) end; fun display_names thy = let val name = display_name (theory_id thy); val ancestor_names = map theory_long_name (ancestors_of thy); in rev (name :: ancestor_names) end; val pretty_thy = Pretty.str_list "{" "}" o display_names; val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty_thy); fun pretty_abbrev_thy thy = let val names = display_names thy; val n = length names; val abbrev = if n > 5 then "..." :: List.drop (names, n - 5) else names; in Pretty.str_list "{" "}" abbrev end; fun get_theory long thy name = if theory_name long thy <> name then (case find_first (fn thy' => theory_name long thy' = name) (ancestors_of thy) of SOME thy' => thy' | NONE => error ("Unknown ancestor theory " ^ quote name)) else if theory_id_final (theory_id thy) then thy else error ("Unfinished theory " ^ quote name); (* identity *) fun merge_ids thys = fold (identity_of #> (fn {id, ids, ...} => fn acc => Intset.merge (acc, ids) |> Intset.insert id)) thys Intset.empty; val eq_thy_id = op = o apply2 (#id o rep_theory_id); val eq_thy = op = o apply2 (#id o identity_of); val proper_subthy_id = apply2 rep_theory_id #> (fn ({id, ...}, {ids, ...}) => Intset.member ids id); val proper_subthy = proper_subthy_id o apply2 theory_id; fun subthy_id p = eq_thy_id p orelse proper_subthy_id p; val subthy = subthy_id o apply2 theory_id; (* consistent ancestors *) fun eq_thy_consistent (thy1, thy2) = eq_thy (thy1, thy2) orelse (theory_base_name thy1 = theory_base_name thy2 andalso raise THEORY ("Duplicate theory name", [thy1, thy2])); fun extend_ancestors thy thys = if member eq_thy_consistent thys thy then raise THEORY ("Duplicate theory node", thy :: thys) else thy :: thys; val merge_ancestors = merge eq_thy_consistent; val eq_ancestry = apply2 ancestry_of #> (fn ({parents, ancestors}, {parents = parents', ancestors = ancestors'}) => eq_list eq_thy (parents, parents') andalso eq_list eq_thy (ancestors, ancestors')); (** theory data **) (* data kinds and access methods *) -val timing = Unsynchronized.ref false; +val data_timing = Unsynchronized.ref false; local type kind = {pos: Position.T, empty: Any.T, merge: (theory * Any.T) list -> Any.T}; val kinds = Synchronized.var "Theory_Data" (Datatab.empty: kind Datatab.table); fun the_kind k = (case Datatab.lookup (Synchronized.value kinds) k of SOME kind => kind | NONE => raise Fail "Invalid theory data identifier"); in fun data_kinds () = Datatab.fold_rev (fn (k, {pos, ...}) => cons (k, pos)) (Synchronized.value kinds) []; val invoke_pos = #pos o the_kind; val invoke_empty = #empty o the_kind; fun invoke_merge kind args = - if ! timing then + if ! data_timing then Timing.cond_timeit true ("Theory_Data.merge" ^ Position.here (#pos kind)) (fn () => #merge kind args) else #merge kind args; fun declare_data pos empty merge = let val k = data_kind (); val kind = {pos = pos, empty = empty, merge = merge}; val _ = Synchronized.change kinds (Datatab.update (k, kind)); in k end; fun lookup_data k thy = Datatab.lookup (data_of thy) k; fun get_data k thy = (case lookup_data k thy of SOME x => x | NONE => invoke_empty k); fun merge_data [] = Datatab.empty | merge_data [thy] = data_of thy | merge_data thys = let fun merge (k, kind) data = (case map_filter (fn thy => lookup_data k thy |> Option.map (pair thy)) thys of [] => data | [(_, x)] => Datatab.default (k, x) data | args => Datatab.update (k, invoke_merge kind args) data); in Datatab.fold merge (Synchronized.value kinds) (data_of (hd thys)) end; end; (** build theories **) (* create theory *) fun create_thy state ids name stage ancestry data = let val theory_id = make_theory_id {id = new_id (), ids = ids, name = name, stage = stage}; val (token, pos, assign) = theory_token (); val identity = {theory_id = theory_id, theory_token = token, theory_token_pos = pos}; in assign (Thy (state, identity, ancestry, data)) end; (* primitives *) val pre_pure_thy = let val state = make_state (); val stage = next_stage state; in create_thy state Intset.empty PureN stage (make_ancestry [] []) Datatab.empty end; local fun change_thy finish f thy = let val {name, stage, ...} = identity_of thy; val Thy (state, _, ancestry, data) = thy; val ancestry' = if stage_final stage then make_ancestry [thy] (extend_ancestors thy (ancestors_of thy)) else ancestry; val ids' = merge_ids [thy]; val stage' = if finish then 0 else next_stage state; val data' = f data; in create_thy state ids' name stage' ancestry' data' end; in val update_thy = change_thy false; val finish_thy = change_thy true I; end; (* join: unfinished theory nodes *) fun join_thys [] = raise List.Empty | join_thys thys = let val thy0 = hd thys; val name0 = theory_long_name thy0; val state0 = state_of thy0; fun ok thy = not (theory_id_final (theory_id thy)) andalso theory_long_name thy = name0 andalso eq_ancestry (thy0, thy); val _ = (case filter_out ok thys of [] => () | bad => raise THEORY ("Cannot join theories", bad)); val stage = next_stage state0; val ids = merge_ids thys; val data = merge_data thys; in create_thy state0 ids name0 stage (ancestry_of thy0) data end; (* merge: finished theory nodes *) fun make_parents thys = let val thys' = distinct eq_thy thys in thys' |> filter_out (fn thy => exists (fn thy' => proper_subthy (thy, thy')) thys') end; fun begin_thy name imports = if name = "" then error ("Bad theory name: " ^ quote name) else if null imports then error "Missing theory imports" else let val parents = make_parents imports; val ancestors = Library.foldl1 merge_ancestors (map ancestors_of parents) |> fold extend_ancestors parents; val ancestry = make_ancestry parents ancestors; val state = make_state (); val stage = next_stage state; val ids = merge_ids parents; val data = merge_data parents; in create_thy state ids name stage ancestry data |> tap finish_thy end; (* theory data *) structure Theory_Data = struct val declare = declare_data; fun get k dest thy = dest (get_data k thy); fun put k make x = update_thy (Datatab.update (k, make x)); fun sizeof1 k thy = Datatab.lookup (data_of thy) k |> Option.map ML_Heap.sizeof1; end; fun theory_data_sizeof1 thy = build (data_of thy |> Datatab.fold_rev (fn (k, _) => (case Theory_Data.sizeof1 k thy of NONE => I | SOME n => (cons (invoke_pos k, n))))); (*** proof context ***) (* proof data kinds *) local val kinds = Synchronized.var "Proof_Data" (Datatab.empty: (theory -> Any.T) Datatab.table); fun init_data thy = Synchronized.value kinds |> Datatab.map (fn _ => fn init => init thy); fun init_new_data thy = Synchronized.value kinds |> Datatab.fold (fn (k, init) => fn data => if Datatab.defined data k then data else Datatab.update (k, init thy) data); fun init_fallback k thy = (case Datatab.lookup (Synchronized.value kinds) k of SOME init => init thy | NONE => raise Fail "Invalid proof data identifier"); in fun raw_transfer thy' (ctxt as Prf (_, _, thy, data)) = if eq_thy (thy, thy') then ctxt else if proper_subthy (thy, thy') then let val (token', pos', assign) = proof_token (); val data' = init_new_data thy' data; in assign (Prf (token', pos', thy', data')) end else error "Cannot transfer proof context: not a super theory"; structure Proof_Context = struct fun theory_of (Prf (_, _, thy, _)) = thy; fun init_global thy = let val (token, pos, assign) = proof_token () in assign (Prf (token, pos, thy, init_data thy)) end; fun get_global long thy name = init_global (get_theory long thy name); end; structure Proof_Data = struct fun declare init = let val k = data_kind (); val _ = Synchronized.change kinds (Datatab.update (k, init)); in k end; fun get k dest (Prf (_, _, thy, data)) = (case Datatab.lookup data k of SOME x => x | NONE => init_fallback k thy) |> dest; fun put k make x (Prf (_, _, thy, data)) = let val (token', pos', assign) = proof_token (); val data' = Datatab.update (k, make x) data; in assign (Prf (token', pos', thy, data')) end; end; end; (*** theory certificate ***) datatype certificate = Certificate of theory | Certificate_Id of theory_id; fun certificate_theory (Certificate thy) = thy | certificate_theory (Certificate_Id thy_id) = error ("No content for theory certificate " ^ display_name thy_id); fun certificate_theory_id (Certificate thy) = theory_id thy | certificate_theory_id (Certificate_Id thy_id) = thy_id; fun eq_certificate (Certificate thy1, Certificate thy2) = eq_thy (thy1, thy2) | eq_certificate (Certificate_Id thy_id1, Certificate_Id thy_id2) = eq_thy_id (thy_id1, thy_id2) | eq_certificate _ = false; fun join_certificate (cert1, cert2) = let val (thy_id1, thy_id2) = apply2 certificate_theory_id (cert1, cert2) in if eq_thy_id (thy_id1, thy_id2) then (case cert1 of Certificate _ => cert1 | _ => cert2) else if proper_subthy_id (thy_id2, thy_id1) then cert1 else if proper_subthy_id (thy_id1, thy_id2) then cert2 else error ("Cannot join unrelated theory certificates " ^ display_name thy_id1 ^ " and " ^ display_name thy_id2) end; (*** generic context ***) fun cases f _ (Theory thy) = f thy | cases _ g (Proof prf) = g prf; fun mapping f g = cases (Theory o f) (Proof o g); fun mapping_result f g = cases (apsnd Theory o f) (apsnd Proof o g); val the_theory = cases I (fn _ => error "Ill-typed context: theory expected"); val the_proof = cases (fn _ => error "Ill-typed context: proof expected") I; fun map_theory f = Theory o f o the_theory; fun map_proof f = Proof o f o the_proof; fun map_theory_result f = apsnd Theory o f o the_theory; fun map_proof_result f = apsnd Proof o f o the_proof; fun theory_map f = the_theory o f o Theory; fun proof_map f = the_proof o f o Proof; val theory_of = cases I Proof_Context.theory_of; val proof_of = cases Proof_Context.init_global I; (** thread data **) local val generic_context_var = Thread_Data.var () : generic Thread_Data.var in fun get_generic_context () = Thread_Data.get generic_context_var; val put_generic_context = Thread_Data.put generic_context_var; fun setmp_generic_context opt_context = Thread_Data.setmp generic_context_var opt_context; fun the_generic_context () = (case get_generic_context () of SOME context => context | _ => error "Unknown context"); val the_global_context = theory_of o the_generic_context; val the_local_context = proof_of o the_generic_context; end; fun >>> f = let val (res, context') = f (the_generic_context ()); val _ = put_generic_context (SOME context'); in res end; nonfix >>; fun >> f = >>> (fn context => ((), f context)); val _ = put_generic_context (SOME (Theory pre_pure_thy)); end; structure Basic_Context: BASIC_CONTEXT = Context; open Basic_Context; (*** type-safe interfaces for data declarations ***) (** theory data **) signature THEORY_DATA'_ARGS = sig type T val empty: T val merge: (theory * T) list -> T end; signature THEORY_DATA_ARGS = sig type T val empty: T val merge: T * T -> T end; signature THEORY_DATA = sig type T val get: theory -> T val put: T -> theory -> theory val map: (T -> T) -> theory -> theory end; functor Theory_Data'(Data: THEORY_DATA'_ARGS): THEORY_DATA = struct type T = Data.T; exception Data of T; val kind = let val pos = Position.thread_data () in Context.Theory_Data.declare pos (Data Data.empty) (Data o Data.merge o map (fn (thy, Data x) => (thy, x))) end; val get = Context.Theory_Data.get kind (fn Data x => x); val put = Context.Theory_Data.put kind Data; fun map f thy = put (f (get thy)) thy; end; functor Theory_Data(Data: THEORY_DATA_ARGS): THEORY_DATA = Theory_Data' ( type T = Data.T; val empty = Data.empty; fun merge args = Library.foldl (fn (a, (_, b)) => Data.merge (a, b)) (#2 (hd args), tl args) ); (** proof data **) signature PROOF_DATA_ARGS = sig type T val init: theory -> T end; signature PROOF_DATA = sig type T val get: Proof.context -> T val put: T -> Proof.context -> Proof.context val map: (T -> T) -> Proof.context -> Proof.context end; functor Proof_Data(Data: PROOF_DATA_ARGS): PROOF_DATA = struct type T = Data.T; exception Data of T; val kind = Context.Proof_Data.declare (Data o Data.init); val get = Context.Proof_Data.get kind (fn Data x => x); val put = Context.Proof_Data.put kind Data; fun map f prf = put (f (get prf)) prf; end; (** generic data **) signature GENERIC_DATA_ARGS = sig type T val empty: T val merge: T * T -> T end; signature GENERIC_DATA = sig type T val get: Context.generic -> T val put: T -> Context.generic -> Context.generic val map: (T -> T) -> Context.generic -> Context.generic end; functor Generic_Data(Data: GENERIC_DATA_ARGS): GENERIC_DATA = struct structure Thy_Data = Theory_Data(Data); structure Prf_Data = Proof_Data(type T = Data.T val init = Thy_Data.get); type T = Data.T; fun get (Context.Theory thy) = Thy_Data.get thy | get (Context.Proof prf) = Prf_Data.get prf; fun put x (Context.Theory thy) = Context.Theory (Thy_Data.put x thy) | put x (Context.Proof prf) = Context.Proof (Prf_Data.put x prf); fun map f ctxt = put (f (get ctxt)) ctxt; end; (*hide private interface*) structure Context: CONTEXT = Context; diff --git a/src/Pure/thm.ML b/src/Pure/thm.ML --- a/src/Pure/thm.ML +++ b/src/Pure/thm.ML @@ -1,2414 +1,2421 @@ (* Title: Pure/thm.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Makarius The very core of Isabelle's Meta Logic: certified types and terms, derivations, theorems, inference rules (including lifting and resolution), oracles. *) infix 0 RS RSN; signature BASIC_THM = sig type ctyp type cterm exception CTERM of string * cterm list type thm type conv = cterm -> thm exception THM of string * int * thm list val RSN: thm * (int * thm) -> thm val RS: thm * thm -> thm end; signature THM = sig include BASIC_THM (*certified types*) val typ_of: ctyp -> typ val global_ctyp_of: theory -> typ -> ctyp val ctyp_of: Proof.context -> typ -> ctyp val dest_ctyp: ctyp -> ctyp list val dest_ctypN: int -> ctyp -> ctyp val dest_ctyp0: ctyp -> ctyp val dest_ctyp1: ctyp -> ctyp val make_ctyp: ctyp -> ctyp list -> ctyp (*certified terms*) val term_of: cterm -> term val typ_of_cterm: cterm -> typ val ctyp_of_cterm: cterm -> ctyp val maxidx_of_cterm: cterm -> int val global_cterm_of: theory -> term -> cterm val cterm_of: Proof.context -> term -> cterm val renamed_term: term -> cterm -> cterm val fast_term_ord: cterm ord val term_ord: cterm ord val dest_comb: cterm -> cterm * cterm val dest_fun: cterm -> cterm val dest_arg: cterm -> cterm val dest_fun2: cterm -> cterm val dest_arg1: cterm -> cterm val dest_abs_fresh: string -> cterm -> cterm * cterm val dest_abs_global: cterm -> cterm * cterm val rename_tvar: indexname -> ctyp -> ctyp val var: indexname * ctyp -> cterm val apply: cterm -> cterm -> cterm val lambda_name: string * cterm -> cterm -> cterm val lambda: cterm -> cterm -> cterm val adjust_maxidx_cterm: int -> cterm -> cterm val incr_indexes_cterm: int -> cterm -> cterm val match: cterm * cterm -> ctyp TVars.table * cterm Vars.table val first_order_match: cterm * cterm -> ctyp TVars.table * cterm Vars.table (*theorems*) val fold_terms: {hyps: bool} -> (term -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_ctyps: {hyps: bool} -> (typ -> bool) -> (ctyp -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_cterms: {hyps: bool} -> (term -> bool) -> (cterm -> 'a -> 'a) -> thm -> 'a -> 'a val terms_of_tpairs: (term * term) list -> term list val full_prop_of: thm -> term val theory_id: thm -> Context.theory_id val theory_name: {long: bool} -> thm -> string val theory_long_name: thm -> string val theory_base_name: thm -> string val maxidx_of: thm -> int val maxidx_thm: thm -> int -> int val shyps_of: thm -> sort Ord_List.T val hyps_of: thm -> term list val prop_of: thm -> term val tpairs_of: thm -> (term * term) list val concl_of: thm -> term val prems_of: thm -> term list val nprems_of: thm -> int val no_prems: thm -> bool val major_prem_of: thm -> term val cprop_of: thm -> cterm val cprem_of: thm -> int -> cterm val cconcl_of: thm -> cterm val cprems_of: thm -> cterm list val chyps_of: thm -> cterm list val thm_ord: thm ord exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option val theory_of_cterm: cterm -> theory val theory_of_thm: thm -> theory val trim_context_ctyp: ctyp -> ctyp val trim_context_cterm: cterm -> cterm val transfer_ctyp: theory -> ctyp -> ctyp val transfer_cterm: theory -> cterm -> cterm val transfer: theory -> thm -> thm val transfer': Proof.context -> thm -> thm val transfer'': Context.generic -> thm -> thm val join_transfer: theory -> thm -> thm val join_transfer_context: Proof.context * thm -> Proof.context * thm val renamed_prop: term -> thm -> thm val weaken: cterm -> thm -> thm val weaken_sorts: sort list -> cterm -> cterm val proof_bodies_of: thm list -> proof_body list val proof_body_of: thm -> proof_body val proof_of: thm -> proof val reconstruct_proof_of: thm -> Proofterm.proof val consolidate: thm list -> unit val expose_proofs: theory -> thm list -> unit val expose_proof: theory -> thm -> unit val future: thm future -> cterm -> thm val thm_deps: thm -> Proofterm.thm Ord_List.T val extra_shyps: thm -> sort list val strip_shyps: thm -> thm val derivation_closed: thm -> bool val derivation_name: thm -> string val derivation_id: thm -> Proofterm.thm_id option val raw_derivation_name: thm -> string val expand_name: thm -> Proofterm.thm_header -> string option val name_derivation: string * Position.T -> thm -> thm val close_derivation: Position.T -> thm -> thm val trim_context: thm -> thm val axiom: theory -> string -> thm val all_axioms_of: theory -> (string * thm) list val get_tags: thm -> Properties.T val map_tags: (Properties.T -> Properties.T) -> thm -> thm val norm_proof: thm -> thm val adjust_maxidx_thm: int -> thm -> thm (*type classes*) val the_classrel: theory -> class * class -> thm val the_arity: theory -> string * sort list * class -> thm val classrel_proof: theory -> class * class -> proof val arity_proof: theory -> string * sort list * class -> proof (*oracles*) val add_oracle: binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory val oracle_space: theory -> Name_Space.T val pretty_oracle: Proof.context -> string -> Pretty.T val extern_oracles: bool -> Proof.context -> (Markup.T * xstring) list val check_oracle: Proof.context -> xstring * Position.T -> string (*inference rules*) val assume: cterm -> thm val implies_intr: cterm -> thm -> thm val implies_elim: thm -> thm -> thm val forall_intr: cterm -> thm -> thm val forall_elim: cterm -> thm -> thm val reflexive: cterm -> thm val symmetric: thm -> thm val transitive: thm -> thm -> thm val beta_conversion: bool -> conv val eta_conversion: conv val eta_long_conversion: conv val abstract_rule: string -> cterm -> thm -> thm val combination: thm -> thm -> thm val equal_intr: thm -> thm -> thm val equal_elim: thm -> thm -> thm val solve_constraints: thm -> thm val flexflex_rule: Proof.context option -> thm -> thm Seq.seq val generalize: Names.set * Names.set -> int -> thm -> thm val generalize_cterm: Names.set * Names.set -> int -> cterm -> cterm val generalize_ctyp: Names.set -> int -> ctyp -> ctyp val instantiate: ctyp TVars.table * cterm Vars.table -> thm -> thm val instantiate_beta: ctyp TVars.table * cterm Vars.table -> thm -> thm val instantiate_cterm: ctyp TVars.table * cterm Vars.table -> cterm -> cterm val instantiate_beta_cterm: ctyp TVars.table * cterm Vars.table -> cterm -> cterm val trivial: cterm -> thm val of_class: ctyp * class -> thm val unconstrainT: thm -> thm val varifyT_global': TFrees.set -> thm -> ((string * sort) * indexname) list * thm val varifyT_global: thm -> thm val legacy_freezeT: thm -> thm val plain_prop_of: thm -> term val dest_state: thm * int -> (term * term) list * term list * term * term val lift_rule: cterm -> thm -> thm val incr_indexes: int -> thm -> thm val assumption: Proof.context option -> int -> thm -> thm Seq.seq val eq_assumption: int -> thm -> thm val rotate_rule: int -> int -> thm -> thm val permute_prems: int -> int -> thm -> thm val bicompose: Proof.context option -> {flatten: bool, match: bool, incremented: bool} -> bool * thm * int -> int -> thm -> thm Seq.seq val biresolution: Proof.context option -> bool -> (bool * thm) list -> int -> thm -> thm Seq.seq val theory_names_of_arity: {long: bool} -> theory -> string * class -> string list val add_classrel: thm -> theory -> theory val add_arity: thm -> theory -> theory end; structure Thm: THM = struct (*** Certified terms and types ***) (** certified types **) datatype ctyp = Ctyp of {cert: Context.certificate, T: typ, maxidx: int, sorts: sort Ord_List.T}; fun typ_of (Ctyp {T, ...}) = T; fun global_ctyp_of thy raw_T = let val T = Sign.certify_typ thy raw_T; val maxidx = Term.maxidx_of_typ T; val sorts = Sorts.insert_typ T []; in Ctyp {cert = Context.Certificate thy, T = T, maxidx = maxidx, sorts = sorts} end; val ctyp_of = global_ctyp_of o Proof_Context.theory_of; fun dest_ctyp (Ctyp {cert, T = Type (_, Ts), maxidx, sorts}) = map (fn T => Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts}) Ts | dest_ctyp cT = raise TYPE ("dest_ctyp", [typ_of cT], []); fun dest_ctypN n (Ctyp {cert, T, maxidx, sorts}) = let fun err () = raise TYPE ("dest_ctypN", [T], []) in (case T of Type (_, Ts) => Ctyp {cert = cert, T = nth Ts n handle General.Subscript => err (), maxidx = maxidx, sorts = sorts} | _ => err ()) end; val dest_ctyp0 = dest_ctypN 0; val dest_ctyp1 = dest_ctypN 1; fun join_certificate_ctyp (Ctyp {cert, ...}) cert0 = Context.join_certificate (cert0, cert); fun union_sorts_ctyp (Ctyp {sorts, ...}) sorts0 = Sorts.union sorts0 sorts; fun maxidx_ctyp (Ctyp {maxidx, ...}) maxidx0 = Int.max (maxidx0, maxidx); fun make_ctyp (Ctyp {cert, T, maxidx = _, sorts = _}) cargs = let val As = map typ_of cargs; fun err () = raise TYPE ("make_ctyp", T :: As, []); in (case T of Type (a, args) => Ctyp { cert = fold join_certificate_ctyp cargs cert, maxidx = fold maxidx_ctyp cargs ~1, sorts = fold union_sorts_ctyp cargs [], T = if length args = length cargs then Type (a, As) else err ()} | _ => err ()) end; (** certified terms **) (*certified terms with checked typ, maxidx, and sorts*) datatype cterm = Cterm of {cert: Context.certificate, t: term, T: typ, maxidx: int, sorts: sort Ord_List.T}; exception CTERM of string * cterm list; fun term_of (Cterm {t, ...}) = t; fun typ_of_cterm (Cterm {T, ...}) = T; fun ctyp_of_cterm (Cterm {cert, T, maxidx, sorts, ...}) = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts}; fun maxidx_of_cterm (Cterm {maxidx, ...}) = maxidx; fun global_cterm_of thy tm = let val (t, T, maxidx) = Sign.certify_term thy tm; val sorts = Sorts.insert_term t []; in Cterm {cert = Context.Certificate thy, t = t, T = T, maxidx = maxidx, sorts = sorts} end; val cterm_of = global_cterm_of o Proof_Context.theory_of; fun join_certificate0 (Cterm {cert = cert1, ...}, Cterm {cert = cert2, ...}) = Context.join_certificate (cert1, cert2); fun renamed_term t' (Cterm {cert, t, T, maxidx, sorts}) = if t aconv t' then Cterm {cert = cert, t = t', T = T, maxidx = maxidx, sorts = sorts} else raise TERM ("renamed_term: terms disagree", [t, t']); val fast_term_ord = Term_Ord.fast_term_ord o apply2 term_of; val term_ord = Term_Ord.term_ord o apply2 term_of; (* destructors *) fun dest_comb (Cterm {t = c $ a, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in (Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts}, Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts}) end | dest_comb ct = raise CTERM ("dest_comb", [ct]); fun dest_fun (Cterm {t = c $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun ct = raise CTERM ("dest_fun", [ct]); fun dest_arg (Cterm {t = c $ a, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg ct = raise CTERM ("dest_arg", [ct]); fun dest_fun2 (Cterm {t = c $ _ $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0; val B = Term.argument_type_of c 1; in Cterm {t = c, T = A --> B --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun2 ct = raise CTERM ("dest_fun2", [ct]); fun dest_arg1 (Cterm {t = c $ a $ _, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg1 ct = raise CTERM ("dest_arg1", [ct]); fun gen_dest_abs dest ct = (case ct of Cterm {t = t as Abs _, T = Type ("fun", [_, U]), cert, maxidx, sorts} => let val ((x', T), t') = dest t; val v = Cterm {t = Free (x', T), T = T, cert = cert, maxidx = maxidx, sorts = sorts}; val body = Cterm {t = t', T = U, cert = cert, maxidx = maxidx, sorts = sorts}; in (v, body) end | _ => raise CTERM ("dest_abs", [ct])); val dest_abs_fresh = gen_dest_abs o Term.dest_abs_fresh; val dest_abs_global = gen_dest_abs Term.dest_abs_global; (* constructors *) fun rename_tvar (a, i) (Ctyp {cert, T, maxidx, sorts}) = let val S = (case T of TFree (_, S) => S | TVar (_, S) => S | _ => raise TYPE ("rename_tvar: no variable", [T], [])); val _ = if i < 0 then raise TYPE ("rename_tvar: bad index", [TVar ((a, i), S)], []) else (); in Ctyp {cert = cert, T = TVar ((a, i), S), maxidx = Int.max (i, maxidx), sorts = sorts} end; fun var ((x, i), Ctyp {cert, T, maxidx, sorts}) = if i < 0 then raise TERM ("var: bad index", [Var ((x, i), T)]) else Cterm {cert = cert, t = Var ((x, i), T), T = T, maxidx = Int.max (i, maxidx), sorts = sorts}; fun apply (cf as Cterm {t = f, T = Type ("fun", [dty, rty]), maxidx = maxidx1, sorts = sorts1, ...}) (cx as Cterm {t = x, T, maxidx = maxidx2, sorts = sorts2, ...}) = if T = dty then Cterm {cert = join_certificate0 (cf, cx), t = f $ x, T = rty, maxidx = Int.max (maxidx1, maxidx2), sorts = Sorts.union sorts1 sorts2} else raise CTERM ("apply: types don't agree", [cf, cx]) | apply cf cx = raise CTERM ("apply: first arg is not a function", [cf, cx]); fun lambda_name (x, ct1 as Cterm {t = t1, T = T1, maxidx = maxidx1, sorts = sorts1, ...}) (ct2 as Cterm {t = t2, T = T2, maxidx = maxidx2, sorts = sorts2, ...}) = let val t = Term.lambda_name (x, t1) t2 in Cterm {cert = join_certificate0 (ct1, ct2), t = t, T = T1 --> T2, maxidx = Int.max (maxidx1, maxidx2), sorts = Sorts.union sorts1 sorts2} end; fun lambda t u = lambda_name ("", t) u; (* indexes *) fun adjust_maxidx_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) = if maxidx = i then ct else if maxidx < i then Cterm {maxidx = i, cert = cert, t = t, T = T, sorts = sorts} else Cterm {maxidx = Int.max (maxidx_of_term t, i), cert = cert, t = t, T = T, sorts = sorts}; fun incr_indexes_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) = if i < 0 then raise CTERM ("negative increment", [ct]) else if i = 0 then ct else Cterm {cert = cert, t = Logic.incr_indexes ([], [], i) t, T = Logic.incr_tvar i T, maxidx = maxidx + i, sorts = sorts}; (*** Derivations and Theorems ***) (* sort constraints *) type constraint = {theory: theory, typ: typ, sort: sort}; local val constraint_ord : constraint ord = Context.theory_id_ord o apply2 (Context.theory_id o #theory) ||| Term_Ord.typ_ord o apply2 #typ ||| Term_Ord.sort_ord o apply2 #sort; val smash_atyps = map_atyps (fn TVar (_, S) => Term.aT S | TFree (_, S) => Term.aT S | T => T); in val union_constraints = Ord_List.union constraint_ord; fun insert_constraints thy (T, S) = let val ignored = S = [] orelse (case T of TFree (_, S') => S = S' | TVar (_, S') => S = S' | _ => false); in if ignored then I else Ord_List.insert constraint_ord {theory = thy, typ = smash_atyps T, sort = S} end; fun insert_constraints_env thy env = let val tyenv = Envir.type_env env; fun insert ([], _) = I | insert (S, T) = insert_constraints thy (Envir.norm_type tyenv T, S); in tyenv |> Vartab.fold (insert o #2) end; end; (* datatype thm *) datatype thm = Thm of deriv * (*derivation*) {cert: Context.certificate, (*background theory certificate*) tags: Properties.T, (*additional annotations/comments*) maxidx: int, (*maximum index of any Var or TVar*) constraints: constraint Ord_List.T, (*implicit proof obligations for sort constraints*) shyps: sort Ord_List.T, (*sort hypotheses*) hyps: term Ord_List.T, (*hypotheses*) tpairs: (term * term) list, (*flex-flex pairs*) prop: term} (*conclusion*) and deriv = Deriv of {promises: (serial * thm future) Ord_List.T, body: Proofterm.proof_body}; type conv = cterm -> thm; (*errors involving theorems*) exception THM of string * int * thm list; fun rep_thm (Thm (_, args)) = args; fun fold_terms h f (Thm (_, {tpairs, prop, hyps, ...})) = fold (fn (t, u) => f t #> f u) tpairs #> f prop #> #hyps h ? fold f hyps; fun fold_atomic_ctyps h g f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun ctyp T = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = shyps} in (fold_terms h o fold_types o fold_atyps) (fn T => if g T then f (ctyp T) else I) th end; fun fold_atomic_cterms h g f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun cterm t T = Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = shyps}; fun apply t T = if g t then f (cterm t T) else I; in (fold_terms h o fold_aterms) (fn t as Const (_, T) => apply t T | t as Free (_, T) => apply t T | t as Var (_, T) => apply t T | _ => I) th end; fun terms_of_tpairs tpairs = fold_rev (fn (t, u) => cons t o cons u) tpairs []; fun eq_tpairs ((t, u), (t', u')) = t aconv t' andalso u aconv u'; fun union_tpairs ts us = Library.merge eq_tpairs (ts, us); val maxidx_tpairs = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u); fun attach_tpairs tpairs prop = Logic.list_implies (map Logic.mk_equals tpairs, prop); fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop; val union_hyps = Ord_List.union Term_Ord.fast_term_ord; val insert_hyps = Ord_List.insert Term_Ord.fast_term_ord; val remove_hyps = Ord_List.remove Term_Ord.fast_term_ord; fun join_certificate1 (Cterm {cert = cert1, ...}, Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); fun join_certificate2 (Thm (_, {cert = cert1, ...}), Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); (* basic components *) val cert_of = #cert o rep_thm; val theory_id = Context.certificate_theory_id o cert_of; fun theory_name long = Context.theory_id_name long o theory_id; val theory_long_name = theory_name {long = true}; val theory_base_name = theory_name {long = false}; val maxidx_of = #maxidx o rep_thm; fun maxidx_thm th i = Int.max (maxidx_of th, i); val shyps_of = #shyps o rep_thm; val hyps_of = #hyps o rep_thm; val prop_of = #prop o rep_thm; val tpairs_of = #tpairs o rep_thm; val concl_of = Logic.strip_imp_concl o prop_of; val prems_of = Logic.strip_imp_prems o prop_of; val nprems_of = Logic.count_prems o prop_of; val no_prems = Logic.no_prems o prop_of; fun major_prem_of th = (case prems_of th of prem :: _ => Logic.strip_assums_concl prem | [] => raise THM ("major_prem_of: rule with no premises", 0, [th])); fun cprop_of (Thm (_, {cert, maxidx, shyps, prop, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, t = prop, sorts = shyps}; fun cprem_of (th as Thm (_, {cert, maxidx, shyps, prop, ...})) i = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = Logic.nth_prem (i, prop) handle TERM _ => raise THM ("cprem_of", i, [th])}; fun cconcl_of (th as Thm (_, {cert, maxidx, shyps, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = concl_of th}; fun cprems_of (th as Thm (_, {cert, maxidx, shyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = t}) (prems_of th); fun chyps_of (Thm (_, {cert, shyps, hyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = ~1, T = propT, sorts = shyps, t = t}) hyps; (* thm order: ignores theory context! *) val thm_ord = pointer_eq_ord (Term_Ord.fast_term_ord o apply2 prop_of ||| list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) o apply2 tpairs_of ||| list_ord Term_Ord.fast_term_ord o apply2 hyps_of ||| list_ord Term_Ord.sort_ord o apply2 shyps_of); (* implicit theory context *) exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option; fun theory_of_cterm (ct as Cterm {cert, ...}) = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct], [], NONE); fun theory_of_thm th = Context.certificate_theory (cert_of th) handle ERROR msg => raise CONTEXT (msg, [], [], [th], NONE); fun trim_context_ctyp cT = (case cT of Ctyp {cert = Context.Certificate_Id _, ...} => cT | Ctyp {cert = Context.Certificate thy, T, maxidx, sorts} => Ctyp {cert = Context.Certificate_Id (Context.theory_id thy), T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_cterm ct = (case ct of Cterm {cert = Context.Certificate_Id _, ...} => ct | Cterm {cert = Context.Certificate thy, t, T, maxidx, sorts} => Cterm {cert = Context.Certificate_Id (Context.theory_id thy), t = t, T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_thm th = (case th of Thm (_, {constraints = _ :: _, ...}) => raise THM ("trim_context: pending sort constraints", 0, [th]) | Thm (_, {cert = Context.Certificate_Id _, ...}) => th | Thm (der, {cert = Context.Certificate thy, tags, maxidx, constraints = [], shyps, hyps, tpairs, prop}) => Thm (der, {cert = Context.Certificate_Id (Context.theory_id thy), tags = tags, maxidx = maxidx, constraints = [], shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop})); fun transfer_ctyp thy' cT = let val Ctyp {cert, T, maxidx, sorts} = cT; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [cT], [], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then cT else Ctyp {cert = cert', T = T, maxidx = maxidx, sorts = sorts} end; fun transfer_cterm thy' ct = let val Cterm {cert, t, T, maxidx, sorts} = ct; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [ct], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then ct else Cterm {cert = cert', t = t, T = T, maxidx = maxidx, sorts = sorts} end; fun transfer thy' th = let val Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) = th; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [], [th], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then th else Thm (der, {cert = cert', tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; val transfer' = transfer o Proof_Context.theory_of; val transfer'' = transfer o Context.theory_of; fun join_transfer thy th = (Context.subthy_id (theory_id th, Context.theory_id thy) ? transfer thy) th; fun join_transfer_context (ctxt, th) = if Context.subthy_id (theory_id th, Context.theory_id (Proof_Context.theory_of ctxt)) then (ctxt, transfer' ctxt th) else (Context.raw_transfer (theory_of_thm th) ctxt, th); (* matching *) local fun gen_match match (ct1 as Cterm {t = t1, sorts = sorts1, ...}, ct2 as Cterm {t = t2, sorts = sorts2, maxidx = maxidx2, ...}) = let val cert = join_certificate0 (ct1, ct2); val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct1, ct2], [], NONE); val (Tinsts, tinsts) = match thy (t1, t2) (Vartab.empty, Vartab.empty); val sorts = Sorts.union sorts1 sorts2; fun mk_cTinst ((a, i), (S, T)) = (((a, i), S), Ctyp {T = T, cert = cert, maxidx = maxidx2, sorts = sorts}); fun mk_ctinst ((x, i), (U, t)) = let val T = Envir.subst_type Tinsts U in (((x, i), T), Cterm {t = t, T = T, cert = cert, maxidx = maxidx2, sorts = sorts}) end; in (TVars.build (Vartab.fold (TVars.add o mk_cTinst) Tinsts), Vars.build (Vartab.fold (Vars.add o mk_ctinst) tinsts)) end; in val match = gen_match Pattern.match; val first_order_match = gen_match Pattern.first_order_match; end; (*implicit alpha-conversion*) fun renamed_prop prop' (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if prop aconv prop' then Thm (der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) else raise TERM ("renamed_prop: props disagree", [prop, prop']); fun make_context ths NONE cert = (Context.Theory (Context.certificate_theory cert) handle ERROR msg => raise CONTEXT (msg, [], [], ths, NONE)) | make_context ths (SOME ctxt) cert = let val thy_id = Context.certificate_theory_id cert; val thy_id' = Context.theory_id (Proof_Context.theory_of ctxt); in if Context.subthy_id (thy_id, thy_id') then Context.Proof ctxt else raise CONTEXT ("Bad context", [], [], ths, SOME (Context.Proof ctxt)) end; fun make_context_certificate ths opt_ctxt cert = let val context = make_context ths opt_ctxt cert; val cert' = Context.Certificate (Context.theory_of context); in (context, cert') end; (*explicit weakening: maps |- B to A |- B*) fun weaken raw_ct th = let val ct as Cterm {t = A, T, sorts, maxidx = maxidxA, ...} = adjust_maxidx_cterm ~1 raw_ct; val Thm (der, {tags, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; in if T <> propT then raise THM ("weaken: assumptions must have type prop", 0, []) else if maxidxA <> ~1 then raise THM ("weaken: assumptions may not contain schematic variables", maxidxA, []) else Thm (der, {cert = join_certificate1 (ct, th), tags = tags, maxidx = maxidx, constraints = constraints, shyps = Sorts.union sorts shyps, hyps = insert_hyps A hyps, tpairs = tpairs, prop = prop}) end; fun weaken_sorts raw_sorts ct = let val Cterm {cert, t, T, maxidx, sorts} = ct; val thy = theory_of_cterm ct; val more_sorts = Sorts.make (map (Sign.certify_sort thy) raw_sorts); val sorts' = Sorts.union sorts more_sorts; in Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = sorts'} end; (** derivations and promised proofs **) fun make_deriv promises oracles thms proof = Deriv {promises = promises, body = PBody {oracles = oracles, thms = thms, proof = proof}}; val empty_deriv = make_deriv [] [] [] MinProof; (* inference rules *) val promise_ord: (serial * thm future) ord = fn ((i, _), (j, _)) => int_ord (j, i); fun bad_proofs i = error ("Illegal level of detail for proof objects: " ^ string_of_int i); fun deriv_rule2 f (Deriv {promises = ps1, body = PBody {oracles = oracles1, thms = thms1, proof = prf1}}) (Deriv {promises = ps2, body = PBody {oracles = oracles2, thms = thms2, proof = prf2}}) = let val ps = Ord_List.union promise_ord ps1 ps2; val oracles = Proofterm.union_oracles oracles1 oracles2; val thms = Proofterm.union_thms thms1 thms2; val prf = (case ! Proofterm.proofs of 2 => f prf1 prf2 | 1 => MinProof | 0 => MinProof | i => bad_proofs i); in make_deriv ps oracles thms prf end; fun deriv_rule1 f = deriv_rule2 (K f) empty_deriv; fun deriv_rule0 make_prf = if ! Proofterm.proofs <= 1 then empty_deriv else deriv_rule1 I (make_deriv [] [] [] (make_prf ())); fun deriv_rule_unconditional f (Deriv {promises, body = PBody {oracles, thms, proof}}) = make_deriv promises oracles thms (f proof); (* fulfilled proofs *) fun raw_promises_of (Thm (Deriv {promises, ...}, _)) = promises; fun join_promises [] = () | join_promises promises = join_promises_of (Future.joins (map snd promises)) and join_promises_of thms = join_promises (Ord_List.make promise_ord (maps raw_promises_of thms)); fun fulfill_body (th as Thm (Deriv {promises, body}, _)) = let val fulfilled_promises = map #1 promises ~~ map fulfill_body (Future.joins (map #2 promises)) in Proofterm.fulfill_norm_proof (theory_of_thm th) fulfilled_promises body end; fun proof_bodies_of thms = (join_promises_of thms; map fulfill_body thms); val proof_body_of = singleton proof_bodies_of; val proof_of = Proofterm.proof_of o proof_body_of; fun reconstruct_proof_of thm = Proofterm.reconstruct_proof (theory_of_thm thm) (prop_of thm) (proof_of thm); val consolidate = ignore o proof_bodies_of; fun expose_proofs thy thms = if Proofterm.export_proof_boxes_required thy then Proofterm.export_proof_boxes (proof_bodies_of (map (transfer thy) thms)) else (); fun expose_proof thy = expose_proofs thy o single; (* future rule *) fun future_result i orig_cert orig_shyps orig_prop thm = let fun err msg = raise THM ("future_result: " ^ msg, 0, [thm]); val Thm (Deriv {promises, ...}, {cert, constraints, shyps, hyps, tpairs, prop, ...}) = thm; val _ = Context.eq_certificate (cert, orig_cert) orelse err "bad theory"; val _ = prop aconv orig_prop orelse err "bad prop"; val _ = null constraints orelse err "bad sort constraints"; val _ = null tpairs orelse err "bad flex-flex constraints"; val _ = null hyps orelse err "bad hyps"; val _ = Sorts.subset (shyps, orig_shyps) orelse err "bad shyps"; val _ = forall (fn (j, _) => i <> j) promises orelse err "bad dependencies"; val _ = join_promises promises; in thm end; fun future future_thm ct = let val Cterm {cert = cert, t = prop, T, maxidx, sorts} = ct; val _ = T <> propT andalso raise CTERM ("future: prop expected", [ct]); val _ = if Proofterm.proofs_enabled () then raise CTERM ("future: proof terms enabled", [ct]) else (); val i = serial (); val future = future_thm |> Future.map (future_result i cert sorts prop); in Thm (make_deriv [(i, future)] [] [] MinProof, {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) end; (** Axioms **) fun axiom thy name = (case Name_Space.lookup (Theory.axiom_table thy) name of SOME prop => let val der = deriv_rule0 (fn () => Proofterm.axm_proof name prop); val cert = Context.Certificate thy; val maxidx = maxidx_of_term prop; val shyps = Sorts.insert_term prop []; in Thm (der, {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = shyps, hyps = [], tpairs = [], prop = prop}) end | NONE => raise THEORY ("No axiom " ^ quote name, [thy])); fun all_axioms_of thy = map (fn (name, _) => (name, axiom thy name)) (Theory.all_axioms_of thy); (* tags *) val get_tags = #tags o rep_thm; fun map_tags f (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = Thm (der, {cert = cert, tags = f tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (* technical adjustments *) fun norm_proof (th as Thm (der, args)) = Thm (deriv_rule1 (Proofterm.rew_proof (theory_of_thm th)) der, args); fun adjust_maxidx_thm i (th as Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if maxidx = i then th else if maxidx < i then Thm (der, {maxidx = i, cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) else Thm (der, {maxidx = Int.max (maxidx_tpairs tpairs (maxidx_of_term prop), i), cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (*** Theory data ***) (* type classes *) structure Aritytab = Table( type key = string * sort list * class; val ord = fast_string_ord o apply2 #1 ||| fast_string_ord o apply2 #3 ||| list_ord Term_Ord.sort_ord o apply2 #2; ); datatype classes = Classes of {classrels: thm Symreltab.table, arities: (thm * string * serial) Aritytab.table}; fun make_classes (classrels, arities) = Classes {classrels = classrels, arities = arities}; val empty_classes = make_classes (Symreltab.empty, Aritytab.empty); (*see Theory.at_begin hook for transitive closure of classrels and arity completion*) fun merge_classes (Classes {classrels = classrels1, arities = arities1}, Classes {classrels = classrels2, arities = arities2}) = let val classrels' = Symreltab.merge (K true) (classrels1, classrels2); val arities' = Aritytab.merge (K true) (arities1, arities2); in make_classes (classrels', arities') end; (* data *) structure Data = Theory_Data ( type T = unit Name_Space.table * (*oracles: authentic derivation names*) classes; (*type classes within the logic*) val empty : T = (Name_Space.empty_table Markup.oracleN, empty_classes); fun merge ((oracles1, sorts1), (oracles2, sorts2)) : T = (Name_Space.merge_tables (oracles1, oracles2), merge_classes (sorts1, sorts2)); ); val get_oracles = #1 o Data.get; val map_oracles = Data.map o apfst; val get_classes = (fn (_, Classes args) => args) o Data.get; val get_classrels = #classrels o get_classes; val get_arities = #arities o get_classes; fun map_classes f = (Data.map o apsnd) (fn Classes {classrels, arities} => make_classes (f (classrels, arities))); fun map_classrels f = map_classes (fn (classrels, arities) => (f classrels, arities)); fun map_arities f = map_classes (fn (classrels, arities) => (classrels, f arities)); (* type classes *) fun the_classrel thy (c1, c2) = (case Symreltab.lookup (get_classrels thy) (c1, c2) of SOME thm => transfer thy thm | NONE => error ("Unproven class relation " ^ Syntax.string_of_classrel (Proof_Context.init_global thy) [c1, c2])); fun the_arity thy (a, Ss, c) = (case Aritytab.lookup (get_arities thy) (a, Ss, c) of SOME (thm, _, _) => transfer thy thm | NONE => error ("Unproven type arity " ^ Syntax.string_of_arity (Proof_Context.init_global thy) (a, Ss, [c]))); val classrel_proof = proof_of oo the_classrel; val arity_proof = proof_of oo the_arity; (* solve sort constraints by pro-forma proof *) local fun union_digest (oracles1, thms1) (oracles2, thms2) = (Proofterm.union_oracles oracles1 oracles2, Proofterm.union_thms thms1 thms2); fun thm_digest (Thm (Deriv {body = PBody {oracles, thms, ...}, ...}, _)) = (oracles, thms); fun constraint_digest ({theory = thy, typ, sort, ...}: constraint) = Sorts.of_sort_derivation (Sign.classes_of thy) {class_relation = fn _ => fn _ => fn (digest, c1) => fn c2 => if c1 = c2 then ([], []) else union_digest digest (thm_digest (the_classrel thy (c1, c2))), type_constructor = fn (a, _) => fn dom => fn c => let val arity_digest = thm_digest (the_arity thy (a, (map o map) #2 dom, c)) in (fold o fold) (union_digest o #1) dom arity_digest end, type_variable = fn T => map (pair ([], [])) (Type.sort_of_atyp T)} (typ, sort); in fun solve_constraints (thm as Thm (_, {constraints = [], ...})) = thm | solve_constraints (thm as Thm (der, args)) = let val {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop} = args; - val thy = Context.certificate_theory cert; + val thy = Context.certificate_theory cert + handle ERROR msg => raise CONTEXT (msg, [], [], [thm], NONE); val bad_thys = constraints |> map_filter (fn {theory = thy', ...} => if Context.eq_thy (thy, thy') then NONE else SOME thy'); val () = if null bad_thys then () else raise THEORY ("solve_constraints: bad theories for theorem\n" ^ Syntax.string_of_term_global thy (prop_of thm), thy :: bad_thys); val Deriv {promises, body = PBody {oracles, thms, proof}} = der; val (oracles', thms') = (oracles, thms) |> fold (fold union_digest o constraint_digest) constraints; val body' = PBody {oracles = oracles', thms = thms', proof = proof}; in Thm (Deriv {promises = promises, body = body'}, {constraints = [], cert = cert, tags = tags, maxidx = maxidx, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; end; (*Dangling sort constraints of a thm*) fun extra_shyps (th as Thm (_, {shyps, ...})) = Sorts.subtract (fold_terms {hyps = true} Sorts.insert_term th []) shyps; (*Remove extra sorts that are witnessed by type signature information*) fun strip_shyps thm = (case thm of Thm (_, {shyps = [], ...}) => thm | Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) => let val thy = theory_of_thm thm; val algebra = Sign.classes_of thy; val minimize = Sorts.minimize_sort algebra; val le = Sorts.sort_le algebra; fun lt (S1, S2) = le (S1, S2) andalso not (le (S2, S1)); fun rel (S1, S2) = if S1 = S2 then [] else [(Term.aT S1, S2)]; val present = (fold_terms {hyps = true} o fold_types o fold_atyps_sorts) (insert (eq_fst op =)) thm []; val extra = fold (Sorts.remove_sort o #2) present shyps; val witnessed = Sign.witness_sorts thy present extra; val non_witnessed = fold (Sorts.remove_sort o #2) witnessed extra |> map (`minimize); val extra' = non_witnessed |> map_filter (fn (S, _) => if non_witnessed |> exists (fn (S', _) => lt (S', S)) then NONE else SOME S) |> Sorts.make; val constrs' = non_witnessed |> maps (fn (S1, S2) => let val S0 = the (find_first (fn S => le (S, S1)) extra') in rel (S0, S1) @ rel (S1, S2) end); val constraints' = fold (insert_constraints thy) (witnessed @ constrs') constraints; val shyps' = fold (Sorts.insert_sort o #2) present extra'; in Thm (deriv_rule_unconditional (Proofterm.strip_shyps_proof algebra present witnessed extra') der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints', shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop}) end) |> solve_constraints; (*** Closed theorems with official name ***) (*non-deterministic, depends on unknown promises*) fun derivation_closed (Thm (Deriv {body, ...}, _)) = Proofterm.compact_proof (Proofterm.proof_of body); (*non-deterministic, depends on unknown promises*) fun raw_derivation_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) = Proofterm.get_approximative_name shyps hyps prop (Proofterm.proof_of body); fun expand_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) = let val self_id = (case Proofterm.get_identity shyps hyps prop (Proofterm.proof_of body) of NONE => K false | SOME {serial, ...} => fn (header: Proofterm.thm_header) => serial = #serial header); fun expand header = if self_id header orelse #name header = "" then SOME "" else NONE; in expand end; (*deterministic name of finished proof*) fun derivation_name (thm as Thm (_, {shyps, hyps, prop, ...})) = Proofterm.get_approximative_name shyps hyps prop (proof_of thm); (*identified PThm node*) fun derivation_id (thm as Thm (_, {shyps, hyps, prop, ...})) = Proofterm.get_id shyps hyps prop (proof_of thm); (*dependencies of PThm node*) fun thm_deps (thm as Thm (Deriv {promises = [], body = PBody {thms, ...}, ...}, _)) = (case (derivation_id thm, thms) of (SOME {serial = i, ...}, [(j, thm_node)]) => if i = j then Proofterm.thm_node_thms thm_node else thms | _ => thms) | thm_deps thm = raise THM ("thm_deps: bad promises", 0, [thm]); fun name_derivation name_pos = strip_shyps #> (fn thm as Thm (der, args) => let val thy = theory_of_thm thm; val Deriv {promises, body} = der; val {shyps, hyps, prop, tpairs, ...} = args; val _ = null tpairs orelse raise THM ("name_derivation: bad flex-flex constraints", 0, [thm]); val ps = map (apsnd (Future.map fulfill_body)) promises; val (pthm, proof) = Proofterm.thm_proof thy (classrel_proof thy) (arity_proof thy) name_pos shyps hyps prop ps body; val der' = make_deriv [] [] [pthm] proof; in Thm (der', args) end); fun close_derivation pos = solve_constraints #> (fn thm => if not (null (tpairs_of thm)) orelse derivation_closed thm then thm else name_derivation ("", pos) thm); val trim_context = solve_constraints #> trim_context_thm; (*** Oracles ***) fun add_oracle (b, oracle_fn) thy = let val (name, oracles') = Name_Space.define (Context.Theory thy) true (b, ()) (get_oracles thy); val thy' = map_oracles (K oracles') thy; fun invoke_oracle arg = let val Cterm {cert = cert2, t = prop, T, maxidx, sorts} = oracle_fn arg in if T <> propT then raise THM ("Oracle's result must have type prop: " ^ name, 0, []) else let val (oracle, prf) = (case ! Proofterm.proofs of 2 => (((name, Position.thread_data ()), SOME prop), Proofterm.oracle_proof name prop) | 1 => (((name, Position.thread_data ()), SOME prop), MinProof) | 0 => (((name, Position.none), NONE), MinProof) | i => bad_proofs i); in Thm (make_deriv [] [oracle] [] prf, {cert = Context.join_certificate (Context.Certificate thy', cert2), tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) end end; in ((name, invoke_oracle), thy') end; val oracle_space = Name_Space.space_of_table o get_oracles; fun pretty_oracle ctxt = Name_Space.pretty ctxt (oracle_space (Proof_Context.theory_of ctxt)); fun extern_oracles verbose ctxt = map #1 (Name_Space.markup_table verbose ctxt (get_oracles (Proof_Context.theory_of ctxt))); fun check_oracle ctxt = Name_Space.check (Context.Proof ctxt) (get_oracles (Proof_Context.theory_of ctxt)) #> #1; (*** Meta rules ***) (** primitive rules **) (*The assumption rule A |- A*) fun assume raw_ct = let val Cterm {cert, t = prop, T, maxidx, sorts} = adjust_maxidx_cterm ~1 raw_ct in if T <> propT then raise THM ("assume: prop", 0, []) else if maxidx <> ~1 then raise THM ("assume: variables", maxidx, []) else Thm (deriv_rule0 (fn () => Proofterm.Hyp prop), {cert = cert, tags = [], maxidx = ~1, constraints = [], shyps = sorts, hyps = [prop], tpairs = [], prop = prop}) end; (*Implication introduction [A] : B ------- A \ B *) fun implies_intr (ct as Cterm {t = A, T, maxidx = maxidx1, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...})) = if T <> propT then raise THM ("implies_intr: assumptions must have type prop", 0, [th]) else Thm (deriv_rule1 (Proofterm.implies_intr_proof A) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = remove_hyps A hyps, tpairs = tpairs, prop = Logic.mk_implies (A, prop)}); (*Implication elimination A \ B A ------------ B *) fun implies_elim thAB thA = let val Thm (derA, {maxidx = maxidx1, hyps = hypsA, constraints = constraintsA, shyps = shypsA, tpairs = tpairsA, prop = propA, ...}) = thA and Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...}) = thAB; fun err () = raise THM ("implies_elim: major premise", 0, [thAB, thA]); in (case prop of Const ("Pure.imp", _) $ A $ B => if A aconv propA then Thm (deriv_rule2 (curry Proofterm.%%) der derA, {cert = join_certificate2 (thAB, thA), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraintsA constraints, shyps = Sorts.union shypsA shyps, hyps = union_hyps hypsA hyps, tpairs = union_tpairs tpairsA tpairs, prop = B}) else err () | _ => err ()) end; (*Forall introduction. The Free or Var x must not be free in the hypotheses. [x] : A ------ \x. A *) fun occs x ts tpairs = let fun occs t = Logic.occs (x, t) in exists occs ts orelse exists (occs o fst) tpairs orelse exists (occs o snd) tpairs end; fun forall_intr (ct as Cterm {maxidx = maxidx1, t = x, T, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) = let fun check_result a ts = if occs x ts tpairs then raise THM ("forall_intr: variable " ^ quote a ^ " free in assumptions", 0, [th]) else Thm (deriv_rule1 (Proofterm.forall_intr_proof (a, x) NONE) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Logic.all_const T $ Abs (a, T, abstract_over (x, prop))}); in (case x of Free (a, _) => check_result a hyps | Var ((a, _), _) => check_result a [] | _ => raise THM ("forall_intr: not a variable", 0, [th])) end; (*Forall elimination \x. A ------ A[t/x] *) fun forall_elim (ct as Cterm {t, T, maxidx = maxidx1, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) = (case prop of Const ("Pure.all", Type ("fun", [Type ("fun", [qary, _]), _])) $ A => if T <> qary then raise THM ("forall_elim: type mismatch", 0, [th]) else Thm (deriv_rule1 (Proofterm.% o rpair (SOME t)) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Term.betapply (A, t)}) | _ => raise THM ("forall_elim: not quantified", 0, [th])); (* Equality *) (*Reflexivity t \ t *) fun reflexive (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, t)}); (*Symmetry t \ u ------ u \ t *) fun symmetric (th as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = (case prop of (eq as Const ("Pure.eq", _)) $ t $ u => Thm (deriv_rule1 Proofterm.symmetric_proof der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = eq $ u $ t}) | _ => raise THM ("symmetric", 0, [th])); (*Transitivity t1 \ u u \ t2 ------------------ t1 \ t2 *) fun transitive th1 th2 = let val Thm (der1, {maxidx = maxidx1, hyps = hyps1, constraints = constraints1, shyps = shyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, hyps = hyps2, constraints = constraints2, shyps = shyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("transitive: " ^ msg, 0, [th1, th2]); in case (prop1, prop2) of ((eq as Const ("Pure.eq", Type (_, [U, _]))) $ t1 $ u, Const ("Pure.eq", _) $ u' $ t2) => if not (u aconv u') then err "middle term" else Thm (deriv_rule2 (Proofterm.transitive_proof U u) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = eq $ t1 $ t2}) | _ => err "premises" end; (*Beta-conversion (\x. t) u \ t[u/x] fully beta-reduces the term if full = true *) fun beta_conversion full (Cterm {cert, t, T = _, maxidx, sorts}) = let val t' = if full then Envir.beta_norm t else (case t of Abs (_, _, bodt) $ u => subst_bound (u, bodt) | _ => raise THM ("beta_conversion: not a redex", 0, [])); in Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, t')}) end; fun eta_conversion (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, Envir.eta_contract t)}); fun eta_long_conversion (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, Envir.eta_long [] t)}); (*The abstraction rule. The Free or Var x must not be free in the hypotheses. The bound variable will be named "a" (since x will be something like x320) t \ u -------------- \x. t \ \x. u *) fun abstract_rule b (Cterm {t = x, T, sorts, ...}) (th as Thm (der, {cert, maxidx, hyps, constraints, shyps, tpairs, prop, ...})) = let val (t, u) = Logic.dest_equals prop handle TERM _ => raise THM ("abstract_rule: premise not an equality", 0, [th]); fun check_result a ts = if occs x ts tpairs then raise THM ("abstract_rule: variable " ^ quote a ^ " free in assumptions", 0, [th]) else Thm (deriv_rule1 (Proofterm.abstract_rule_proof (b, x)) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Logic.mk_equals (Abs (b, T, abstract_over (x, t)), Abs (b, T, abstract_over (x, u)))}); in (case x of Free (a, _) => check_result a hyps | Var ((a, _), _) => check_result a [] | _ => raise THM ("abstract_rule: not a variable", 0, [th])) end; (*The combination rule f \ g t \ u ------------- f t \ g u *) fun combination th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun chktypes fT tT = (case fT of Type ("fun", [T1, _]) => if T1 <> tT then raise THM ("combination: types", 0, [th1, th2]) else () | _ => raise THM ("combination: not function type", 0, [th1, th2])); in (case (prop1, prop2) of (Const ("Pure.eq", Type ("fun", [fT, _])) $ f $ g, Const ("Pure.eq", Type ("fun", [tT, _])) $ t $ u) => (chktypes fT tT; Thm (deriv_rule2 (Proofterm.combination_proof f g t u) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = Logic.mk_equals (f $ t, g $ u)})) | _ => raise THM ("combination: premises", 0, [th1, th2])) end; (*Equality introduction A \ B B \ A ---------------- A \ B *) fun equal_intr th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("equal_intr: " ^ msg, 0, [th1, th2]); in (case (prop1, prop2) of (Const("Pure.imp", _) $ A $ B, Const("Pure.imp", _) $ B' $ A') => if A aconv A' andalso B aconv B' then Thm (deriv_rule2 (Proofterm.equal_intr_proof A B) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = Logic.mk_equals (A, B)}) else err "not equal" | _ => err "premises") end; (*The equal propositions rule A \ B A --------- B *) fun equal_elim th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("equal_elim: " ^ msg, 0, [th1, th2]); in (case prop1 of Const ("Pure.eq", _) $ A $ B => if prop2 aconv A then Thm (deriv_rule2 (Proofterm.equal_elim_proof A B) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = B}) else err "not equal" | _ => err "major premise") end; (**** Derived rules ****) (*Smash unifies the list of term pairs leaving no flex-flex pairs. Instantiates the theorem and deletes trivial tpairs. Resulting sequence may contain multiple elements if the tpairs are not all flex-flex.*) fun flexflex_rule opt_ctxt = solve_constraints #> (fn th => let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val (context, cert') = make_context_certificate [th] opt_ctxt cert; in Unify.smash_unifiers context tpairs (Envir.empty maxidx) |> Seq.map (fn env => if Envir.is_empty env then th else let val tpairs' = tpairs |> map (apply2 (Envir.norm_term env)) (*remove trivial tpairs, of the form t \ t*) |> filter_out (op aconv); val der' = deriv_rule1 (Proofterm.norm_proof' env) der; - val constraints' = - insert_constraints_env (Context.certificate_theory cert') env constraints; + val thy' = Context.certificate_theory cert' handle ERROR msg => + raise CONTEXT (msg, [], [], [th], Option.map Context.Proof opt_ctxt); + val constraints' = insert_constraints_env thy' env constraints; val prop' = Envir.norm_term env prop; val maxidx = maxidx_tpairs tpairs' (maxidx_of_term prop'); val shyps = Envir.insert_sorts env shyps; in Thm (der', {cert = cert', tags = [], maxidx = maxidx, constraints = constraints', shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end) end); (*Generalization of fixed variables A -------------------- A[?'a/'a, ?x/x, ...] *) fun generalize (tfrees, frees) idx th = if Names.is_empty tfrees andalso Names.is_empty frees then th else let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val _ = idx <= maxidx andalso raise THM ("generalize: bad index", idx, [th]); val bad_type = if Names.is_empty tfrees then K false else Term.exists_subtype (fn TFree (a, _) => Names.defined tfrees a | _ => false); fun bad_term (Free (x, T)) = bad_type T orelse Names.defined frees x | bad_term (Var (_, T)) = bad_type T | bad_term (Const (_, T)) = bad_type T | bad_term (Abs (_, T, t)) = bad_type T orelse bad_term t | bad_term (t $ u) = bad_term t orelse bad_term u | bad_term (Bound _) = false; val _ = exists bad_term hyps andalso raise THM ("generalize: variable free in assumptions", 0, [th]); val generalize = Term_Subst.generalize (tfrees, frees) idx; val prop' = generalize prop; val tpairs' = map (apply2 generalize) tpairs; val maxidx' = maxidx_tpairs tpairs' (maxidx_of_term prop'); in Thm (deriv_rule1 (Proofterm.generalize_proof (tfrees, frees) idx prop) der, {cert = cert, tags = [], maxidx = maxidx', constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end; fun generalize_cterm (tfrees, frees) idx (ct as Cterm {cert, t, T, maxidx, sorts}) = if Names.is_empty tfrees andalso Names.is_empty frees then ct else if idx <= maxidx then raise CTERM ("generalize_cterm: bad index", [ct]) else Cterm {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, t = Term_Subst.generalize (tfrees, frees) idx t, maxidx = Int.max (maxidx, idx)}; fun generalize_ctyp tfrees idx (cT as Ctyp {cert, T, maxidx, sorts}) = if Names.is_empty tfrees then cT else if idx <= maxidx then raise CTERM ("generalize_ctyp: bad index", []) else Ctyp {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, maxidx = Int.max (maxidx, idx)}; (*Instantiation of schematic variables A -------------------- A[t1/v1, ..., tn/vn] *) local fun add_cert cert_of (_, c) cert = Context.join_certificate (cert, cert_of c); val add_instT_cert = add_cert (fn Ctyp {cert, ...} => cert); val add_inst_cert = add_cert (fn Cterm {cert, ...} => cert); fun add_sorts sorts_of (_, c) sorts = Sorts.union (sorts_of c) sorts; val add_instT_sorts = add_sorts (fn Ctyp {sorts, ...} => sorts); val add_inst_sorts = add_sorts (fn Cterm {sorts, ...} => sorts); fun make_instT thy (_: indexname, S) (Ctyp {T = U, maxidx, ...}) = if Sign.of_sort thy (U, S) then (U, maxidx) else raise TYPE ("Type not of sort " ^ Syntax.string_of_sort_global thy S, [U], []); fun make_inst thy (v as (_, T)) (Cterm {t = u, T = U, maxidx, ...}) = if T = U then (u, maxidx) else let fun pretty_typing t ty = Pretty.block [Syntax.pretty_term_global thy t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ_global thy ty]; val msg = Pretty.string_of (Pretty.block [Pretty.str "instantiate: type conflict", Pretty.fbrk, pretty_typing (Var v) T, Pretty.fbrk, pretty_typing u U]) in raise TYPE (msg, [T, U], [Var v, u]) end; fun prep_insts (instT, inst) (cert, sorts) = let val cert' = cert |> TVars.fold add_instT_cert instT |> Vars.fold add_inst_cert inst; val thy' = Context.certificate_theory cert' handle ERROR msg => raise CONTEXT (msg, TVars.dest instT |> map #2, Vars.dest inst |> map #2, [], NONE); val sorts' = sorts |> TVars.fold add_instT_sorts instT |> Vars.fold add_inst_sorts inst; val instT' = TVars.map (make_instT thy') instT; val inst' = Vars.map (make_inst thy') inst; in ((instT', inst'), (cert', sorts')) end; in (*Left-to-right replacements: ctpairs = [..., (vi, ti), ...]. Instantiates distinct Vars by terms of same type. Does NOT normalize the resulting theorem!*) fun gen_instantiate inst_fn (instT, inst) th = if TVars.is_empty instT andalso Vars.is_empty inst then th else let val Thm (der, {cert, hyps, constraints, shyps, tpairs, prop, ...}) = th; val ((instT', inst'), (cert', shyps')) = prep_insts (instT, inst) (cert, shyps) handle CONTEXT (msg, cTs, cts, ths, context) => raise CONTEXT (msg, cTs, cts, th :: ths, context); val subst = inst_fn (instT', inst'); val (prop', maxidx1) = subst prop ~1; val (tpairs', maxidx') = fold_map (fn (t, u) => fn i => subst t i ||>> subst u) tpairs maxidx1; - val thy' = Context.certificate_theory cert'; + val thy' = Context.certificate_theory cert' + handle ERROR msg => raise CONTEXT (msg, [], [], [th], NONE); val constraints' = TVars.fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S)) instT' constraints; in Thm (deriv_rule1 (fn d => Proofterm.instantiate (TVars.map (K #1) instT', Vars.map (K #1) inst') d) der, {cert = cert', tags = [], maxidx = maxidx', constraints = constraints', shyps = shyps', hyps = hyps, tpairs = tpairs', prop = prop'}) |> solve_constraints end handle TYPE (msg, _, _) => raise THM (msg, 0, [th]); val instantiate = gen_instantiate Term_Subst.instantiate_maxidx; val instantiate_beta = gen_instantiate Term_Subst.instantiate_beta_maxidx; fun gen_instantiate_cterm inst_fn (instT, inst) ct = if TVars.is_empty instT andalso Vars.is_empty inst then ct else let val Cterm {cert, t, T, sorts, ...} = ct; val ((instT', inst'), (cert', sorts')) = prep_insts (instT, inst) (cert, sorts); val subst = inst_fn (instT', inst'); val substT = Term_Subst.instantiateT_maxidx instT'; val (t', maxidx1) = subst t ~1; val (T', maxidx') = substT T maxidx1; in Cterm {cert = cert', t = t', T = T', sorts = sorts', maxidx = maxidx'} end handle TYPE (msg, _, _) => raise CTERM (msg, [ct]); val instantiate_cterm = gen_instantiate_cterm Term_Subst.instantiate_maxidx; val instantiate_beta_cterm = gen_instantiate_cterm Term_Subst.instantiate_beta_maxidx; end; (*The trivial implication A \ A, justified by assume and forall rules. A can contain Vars, not so for assume!*) fun trivial (Cterm {cert, t = A, T, maxidx, sorts}) = if T <> propT then raise THM ("trivial: the term must have type prop", 0, []) else Thm (deriv_rule0 (fn () => Proofterm.trivial_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_implies (A, A)}); (*Axiom-scheme reflecting signature contents T :: c ------------------- OFCLASS(T, c_class) *) fun of_class (cT, raw_c) = let val Ctyp {cert, T, ...} = cT; val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [cT], [], [], NONE); val c = Sign.certify_class thy raw_c; val Cterm {t = prop, maxidx, sorts, ...} = global_cterm_of thy (Logic.mk_of_class (T, c)); in if Sign.of_sort thy (T, [c]) then Thm (deriv_rule0 (fn () => Proofterm.PClass (T, c)), {cert = cert, tags = [], maxidx = maxidx, constraints = insert_constraints thy (T, [c]) [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) else raise THM ("of_class: type not of class " ^ Syntax.string_of_sort_global thy [c], 0, []) end |> solve_constraints; (*Internalize sort constraints of type variables*) val unconstrainT = strip_shyps #> (fn thm as Thm (der, args) => let val Deriv {promises, body} = der; val {cert, shyps, hyps, tpairs, prop, ...} = args; val thy = theory_of_thm thm; fun err msg = raise THM ("unconstrainT: " ^ msg, 0, [thm]); val _ = null hyps orelse err "bad hyps"; val _ = null tpairs orelse err "bad flex-flex constraints"; val tfrees = build_rev (Term.add_tfree_names prop); val _ = null tfrees orelse err ("illegal free type variables " ^ commas_quote tfrees); val ps = map (apsnd (Future.map fulfill_body)) promises; val (pthm, proof) = Proofterm.unconstrain_thm_proof thy (classrel_proof thy) (arity_proof thy) shyps prop ps body; val der' = make_deriv [] [] [pthm] proof; val prop' = Proofterm.thm_node_prop (#2 pthm); in Thm (der', {cert = cert, tags = [], maxidx = maxidx_of_term prop', constraints = [], shyps = [[]], (*potentially redundant*) hyps = [], tpairs = [], prop = prop'}) end); (*Replace all TFrees not fixed or in the hyps by new TVars*) fun varifyT_global' fixed (Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = let val tfrees = fold TFrees.add_tfrees hyps fixed; val prop1 = attach_tpairs tpairs prop; val (al, prop2) = Type.varify_global tfrees prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in (al, Thm (deriv_rule1 (Proofterm.varify_proof prop tfrees) der, {cert = cert, tags = [], maxidx = Int.max (0, maxidx), constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3})) end; val varifyT_global = #2 o varifyT_global' TFrees.empty; (*Replace all TVars by TFrees that are often new*) fun legacy_freezeT (Thm (der, {cert, constraints, shyps, hyps, tpairs, prop, ...})) = let val prop1 = attach_tpairs tpairs prop; val prop2 = Type.legacy_freeze prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in Thm (deriv_rule1 (Proofterm.legacy_freezeT prop1) der, {cert = cert, tags = [], maxidx = maxidx_of_term prop2, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3}) end; fun plain_prop_of raw_thm = let val thm = strip_shyps raw_thm; fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]); in if not (null (hyps_of thm)) then err "theorem may not contain hypotheses" else if not (null (extra_shyps thm)) then err "theorem may not contain sort hypotheses" else if not (null (tpairs_of thm)) then err "theorem may not contain flex-flex pairs" else prop_of thm end; (*** Inference rules for tactics ***) (*Destruct proof state into constraints, other goals, goal(i), rest *) fun dest_state (state as Thm (_, {prop,tpairs,...}), i) = (case Logic.strip_prems(i, [], prop) of (B::rBs, C) => (tpairs, rev rBs, B, C) | _ => raise THM("dest_state", i, [state])) handle TERM _ => raise THM("dest_state", i, [state]); (*Prepare orule for resolution by lifting it over the parameters and assumptions of goal.*) fun lift_rule goal orule = let val Cterm {t = gprop, T, maxidx = gmax, sorts, ...} = goal; val inc = gmax + 1; val lift_abs = Logic.lift_abs inc gprop; val lift_all = Logic.lift_all inc gprop; val Thm (der, {maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = orule; val (As, B) = Logic.strip_horn prop; in if T <> propT then raise THM ("lift_rule: the term must have type prop", 0, []) else Thm (deriv_rule1 (Proofterm.lift_proof gprop inc prop) der, {cert = join_certificate1 (goal, orule), tags = [], maxidx = maxidx + inc, constraints = constraints, shyps = Sorts.union shyps sorts, (*sic!*) hyps = hyps, tpairs = map (apply2 lift_abs) tpairs, prop = Logic.list_implies (map lift_all As, lift_all B)}) end; fun incr_indexes i (thm as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = if i < 0 then raise THM ("negative increment", 0, [thm]) else if i = 0 then thm else Thm (deriv_rule1 (Proofterm.incr_indexes i) der, {cert = cert, tags = [], maxidx = maxidx + i, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = map (apply2 (Logic.incr_indexes ([], [], i))) tpairs, prop = Logic.incr_indexes ([], [], i) prop}); (*Solve subgoal Bi of proof state B1...Bn/C by assumption. *) fun assumption opt_ctxt i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (context, cert') = make_context_certificate [state] opt_ctxt cert; val (tpairs, Bs, Bi, C) = dest_state (state, i); fun newth n (env, tpairs) = let val normt = Envir.norm_term env; fun assumption_proof prf = Proofterm.assumption_proof (map normt Bs) (normt Bi) n prf; + val thy' = Context.certificate_theory cert' handle ERROR msg => + raise CONTEXT (msg, [], [], [state], Option.map Context.Proof opt_ctxt); in Thm (deriv_rule1 (assumption_proof #> not (Envir.is_empty env) ? Proofterm.norm_proof' env) der, {tags = [], maxidx = Envir.maxidx_of env, - constraints = insert_constraints_env (Context.certificate_theory cert') env constraints, + constraints = insert_constraints_env thy' env constraints, shyps = Envir.insert_sorts env shyps, hyps = hyps, tpairs = if Envir.is_empty env then tpairs else map (apply2 normt) tpairs, prop = if Envir.is_empty env then Logic.list_implies (Bs, C) (*avoid wasted normalizations*) else normt (Logic.list_implies (Bs, C)) (*normalize the new rule fully*), cert = cert'}) end; val (close, asms, concl) = Logic.assum_problems (~1, Bi); val concl' = close concl; fun addprfs [] _ = Seq.empty | addprfs (asm :: rest) n = Seq.make (fn () => Seq.pull (Seq.mapp (newth n) (if Term.could_unify (asm, concl) then (Unify.unifiers (context, Envir.empty maxidx, (close asm, concl') :: tpairs)) else Seq.empty) (addprfs rest (n + 1)))) in addprfs asms 1 end; (*Solve subgoal Bi of proof state B1...Bn/C by assumption. Checks if Bi's conclusion is alpha/eta-convertible to one of its assumptions*) fun eq_assumption i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val (_, asms, concl) = Logic.assum_problems (~1, Bi); in (case find_index (fn asm => Envir.aeconv (asm, concl)) asms of ~1 => raise THM ("eq_assumption", 0, [state]) | n => Thm (deriv_rule1 (Proofterm.assumption_proof Bs Bi (n + 1)) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs, C)})) end; (*For rotate_tac: fast rotation of assumptions of subgoal i*) fun rotate_rule k i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val params = Term.strip_all_vars Bi; val rest = Term.strip_all_body Bi; val asms = Logic.strip_imp_prems rest val concl = Logic.strip_imp_concl rest; val n = length asms; val m = if k < 0 then n + k else k; val Bi' = if 0 = m orelse m = n then Bi else if 0 < m andalso m < n then let val (ps, qs) = chop m asms in Logic.list_all (params, Logic.list_implies (qs @ ps, concl)) end else raise THM ("rotate_rule", k, [state]); in Thm (deriv_rule1 (Proofterm.rotate_proof Bs Bi' params asms m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs @ [Bi'], C)}) end; (*Rotates a rule's premises to the left by k, leaving the first j premises unchanged. Does nothing if k=0 or if k equals n-j, where n is the number of premises. Useful with eresolve_tac and underlies defer_tac*) fun permute_prems j k rl = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = rl; val prems = Logic.strip_imp_prems prop and concl = Logic.strip_imp_concl prop; val moved_prems = List.drop (prems, j) and fixed_prems = List.take (prems, j) handle General.Subscript => raise THM ("permute_prems: j", j, [rl]); val n_j = length moved_prems; val m = if k < 0 then n_j + k else k; val (prems', prop') = if 0 = m orelse m = n_j then (prems, prop) else if 0 < m andalso m < n_j then let val (ps, qs) = chop m moved_prems; val prems' = fixed_prems @ qs @ ps; in (prems', Logic.list_implies (prems', concl)) end else raise THM ("permute_prems: k", k, [rl]); in Thm (deriv_rule1 (Proofterm.permute_prems_proof prems' j m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) end; (* strip_apply f B A strips off all assumptions/parameters from A introduced by lifting over B, and applies f to remaining part of A*) fun strip_apply f = let fun strip (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ A2 $ B2) = Logic.mk_implies (A2, strip B1 B2) | strip ((c as Const ("Pure.all", _)) $ Abs (_, _, t1)) ( Const ("Pure.all", _) $ Abs (a, T, t2)) = c $ Abs (a, T, strip t1 t2) | strip _ A = f A in strip end; fun strip_lifted (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ _ $ B2) = strip_lifted B1 B2 | strip_lifted (Const ("Pure.all", _) $ Abs (_, _, t1)) (Const ("Pure.all", _) $ Abs (_, _, t2)) = strip_lifted t1 t2 | strip_lifted _ A = A; (*Use the alist to rename all bound variables and some unknowns in a term dpairs = current disagreement pairs; tpairs = permanent ones (flexflex); Preserves unknowns in tpairs and on lhs of dpairs. *) fun rename_bvs [] _ _ _ _ = K I | rename_bvs al dpairs tpairs B As = let val add_var = fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I); val vids = [] |> fold (add_var o fst) dpairs |> fold (add_var o fst) tpairs |> fold (add_var o snd) tpairs; val vids' = fold (add_var o strip_lifted B) As []; (*unknowns appearing elsewhere be preserved!*) val al' = distinct ((op =) o apply2 fst) (filter_out (fn (x, y) => not (member (op =) vids' x) orelse member (op =) vids x orelse member (op =) vids y) al); val unchanged = filter_out (AList.defined (op =) al') vids'; fun del_clashing clash xs _ [] qs = if clash then del_clashing false xs xs qs [] else qs | del_clashing clash xs ys ((p as (x, y)) :: ps) qs = if member (op =) ys y then del_clashing true (x :: xs) (x :: ys) ps qs else del_clashing clash xs (y :: ys) ps (p :: qs); val al'' = del_clashing false unchanged unchanged al' []; fun rename (t as Var ((x, i), T)) = (case AList.lookup (op =) al'' x of SOME y => Var ((y, i), T) | NONE => t) | rename (Abs (x, T, t)) = Abs (the_default x (AList.lookup (op =) al x), T, rename t) | rename (f $ t) = rename f $ rename t | rename t = t; fun strip_ren f Ai = f rename B Ai in strip_ren end; (*Function to rename bounds/unknowns in the argument, lifted over B*) fun rename_bvars dpairs = rename_bvs (fold_rev Term.match_bvars dpairs []) dpairs; (*** RESOLUTION ***) (** Lifting optimizations **) (*strip off pairs of assumptions/parameters in parallel -- they are identical because of lifting*) fun strip_assums2 (Const("Pure.imp", _) $ _ $ B1, Const("Pure.imp", _) $ _ $ B2) = strip_assums2 (B1,B2) | strip_assums2 (Const("Pure.all",_)$Abs(a,T,t1), Const("Pure.all",_)$Abs(_,_,t2)) = let val (B1,B2) = strip_assums2 (t1,t2) in (Abs(a,T,B1), Abs(a,T,B2)) end | strip_assums2 BB = BB; (*Faster normalization: skip assumptions that were lifted over*) fun norm_term_skip env 0 t = Envir.norm_term env t | norm_term_skip env n (Const ("Pure.all", _) $ Abs (a, T, t)) = let val T' = Envir.norm_type (Envir.type_env env) T (*Must instantiate types of parameters because they are flattened; this could be a NEW parameter*) in Logic.all_const T' $ Abs (a, T', norm_term_skip env n t) end | norm_term_skip env n (Const ("Pure.imp", _) $ A $ B) = Logic.mk_implies (A, norm_term_skip env (n - 1) B) | norm_term_skip _ _ _ = error "norm_term_skip: too few assumptions??"; (*unify types of schematic variables (non-lifted case)*) fun unify_var_types context (th1, th2) env = let fun unify_vars (T :: Us) = fold (fn U => Pattern.unify_types context (T, U)) Us | unify_vars _ = I; val add_vars = full_prop_of #> fold_aterms (fn Var v => Vartab.insert_list (op =) v | _ => I); val vars = Vartab.build (add_vars th1 #> add_vars th2); in SOME (Vartab.fold (unify_vars o #2) vars env) end handle Pattern.Unif => NONE; (*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C) Unifies B with Bi, replacing subgoal i (1 <= i <= n) If match then forbid instantiations in proof state If lifted then shorten the dpair using strip_assums2. If eres_flg then simultaneously proves A1 by assumption. nsubgoal is the number of new subgoals (written m above). Curried so that resolution calls dest_state only once. *) local exception COMPOSE in fun bicompose_aux opt_ctxt {flatten, match, incremented} (state, (stpairs, Bs, Bi, C), lifted) (eres_flg, orule, nsubgoal) = let val Thm (sder, {maxidx=smax, constraints = constraints2, shyps = shyps2, hyps = hyps2, ...}) = state and Thm (rder, {maxidx=rmax, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs=rtpairs, prop=rprop,...}) = orule (*How many hyps to skip over during normalization*) and nlift = Logic.count_prems (strip_all_body Bi) + (if eres_flg then ~1 else 0) val (context, cert) = make_context_certificate [state, orule] opt_ctxt (join_certificate2 (state, orule)); (*Add new theorem with prop = "\Bs; As\ \ C" to thq*) fun addth A (As, oldAs, rder', n) ((env, tpairs), thq) = let val normt = Envir.norm_term env; (*perform minimal copying here by examining env*) val (ntpairs, normp) = if Envir.is_empty env then (tpairs, (Bs @ As, C)) else let val ntps = map (apply2 normt) tpairs in if Envir.above env smax then (*no assignments in state; normalize the rule only*) if lifted then (ntps, (Bs @ map (norm_term_skip env nlift) As, C)) else (ntps, (Bs @ map normt As, C)) else if match then raise COMPOSE else (*normalize the new rule fully*) (ntps, (map normt (Bs @ As), normt C)) end + val thy' = Context.certificate_theory cert handle ERROR msg => + raise CONTEXT (msg, [], [], [state, orule], Option.map Context.Proof opt_ctxt); val constraints' = union_constraints constraints1 constraints2 - |> insert_constraints_env (Context.certificate_theory cert) env; + |> insert_constraints_env thy' env; fun bicompose_proof prf1 prf2 = Proofterm.bicompose_proof flatten (map normt Bs) (map normt As) A oldAs n (nlift+1) prf1 prf2 val th = Thm (deriv_rule2 (if Envir.is_empty env then bicompose_proof else if Envir.above env smax then bicompose_proof o Proofterm.norm_proof' env else Proofterm.norm_proof' env oo bicompose_proof) rder' sder, {tags = [], maxidx = Envir.maxidx_of env, constraints = constraints', shyps = Envir.insert_sorts env (Sorts.union shyps1 shyps2), hyps = union_hyps hyps1 hyps2, tpairs = ntpairs, prop = Logic.list_implies normp, cert = cert}) in Seq.cons th thq end handle COMPOSE => thq; val (rAs,B) = Logic.strip_prems(nsubgoal, [], rprop) handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]); (*Modify assumptions, deleting n-th if n>0 for e-resolution*) fun newAs(As0, n, dpairs, tpairs) = let val (As1, rder') = if not lifted then (As0, rder) else let val rename = rename_bvars dpairs tpairs B As0 in (map (rename strip_apply) As0, deriv_rule1 (Proofterm.map_proof_terms (rename K) I) rder) end; in (map (if flatten then (Logic.flatten_params n) else I) As1, As1, rder', n) handle TERM _ => raise THM("bicompose: 1st premise", 0, [orule]) end; val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi); val dpairs = BBi :: (rtpairs@stpairs); (*elim-resolution: try each assumption in turn*) fun eres _ [] = raise THM ("bicompose: no premises", 0, [orule, state]) | eres env (A1 :: As) = let val A = SOME A1; val (close, asms, concl) = Logic.assum_problems (nlift + 1, A1); val concl' = close concl; fun tryasms [] _ = Seq.empty | tryasms (asm :: rest) n = if Term.could_unify (asm, concl) then let val asm' = close asm in (case Seq.pull (Unify.unifiers (context, env, (asm', concl') :: dpairs)) of NONE => tryasms rest (n + 1) | cell as SOME ((_, tpairs), _) => Seq.it_right (addth A (newAs (As, n, [BBi, (concl', asm')], tpairs))) (Seq.make (fn () => cell), Seq.make (fn () => Seq.pull (tryasms rest (n + 1))))) end else tryasms rest (n + 1); in tryasms asms 1 end; (*ordinary resolution*) fun res env = (case Seq.pull (Unify.unifiers (context, env, dpairs)) of NONE => Seq.empty | cell as SOME ((_, tpairs), _) => Seq.it_right (addth NONE (newAs (rev rAs, 0, [BBi], tpairs))) (Seq.make (fn () => cell), Seq.empty)); val env0 = Envir.empty (Int.max (rmax, smax)); in (case if incremented then SOME env0 else unify_var_types context (state, orule) env0 of NONE => Seq.empty | SOME env => if eres_flg then eres env (rev rAs) else res env) end; end; fun bicompose opt_ctxt flags arg i state = bicompose_aux opt_ctxt flags (state, dest_state (state,i), false) arg; (*Quick test whether rule is resolvable with the subgoal with hyps Hs and conclusion B. If eres_flg then checks 1st premise of rule also*) fun could_bires (Hs, B, eres_flg, rule) = let fun could_reshyp (A1::_) = exists (fn H => Term.could_unify (A1, H)) Hs | could_reshyp [] = false; (*no premise -- illegal*) in Term.could_unify(concl_of rule, B) andalso (not eres_flg orelse could_reshyp (prems_of rule)) end; (*Bi-resolution of a state with a list of (flag,rule) pairs. Puts the rule above: rule/state. Renames vars in the rules. *) fun biresolution opt_ctxt match brules i state = let val (stpairs, Bs, Bi, C) = dest_state(state,i); val lift = lift_rule (cprem_of state i); val B = Logic.strip_assums_concl Bi; val Hs = Logic.strip_assums_hyp Bi; val compose = bicompose_aux opt_ctxt {flatten = true, match = match, incremented = true} (state, (stpairs, Bs, Bi, C), true); fun res [] = Seq.empty | res ((eres_flg, rule)::brules) = if Config.get_generic (make_context [state] opt_ctxt (cert_of state)) Pattern.unify_trace_failure orelse could_bires (Hs, B, eres_flg, rule) then Seq.make (*delay processing remainder till needed*) (fn()=> SOME(compose (eres_flg, lift rule, nprems_of rule), res brules)) else res brules in Seq.flat (res brules) end; (*Resolution: exactly one resolvent must be produced*) fun tha RSN (i, thb) = (case Seq.chop 2 (biresolution NONE false [(false, tha)] i thb) of ([th], _) => solve_constraints th | ([], _) => raise THM ("RSN: no unifiers", i, [tha, thb]) | _ => raise THM ("RSN: multiple unifiers", i, [tha, thb])); (*Resolution: P \ Q, Q \ R gives P \ R*) fun tha RS thb = tha RSN (1,thb); (**** Type classes ****) fun standard_tvars thm = let val thy = theory_of_thm thm; val tvars = build_rev (Term.add_tvars (prop_of thm)); val names = Name.invent Name.context Name.aT (length tvars); val tinst = map2 (fn (ai, S) => fn b => ((ai, S), global_ctyp_of thy (TVar ((b, 0), S)))) tvars names; in instantiate (TVars.make tinst, Vars.empty) thm end (* class relations *) val is_classrel = Symreltab.defined o get_classrels; fun complete_classrels thy = let fun complete (c, (_, (all_preds, all_succs))) (finished1, thy1) = let fun compl c1 c2 (finished2, thy2) = if is_classrel thy2 (c1, c2) then (finished2, thy2) else (false, thy2 |> (map_classrels o Symreltab.update) ((c1, c2), (the_classrel thy2 (c1, c) RS the_classrel thy2 (c, c2)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy2) |> trim_context)); val proven = is_classrel thy1; val preds = Graph.Keys.fold (fn c1 => proven (c1, c) ? cons c1) all_preds []; val succs = Graph.Keys.fold (fn c2 => proven (c, c2) ? cons c2) all_succs []; in fold_product compl preds succs (finished1, thy1) end; in (case Graph.fold complete (Sorts.classes_of (Sign.classes_of thy)) (true, thy) of (true, _) => NONE | (_, thy') => SOME thy') end; (* type arities *) fun theory_names_of_arity {long} thy (a, c) = build (get_arities thy |> Aritytab.fold (fn ((a', _, c'), (_, name, ser)) => (a = a' andalso c = c') ? cons (name, ser))) |> sort (int_ord o apply2 #2) |> map (if long then #1 else Long_Name.base_name o #1); fun insert_arity_completions thy ((t, Ss, c), (th, thy_name, ser)) (finished, arities) = let val completions = Sign.super_classes thy c |> map_filter (fn c1 => if Aritytab.defined arities (t, Ss, c1) then NONE else let val th1 = (th RS the_classrel thy (c, c1)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy) |> trim_context; in SOME ((t, Ss, c1), (th1, thy_name, ser)) end); val finished' = finished andalso null completions; val arities' = fold Aritytab.update completions arities; in (finished', arities') end; fun complete_arities thy = let val arities = get_arities thy; val (finished, arities') = Aritytab.fold (insert_arity_completions thy) arities (true, get_arities thy); in if finished then NONE else SOME (map_arities (K arities') thy) end; val _ = Theory.setup (Theory.at_begin complete_classrels #> Theory.at_begin complete_arities); (* primitive rules *) fun add_classrel raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (c1, c2) = Logic.dest_classrel prop; in thy |> Sign.primitive_classrel (c1, c2) |> map_classrels (Symreltab.update ((c1, c2), th')) |> perhaps complete_classrels |> perhaps complete_arities end; fun add_arity raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (t, Ss, c) = Logic.dest_arity prop; val ar = ((t, Ss, c), (th', Context.theory_long_name thy, serial ())); in thy |> Sign.primitive_arity (t, Ss, [c]) |> map_arities (Aritytab.update ar #> curry (insert_arity_completions thy ar) true #> #2) end; end; structure Basic_Thm: BASIC_THM = Thm; open Basic_Thm; diff --git a/src/Tools/Haskell/Haskell.thy b/src/Tools/Haskell/Haskell.thy --- a/src/Tools/Haskell/Haskell.thy +++ b/src/Tools/Haskell/Haskell.thy @@ -1,3909 +1,3919 @@ (* Title: Tools/Haskell/Haskell.thy Author: Makarius Support for Isabelle tools in Haskell. *) theory Haskell imports Main begin generate_file "Isabelle/Bytes.hs" = \ {- Title: Isabelle/Bytes.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Compact byte strings. See \<^file>\$ISABELLE_HOME/src/Pure/General/bytes.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/General/bytes.scala\. -} {-# LANGUAGE TypeApplications #-} module Isabelle.Bytes ( Bytes, make, unmake, pack, unpack, empty, null, length, index, all, any, head, last, take, drop, isPrefixOf, isSuffixOf, try_unprefix, try_unsuffix, concat, space, spaces, char, all_char, any_char, byte, singleton ) where import Prelude hiding (null, length, all, any, head, last, take, drop, concat) import qualified Data.ByteString.Short as ShortByteString import Data.ByteString.Short (ShortByteString) import qualified Data.ByteString as ByteString import Data.ByteString (ByteString) import qualified Data.List as List import Data.Word (Word8) import Data.Array (Array, array, (!)) type Bytes = ShortByteString make :: ByteString -> Bytes make = ShortByteString.toShort unmake :: Bytes -> ByteString unmake = ShortByteString.fromShort pack :: [Word8] -> Bytes pack = ShortByteString.pack unpack :: Bytes -> [Word8] unpack = ShortByteString.unpack empty :: Bytes empty = ShortByteString.empty null :: Bytes -> Bool null = ShortByteString.null length :: Bytes -> Int length = ShortByteString.length index :: Bytes -> Int -> Word8 index = ShortByteString.index all :: (Word8 -> Bool) -> Bytes -> Bool all p = List.all p . unpack any :: (Word8 -> Bool) -> Bytes -> Bool any p = List.any p . unpack head :: Bytes -> Word8 head bytes = index bytes 0 last :: Bytes -> Word8 last bytes = index bytes (length bytes - 1) take :: Int -> Bytes -> Bytes take n bs | n == 0 = empty | n >= length bs = bs | otherwise = pack (List.take n (unpack bs)) drop :: Int -> Bytes -> Bytes drop n bs | n == 0 = bs | n >= length bs = empty | otherwise = pack (List.drop n (unpack bs)) isPrefixOf :: Bytes -> Bytes -> Bool isPrefixOf bs1 bs2 = n1 <= n2 && List.all (\i -> index bs1 i == index bs2 i) [0 .. n1 - 1] where n1 = length bs1; n2 = length bs2 isSuffixOf :: Bytes -> Bytes -> Bool isSuffixOf bs1 bs2 = n1 <= n2 && List.all (\i -> index bs1 i == index bs2 (i + k)) [0 .. n1 - 1] where n1 = length bs1; n2 = length bs2; k = n2 - n1 try_unprefix :: Bytes -> Bytes -> Maybe Bytes try_unprefix bs1 bs2 = if isPrefixOf bs1 bs2 then Just (drop (length bs1) bs2) else Nothing try_unsuffix :: Bytes -> Bytes -> Maybe Bytes try_unsuffix bs1 bs2 = if isSuffixOf bs1 bs2 then Just (take (length bs2 - length bs1) bs2) else Nothing concat :: [Bytes] -> Bytes concat = mconcat space :: Word8 space = 32 small_spaces :: Array Int Bytes small_spaces = array (0, 64) [(i, pack (replicate i space)) | i <- [0 .. 64]] spaces :: Int -> Bytes spaces n = if n < 64 then small_spaces ! n else concat ((small_spaces ! (n `mod` 64)) : replicate (n `div` 64) (small_spaces ! 64)) char :: Word8 -> Char char = toEnum . fromEnum all_char :: (Char -> Bool) -> Bytes -> Bool all_char pred = all (pred . char) any_char :: (Char -> Bool) -> Bytes -> Bool any_char pred = any (pred . char) byte :: Char -> Word8 byte = toEnum . fromEnum singletons :: Array Word8 Bytes singletons = array (minBound, maxBound) [(i, make (ByteString.singleton i)) | i <- [minBound .. maxBound]] singleton :: Word8 -> Bytes singleton b = singletons ! b \ generate_file "Isabelle/UTF8.hs" = \ {- Title: Isabelle/UTF8.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Variations on UTF-8. See \<^file>\$ISABELLE_HOME/src/Pure/General/utf8.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/General/utf8.scala\. -} {-# LANGUAGE MultiParamTypeClasses #-} {-# LANGUAGE TypeSynonymInstances #-} {-# LANGUAGE FlexibleInstances #-} {-# LANGUAGE InstanceSigs #-} module Isabelle.UTF8 ( setup, setup3, Recode (..) ) where import qualified System.IO as IO import Data.Text (Text) import qualified Data.Text as Text import qualified Data.Text.Encoding as Encoding import qualified Data.Text.Encoding.Error as Error import Data.ByteString (ByteString) import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) setup :: IO.Handle -> IO () setup h = do IO.hSetEncoding h IO.utf8 IO.hSetNewlineMode h IO.noNewlineTranslation setup3 :: IO.Handle -> IO.Handle -> IO.Handle -> IO () setup3 h1 h2 h3 = do setup h1 setup h2 setup h3 class Recode a b where encode :: a -> b decode :: b -> a instance Recode Text ByteString where encode :: Text -> ByteString encode = Encoding.encodeUtf8 decode :: ByteString -> Text decode = Encoding.decodeUtf8With Error.lenientDecode instance Recode Text Bytes where encode :: Text -> Bytes encode = Bytes.make . encode decode :: Bytes -> Text decode = decode . Bytes.unmake instance Recode String ByteString where encode :: String -> ByteString encode = encode . Text.pack decode :: ByteString -> String decode = Text.unpack . decode instance Recode String Bytes where encode :: String -> Bytes encode = encode . Text.pack decode :: Bytes -> String decode = Text.unpack . decode \ generate_file "Isabelle/Library.hs" = \ {- Title: Isabelle/Library.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Basic library of Isabelle idioms. See \<^file>\$ISABELLE_HOME/src/Pure/General/basics.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/library.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# LANGUAGE TypeSynonymInstances #-} {-# LANGUAGE FlexibleInstances #-} {-# LANGUAGE InstanceSigs #-} module Isabelle.Library ( (|>), (|->), (#>), (#->), fold, fold_rev, fold_map, single, the_single, singletonM, map_index, get_index, separate, StringLike, STRING (..), TEXT (..), BYTES (..), show_bytes, show_text, proper_string, enclose, quote, space_implode, commas, commas_quote, cat_lines, space_explode, split_lines, trim_line, trim_split_lines, getenv, getenv_strict) where import System.Environment (lookupEnv) import Data.Maybe (fromMaybe) import qualified Data.Text as Text import Data.Text (Text) import qualified Data.Text.Lazy as Lazy import Data.String (IsString) import qualified Data.List.Split as Split import qualified Isabelle.Symbol as Symbol import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.UTF8 as UTF8 {- functions -} (|>) :: a -> (a -> b) -> b x |> f = f x (|->) :: (a, b) -> (a -> b -> c) -> c (x, y) |-> f = f x y (#>) :: (a -> b) -> (b -> c) -> a -> c (f #> g) x = x |> f |> g (#->) :: (a -> (c, b)) -> (c -> b -> d) -> a -> d (f #-> g) x = x |> f |-> g {- lists -} fold :: (a -> b -> b) -> [a] -> b -> b fold _ [] y = y fold f (x : xs) y = fold f xs (f x y) fold_rev :: (a -> b -> b) -> [a] -> b -> b fold_rev _ [] y = y fold_rev f (x : xs) y = f x (fold_rev f xs y) fold_map :: (a -> b -> (c, b)) -> [a] -> b -> ([c], b) fold_map _ [] y = ([], y) fold_map f (x : xs) y = let (x', y') = f x y (xs', y'') = fold_map f xs y' in (x' : xs', y'') single :: a -> [a] single x = [x] the_single :: [a] -> a the_single [x] = x the_single _ = undefined singletonM :: Monad m => ([a] -> m [b]) -> a -> m b singletonM f x = the_single <$> f [x] map_index :: ((Int, a) -> b) -> [a] -> [b] map_index f = map_aux 0 where map_aux _ [] = [] map_aux i (x : xs) = f (i, x) : map_aux (i + 1) xs get_index :: (a -> Maybe b) -> [a] -> Maybe (Int, b) get_index f = get_aux 0 where get_aux _ [] = Nothing get_aux i (x : xs) = case f x of Nothing -> get_aux (i + 1) xs Just y -> Just (i, y) separate :: a -> [a] -> [a] separate s (x : xs@(_ : _)) = x : s : separate s xs separate _ xs = xs; {- string-like interfaces -} class (IsString a, Monoid a, Eq a, Ord a) => StringLike a where space_explode :: Char -> a -> [a] trim_line :: a -> a gen_trim_line :: Int -> (Int -> Char) -> (Int -> a -> a) -> a -> a gen_trim_line n at trim s = if n >= 2 && at (n - 2) == '\r' && at (n - 1) == '\n' then trim (n - 2) s else if n >= 1 && Symbol.is_ascii_line_terminator (at (n - 1)) then trim (n - 1) s else s instance StringLike String where space_explode :: Char -> String -> [String] space_explode c = Split.split (Split.dropDelims (Split.whenElt (== c))) trim_line :: String -> String trim_line s = gen_trim_line (length s) (s !!) take s instance StringLike Text where space_explode :: Char -> Text -> [Text] space_explode c str = if Text.null str then [] else if Text.all (/= c) str then [str] else map Text.pack $ space_explode c $ Text.unpack str trim_line :: Text -> Text trim_line s = gen_trim_line (Text.length s) (Text.index s) Text.take s instance StringLike Lazy.Text where space_explode :: Char -> Lazy.Text -> [Lazy.Text] space_explode c str = if Lazy.null str then [] else if Lazy.all (/= c) str then [str] else map Lazy.pack $ space_explode c $ Lazy.unpack str trim_line :: Lazy.Text -> Lazy.Text trim_line = Lazy.fromStrict . trim_line . Lazy.toStrict instance StringLike Bytes where space_explode :: Char -> Bytes -> [Bytes] space_explode c str = if Bytes.null str then [] else if Bytes.all_char (/= c) str then [str] else explode (Bytes.unpack str) where explode rest = case span (/= (Bytes.byte c)) rest of (_, []) -> [Bytes.pack rest] (prfx, _ : rest') -> Bytes.pack prfx : explode rest' trim_line :: Bytes -> Bytes trim_line s = gen_trim_line (Bytes.length s) (Bytes.char . Bytes.index s) Bytes.take s class StringLike a => STRING a where make_string :: a -> String instance STRING String where make_string = id instance STRING Text where make_string = Text.unpack instance STRING Lazy.Text where make_string = Lazy.unpack instance STRING Bytes where make_string = UTF8.decode class StringLike a => TEXT a where make_text :: a -> Text instance TEXT String where make_text = Text.pack instance TEXT Text where make_text = id instance TEXT Lazy.Text where make_text = Lazy.toStrict instance TEXT Bytes where make_text = UTF8.decode class StringLike a => BYTES a where make_bytes :: a -> Bytes instance BYTES String where make_bytes = UTF8.encode instance BYTES Text where make_bytes = UTF8.encode instance BYTES Lazy.Text where make_bytes = UTF8.encode . Lazy.toStrict instance BYTES Bytes where make_bytes = id show_bytes :: Show a => a -> Bytes show_bytes = make_bytes . show show_text :: Show a => a -> Text show_text = make_text . show {- strings -} proper_string :: StringLike a => a -> Maybe a proper_string s = if s == "" then Nothing else Just s enclose :: StringLike a => a -> a -> a -> a enclose lpar rpar str = lpar <> str <> rpar quote :: StringLike a => a -> a quote = enclose "\"" "\"" space_implode :: StringLike a => a -> [a] -> a space_implode s = mconcat . separate s commas, commas_quote :: StringLike a => [a] -> a commas = space_implode ", " commas_quote = commas . map quote split_lines :: StringLike a => a -> [a] split_lines = space_explode '\n' cat_lines :: StringLike a => [a] -> a cat_lines = space_implode "\n" trim_split_lines :: StringLike a => a -> [a] trim_split_lines = trim_line #> split_lines #> map trim_line {- getenv -} getenv :: Bytes -> IO Bytes getenv x = do y <- lookupEnv (make_string x) return $ make_bytes $ fromMaybe "" y getenv_strict :: Bytes -> IO Bytes getenv_strict x = do y <- getenv x if Bytes.null y then errorWithoutStackTrace $ make_string ("Undefined Isabelle environment variable: " <> quote x) else return y \ generate_file "Isabelle/Symbol.hs" = \ {- Title: Isabelle/Symbols.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Isabelle text symbols. See \<^file>\$ISABELLE_HOME/src/Pure/General/symbol.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/General/symbol_explode.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Symbol ( Symbol, eof, is_eof, not_eof, is_ascii_letter, is_ascii_digit, is_ascii_hex, is_ascii_quasi, is_ascii_blank, is_ascii_line_terminator, is_ascii_letdig, is_ascii_identifier, explode ) where import Data.Word (Word8) import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) {- type -} type Symbol = Bytes eof :: Symbol eof = "" is_eof, not_eof :: Symbol -> Bool is_eof = Bytes.null not_eof = not . is_eof {- ASCII characters -} is_ascii_letter :: Char -> Bool is_ascii_letter c = 'A' <= c && c <= 'Z' || 'a' <= c && c <= 'z' is_ascii_digit :: Char -> Bool is_ascii_digit c = '0' <= c && c <= '9' is_ascii_hex :: Char -> Bool is_ascii_hex c = '0' <= c && c <= '9' || 'A' <= c && c <= 'F' || 'a' <= c && c <= 'f' is_ascii_quasi :: Char -> Bool is_ascii_quasi c = c == '_' || c == '\'' is_ascii_blank :: Char -> Bool is_ascii_blank c = c `elem` (" \t\n\11\f\r" :: String) is_ascii_line_terminator :: Char -> Bool is_ascii_line_terminator c = c == '\r' || c == '\n' is_ascii_letdig :: Char -> Bool is_ascii_letdig c = is_ascii_letter c || is_ascii_digit c || is_ascii_quasi c is_ascii_identifier :: String -> Bool is_ascii_identifier s = not (null s) && is_ascii_letter (head s) && all is_ascii_letdig s {- explode symbols: ASCII, UTF8, named -} is_utf8 :: Word8 -> Bool is_utf8 b = b >= 128 is_utf8_trailer :: Word8 -> Bool is_utf8_trailer b = 128 <= b && b < 192 is_utf8_control :: Word8 -> Bool is_utf8_control b = 128 <= b && b < 160 (|>) :: a -> (a -> b) -> b x |> f = f x explode :: Bytes -> [Symbol] explode string = scan 0 where byte = Bytes.index string substring i j = if i == j - 1 then Bytes.singleton (byte i) else Bytes.pack (map byte [i .. j - 1]) n = Bytes.length string test pred i = i < n && pred (byte i) test_char pred i = i < n && pred (Bytes.char (byte i)) many pred i = if test pred i then many pred (i + 1) else i maybe_char c i = if test_char (== c) i then i + 1 else i maybe_ascii_id i = if test_char is_ascii_letter i then many (is_ascii_letdig . Bytes.char) (i + 1) else i scan i = if i < n then let b = byte i c = Bytes.char b in {-encoded newline-} if c == '\r' then "\n" : scan (maybe_char '\n' (i + 1)) {-pseudo utf8: encoded ascii control-} else if b == 192 && test is_utf8_control (i + 1) && not (test is_utf8 (i + 2)) then Bytes.singleton (byte (i + 1) - 128) : scan (i + 2) {-utf8-} else if is_utf8 b then let j = many is_utf8_trailer (i + 1) in substring i j : scan j {-named symbol-} else if c == '\\' && test_char (== '<') (i + 1) then let j = (i + 2) |> maybe_char '^' |> maybe_ascii_id |> maybe_char '>' in substring i j : scan j {-single character-} else Bytes.singleton b : scan (i + 1) else [] \ generate_file "Isabelle/Buffer.hs" = \ {- Title: Isabelle/Buffer.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Efficient buffer of byte strings. See \<^file>\$ISABELLE_HOME/src/Pure/General/buffer.ML\. -} module Isabelle.Buffer (T, empty, add, content, build, build_content) where import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import Isabelle.Library newtype T = Buffer [Bytes] empty :: T empty = Buffer [] add :: Bytes -> T -> T add b (Buffer bs) = Buffer (if Bytes.null b then bs else b : bs) content :: T -> Bytes content (Buffer bs) = Bytes.concat (reverse bs) build :: (T -> T) -> T build f = f empty build_content :: (T -> T) -> Bytes build_content f = build f |> content \ generate_file "Isabelle/Value.hs" = \ {- Title: Isabelle/Value.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Plain values, represented as string. See \<^file>\$ISABELLE_HOME/src/Pure/General/value.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Value (print_bool, parse_bool, parse_nat, print_int, parse_int, print_real, parse_real) where import qualified Data.List as List import qualified Text.Read as Read import Isabelle.Bytes (Bytes) import Isabelle.Library {- bool -} print_bool :: Bool -> Bytes print_bool True = "true" print_bool False = "false" parse_bool :: Bytes -> Maybe Bool parse_bool "true" = Just True parse_bool "false" = Just False parse_bool _ = Nothing {- nat -} parse_nat :: Bytes -> Maybe Int parse_nat s = case Read.readMaybe (make_string s) of Just n | n >= 0 -> Just n _ -> Nothing {- int -} print_int :: Int -> Bytes print_int = show_bytes parse_int :: Bytes -> Maybe Int parse_int = Read.readMaybe . make_string {- real -} print_real :: Double -> Bytes print_real x = let s = show x in case span (/= '.') s of (a, '.' : b) | List.all (== '0') b -> make_bytes a _ -> make_bytes s parse_real :: Bytes -> Maybe Double parse_real = Read.readMaybe . make_string \ generate_file "Isabelle/Properties.hs" = \ {- Title: Isabelle/Properties.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Property lists. See \<^file>\$ISABELLE_HOME/src/Pure/General/properties.ML\. -} module Isabelle.Properties (Entry, T, defined, get, get_value, put, remove) where import qualified Data.List as List import Isabelle.Bytes (Bytes) type Entry = (Bytes, Bytes) type T = [Entry] defined :: T -> Bytes -> Bool defined props name = any (\(a, _) -> a == name) props get :: T -> Bytes -> Maybe Bytes get props name = List.lookup name props get_value :: (Bytes -> Maybe a) -> T -> Bytes -> Maybe a get_value parse props name = maybe Nothing parse (get props name) put :: Entry -> T -> T put entry props = entry : remove (fst entry) props remove :: Bytes -> T -> T remove name props = if defined props name then filter (\(a, _) -> a /= name) props else props \ generate_file "Isabelle/Markup.hs" = \ {- Title: Isabelle/Markup.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Quasi-abstract markup elements. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/markup.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures #-} module Isabelle.Markup ( T, empty, is_empty, properties, nameN, name, xnameN, xname, kindN, bindingN, binding, entityN, entity, defN, refN, completionN, completion, no_completionN, no_completion, - lineN, end_lineN, offsetN, end_offsetN, fileN, idN, positionN, position, + lineN, end_lineN, offsetN, end_offsetN, labelN, fileN, idN, positionN, position, position_properties, def_name, expressionN, expression, pathN, path, urlN, url, docN, doc, markupN, consistentN, unbreakableN, indentN, widthN, blockN, block, breakN, break, fbreakN, fbreak, itemN, item, wordsN, words, tfreeN, tfree, tvarN, tvar, freeN, free, skolemN, skolem, boundN, bound, varN, var, numeralN, numeral, literalN, literal, delimiterN, delimiter, inner_stringN, inner_string, inner_cartoucheN, inner_cartouche, token_rangeN, token_range, sortingN, sorting, typingN, typing, class_parameterN, class_parameter, antiquotedN, antiquoted, antiquoteN, antiquote, paragraphN, paragraph, text_foldN, text_fold, keyword1N, keyword1, keyword2N, keyword2, keyword3N, keyword3, quasi_keywordN, quasi_keyword, improperN, improper, operatorN, operator, stringN, string, alt_stringN, alt_string, verbatimN, verbatim, cartoucheN, cartouche, commentN, comment, comment1N, comment1, comment2N, comment2, comment3N, comment3, forkedN, forked, joinedN, joined, runningN, running, finishedN, finished, failedN, failed, canceledN, canceled, initializedN, initialized, finalizedN, finalized, consolidatedN, consolidated, writelnN, writeln, stateN, state, informationN, information, tracingN, tracing, warningN, warning, legacyN, legacy, errorN, error, reportN, report, no_reportN, no_report, intensifyN, intensify, Output, no_output) where import Prelude hiding (words, error, break) import Data.Map.Strict (Map) import qualified Data.Map.Strict as Map import Isabelle.Library import qualified Isabelle.Properties as Properties import qualified Isabelle.Value as Value import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) {- basic markup -} type T = (Bytes, Properties.T) empty :: T empty = ("", []) is_empty :: T -> Bool is_empty ("", _) = True is_empty _ = False properties :: Properties.T -> T -> T properties more_props (elem, props) = (elem, fold_rev Properties.put more_props props) markup_elem :: Bytes -> T markup_elem name = (name, []) markup_string :: Bytes -> Bytes -> Bytes -> T markup_string name prop = \s -> (name, [(prop, s)]) {- misc properties -} nameN :: Bytes nameN = \Markup.nameN\ name :: Bytes -> T -> T name a = properties [(nameN, a)] xnameN :: Bytes xnameN = \Markup.xnameN\ xname :: Bytes -> T -> T xname a = properties [(xnameN, a)] kindN :: Bytes kindN = \Markup.kindN\ {- formal entities -} bindingN :: Bytes bindingN = \Markup.bindingN\ binding :: T binding = markup_elem bindingN entityN :: Bytes entityN = \Markup.entityN\ entity :: Bytes -> Bytes -> T entity kind name = (entityN, (if Bytes.null name then [] else [(nameN, name)]) <> (if Bytes.null kind then [] else [(kindN, kind)])) defN :: Bytes defN = \Markup.defN\ refN :: Bytes refN = \Markup.refN\ {- completion -} completionN :: Bytes completionN = \Markup.completionN\ completion :: T completion = markup_elem completionN no_completionN :: Bytes no_completionN = \Markup.no_completionN\ no_completion :: T no_completion = markup_elem no_completionN {- position -} lineN, end_lineN :: Bytes lineN = \Markup.lineN\ end_lineN = \Markup.end_lineN\ offsetN, end_offsetN :: Bytes offsetN = \Markup.offsetN\ end_offsetN = \Markup.end_offsetN\ -fileN, idN :: Bytes +labelN, fileN, idN :: Bytes +labelN = \Markup.labelN\ fileN = \Markup.fileN\ idN = \Markup.idN\ positionN :: Bytes positionN = \Markup.positionN\ position :: T position = markup_elem positionN position_properties :: [Bytes] -position_properties = [lineN, offsetN, end_offsetN, fileN, idN] +position_properties = [lineN, offsetN, end_offsetN, labelN, fileN, idN] {- position "def" names -} make_def :: Bytes -> Bytes make_def a = "def_" <> a def_names :: Map Bytes Bytes def_names = Map.fromList $ map (\a -> (a, make_def a)) position_properties def_name :: Bytes -> Bytes def_name a = case Map.lookup a def_names of Just b -> b Nothing -> make_def a {- expression -} expressionN :: Bytes expressionN = \Markup.expressionN\ expression :: Bytes -> T expression kind = (expressionN, if kind == "" then [] else [(kindN, kind)]) {- external resources -} pathN :: Bytes pathN = \Markup.pathN\ path :: Bytes -> T path = markup_string pathN nameN urlN :: Bytes urlN = \Markup.urlN\ url :: Bytes -> T url = markup_string urlN nameN docN :: Bytes docN = \Markup.docN\ doc :: Bytes -> T doc = markup_string docN nameN {- pretty printing -} markupN, consistentN, unbreakableN, indentN :: Bytes markupN = \Markup.markupN\ consistentN = \Markup.consistentN\ unbreakableN = \Markup.unbreakableN\ indentN = \Markup.indentN\ widthN :: Bytes widthN = \Markup.widthN\ blockN :: Bytes blockN = \Markup.blockN\ block :: Bool -> Int -> T block c i = (blockN, (if c then [(consistentN, Value.print_bool c)] else []) <> (if i /= 0 then [(indentN, Value.print_int i)] else [])) breakN :: Bytes breakN = \Markup.breakN\ break :: Int -> Int -> T break w i = (breakN, (if w /= 0 then [(widthN, Value.print_int w)] else []) <> (if i /= 0 then [(indentN, Value.print_int i)] else [])) fbreakN :: Bytes fbreakN = \Markup.fbreakN\ fbreak :: T fbreak = markup_elem fbreakN itemN :: Bytes itemN = \Markup.itemN\ item :: T item = markup_elem itemN {- text properties -} wordsN :: Bytes wordsN = \Markup.wordsN\ words :: T words = markup_elem wordsN {- inner syntax -} tfreeN :: Bytes tfreeN = \Markup.tfreeN\ tfree :: T tfree = markup_elem tfreeN tvarN :: Bytes tvarN = \Markup.tvarN\ tvar :: T tvar = markup_elem tvarN freeN :: Bytes freeN = \Markup.freeN\ free :: T free = markup_elem freeN skolemN :: Bytes skolemN = \Markup.skolemN\ skolem :: T skolem = markup_elem skolemN boundN :: Bytes boundN = \Markup.boundN\ bound :: T bound = markup_elem boundN varN :: Bytes varN = \Markup.varN\ var :: T var = markup_elem varN numeralN :: Bytes numeralN = \Markup.numeralN\ numeral :: T numeral = markup_elem numeralN literalN :: Bytes literalN = \Markup.literalN\ literal :: T literal = markup_elem literalN delimiterN :: Bytes delimiterN = \Markup.delimiterN\ delimiter :: T delimiter = markup_elem delimiterN inner_stringN :: Bytes inner_stringN = \Markup.inner_stringN\ inner_string :: T inner_string = markup_elem inner_stringN inner_cartoucheN :: Bytes inner_cartoucheN = \Markup.inner_cartoucheN\ inner_cartouche :: T inner_cartouche = markup_elem inner_cartoucheN token_rangeN :: Bytes token_rangeN = \Markup.token_rangeN\ token_range :: T token_range = markup_elem token_rangeN sortingN :: Bytes sortingN = \Markup.sortingN\ sorting :: T sorting = markup_elem sortingN typingN :: Bytes typingN = \Markup.typingN\ typing :: T typing = markup_elem typingN class_parameterN :: Bytes class_parameterN = \Markup.class_parameterN\ class_parameter :: T class_parameter = markup_elem class_parameterN {- antiquotations -} antiquotedN :: Bytes antiquotedN = \Markup.antiquotedN\ antiquoted :: T antiquoted = markup_elem antiquotedN antiquoteN :: Bytes antiquoteN = \Markup.antiquoteN\ antiquote :: T antiquote = markup_elem antiquoteN {- text structure -} paragraphN :: Bytes paragraphN = \Markup.paragraphN\ paragraph :: T paragraph = markup_elem paragraphN text_foldN :: Bytes text_foldN = \Markup.text_foldN\ text_fold :: T text_fold = markup_elem text_foldN {- outer syntax -} keyword1N :: Bytes keyword1N = \Markup.keyword1N\ keyword1 :: T keyword1 = markup_elem keyword1N keyword2N :: Bytes keyword2N = \Markup.keyword2N\ keyword2 :: T keyword2 = markup_elem keyword2N keyword3N :: Bytes keyword3N = \Markup.keyword3N\ keyword3 :: T keyword3 = markup_elem keyword3N quasi_keywordN :: Bytes quasi_keywordN = \Markup.quasi_keywordN\ quasi_keyword :: T quasi_keyword = markup_elem quasi_keywordN improperN :: Bytes improperN = \Markup.improperN\ improper :: T improper = markup_elem improperN operatorN :: Bytes operatorN = \Markup.operatorN\ operator :: T operator = markup_elem operatorN stringN :: Bytes stringN = \Markup.stringN\ string :: T string = markup_elem stringN alt_stringN :: Bytes alt_stringN = \Markup.alt_stringN\ alt_string :: T alt_string = markup_elem alt_stringN verbatimN :: Bytes verbatimN = \Markup.verbatimN\ verbatim :: T verbatim = markup_elem verbatimN cartoucheN :: Bytes cartoucheN = \Markup.cartoucheN\ cartouche :: T cartouche = markup_elem cartoucheN commentN :: Bytes commentN = \Markup.commentN\ comment :: T comment = markup_elem commentN {- comments -} comment1N :: Bytes comment1N = \Markup.comment1N\ comment1 :: T comment1 = markup_elem comment1N comment2N :: Bytes comment2N = \Markup.comment2N\ comment2 :: T comment2 = markup_elem comment2N comment3N :: Bytes comment3N = \Markup.comment3N\ comment3 :: T comment3 = markup_elem comment3N {- command status -} forkedN, joinedN, runningN, finishedN, failedN, canceledN, initializedN, finalizedN, consolidatedN :: Bytes forkedN = \Markup.forkedN\ joinedN = \Markup.joinedN\ runningN = \Markup.runningN\ finishedN = \Markup.finishedN\ failedN = \Markup.failedN\ canceledN = \Markup.canceledN\ initializedN = \Markup.initializedN\ finalizedN = \Markup.finalizedN\ consolidatedN = \Markup.consolidatedN\ forked, joined, running, finished, failed, canceled, initialized, finalized, consolidated :: T forked = markup_elem forkedN joined = markup_elem joinedN running = markup_elem runningN finished = markup_elem finishedN failed = markup_elem failedN canceled = markup_elem canceledN initialized = markup_elem initializedN finalized = markup_elem finalizedN consolidated = markup_elem consolidatedN {- messages -} writelnN :: Bytes writelnN = \Markup.writelnN\ writeln :: T writeln = markup_elem writelnN stateN :: Bytes stateN = \Markup.stateN\ state :: T state = markup_elem stateN informationN :: Bytes informationN = \Markup.informationN\ information :: T information = markup_elem informationN tracingN :: Bytes tracingN = \Markup.tracingN\ tracing :: T tracing = markup_elem tracingN warningN :: Bytes warningN = \Markup.warningN\ warning :: T warning = markup_elem warningN legacyN :: Bytes legacyN = \Markup.legacyN\ legacy :: T legacy = markup_elem legacyN errorN :: Bytes errorN = \Markup.errorN\ error :: T error = markup_elem errorN reportN :: Bytes reportN = \Markup.reportN\ report :: T report = markup_elem reportN no_reportN :: Bytes no_reportN = \Markup.no_reportN\ no_report :: T no_report = markup_elem no_reportN intensifyN :: Bytes intensifyN = \Markup.intensifyN\ intensify :: T intensify = markup_elem intensifyN {- output -} type Output = (Bytes, Bytes) no_output :: Output no_output = ("", "") \ generate_file "Isabelle/Position.hs" = \ {- Title: Isabelle/Position.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Source positions starting from 1; values <= 0 mean "absent". Count Isabelle symbols, not UTF8 bytes nor UTF16 characters. Position range specifies a right-open interval offset .. end_offset (exclusive). See \<^file>\$ISABELLE_HOME/src/Pure/General/position.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Position ( T, line_of, column_of, offset_of, end_offset_of, file_of, id_of, start, none, put_file, file, file_only, put_id, id, id_only, symbol, symbol_explode, symbol_explode_string, shift_offsets, of_properties, properties_of, def_properties_of, entity_markup, make_entity_markup, Report, Report_Text, is_reported, is_reported_range, here, Range, no_range, no_range_position, range_position, range ) where import Prelude hiding (id) import Data.Maybe (isJust, fromMaybe) import Data.Bifunctor (first) import qualified Isabelle.Properties as Properties import qualified Isabelle.Bytes as Bytes import qualified Isabelle.Value as Value import Isabelle.Bytes (Bytes) import qualified Isabelle.Markup as Markup import qualified Isabelle.YXML as YXML import Isabelle.Library import qualified Isabelle.Symbol as Symbol import Isabelle.Symbol (Symbol) {- position -} data T = Position { _line :: Int, _column :: Int, _offset :: Int, _end_offset :: Int, + _label :: Bytes, _file :: Bytes, _id :: Bytes } deriving (Eq, Ord) valid, invalid :: Int -> Bool valid i = i > 0 invalid = not . valid maybe_valid :: Int -> Maybe Int maybe_valid i = if valid i then Just i else Nothing if_valid :: Int -> Int -> Int if_valid i i' = if valid i then i' else i {- fields -} line_of, column_of, offset_of, end_offset_of :: T -> Maybe Int line_of = maybe_valid . _line column_of = maybe_valid . _column offset_of = maybe_valid . _offset end_offset_of = maybe_valid . _end_offset +label_of :: T -> Maybe Bytes +label_of = proper_string . _label + file_of :: T -> Maybe Bytes file_of = proper_string . _file id_of :: T -> Maybe Bytes id_of = proper_string . _id {- make position -} start :: T -start = Position 1 1 1 0 Bytes.empty Bytes.empty +start = Position 1 1 1 0 Bytes.empty Bytes.empty Bytes.empty none :: T -none = Position 0 0 0 0 Bytes.empty Bytes.empty +none = Position 0 0 0 0 Bytes.empty Bytes.empty Bytes.empty + +label :: Bytes -> T -> T +label label pos = pos { _label = label } put_file :: Bytes -> T -> T put_file file pos = pos { _file = file } file :: Bytes -> T file file = put_file file start file_only :: Bytes -> T file_only file = put_file file none put_id :: Bytes -> T -> T put_id id pos = pos { _id = id } id :: Bytes -> T id id = put_id id start id_only :: Bytes -> T id_only id = put_id id none {- count position -} count_line :: Symbol -> Int -> Int count_line "\n" line = if_valid line (line + 1) count_line _ line = line count_column :: Symbol -> Int -> Int count_column "\n" column = if_valid column 1 count_column s column = if Symbol.not_eof s then if_valid column (column + 1) else column count_offset :: Symbol -> Int -> Int count_offset s offset = if Symbol.not_eof s then if_valid offset (offset + 1) else offset symbol :: Symbol -> T -> T symbol s pos = pos { _line = count_line s (_line pos), _column = count_column s (_column pos), _offset = count_offset s (_offset pos) } symbol_explode :: BYTES a => a -> T -> T symbol_explode = fold symbol . Symbol.explode . make_bytes symbol_explode_string :: String -> T -> T symbol_explode_string = symbol_explode {- shift offsets -} shift_offsets :: Int -> T -> T shift_offsets shift pos = pos { _offset = offset', _end_offset = end_offset' } where offset = _offset pos end_offset = _end_offset pos offset' = if invalid offset || invalid shift then offset else offset + shift end_offset' = if invalid end_offset || invalid shift then end_offset else end_offset + shift {- markup properties -} get_string :: Properties.T -> Bytes -> Bytes get_string props name = fromMaybe "" (Properties.get_value Just props name) get_int :: Properties.T -> Bytes -> Int get_int props name = fromMaybe 0 (Properties.get_value Value.parse_int props name) of_properties :: Properties.T -> T of_properties props = none { _line = get_int props Markup.lineN, _offset = get_int props Markup.offsetN, _end_offset = get_int props Markup.end_offsetN, + _label = get_string props Markup.labelN, _file = get_string props Markup.fileN, _id = get_string props Markup.idN } string_entry :: Bytes -> Bytes -> Properties.T string_entry k s = if Bytes.null s then [] else [(k, s)] int_entry :: Bytes -> Int -> Properties.T int_entry k i = if invalid i then [] else [(k, Value.print_int i)] properties_of :: T -> Properties.T properties_of pos = int_entry Markup.lineN (_line pos) ++ int_entry Markup.offsetN (_offset pos) ++ int_entry Markup.end_offsetN (_end_offset pos) ++ + string_entry Markup.labelN (_label pos) ++ string_entry Markup.fileN (_file pos) ++ string_entry Markup.idN (_id pos) def_properties_of :: T -> Properties.T def_properties_of = properties_of #> map (first Markup.def_name) entity_markup :: Bytes -> (Bytes, T) -> Markup.T entity_markup kind (name, pos) = Markup.entity kind name |> Markup.properties (def_properties_of pos) make_entity_markup :: Bool -> Int -> Bytes -> (Bytes, T) -> Markup.T make_entity_markup def serial kind (name, pos) = let props = if def then (Markup.defN, Value.print_int serial) : properties_of pos else (Markup.refN, Value.print_int serial) : def_properties_of pos in Markup.entity kind name |> Markup.properties props {- reports -} type Report = (T, Markup.T) type Report_Text = (Report, Bytes) is_reported :: T -> Bool is_reported pos = isJust (offset_of pos) && isJust (id_of pos) is_reported_range :: T -> Bool is_reported_range pos = is_reported pos && isJust (end_offset_of pos) {- here: user output -} here :: T -> Bytes here pos = if Bytes.null s2 then "" else s1 <> m1 <> s2 <> m2 where props = properties_of pos (m1, m2) = YXML.output_markup (Markup.properties props Markup.position) (s1, s2) = case (line_of pos, file_of pos) of (Just i, Nothing) -> (" ", "(line " <> Value.print_int i <> ")") (Just i, Just name) -> (" ", "(line " <> Value.print_int i <> " of " <> quote name <> ")") (Nothing, Just name) -> (" ", "(file " <> quote name <> ")") _ -> if is_reported pos then ("", "\092<^here>") else ("", "") {- range -} type Range = (T, T) no_range :: Range no_range = (none, none) no_range_position :: T -> T no_range_position pos = pos { _end_offset = 0 } range_position :: Range -> T range_position (pos, pos') = pos { _end_offset = _offset pos' } range :: Range -> Range range (pos, pos') = (range_position (pos, pos'), no_range_position pos') \ generate_file "Isabelle/XML.hs" = \ {- Title: Isabelle/XML.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Untyped XML trees and representation of ML values. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/xml.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures #-} module Isabelle.XML (Attributes, Body, Tree(..), wrap_elem, unwrap_elem, content_of) where import Isabelle.Library import qualified Isabelle.Properties as Properties import qualified Isabelle.Markup as Markup import qualified Isabelle.Buffer as Buffer import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) {- types -} type Attributes = Properties.T type Body = [Tree] data Tree = Elem (Markup.T, Body) | Text Bytes {- wrapped elements -} wrap_elem :: ((Markup.T, Body), [Tree]) -> Tree wrap_elem (((a, atts), body1), body2) = Elem ((\XML.xml_elemN\, (\XML.xml_nameN\, a) : atts), Elem ((\XML.xml_bodyN\, []), body1) : body2) unwrap_elem :: Tree -> Maybe ((Markup.T, Body), [Tree]) unwrap_elem (Elem ((\XML.xml_elemN\, (\XML.xml_nameN\, a) : atts), Elem ((\XML.xml_bodyN\, []), body1) : body2)) = Just (((a, atts), body1), body2) unwrap_elem _ = Nothing {- text content -} add_content :: Tree -> Buffer.T -> Buffer.T add_content tree = case unwrap_elem tree of Just (_, ts) -> fold add_content ts Nothing -> case tree of Elem (_, ts) -> fold add_content ts Text s -> Buffer.add s content_of :: Body -> Bytes content_of = Buffer.build_content . fold add_content {- string representation -} encode_char :: Char -> String encode_char '<' = "<" encode_char '>' = ">" encode_char '&' = "&" encode_char '\'' = "'" encode_char '\"' = """ encode_char c = [c] encode_text :: Bytes -> Bytes encode_text = make_bytes . concatMap (encode_char . Bytes.char) . Bytes.unpack instance Show Tree where show tree = make_string $ Buffer.build_content (show_tree tree) where show_tree (Elem ((name, atts), [])) = Buffer.add "<" #> Buffer.add (show_elem name atts) #> Buffer.add "/>" show_tree (Elem ((name, atts), ts)) = Buffer.add "<" #> Buffer.add (show_elem name atts) #> Buffer.add ">" #> fold show_tree ts #> Buffer.add " Buffer.add name #> Buffer.add ">" show_tree (Text s) = Buffer.add (encode_text s) show_elem name atts = space_implode " " (name : map (\(a, x) -> a <> "=\"" <> encode_text x <> "\"") atts) \ generate_file "Isabelle/XML/Encode.hs" = \ {- Title: Isabelle/XML/Encode.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) XML as data representation language. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/xml.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures #-} module Isabelle.XML.Encode ( A, T, V, P, int_atom, bool_atom, unit_atom, tree, properties, string, int, bool, unit, pair, triple, list, option, variant ) where import Data.Maybe (fromJust) import Isabelle.Library import Isabelle.Bytes (Bytes) import qualified Isabelle.Value as Value import qualified Isabelle.Properties as Properties import qualified Isabelle.XML as XML type A a = a -> Bytes type T a = a -> XML.Body type V a = a -> Maybe ([Bytes], XML.Body) type P a = a -> [Bytes] -- atomic values int_atom :: A Int int_atom = Value.print_int bool_atom :: A Bool bool_atom False = "0" bool_atom True = "1" unit_atom :: A () unit_atom () = "" -- structural nodes node ts = XML.Elem ((":", []), ts) vector = map_index (\(i, x) -> (int_atom i, x)) tagged (tag, (xs, ts)) = XML.Elem ((int_atom tag, vector xs), ts) -- representation of standard types tree :: T XML.Tree tree t = [t] properties :: T Properties.T properties props = [XML.Elem ((":", props), [])] string :: T Bytes string "" = [] string s = [XML.Text s] int :: T Int int = string . int_atom bool :: T Bool bool = string . bool_atom unit :: T () unit = string . unit_atom pair :: T a -> T b -> T (a, b) pair f g (x, y) = [node (f x), node (g y)] triple :: T a -> T b -> T c -> T (a, b, c) triple f g h (x, y, z) = [node (f x), node (g y), node (h z)] list :: T a -> T [a] list f xs = map (node . f) xs option :: T a -> T (Maybe a) option _ Nothing = [] option f (Just x) = [node (f x)] variant :: [V a] -> T a variant fs x = [tagged (fromJust (get_index (\f -> f x) fs))] \ generate_file "Isabelle/XML/Decode.hs" = \ {- Title: Isabelle/XML/Decode.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) XML as data representation language. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/xml.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures #-} module Isabelle.XML.Decode ( A, T, V, P, int_atom, bool_atom, unit_atom, tree, properties, string, int, bool, unit, pair, triple, list, option, variant ) where import Isabelle.Library import Isabelle.Bytes (Bytes) import qualified Isabelle.Value as Value import qualified Isabelle.Properties as Properties import qualified Isabelle.XML as XML type A a = Bytes -> a type T a = XML.Body -> a type V a = ([Bytes], XML.Body) -> a type P a = [Bytes] -> a err_atom = error "Malformed XML atom" err_body = error "Malformed XML body" {- atomic values -} int_atom :: A Int int_atom s = case Value.parse_int s of Just i -> i Nothing -> err_atom bool_atom :: A Bool bool_atom "0" = False bool_atom "1" = True bool_atom _ = err_atom unit_atom :: A () unit_atom "" = () unit_atom _ = err_atom {- structural nodes -} node (XML.Elem ((":", []), ts)) = ts node _ = err_body vector atts = map_index (\(i, (a, x)) -> if int_atom a == i then x else err_atom) atts tagged (XML.Elem ((name, atts), ts)) = (int_atom name, (vector atts, ts)) tagged _ = err_body {- representation of standard types -} tree :: T XML.Tree tree [t] = t tree _ = err_body properties :: T Properties.T properties [XML.Elem ((":", props), [])] = props properties _ = err_body string :: T Bytes string [] = "" string [XML.Text s] = s string _ = err_body int :: T Int int = int_atom . string bool :: T Bool bool = bool_atom . string unit :: T () unit = unit_atom . string pair :: T a -> T b -> T (a, b) pair f g [t1, t2] = (f (node t1), g (node t2)) pair _ _ _ = err_body triple :: T a -> T b -> T c -> T (a, b, c) triple f g h [t1, t2, t3] = (f (node t1), g (node t2), h (node t3)) triple _ _ _ _ = err_body list :: T a -> T [a] list f ts = map (f . node) ts option :: T a -> T (Maybe a) option _ [] = Nothing option f [t] = Just (f (node t)) option _ _ = err_body variant :: [V a] -> T a variant fs [t] = (fs !! tag) (xs, ts) where (tag, (xs, ts)) = tagged t variant _ _ = err_body \ generate_file "Isabelle/YXML.hs" = \ {- Title: Isabelle/YXML.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Efficient text representation of XML trees. Suitable for direct inlining into plain text. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/yxml.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures -fno-warn-incomplete-patterns #-} module Isabelle.YXML (charX, charY, strX, strY, detect, output_markup, buffer_body, buffer, string_of_body, string_of, parse_body, parse) where import qualified Data.List as List import Data.Word (Word8) import Isabelle.Library import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.Markup as Markup import qualified Isabelle.XML as XML import qualified Isabelle.Buffer as Buffer {- markers -} charX, charY :: Word8 charX = 5 charY = 6 strX, strY, strXY, strXYX :: Bytes strX = Bytes.singleton charX strY = Bytes.singleton charY strXY = strX <> strY strXYX = strXY <> strX detect :: Bytes -> Bool detect = Bytes.any (\c -> c == charX || c == charY) {- output -} output_markup :: Markup.T -> Markup.Output output_markup markup@(name, atts) = if Markup.is_empty markup then Markup.no_output else (strXY <> name <> Bytes.concat (map (\(a, x) -> strY <> a <> "=" <> x) atts) <> strX, strXYX) buffer_attrib (a, x) = Buffer.add strY #> Buffer.add a #> Buffer.add "=" #> Buffer.add x buffer_body :: XML.Body -> Buffer.T -> Buffer.T buffer_body = fold buffer buffer :: XML.Tree -> Buffer.T -> Buffer.T buffer (XML.Elem ((name, atts), ts)) = Buffer.add strXY #> Buffer.add name #> fold buffer_attrib atts #> Buffer.add strX #> buffer_body ts #> Buffer.add strXYX buffer (XML.Text s) = Buffer.add s string_of_body :: XML.Body -> Bytes string_of_body = Buffer.build_content . buffer_body string_of :: XML.Tree -> Bytes string_of = string_of_body . single {- parse -} -- split: fields or non-empty tokens split :: Bool -> Word8 -> [Word8] -> [[Word8]] split _ _ [] = [] split fields sep str = splitting str where splitting rest = case span (/= sep) rest of (_, []) -> cons rest [] (prfx, _ : rest') -> cons prfx (splitting rest') cons item = if fields || not (null item) then (:) item else id -- structural errors err :: Bytes -> a err msg = error (make_string ("Malformed YXML: " <> msg)) err_attribute = err "bad attribute" err_element = err "bad element" err_unbalanced :: Bytes -> a err_unbalanced name = if Bytes.null name then err "unbalanced element" else err ("unbalanced element " <> quote name) -- stack operations add x ((elem, body) : pending) = (elem, x : body) : pending push name atts pending = if Bytes.null name then err_element else ((name, atts), []) : pending pop (((name, atts), body) : pending) = if Bytes.null name then err_unbalanced name else add (XML.Elem ((name, atts), reverse body)) pending -- parsing parse_attrib s = case List.elemIndex (Bytes.byte '=') s of Just i | i > 0 -> (Bytes.pack $ take i s, Bytes.pack $ drop (i + 1) s) _ -> err_attribute parse_chunk [[], []] = pop parse_chunk ([] : name : atts) = push (Bytes.pack name) (map parse_attrib atts) parse_chunk txts = fold (add . XML.Text . Bytes.pack) txts parse_body :: Bytes -> XML.Body parse_body source = case fold parse_chunk chunks [((Bytes.empty, []), [])] of [((name, _), result)] | Bytes.null name -> reverse result ((name, _), _) : _ -> err_unbalanced name where chunks = source |> Bytes.unpack |> split False charX |> map (split True charY) parse :: Bytes -> XML.Tree parse source = case parse_body source of [result] -> result [] -> XML.Text "" _ -> err "multiple results" \ generate_file "Isabelle/Completion.hs" = \ {- Title: Isabelle/Completion.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Completion of names. See \<^file>\$ISABELLE_HOME/src/Pure/General/completion.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Completion ( Name, T, names, none, make, markup_element, markup_report, make_report ) where import qualified Isabelle.Bytes as Bytes import qualified Isabelle.Name as Name import Isabelle.Name (Name) import qualified Isabelle.Properties as Properties import qualified Isabelle.Markup as Markup import Isabelle.XML.Classes import qualified Isabelle.XML as XML import qualified Isabelle.YXML as YXML type Names = [(Name, (Name, Name))] -- external name, kind, internal name data T = Completion Properties.T Int Names -- position, total length, names names :: Int -> Properties.T -> Names -> T names limit props names = Completion props (length names) (take limit names) none :: T none = names 0 [] [] make :: Int -> (Name, Properties.T) -> ((Name -> Bool) -> Names) -> T make limit (name, props) make_names = if name /= "" && name /= "_" then names limit props (make_names (Bytes.isPrefixOf (Name.clean name))) else none markup_element :: T -> (Markup.T, XML.Body) markup_element (Completion props total names) = if not (null names) then (Markup.properties props Markup.completion, encode (total, names)) else (Markup.empty, []) markup_report :: [T] -> Name markup_report [] = Bytes.empty markup_report elems = YXML.string_of $ XML.Elem (Markup.report, map (XML.Elem . markup_element) elems) make_report :: Int -> (Name, Properties.T) -> ((Name -> Bool) -> Names) -> Name make_report limit name_props make_names = markup_report [make limit name_props make_names] \ generate_file "Isabelle/File.hs" = \ {- Title: Isabelle/File.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) File-system operations. See \<^file>\$ISABELLE_HOME/src/Pure/General/file.ML\. -} module Isabelle.File (read, write, append) where import Prelude hiding (read) import qualified System.IO as IO import qualified Data.ByteString as ByteString import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) read :: IO.FilePath -> IO Bytes read path = Bytes.make <$> IO.withFile path IO.ReadMode ByteString.hGetContents write :: IO.FilePath -> Bytes -> IO () write path bs = IO.withFile path IO.WriteMode (\h -> ByteString.hPut h (Bytes.unmake bs)) append :: IO.FilePath -> Bytes -> IO () append path bs = IO.withFile path IO.AppendMode (\h -> ByteString.hPut h (Bytes.unmake bs)) \ generate_file "Isabelle/Pretty.hs" = \ {- Title: Isabelle/Pretty.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Generic pretty printing module. See \<^file>\$ISABELLE_HOME/src/Pure/General/pretty.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures #-} module Isabelle.Pretty ( T, symbolic, formatted, unformatted, str, brk_indent, brk, fbrk, breaks, fbreaks, blk, block, strs, markup, mark, mark_str, marks_str, item, text_fold, keyword1, keyword2, text, paragraph, para, quote, cartouche, separate, commas, enclose, enum, list, str_list, big_list) where import qualified Data.List as List import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import Isabelle.Library hiding (enclose, quote, separate, commas) import qualified Isabelle.Buffer as Buffer import qualified Isabelle.Markup as Markup import qualified Isabelle.XML as XML import qualified Isabelle.YXML as YXML data T = Block Markup.T Bool Int [T] | Break Int Int | Str Bytes {- output -} symbolic_text s = if Bytes.null s then [] else [XML.Text s] symbolic_markup markup body = if Markup.is_empty markup then body else [XML.Elem (markup, body)] symbolic :: T -> XML.Body symbolic (Block markup consistent indent prts) = concatMap symbolic prts |> symbolic_markup block_markup |> symbolic_markup markup where block_markup = if null prts then Markup.empty else Markup.block consistent indent symbolic (Break wd ind) = [XML.Elem (Markup.break wd ind, symbolic_text (Bytes.spaces wd))] symbolic (Str s) = symbolic_text s formatted :: T -> Bytes formatted = YXML.string_of_body . symbolic unformatted :: T -> Bytes unformatted = Buffer.build_content . out where out (Block markup _ _ prts) = let (bg, en) = YXML.output_markup markup in Buffer.add bg #> fold out prts #> Buffer.add en out (Break _ wd) = Buffer.add (Bytes.spaces wd) out (Str s) = Buffer.add s {- derived operations to create formatting expressions -} force_nat n | n < 0 = 0 force_nat n = n str :: BYTES a => a -> T str = Str . make_bytes brk_indent :: Int -> Int -> T brk_indent wd ind = Break (force_nat wd) ind brk :: Int -> T brk wd = brk_indent wd 0 fbrk :: T fbrk = Str "\n" breaks, fbreaks :: [T] -> [T] breaks = List.intersperse (brk 1) fbreaks = List.intersperse fbrk blk :: (Int, [T]) -> T blk (indent, es) = Block Markup.empty False (force_nat indent) es block :: [T] -> T block prts = blk (2, prts) strs :: BYTES a => [a] -> T strs = block . breaks . map str markup :: Markup.T -> [T] -> T markup m = Block m False 0 mark :: Markup.T -> T -> T mark m prt = if m == Markup.empty then prt else markup m [prt] mark_str :: BYTES a => (Markup.T, a) -> T mark_str (m, s) = mark m (str s) marks_str :: BYTES a => ([Markup.T], a) -> T marks_str (ms, s) = fold_rev mark ms (str s) item :: [T] -> T item = markup Markup.item text_fold :: [T] -> T text_fold = markup Markup.text_fold keyword1, keyword2 :: BYTES a => a -> T keyword1 name = mark_str (Markup.keyword1, name) keyword2 name = mark_str (Markup.keyword2, name) text :: BYTES a => a -> [T] text = breaks . map str . filter (not . Bytes.null) . space_explode ' ' . make_bytes paragraph :: [T] -> T paragraph = markup Markup.paragraph para :: BYTES a => a -> T para = paragraph . text quote :: T -> T quote prt = blk (1, [Str "\"", prt, Str "\""]) cartouche :: T -> T cartouche prt = blk (1, [Str "\92", prt, Str "\92"]) separate :: BYTES a => a -> [T] -> [T] separate sep = List.intercalate [str sep, brk 1] . map single commas :: [T] -> [T] commas = separate ("," :: Bytes) enclose :: BYTES a => a -> a -> [T] -> T enclose lpar rpar prts = block (str lpar : prts <> [str rpar]) enum :: BYTES a => a -> a -> a -> [T] -> T enum sep lpar rpar = enclose lpar rpar . separate sep list :: BYTES a => a -> a -> [T] -> T list = enum "," str_list :: BYTES a => a -> a -> [a] -> T str_list lpar rpar = list lpar rpar . map str big_list :: BYTES a => a -> [T] -> T big_list name prts = block (fbreaks (str name : prts)) \ generate_file "Isabelle/Name.hs" = \ {- Title: Isabelle/Name.hs Author: Makarius Names of basic logical entities (variables etc.). See \<^file>\$ISABELLE_HOME/src/Pure/name.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Name ( Name, uu, uu_, aT, clean_index, clean, internal, skolem, is_internal, is_skolem, dest_internal, dest_skolem, Context, declare, declare_renaming, is_declared, declared, context, make_context, variant, variant_list ) where import Data.Maybe (fromMaybe) import Data.Map.Strict (Map) import qualified Data.Map.Strict as Map import Data.Word (Word8) import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.Symbol as Symbol import Isabelle.Library type Name = Bytes {- common defaults -} uu, uu_, aT :: Name uu = "uu" uu_ = "uu_" aT = "'a" {- internal names -- NB: internal subsumes skolem -} underscore :: Word8 underscore = Bytes.byte '_' internal, skolem :: Name -> Name internal x = x <> "_" skolem x = x <> "__" is_internal, is_skolem :: Name -> Bool is_internal = Bytes.isSuffixOf "_" is_skolem = Bytes.isSuffixOf "__" dest_internal, dest_skolem :: Name -> Maybe Name dest_internal = Bytes.try_unsuffix "_" dest_skolem = Bytes.try_unsuffix "__" clean_index :: (Name, Int) -> (Name, Int) clean_index (x, i) = case dest_internal x of Nothing -> (x, i) Just x' -> clean_index (x', i + 1) clean :: Name -> Name clean x = fst (clean_index (x, 0)) {- context for used names -} newtype Context = Context (Map Name (Maybe Name)) {-declared names with latest renaming-} declare :: Name -> Context -> Context declare x (Context names) = Context ( let a = clean x in if Map.member a names then names else Map.insert a Nothing names) declare_renaming :: (Name, Name) -> Context -> Context declare_renaming (x, x') (Context names) = Context (Map.insert (clean x) (Just (clean x')) names) is_declared :: Context -> Name -> Bool is_declared (Context names) x = Map.member x names declared :: Context -> Name -> Maybe (Maybe Name) declared (Context names) a = Map.lookup a names context :: Context context = Context Map.empty |> fold declare ["", "'"] make_context :: [Name] -> Context make_context used = fold declare used context {- generating fresh names -} bump_init :: Name -> Name bump_init str = str <> "a" bump_string :: Name -> Name bump_string str = let a = Bytes.byte 'a' z = Bytes.byte 'z' bump (b : bs) | b == z = a : bump bs bump (b : bs) | a <= b && b < z = b + 1 : bs bump bs = a : bs rev = reverse (Bytes.unpack str) part2 = reverse (takeWhile (Symbol.is_ascii_quasi . Bytes.char) rev) part1 = reverse (bump (drop (length part2) rev)) in Bytes.pack (part1 <> part2) variant :: Name -> Context -> (Name, Context) variant name ctxt = let vary x = case declared ctxt x of Nothing -> x Just x' -> vary (bump_string (fromMaybe x x')) (x, n) = clean_index (name, 0) (x', ctxt') = if not (is_declared ctxt x) then (x, declare x ctxt) else let x0 = bump_init x x' = vary x0 ctxt' = ctxt |> (if x0 /= x' then declare_renaming (x0, x') else id) |> declare x' in (x', ctxt') in (x' <> Bytes.pack (replicate n underscore), ctxt') variant_list :: [Name] -> [Name] -> [Name] variant_list used names = fst (make_context used |> fold_map variant names) \ generate_file "Isabelle/Term.hs" = \ {- Title: Isabelle/Term.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Lambda terms, types, sorts. See \<^file>\$ISABELLE_HOME/src/Pure/term.scala\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Term ( Indexname, Sort, Typ(..), Term(..), Free, lambda, declare_frees, incr_boundvars, subst_bound, dest_lambda, strip_lambda, type_op0, type_op1, op0, op1, op2, typed_op0, typed_op1, typed_op2, binder, dummyS, dummyT, is_dummyT, propT, is_propT, (-->), dest_funT, (--->), aconv, list_comb, strip_comb, head_of ) where import Isabelle.Library import qualified Isabelle.Name as Name import Isabelle.Name (Name) infixr 5 --> infixr ---> {- types and terms -} type Indexname = (Name, Int) type Sort = [Name] data Typ = Type (Name, [Typ]) | TFree (Name, Sort) | TVar (Indexname, Sort) deriving (Show, Eq, Ord) data Term = Const (Name, [Typ]) | Free (Name, Typ) | Var (Indexname, Typ) | Bound Int | Abs (Name, Typ, Term) | App (Term, Term) deriving (Show, Eq, Ord) {- free and bound variables -} type Free = (Name, Typ) lambda :: Free -> Term -> Term lambda (name, typ) body = Abs (name, typ, abstract 0 body) where abstract lev (Free (x, ty)) | name == x && typ == ty = Bound lev abstract lev (Abs (a, ty, t)) = Abs (a, ty, abstract (lev + 1) t) abstract lev (App (t, u)) = App (abstract lev t, abstract lev u) abstract _ t = t declare_frees :: Term -> Name.Context -> Name.Context declare_frees (Free (x, _)) = Name.declare x declare_frees (Abs (_, _, b)) = declare_frees b declare_frees (App (t, u)) = declare_frees t #> declare_frees u declare_frees _ = id incr_boundvars :: Int -> Term -> Term incr_boundvars inc = if inc == 0 then id else incr 0 where incr lev (Bound i) = if i >= lev then Bound (i + inc) else Bound i incr lev (Abs (a, ty, b)) = Abs (a, ty, incr (lev + 1) b) incr lev (App (t, u)) = App (incr lev t, incr lev u) incr _ t = t subst_bound :: Term -> Term -> Term subst_bound arg = subst 0 where subst lev (Bound i) = if i < lev then Bound i else if i == lev then incr_boundvars lev arg else Bound (i - 1) subst lev (Abs (a, ty, b)) = Abs (a, ty, subst (lev + 1) b) subst lev (App (t, u)) = App (subst lev t, subst lev u) subst _ t = t dest_lambda :: Name.Context -> Term -> Maybe (Free, Term) dest_lambda names (Abs (x, ty, b)) = let (x', _) = Name.variant x (declare_frees b names) v = (x', ty) in Just (v, subst_bound (Free v) b) dest_lambda _ _ = Nothing strip_lambda :: Name.Context -> Term -> ([Free], Term) strip_lambda names tm = case dest_lambda names tm of Just (v, t) -> let (vs, t') = strip_lambda names t' in (v : vs, t') Nothing -> ([], tm) {- type and term operators -} type_op0 :: Name -> (Typ, Typ -> Bool) type_op0 name = (mk, is) where mk = Type (name, []) is (Type (c, _)) = c == name is _ = False type_op1 :: Name -> (Typ -> Typ, Typ -> Maybe Typ) type_op1 name = (mk, dest) where mk ty = Type (name, [ty]) dest (Type (c, [ty])) | c == name = Just ty dest _ = Nothing type_op2 :: Name -> (Typ -> Typ -> Typ, Typ -> Maybe (Typ, Typ)) type_op2 name = (mk, dest) where mk ty1 ty2 = Type (name, [ty1, ty2]) dest (Type (c, [ty1, ty2])) | c == name = Just (ty1, ty2) dest _ = Nothing op0 :: Name -> (Term, Term -> Bool) op0 name = (mk, is) where mk = Const (name, []) is (Const (c, _)) = c == name is _ = False op1 :: Name -> (Term -> Term, Term -> Maybe Term) op1 name = (mk, dest) where mk t = App (Const (name, []), t) dest (App (Const (c, _), t)) | c == name = Just t dest _ = Nothing op2 :: Name -> (Term -> Term -> Term, Term -> Maybe (Term, Term)) op2 name = (mk, dest) where mk t u = App (App (Const (name, []), t), u) dest (App (App (Const (c, _), t), u)) | c == name = Just (t, u) dest _ = Nothing typed_op0 :: Name -> (Typ -> Term, Term -> Maybe Typ) typed_op0 name = (mk, dest) where mk ty = Const (name, [ty]) dest (Const (c, [ty])) | c == name = Just ty dest _ = Nothing typed_op1 :: Name -> (Typ -> Term -> Term, Term -> Maybe (Typ, Term)) typed_op1 name = (mk, dest) where mk ty t = App (Const (name, [ty]), t) dest (App (Const (c, [ty]), t)) | c == name = Just (ty, t) dest _ = Nothing typed_op2 :: Name -> (Typ -> Term -> Term -> Term, Term -> Maybe (Typ, Term, Term)) typed_op2 name = (mk, dest) where mk ty t u = App (App (Const (name, [ty]), t), u) dest (App (App (Const (c, [ty]), t), u)) | c == name = Just (ty, t, u) dest _ = Nothing binder :: Name -> (Free -> Term -> Term, Name.Context -> Term -> Maybe (Free, Term)) binder name = (mk, dest) where mk (a, ty) b = App (Const (name, [ty]), lambda (a, ty) b) dest names (App (Const (c, _), t)) | c == name = dest_lambda names t dest _ _ = Nothing {- type operations -} dummyS :: Sort dummyS = [""] dummyT :: Typ; is_dummyT :: Typ -> Bool (dummyT, is_dummyT) = type_op0 \\<^type_name>\dummy\\ propT :: Typ; is_propT :: Typ -> Bool (propT, is_propT) = type_op0 \\<^type_name>\prop\\ (-->) :: Typ -> Typ -> Typ; dest_funT :: Typ -> Maybe (Typ, Typ) ((-->), dest_funT) = type_op2 \\<^type_name>\fun\\ (--->) :: [Typ] -> Typ -> Typ [] ---> b = b (a : as) ---> b = a --> (as ---> b) {- term operations -} aconv :: Term -> Term -> Bool aconv (App (t1, u1)) (App (t2, u2)) = aconv t1 t2 && aconv u1 u2 aconv (Abs (_, ty1, t1)) (Abs (_, ty2, t2)) = aconv t1 t2 && ty1 == ty2 aconv a1 a2 = a1 == a2 list_comb :: Term -> [Term] -> Term list_comb f [] = f list_comb f (t : ts) = list_comb (App (f, t)) ts strip_comb :: Term -> (Term, [Term]) strip_comb tm = strip (tm, []) where strip (App (f, t), ts) = strip (f, t : ts) strip x = x head_of :: Term -> Term head_of (App (f, _)) = head_of f head_of u = u \ generate_file "Isabelle/Pure.hs" = \ {- Title: Isabelle/Term.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Support for Isabelle/Pure logic. See \<^file>\$ISABELLE_HOME/src/Pure/logic.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Pure ( mk_forall_op, dest_forall_op, mk_forall, dest_forall, mk_equals, dest_equals, mk_implies, dest_implies ) where import qualified Isabelle.Name as Name import Isabelle.Term mk_forall_op :: Typ -> Term -> Term; dest_forall_op :: Term -> Maybe (Typ, Term) (mk_forall_op, dest_forall_op) = typed_op1 \\<^const_name>\Pure.all\\ mk_forall :: Free -> Term -> Term; dest_forall :: Name.Context -> Term -> Maybe (Free, Term) (mk_forall, dest_forall) = binder \\<^const_name>\Pure.all\\ mk_equals :: Typ -> Term -> Term -> Term; dest_equals :: Term -> Maybe (Typ, Term, Term) (mk_equals, dest_equals) = typed_op2 \\<^const_name>\Pure.eq\\ mk_implies :: Term -> Term -> Term; dest_implies :: Term -> Maybe (Term, Term) (mk_implies, dest_implies) = op2 \\<^const_name>\Pure.imp\\ \ generate_file "Isabelle/HOL.hs" = \ {- Title: Isabelle/Term.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Support for Isabelle/HOL logic. See \<^file>\$ISABELLE_HOME/src/HOL/Tools/hologic.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.HOL ( boolT, is_boolT, mk_trueprop, dest_trueprop, mk_setT, dest_setT, mk_mem, dest_mem, mk_eq, dest_eq, true, is_true, false, is_false, mk_not, dest_not, mk_conj, dest_conj, mk_disj, dest_disj, mk_imp, dest_imp, mk_iff, dest_iff, mk_all_op, dest_all_op, mk_ex_op, dest_ex_op, mk_all, dest_all, mk_ex, dest_ex, mk_undefined, dest_undefined ) where import qualified Isabelle.Name as Name import Isabelle.Term boolT :: Typ; is_boolT :: Typ -> Bool (boolT, is_boolT) = type_op0 \\<^type_name>\bool\\ mk_trueprop :: Term -> Term; dest_trueprop :: Term -> Maybe Term (mk_trueprop, dest_trueprop) = op1 \\<^const_name>\Trueprop\\ mk_setT :: Typ -> Typ; dest_setT :: Typ -> Maybe Typ (mk_setT, dest_setT) = type_op1 \\<^type_name>\set\\ mk_mem :: Typ -> Term -> Term -> Term; dest_mem :: Term -> Maybe (Typ, Term, Term) (mk_mem, dest_mem) = typed_op2 \\<^const_name>\Set.member\\ mk_eq :: Typ -> Term -> Term -> Term; dest_eq :: Term -> Maybe (Typ, Term, Term) (mk_eq, dest_eq) = typed_op2 \\<^const_name>\HOL.eq\\ true :: Term; is_true :: Term -> Bool (true, is_true) = op0 \\<^const_name>\True\\ false :: Term; is_false :: Term -> Bool (false, is_false) = op0 \\<^const_name>\False\\ mk_not :: Term -> Term; dest_not :: Term -> Maybe Term (mk_not, dest_not) = op1 \\<^const_name>\Not\\ mk_conj :: Term -> Term -> Term; dest_conj :: Term -> Maybe (Term, Term) (mk_conj, dest_conj) = op2 \\<^const_name>\conj\\ mk_disj :: Term -> Term -> Term; dest_disj :: Term -> Maybe (Term, Term) (mk_disj, dest_disj) = op2 \\<^const_name>\disj\\ mk_imp :: Term -> Term -> Term; dest_imp :: Term -> Maybe (Term, Term) (mk_imp, dest_imp) = op2 \\<^const_name>\implies\\ mk_iff :: Term -> Term -> Term mk_iff = mk_eq boolT dest_iff :: Term -> Maybe (Term, Term) dest_iff tm = case dest_eq tm of Just (ty, t, u) | ty == boolT -> Just (t, u) _ -> Nothing mk_all_op :: Typ -> Term -> Term; dest_all_op :: Term -> Maybe (Typ, Term) (mk_all_op, dest_all_op) = typed_op1 \\<^const_name>\All\\ mk_ex_op :: Typ -> Term -> Term; dest_ex_op :: Term -> Maybe (Typ, Term) (mk_ex_op, dest_ex_op) = typed_op1 \\<^const_name>\Ex\\ mk_all :: Free -> Term -> Term; dest_all :: Name.Context -> Term -> Maybe (Free, Term) (mk_all, dest_all) = binder \\<^const_name>\All\\ mk_ex :: Free -> Term -> Term; dest_ex :: Name.Context -> Term -> Maybe (Free, Term) (mk_ex, dest_ex) = binder \\<^const_name>\Ex\\ mk_undefined :: Typ -> Term; dest_undefined :: Term -> Maybe Typ (mk_undefined, dest_undefined) = typed_op0 \\<^const_name>\undefined\\ \ generate_file "Isabelle/Term_XML/Encode.hs" = \ {- Title: Isabelle/Term_XML/Encode.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) XML data representation of lambda terms. See \<^file>\$ISABELLE_HOME/src/Pure/term_xml.ML\. -} {-# LANGUAGE LambdaCase #-} module Isabelle.Term_XML.Encode (indexname, sort, typ, typ_body, term) where import Isabelle.Library import Isabelle.XML.Encode import Isabelle.Term indexname :: P Indexname indexname (a, b) = if b == 0 then [a] else [a, int_atom b] sort :: T Sort sort = list string typ :: T Typ typ ty = ty |> variant [\case { Type (a, b) -> Just ([a], list typ b); _ -> Nothing }, \case { TFree (a, b) -> Just ([a], sort b); _ -> Nothing }, \case { TVar (a, b) -> Just (indexname a, sort b); _ -> Nothing }] typ_body :: T Typ typ_body ty = if is_dummyT ty then [] else typ ty term :: T Term term t = t |> variant [\case { Const (a, b) -> Just ([a], list typ b); _ -> Nothing }, \case { Free (a, b) -> Just ([a], typ_body b); _ -> Nothing }, \case { Var (a, b) -> Just (indexname a, typ_body b); _ -> Nothing }, \case { Bound a -> Just ([], int a); _ -> Nothing }, \case { Abs (a, b, c) -> Just ([a], pair typ term (b, c)); _ -> Nothing }, \case { App a -> Just ([], pair term term a); _ -> Nothing }] \ generate_file "Isabelle/Term_XML/Decode.hs" = \ {- Title: Isabelle/Term_XML/Decode.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) XML data representation of lambda terms. See \<^file>\$ISABELLE_HOME/src/Pure/term_xml.ML\. -} {-# OPTIONS_GHC -fno-warn-incomplete-patterns #-} module Isabelle.Term_XML.Decode (indexname, sort, typ, typ_body, term) where import Isabelle.Library import Isabelle.XML.Decode import Isabelle.Term indexname :: P Indexname indexname [a] = (a, 0) indexname [a, b] = (a, int_atom b) sort :: T Sort sort = list string typ :: T Typ typ ty = ty |> variant [\([a], b) -> Type (a, list typ b), \([a], b) -> TFree (a, sort b), \(a, b) -> TVar (indexname a, sort b)] typ_body :: T Typ typ_body [] = dummyT typ_body body = typ body term :: T Term term t = t |> variant [\([a], b) -> Const (a, list typ b), \([a], b) -> Free (a, typ_body b), \(a, b) -> Var (indexname a, typ_body b), \([], a) -> Bound (int a), \([a], b) -> let (c, d) = pair typ term b in Abs (a, c, d), \([], a) -> App (pair term term a)] \ generate_file "Isabelle/XML/Classes.hs" = \ {- generated by Isabelle -} {- Title: Isabelle/XML/Classes.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Type classes for XML data representation. -} {-# LANGUAGE FlexibleInstances #-} module Isabelle.XML.Classes (Encode_Atom(..), Decode_Atom(..), Encode (..), Decode (..)) where import qualified Isabelle.XML as XML import qualified Isabelle.XML.Encode as Encode import qualified Isabelle.XML.Decode as Decode import qualified Isabelle.Term_XML.Encode as Encode import qualified Isabelle.Term_XML.Decode as Decode import qualified Isabelle.Properties as Properties import Isabelle.Bytes (Bytes) import Isabelle.Term (Typ, Term) class Encode_Atom a where encode_atom :: Encode.A a class Decode_Atom a where decode_atom :: Decode.A a instance Encode_Atom Int where encode_atom = Encode.int_atom instance Decode_Atom Int where decode_atom = Decode.int_atom instance Encode_Atom Bool where encode_atom = Encode.bool_atom instance Decode_Atom Bool where decode_atom = Decode.bool_atom instance Encode_Atom () where encode_atom = Encode.unit_atom instance Decode_Atom () where decode_atom = Decode.unit_atom class Encode a where encode :: Encode.T a class Decode a where decode :: Decode.T a instance Encode Bytes where encode = Encode.string instance Decode Bytes where decode = Decode.string instance Encode Int where encode = Encode.int instance Decode Int where decode = Decode.int instance Encode Bool where encode = Encode.bool instance Decode Bool where decode = Decode.bool instance Encode () where encode = Encode.unit instance Decode () where decode = Decode.unit instance (Encode a, Encode b) => Encode (a, b) where encode = Encode.pair encode encode instance (Decode a, Decode b) => Decode (a, b) where decode = Decode.pair decode decode instance (Encode a, Encode b, Encode c) => Encode (a, b, c) where encode = Encode.triple encode encode encode instance (Decode a, Decode b, Decode c) => Decode (a, b, c) where decode = Decode.triple decode decode decode instance Encode a => Encode [a] where encode = Encode.list encode instance Decode a => Decode [a] where decode = Decode.list decode instance Encode a => Encode (Maybe a) where encode = Encode.option encode instance Decode a => Decode (Maybe a) where decode = Decode.option decode instance Encode XML.Tree where encode = Encode.tree instance Decode XML.Tree where decode = Decode.tree instance Encode Properties.T where encode = Encode.properties instance Decode Properties.T where decode = Decode.properties instance Encode Typ where encode = Encode.typ instance Decode Typ where decode = Decode.typ instance Encode Term where encode = Encode.term instance Decode Term where decode = Decode.term \ generate_file "Isabelle/UUID.hs" = \ {- Title: Isabelle/UUID.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Universally unique identifiers. See \<^file>\$ISABELLE_HOME/src/Pure/General/uuid.scala\. -} module Isabelle.UUID (T, random, print, parse) where import Prelude hiding (print) import Data.UUID (UUID) import qualified Data.UUID as UUID import Data.UUID.V4 (nextRandom) import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) type T = UUID random :: IO T random = nextRandom print :: T -> Bytes print = Bytes.make . UUID.toASCIIBytes parse :: Bytes -> Maybe T parse = UUID.fromASCIIBytes . Bytes.unmake \ generate_file "Isabelle/Byte_Message.hs" = \ {- Title: Isabelle/Byte_Message.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Byte-oriented messages. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/byte_message.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/PIDE/byte_message.scala\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-incomplete-patterns #-} module Isabelle.Byte_Message ( write, write_line, read, read_block, read_line, make_message, write_message, read_message, exchange_message, exchange_message0, make_line_message, write_line_message, read_line_message, read_yxml, write_yxml ) where import Prelude hiding (read) import Data.Maybe import qualified Data.ByteString as ByteString import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.Symbol as Symbol import qualified Isabelle.UTF8 as UTF8 import qualified Isabelle.XML as XML import qualified Isabelle.YXML as YXML import Network.Socket (Socket) import qualified Network.Socket.ByteString as Socket import Isabelle.Library import qualified Isabelle.Value as Value {- output operations -} write :: Socket -> [Bytes] -> IO () write socket = Socket.sendMany socket . map Bytes.unmake write_line :: Socket -> Bytes -> IO () write_line socket s = write socket [s, "\n"] {- input operations -} read :: Socket -> Int -> IO Bytes read socket n = read_body 0 [] where result = Bytes.concat . reverse read_body len ss = if len >= n then return (result ss) else (do s <- Socket.recv socket (min (n - len) 8192) case ByteString.length s of 0 -> return (result ss) m -> read_body (len + m) (Bytes.make s : ss)) read_block :: Socket -> Int -> IO (Maybe Bytes, Int) read_block socket n = do msg <- read socket n let len = Bytes.length msg return (if len == n then Just msg else Nothing, len) read_line :: Socket -> IO (Maybe Bytes) read_line socket = read_body [] where result = trim_line . Bytes.pack . reverse read_body bs = do s <- Socket.recv socket 1 case ByteString.length s of 0 -> return (if null bs then Nothing else Just (result bs)) 1 -> case ByteString.head s of 10 -> return (Just (result bs)) b -> read_body (b : bs) {- messages with multiple chunks (arbitrary content) -} make_header :: [Int] -> [Bytes] make_header ns = [space_implode "," (map Value.print_int ns), "\n"] make_message :: [Bytes] -> [Bytes] make_message chunks = make_header (map Bytes.length chunks) <> chunks write_message :: Socket -> [Bytes] -> IO () write_message socket = write socket . make_message parse_header :: Bytes -> [Int] parse_header line = let res = map Value.parse_nat (space_explode ',' line) in if all isJust res then map fromJust res else error ("Malformed message header: " <> quote (UTF8.decode line)) read_chunk :: Socket -> Int -> IO Bytes read_chunk socket n = do res <- read_block socket n return $ case res of (Just chunk, _) -> chunk (Nothing, len) -> error ("Malformed message chunk: unexpected EOF after " <> show len <> " of " <> show n <> " bytes") read_message :: Socket -> IO (Maybe [Bytes]) read_message socket = do res <- read_line socket case res of Just line -> Just <$> mapM (read_chunk socket) (parse_header line) Nothing -> return Nothing exchange_message :: Socket -> [Bytes] -> IO (Maybe [Bytes]) exchange_message socket msg = do write_message socket msg read_message socket exchange_message0 :: Socket -> [Bytes] -> IO () exchange_message0 socket msg = do _ <- exchange_message socket msg return () -- hybrid messages: line or length+block (with content restriction) is_length :: Bytes -> Bool is_length msg = not (Bytes.null msg) && Bytes.all_char (\c -> '0' <= c && c <= '9') msg is_terminated :: Bytes -> Bool is_terminated msg = not (Bytes.null msg) && Symbol.is_ascii_line_terminator (Bytes.char $ Bytes.last msg) make_line_message :: Bytes -> [Bytes] make_line_message msg = let n = Bytes.length msg in if is_length msg || is_terminated msg then error ("Bad content for line message:\n" <> take 100 (UTF8.decode msg)) else (if n > 100 || Bytes.any_char (== '\n') msg then make_header [n + 1] else []) <> [msg, "\n"] write_line_message :: Socket -> Bytes -> IO () write_line_message socket = write socket . make_line_message read_line_message :: Socket -> IO (Maybe Bytes) read_line_message socket = do opt_line <- read_line socket case opt_line of Nothing -> return Nothing Just line -> case Value.parse_nat line of Nothing -> return $ Just line Just n -> fmap trim_line . fst <$> read_block socket n read_yxml :: Socket -> IO (Maybe XML.Body) read_yxml socket = do res <- read_line_message socket return (YXML.parse_body <$> res) write_yxml :: Socket -> XML.Body -> IO () write_yxml socket body = write_line_message socket (YXML.string_of_body body) \ generate_file "Isabelle/Isabelle_Thread.hs" = \ {- Title: Isabelle/Isabelle_Thread.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Isabelle-specific thread management. See \<^file>\$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.scala\. -} {-# LANGUAGE NamedFieldPuns #-} module Isabelle.Isabelle_Thread ( ThreadId, Result, find_id, properties, change_properties, add_resource, del_resource, bracket_resource, is_stopped, expose_stopped, stop, my_uuid, stop_uuid, Fork, fork_finally, fork) where import Data.Unique import Data.IORef import System.IO.Unsafe import qualified Data.List as List import Control.Monad (when, forM_) import Data.Map.Strict (Map) import qualified Data.Map.Strict as Map import Control.Exception as Exception import Control.Concurrent (ThreadId) import qualified Control.Concurrent as Concurrent import Control.Concurrent.Thread (Result) import qualified Control.Concurrent.Thread as Thread import qualified Isabelle.UUID as UUID import qualified Isabelle.Properties as Properties {- thread info -} type Resources = Map Unique (IO ()) data Info = Info {uuid :: UUID.T, props :: Properties.T, stopped :: Bool, resources :: Resources} type Infos = Map ThreadId Info lookup_info :: Infos -> ThreadId -> Maybe Info lookup_info infos id = Map.lookup id infos init_info :: ThreadId -> UUID.T -> Infos -> (Infos, ()) init_info id uuid infos = (Map.insert id (Info uuid [] False Map.empty) infos, ()) {- global state -} {-# NOINLINE global_state #-} global_state :: IORef Infos global_state = unsafePerformIO (newIORef Map.empty) find_id :: UUID.T -> IO (Maybe ThreadId) find_id uuid = do state <- readIORef global_state return $ fst <$> List.find (\(_, Info{uuid = uuid'}) -> uuid == uuid') (Map.assocs state) get_info :: ThreadId -> IO (Maybe Info) get_info id = do state <- readIORef global_state return $ lookup_info state id map_info :: ThreadId -> (Info -> Info) -> IO (Maybe Info) map_info id f = atomicModifyIORef' global_state (\infos -> case lookup_info infos id of Nothing -> (infos, Nothing) Just info -> let info' = f info in (Map.insert id info' infos, Just info')) delete_info :: ThreadId -> IO () delete_info id = atomicModifyIORef' global_state (\infos -> (Map.delete id infos, ())) {- thread properties -} my_info :: IO (Maybe Info) my_info = do id <- Concurrent.myThreadId get_info id properties :: IO Properties.T properties = maybe [] props <$> my_info change_properties :: (Properties.T -> Properties.T) -> IO () change_properties f = do id <- Concurrent.myThreadId map_info id (\info -> info {props = f (props info)}) return () {- managed resources -} add_resource :: IO () -> IO Unique add_resource resource = do id <- Concurrent.myThreadId u <- newUnique map_info id (\info -> info {resources = Map.insert u resource (resources info)}) return u del_resource :: Unique -> IO () del_resource u = do id <- Concurrent.myThreadId map_info id (\info -> info {resources = Map.delete u (resources info)}) return () bracket_resource :: IO () -> IO a -> IO a bracket_resource resource body = Exception.bracket (add_resource resource) del_resource (const body) {- stop -} is_stopped :: IO Bool is_stopped = maybe False stopped <$> my_info expose_stopped :: IO () expose_stopped = do stopped <- is_stopped when stopped $ throw ThreadKilled stop :: ThreadId -> IO () stop id = do info <- map_info id (\info -> info {stopped = True}) let ops = case info of Nothing -> []; Just Info{resources} -> map snd (Map.toDescList resources) sequence_ ops {- UUID -} my_uuid :: IO (Maybe UUID.T) my_uuid = fmap uuid <$> my_info stop_uuid :: UUID.T -> IO () stop_uuid uuid = do id <- find_id uuid forM_ id stop {- fork -} type Fork a = (ThreadId, UUID.T, IO (Result a)) fork_finally :: IO a -> (Either SomeException a -> IO b) -> IO (Fork b) fork_finally body finally = do uuid <- UUID.random (id, result) <- Exception.mask (\restore -> Thread.forkIO (Exception.try (do id <- Concurrent.myThreadId atomicModifyIORef' global_state (init_info id uuid) restore body) >>= (\res -> do id <- Concurrent.myThreadId; delete_info id; finally res))) return (id, uuid, result) fork :: IO a -> IO (Fork a) fork body = fork_finally body Thread.result \ generate_file "Isabelle/Server.hs" = \ {- Title: Isabelle/Server.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) TCP server on localhost. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Server ( localhost_name, localhost_prefix, localhost, publish_text, publish_stdout, server, connection ) where import Control.Monad (forever, when) import qualified Control.Exception as Exception import Network.Socket (Socket) import qualified Network.Socket as Socket import qualified System.IO as IO import qualified Data.ByteString.Char8 as Char8 import Isabelle.Library import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.UUID as UUID import qualified Isabelle.Byte_Message as Byte_Message import qualified Isabelle.Isabelle_Thread as Isabelle_Thread {- server address -} localhost_name :: Bytes localhost_name = "127.0.0.1" localhost_prefix :: Bytes localhost_prefix = localhost_name <> ":" localhost :: Socket.HostAddress localhost = Socket.tupleToHostAddress (127, 0, 0, 1) publish_text :: Bytes -> Bytes -> UUID.T -> Bytes publish_text name address password = "server " <> quote name <> " = " <> address <> " (password " <> quote (show_bytes password) <> ")" publish_stdout :: Bytes -> Bytes -> UUID.T -> IO () publish_stdout name address password = Char8.putStrLn (Bytes.unmake $ publish_text name address password) {- server -} server :: (Bytes -> UUID.T -> IO ()) -> (Socket -> IO ()) -> IO () server publish handle = Socket.withSocketsDo $ Exception.bracket open (Socket.close . fst) (uncurry loop) where open :: IO (Socket, Bytes) open = do server_socket <- Socket.socket Socket.AF_INET Socket.Stream Socket.defaultProtocol Socket.bind server_socket (Socket.SockAddrInet 0 localhost) Socket.listen server_socket 50 port <- Socket.socketPort server_socket let address = localhost_name <> ":" <> show_bytes port password <- UUID.random publish address password return (server_socket, UUID.print password) loop :: Socket -> Bytes -> IO () loop server_socket password = forever $ do (connection, _) <- Socket.accept server_socket Isabelle_Thread.fork_finally (do line <- Byte_Message.read_line connection when (line == Just password) $ handle connection) (\finally -> do Socket.close connection case finally of Left exn -> IO.hPutStrLn IO.stderr $ Exception.displayException exn Right () -> return ()) return () {- client connection -} connection :: String -> Bytes -> (Socket -> IO a) -> IO a connection port password client = Socket.withSocketsDo $ do addr <- resolve Exception.bracket (open addr) Socket.close body where resolve = do let hints = Socket.defaultHints { Socket.addrFlags = [Socket.AI_NUMERICHOST, Socket.AI_NUMERICSERV], Socket.addrSocketType = Socket.Stream } head <$> Socket.getAddrInfo (Just hints) (Just $ make_string localhost_name) (Just port) open addr = do socket <- Socket.socket (Socket.addrFamily addr) (Socket.addrSocketType addr) (Socket.addrProtocol addr) Socket.connect socket $ Socket.addrAddress addr return socket body socket = do Byte_Message.write_line socket password client socket \ generate_file "Isabelle/Time.hs" = \ {- Title: Isabelle/Time.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Time based on milliseconds. See \<^file>\~~/src/Pure/General/time.scala\ -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Time ( Time, seconds, minutes, ms, zero, is_zero, is_relevant, get_seconds, get_minutes, get_ms, message, now ) where import Text.Printf (printf) import Data.Time.Clock.POSIX (getPOSIXTime) import Isabelle.Bytes (Bytes) import Isabelle.Library newtype Time = Time Int instance Eq Time where Time a == Time b = a == b instance Ord Time where compare (Time a) (Time b) = compare a b instance Num Time where fromInteger = Time . fromInteger Time a + Time b = Time (a + b) Time a - Time b = Time (a - b) Time a * Time b = Time (a * b) abs (Time a) = Time (abs a) signum (Time a) = Time (signum a) seconds :: Double -> Time seconds s = Time (round (s * 1000.0)) minutes :: Double -> Time minutes m = Time (round (m * 60000.0)) ms :: Int -> Time ms = Time zero :: Time zero = ms 0 is_zero :: Time -> Bool is_zero (Time ms) = ms == 0 is_relevant :: Time -> Bool is_relevant (Time ms) = ms >= 1 get_seconds :: Time -> Double get_seconds (Time ms) = fromIntegral ms / 1000.0 get_minutes :: Time -> Double get_minutes (Time ms) = fromIntegral ms / 60000.0 get_ms :: Time -> Int get_ms (Time ms) = ms instance Show Time where show t = printf "%.3f" (get_seconds t) message :: Time -> Bytes message t = make_bytes (show t) <> "s" now :: IO Time now = do t <- getPOSIXTime return $ Time (round (realToFrac t * 1000.0 :: Double)) \ generate_file "Isabelle/Timing.hs" = \ {- Title: Isabelle/Timing.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Support for time measurement. See \<^file>\~~/src/Pure/General/timing.ML\ and \<^file>\~~/src/Pure/General/timing.scala\ -} module Isabelle.Timing ( Timing (..), zero, is_zero, is_relevant ) where import qualified Isabelle.Time as Time import Isabelle.Time (Time) data Timing = Timing {elapsed :: Time, cpu :: Time, gc :: Time} deriving (Show, Eq) zero :: Timing zero = Timing Time.zero Time.zero Time.zero is_zero :: Timing -> Bool is_zero t = Time.is_zero (elapsed t) && Time.is_zero (cpu t) && Time.is_zero (gc t) is_relevant :: Timing -> Bool is_relevant t = Time.is_relevant (elapsed t) || Time.is_relevant (cpu t) || Time.is_relevant (gc t) \ generate_file "Isabelle/Bash.hs" = \ {- Title: Isabelle/Bash.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Support for GNU bash. See \<^file>\$ISABELLE_HOME/src/Pure/System/bash.ML\ -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Bash ( string, strings, Params, get_script, get_input, get_cwd, get_putenv, get_redirect, get_timeout, get_description, script, input, cwd, putenv, redirect, timeout, description, server_run, server_kill, server_uuid, server_interrupt, server_failure, server_result ) where import Text.Printf (printf) import qualified Isabelle.Symbol as Symbol import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.Time as Time import Isabelle.Time (Time) import Isabelle.Library {- concrete syntax -} string :: Bytes -> Bytes string str = if Bytes.null str then "\"\"" else str |> Bytes.unpack |> map trans |> Bytes.concat where trans b = case Bytes.char b of '\t' -> "$'\\t'" '\n' -> "$'\\n'" '\f' -> "$'\\f'" '\r' -> "$'\\r'" c -> if Symbol.is_ascii_letter c || Symbol.is_ascii_digit c || c `elem` ("+,-./:_" :: String) then Bytes.singleton b else if b < 32 || b >= 127 then make_bytes (printf "$'\\x%02x'" b :: String) else "\\" <> Bytes.singleton b strings :: [Bytes] -> Bytes strings = space_implode " " . map string {- server parameters -} data Params = Params { _script :: Bytes, _input :: Bytes, _cwd :: Maybe Bytes, _putenv :: [(Bytes, Bytes)], _redirect :: Bool, _timeout :: Time, _description :: Bytes} deriving (Show, Eq) get_script :: Params -> Bytes get_script = _script get_input :: Params -> Bytes get_input = _input get_cwd :: Params -> Maybe Bytes get_cwd = _cwd get_putenv :: Params -> [(Bytes, Bytes)] get_putenv = _putenv get_redirect :: Params -> Bool get_redirect = _redirect get_timeout :: Params -> Time get_timeout = _timeout get_description :: Params -> Bytes get_description = _description script :: Bytes -> Params script script = Params script "" Nothing [] False Time.zero "" input :: Bytes -> Params -> Params input input params = params { _input = input } cwd :: Bytes -> Params -> Params cwd cwd params = params { _cwd = Just cwd } putenv :: [(Bytes, Bytes)] -> Params -> Params putenv putenv params = params { _putenv = putenv } redirect :: Params -> Params redirect params = params { _redirect = True } timeout :: Time -> Params -> Params timeout timeout params = params { _timeout = timeout } description :: Bytes -> Params -> Params description description params = params { _description = description } {- server messages -} server_run, server_kill :: Bytes server_run = \Bash.server_run\; server_kill = \Bash.server_kill\; server_uuid, server_interrupt, server_failure, server_result :: Bytes server_uuid = \Bash.server_uuid\; server_interrupt = \Bash.server_interrupt\; server_failure = \Bash.server_failure\; server_result = \Bash.server_result\; \ generate_file "Isabelle/Process_Result.hs" = \ {- Title: Isabelle/Process_Result.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Result of system process. See \<^file>\~~/src/Pure/System/process_result.ML\ and \<^file>\~~/src/Pure/System/process_result.scala\ -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Process_Result ( ok_rc, error_rc, failure_rc, interrupt_rc , timeout_rc, T, make, rc, out_lines, err_lines, timing, timing_elapsed, out, err, ok, check ) where import Isabelle.Time (Time) import qualified Isabelle.Timing as Timing import Isabelle.Timing (Timing) import Isabelle.Bytes (Bytes) import Isabelle.Library ok_rc, error_rc, failure_rc, interrupt_rc , timeout_rc :: Int ok_rc = 0 error_rc = 1 failure_rc = 2 interrupt_rc = 130 timeout_rc = 142 data T = Process_Result { _rc :: Int, _out_lines :: [Bytes], _err_lines :: [Bytes], _timing :: Timing} deriving (Show, Eq) make :: Int -> [Bytes] -> [Bytes] -> Timing -> T make = Process_Result rc :: T -> Int rc = _rc out_lines :: T -> [Bytes] out_lines = _out_lines err_lines :: T -> [Bytes] err_lines = _err_lines timing :: T -> Timing timing = _timing timing_elapsed :: T -> Time timing_elapsed = Timing.elapsed . timing out :: T -> Bytes out = trim_line . cat_lines . out_lines err :: T -> Bytes err = trim_line . cat_lines . err_lines ok :: T -> Bool ok result = rc result == ok_rc check :: T -> T check result = if ok result then result else error (make_string $ err result) \ generate_file "Isabelle/Options.hs" = \ {- Title: Isabelle/Options.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) System options with external string representation. See \<^file>\~~/src/Pure/System/options.ML\ and \<^file>\~~/src/Pure/System/options.scala\ -} {-# LANGUAGE OverloadedStrings #-} {-# LANGUAGE InstanceSigs #-} module Isabelle.Options ( boolT, intT, realT, stringT, unknownT, T, typ, bool, int, real, seconds, string, decode ) where import qualified Data.Map.Strict as Map import Data.Map.Strict (Map) import qualified Isabelle.Properties as Properties import Isabelle.Bytes (Bytes) import qualified Isabelle.Value as Value import qualified Isabelle.Time as Time import Isabelle.Time (Time) import Isabelle.Library import qualified Isabelle.XML.Decode as Decode import Isabelle.XML.Classes (Decode (..)) {- representation -} boolT :: Bytes boolT = "bool" intT :: Bytes intT = "int" realT :: Bytes realT = "real" stringT :: Bytes stringT = "string" unknownT :: Bytes unknownT = "unknown" data Opt = Opt { _pos :: Properties.T, _name :: Bytes, _typ :: Bytes, _value :: Bytes } newtype T = Options (Map Bytes Opt) {- check -} check_name :: T -> Bytes -> Opt check_name (Options map) name = case Map.lookup name map of Just opt | _typ opt /= unknownT -> opt _ -> error (make_string ("Unknown system option " <> quote name)) check_type :: T -> Bytes -> Bytes -> Opt check_type options name typ = let opt = check_name options name t = _typ opt in if t == typ then opt else error (make_string ("Ill-typed system option " <> quote name <> " : " <> t <> " vs. " <> typ)) {- get typ -} typ :: T -> Bytes -> Bytes typ options name = _typ (check_name options name) {- get value -} get :: Bytes -> (Bytes -> Maybe a) -> T -> Bytes -> a get typ parse options name = let opt = check_type options name typ in case parse (_value opt) of Just x -> x Nothing -> error (make_string ("Malformed value for system option " <> quote name <> " : " <> typ <> " =\n" <> quote (_value opt))) bool :: T -> Bytes -> Bool bool = get boolT Value.parse_bool int :: T -> Bytes -> Int int = get intT Value.parse_int real :: T -> Bytes -> Double real = get realT Value.parse_real seconds :: T -> Bytes -> Time seconds options = Time.seconds . real options string :: T -> Bytes -> Bytes string = get stringT Just {- decode -} instance Decode T where decode :: Decode.T T decode = let decode_entry :: Decode.T (Bytes, Opt) decode_entry body = let (pos, (name, (typ, value))) = Decode.pair Decode.properties (Decode.pair Decode.string (Decode.pair Decode.string Decode.string)) body in (name, Opt { _pos = pos, _name = name, _typ = typ, _value = value }) in Options . Map.fromList . Decode.list decode_entry \ generate_file "Isabelle/Isabelle_System.hs" = \ {- Title: Isabelle/Isabelle_System.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Isabelle system support. See \<^file>\~~/src/Pure/System/isabelle_system.ML\ and \<^file>\~~/src/Pure/System/isabelle_system.scala\ -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Isabelle_System ( bash_process, bash_process0 ) where import Data.Maybe (fromMaybe) import Control.Exception (throw, AsyncException (UserInterrupt)) import Network.Socket (Socket) import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.Byte_Message as Byte_Message import qualified Isabelle.Time as Time import Isabelle.Timing (Timing (..)) import qualified Isabelle.Options as Options import qualified Isabelle.Bash as Bash import qualified Isabelle.Process_Result as Process_Result import qualified Isabelle.XML.Encode as Encode import qualified Isabelle.YXML as YXML import qualified Isabelle.Value as Value import qualified Isabelle.Server as Server import qualified Isabelle.Isabelle_Thread as Isabelle_Thread import Isabelle.Library {- bash_process -} absolute_path :: Bytes -> Bytes -- FIXME dummy absolute_path = id bash_process :: Options.T -> Bash.Params -> IO Process_Result.T bash_process options = bash_process0 address password where address = Options.string options "bash_process_address" password = Options.string options "bash_process_password" bash_process0 :: Bytes -> Bytes -> Bash.Params -> IO Process_Result.T bash_process0 address password params = do Server.connection port password (\socket -> do isabelle_tmp <- getenv "ISABELLE_TMP" Byte_Message.write_message socket (run isabelle_tmp) loop Nothing socket) where port = case Bytes.try_unprefix Server.localhost_prefix address of Just port -> make_string port Nothing -> errorWithoutStackTrace "Bad bash_process server address" script = Bash.get_script params input = Bash.get_input params cwd = Bash.get_cwd params putenv = Bash.get_putenv params redirect = Bash.get_redirect params timeout = Bash.get_timeout params description = Bash.get_description params run :: Bytes -> [Bytes] run isabelle_tmp = [Bash.server_run, script, input, YXML.string_of_body (Encode.option (Encode.string . absolute_path) cwd), YXML.string_of_body (Encode.list (Encode.pair Encode.string Encode.string) (("ISABELLE_TMP", isabelle_tmp) : putenv)), Value.print_bool redirect, Value.print_real (Time.get_seconds timeout), description] kill :: Maybe Bytes -> IO () kill maybe_uuid = do case maybe_uuid of Just uuid -> Server.connection port password (\socket -> Byte_Message.write_message socket [Bash.server_kill, uuid]) Nothing -> return () err = errorWithoutStackTrace "Malformed result from bash_process server" the = fromMaybe err loop :: Maybe Bytes -> Socket -> IO Process_Result.T loop maybe_uuid socket = do result <- Isabelle_Thread.bracket_resource (kill maybe_uuid) (Byte_Message.read_message socket) case result of Just [head, uuid] | head == Bash.server_uuid -> loop (Just uuid) socket Just [head] | head == Bash.server_interrupt -> throw UserInterrupt Just [head, msg] | head == Bash.server_failure -> errorWithoutStackTrace $ make_string msg Just (head : a : b : c : d : lines) | head == Bash.server_result -> let rc = the $ Value.parse_int a elapsed = Time.ms $ the $ Value.parse_int b cpu = Time.ms $ the $ Value.parse_int c timing = Timing elapsed cpu Time.zero n = the $ Value.parse_int d out_lines = take n lines err_lines = drop n lines in return $ Process_Result.make rc out_lines err_lines timing _ -> err \ generate_file "Isabelle/Cache.hs" = \ {- Title: Isabelle/Cache.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Cache for slow computations. -} module Isabelle.Cache ( T, init, apply, prune ) where import Prelude hiding (init) import Data.IORef import Data.Map.Strict (Map) import qualified Data.Map.Strict as Map import qualified Data.List as List import Isabelle.Time (Time) import qualified Isabelle.Time as Time data Entry v = Entry {_value :: v, _access :: Time, _timing :: Time} newtype T k v = Cache (IORef (Map k (Entry v))) init :: IO (T k v) init = Cache <$> newIORef Map.empty commit :: Ord k => T k v -> k -> Entry v -> IO v commit (Cache ref) x e = do atomicModifyIORef' ref (\entries -> let entry = case Map.lookup x entries of Just e' | _access e' > _access e -> e' _ -> e in (Map.insert x entry entries, _value entry)) apply :: Ord k => T k v -> k -> IO v -> IO v apply cache@(Cache ref) x body = do start <- Time.now entries <- readIORef ref case Map.lookup x entries of Just entry -> do commit cache x (entry {_access = start}) Nothing -> do y <- body stop <- Time.now commit cache x (Entry y start (stop - start)) prune :: Ord k => T k v -> Int -> Time -> IO () prune (Cache ref) max_size min_timing = do atomicModifyIORef' ref (\entries -> let trim x e = if _timing e < min_timing then Map.delete x else id sort = List.sortBy (\(_, e1) (_, e2) -> compare (_access e2) (_access e1)) entries1 = Map.foldrWithKey trim entries entries entries2 = if Map.size entries1 <= max_size then entries1 else Map.fromList $ List.take max_size $ sort $ Map.toList entries1 in (entries2, ())) \ export_generated_files _ end diff --git a/src/ZF/Tools/induct_tacs.ML b/src/ZF/Tools/induct_tacs.ML --- a/src/ZF/Tools/induct_tacs.ML +++ b/src/ZF/Tools/induct_tacs.ML @@ -1,202 +1,202 @@ (* Title: ZF/Tools/induct_tacs.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Induction and exhaustion tactics for Isabelle/ZF. The theory information needed to support them (and to support primrec). Also a function to install other sets as if they were datatypes. *) signature DATATYPE_TACTICS = sig val exhaust_tac: Proof.context -> string -> (binding * string option * mixfix) list -> int -> tactic val induct_tac: Proof.context -> string -> (binding * string option * mixfix) list -> int -> tactic val rep_datatype_i: thm -> thm -> thm list -> thm list -> theory -> theory val rep_datatype: Facts.ref * Token.src list -> Facts.ref * Token.src list -> (Facts.ref * Token.src list) list -> (Facts.ref * Token.src list) list -> theory -> theory end; (** Datatype information, e.g. associated theorems **) type datatype_info = {inductive: bool, (*true if inductive, not coinductive*) constructors : term list, (*the constructors, as Consts*) rec_rewrites : thm list, (*recursor equations*) case_rewrites : thm list, (*case equations*) induct : thm, mutual_induct : thm, exhaustion : thm}; structure DatatypesData = Theory_Data ( type T = datatype_info Symtab.table; val empty = Symtab.empty; fun merge data : T = Symtab.merge (K true) data; ); (** Constructor information: needed to map constructors to datatypes **) type constructor_info = {big_rec_name : string, (*name of the mutually recursive set*) constructors : term list, (*the constructors, as Consts*) free_iffs : thm list, (*freeness simprules*) rec_rewrites : thm list}; (*recursor equations*) structure ConstructorsData = Theory_Data ( type T = constructor_info Symtab.table val empty = Symtab.empty fun merge data = Symtab.merge (K true) data; ); structure DatatypeTactics : DATATYPE_TACTICS = struct fun datatype_info thy name = (case Symtab.lookup (DatatypesData.get thy) name of SOME info => info | NONE => error ("Unknown datatype " ^ quote name)); (*Given a variable, find the inductive set associated it in the assumptions*) exception Find_tname of string fun find_tname ctxt var As = let fun mk_pair \<^Const_>\mem for \Free (v,_)\ A\ = (v, #1 (dest_Const (head_of A))) | mk_pair _ = raise Match val pairs = map_filter (try (mk_pair o \<^dest_judgment>)) As val x = (case try (dest_Free o Syntax.read_term ctxt) var of SOME (x, _) => x | _ => raise Find_tname ("Bad variable " ^ quote var)) in case AList.lookup (op =) pairs x of NONE => raise Find_tname ("Cannot determine datatype of " ^ quote var) | SOME t => t end; (** generic exhaustion and induction tactic for datatypes Differences from HOL: (1) no checking if the induction var occurs in premises, since it always appears in one of them, and it's hard to check for other occurrences (2) exhaustion works for VARIABLES in the premises, not general terms **) fun exhaust_induct_tac exh ctxt var fixes i state = SUBGOAL (fn _ => let val thy = Proof_Context.theory_of ctxt val ({context = ctxt', asms, ...}, _) = Subgoal.focus ctxt i NONE state val tn = find_tname ctxt' var (map Thm.term_of asms) val rule = datatype_info thy tn |> (if exh then #exhaustion else #induct) |> Thm.transfer thy; val \<^Const_>\mem for \Var(ixn,_)\ _\ = (case Thm.prems_of rule of [] => error "induction is not available for this datatype" | major::_ => \<^dest_judgment> major) in Rule_Insts.eres_inst_tac ctxt [((ixn, Position.none), var)] fixes rule i end handle Find_tname msg => if exh then (*try boolean case analysis instead*) case_tac ctxt var fixes i else error msg) i state; val exhaust_tac = exhaust_induct_tac true; val induct_tac = exhaust_induct_tac false; (**** declare non-datatype as datatype ****) fun rep_datatype_i elim induct case_eqns recursor_eqns thy = let (*analyze the LHS of a case equation to get a constructor*) fun const_of \<^Const_>\IFOL.eq _ for \_ $ c\ _\ = c | const_of eqn = error ("Ill-formed case equation: " ^ Syntax.string_of_term_global thy eqn); val constructors = map (head_of o const_of o \<^dest_judgment> o Thm.prop_of) case_eqns; val \<^Const_>\mem for _ data\ = \<^dest_judgment> (hd (Thm.prems_of elim)); val Const(big_rec_name, _) = head_of data; val simps = case_eqns @ recursor_eqns; val dt_info = {inductive = true, constructors = constructors, rec_rewrites = map Thm.trim_context recursor_eqns, case_rewrites = map Thm.trim_context case_eqns, induct = Thm.trim_context induct, mutual_induct = Thm.trim_context @{thm TrueI}, (*No need for mutual induction*) exhaustion = Thm.trim_context elim}; val con_info = {big_rec_name = big_rec_name, constructors = constructors, (*let primrec handle definition by cases*) free_iffs = [], (*thus we expect the necessary freeness rewrites to be in the simpset already, as is the case for Nat and disjoint sum*) rec_rewrites = (case recursor_eqns of [] => case_eqns | _ => recursor_eqns) |> map Thm.trim_context}; (*associate with each constructor the datatype name and rewrites*) val con_pairs = map (fn c => (#1 (dest_Const c), con_info)) constructors in thy |> Sign.add_path (Long_Name.base_name big_rec_name) |> Global_Theory.add_thmss [((Binding.name "simps", simps), [Simplifier.simp_add])] |> snd - |> DatatypesData.put (Symtab.update (big_rec_name, dt_info) (DatatypesData.get thy)) - |> ConstructorsData.put (fold_rev Symtab.update con_pairs (ConstructorsData.get thy)) + |> DatatypesData.map (Symtab.update (big_rec_name, dt_info)) + |> ConstructorsData.map (fold_rev Symtab.update con_pairs) |> Sign.parent_path end; fun rep_datatype raw_elim raw_induct raw_case_eqns raw_recursor_eqns thy = let val ctxt = Proof_Context.init_global thy; val elim = Facts.the_single ("elimination", Position.none) (Attrib.eval_thms ctxt [raw_elim]); val induct = Facts.the_single ("induction", Position.none) (Attrib.eval_thms ctxt [raw_induct]); val case_eqns = Attrib.eval_thms ctxt raw_case_eqns; val recursor_eqns = Attrib.eval_thms ctxt raw_recursor_eqns; in rep_datatype_i elim induct case_eqns recursor_eqns thy end; (* theory setup *) val _ = Theory.setup (Method.setup \<^binding>\induct_tac\ (Args.goal_spec -- Scan.lift (Parse.embedded -- Parse.for_fixes) >> (fn (quant, (s, xs)) => fn ctxt => SIMPLE_METHOD'' quant (induct_tac ctxt s xs))) "induct_tac emulation (dynamic instantiation!)" #> Method.setup \<^binding>\case_tac\ (Args.goal_spec -- Scan.lift (Parse.embedded -- Parse.for_fixes) >> (fn (quant, (s, xs)) => fn ctxt => SIMPLE_METHOD'' quant (exhaust_tac ctxt s xs))) "datatype case_tac emulation (dynamic instantiation!)"); (* outer syntax *) val _ = Outer_Syntax.command \<^command_keyword>\rep_datatype\ "represent existing set inductively" ((\<^keyword>\elimination\ |-- Parse.!!! Parse.thm) -- (\<^keyword>\induction\ |-- Parse.!!! Parse.thm) -- (\<^keyword>\case_eqns\ |-- Parse.!!! Parse.thms1) -- Scan.optional (\<^keyword>\recursor_eqns\ |-- Parse.!!! Parse.thms1) [] >> (fn (((x, y), z), w) => Toplevel.theory (rep_datatype x y z w))); end; val exhaust_tac = DatatypeTactics.exhaust_tac; val induct_tac = DatatypeTactics.induct_tac; diff --git a/src/ZF/Tools/inductive_package.ML b/src/ZF/Tools/inductive_package.ML --- a/src/ZF/Tools/inductive_package.ML +++ b/src/ZF/Tools/inductive_package.ML @@ -1,609 +1,610 @@ (* Title: ZF/Tools/inductive_package.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Fixedpoint definition module -- for Inductive/Coinductive Definitions The functor will be instantiated for normal sums/products (inductive defs) and non-standard sums/products (coinductive defs) Sums are used only for mutual recursion; Products are used only to derive "streamlined" induction rules for relations *) type inductive_result = {defs : thm list, (*definitions made in thy*) bnd_mono : thm, (*monotonicity for the lfp definition*) dom_subset : thm, (*inclusion of recursive set in dom*) intrs : thm list, (*introduction rules*) elim : thm, (*case analysis theorem*) induct : thm, (*main induction rule*) mutual_induct : thm}; (*mutual induction rule*) (*Functor's result signature*) signature INDUCTIVE_PACKAGE = sig (*Insert definitions for the recursive sets, which must *already* be declared as constants in parent theory!*) val add_inductive_i: bool -> term list * term -> ((binding * term) * attribute list) list -> thm list * thm list * thm list * thm list -> theory -> theory * inductive_result val add_inductive: string list * string -> ((binding * string) * Token.src list) list -> (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list -> theory -> theory * inductive_result end; (*Declares functions to add fixedpoint/constructor defs to a theory. Recursive sets must *already* be declared as constants.*) functor Add_inductive_def_Fun (structure Fp: FP and Pr : PR and CP: CARTPROD and Su : SU val coind: bool) : INDUCTIVE_PACKAGE = struct val co_prefix = if coind then "co" else ""; (* utils *) (*make distinct individual variables a1, a2, a3, ..., an. *) fun mk_frees a [] = [] | mk_frees a (T::Ts) = Free(a,T) :: mk_frees (Symbol.bump_string a) Ts; (* add_inductive(_i) *) (*internal version, accepting terms*) fun add_inductive_i verbose (rec_tms, dom_sum) raw_intr_specs (monos, con_defs, type_intrs, type_elims) thy0 = let val ctxt0 = Proof_Context.init_global thy0; val intr_specs = map (apfst (apfst Binding.name_of)) raw_intr_specs; val (intr_names, intr_tms) = split_list (map fst intr_specs); val case_names = Rule_Cases.case_names intr_names; (*recT and rec_params should agree for all mutually recursive components*) val rec_hds = map head_of rec_tms; val dummy = rec_hds |> forall (fn t => is_Const t orelse error ("Recursive set not previously declared as constant: " ^ Syntax.string_of_term ctxt0 t)); (*Now we know they are all Consts, so get their names, type and params*) val rec_names = map (#1 o dest_Const) rec_hds and (Const(_,recT),rec_params) = strip_comb (hd rec_tms); val rec_base_names = map Long_Name.base_name rec_names; val dummy = rec_base_names |> forall (fn a => Symbol_Pos.is_identifier a orelse error ("Base name of recursive set not an identifier: " ^ a)); local (*Checking the introduction rules*) val intr_sets = map (#2 o Ind_Syntax.rule_concl_msg thy0) intr_tms; fun intr_ok set = case head_of set of Const(a,recT) => member (op =) rec_names a | _ => false; in val dummy = intr_sets |> forall (fn t => intr_ok t orelse error ("Conclusion of rule does not name a recursive set: " ^ Syntax.string_of_term ctxt0 t)); end; val dummy = rec_params |> forall (fn t => is_Free t orelse error ("Param in recursion term not a free variable: " ^ Syntax.string_of_term ctxt0 t)); (*** Construct the fixedpoint definition ***) val mk_variant = singleton (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] intr_tms)); val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w"; fun dest_tprop \<^Const_>\Trueprop for P\ = P | dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^ Syntax.string_of_term ctxt0 Q); (*Makes a disjunct from an introduction rule*) fun fp_part intr = (*quantify over rule's free vars except parameters*) let val prems = map dest_tprop (Logic.strip_imp_prems intr) val dummy = List.app (fn rec_hd => List.app (Ind_Syntax.chk_prem rec_hd) prems) rec_hds val exfrees = subtract (op =) rec_params (Misc_Legacy.term_frees intr) val zeq = FOLogic.mk_eq (Free(z', \<^Type>\i\), #1 (Ind_Syntax.rule_concl intr)) in fold_rev (FOLogic.mk_exists o Term.dest_Free) exfrees (Balanced_Tree.make FOLogic.mk_conj (zeq::prems)) end; (*The Part(A,h) terms -- compose injections to make h*) fun mk_Part (Bound 0) = Free(X', \<^Type>\i\) (*no mutual rec, no Part needed*) | mk_Part h = \<^Const>\Part\ $ Free(X', \<^Type>\i\) $ Abs (w', \<^Type>\i\, h); (*Access to balanced disjoint sums via injections*) val parts = map mk_Part (Balanced_Tree.accesses {left = fn t => Su.inl $ t, right = fn t => Su.inr $ t, init = Bound 0} (length rec_tms)); (*replace each set by the corresponding Part(A,h)*) val part_intrs = map (subst_free (rec_tms ~~ parts) o fp_part) intr_tms; val fp_abs = absfree (X', \<^Type>\i\) (Ind_Syntax.mk_Collect (z', dom_sum, Balanced_Tree.make FOLogic.mk_disj part_intrs)); val fp_rhs = Fp.oper $ dom_sum $ fp_abs val dummy = List.app (fn rec_hd => (Logic.occs (rec_hd, fp_rhs) andalso error "Illegal occurrence of recursion operator"; ())) rec_hds; (*** Make the new theory ***) (*A key definition: If no mutual recursion then it equals the one recursive set. If mutual recursion then it differs from all the recursive sets. *) val big_rec_base_name = space_implode "_" rec_base_names; val big_rec_name = Proof_Context.intern_const ctxt0 big_rec_base_name; val _ = if verbose then writeln ((if coind then "Coind" else "Ind") ^ "uctive definition " ^ quote big_rec_name) else (); (*Big_rec... is the union of the mutually recursive sets*) val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); (*The individual sets must already be declared*) val axpairs = map Misc_Legacy.mk_defpair ((big_rec_tm, fp_rhs) :: (case parts of [_] => [] (*no mutual recursion*) | _ => rec_tms ~~ (*define the sets as Parts*) map (subst_atomic [(Free (X', \<^Type>\i\), big_rec_tm)]) parts)); (*tracing: print the fixedpoint definition*) val dummy = if !Ind_Syntax.trace then writeln (cat_lines (map (Syntax.string_of_term ctxt0 o #2) axpairs)) else () (*add definitions of the inductive sets*) val (_, thy1) = thy0 |> Sign.add_path big_rec_base_name |> Global_Theory.add_defs false (map (Thm.no_attributes o apfst Binding.name) axpairs); (*fetch fp definitions from the theory*) val big_rec_def::part_rec_defs = map (Misc_Legacy.get_def thy1) (case rec_names of [_] => rec_names | _ => big_rec_base_name::rec_names); (********) val dummy = writeln " Proving monotonicity..."; val bnd_mono0 = Goal.prove_global thy1 [] [] (\<^make_judgment> (Fp.bnd_mono $ dom_sum $ fp_abs)) (fn {context = ctxt, ...} => EVERY [resolve_tac ctxt [@{thm Collect_subset} RS @{thm bnd_monoI}] 1, REPEAT (ares_tac ctxt (@{thms basic_monos} @ monos) 1)]); val dom_subset0 = Drule.export_without_context (big_rec_def RS Fp.subs); val ([bnd_mono, dom_subset], thy2) = thy1 |> Global_Theory.add_thms [((Binding.name "bnd_mono", bnd_mono0), []), ((Binding.name "dom_subset", dom_subset0), [])]; val unfold = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.Tarski); (********) val dummy = writeln " Proving the introduction rules..."; (*Mutual recursion? Helps to derive subset rules for the individual sets.*) val Part_trans = case rec_names of [_] => asm_rl | _ => Drule.export_without_context (@{thm Part_subset} RS @{thm subset_trans}); (*To type-check recursive occurrences of the inductive sets, possibly enclosed in some monotonic operator M.*) val rec_typechecks = [dom_subset] RL (asm_rl :: ([Part_trans] RL monos)) RL [@{thm subsetD}]; (*Type-checking is hardest aspect of proof; disjIn selects the correct disjunct after unfolding*) fun intro_tacsf disjIn ctxt = [DETERM (stac ctxt unfold 1), REPEAT (resolve_tac ctxt [@{thm Part_eqI}, @{thm CollectI}] 1), (*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*) resolve_tac ctxt [disjIn] 2, (*Not ares_tac, since refl must be tried before equality assumptions; backtracking may occur if the premises have extra variables!*) DEPTH_SOLVE_1 (resolve_tac ctxt [@{thm refl}, @{thm exI}, @{thm conjI}] 2 APPEND assume_tac ctxt 2), (*Now solve the equations like Tcons(a,f) = Inl(?b4)*) rewrite_goals_tac ctxt con_defs, REPEAT (resolve_tac ctxt @{thms refl} 2), (*Typechecking; this can fail*) if !Ind_Syntax.trace then print_tac ctxt "The type-checking subgoal:" else all_tac, REPEAT (FIRSTGOAL (dresolve_tac ctxt rec_typechecks ORELSE' eresolve_tac ctxt (asm_rl :: @{thm PartE} :: @{thm SigmaE2} :: type_elims) ORELSE' hyp_subst_tac ctxt)), if !Ind_Syntax.trace then print_tac ctxt "The subgoal after monos, type_elims:" else all_tac, DEPTH_SOLVE (swap_res_tac ctxt (@{thm SigmaI} :: @{thm subsetI} :: type_intrs) 1)]; (*combines disjI1 and disjI2 to get the corresponding nested disjunct...*) val mk_disj_rls = Balanced_Tree.accesses {left = fn rl => rl RS @{thm disjI1}, right = fn rl => rl RS @{thm disjI2}, init = @{thm asm_rl}}; val intrs0 = (intr_tms, map intro_tacsf (mk_disj_rls (length intr_tms))) |> ListPair.map (fn (t, tacs) => Goal.prove_global thy2 [] [] t (fn {context = ctxt, ...} => EVERY (rewrite_goals_tac ctxt part_rec_defs :: tacs ctxt))); val ([intrs], thy3) = thy2 |> Global_Theory.add_thmss [((Binding.name "intros", intrs0), [])]; val ctxt3 = Proof_Context.init_global thy3; (********) val dummy = writeln " Proving the elimination rule..."; (*Breaks down logical connectives in the monotonic function*) fun basic_elim_tac ctxt = REPEAT (SOMEGOAL (eresolve_tac ctxt (Ind_Syntax.elim_rls @ Su.free_SEs) ORELSE' bound_hyp_subst_tac ctxt)) THEN prune_params_tac ctxt (*Mutual recursion: collapse references to Part(D,h)*) THEN (PRIMITIVE (fold_rule ctxt part_rec_defs)); (*Elimination*) val (elim, thy4) = thy3 |> Global_Theory.add_thm ((Binding.name "elim", rule_by_tactic ctxt3 (basic_elim_tac ctxt3) (unfold RS Ind_Syntax.equals_CollectD)), []); + val elim' = Thm.trim_context elim; val ctxt4 = Proof_Context.init_global thy4; (*Applies freeness of the given constructors, which *must* be unfolded by the given defs. Cannot simply use the local con_defs because con_defs=[] for inference systems. Proposition A should have the form t:Si where Si is an inductive set*) fun make_cases ctxt A = rule_by_tactic ctxt (basic_elim_tac ctxt THEN ALLGOALS (asm_full_simp_tac ctxt) THEN basic_elim_tac ctxt) - (Thm.assume A RS elim) + (Thm.assume A RS Thm.transfer' ctxt elim') |> Drule.export_without_context_open; fun induction_rules raw_induct = let val dummy = writeln " Proving the induction rule..."; (*** Prove the main induction rule ***) val pred_name = "P"; (*name for predicate variables*) (*Used to make induction rules; ind_alist = [(rec_tm1,pred1),...] associates predicates with rec ops prem is a premise of an intr rule*) fun add_induct_prem ind_alist (prem as \<^Const_>\Trueprop for \<^Const_>\mem for t X\\, iprems) = (case AList.lookup (op aconv) ind_alist X of SOME pred => prem :: \<^make_judgment> (pred $ t) :: iprems | NONE => (*possibly membership in M(rec_tm), for M monotone*) let fun mk_sb (rec_tm,pred) = (rec_tm, \<^Const>\Collect\ $ rec_tm $ pred) in subst_free (map mk_sb ind_alist) prem :: iprems end) | add_induct_prem ind_alist (prem,iprems) = prem :: iprems; (*Make a premise of the induction rule.*) fun induct_prem ind_alist intr = let val xs = subtract (op =) rec_params (Misc_Legacy.term_frees intr) val iprems = List.foldr (add_induct_prem ind_alist) [] (Logic.strip_imp_prems intr) val (t,X) = Ind_Syntax.rule_concl intr val (SOME pred) = AList.lookup (op aconv) ind_alist X val concl = \<^make_judgment> (pred $ t) in fold_rev Logic.all xs (Logic.list_implies (iprems,concl)) end handle Bind => error"Recursion term not found in conclusion"; (*Minimizes backtracking by delivering the correct premise to each goal. Intro rules with extra Vars in premises still cause some backtracking *) fun ind_tac _ [] 0 = all_tac | ind_tac ctxt (prem::prems) i = DEPTH_SOLVE_1 (ares_tac ctxt [prem, @{thm refl}] i) THEN ind_tac ctxt prems (i-1); val pred = Free(pred_name, \<^Type>\i\ --> \<^Type>\o\); val ind_prems = map (induct_prem (map (rpair pred) rec_tms)) intr_tms; val dummy = if ! Ind_Syntax.trace then writeln (cat_lines (["ind_prems:"] @ map (Syntax.string_of_term ctxt4) ind_prems @ ["raw_induct:", Thm.string_of_thm ctxt4 raw_induct])) else (); (*We use a MINIMAL simpset. Even FOL_ss contains too many simpules. If the premises get simplified, then the proofs could fail.*) val min_ss = (empty_simpset ctxt4 |> Simplifier.set_mksimps (fn ctxt => map mk_eq o ZF_atomize o Variable.gen_all ctxt)) setSolver (mk_solver "minimal" (fn ctxt => resolve_tac ctxt (triv_rls @ Simplifier.prems_of ctxt) ORELSE' assume_tac ctxt ORELSE' eresolve_tac ctxt @{thms FalseE})); val quant_induct = Goal.prove_global thy4 [] ind_prems (\<^make_judgment> (Ind_Syntax.mk_all_imp (big_rec_tm, pred))) (fn {context = ctxt, prems} => EVERY [rewrite_goals_tac ctxt part_rec_defs, resolve_tac ctxt [@{thm impI} RS @{thm allI}] 1, DETERM (eresolve_tac ctxt [raw_induct] 1), (*Push Part inside Collect*) full_simp_tac (min_ss addsimps [@{thm Part_Collect}]) 1, (*This CollectE and disjE separates out the introduction rules*) REPEAT (FIRSTGOAL (eresolve_tac ctxt [@{thm CollectE}, @{thm disjE}])), (*Now break down the individual cases. No disjE here in case some premise involves disjunction.*) REPEAT (FIRSTGOAL (eresolve_tac ctxt [@{thm CollectE}, @{thm exE}, @{thm conjE}] ORELSE' (bound_hyp_subst_tac ctxt))), ind_tac ctxt (rev (map (rewrite_rule ctxt part_rec_defs) prems)) (length prems)]); val dummy = if ! Ind_Syntax.trace then writeln ("quant_induct:\n" ^ Thm.string_of_thm ctxt4 quant_induct) else (); (*** Prove the simultaneous induction rule ***) (*Make distinct predicates for each inductive set*) (*The components of the element type, several if it is a product*) val elem_type = CP.pseudo_type dom_sum; val elem_factors = CP.factors elem_type; val elem_frees = mk_frees "za" elem_factors; val elem_tuple = CP.mk_tuple Pr.pair elem_type elem_frees; (*Given a recursive set and its domain, return the "fsplit" predicate and a conclusion for the simultaneous induction rule. NOTE. This will not work for mutually recursive predicates. Previously a summand 'domt' was also an argument, but this required the domain of mutual recursion to invariably be a disjoint sum.*) fun mk_predpair rec_tm = let val rec_name = (#1 o dest_Const o head_of) rec_tm val pfree = Free(pred_name ^ "_" ^ Long_Name.base_name rec_name, elem_factors ---> \<^Type>\o\) val qconcl = fold_rev (FOLogic.mk_all o Term.dest_Free) elem_frees \<^Const>\imp for \<^Const>\mem for elem_tuple rec_tm\ \list_comb (pfree, elem_frees)\\ in (CP.ap_split elem_type \<^Type>\o\ pfree, qconcl) end; val (preds,qconcls) = split_list (map mk_predpair rec_tms); (*Used to form simultaneous induction lemma*) fun mk_rec_imp (rec_tm,pred) = \<^Const>\imp for \<^Const>\mem for \Bound 0\ rec_tm\ \pred $ Bound 0\\; (*To instantiate the main induction rule*) val induct_concl = \<^make_judgment> (Ind_Syntax.mk_all_imp (big_rec_tm, Abs("z", \<^Type>\i\, Balanced_Tree.make FOLogic.mk_conj (ListPair.map mk_rec_imp (rec_tms, preds))))) and mutual_induct_concl = \<^make_judgment> (Balanced_Tree.make FOLogic.mk_conj qconcls); val dummy = if !Ind_Syntax.trace then (writeln ("induct_concl = " ^ Syntax.string_of_term ctxt4 induct_concl); writeln ("mutual_induct_concl = " ^ Syntax.string_of_term ctxt4 mutual_induct_concl)) else (); fun lemma_tac ctxt = FIRST' [eresolve_tac ctxt [@{thm asm_rl}, @{thm conjE}, @{thm PartE}, @{thm mp}], resolve_tac ctxt [@{thm allI}, @{thm impI}, @{thm conjI}, @{thm Part_eqI}], dresolve_tac ctxt [@{thm spec}, @{thm mp}, Pr.fsplitD]]; val need_mutual = length rec_names > 1; val lemma = (*makes the link between the two induction rules*) if need_mutual then (writeln " Proving the mutual induction rule..."; Goal.prove_global thy4 [] [] (Logic.mk_implies (induct_concl, mutual_induct_concl)) (fn {context = ctxt, ...} => EVERY [rewrite_goals_tac ctxt part_rec_defs, REPEAT (rewrite_goals_tac ctxt [Pr.split_eq] THEN lemma_tac ctxt 1)])) else (writeln " [ No mutual induction rule needed ]"; @{thm TrueI}); val dummy = if ! Ind_Syntax.trace then writeln ("lemma: " ^ Thm.string_of_thm ctxt4 lemma) else (); (*Mutual induction follows by freeness of Inl/Inr.*) (*Simplification largely reduces the mutual induction rule to the standard rule*) val mut_ss = min_ss addsimps [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff]; val all_defs = con_defs @ part_rec_defs; (*Removes Collects caused by M-operators in the intro rules. It is very hard to simplify list({v: tf. (v : t --> P_t(v)) & (v : f --> P_f(v))}) where t==Part(tf,Inl) and f==Part(tf,Inr) to list({v: tf. P_t(v)}). Instead the following rules extract the relevant conjunct. *) val cmonos = [@{thm subset_refl} RS @{thm Collect_mono}] RL monos RLN (2,[@{thm rev_subsetD}]); (*Minimizes backtracking by delivering the correct premise to each goal*) fun mutual_ind_tac _ [] 0 = all_tac | mutual_ind_tac ctxt (prem::prems) i = DETERM (SELECT_GOAL ( (*Simplify the assumptions and goal by unfolding Part and using freeness of the Sum constructors; proves all but one conjunct by contradiction*) rewrite_goals_tac ctxt all_defs THEN simp_tac (mut_ss addsimps [@{thm Part_iff}]) 1 THEN IF_UNSOLVED (*simp_tac may have finished it off!*) ((*simplify assumptions*) (*some risk of excessive simplification here -- might have to identify the bare minimum set of rewrites*) full_simp_tac (mut_ss addsimps @{thms conj_simps} @ @{thms imp_simps} @ @{thms quant_simps}) 1 THEN (*unpackage and use "prem" in the corresponding place*) REPEAT (resolve_tac ctxt @{thms impI} 1) THEN resolve_tac ctxt [rewrite_rule ctxt all_defs prem] 1 THEN (*prem must not be REPEATed below: could loop!*) DEPTH_SOLVE (FIRSTGOAL (ares_tac ctxt [@{thm impI}] ORELSE' eresolve_tac ctxt (@{thm conjE} :: @{thm mp} :: cmonos)))) ) i) THEN mutual_ind_tac ctxt prems (i-1); val mutual_induct_fsplit = if need_mutual then Goal.prove_global thy4 [] (map (induct_prem (rec_tms~~preds)) intr_tms) mutual_induct_concl (fn {context = ctxt, prems} => EVERY [resolve_tac ctxt [quant_induct RS lemma] 1, mutual_ind_tac ctxt (rev prems) (length prems)]) else @{thm TrueI}; (** Uncurrying the predicate in the ordinary induction rule **) (*instantiate the variable to a tuple, if it is non-trivial, in order to allow the predicate to be "opened up". The name "x.1" comes from the "RS spec" !*) val inst = case elem_frees of [_] => I | _ => Drule.instantiate_normalize (TVars.empty, Vars.make1 ((("x", 1), \<^Type>\i\), Thm.global_cterm_of thy4 elem_tuple)); (*strip quantifier and the implication*) val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp})); val \<^Const_>\Trueprop for \pred_var $ _\\ = Thm.concl_of induct0 val induct0 = CP.split_rule_var ctxt4 (pred_var, elem_type --> \<^Type>\o\, induct0) |> Drule.export_without_context and mutual_induct = CP.remove_split ctxt4 mutual_induct_fsplit val ([induct, mutual_induct], thy5) = thy4 |> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), induct0), [case_names, Induct.induct_pred big_rec_name]), ((Binding.name "mutual_induct", mutual_induct), [case_names])]; in ((induct, mutual_induct), thy5) end; (*of induction_rules*) val raw_induct = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.induct) val ((induct, mutual_induct), thy5) = if not coind then induction_rules raw_induct else thy4 |> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), raw_induct), [])] |> apfst (hd #> pair @{thm TrueI}); val (([cases], [defs]), thy6) = thy5 |> IndCases.declare big_rec_name make_cases |> Global_Theory.add_thms [((Binding.name "cases", elim), [case_names, Induct.cases_pred big_rec_name])] ||>> (Global_Theory.add_thmss o map Thm.no_attributes) [(Binding.name "defs", big_rec_def :: part_rec_defs)]; val (named_intrs, thy7) = thy6 |> Global_Theory.add_thms ((map Binding.name intr_names ~~ intrs) ~~ map #2 intr_specs) ||> Sign.parent_path; in (thy7, {defs = defs, bnd_mono = bnd_mono, dom_subset = dom_subset, intrs = named_intrs, elim = cases, induct = induct, mutual_induct = mutual_induct}) end; (*source version*) fun add_inductive (srec_tms, sdom_sum) intr_srcs (raw_monos, raw_con_defs, raw_type_intrs, raw_type_elims) thy = let val ctxt = Proof_Context.init_global thy; val read_terms = map (Syntax.parse_term ctxt #> Type.constraint \<^Type>\i\) #> Syntax.check_terms ctxt; val intr_atts = map (map (Attrib.attribute_cmd ctxt) o snd) intr_srcs; val sintrs = map fst intr_srcs ~~ intr_atts; val rec_tms = read_terms srec_tms; val dom_sum = singleton read_terms sdom_sum; val intr_tms = Syntax.read_props ctxt (map (snd o fst) sintrs); val intr_specs = (map (fst o fst) sintrs ~~ intr_tms) ~~ map snd sintrs; val monos = Attrib.eval_thms ctxt raw_monos; val con_defs = Attrib.eval_thms ctxt raw_con_defs; val type_intrs = Attrib.eval_thms ctxt raw_type_intrs; val type_elims = Attrib.eval_thms ctxt raw_type_elims; in thy |> add_inductive_i true (rec_tms, dom_sum) intr_specs (monos, con_defs, type_intrs, type_elims) end; (* outer syntax *) fun mk_ind (((((doms, intrs), monos), con_defs), type_intrs), type_elims) = #1 o add_inductive doms (map (fn ((x, y), z) => ((x, z), y)) intrs) (monos, con_defs, type_intrs, type_elims); val ind_decl = (\<^keyword>\domains\ |-- Parse.!!! (Parse.enum1 "+" Parse.term -- ((\<^keyword>\\\ || \<^keyword>\<=\) |-- Parse.term))) -- (\<^keyword>\intros\ |-- Parse.!!! (Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop))) -- Scan.optional (\<^keyword>\monos\ |-- Parse.!!! Parse.thms1) [] -- Scan.optional (\<^keyword>\con_defs\ |-- Parse.!!! Parse.thms1) [] -- Scan.optional (\<^keyword>\type_intros\ |-- Parse.!!! Parse.thms1) [] -- Scan.optional (\<^keyword>\type_elims\ |-- Parse.!!! Parse.thms1) [] >> (Toplevel.theory o mk_ind); val _ = Outer_Syntax.command (if coind then \<^command_keyword>\coinductive\ else \<^command_keyword>\inductive\) ("define " ^ co_prefix ^ "inductive sets") ind_decl; end;