diff --git a/thys/Topological_Semantics/boolean_algebra.thy b/thys/Topological_Semantics/boolean_algebra.thy --- a/thys/Topological_Semantics/boolean_algebra.thy +++ b/thys/Topological_Semantics/boolean_algebra.thy @@ -1,203 +1,203 @@ theory boolean_algebra imports Main begin (*-----------------------------------*) text\Technical configuration\ declare[[smt_timeout=30]] declare[[show_types]] (* declare[[syntax_ambiguity_warning=false]] *) sledgehammer_params[isar_proof=false] nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3, atoms=a b c d] (*default Nitpick settings*) text\We hide some Isabelle/HOL notation from the libraries (which we don't use) to avoid overloading\ hide_const(open) List.list.Nil no_notation List.list.Nil ("[]") (*We have no use for lists... *) hide_const(open) Relation.converse no_notation Relation.converse ("(_\)" [1000] 999) (*..nor for relations in this work*) hide_const(open) Fun.comp no_notation Fun.comp (infixl "\" 55) (*we redefine function composition below*) hide_const(open) Groups.plus_class.plus no_notation Groups.plus_class.plus (infixl "+" 65) (*we don't use this*) hide_const(open) Groups.times_class.times no_notation Groups.times_class.times (infixl "*" 70) (*we don't use this*) hide_const(open) Groups.minus_class.minus no_notation Groups.minus_class.minus (infixl "-" 65) (*we don't use this*) hide_const(open) Groups.uminus_class.uminus no_notation Groups.uminus_class.uminus ("- _" [81] 80) (*we don't use this*) (*-----------------------------------*) section \Shallow semantical embedding of (a logic of) Boolean algebras\ text\We encode Boolean algebras via their (Stone) representation as algebras of sets ('fields of sets'). This means that each element of (the carrier of) the algebra will be a set of 'points'. Inspired by the 'propositions as sets of worlds' paradigm from modal logic, we may think of points as being 'worlds', and thus of the elements of our Boolean algebras as 'propositions'. Of course, this is just one among many possible interpretations, and nothing stops us from thinking of points as any other kind of object (e.g., they can be sets, functions, sets of functions, etc.). \ text\We utilize a particular naming convention: The type parameter 'w is employed for the -domain/universe of 'points'. We conveniently introduce the (parametric) type-alias ('w)\ -as shorthand for 'w\bool. Hence, the elements of our algebra are objects of type ('w)\, +domain/universe of 'points'. We conveniently introduce the (parametric) type-alias @{text "('w)\"} +as shorthand for @{text "'w\bool"}. Hence, the elements of our algebra are objects of type @{text "('w)\"}, and thus correspond to (characteristic functions of) sets of 'points'. Set-valued (resp. set-domain) functions are thus functions that have sets as their codomain (resp. domain), -they are basically anything with a (parametric) type 'a\('w)\ (resp. ('w)\\'a).\ +they are basically anything with a (parametric) type @{text "'a\('w)\"} (resp. @{text "('w)\\'a)"}.\ text\Type for (characteristic functions of) sets (of 'points')\ type_synonym 'w \ = \'w \ bool\ text\In the sequel, we will (try to) enforce the following naming convention:\ -text\(i) Upper-case latin letters (A, B, D, M, P, S, X, etc.) denote arbitrary sets (type ('w)\). +text\(i) Upper-case latin letters (A, B, D, M, P, S, X, etc.) denote arbitrary sets (type @{text "('w)\"}). We will employ lower-case letters (p, q, x, w, etc.) to denote variables playing the role of 'points'. In some contexts, the letters S and D will be employed to denote sets/domains of sets (of 'points').\ -text\(ii) Greek letters denote arbitrary set-valued functions (type 'a\('w)\). -We employ the letters \, \ and \ to denote arbitrary unary operations -(with type ('w)\\('w)\).\ +text\(ii) Greek letters denote arbitrary set-valued functions (type @{text "'a\('w)\"}). +We employ the letters @{text "\, \ and \"} to denote arbitrary unary operations +(with type @{text "('w)\\('w)\)"}.\ -text\(iii) Upper-case calligraphic letters (\, \, \, \, etc.) are reserved for unary operations that are +text\(iii) Upper-case calligraphic letters @{text "(\, \, \, \, etc.)"} are reserved for unary operations that are intended to act as 'topological operators' in the given context.\ subsection \Encoding Boolean operations\ text\Standard inclusion-based order structure on sets.\ definition subset::"'w \ \ 'w \ \ bool" (infixr "\<^bold>\" 45) where "A \<^bold>\ B \ \p. A p \ B p" definition setequ::"'w \ \ 'w \ \ bool" (infixr "\<^bold>=" 45) where "A \<^bold>= B \ \p. A p \ B p" named_theorems order (*to group together order-related definitions*) declare setequ_def[order] subset_def[order] lemma subset_char1: "(A \<^bold>\ B) = (\C. B \<^bold>\ C \ A \<^bold>\ C)" by (metis subset_def) lemma subset_char2: "(A \<^bold>\ B) = (\C. C \<^bold>\ A \ C \<^bold>\ B)" by (metis subset_def) text\These (trivial) lemmas are intended to help automated tools.\ lemma setequ_char: "(A \<^bold>= B) = (A \<^bold>\ B \ B \<^bold>\ A)" unfolding order by blast lemma setequ_ext: "(A \<^bold>= B) = (A = B)" unfolding order by blast text\We now encode connectives for (distributive and complemented) bounded lattices, mostly by reusing their counterpart meta-logical HOL connectives.\ definition meet::"'w \ \ 'w \ \ 'w \" (infixr "\<^bold>\" 54) where "A \<^bold>\ B \ \p. (A p) \ (B p)" \\ intersection \ definition join::"'w \ \ 'w \ \ 'w \" (infixr "\<^bold>\" 53) where "A \<^bold>\ B \ \p. (A p) \ (B p)" \\ union \ definition top::"'w \" ("\<^bold>\") where "\<^bold>\ \ \w. True" \\ universe \ definition bottom::"'w \" ("\<^bold>\") where "\<^bold>\ \ \w. False" \\ empty-set \ text\And introduce further operations to obtain a Boolean algebra (of sets).\ definition compl::"'w \ \ 'w \" ("\<^bold>\_" [57]58) where "\<^bold>\A \ \p. \(A p)" \\ (set-)complement \ definition impl::"'w \ \ 'w \ \ 'w \" (infixr "\<^bold>\" 51) where "A \<^bold>\ B \ \p. (A p) \ (B p)" \\ (set-)implication \ definition diff::"'w \ \ 'w \ \ 'w \" (infixr "\<^bold>\" 51) where "A \<^bold>\ B \ \p. (A p) \ \(B p)" \\ (set-)difference \ definition dimpl::"'w \ \ 'w \ \ 'w \" (infixr "\<^bold>\" 51) where "A \<^bold>\ B \ \p. (A p) = (B p)" \\ double implication \ definition sdiff::"'w \ \ 'w \ \ 'w \" (infixr "\<^bold>\" 51) where "A \<^bold>\ B \ \p. (A p) \ (B p)" \\ symmetric difference (aka. xor) \ named_theorems conn (*to group together definitions for algebraic connectives*) declare meet_def[conn] join_def[conn] top_def[conn] bottom_def[conn] impl_def[conn] dimpl_def[conn] diff_def[conn] sdiff_def[conn] compl_def[conn] text\Verify characterization for some connectives.\ lemma compl_char: "\<^bold>\A = A \<^bold>\ \<^bold>\" unfolding conn by simp lemma impl_char: "A \<^bold>\ B = \<^bold>\A \<^bold>\ B" unfolding conn by simp lemma dimpl_char: "A \<^bold>\ B = (A \<^bold>\ B) \<^bold>\ (B \<^bold>\ A)" unfolding conn by blast lemma diff_char1: "A \<^bold>\ B = A \<^bold>\ \<^bold>\B" unfolding conn by simp lemma diff_char2: "A \<^bold>\ B = \<^bold>\(A \<^bold>\ B)" unfolding conn by simp lemma sdiff_char1: "A \<^bold>\ B = (A \<^bold>\ B) \<^bold>\ (B \<^bold>\ A)" unfolding conn by blast lemma sdiff_char2: "A \<^bold>\ B = \<^bold>\(A \<^bold>\ B)" unfolding conn by simp text\We can verify that (quite trivially) this algebra satisfies some properties of lattices.\ lemma L1: "A \<^bold>= A \<^bold>\ A" unfolding conn order by simp lemma L2: "A \<^bold>= A \<^bold>\ A" unfolding conn order by simp lemma L3: "A \<^bold>\ A \<^bold>\ B" unfolding conn order by simp lemma L4: "A \<^bold>\ B \<^bold>\ A" unfolding conn order by simp lemma L5: "(A \<^bold>\ B) \<^bold>\ B \<^bold>= B" unfolding setequ_char conn order by simp lemma L6: "A \<^bold>\ (A \<^bold>\ B) \<^bold>= A" unfolding setequ_char conn order by simp lemma L7: "A \<^bold>\ C \ B \<^bold>\ C \ A \<^bold>\ B \<^bold>\ C" unfolding conn order by simp lemma L8: "C \<^bold>\ A \ C \<^bold>\ B \ C \<^bold>\ A \<^bold>\ B" unfolding conn order by simp lemma L9: "A \<^bold>\ B \ (A \<^bold>\ B) \<^bold>= B" unfolding setequ_char conn order by simp lemma L10: "B \<^bold>\ A \ (A \<^bold>\ B) \<^bold>= B" unfolding setequ_char conn order by simp lemma L11: "A \<^bold>\ B \ C \<^bold>\ D \ A \<^bold>\ C \<^bold>\ B \<^bold>\ D" unfolding conn order by simp lemma L12: "A \<^bold>\ B \ C \<^bold>\ D \ A \<^bold>\ C \<^bold>\ B \<^bold>\ D" unfolding conn order by simp lemma L13: "A \<^bold>\ \<^bold>\ \<^bold>= A" unfolding conn order by simp lemma L14: "A \<^bold>\ \<^bold>\ \<^bold>= A" unfolding conn order by simp lemma L15: "A \<^bold>\ B \ (\C. C \<^bold>\ A \<^bold>\ C \<^bold>\ B)" by (metis L3 L4 L5 L8 setequ_char subset_char1) lemma L16: "A \<^bold>\ B \ (\C. C \<^bold>\ A \<^bold>\ C \<^bold>\ B)" by (metis L11 L4 L5 setequ_char setequ_ext) text\These properties below hold in particular also for Boolean algebras.\ lemma BA_impl: "A \<^bold>\ B \ A \<^bold>\ B \<^bold>= \<^bold>\" unfolding conn order by simp lemma BA_distr1: "(A \<^bold>\ (B \<^bold>\ C)) \<^bold>= ((A \<^bold>\ B) \<^bold>\ (A \<^bold>\ C))" unfolding setequ_char conn order by simp lemma BA_distr2: "(A \<^bold>\ (B \<^bold>\ C)) \<^bold>= ((A \<^bold>\ B) \<^bold>\ (A \<^bold>\ C))" unfolding conn order by blast lemma BA_cp: "A \<^bold>\ B \ \<^bold>\B \<^bold>\ \<^bold>\A" unfolding conn order by blast lemma BA_deMorgan1: "\<^bold>\(A \<^bold>\ B) \<^bold>= (\<^bold>\A \<^bold>\ \<^bold>\B)" unfolding conn order by simp lemma BA_deMorgan2: "\<^bold>\(A \<^bold>\ B) \<^bold>= (\<^bold>\A \<^bold>\ \<^bold>\B)" unfolding conn order by simp lemma BA_dn: "\<^bold>\\<^bold>\A \<^bold>= A" unfolding conn order by simp lemma BA_cmpl_equ: "(\<^bold>\A \<^bold>= B) = (A \<^bold>= \<^bold>\B)" unfolding conn order by blast text\We conveniently introduce these properties of sets of sets (of points).\ definition meet_closed::"('w \)\ \ bool" where "meet_closed S \ \X Y. (S X \ S Y) \ S(X \<^bold>\ Y)" definition join_closed::"('w \)\ \ bool" where "join_closed S \ \X Y. (S X \ S Y) \ S(X \<^bold>\ Y)" definition upwards_closed::"('w \)\ \ bool" where "upwards_closed S \ \X Y. S X \ X \<^bold>\ Y \ S Y" definition downwards_closed::"('w \)\ \ bool" where "downwards_closed S \ \X Y. S X \ Y \<^bold>\ X \ S Y" subsection \Atomicity and primitive equality\ text\We can verify indeed that the algebra is atomic (in three different ways) by relying on the presence of primitive equality in HOL.\ lemma atomic1: "\w. \Q. Q w \ (\P. P w \ Q \<^bold>\ P)" unfolding order using the_sym_eq_trivial by metis definition "atom A \ \(A \<^bold>= \<^bold>\) \ (\P. A \<^bold>\ P \ A \<^bold>\ \<^bold>\P)" lemma atomic2: "\w. \Q. Q w \ atom Q" unfolding atom_def order conn using the_sym_eq_trivial by metis lemma atomic3: "\P. \(P \<^bold>= \<^bold>\) \ (\Q. atom Q \ Q \<^bold>\ P)" unfolding atom_def order conn by fastforce text\Now with interactive proof:\ lemma "\P. \(P \<^bold>= \<^bold>\) \ (\Q. atom Q \ Q \<^bold>\ P)" proof - { fix P::"'w \" { assume "\(P \<^bold>= \<^bold>\)" hence "\v. P v" unfolding conn order by simp then obtain w where 1:"P w" by (rule exE) let ?Q="\v. v = w" \\ using HOL primitive equality \ have 2: "atom ?Q" unfolding atom_def unfolding conn order by simp have "\v. ?Q v \ P v" using 1 by simp hence 3: "?Q \<^bold>\ P" unfolding order by simp from 2 3 have "\Q. atom Q \ Q \<^bold>\ P" by blast } hence "\(P \<^bold>= \<^bold>\) \ (\Q. atom Q \ Q \<^bold>\ P)" by (rule impI) } thus ?thesis by (rule allI) qed subsection \Miscellaneous notions\ text\We add some miscellaneous notions that will be useful later.\ abbreviation "isEmpty S \ \x. \S x" abbreviation "nonEmpty S \ \x. S x" text\Function composition.\ definition fun_comp :: "('b \ 'c) \ ( 'a \ 'b) \ 'a \ 'c" (infixl "\" 75) where "\ \ \ \ \x. \ (\ x)" text\Inverse projection maps a unary function to a 'projected' binary function wrt. its 1st argument.\ abbreviation inv_proj::\('a \ 'c) \ ('a \ 'b \ 'c)\ ("(_)\") where "D\ \ \A B. D A" -text\Image of a mapping \, with range as an special case.\ +text\Image of a mapping @{text "\"}, with range as an special case.\ definition image::"('a \ 'b) \ ('a \ bool) \ ('b \ bool)" ("\_ _\") where "\\ S\ \ \y. \x. (S x) \ (\ x) = y" definition range::"('a \ 'b) \ ('b \ bool)" ("\_'_\") where "\\ _\ \ \Y. \x. (\ x) = Y" lemma range_char1: "\\ _\ = \\ (\x. True)\" by (simp add: image_def range_def) lemma range_char2: "\\ _\ = (\X. \S. \\ S\ X)" unfolding range_def image_def by blast lemma image_comp: "\(f \ g) S\ = \f \g S\\" unfolding fun_comp_def image_def by metis end diff --git a/thys/Topological_Semantics/boolean_algebra_infinitary.thy b/thys/Topological_Semantics/boolean_algebra_infinitary.thy --- a/thys/Topological_Semantics/boolean_algebra_infinitary.thy +++ b/thys/Topological_Semantics/boolean_algebra_infinitary.thy @@ -1,130 +1,130 @@ theory boolean_algebra_infinitary imports boolean_algebra_operators begin subsection \Encoding infinitary Boolean operations\ text\Our aim is to encode complete Boolean algebras (of sets) which we can later be used to interpret quantified formulas (somewhat in the spirit of Boolean-valued models for set theory).\ text\We start by defining infinite meet (infimum) and infinite join (supremum) operations.\ definition infimum:: "('w \)\ \ 'w \" ("\<^bold>\_") where "\<^bold>\S \ \w. \X. S X \ X w" definition supremum::"('w \)\ \ 'w \" ("\<^bold>\_") where "\<^bold>\S \ \w. \X. S X \ X w" declare infimum_def[conn] supremum_def[conn] (*add infimum and supremum to definition set of algebra connectives*) text\Infimum and supremum satisfy an infinite variant of the De Morgan laws.\ lemma iDM_a: "\<^bold>\(\<^bold>\S) \<^bold>= \<^bold>\(S\<^sup>d\<^sup>-)" unfolding order conn conn2 by force lemma iDM_b:" \<^bold>\(\<^bold>\S) \<^bold>= \<^bold>\(S\<^sup>d\<^sup>-)" unfolding order conn conn2 by force text\We show that our encoded Boolean algebras are lattice-complete. The functions below return the set of upper-/lower-bounds of a set of sets S (wrt. domain D).\ definition upper_bounds::"('w \)\ \ ('w \)\" ("ub") where "ub S \ \U. \X. S X \ X \<^bold>\ U" definition upper_bounds_restr::"('w \)\ \ ('w \)\ \ ('w \)\" ("ub\<^sup>_") where "ub\<^sup>D S \ \U. D U \ (\X. S X \ X \<^bold>\ U)" definition lower_bounds::"('w \)\ \ ('w \)\" ("lb") where "lb S \ \L. \X. S X \ L \<^bold>\ X" definition lower_bounds_restr::"('w \)\ \ ('w \)\ \ ('w \)\" ("lb\<^sup>_") where "lb\<^sup>D S \ \L. D L \ (\X. S X \ L \<^bold>\ X)" lemma ub_char: "ub S = (let D=\<^bold>\ in ub\<^sup>D S) " by (simp add: top_def upper_bounds_def upper_bounds_restr_def) lemma lb_char: "lb S = (let D=\<^bold>\ in lb\<^sup>D S) " by (simp add: top_def lower_bounds_def lower_bounds_restr_def) text\Similarly, the functions below return the set of least/greatest upper-/lower-bounds for S (wrt. D).\ definition lub::"('w \)\ \ ('w \)\" ("lub") where "lub S \ \U. ub S U \ (\X. ub S X \ U \<^bold>\ X)" definition lub_restr::"('w \)\ \ ('w \)\ \ ('w \)\" ("lub\<^sup>_") where "lub\<^sup>D S \ \U. ub\<^sup>D S U \ (\X. ub\<^sup>D S X \ U \<^bold>\ X)" definition glb::"('w \)\ \ ('w \)\" ("glb") where "glb S \ \L. lb S L \ (\X. lb S X \ X \<^bold>\ L)" definition glb_restr::"('w \)\ \ ('w \)\ \ ('w \)\" ("glb\<^sup>_") where "glb\<^sup>D S \ \L. lb\<^sup>D S L \ (\X. lb\<^sup>D S X \ X \<^bold>\ L)" text\Both pairs of definitions above are suitably related. -(Note that the term \ below denotes the top element in the algebra of sets of sets (i.e. the powerset).)\ +(Note that the term @{text "\"} below denotes the top element in the algebra of sets of sets (i.e. the powerset).)\ lemma lub_char: "lub S = (let D=\<^bold>\ in lub\<^sup>D S) " by (simp add: lub_def lub_restr_def ub_char) lemma glb_char: "glb S = (let D=\<^bold>\ in glb\<^sup>D S) " by (simp add: glb_def glb_restr_def lb_char) text\Clearly, the notions of infimum/supremum correspond to least/greatest upper-/lower-bound.\ lemma sup_lub: "lub S \<^bold>\S" unfolding lub_def upper_bounds_def supremum_def subset_def by blast lemma sup_exist_unique: "\S. \!X. lub S X" by (meson lub_def setequ_char setequ_ext sup_lub) lemma inf_glb: "glb S \<^bold>\S" unfolding glb_def lower_bounds_def infimum_def subset_def by blast lemma inf_exist_unique: "\S. \!X. glb S X" by (meson glb_def inf_glb setequ_char setequ_ext) text\The property of being closed under arbitrary (resp. nonempty) supremum/infimum.\ definition infimum_closed :: "('w \)\ \ bool" where "infimum_closed S \ \D. D \<^bold>\ S \ S(\<^bold>\D)" (*observe that D can be empty*) definition supremum_closed :: "('w \)\ \ bool" where "supremum_closed S \ \D. D \<^bold>\ S \ S(\<^bold>\D)" (*observe that D can be empty*) definition infimum_closed' :: "('w \)\ \ bool" where"infimum_closed' S \ \D. nonEmpty D \ D \<^bold>\ S \ S(\<^bold>\D)" definition supremum_closed' :: "('w \)\ \ bool" where "supremum_closed' S \ \D. nonEmpty D \ D \<^bold>\ S \ S(\<^bold>\D)" lemma inf_empty: "isEmpty S \ \<^bold>\S \<^bold>= \<^bold>\" unfolding order conn by simp lemma sup_empty: "isEmpty S \ \<^bold>\S \<^bold>= \<^bold>\" unfolding order conn by simp text\Note that arbitrary infimum- (resp. supremum-) closed sets include the top (resp. bottom) element.\ lemma "infimum_closed S \ S \<^bold>\" unfolding infimum_closed_def conn order by force lemma "supremum_closed S \ S \<^bold>\" unfolding supremum_closed_def conn order by force text\However, the above does not hold for non-empty infimum- (resp. supremum-) closed sets.\ lemma "infimum_closed' S \ S \<^bold>\" nitpick oops \\ countermodel \ lemma "supremum_closed' S \ S \<^bold>\" nitpick oops \\ countermodel \ text\We have in fact the following characterizations for the notions above.\ lemma inf_closed_char: "infimum_closed S = (infimum_closed' S \ S \<^bold>\)" unfolding infimum_closed'_def infimum_closed_def by (metis bottom_def infimum_closed_def infimum_def setequ_char setequ_ext subset_def top_def) lemma sup_closed_char: "supremum_closed S = (supremum_closed' S \ S \<^bold>\)" unfolding supremum_closed'_def supremum_closed_def by (metis (no_types, opaque_lifting) L14 L9 bottom_def setequ_ext subset_def supremum_def) lemma inf_sup_closed_dc: "infimum_closed S = supremum_closed S\<^sup>d\<^sup>-" by (smt (verit) BA_dn iDM_a iDM_b infimum_closed_def setequ_ext sdfun_dcompl_def subset_def supremum_closed_def) lemma inf_sup_closed_dc': "infimum_closed' S = supremum_closed' S\<^sup>d\<^sup>-" by (smt (verit) dualcompl_invol iDM_a infimum_closed'_def sdfun_dcompl_def setequ_ext subset_def supremum_closed'_def) text\We check some further properties.\ lemma fp_inf_sup_closed_dual: "infimum_closed (fp \) = supremum_closed (fp \\<^sup>d)" by (simp add: fp_dual inf_sup_closed_dc) lemma fp_inf_sup_closed_dual': "infimum_closed' (fp \) = supremum_closed' (fp \\<^sup>d)" by (simp add: fp_dual inf_sup_closed_dc') text\We verify that being infimum-closed' (resp. supremum-closed') entails being meet-closed (resp. join-closed).\ lemma inf_meet_closed: "\S. infimum_closed' S \ meet_closed S" proof - { fix S::"'w \ \ bool" { assume inf_closed: "infimum_closed' S" hence "meet_closed S" proof - { fix X::"'w \" and Y::"'w \" let ?D="\Z. Z=X \ Z=Y" { assume "S X \ S Y" hence "?D \<^bold>\ S" using subset_def by blast moreover have "nonEmpty ?D" by auto ultimately have "S(\<^bold>\?D)" using inf_closed infimum_closed'_def by (smt (z3)) hence "S(\w. \Z. (Z=X \ Z=Y) \ Z w)" unfolding infimum_def by simp moreover have "(\w. \Z. (Z=X \ Z=Y) \ Z w) = (\w. X w \ Y w)" by auto ultimately have "S(\w. X w \ Y w)" by simp } hence "(S X \ S Y) \ S(X \<^bold>\ Y)" unfolding conn by (rule impI) } thus ?thesis unfolding meet_closed_def by simp qed } hence "infimum_closed' S \ meet_closed S" by simp } thus ?thesis by (rule allI) qed lemma sup_join_closed: "\P. supremum_closed' P \ join_closed P" proof - { fix S::"'w \ \ bool" { assume sup_closed: "supremum_closed' S" hence "join_closed S" proof - { fix X::"'w \" and Y::"'w \" let ?D="\Z. Z=X \ Z=Y" { assume "S X \ S Y" hence "?D \<^bold>\ S" using subset_def by blast moreover have "nonEmpty ?D" by auto ultimately have "S(\<^bold>\?D)" using sup_closed supremum_closed'_def by (smt (z3)) hence "S(\w. \Z. (Z=X \ Z=Y) \ Z w)" unfolding supremum_def by simp moreover have "(\w. \Z. (Z=X \ Z=Y) \ Z w) = (\w. X w \ Y w)" by auto ultimately have "S(\w. X w \ Y w)" by simp } hence "(S X \ S Y) \ S(X \<^bold>\ Y)" unfolding conn by (rule impI) } thus ?thesis unfolding join_closed_def by simp qed } hence "supremum_closed' S \ join_closed S" by simp } thus ?thesis by (rule allI) qed end diff --git a/thys/Topological_Semantics/boolean_algebra_operators.thy b/thys/Topological_Semantics/boolean_algebra_operators.thy --- a/thys/Topological_Semantics/boolean_algebra_operators.thy +++ b/thys/Topological_Semantics/boolean_algebra_operators.thy @@ -1,155 +1,155 @@ theory boolean_algebra_operators imports boolean_algebra begin subsection \Operations on set-valued functions\ text\Functions with sets in their codomain will be called here 'set-valued functions'. We conveniently define some (2nd-order) Boolean operations on them.\ text\The 'meet' and 'join' of two set-valued functions.\ definition svfun_meet::"('a \ 'w \) \ ('a \ 'w \) \ ('a \ 'w \)" (infixr "\<^bold>\\<^sup>:" 62) where "\ \<^bold>\\<^sup>: \ \ \x. (\ x) \<^bold>\ (\ x)" definition svfun_join::"('a \ 'w \) \ ('a \ 'w \) \ ('a \ 'w \)" (infixr "\<^bold>\\<^sup>:" 61) where "\ \<^bold>\\<^sup>: \ \ \x. (\ x) \<^bold>\ (\ x)" text\Analogously, we can define an 'implication' and a 'complement'.\ definition svfun_impl::"('a \ 'w \) \ ('a \ 'w \) \ ('a \ 'w \)" (infixr "\<^bold>\\<^sup>:" 61) where "\ \<^bold>\\<^sup>: \ \ \x. (\ x) \<^bold>\ (\ x)" definition svfun_compl::"('a \ 'w \) \ ('a \ 'w \)" ("(_\<^sup>-)") where "\\<^sup>- \ \x. \<^bold>\(\ x)" text\There are two natural 0-ary connectives (aka. constants). \ definition svfun_top::"'a \ 'w \" ("\<^bold>\\<^sup>:") where "\<^bold>\\<^sup>: \ \x. \<^bold>\" definition svfun_bot::"'a \ 'w \" ("\<^bold>\\<^sup>:") where "\<^bold>\\<^sup>: \ \x. \<^bold>\" named_theorems conn2 (*to group together definitions for 2nd-order algebraic connectives*) declare svfun_meet_def[conn2] svfun_join_def[conn2] svfun_impl_def[conn2] svfun_compl_def[conn2] svfun_top_def[conn2] svfun_bot_def[conn2] text\And, of course, set-valued functions are naturally ordered in the expected way:\ definition svfun_sub::"('a \ 'w \) \ ('a \ 'w \) \ bool" (infixr "\<^bold>\\<^sup>:" 55) where "\ \<^bold>\\<^sup>: \ \ \x. (\ x) \<^bold>\ (\ x)" definition svfun_equ::"('a \ 'w \) \ ('a \ 'w \) \ bool" (infixr "\<^bold>=\<^sup>:" 55) where "\ \<^bold>=\<^sup>: \ \ \x. (\ x) \<^bold>= (\ x)" named_theorems order2 (*to group together definitions for 2nd-order algebraic connectives*) declare svfun_sub_def[order2] svfun_equ_def[order2] text\These (trivial) lemmas are intended to help automated tools.\ lemma svfun_sub_char: "(\ \<^bold>\\<^sup>: \) = (\ \<^bold>\\<^sup>: \ \<^bold>=\<^sup>: \<^bold>\\<^sup>:)" by (simp add: BA_impl svfun_equ_def svfun_impl_def svfun_sub_def svfun_top_def) lemma svfun_equ_char: "(\ \<^bold>=\<^sup>: \) = (\ \<^bold>\\<^sup>: \ \ \ \<^bold>\\<^sup>: \)" unfolding order2 setequ_char by blast lemma svfun_equ_ext: "(\ \<^bold>=\<^sup>: \) = (\ = \)" by (meson ext setequ_ext svfun_equ_def) text\Clearly, set-valued functions form a Boolean algebra. We can prove some interesting relationships:\ lemma svfun_compl_char: "\\<^sup>- = (\ \<^bold>\\<^sup>: \<^bold>\\<^sup>:)" unfolding conn conn2 by simp lemma svfun_impl_char1: "(\ \<^bold>\\<^sup>: \) = (\\<^sup>- \<^bold>\\<^sup>: \)" unfolding conn conn2 by simp lemma svfun_impl_char2: "(\ \<^bold>\\<^sup>: \) = (\ \<^bold>\\<^sup>: (\\<^sup>-))\<^sup>-" unfolding conn conn2 by simp lemma svfun_deMorgan1: "(\ \<^bold>\\<^sup>: \)\<^sup>- = (\\<^sup>-) \<^bold>\\<^sup>: (\\<^sup>-)" unfolding conn conn2 by simp lemma svfun_deMorgan2: "(\ \<^bold>\\<^sup>: \)\<^sup>- = (\\<^sup>-) \<^bold>\\<^sup>: (\\<^sup>-)" unfolding conn conn2 by simp subsection \Operators and their transformations\ text\Dual to set-valued functions we can have set-domain functions. For them we can define the 'dual-complement'.\ definition sdfun_dcompl::"('w \ \ 'a) \ ('w \ \ 'a)" ("(_\<^sup>d\<^sup>-)") where "\\<^sup>d\<^sup>- \ \X. \(\<^bold>\X)" lemma sdfun_dcompl_char: "\\<^sup>d\<^sup>- = (\X. \Y. (\ Y) \ (X = \<^bold>\Y))" by (metis BA_dn setequ_ext sdfun_dcompl_def) text\Operators are a particularly important kind of functions. They are both set-valued and set-domain. Thus our algebra of operators inherits the connectives defined above plus the ones below. \ text\We conveniently define the 'dual' of an operator.\ definition op_dual::"('w \ \ 'w \) \ ('w \ \ 'w \)" ("(_\<^sup>d)") where "\\<^sup>d \ \X. \<^bold>\(\(\<^bold>\X))" text\The following two 0-ary connectives (i.e. operator 'constants') exist already (but somehow implicitly). We just make them explicit by introducing some convenient notation.\ definition id_op::"'w \ \ 'w \" ("\<^bold>e") where "\<^bold>e \ \X. X" (*introduces notation to refer to 'identity' operator*) definition compl_op::"'w \ \ 'w \" ("\<^bold>n") where "\<^bold>n \ \X. \<^bold>\X" (*to refer to 'complement' operator*) declare sdfun_dcompl_def[conn2] op_dual_def[conn2] id_op_def[conn2] compl_op_def[conn2] text\We now prove some lemmas (some of them might help provers in their hard work).\ lemma dual_compl_char1: "\\<^sup>d\<^sup>- = (\\<^sup>d)\<^sup>-" unfolding conn2 conn order by simp lemma dual_compl_char2: "\\<^sup>d\<^sup>- = (\\<^sup>-)\<^sup>d" unfolding conn2 conn order by simp lemma sfun_compl_invol: "\\<^sup>-\<^sup>- = \" unfolding conn2 conn order by simp lemma dual_invol: "\\<^sup>d\<^sup>d = \" unfolding conn2 conn order by simp lemma dualcompl_invol: "(\\<^sup>d\<^sup>-)\<^sup>d\<^sup>- = \" unfolding conn2 conn order by simp lemma op_prop1: "\<^bold>e\<^sup>d = \<^bold>e" unfolding conn2 conn by simp lemma op_prop2: "\<^bold>n\<^sup>d = \<^bold>n" unfolding conn2 conn by simp lemma op_prop3: "\<^bold>e\<^sup>- = \<^bold>n" unfolding conn2 conn by simp lemma op_prop4: "(\ \<^bold>\\<^sup>: \)\<^sup>d = (\\<^sup>d) \<^bold>\\<^sup>: (\\<^sup>d)" unfolding conn2 conn by simp lemma op_prop5: "(\ \<^bold>\\<^sup>: \)\<^sup>- = (\\<^sup>-) \<^bold>\\<^sup>: (\\<^sup>-)" unfolding conn2 conn by simp lemma op_prop6: "(\ \<^bold>\\<^sup>: \)\<^sup>d = (\\<^sup>d) \<^bold>\\<^sup>: (\\<^sup>d)" unfolding conn2 conn by simp lemma op_prop7: "(\ \<^bold>\\<^sup>: \)\<^sup>- = (\\<^sup>-) \<^bold>\\<^sup>: (\\<^sup>-)" unfolding conn2 conn by simp lemma op_prop8: "\<^bold>\\<^sup>: = \<^bold>n \<^bold>\\<^sup>: \<^bold>e" unfolding conn2 conn by simp lemma op_prop9: "\<^bold>\\<^sup>: = \<^bold>n \<^bold>\\<^sup>: \<^bold>e" unfolding conn2 conn by simp text\The notion of a fixed-point is fundamental. We speak of sets being fixed-points of operators. We define a function that given an operator returns the set of all its fixed-points.\ definition fixpoints::"('w \ \ 'w \) \ ('w \)\" ("fp") where "fp \ \ \X. (\ X) \<^bold>= X" text\We can in fact 'operationalize' the function above, thus obtaining a 'fixed-point operation'.\ definition op_fixpoint::"('w \ \ 'w \) \ ('w \ \ 'w \)" ("(_\<^sup>f\<^sup>p)") where "\\<^sup>f\<^sup>p \ \X. (\ X) \<^bold>\ X" declare fixpoints_def[conn2] op_fixpoint_def[conn2] text\Interestingly, the fixed-point operation (or transformation) is definable in terms of the others.\ lemma op_fixpoint_char: "\\<^sup>f\<^sup>p = (\ \<^bold>\\<^sup>: \<^bold>e) \<^bold>\\<^sup>: (\\<^sup>- \<^bold>\\<^sup>: \<^bold>n)" unfolding conn2 order conn by blast -text\Given an operator \ the fixed-points of \'s dual is the set of complements of \'s fixed-points.\ +text\Given an operator @{text "\"} the fixed-points of it's dual is the set of complements of @{text "\'s"} fixed-points.\ lemma fp_dual: "fp \\<^sup>d = (fp \)\<^sup>d\<^sup>-" unfolding order conn conn2 by blast -text\The fixed-points of \'s complement is the set of complements of the fixed-points of \'s dual-complement.\ +text\The fixed-points of @{text "\'s"} complement is the set of complements of the fixed-points of @{text "\'s"} dual-complement.\ lemma fp_compl: "fp \\<^sup>- = (fp (\\<^sup>d\<^sup>-))\<^sup>d\<^sup>-" by (simp add: dual_compl_char2 dualcompl_invol fp_dual) -text\The fixed-points of \'s dual-complement is the set of complements of the fixed-points of \'s complement.\ +text\The fixed-points of @{text "\'s"} dual-complement is the set of complements of the fixed-points of @{text "\'s"} complement.\ lemma fp_dcompl: "fp (\\<^sup>d\<^sup>-) = (fp \\<^sup>-)\<^sup>d\<^sup>-" by (simp add: dualcompl_invol fp_compl) text\The fixed-points function and the fixed-point operation are essentially related.\ lemma fp_rel: "fp \ A \ (\\<^sup>f\<^sup>p A) \<^bold>= \<^bold>\" unfolding conn2 order conn by simp lemma fp_d_rel: "fp \\<^sup>d A \ \\<^sup>f\<^sup>p(\<^bold>\A) \<^bold>= \<^bold>\" unfolding conn2 order conn by blast lemma fp_c_rel: "fp \\<^sup>- A \ \\<^sup>f\<^sup>p A \<^bold>= \<^bold>\" unfolding conn2 order conn by blast lemma fp_dc_rel: "fp (\\<^sup>d\<^sup>-) A \ \\<^sup>f\<^sup>p(\<^bold>\A) \<^bold>= \<^bold>\" unfolding conn2 order conn by simp text\The fixed-point operation is involutive.\ lemma ofp_invol: "(\\<^sup>f\<^sup>p)\<^sup>f\<^sup>p = \" unfolding conn2 order conn by blast text\And commutes the dual with the dual-complement operations.\ lemma ofp_comm_dc1: "(\\<^sup>d)\<^sup>f\<^sup>p = (\\<^sup>f\<^sup>p)\<^sup>d\<^sup>-" unfolding conn2 order conn by blast lemma ofp_comm_dc2:"(\\<^sup>d\<^sup>-)\<^sup>f\<^sup>p = (\\<^sup>f\<^sup>p)\<^sup>d" unfolding conn2 order conn by simp text\The fixed-point operation commutes with the complement.\ lemma ofp_comm_compl: "(\\<^sup>-)\<^sup>f\<^sup>p = (\\<^sup>f\<^sup>p)\<^sup>-" unfolding conn2 order conn by blast text\The above motivates the following alternative definition for a 'complemented-fixed-point' operation.\ lemma ofp_fixpoint_compl_def: "\\<^sup>f\<^sup>p\<^sup>- = (\X. (\ X) \<^bold>\ X)" unfolding conn2 conn by simp text\Analogously, the 'complemented fixed-point' operation is also definable in terms of the others.\ lemma op_fixpoint_compl_char: "\\<^sup>f\<^sup>p\<^sup>- = (\ \<^bold>\\<^sup>: \<^bold>e) \<^bold>\\<^sup>: (\\<^sup>- \<^bold>\\<^sup>: \<^bold>n)" unfolding conn2 conn by blast text\In fact, function composition can be seen as an additional binary connective for operators. We show below some interesting relationships that hold. \ lemma op_prop10: "\ = (\<^bold>e \ \)" unfolding conn2 fun_comp_def by simp lemma op_prop11: "\ = (\ \ \<^bold>e)" unfolding conn2 fun_comp_def by simp lemma op_prop12: "\<^bold>e = (\<^bold>n \ \<^bold>n)" unfolding conn2 conn fun_comp_def by simp lemma op_prop13: "\\<^sup>- = (\<^bold>n \ \)" unfolding conn2 fun_comp_def by simp lemma op_prop14: "\\<^sup>d\<^sup>- = (\ \ \<^bold>n)" unfolding conn2 fun_comp_def by simp lemma op_prop15: "\\<^sup>d = (\<^bold>n \ \ \ \<^bold>n)" unfolding conn2 fun_comp_def by simp text\There are also some useful properties regarding the images of operators.\ lemma im_prop1: "\\ D\\<^sup>d\<^sup>- = \\\<^sup>d D\<^sup>d\<^sup>-\" unfolding image_def op_dual_def sdfun_dcompl_def by (metis BA_dn setequ_ext) lemma im_prop2: "\\\<^sup>- D\\<^sup>d\<^sup>- = \\ D\" unfolding image_def svfun_compl_def sdfun_dcompl_def by (metis BA_dn setequ_ext) lemma im_prop3: "\\\<^sup>d D\\<^sup>d\<^sup>- = \\ D\<^sup>d\<^sup>-\" unfolding image_def op_dual_def sdfun_dcompl_def by (metis BA_dn setequ_ext) lemma im_prop4: "\\\<^sup>d\<^sup>- D\\<^sup>d\<^sup>- = \\\<^sup>d D\" unfolding image_def op_dual_def sdfun_dcompl_def by (metis BA_dn setequ_ext) lemma im_prop5: "\\\<^sup>- D\<^sup>d\<^sup>-\ = \\\<^sup>d\<^sup>- D\\<^sup>d\<^sup>-" unfolding image_def svfun_compl_def sdfun_dcompl_def by (metis (no_types, opaque_lifting) BA_dn setequ_ext) lemma im_prop6: "\\\<^sup>d\<^sup>- D\<^sup>d\<^sup>-\ = \\ D\" unfolding image_def sdfun_dcompl_def by (metis BA_dn setequ_ext) text\Observe that all results obtained by assuming fixed-point predicates extend to their associated operators.\ lemma "\\<^sup>f\<^sup>p(A) \<^bold>\ \(A) \<^bold>\ \(A) \ (fp \)(A) \ \(A) \<^bold>\ \(A)" by (simp add: fp_rel meet_def setequ_ext subset_def top_def) lemma "\\<^sup>f\<^sup>p(A) \<^bold>\ \\<^sup>f\<^sup>p(B) \<^bold>\ (\ A B) \<^bold>\ (\ A B) \ (fp \)(A) \ (fp \)(B) \ (\ A B) \<^bold>\ (\ A B)" by (simp add: fp_rel meet_def setequ_ext subset_def top_def) end diff --git a/thys/Topological_Semantics/conditions_positive.thy b/thys/Topological_Semantics/conditions_positive.thy --- a/thys/Topological_Semantics/conditions_positive.thy +++ b/thys/Topological_Semantics/conditions_positive.thy @@ -1,140 +1,140 @@ theory conditions_positive imports boolean_algebra_operators begin section \Topological Conditions\ text\We define and interrelate some useful axiomatic conditions on unary operations (operators) -having a 'w-parametric type ('w)\\('w)\. +having a 'w-parametric type @{text "('w)\\('w)\"}. Boolean algebras extended with such operators give us different sorts of topological Boolean algebras.\ subsection \Positive Conditions\ text\Monotonicity (MONO).\ definition MONO::"('w \ \ 'w \) \ bool" ("MONO") where "MONO \ \ \A B. A \<^bold>\ B \ \ A \<^bold>\ \ B" named_theorems cond (*to group together axiomatic conditions*) declare MONO_def[cond] text\MONO is self-dual.\ lemma MONO_dual: "MONO \ = MONO \\<^sup>d" by (smt (verit) BA_cp MONO_def dual_invol op_dual_def) text\Expansive/extensive (EXPN) and its dual contractive (CNTR).\ definition EXPN::"('w \ \ 'w \) \ bool" ("EXPN") where "EXPN \ \ \A. A \<^bold>\ \ A" definition CNTR::"('w \ \ 'w \) \ bool" ("CNTR") where "CNTR \ \ \A. \ A \<^bold>\ A" declare EXPN_def[cond] CNTR_def[cond] text\EXPN and CNTR are dual to each other.\ lemma EXPN_CNTR_dual1: "EXPN \ = CNTR \\<^sup>d" unfolding cond by (metis BA_cp BA_dn op_dual_def setequ_ext) lemma EXPN_CNTR_dual2: "CNTR \ = EXPN \\<^sup>d" by (simp add: EXPN_CNTR_dual1 dual_invol) text\Normality (NORM) and its dual (DNRM).\ definition NORM::"('w \ \ 'w \) \ bool" ("NORM") where "NORM \ \ (\ \<^bold>\) \<^bold>= \<^bold>\" definition DNRM::"('w \ \ 'w \) \ bool" ("DNRM") where "DNRM \ \ (\ \<^bold>\) \<^bold>= \<^bold>\" declare NORM_def[cond] DNRM_def[cond] text\NORM and DNRM are dual to each other.\ lemma NOR_dual1: "NORM \ = DNRM \\<^sup>d" unfolding cond by (simp add: bottom_def compl_def op_dual_def setequ_def top_def) lemma NOR_dual2: "DNRM \ = NORM \\<^sup>d" by (simp add: NOR_dual1 dual_invol) text\EXPN (CNTR) entails DNRM (NORM).\ lemma EXPN_impl_DNRM: "EXPN \ \ DNRM \" unfolding cond by (simp add: setequ_def subset_def top_def) lemma CNTR_impl_NORM: "CNTR \ \ NORM \" by (simp add: EXPN_CNTR_dual2 EXPN_impl_DNRM NOR_dual1 dual_invol) text\Idempotence (IDEM).\ definition IDEM::"('w \ \ 'w \) \ bool" ("IDEM") where "IDEM \ \ \A. \(\ A) \<^bold>= (\ A)" definition IDEM_a::"('w \ \ 'w \) \ bool" ("IDEM\<^sup>a") where "IDEM\<^sup>a \ \ \A. \(\ A) \<^bold>\ (\ A)" definition IDEM_b::"('w \ \ 'w \) \ bool" ("IDEM\<^sup>b") where "IDEM\<^sup>b \ \ \A. (\ A) \<^bold>\ \(\ A)" declare IDEM_def[cond] IDEM_a_def[cond] IDEM_b_def[cond] text\IDEM-a and IDEM-b are dual to each other.\ lemma IDEM_dual1: "IDEM\<^sup>a \ = IDEM\<^sup>b \\<^sup>d" unfolding cond by (metis (mono_tags, opaque_lifting) BA_cp BA_dn op_dual_def setequ_ext) lemma IDEM_dual2: "IDEM\<^sup>b \ = IDEM\<^sup>a \\<^sup>d" by (simp add: IDEM_dual1 dual_invol) lemma IDEM_char: "IDEM \ = (IDEM\<^sup>a \ \ IDEM\<^sup>b \)" unfolding cond setequ_char by blast lemma IDEM_dual: "IDEM \ = IDEM \\<^sup>d" using IDEM_char IDEM_dual1 IDEM_dual2 by blast text\EXPN (CNTR) entail IDEM-b (IDEM-a).\ lemma EXPN_impl_IDEM_b: "EXPN \ \ IDEM\<^sup>b \" by (simp add: EXPN_def IDEM_b_def) lemma CNTR_impl_IDEM_a: "CNTR \ \ IDEM\<^sup>a \" by (simp add: CNTR_def IDEM_a_def) text\Moreover, IDEM has some other interesting characterizations. For example, via function composition:\ lemma IDEM_fun_comp_char: "IDEM \ = (\ = \ \ \)" unfolding cond fun_comp_def by (metis setequ_ext) text\Or having the property of collapsing the range and the set of fixed-points of an operator:\ lemma IDEM_range_fp_char: "IDEM \ = (\\ _\ = fp \)" unfolding cond range_def fixpoints_def by (metis setequ_ext) text\Distribution over joins or additivity (ADDI).\ definition ADDI::"('w \ \ 'w \) \ bool" ("ADDI") where "ADDI \ \ \A B. \(A \<^bold>\ B) \<^bold>= (\ A) \<^bold>\ (\ B)" definition ADDI_a::"('w \ \ 'w \) \ bool" ("ADDI\<^sup>a") where "ADDI\<^sup>a \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" definition ADDI_b::"('w \ \ 'w \) \ bool" ("ADDI\<^sup>b") where "ADDI\<^sup>b \ \ \A B. (\ A) \<^bold>\ (\ B) \<^bold>\ \(A \<^bold>\ B)" text\Distribution over meets or multiplicativity (MULT).\ definition MULT::"('w \ \ 'w \) \ bool" ("MULT") where "MULT \ \ \A B. \(A \<^bold>\ B) \<^bold>= (\ A) \<^bold>\ (\ B)" definition MULT_a::"('w \ \ 'w \) \ bool" ("MULT\<^sup>a") where "MULT\<^sup>a \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" definition MULT_b::"('w \ \ 'w \) \ bool" ("MULT\<^sup>b") where "MULT\<^sup>b \ \ \A B. (\ A) \<^bold>\ (\ B) \<^bold>\ \(A \<^bold>\ B)" declare ADDI_def[cond] ADDI_a_def[cond] ADDI_b_def[cond] MULT_def[cond] MULT_a_def[cond] MULT_b_def[cond] lemma ADDI_char: "ADDI \ = (ADDI\<^sup>a \ \ ADDI\<^sup>b \)" unfolding cond using setequ_char by blast lemma MULT_char: "MULT \ = (MULT\<^sup>a \ \ MULT\<^sup>b \)" unfolding cond using setequ_char by blast text\MONO, MULT-a and ADDI-b are equivalent.\ lemma MONO_MULTa: "MULT\<^sup>a \ = MONO \" unfolding cond by (metis L10 L3 L4 L5 L8 setequ_char setequ_ext) lemma MONO_ADDIb: "ADDI\<^sup>b \ = MONO \" unfolding cond by (metis (mono_tags, lifting) L7 L9 join_def setequ_ext subset_def) text\Below we prove several duality relationships between ADDI(a/b) and MULT(a/b).\ text\Duality between MULT-a and ADDI-b (an easy corollary from the self-duality of MONO).\ lemma MULTa_ADDIb_dual1: "MULT\<^sup>a \ = ADDI\<^sup>b \\<^sup>d" by (metis MONO_ADDIb MONO_MULTa MONO_dual) lemma MULTa_ADDIb_dual2: "ADDI\<^sup>b \ = MULT\<^sup>a \\<^sup>d" by (simp add: MULTa_ADDIb_dual1 dual_invol) text\Duality between ADDI-a and MULT-b.\ lemma ADDIa_MULTb_dual1: "ADDI\<^sup>a \ = MULT\<^sup>b \\<^sup>d" unfolding cond op_dual_def by (metis BA_cp BA_deMorgan1 BA_dn setequ_ext) lemma ADDIa_MULTb_dual2: "MULT\<^sup>b \ = ADDI\<^sup>a \\<^sup>d" by (simp add: ADDIa_MULTb_dual1 dual_invol) text\Duality between ADDI and MULT.\ lemma ADDI_MULT_dual1: "ADDI \ = MULT \\<^sup>d" using ADDI_char ADDIa_MULTb_dual1 MULT_char MULTa_ADDIb_dual2 by blast lemma ADDI_MULT_dual2: "MULT \ = ADDI \\<^sup>d" by (simp add: ADDI_MULT_dual1 dual_invol) text\We verify properties regarding closure over meets/joins for fixed-points.\ text\MULT implies meet-closedness of the set of fixed-points (the converse requires additional assumptions).\ lemma MULT_meetclosed: "MULT \ \ meet_closed (fp \)" by (simp add: MULT_def fixpoints_def meet_closed_def setequ_ext) lemma "meet_closed (fp \) \ MULT \" nitpick oops \\ countermodel found: needs further assumptions. \ lemma meetclosed_MULT: "MONO \ \ CNTR \ \ IDEM\<^sup>b \ \ meet_closed (fp \) \ MULT \" by (smt (z3) CNTR_def IDEM_b_def MONO_MULTa MONO_def MULT_a_def MULT_def fixpoints_def meet_closed_def meet_def setequ_char setequ_ext subset_def) text\ADDI implies join-closedness of the set of fixed-points (the converse requires additional assumptions).\ lemma ADDI_joinclosed: "ADDI \ \ join_closed (fp \)" by (simp add: ADDI_def fixpoints_def join_closed_def setequ_ext) lemma "join_closed (fp \) \ ADDI \" nitpick oops \\ countermodel found: needs further assumptions \ lemma joinclosed_ADDI: "MONO \ \ EXPN \ \ IDEM\<^sup>a \ \ join_closed (fp \) \ ADDI \" by (smt (verit, ccfv_threshold) ADDI_MULT_dual1 BA_deMorgan2 EXPN_CNTR_dual1 IDEM_dual1 MONO_dual fp_dual join_closed_def meet_closed_def meetclosed_MULT sdfun_dcompl_def setequ_ext) text\Assuming MONO, we have that EXPN (CNTR) implies meet-closed (join-closed) for the set of fixed-points.\ lemma EXPN_meetclosed: "MONO \ \ EXPN \ \ meet_closed (fp \)" by (smt (verit) EXPN_def MONO_MULTa MULT_a_def fixpoints_def meet_closed_def setequ_char setequ_ext) lemma CNTR_joinclosed: "MONO \ \ CNTR \ \ join_closed (fp \)" by (smt (verit, best) ADDI_b_def CNTR_def MONO_ADDIb fixpoints_def join_closed_def setequ_char setequ_ext) text\Further assuming IDEM the above results can be stated to the whole range of an operator.\ lemma "MONO \ \ EXPN \ \ IDEM \ \ meet_closed (\\ _\)" by (simp add: EXPN_meetclosed IDEM_range_fp_char) lemma "MONO \ \ CNTR \ \ IDEM \ \ join_closed (\\ _\)" by (simp add: CNTR_joinclosed IDEM_range_fp_char) end diff --git a/thys/Topological_Semantics/ex_LFIs.thy b/thys/Topological_Semantics/ex_LFIs.thy deleted file mode 100644 --- a/thys/Topological_Semantics/ex_LFIs.thy +++ /dev/null @@ -1,157 +0,0 @@ -theory ex_LFIs - imports topo_negation_conditions -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Example application: Logics of Formal Inconsistency (LFIs)\ -text\\noindent{The LFIs \<^cite>\LFI\ \<^cite>\RLFI\ are a family of paraconsistent logics featuring a 'consistency' -operator @{text "\<^bold>\"} that can be used to recover some classical properties of negation (in particular ECQ). -We show how to semantically embed LFIs as extensions of Boolean algebras (here as frontier algebras).}\ - -text\\noindent{Logical validity can be defined as truth in all worlds/points (i.e. as denoting the top element)}\ -abbreviation gtrue::"\\bool" ("[\<^bold>\ _]") where "[\<^bold>\ A] \ \w. A w" -lemma gtrue_def: "[\<^bold>\ A] \ A \<^bold>\ \<^bold>\" by (simp add: top_def) - -text\\noindent{When defining a logic over an existing algebra we have two choices: a global (truth preserving) -and a local (truth-degree preserving) notion of logical consequence. For LFIs we prefer the latter.}\ -abbreviation conseq_global1::"\\\\bool" ("[_ \<^bold>\\<^sub>g _]") where "[A \<^bold>\\<^sub>g B] \ [\<^bold>\ A] \ [\<^bold>\ B]" -abbreviation conseq_global2::"\\\\\\bool" ("[_,_ \<^bold>\\<^sub>g _]") where "[A\<^sub>1, A\<^sub>2 \<^bold>\\<^sub>g B] \ [\<^bold>\ A\<^sub>1] \ [\<^bold>\ A\<^sub>2] \ [\<^bold>\ B]" -abbreviation conseq_global3::"\\\\\\\\bool" ("[_,_,_ \<^bold>\\<^sub>g _]") where "[A\<^sub>1, A\<^sub>2, A\<^sub>3 \<^bold>\\<^sub>g B] \ [\<^bold>\ A\<^sub>1] \ [\<^bold>\ A\<^sub>2] \ [\<^bold>\ A\<^sub>3] \ [\<^bold>\ B]" -abbreviation conseq_local1::"\\\\bool" ("[_ \<^bold>\ _]") where "[A \<^bold>\ B] \ A \<^bold>\ B" -abbreviation conseq_local2::"\\\\\\bool" ("[_,_ \<^bold>\ _]") where "[A\<^sub>1, A\<^sub>2 \<^bold>\ B] \ A\<^sub>1 \<^bold>\ A\<^sub>2 \<^bold>\ B" -abbreviation conseq_local3::"\\\\\\\\bool" ("[_,_,_ \<^bold>\ _]") where "[A\<^sub>1, A\<^sub>2, A\<^sub>3 \<^bold>\ B] \ A\<^sub>1 \<^bold>\ A\<^sub>2 \<^bold>\ A\<^sub>3 \<^bold>\ B" -(*add more as the need arises...*) - -text\\noindent{For LFIs we use the (paraconsistent) closure-based negation previously defined (taking frontier as primitive). }\ -abbreviation cneg::"\\\" ("\<^bold>\") where "\<^bold>\A \ \<^bold>\\<^sup>CA" - -text\\noindent{In terms of the frontier operator the negation looks as follows:}\ -lemma "\<^bold>\A \<^bold>\ \<^bold>\A \<^bold>\ \(\<^bold>\A)" by (simp add: neg_C_def pC2) -lemma cneg_prop: "Fr_2 \ \ \<^bold>\A \<^bold>\ \<^bold>\A \<^bold>\ \(A)" using pC2 pF2 unfolding conn by blast - -text\\noindent{This negation is of course boldly paraconsistent.}\ -lemma "[a, \<^bold>\a \<^bold>\ b]" nitpick oops (*countermodel*) -lemma "[a, \<^bold>\a \<^bold>\ \<^bold>\b]" nitpick oops (*countermodel*) -lemma "[a, \<^bold>\a \<^bold>\\<^sub>g b]" nitpick oops (*countermodel*) -lemma "[a, \<^bold>\a \<^bold>\\<^sub>g \<^bold>\b]" nitpick oops (*countermodel*) - -text\\noindent{We define two pairs of in/consistency operators and show how they relate to each other. -Using LFIs terminology, the minimal logic so encoded corresponds to 'RmbC-ciw' (cf. \<^cite>\RLFI\).}\ -abbreviation op_inc_a::"\\\" ("\\<^sup>A_" [57]58) where "\\<^sup>AA \ A \<^bold>\ \<^bold>\A" -abbreviation op_con_a::"\\\" ("\<^bold>\\<^sup>A_" [57]58) where "\<^bold>\\<^sup>AA \ \<^bold>\\\<^sup>AA" -abbreviation op_inc_b::"\\\" ("\\<^sup>B_" [57]58) where "\\<^sup>BA \ \ A" -abbreviation op_con_b::"\\\" ("\<^bold>\\<^sup>B_" [57]58) where "\<^bold>\\<^sup>BA \ \\<^sup>c A" - -text\\noindent{Observe that assumming condition Fr-2 are we allowed to exchange A and B variants.}\ -lemma pincAB: "Fr_2 \ \ \\<^sup>AA \<^bold>\ \\<^sup>BA" using Br_fr_def Cl_fr_def pF2 conn by auto -lemma pconAB: "Fr_2 \ \ \<^bold>\\<^sup>AA \<^bold>\ \<^bold>\\<^sup>BA" using pincAB unfolding conn by simp - -text\\noindent{Variants A and B give us slightly different properties.}\ -lemma Prop1: "\<^bold>\\<^sup>BA \<^bold>\ \\<^sup>f\<^sup>p A" using fp1 unfolding conn equal_op_def by metis -lemma "\<^bold>\\<^sup>AA \<^bold>\ A \<^bold>\ \ A" nitpick oops (*countermodel*) -lemma Prop2: "Cl A \ \<^bold>\\<^sup>A\<^bold>\A \<^bold>\ \<^bold>\" using pC2 unfolding conn by auto -lemma "Cl A \ \<^bold>\\<^sup>B\<^bold>\A \<^bold>\ \<^bold>\" nitpick oops (*countermodel*) -lemma Prop3: "Cl A \ \\<^sup>A\<^bold>\A \<^bold>\ \<^bold>\" using Cl_fr_def unfolding conn by auto -lemma "Cl A \ \\<^sup>B\<^bold>\A \<^bold>\ \<^bold>\" nitpick oops (*countermodel*) -lemma Prop4: "Op A \ \<^bold>\\<^sup>BA \<^bold>\ \<^bold>\" using Op_Bzero unfolding conn by simp -lemma "Op A \ \<^bold>\\<^sup>AA \<^bold>\ \<^bold>\" nitpick oops (*countermodel*) -lemma Prop5: "Op A \ \\<^sup>BA \<^bold>\ \<^bold>\" using Op_Bzero by simp -lemma "Op A \ \\<^sup>AA \<^bold>\ \<^bold>\" nitpick oops (*countermodel*) - -text\\noindent{Importantly, LFIs must satisfy the so-called 'principle of gentle explosion'. Only variant A works here:}\ -lemma "[\<^bold>\\<^sup>Aa, a, \<^bold>\a \<^bold>\ b]" using compl_def meet_def by auto -lemma "[\<^bold>\\<^sup>Aa, a, \<^bold>\a \<^bold>\\<^sub>g b]" using compl_def meet_def by auto -lemma "[\<^bold>\\<^sup>Ba, a, \<^bold>\a \<^bold>\ b]" nitpick oops (*countermodel*) -lemma "[\<^bold>\\<^sup>Ba, a, \<^bold>\a \<^bold>\\<^sub>g b]" nitpick oops (*countermodel*) - -text\\noindent{In what follows we employ the (minimal) A-variant and verify some properties.}\ -abbreviation op_inc :: "\\\" ("\_" [57]58) where "\A \ \\<^sup>AA" -abbreviation op_con :: "\\\" ("\<^bold>\_" [57]58) where "\<^bold>\A \ \<^bold>\\A" - -lemma "TND(\<^bold>\)" by (simp add: TND_C) -lemma "ECQm(\<^bold>\)" nitpick oops (*countermodel*) -lemma "Fr_1b \ \ Fr_2 \ \ LNC(\<^bold>\)" by (simp add: LNC_C PF6) -lemma "\ \ \ DNI(\<^bold>\)" nitpick oops (*countermodel*) -lemma "\ \ \ DNE(\<^bold>\)" nitpick oops (*countermodel*) -lemma "Fr_1b \ \ Fr_2 \ \ CoPw(\<^bold>\)" by (simp add: CoPw_C PF6) -lemma "\ \ \ CoP1(\<^bold>\)" nitpick oops (*countermodel*) -lemma "\ \ \ CoP2(\<^bold>\)" nitpick oops (*countermodel*) -lemma "\ \ \ CoP3(\<^bold>\)" nitpick oops (*countermodel*) -lemma "Fr_1a \ \ Fr_2 \ \ DM3(\<^bold>\)" by (simp add: DM3_C) -lemma "\ \ \ DM4(\<^bold>\)" nitpick oops (*countermodel*) -lemma "nNor(\<^bold>\)" by (simp add: nNor_C) -lemma "Fr_3 \ \ nDNor(\<^bold>\)" by (simp add: nDNor_C) -lemma "Fr_1b \ \ Fr_2 \ \ MT0(\<^bold>\)" by (simp add: MT0_C PF6) -lemma "Fr_1b \ \ Fr_2 \ \ MT1(\<^bold>\)" by (simp add: MT1_C PF6) -lemma "\ \ \ MT2(\<^bold>\)" nitpick oops (*countermodel*) -lemma "Fr_1b \ \ Fr_2 \ \ Fr_3 \ \ MT3(\<^bold>\)" using MT3_C by auto - -text\\noindent{We show how all local contraposition variants (lCoP) can be recovered using the consistency operator. -Observe that we can recover in the same way other (weaker) properties: CoP, MT and DNI/DNE (local \& global).}\ -lemma "\ \ \ lCoPw(\<^bold>\)(\<^bold>\)" nitpick oops (*countermodel*) -lemma cons_lcop1: "[\<^bold>\b, a \<^bold>\ b \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\a]" using Cl_fr_def conn by auto -lemma "\ \ \ lCoP1(\<^bold>\)(\<^bold>\)" nitpick oops (*countermodel*) -lemma cons_lcop2: "[\<^bold>\b, a \<^bold>\ \<^bold>\b \<^bold>\ b \<^bold>\ \<^bold>\a]" using Cl_fr_def conn by auto -lemma "\ \ \ lCoP2(\<^bold>\)(\<^bold>\)" nitpick oops (*countermodel*) -lemma cons_lcop3: "[\<^bold>\b, \<^bold>\a \<^bold>\ b \<^bold>\ \<^bold>\b \<^bold>\ a]" using Cl_fr_def conn by auto -lemma "\ \ \ lCoP3(\<^bold>\)(\<^bold>\)" nitpick oops (*countermodel*) -lemma cons_lcop4: "[\<^bold>\b, \<^bold>\a \<^bold>\ \<^bold>\b \<^bold>\ b \<^bold>\ a]" using Cl_fr_def conn by auto -text\\noindent{Disjunctive syllogism (DS).}\ -lemma "\ \ \ DS1(\<^bold>\)(\<^bold>\)" nitpick oops (*countermodel*) -lemma cons_ds1: "[\<^bold>\a, a \<^bold>\ b \<^bold>\ \<^bold>\a \<^bold>\ b]" using conn by auto -lemma "DS2(\<^bold>\)(\<^bold>\)" by (metis Cl_fr_def DS2_def compl_def impl_def join_def neg_C_def) -text\\noindent{Further properties.}\ -lemma "[a \<^bold>\ \<^bold>\a \<^bold>\ \<^bold>\(\<^bold>\a)]" by (simp add: pC2 conn) -lemma "\ \ \ [\<^bold>\(\<^bold>\a) \<^bold>\ a \<^bold>\ \<^bold>\a]" nitpick oops (* countermodel found *) -lemma "[\<^bold>\a \<^bold>\ \<^bold>\(a \<^bold>\ \<^bold>\a)]" by (simp add: pC2 conn) -lemma "\ \ \ [\<^bold>\(a \<^bold>\ \<^bold>\a) \<^bold>\ \<^bold>\a]" nitpick oops (* countermodel found *) - -text\\noindent{The following axioms are commonly employed in the literature on LFIs to obtain stronger logics. -We explore under which conditions they can be assumed while keeping the logic boldly paraconsistent.}\ -abbreviation cf where "cf \ DNE(\<^bold>\)" -abbreviation ce where "ce \ DNI(\<^bold>\)" -abbreviation ciw where "ciw \ \P. [\<^bold>\ \<^bold>\P \<^bold>\ \P]" -abbreviation ci where "ci \ \P. [\<^bold>\(\<^bold>\P) \<^bold>\ \P]" -abbreviation cl where "cl \ \P. [\<^bold>\(\P) \<^bold>\ \<^bold>\P]" -abbreviation ca_conj where "ca_conj \ \P Q. [\<^bold>\P,\<^bold>\Q \<^bold>\ \<^bold>\(P \<^bold>\ Q)]" -abbreviation ca_disj where "ca_disj \ \P Q. [\<^bold>\P,\<^bold>\Q \<^bold>\ \<^bold>\(P \<^bold>\ Q)]" -abbreviation ca_impl where "ca_impl \ \P Q. [\<^bold>\P,\<^bold>\Q \<^bold>\ \<^bold>\(P \<^bold>\ Q)]" -abbreviation ca where "ca \ ca_conj \ ca_disj \ ca_impl" - -text\\noindent{cf}\ -lemma "\ \ \ cf" nitpick oops -lemma "\ \ \ cf \ \ECQm(\<^bold>\)" nitpick[satisfy] oops (*model found*) -text\\noindent{ce}\ -lemma "\ \ \ ce" nitpick oops -lemma "Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ ce \ \ECQ(\<^bold>\)" nitpick[satisfy] oops (*model found*) -lemma "Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ ce \ \ECQm(\<^bold>\)" nitpick[satisfy] oops (*model found*) -lemma "Fr_1a \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ ce \ \ECQm(\<^bold>\)" nitpick[satisfy] oops (*model found*) -lemma "Fr_1b \ \ Fr_2 \ \ ce \ ECQm(\<^bold>\)" unfolding Defs using CoP1_XCoP CoP1_def2 CoPw_C DNI_def ECQw_def PF6 XCoP_def2 by auto -text\\noindent{ciw}\ -lemma ciw by (simp add:conn) -text\\noindent{ci}\ -lemma "\ \ \ ci" nitpick oops -lemma "\ \ \ ci \ \ECQm(\<^bold>\)" nitpick[satisfy] oops (*model found*) -text\\noindent{cl}\ -lemma "\ \ \ cl" nitpick oops -lemma "Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ cl \ \ECQm(\<^bold>\)" nitpick[satisfy] oops (*model found*) -lemma "Fr_1a \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ cl \ \ECQm(\<^bold>\)" nitpick[satisfy] oops (*model found*) -lemma "Fr_1b \ \ Fr_2 \ \ cl \ ECQ(\<^bold>\)" unfolding Defs by (metis BC_rel Br_Border Br_cl_def bottom_def compl_def eq_ext' meet_def neg_C_def) -text\\noindent{ca-conj/disj}\ -lemma "Fr_1a \ \ Fr_2 \ \ ca_conj" using DM3_C DM3_def conn by auto -lemma "Fr_1b \ \ Fr_2 \ \ ca_disj" using ADDI_b_def MONO_ADDIb monI pB1 pincAB unfolding conn by metis -lemma "\ \ \ ca_impl" nitpick oops -text\\noindent{ca-impl}\ -lemma "ca_impl \ \ECQ(\<^bold>\)" (*nitpick[satisfy]*) oops (*cannot find finite model*) -lemma "ca_impl \ ECQm(\<^bold>\)" oops (*nor proof yet*) -text\\noindent{cf \& ci}\ -lemma "Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ cf \ ci \ \ECQm(\<^bold>\)" nitpick[satisfy] oops (*model found*) -lemma "Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ cf \ ci \ \ECQm(\<^bold>\)" nitpick[satisfy] oops (*model found*) -lemma "Fr_1b \ \ Fr_2 \ \ cf \ ci \ \ECQ(\<^bold>\)" (*nitpick[satisfy]*) oops (*cannot find finite model*) -lemma "\ \ \ cf \ ci \ ECQm(\<^bold>\)" oops (*nor proof yet*) -text\\noindent{cf \& cl}\ -lemma "Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ cf \ cl \ \ECQm(\<^bold>\)" nitpick[satisfy] oops (*model found*) -lemma "Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ cf \ cl \ \ECQm(\<^bold>\)" nitpick[satisfy] oops (*model found*) -lemma "Fr_1b \ \ Fr_2 \ \ cf \ cl \ ECQ(\<^bold>\)" unfolding Defs by (smt Br_fr_def Fr_1b_def Prop2 Prop3 pF3 cneg_prop conn) - -end diff --git a/thys/Topological_Semantics/ex_LFUs.thy b/thys/Topological_Semantics/ex_LFUs.thy deleted file mode 100644 --- a/thys/Topological_Semantics/ex_LFUs.thy +++ /dev/null @@ -1,84 +0,0 @@ -theory ex_LFUs - imports topo_derivative_algebra sse_operation_negative -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Example application: Logics of Formal Undeterminedness (LFUs)\ -text\\noindent{The LFUs \<^cite>\LFU\ \<^cite>\LFI\ are a family of paracomplete logics featuring a 'determinedness' -operator @{text "\<^bold>\"} that can be used to recover some classical properties of negation (in particular TND). -LFUs behave in a sense dually to LFIs. Both can be semantically embedded as extensions of Boolean algebras. -Here we show how to semantically embed LFUs as derivative algebras.}\ - -text\\noindent{(We rename (classical) meta-logical negation to avoid terminological confusion)}\ -abbreviation cneg::"bool\bool" ("\_" [40]40) where "\\ \ \\" - -text\\noindent{Logical validity can be defined as truth in all worlds/points (i.e. as denoting the top element)}\ -abbreviation gtrue::"\\bool" ("[\<^bold>\ _]") where "[\<^bold>\ A] \ \w. A w" -lemma gtrue_def: "[\<^bold>\ A] \ A \<^bold>\ \<^bold>\" by (simp add: top_def) - -text\\noindent{As for LFIs, we focus on the local (truth-degree preserving) notion of logical consequence.}\ -abbreviation conseq_local1::"\\\\bool" ("[_ \<^bold>\ _]") where "[A \<^bold>\ B] \ A \<^bold>\ B" -abbreviation conseq_local2::"\\\\\\bool" ("[_,_ \<^bold>\ _]") where "[A\<^sub>1, A\<^sub>2 \<^bold>\ B] \ A\<^sub>1 \<^bold>\ A\<^sub>2 \<^bold>\ B" -abbreviation conseq_local12::"\\\\\\bool" ("[_ \<^bold>\ _,_]") where "[A \<^bold>\ B\<^sub>1, B\<^sub>2] \ A \<^bold>\ B\<^sub>1 \<^bold>\ B\<^sub>2" -(*add more as the need arises...*) - -text\\noindent{For LFUs we use the interior-based negation previously defined (taking derivative as primitive). }\ -definition ineg::"\\\" ("\<^bold>\") where "\<^bold>\A \ \(\<^bold>\A)" -declare ineg_def[conn] - -text\\noindent{In terms of the derivative operator the negation looks as follows:}\ -lemma ineg_prop: "\<^bold>\A \<^bold>\ \<^bold>\(\ A) \<^bold>\ A" using Cl_der_def IB_rel Int_br_def eq_ext' pB4 conn by fastforce - -text\\noindent{This negation is of course paracomplete.}\ -lemma "[\<^bold>\ a \<^bold>\ \<^bold>\a]" nitpick oops (*countermodel*) - -text\\noindent{We define two pairs of in/determinedness operators and show how they relate to each other.}\ -abbreviation op_det::"\\\" ("\<^bold>\_" [57]58) where "\<^bold>\A \ \\<^sup>d A" -abbreviation op_ind::"\\\" ("\_" [57]58) where "\A \ \<^bold>\\<^bold>\A" - -lemma op_det_def: "\<^bold>\a \<^bold>\ a \<^bold>\ \<^bold>\a" by (simp add: compl_def diff_def dual_def ineg_def join_def pB1) -lemma Prop1: "\<^bold>\A \<^bold>\ \\<^sup>f\<^sup>p A" by (metis dimp_def eq_ext' fp3) -lemma Prop2: "Op A \ \<^bold>\\<^bold>\A \<^bold>\ \<^bold>\" by (metis dual_def dual_symm pB1 pI1 top_def compl_def diff_def) -lemma Prop3: "Op A \ \\<^bold>\A \<^bold>\ \<^bold>\" by (metis Op_Bzero dual_def dual_symm) -lemma Prop4: "Cl A \ \<^bold>\A \<^bold>\ \<^bold>\" by (metis Prop1 dimp_def top_def) -lemma Prop5: "Cl A \ \A \<^bold>\ \<^bold>\" by (simp add: Prop4 bottom_def compl_def top_def) - -text\\noindent{Analogously as for LFIs, LFUs provide means for recovering the principle of excluded middle.}\ -lemma "[\ \<^bold>\ \a, a \<^bold>\ \<^bold>\a]" using IB_rel Int_br_def compl_def diff_def dual_def eq_ext' ineg_def join_def by fastforce -lemma "[\, \<^bold>\a \<^bold>\ a \<^bold>\ \<^bold>\a]" using dual_def pB1 unfolding conn by auto - -lemma "TNDm(\<^bold>\)" nitpick oops (*countermodel*) -lemma "ECQ(\<^bold>\)" by (simp add: ECQ_def bottom_def diff_def ineg_prop meet_def) -lemma "Der_3 \ \ LNC(\<^bold>\)" using ineg_prop ECQ_def ID3 LNC_def dNOR_def unfolding conn by auto -lemma "\
\ \ DNI(\<^bold>\)" nitpick oops (*countermodel*) -lemma "\
\ \ DNE(\<^bold>\)" nitpick oops (*countermodel*) -lemma "Der_1b \ \ CoPw(\<^bold>\)" by (smt CoPw_def MONO_ADDIb PD1 compl_def ineg_def monI) -lemma "\
\ \ CoP1(\<^bold>\)" nitpick oops (*countermodel*) -lemma "\
\ \ CoP2(\<^bold>\)" nitpick oops (*countermodel*) -lemma "\
\ \ CoP3(\<^bold>\)" nitpick oops (*countermodel*) -lemma "\
\ \ DM3(\<^bold>\)" nitpick oops (*countermodel*) -lemma "Der_1a \ \ DM4(\<^bold>\)" unfolding Defs using ADDI_a_def ineg_prop compl_def diff_def join_def meet_def by auto -lemma "Der_3 \ \ nNor(\<^bold>\)" by (simp add: NOR_def ineg_prop nNor_def bottom_def compl_def diff_def top_def) -lemma "nDNor(\<^bold>\)" by (simp add: bottom_def diff_def ineg_prop nDNor_def top_def) -lemma "Der_1b \ \ MT0(\<^bold>\)" unfolding Defs by (metis (mono_tags, opaque_lifting) CD1b Disj_I OpCldual PD1 bottom_def compl_def ineg_def meet_def top_def) -lemma "Der_1b \ \ Der_3 \ \ MT1(\<^bold>\)" unfolding Defs by (metis (full_types) NOR_def PD1 bottom_def compl_def diff_def ineg_prop top_def) -lemma "\
\ \ MT2(\<^bold>\)" nitpick oops (*countermodel*) -lemma "\
\ \ MT3(\<^bold>\)" nitpick oops (*countermodel*) - -text\\noindent{We show how all local contraposition variants (lCoP) can be recovered using the determinedness operator. -Observe that we can recover in the same way other (weaker) properties: CoP, MT and DNI/DNE (local \& global).}\ -lemma "\
\ \ lCoPw(\<^bold>\)(\<^bold>\)" nitpick oops (*countermodel*) -lemma det_lcop1: "[\<^bold>\a, a \<^bold>\ b \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\a]" using dual_def pI1 conn by auto -lemma "\
\ \ lCoP1(\<^bold>\)(\<^bold>\)" nitpick oops (*countermodel*) -lemma det_lcop2: "[\<^bold>\a, a \<^bold>\ \<^bold>\b \<^bold>\ b \<^bold>\ \<^bold>\a]" using dual_def pI1 conn by auto -lemma "\
\ \ lCoP2(\<^bold>\)(\<^bold>\)" nitpick oops (*countermodel*) -lemma det_lcop3: "[\<^bold>\a, \<^bold>\a \<^bold>\ b \<^bold>\ \<^bold>\b \<^bold>\ a]" using dual_def pI1 conn by auto -lemma "\
\ \ lCoP3(\<^bold>\)(\<^bold>\)" nitpick oops (*countermodel*) -lemma det_lcop4: "[\<^bold>\a, \<^bold>\a \<^bold>\ \<^bold>\b \<^bold>\ b \<^bold>\ a]" using dual_def pI1 conn by auto - -text\\noindent{Disjunctive syllogism (DS).}\ -lemma "DS1(\<^bold>\)(\<^bold>\)" by (simp add: DS1_def diff_def impl_def ineg_prop join_def) -lemma "\
\ \ DS2(\<^bold>\)(\<^bold>\)" nitpick oops (*countermodel*) -lemma det_ds2: "[\<^bold>\a, \<^bold>\a \<^bold>\ b \<^bold>\ a \<^bold>\ b]" using pB1 dual_def conn by auto - -end diff --git a/thys/Topological_Semantics/ex_subminimal_logics.thy b/thys/Topological_Semantics/ex_subminimal_logics.thy deleted file mode 100644 --- a/thys/Topological_Semantics/ex_subminimal_logics.thy +++ /dev/null @@ -1,220 +0,0 @@ -theory ex_subminimal_logics - imports topo_negation_conditions topo_strict_implication -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Example application: Subintuitionistic and subminimal logics\ - -text\\noindent{In this section we examine some special paracomplete logics. The idea is to illustrate an approach by -means of which we can obtain logics which are boldly paracomplete and (non-boldly) paraconsistent at the -same time, Johansson's 'minimal logic' \<^cite>\JML\ being the paradigmatic case we aim at modelling.}\ - -text\\noindent{Drawing upon the literature on Johanson's minimal logic, we introduce an unconstrained propositional -constant Q, which we then employ to define a 'rigid' frontier operation @{text "\'"}.}\ -consts Q::"\" -abbreviation "\' \ \X. Q" -abbreviation "\' \ \\<^sub>F \'" -abbreviation "\' \ \\<^sub>F \'" -abbreviation "\' \ \\<^sub>F \'" - -text\\noindent{As defined, @{text "\'"} (and its corresponding closure operation) satisfies several semantic conditions.}\ -lemma "Fr_1 \' \ Fr_2 \' \ Fr_4 \'" by (simp add: Fr_1_def Fr_2_def Fr_4_def conn) -lemma "Cl_1 \' \ Cl_2 \' \ Cl_4 \'" using ADDI_def CF2 IDEMb_def Cl_fr_def PC4 unfolding conn by auto -text\\noindent{However Fr-3 is not valid. In fact, adding it by hand would collapse into classical logic (making all sets clopen).}\ -lemma "Fr_3 \'" nitpick oops (*counterexample found*) -lemma "Cl_3 \'" nitpick oops (*counterexample found*) -lemma "Fr_3 \' \ \A. \'(A) \<^bold>\ \<^bold>\" by (simp add: NOR_def) - -text\\noindent{In order to obtain a paracomplete logic not validating ECQ, we define negation as follows,}\ -abbreviation neg_IC::"\\\" ("\<^bold>\") where "\<^bold>\A \ \'(\(\<^bold>\A))" - -text\\noindent{and observe that some plausible semantic properties obtain:}\ -lemma Q_def1: "\A. Q \<^bold>\ \<^bold>\A \<^bold>\ \<^bold>\(\<^bold>\A)" using Cl_fr_def IF2 dEXP_def conn by auto -lemma Q_def2: "Fr_1b \ \ \A. Q \<^bold>\ \<^bold>\(A \<^bold>\ \<^bold>\A)" by (smt Cl_fr_def IF2 dEXP_def MONO_def monI conn) -lemma neg_Idef: "\A. \<^bold>\A \<^bold>\ \(\<^bold>\A) \<^bold>\ Q" by (simp add: Cl_fr_def) -lemma neg_Cdef: "Fr_2 \ \ \A. \<^bold>\A \<^bold>\ \(A) \<^bold>\ Q" using Cl_fr_def Fr_2_def Int_fr_def conn by auto - -text\\noindent{The negation so defined validates some properties corresponding to a (rather weak) paracomplete logic:}\ -lemma "\ \ \ TND \<^bold>\" nitpick oops (*counterexample found: negation is paracomplete*) -lemma "\ \ \ TNDw \<^bold>\" nitpick oops -lemma "\ \ \ TNDm \<^bold>\" nitpick oops -lemma "\ \ \ ECQ \<^bold>\" nitpick oops (*counterexample found: negation is paraconsistent...*) -lemma ECQw: "ECQw \<^bold>\" using Cl_fr_def Disj_I ECQw_def unfolding conn by auto (*...but not 'boldly' paraconsistent*) -lemma ECQm: "ECQm \<^bold>\" using Cl_fr_def Disj_I ECQm_def unfolding conn by auto -lemma "\ \ \ LNC \<^bold>\" nitpick oops -lemma "\ \ \ DNI \<^bold>\" nitpick oops -lemma "\ \ \ DNE \<^bold>\" nitpick oops -lemma CoPw: "Fr_1b \ \ CoPw \<^bold>\" using Cl_fr_def MONO_def monI unfolding Defs conn by smt -lemma "\ \ \ CoP1 \<^bold>\" nitpick oops -lemma "\ \ \ CoP2 \<^bold>\" nitpick oops -lemma "\ \ \ CoP3 \<^bold>\" nitpick oops -lemma "\ \ \ XCoP \<^bold>\" nitpick oops -lemma "\ \ \ DM3 \<^bold>\" nitpick oops -lemma DM4: "Fr_1a \ \ DM4 \<^bold>\" using ADDI_a_def Cl_fr_def DM4_def IC1b IF1b dual_def unfolding conn by smt -lemma Nor: "Fr_2 \ \ Fr_3 \ \ nNor \<^bold>\" using Cl_fr_def nNor_I nNor_def unfolding conn by auto -lemma "\ \ \ nDNor \<^bold>\" nitpick oops -lemma "\ \ \ lCoPw(\<^bold>\) \<^bold>\" nitpick oops -lemma "\ \ \ lCoP1(\<^bold>\) \<^bold>\" nitpick oops -lemma "\ \ \ lCoP2(\<^bold>\) \<^bold>\" nitpick oops -lemma "\ \ \ lCoP3(\<^bold>\) \<^bold>\" nitpick oops -lemma "\ \ \ DS1(\<^bold>\) \<^bold>\" nitpick oops -lemma "\ \ \ DS2(\<^bold>\) \<^bold>\" nitpick oops - -text\\noindent{Moreover, we cannot have both DNI and DNE without validating ECQ (thus losing paraconsistency).}\ -lemma "DNI \<^bold>\ \ DNE \<^bold>\ \ ECQ \<^bold>\" using DNE_def ECQ_def Int_fr_def neg_Idef unfolding conn by (metis (no_types, lifting)) -text\\noindent{However, we can have all of De Morgan laws while keeping (non-bold) paraconsistency.}\ -lemma "\ECQ \<^bold>\ \ DM1 \<^bold>\ \ DM2 \<^bold>\ \ DM3 \<^bold>\ \ DM4 \<^bold>\ \ \ \" nitpick[satisfy,card w=3] oops (*(weakly paraconsistent) model found*) - -text\\noindent{Below we combine negation with strict implication and verify some interesting properties. -For instance, the following are not valid (and cannot become valid by adding semantic restrictions). }\ -lemma "\ \ \ \a b. (\<^bold>\a \<^bold>\ (a \<^bold>\ b)) \<^bold>\ \<^bold>\" nitpick oops (*counterexample found*) -lemma "\ \ \ \a b. (\<^bold>\a \<^bold>\ (a \<^bold>\ b)) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. (a \<^bold>\ \<^bold>\a \<^bold>\ b) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. (a \<^bold>\ \<^bold>\a \<^bold>\ b) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. (a \<^bold>\ (b \<^bold>\ \<^bold>\b)) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. (a \<^bold>\ (b \<^bold>\ \<^bold>\b)) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a. (a \<^bold>\ \<^bold>\a) \<^bold>\ \<^bold>\a \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a. (a \<^bold>\ \<^bold>\a) \<^bold>\ \<^bold>\a \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. (a \<^bold>\ \<^bold>\a) \<^bold>\ b \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. (a \<^bold>\ \<^bold>\a) \<^bold>\ b \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. a \<^bold>\ (b \<^bold>\ \<^bold>\b) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. a \<^bold>\ (b \<^bold>\ \<^bold>\b) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. (a \<^bold>\ b) \<^bold>\ (\<^bold>\a \<^bold>\ \<^bold>\b) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. (a \<^bold>\ b) \<^bold>\ (\<^bold>\a \<^bold>\ \<^bold>\b) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. (a \<^bold>\ b) \<^bold>\ (a \<^bold>\ \<^bold>\b) \<^bold>\ \<^bold>\a \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a b. (a \<^bold>\ b) \<^bold>\ (a \<^bold>\ \<^bold>\b) \<^bold>\ \<^bold>\a \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a. (\<^bold>\a \<^bold>\ \<^bold>\) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a. (\<^bold>\a \<^bold>\ \<^bold>\) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a. (\<^bold>\a \<^bold>\ \<^bold>\(\<^bold>\\<^bold>\)) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a. (\<^bold>\a \<^bold>\ \<^bold>\(\<^bold>\\<^bold>\)) \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a. \<^bold>\(\<^bold>\(\<^bold>\a)) \<^bold>\ \<^bold>\a \<^bold>\ \<^bold>\" nitpick oops -lemma "\ \ \ \a. \<^bold>\(\<^bold>\(\<^bold>\a)) \<^bold>\ \<^bold>\a \<^bold>\ \<^bold>\" nitpick oops - -text\\noindent{The (weak) local contraposition axiom is indeed valid under appropriate conditions.}\ -lemma lCoPw: "Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ lCoPw(\<^bold>\) \<^bold>\" proof - - assume fr1: "Fr_1 \" and fr2: "Fr_2 \" and fr3: "Fr_3 \" and fr4: "Fr_4 \" - { fix a b - from fr2 have "\<^bold>\b \<^bold>\ \<^bold>\a \<^bold>\ (\ a \<^bold>\ \ b) \<^bold>\ Q" using Cl_fr_def Fr_2_def Int_fr_def conn by auto - moreover from fr1 fr2 fr3 have "\(a \<^bold>\ b) \<^bold>\ \ a \<^bold>\ \ b" using IC_imp by simp - ultimately have "\(a \<^bold>\ b) \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\a" unfolding conn by simp - moreover from fr1 fr2 fr4 have "let A=(a \<^bold>\ b); B=(\<^bold>\b \<^bold>\ \<^bold>\a) in \ A \<^bold>\ B \ \ A \<^bold>\ \ B" - using PF1 MONO_MULTa IF1a IF4 PI9 Int_9_def by smt - ultimately have "\(a \<^bold>\ b) \<^bold>\ \(\<^bold>\b \<^bold>\ \<^bold>\a)" by simp - } hence "lCoPw(\<^bold>\) \<^bold>\" unfolding Defs conn by blast - thus ?thesis by simp -qed -lemma lCoPw_strict: "\ \ \ \a b. (a \<^bold>\ b) \<^bold>\ (\<^bold>\b \<^bold>\ \<^bold>\a) \<^bold>\ \<^bold>\" by (metis (no_types, lifting) DTw2 lCoPw lCoPw_def) - -text\\noindent{However, other (local) contraposition axioms are not valid.}\ -lemma "\ \ \ lCoP1(\<^bold>\) \<^bold>\" nitpick oops (*counterexample found*) -lemma "\ \ \ lCoP2(\<^bold>\) \<^bold>\" nitpick oops -lemma "\ \ \ lCoP3(\<^bold>\) \<^bold>\" nitpick oops -text\\noindent{And this time no variant of disjunctive syllogism is valid.}\ -lemma "\ \ \ DS1(\<^bold>\) \<^bold>\" nitpick oops -lemma "\ \ \ DS2(\<^bold>\) \<^bold>\" nitpick oops -lemma "\ \ \ DS2(\<^bold>\) \<^bold>\" nitpick oops -lemma "\ \ \ DS4(\<^bold>\) \<^bold>\" nitpick oops - -text\\noindent{Interestingly, one of the local contraposition axioms (lCoP1) follows from DNI.}\ -lemma DNI_lCoP1: "Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ DNI \<^bold>\ \ lCoP1(\<^bold>\) \<^bold>\" proof - - assume fr1: "Fr_1 \" and fr2: "Fr_2 \" and fr3: "Fr_3 \" and fr4: "Fr_4 \" - { assume dni: "DNI \<^bold>\" - { fix a b - from fr1 fr2 fr3 fr4 have "lCoPw(\<^bold>\) \<^bold>\" using lCoPw by simp - hence 1: "a \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\(\<^bold>\b) \<^bold>\ \<^bold>\a" unfolding lCoPw_def by simp - from fr1 have 2: "let A=b; B=\<^bold>\(\<^bold>\b); C=\<^bold>\a in A \<^bold>\ B \ \(B \<^bold>\ C) \<^bold>\ \(A \<^bold>\ C)" by (simp add: MONO_ant PF1 monI) - from dni have dnib: "b \<^bold>\ \<^bold>\(\<^bold>\b)" unfolding DNI_def by simp - from 1 2 dnib have "a \<^bold>\ \<^bold>\b \<^bold>\ b \<^bold>\ \<^bold>\a" unfolding conn by meson - } hence "lCoP1(\<^bold>\) \<^bold>\" unfolding Defs by blast - } thus ?thesis by simp -qed -text\\noindent{This entails some other interesting results.}\ -lemma DNI_CoP1: "Fr_1b \ \ DNI \<^bold>\ \ CoP1 \<^bold>\" using CoP1_def2 CoPw by blast -lemma CoP1_LNC: "CoP1 \<^bold>\ \ LNC \<^bold>\" using CoP1_def ECQm_def LNC_def Cl_fr_def Disj_I ECQm_def unfolding conn by smt -lemma DNI_LNC: "Fr_1b \ \ DNI \<^bold>\ \ LNC \<^bold>\" by (simp add: CoP1_LNC DNI_CoP1) - -text\\noindent{The following variants of modus tollens also obtain.}\ -lemma MT: "Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ \a b. (a \<^bold>\ b) \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\a" using Cl_fr_def Fr_2_def IC_imp Int_fr_def unfolding conn by metis -lemma MT': "Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ \a b. ((a \<^bold>\ b) \<^bold>\ \<^bold>\b) \<^bold>\ \<^bold>\a \<^bold>\ \<^bold>\" by (simp add: DTw2 MT) - -text\\noindent{We now semantically characterize (an approximation of) Johansson's Minimal Logic along with some -exemplary 'subminimal' logics (observing that many more are possible). We check some relevant properties.}\ -abbreviation "JML \ \ \ \ DNI \<^bold>\" -abbreviation "SML1 \ \ \" (*Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \*) -abbreviation "SML2 \ Fr_1 \ \ Fr_2 \ \ Fr_3 \" -abbreviation "SML3 \ Fr_1 \" -abbreviation "SML4 \ Fr_1b \" - -text\\noindent{TND:}\ -lemma "JML \ TND \<^bold>\" nitpick oops (*counterexample found*) -lemma "JML \ TNDw \<^bold>\" nitpick oops -lemma "JML \ TNDm \<^bold>\" nitpick oops - -text\\noindent{ECQ:}\ -lemma "JML \ ECQ \<^bold>\" nitpick oops -lemma "ECQw \<^bold>\" using Cl_fr_def Disj_I ECQw_def unfolding conn by auto -lemma "ECQm \<^bold>\" using Cl_fr_def Disj_I ECQm_def unfolding conn by auto - -text\\noindent{LNC:}\ -lemma "JML \ LNC \<^bold>\" using DNI_LNC PF1 by blast -lemma "SML1 \ LNC \<^bold>\" nitpick oops - -text\\noindent{(r)DNI/DNE:}\ -lemma "JML \ DNI \<^bold>\" using CoP1_def2 by blast -lemma "SML1 \ rDNI \<^bold>\" nitpick oops -lemma "JML \ rDNE \<^bold>\" nitpick oops - -text\\noindent{CoP/MT:}\ -lemma "SML4 \ CoPw \<^bold>\" unfolding Defs by (smt Cl_fr_def MONO_def monI conn) -lemma "JML \ CoP1 \<^bold>\" using DNI_CoP1 PF1 by blast -lemma "SML1 \ MT1 \<^bold>\" nitpick oops -lemma "JML \ MT2 \<^bold>\" nitpick oops -lemma "JML \ MT3 \<^bold>\" nitpick oops - -text\\noindent{XCoP:}\ -lemma "JML \ XCoP \<^bold>\" nitpick oops - -text\\noindent{DM3/4:}\ -lemma "JML \ DM3 \<^bold>\" nitpick oops -lemma "SML3 \ DM4 \<^bold>\" by (simp add: DM4 PF1) -lemma "SML4 \ DM4 \<^bold>\" nitpick oops - -text\\noindent{nNor/nDNor:}\ -lemma "SML2 \ nNor \<^bold>\" using Cl_fr_def nNor_I nNor_def unfolding conn by auto -lemma "SML3 \ nNor \<^bold>\" nitpick oops -lemma "JML \ nDNor \<^bold>\" nitpick oops - -text\\noindent{lCoP classical:}\ -lemma "JML \ lCoPw(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ lCoP1(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ lCoP2(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ lCoP3(\<^bold>\) \<^bold>\" nitpick oops - -text\\noindent{DS classical:}\ -lemma "JML \ DS1(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ DS2(\<^bold>\) \<^bold>\" nitpick oops - -text\\noindent{lCoP strict:}\ -lemma "SML1 \ lCoPw(\<^bold>\) \<^bold>\" using lCoPw by blast -lemma "SML2 \ lCoPw(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ lCoP1(\<^bold>\) \<^bold>\" using CoP1_def2 DNI_lCoP1 by blast -lemma "SML1 \ lCoP1(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ lCoP2(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ lCoP3(\<^bold>\) \<^bold>\" nitpick oops - -text\\noindent{lMT strict:}\ -lemma "SML2 \ lMT0(\<^bold>\) \<^bold>\" unfolding Defs using MT by auto -lemma "SML3 \ lMT0(\<^bold>\) \<^bold>\" (*nitpick*) oops (*no countermodel found*) -lemma "SML4 \ lMT0(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ lMT1(\<^bold>\) \<^bold>\" by (smt DNI_lCoP1 DT1 lCoP1_def lMT1_def) -lemma "SML1 \ lMT1(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ lMT2(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ lMT3(\<^bold>\) \<^bold>\" nitpick oops - -text\\noindent{DS strict:}\ -lemma "JML \ DS1(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ DS2(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ DS3(\<^bold>\) \<^bold>\" nitpick oops -lemma "JML \ DS4(\<^bold>\) \<^bold>\" nitpick oops - -end diff --git a/thys/Topological_Semantics/logics_LFI.thy b/thys/Topological_Semantics/logics_LFI.thy --- a/thys/Topological_Semantics/logics_LFI.thy +++ b/thys/Topological_Semantics/logics_LFI.thy @@ -1,131 +1,131 @@ theory logics_LFI imports logics_consequence conditions_relativized_infinitary begin subsection \Logics of Formal Inconsistency (LFIs)\ -text\The LFIs are a family of paraconsistent logics featuring a 'consistency' operator \ +text\The LFIs are a family of paraconsistent logics featuring a 'consistency' operator @{text "\"} that can be used to recover some classical properties of negation (in particular ECQ). We show a shallow semantical embedding of a family of self-extensional LFIs using the border operator as primitive.\ text\Let us assume a concrete type w (for 'worlds' or 'points').\ typedecl w -text\Let us assume the following primitive unary operation \ (intended as a border operator).\ +text\Let us assume the following primitive unary operation intended as a border operator.\ consts \::"w \ \ w \" text\From the topological cube of opposition we have that:\ abbreviation "\ \ (\\<^sup>f\<^sup>p)\<^sup>d\<^sup>-" abbreviation "\ \ \\<^sup>f\<^sup>p\<^sup>-" lemma "\\<^sup>d\<^sup>- = \\<^sup>f\<^sup>p" by (simp add: dualcompl_invol) text\Let us recall that:\ lemma expn_cntr: "EXPN \ = CNTR \" by (metis EXPN_CNTR_dual2 EXPN_fp ofp_comm_dc1) text\For LFIs we use the negation previously defined as @{text "\\<^sup>d\<^sup>- = \\<^sup>f\<^sup>p"}.\ abbreviation neg ("\<^bold>\_"[70]71) where "neg \ \\<^sup>f\<^sup>p" text\In terms of the border operator the negation looks as follows (under appropriate assumptions):\ lemma neg_char: "CNTR \ \ \<^bold>\A = (\<^bold>\A \<^bold>\ \ A)" unfolding conn by (metis CNTR_def dimpl_def op_fixpoint_def subset_def) text\This negation is of course boldly paraconsistent (for both local and global consequence).\ lemma "[a, \<^bold>\a \ b]" nitpick oops \\ countermodel \ lemma "[a, \<^bold>\a \ \<^bold>\b]" nitpick oops \\ countermodel \ lemma "[a, \<^bold>\a \\<^sub>g b]" nitpick oops \\ countermodel \ lemma "[a, \<^bold>\a \\<^sub>g \<^bold>\b]" nitpick oops \\ countermodel \ text\We define two pairs of in/consistency operators and show how they relate to each other. Using LFIs terminology, the minimal logic so encoded corresponds to RmbC + 'ciw' axiom.\ abbreviation op_inc_a::"w \ \ w \" ("\\<^sup>A_" [57]58) (* \ as truth-glut *) where "\\<^sup>AA \ A \<^bold>\ \<^bold>\A" abbreviation op_con_a::"w \ \ w \" ("\\<^sup>A_" [57]58) where "\\<^sup>AA \ \<^bold>\\\<^sup>AA" abbreviation op_inc_b::"w \ \ w \" ("\\<^sup>B_" [57]58) (* \ as border *) where "\\<^sup>BA \ \ A" abbreviation op_con_b::"w \ \ w \" ("\\<^sup>B_" [57]58) where "\\<^sup>BA \ \\<^sup>- A" -text\Observe that assumming CNTR \ we are allowed to exchange A and B variants.\ +text\Observe that assumming CNTR for border we are allowed to exchange A and B variants.\ lemma pincAB: "CNTR \ \ \\<^sup>AA = \\<^sup>BA" using neg_char by (metis CNTR_def CNTR_fpc L5 L6 L9 dimpl_char impl_char ofp_invol op_fixpoint_def setequ_ext svfun_compl_def) lemma pconAB: "CNTR \ \ \\<^sup>AA = \\<^sup>BA" by (metis pincAB setequ_ext svfun_compl_def) text\Variants A and B give us slightly different properties (there are countermodels for those not shown).\ lemma Prop1: "\\<^sup>BA = \\<^sup>f\<^sup>p A" by (simp add: ofp_comm_compl ofp_invol) lemma Prop2: "\\<^sup>AA = (A \<^bold>\ \ A)" by (simp add: compl_def impl_def meet_def svfun_compl_def) lemma Prop3: "fp \ A \ [\ \\<^sup>B\<^bold>\A]" by (simp add: fp_rel gtrue_def ofp_comm_dc2 ofp_invol op_dual_def svfun_compl_def) lemma Prop4a: "fp \ A \ [\ \\<^sup>BA]" by (simp add: fp_rel gtrue_def ofp_comm_compl ofp_invol) lemma Prop4b: "fp \ A \ [\ \\<^sup>AA]" by (simp add: Prop2 fixpoints_def impl_def setequ_ext) text\The 'principle of gentle explosion' works for both variants (both locally and globally).\ lemma "[\\<^sup>Aa, a, \<^bold>\a \ b]" by (metis (mono_tags, lifting) compl_def meet_def subset_def) lemma "[\\<^sup>Aa, a, \<^bold>\a \\<^sub>g b]" by (metis compl_def meet_def) lemma "[\\<^sup>Ba, a, \<^bold>\a \ b]" by (smt (z3) meet_def ofp_fixpoint_compl_def ofp_invol sdiff_def subset_def) lemma "[\\<^sup>Ba, a, \<^bold>\a \\<^sub>g b]" by (metis compl_def fixpoints_def fp_rel gtrue_def setequ_ext svfun_compl_def) abbreviation "BORDER \ \ nMULTr \ \ CNTR \ \ nDNRM \ \ nIDEMr\<^sup>b \" text\We show how (local) contraposition variants (among others) can be recovered using the consistency operators.\ lemma "[\\<^sup>Ab, a \<^bold>\ b \ \<^bold>\b \<^bold>\ \<^bold>\a]" nitpick oops \\ countermodel \ lemma cons_lcop0_A: "CNTR \ \ [\\<^sup>Ab, a \<^bold>\ b \ \<^bold>\b \<^bold>\ \<^bold>\a]" by (smt (verit, del_insts) neg_char compl_def impl_char join_def meet_def subset_def) lemma "[\\<^sup>Bb, a \<^bold>\ b \ \<^bold>\b \<^bold>\ \<^bold>\a]" nitpick oops \\ countermodel \ lemma cons_lcop0_B: "CNTR \ \ [\\<^sup>Bb, a \<^bold>\ b \ \<^bold>\b \<^bold>\ \<^bold>\a]" by (metis cons_lcop0_A pconAB) lemma "[\\<^sup>Ab, a \<^bold>\ \<^bold>\b \ b \<^bold>\ \<^bold>\a]" nitpick oops \\ countermodel \ lemma cons_lcop1_A: "CNTR \ \ [\\<^sup>Ab, a \<^bold>\ \<^bold>\b \ b \<^bold>\ \<^bold>\a]" by (smt (verit, del_insts) neg_char compl_def impl_char join_def meet_def subset_def) lemma "[\\<^sup>Bb, a \<^bold>\ \<^bold>\b \ b \<^bold>\ \<^bold>\a]" nitpick oops \\ countermodel \ lemma cons_lcop1_B: "CNTR \ \ [\\<^sup>Bb, a \<^bold>\ \<^bold>\b \ b \<^bold>\ \<^bold>\a]" by (metis cons_lcop1_A pconAB) lemma "[\\<^sup>Ab, \<^bold>\a \<^bold>\ b \ \<^bold>\b \<^bold>\ a]" nitpick oops \\ countermodel \ lemma cons_lcop2_A: "CNTR \ \ [\\<^sup>Ab, \<^bold>\a \<^bold>\ b \ \<^bold>\b \<^bold>\ a]" by (smt (verit, del_insts) neg_char compl_def impl_char join_def meet_def subset_def) lemma "[\\<^sup>Bb, \<^bold>\a \<^bold>\ b \ \<^bold>\b \<^bold>\ a]" nitpick oops \\ countermodel \ lemma cons_lcop2_B: "CNTR \ \ [\\<^sup>Bb, \<^bold>\a \<^bold>\ b \ \<^bold>\b \<^bold>\ a]" by (metis cons_lcop2_A pconAB) text\The following axioms are commonly employed in the literature on LFIs to obtain stronger logics. We explore under which conditions they can be assumed while keeping the logic boldly paraconsistent.\ abbreviation cf where "cf \ \P. [\<^bold>\\<^bold>\P \ P]" abbreviation ce where "ce \ \P. [P \ \<^bold>\\<^bold>\P]" abbreviation ciw_a where "ciw_a \ \P. [\ \\<^sup>AP \<^bold>\ \\<^sup>AP]" abbreviation ciw_b where "ciw_b \ \P. [\ \\<^sup>BP \<^bold>\ \\<^sup>BP]" abbreviation ci_a where "ci_a \ \P. [\<^bold>\(\\<^sup>AP) \ \\<^sup>AP]" abbreviation ci_b where "ci_b \ \P. [\<^bold>\(\\<^sup>BP) \ \\<^sup>BP]" abbreviation cl_a where "cl_a \ \P. [\<^bold>\(\\<^sup>AP) \ \\<^sup>AP]" abbreviation cl_b where "cl_b \ \P. [\<^bold>\(\\<^sup>BP) \ \\<^sup>BP]" abbreviation ca_conj_a where "ca_conj_a \ \P Q. [\\<^sup>AP,\\<^sup>AQ \ \\<^sup>A(P \<^bold>\ Q)]" abbreviation ca_conj_b where "ca_conj_b \ \P Q. [\\<^sup>BP,\\<^sup>BQ \ \\<^sup>B(P \<^bold>\ Q)]" abbreviation ca_disj_a where "ca_disj_a \ \P Q. [\\<^sup>AP,\\<^sup>AQ \ \\<^sup>A(P \<^bold>\ Q)]" abbreviation ca_disj_b where "ca_disj_b \ \P Q. [\\<^sup>BP,\\<^sup>BQ \ \\<^sup>B(P \<^bold>\ Q)]" abbreviation ca_impl_a where "ca_impl_a \ \P Q. [\\<^sup>AP,\\<^sup>AQ \ \\<^sup>A(P \<^bold>\ Q)]" abbreviation ca_impl_b where "ca_impl_b \ \P Q. [\\<^sup>BP,\\<^sup>BQ \ \\<^sup>B(P \<^bold>\ Q)]" abbreviation ca_a where "ca_a \ ca_conj_a \ ca_disj_a \ ca_impl_a" abbreviation ca_b where "ca_b \ ca_conj_b \ ca_disj_b \ ca_impl_b" text\cf\ lemma "BORDER \ \ cf" nitpick oops \\ countermodel \ text\ce\ lemma "BORDER \ \ ce" nitpick oops \\ countermodel \ text\ciw\ lemma prop_ciw_a: "ciw_a" by (simp add: conn) lemma prop_ciw_b: "ciw_b" by (simp add: conn svfun_compl_def) text\ci\ lemma "BORDER \ \ ci_a" nitpick oops \\ countermodel \ lemma "BORDER \ \ ci_b" nitpick oops \\ countermodel \ text\cl\ lemma "BORDER \ \ cl_a" nitpick oops \\ countermodel \ lemma "BORDER \ \ cl_b" nitpick oops \\ countermodel \ -text\ca_conj\ +text\ca-conj\ lemma prop_ca_conj_b: "nMULT\<^sup>b \ = ca_conj_b" by (metis MULT_b_def nMULTb_compl sfun_compl_invol) lemma prop_ca_conj_a: "nMULTr\<^sup>b \ = ca_conj_a" unfolding cond op_fixpoint_def by (smt (z3) compl_def dimpl_def join_def meet_def op_fixpoint_def subset_def subset_in_def) -text\ca_disj\ +text\ca-disj\ lemma prop_ca_disj_b: "ADDI\<^sup>a \ = ca_disj_b" by (simp add: nADDI_a_def nADDIa_compl) lemma prop_ca_disj_a: "nMULTr\<^sup>a \ = ca_disj_a" oops (*TODO*) -text\ca_impl\ +text\ca-impl\ lemma "BORDER \ \ ca_impl_a" nitpick oops \\ countermodel \ lemma "BORDER \ \ ca_impl_b" nitpick oops \\ countermodel \ end diff --git a/thys/Topological_Semantics/logics_LFU.thy b/thys/Topological_Semantics/logics_LFU.thy --- a/thys/Topological_Semantics/logics_LFU.thy +++ b/thys/Topological_Semantics/logics_LFU.thy @@ -1,81 +1,81 @@ theory logics_LFU imports logics_consequence conditions_relativized_infinitary begin subsection \Logics of Formal Undeterminedness (LFUs)\ -text\The LFUs are a family of paracomplete logics featuring a 'determinedness' operator \ +text\The LFUs are a family of paracomplete logics featuring a 'determinedness' operator @{text "\"} that can be used to recover some classical properties of negation (in particular TND). LFUs behave in a sense dually to LFIs. Both can be semantically embedded as extensions of Boolean algebras. We show a shallow semantical embedding of a family of self-extensional LFUs using the closure operator as primitive.\ (*Let us assume a concrete type w (for 'worlds' or 'points').*) typedecl w -(*Let us assume the following primitive unary operation \ (intended as a closure operator).*) +(*Let us assume the following primitive unary operation intended as a closure operator.*) consts \::"w \ \ w \" (*From the topological cube of opposition:*) abbreviation "\ \ \\<^sup>d" abbreviation "\ \ (\\<^sup>f\<^sup>p)\<^sup>d" (*let us recall that: *) lemma "EXPN \ = CNTR \" using EXPN_CNTR_dual1 EXPN_fp by blast lemma "EXPN \ = CNTR \" by (simp add: EXPN_CNTR_dual1) text\For LFUs we use the negation previously defined as @{text "\\<^sup>d\<^sup>- = \\<^sup>-"}.\ abbreviation neg ("\<^bold>\_"[70]71) where "neg \ \\<^sup>-" text\In terms of the border operator the negation looks as follows:\ lemma neg_char: "EXPN \ \ \<^bold>\A = (\<^bold>\A \<^bold>\ \\<^sup>d A)" unfolding conn by (metis EXPN_def compl_def dimpl_def dual_invol op_fixpoint_def subset_def svfun_compl_def) abbreviation "CLOSURE \ \ ADDI \ \ EXPN \ \ NORM \ \ IDEM \" text\This negation is of course paracomplete.\ lemma "CLOSURE \ \ [\ a \<^bold>\ \<^bold>\a]" nitpick oops \\ countermodel \ text\We define two pairs of un/determinedness operators and show how they relate to each other. This logic corresponds to the paracomplete dual of the LFI 'RmbC-ciw'.\ abbreviation op_det_a::"w \ \ w \" ("\\<^sup>A_" [57]58) where "\\<^sup>AA \ A \<^bold>\ \<^bold>\A" abbreviation op_und_a::"w \ \ w \" ("\\<^sup>A_" [57]58) (* \ as truth-gap *) where "\\<^sup>AA \ \<^bold>\\\<^sup>AA" abbreviation op_det_b::"w \ \ w \" ("\\<^sup>B_" [57]58) where "\\<^sup>BA \ \\<^sup>d A" abbreviation op_und_b::"w \ \ w \" ("\\<^sup>B_" [57]58) (* \ as border of the complement *) where "\\<^sup>BA \ \\<^sup>d\<^sup>- A" -text\Observe that assumming EXPN \ we are allowed to exchange A and B variants.\ +text\Observe that assumming EXPN for closure we are allowed to exchange A and B variants.\ lemma pundAB: "EXPN \ \ \\<^sup>AA = \\<^sup>BA" using neg_char by (metis BA_deMorgan1 BA_dn L4 L9 dimpl_char impl_char ofp_comm_dc2 op_fixpoint_def sdfun_dcompl_def setequ_ext svfun_compl_def) lemma pdetAB: "EXPN \ \ \\<^sup>AA = \\<^sup>BA" by (metis dual_compl_char1 pundAB sfun_compl_invol svfun_compl_def) text\Variants A and B give us slightly different properties (there are countermodels for those not shown).\ lemma Prop1: "\\<^sup>BA = \\<^sup>f\<^sup>p A" by (simp add: dual_invol setequ_ext) lemma Prop2: "\\<^sup>AA = (\ A \<^bold>\ A)" unfolding conn by (metis compl_def svfun_compl_def) lemma Prop3: "fp \ A \ [\ \\<^sup>B\<^bold>\A]" by (simp add: dual_invol fp_d_rel gtrue_def) lemma Prop4a: "fp \ A \ [\ \\<^sup>BA]" by (simp add: dual_invol fp_rel gtrue_def) lemma Prop4b: "fp \ A \ [\ \\<^sup>AA]" by (simp add: compl_def fixpoints_def join_def setequ_ext svfun_compl_def) text\Recovering TND works for both variants.\ lemma "[\\<^sup>Aa \ a, \<^bold>\a]" by (simp add: subset_def) lemma "[\ \\<^sup>Aa \<^bold>\ a \<^bold>\ \<^bold>\a]" by (metis compl_def join_def) lemma "[\\<^sup>Ba \ a, \<^bold>\a]" by (simp add: compl_def dimpl_def dual_invol join_def op_fixpoint_def subset_def svfun_compl_def) lemma "[\ \\<^sup>Ba \<^bold>\ a \<^bold>\ \<^bold>\a]" by (metis dimpl_def dual_compl_char1 dual_invol join_def ofp_comm_compl op_fixpoint_def) text\We show how (local) contraposition variants (among others) can be recovered using the determinedness operators.\ lemma "[\\<^sup>Aa, a \<^bold>\ b \ \<^bold>\b \<^bold>\ \<^bold>\a]" nitpick oops lemma det_lcop0_A: "EXPN \ \ [\\<^sup>Aa, a \<^bold>\ b \ \<^bold>\b \<^bold>\ \<^bold>\a]" using neg_char impl_char unfolding conn order by fastforce lemma "[\\<^sup>Ba, a \<^bold>\ b \ \<^bold>\b \<^bold>\ \<^bold>\a]" nitpick oops lemma det_lcop0_B: "EXPN \ \ [\\<^sup>Ba, a \<^bold>\ b \ \<^bold>\b \<^bold>\ \<^bold>\a]" by (metis det_lcop0_A pdetAB) lemma "[\\<^sup>Aa, a \<^bold>\ \<^bold>\b \ b \<^bold>\ \<^bold>\a]" nitpick oops lemma det_lcop1_A: "EXPN \ \ [\\<^sup>Aa, a \<^bold>\ \<^bold>\b \ b \<^bold>\ \<^bold>\a]" by (smt (verit, ccfv_SIG) impl_char impl_def join_def meet_def neg_char subset_def) lemma "[\\<^sup>Ba, a \<^bold>\ \<^bold>\b \ b \<^bold>\ \<^bold>\a]" nitpick oops lemma det_lcop1_B: "EXPN \ \ [\\<^sup>Ba, a \<^bold>\ \<^bold>\b \ b \<^bold>\ \<^bold>\a]" by (metis det_lcop1_A pdetAB) lemma "[\\<^sup>Aa, \<^bold>\a \<^bold>\ b \ \<^bold>\b \<^bold>\ a]" nitpick oops lemma det_lcop2_A: "EXPN \ \ [\\<^sup>Aa, \<^bold>\a \<^bold>\ b \ \<^bold>\b \<^bold>\ a]" by (smt (verit, del_insts) neg_char compl_def impl_char join_def meet_def subset_def) lemma "[\\<^sup>Ba, \<^bold>\a \<^bold>\ b \ \<^bold>\b \<^bold>\ a]" nitpick oops lemma det_lcop2_B: "EXPN \ \ [\\<^sup>Ba, \<^bold>\a \<^bold>\ b \ \<^bold>\b \<^bold>\ a]" by (metis det_lcop2_A pdetAB) end diff --git a/thys/Topological_Semantics/logics_negation.thy b/thys/Topological_Semantics/logics_negation.thy --- a/thys/Topological_Semantics/logics_negation.thy +++ b/thys/Topological_Semantics/logics_negation.thy @@ -1,271 +1,271 @@ theory logics_negation imports logics_consequence conditions_relativized begin subsection \Properties of negation(-like) operators\ text\To avoid visual clutter we introduce convenient notation for type for properties of operators.\ type_synonym 'w \ = "('w \ \ 'w \) \ bool" named_theorems neg (*to group together definitions for properties of negations*) subsubsection \Principles of excluded middle, contradiction and explosion\ text\TND: tertium non datur, aka. law of excluded middle (resp. strong, weak, minimal).\ abbreviation pTND ("TND\<^sup>_ _") where "TND\<^sup>a \ \ [\ a \<^bold>\ \ a]" abbreviation pTNDw ("TNDw\<^sup>_ _") where "TNDw\<^sup>a \ \ \b. [\ b \ a, \ a]" abbreviation pTNDm ("TNDm\<^sup>_ _") where "TNDm\<^sup>a \ \ [\ \<^bold>\ \ a, \ a]" definition TND ::"'w \" where "TND \ \ \\. TND\<^sup>\ \" definition TNDw::"'w \" where "TNDw \ \ \\. TNDw\<^sup>\ \" definition TNDm::"'w \" where "TNDm \ \ \\. TNDm\<^sup>\ \" declare TND_def[neg] TNDw_def[neg] TNDm_def[neg] text\Explore some (non)entailment relations.\ lemma "TND \ \ TNDw \" unfolding neg conn order by simp lemma "TNDw \ \ TND \" nitpick oops \\ counterexample \ lemma "TNDw \ \ TNDm \" unfolding neg by simp lemma "TNDm \ \ TNDw \" nitpick oops \\ counterexample \ text\ECQ: ex contradictione (sequitur) quodlibet (variants: strong, weak, minimal).\ abbreviation pECQ ("ECQ\<^sup>_ _") where "ECQ\<^sup>a \ \ [a, \ a \ \<^bold>\]" abbreviation pECQw ("ECQw\<^sup>_ _") where "ECQw\<^sup>a \ \ \b. [a, \ a \ \ b]" abbreviation pECQm ("ECQm\<^sup>_ _") where "ECQm\<^sup>a \ \ [a, \ a \ \ \<^bold>\]" definition ECQ ::"'w \" where "ECQ \ \ \a. ECQ\<^sup>a \" definition ECQw::"'w \" where "ECQw \ \ \a. ECQw\<^sup>a \" definition ECQm::"'w \" where "ECQm \ \ \a. ECQm\<^sup>a \" declare ECQ_def[neg] ECQw_def[neg] ECQm_def[neg] text\Explore some (non)entailment relations.\ lemma "ECQ \ \ ECQw \" unfolding neg conn order by blast lemma "ECQw \ \ ECQ \" nitpick oops \\ counterexample \ lemma "ECQw \ \ ECQm \" unfolding neg by simp lemma "ECQm \ \ ECQw \" nitpick oops \\ counterexample \ text\LNC: law of non-contradiction.\ abbreviation pLNC ("LNC\<^sup>_ _") where "LNC\<^sup>a \ \ [\ \(a \<^bold>\ \ a)]" definition LNC::"'w \" where "LNC \ \ \a. LNC\<^sup>a \" declare LNC_def[neg] text\ECQ and LNC are in general independent.\ lemma "ECQ \ \ LNC \" nitpick oops \\ counterexample \ lemma "LNC \ \ ECQm \" nitpick oops \\ counterexample \ subsubsection \Contraposition rules\ text\CoP: contraposition (weak 'rule-like' variants). Variant 0 is antitonicity (ANTI). Variants 1-3 are stronger.\ abbreviation pCoP1 ("CoP1\<^sup>_\<^sup>_ _") where "CoP1\<^sup>a\<^sup>b \ \ [a \ \ b] \ [b \ \ a]" abbreviation pCoP2 ("CoP2\<^sup>_\<^sup>_ _") where "CoP2\<^sup>a\<^sup>b \ \ [\ a \ b] \ [\ b \ a]" abbreviation pCoP3 ("CoP3\<^sup>_\<^sup>_ _") where "CoP3\<^sup>a\<^sup>b \ \ [\ a \ \ b] \ [b \ a]" abbreviation CoP0 ::"'w \" where "CoP0 \ \ ANTI \" definition CoP1 ::"'w \" where "CoP1 \ \ \a b. CoP1\<^sup>a\<^sup>b \" definition CoP2 ::"'w \" where "CoP2 \ \ \a b. CoP2\<^sup>a\<^sup>b \" definition CoP3 ::"'w \" where "CoP3 \ \ \a b. CoP3\<^sup>a\<^sup>b \" declare CoP1_def[neg] CoP2_def[neg] CoP3_def[neg] text\Explore some (non)entailment relations.\ lemma "CoP1 \ \ CoP0 \" unfolding ANTI_def CoP1_def using subset_char1 by blast lemma "CoP0 \ \ CoP1 \" nitpick oops \\ counterexample \ lemma "CoP2 \ \ CoP0 \" unfolding ANTI_def CoP2_def using subset_char1 by blast lemma "CoP0 \ \ CoP2 \" nitpick oops \\ counterexample \ lemma "CoP3 \ \ CoP0 \" oops (*TODO*) lemma "CoP0 \ \ CoP3 \" nitpick oops \\ counterexample \ text\All three strong variants are pairwise independent. However, CoP3 follows from CoP1 plus CoP2.\ lemma CoP123: "CoP1 \ \ CoP2 \ \ CoP3 \" unfolding neg order by smt text\Taking all CoP together still leaves room for a boldly paraconsistent resp. paracomplete logic.\ lemma "CoP1 \ \ CoP2 \ \ ECQm \" nitpick oops \\ counterexample \ lemma "CoP1 \ \ CoP2 \ \ TNDm \" nitpick oops \\ counterexample \ subsubsection \Modus tollens rules\ text\MT: modus (tollendo) tollens (weak 'rule-like' variants).\ abbreviation pMT0 ("MT0\<^sup>_\<^sup>_ _") where "MT0\<^sup>a\<^sup>b \ \ [a \ b] \ [\ \ b] \ [\ \ a]" abbreviation pMT1 ("MT1\<^sup>_\<^sup>_ _") where "MT1\<^sup>a\<^sup>b \ \ [a \ \ b] \ [\ b] \ [\ \ a]" abbreviation pMT2 ("MT2\<^sup>_\<^sup>_ _") where "MT2\<^sup>a\<^sup>b \ \ [\ a \ b] \ [\ \ b] \ [\ a]" abbreviation pMT3 ("MT3\<^sup>_\<^sup>_ _") where "MT3\<^sup>a\<^sup>b \ \ [\ a \ \ b] \ [\ b] \ [\ a]" definition MT0::"'w \" where "MT0 \ \ \a b. MT0\<^sup>a\<^sup>b \" definition MT1::"'w \" where "MT1 \ \ \a b. MT1\<^sup>a\<^sup>b \" definition MT2::"'w \" where "MT2 \ \ \a b. MT2\<^sup>a\<^sup>b \" definition MT3::"'w \" where "MT3 \ \ \a b. MT3\<^sup>a\<^sup>b \" declare MT0_def[neg] MT1_def[neg] MT2_def[neg] MT3_def[neg] text\Again, all MT variants are pairwise independent. We explore some (non)entailment relations.\ lemma "CoP0 \ \ MT0 \" unfolding neg order cond conn by blast lemma "CoP1 \ \ MT1 \" unfolding neg order conn by blast lemma "CoP2 \ \ MT2 \" unfolding neg order conn by blast lemma "CoP3 \ \ MT3 \" unfolding neg order conn by blast lemma "MT0 \ \ MT1 \ \ MT2 \ \ MT3 \ \ CoP0 \" nitpick oops \\ counterexample \ lemma "MT0 \ \ MT1 \ \ MT2 \ \ MT3 \ \ ECQm \" nitpick oops \\ counterexample \ lemma "MT0 \ \ MT1 \ \ MT2 \ \ MT3 \ \ TNDm \" nitpick oops \\ counterexample \ lemma MT123: "MT1 \ \ MT2 \ \ MT3 \" unfolding neg order conn by metis subsubsection \Double negation introduction and elimination\ text\DNI/DNE: double negation introduction/elimination (strong 'axiom-like' variants).\ abbreviation pDNI ("DNI\<^sup>_ _") where "DNI\<^sup>a \ \ [a \ \(\ a)]" abbreviation pDNE ("DNE\<^sup>_ _") where "DNE\<^sup>a \ \ [\(\ a) \ a]" definition DNI::"'w \" where "DNI \ \ \a. DNI\<^sup>a \" definition DNE::"'w \" where "DNE \ \ \a. DNE\<^sup>a \" declare DNI_def[neg] DNE_def[neg] text\CoP1 (resp. CoP2) can alternatively be defined as CoPw plus DNI (resp. DNE).\ lemma "DNI \ \ CoP1 \" nitpick oops \\ counterexample \ lemma CoP1_def2: "CoP1 \ = (CoP0 \ \ DNI \)" unfolding neg cond using subset_char2 by blast lemma "DNE \ \ CoP2 \" nitpick oops \\ counterexample \ lemma CoP2_def2: "CoP2 \ = (CoP0 \ \ DNE \)" unfolding neg cond using subset_char1 by blast text\Explore some non-entailment relations:\ lemma "DNI \ \ DNE \ \ CoP0 \" nitpick oops \\ counterexample \ lemma "DNI \ \ DNE \ \ TNDm \" nitpick oops \\ counterexample \ lemma "DNI \ \ DNE \ \ ECQm \" nitpick oops \\ counterexample \ lemma "DNI \ \ DNE \ \ MT0 \" nitpick oops \\ counterexample \ lemma "DNI \ \ DNE \ \ MT1 \" nitpick oops \\ counterexample \ lemma "DNI \ \ DNE \ \ MT2 \" nitpick oops \\ counterexample \ lemma "DNI \ \ DNE \ \ MT3 \" nitpick oops \\ counterexample \ text\DNI/DNE: double negation introduction/elimination (weak 'rule-like' variants).\ abbreviation prDNI ("rDNI\<^sup>_ _") where "rDNI\<^sup>a \ \ [\ a] \ [\ \(\ a)]" abbreviation prDNE ("rDNE\<^sup>_ _") where "rDNE\<^sup>a \ \ [\ \(\ a)] \ [\ a]" definition rDNI::"'w \" where "rDNI \ \ \a. rDNI\<^sup>a \" definition rDNE::"'w \" where "rDNE \ \ \a. rDNE\<^sup>a \" declare rDNI_def[neg] rDNE_def[neg] text\The 'rule-like' variants for DNI/DNE are strictly weaker than the 'axiom-like' ones.\ lemma "DNI \ \ rDNI \" unfolding neg order conn by simp lemma "rDNI \ \ DNI \" nitpick oops \\ counterexample \ lemma "DNE \ \ rDNE \" unfolding neg order conn by blast lemma "rDNE \ \ DNE \" nitpick oops \\ counterexample \ text\The 'rule-like' variants for DNI/DNE follow already from modus tollens.\ lemma MT1_rDNI: "MT1 \ \ rDNI \" unfolding neg order by blast lemma MT2_rDNE: "MT2 \ \ rDNE \" unfolding neg order by blast subsubsection \(Anti)Normality and its dual\ text\nNORM (resp. nDNRM) is entailed by CoP1 (resp. CoP2). \ lemma CoP1_NORM: "CoP1 \ \ nNORM \" unfolding neg cond conn order by simp lemma CoP2_DNRM: "CoP2 \ \ nDNRM \" unfolding neg cond conn by (smt (verit) setequ_char subset_def) lemma "DNI \ \ nNORM \" nitpick oops \\ counterexample \ lemma "DNE \ \ nDNRM \" nitpick oops \\ counterexample \ text\nNORM and nDNRM together entail the rule variant of DNI (rDNI).\ lemma nDNRM_rDNI: "nNORM \ \ nDNRM \ \ rDNI \" unfolding neg cond by (simp add: gtrue_def setequ_ext) lemma "nNORM \ \ nDNRM \ \ rDNE \" nitpick oops \\ counterexample \ subsubsection \De Morgan laws\ text\De Morgan laws correspond to anti-additivity and anti-multiplicativity).\ text\DM3 (resp. DM4) are entailed by CoP0/ANTI together with DNE (resp. DNI).\ lemma CoP0_DNE_nMULTb: "CoP0 \ \ DNE \ \ nMULT\<^sup>b \" unfolding neg cond by (metis ANTI_def ANTI_nADDIb L12 nADDI_b_def subset_char1) lemma CoP0_DNI_nADDIa: "CoP0 \ \ DNI \ \ nADDI\<^sup>a \" unfolding neg cond by (metis ANTI_def ANTI_nMULTa L11 nMULT_a_def subset_char2) text\From this follows that DM3 (resp. DM4) is entailed by CoP2 (resp. CoP1).\ lemma CoP2_nMULTb: "CoP2 \ \ nMULT\<^sup>b \" by (simp add: CoP0_DNE_nMULTb CoP2_def2) lemma CoP1_nADDIa: "CoP1 \ \ nADDI\<^sup>a \" by (simp add: CoP0_DNI_nADDIa CoP1_def2) text\Explore some non-entailment relations:\ lemma "CoP0 \ \ nADDI\<^sup>a \ \ nMULT\<^sup>b \ \ nNORM \ \ nDNRM \ \ DNI \" nitpick oops \\ counterexample \ lemma "CoP0 \ \ nADDI\<^sup>a \ \ nMULT\<^sup>b \ \ nNORM \ \ nDNRM \ \ DNE \" nitpick oops \\ counterexample \ lemma "CoP0 \ \ nADDI\<^sup>a \ \ nMULT\<^sup>b \ \ DNI \ \ DNE \ \ ECQm \" nitpick oops \\ counterexample \ lemma "CoP0 \ \ nADDI\<^sup>a \ \ nMULT\<^sup>b \ \ DNI \ \ DNE \ \ TNDm \" nitpick oops \\ counterexample \ subsubsection \Strong contraposition (axiom-like)\ -text\Observe that the definitions below take implication as an additional parameter: \.\ +text\Observe that the definitions below take implication as an additional parameter: @{text "\"}.\ text\lCoP: (local) contraposition (strong 'axiom-like' variants, using local consequence).\ abbreviation plCoP0 ("lCoP0\<^sup>_\<^sup>_ _ _") where "lCoP0\<^sup>a\<^sup>b \ \ \ [\ a b \ \ (\ b) (\ a)]" abbreviation plCoP1 ("lCoP1\<^sup>_\<^sup>_ _ _") where "lCoP1\<^sup>a\<^sup>b \ \ \ [\ a (\ b) \ \ b (\ a)]" abbreviation plCoP2 ("lCoP2\<^sup>_\<^sup>_ _ _") where "lCoP2\<^sup>a\<^sup>b \ \ \ [\ (\ a) b \ \ (\ b) a]" abbreviation plCoP3 ("lCoP3\<^sup>_\<^sup>_ _ _") where "lCoP3\<^sup>a\<^sup>b \ \ \ [\ (\ a) (\ b) \ \ b a]" definition lCoP0::"('w \\'w \\'w \) \ 'w \" where "lCoP0 \ \ \ \a b. lCoP0\<^sup>a\<^sup>b \ \" definition lCoP1::"('w \\'w \\'w \) \ 'w \" where "lCoP1 \ \ \ \a b. lCoP1\<^sup>a\<^sup>b \ \" definition lCoP2::"('w \\'w \\'w \) \ 'w \" where "lCoP2 \ \ \ \a b. lCoP2\<^sup>a\<^sup>b \ \" definition lCoP3::"('w \\'w \\'w \) \ 'w \" where "lCoP3 \ \ \ \a b. lCoP3\<^sup>a\<^sup>b \ \" declare lCoP0_def[neg] lCoP1_def[neg] lCoP2_def[neg] lCoP3_def[neg] text\All these contraposition variants are in general independent from each other. However if we employ classical implication we can verify some relationships.\ lemma lCoP1_def2: "lCoP1(\<^bold>\) \ = (lCoP0(\<^bold>\) \ \ DNI \)" unfolding neg conn order by metis lemma lCoP2_def2: "lCoP2(\<^bold>\) \ = (lCoP0(\<^bold>\) \ \ DNE \)" unfolding neg conn order by blast lemma "lCoP1(\<^bold>\) \ \ lCoP0(\<^bold>\) \" unfolding neg conn order by blast lemma "lCoP0(\<^bold>\) \ \ lCoP1(\<^bold>\) \" nitpick oops \\ counterexample \ lemma "lCoP2(\<^bold>\) \ \ lCoP0(\<^bold>\) \" unfolding neg conn order by blast lemma "lCoP0(\<^bold>\) \ \ lCoP2(\<^bold>\) \" nitpick oops \\ counterexample \ lemma "lCoP3(\<^bold>\) \ \ lCoP0(\<^bold>\) \" unfolding neg conn order by blast lemma "lCoP0(\<^bold>\) \ \ lCoP3(\<^bold>\) \" nitpick oops \\ counterexample \ lemma lCoP123: "lCoP1(\<^bold>\) \ \ lCoP2(\<^bold>\) \ \ lCoP3(\<^bold>\) \" unfolding neg conn order by metis text\Strong/axiom-like variants imply weak/rule-like ones as expected.\ lemma "lCoP0(\<^bold>\) \ \ CoP0 \" unfolding neg cond conn order by blast lemma "lCoP1(\<^bold>\) \ \ CoP1 \" unfolding neg conn order by blast lemma "lCoP2(\<^bold>\) \ \ CoP2 \" unfolding neg conn order by blast lemma "lCoP3(\<^bold>\) \ \ CoP3 \" unfolding neg conn order by blast text\Explore some (non)entailment relations.\ lemma lCoP1_TND: "lCoP1(\<^bold>\) \ \ TND \" unfolding neg conn by (smt (verit, best) setequ_char subset_def) lemma "TND \ \ lCoP1(\<^bold>\) \" nitpick oops \\ counterexample \ lemma lCoP2_ECQ: "lCoP2(\<^bold>\) \ \ ECQ \" unfolding neg conn by (smt (verit) setequ_def subset_def) lemma "ECQ \ \ lCoP2(\<^bold>\) \" nitpick oops \\ counterexample \ subsubsection \Local modus tollens axioms\ text\lMT: (local) Modus tollens (strong, 'axiom-like' variants, using local consequence).\ abbreviation plMT0 ("lMT0\<^sup>_\<^sup>_ _ _") where "lMT0\<^sup>a\<^sup>b \ \ \ [\ a b, \ b \ \ a]" abbreviation plMT1 ("lMT1\<^sup>_\<^sup>_ _ _") where "lMT1\<^sup>a\<^sup>b \ \ \ [\ a (\ b), b \ \ a]" abbreviation plMT2 ("lMT2\<^sup>_\<^sup>_ _ _") where "lMT2\<^sup>a\<^sup>b \ \ \ [\ (\ a) b, \ b \ a]" abbreviation plMT3 ("lMT3\<^sup>_\<^sup>_ _ _") where "lMT3\<^sup>a\<^sup>b \ \ \ [\ (\ a) (\ b), b \ a]" definition lMT0::"('w \\'w \\'w \) \ 'w \" where "lMT0 \ \ \ \a b. lMT0\<^sup>a\<^sup>b \ \" definition lMT1::"('w \\'w \\'w \) \ 'w \" where "lMT1 \ \ \ \a b. lMT1\<^sup>a\<^sup>b \ \" definition lMT2::"('w \\'w \\'w \) \ 'w \" where "lMT2 \ \ \ \a b. lMT2\<^sup>a\<^sup>b \ \" definition lMT3::"('w \\'w \\'w \) \ 'w \" where "lMT3 \ \ \ \a b. lMT3\<^sup>a\<^sup>b \ \" declare lMT0_def[neg] lMT1_def[neg] lMT2_def[neg] lMT3_def[neg] text\All these MT variants are in general independent from each other and also from (strong) CoP instances. However if we take classical implication we can verify that local MT and CoP are indeed equivalent.\ lemma "lMT0(\<^bold>\) \ = lCoP0(\<^bold>\) \" unfolding neg conn order by blast lemma "lMT1(\<^bold>\) \ = lCoP1(\<^bold>\) \" unfolding neg conn order by blast lemma "lMT2(\<^bold>\) \ = lCoP2(\<^bold>\) \" unfolding neg conn order by blast lemma "lMT3(\<^bold>\) \ = lCoP3(\<^bold>\) \" unfolding neg conn order by blast subsubsection \Disjunctive syllogism\ text\DS: disjunctive syllogism.\ abbreviation pDS1 ("DS1\<^sup>_\<^sup>_ _ _") where "DS1\<^sup>a\<^sup>b \ \ \ [a \<^bold>\ b \ \ (\ a) b]" abbreviation pDS2 ("DS2\<^sup>_\<^sup>_ _ _") where "DS2\<^sup>a\<^sup>b \ \ \ [\ (\ a) b \ a \<^bold>\ b]" abbreviation pDS3 ("DS3\<^sup>_\<^sup>_ _ _") where "DS3\<^sup>a\<^sup>b \ \ \ [\ a \<^bold>\ b \ \ a b]" abbreviation pDS4 ("DS4\<^sup>_\<^sup>_ _ _") where "DS4\<^sup>a\<^sup>b \ \ \ [\ a b \ \ a \<^bold>\ b]" definition DS1::"('w \\'w \\'w \) \ 'w \" where "DS1 \ \ \ \a b. DS1\<^sup>a\<^sup>b \ \" definition DS2::"('w \\'w \\'w \) \ 'w \" where "DS2 \ \ \ \a b. DS2\<^sup>a\<^sup>b \ \" definition DS3::"('w \\'w \\'w \) \ 'w \" where "DS3 \ \ \ \a b. DS3\<^sup>a\<^sup>b \ \" definition DS4::"('w \\'w \\'w \) \ 'w \" where "DS4 \ \ \ \a b. DS4\<^sup>a\<^sup>b \ \" declare DS1_def[neg] DS2_def[neg] DS3_def[neg] DS4_def[neg] text\All DS variants are in general independent from each other. However if we take classical implication we can verify that the pairs DS1-DS3 and DS2-DS4 are indeed equivalent. \ lemma "DS1(\<^bold>\) \ = DS3(\<^bold>\) \" unfolding neg conn order by blast lemma "DS2(\<^bold>\) \ = DS4(\<^bold>\) \" unfolding neg conn order by blast text\Explore some (non)entailment relations.\ lemma DS1_nDNor: "DS1(\<^bold>\) \ \ nDNRM \" unfolding neg cond conn order by metis lemma DS2_nNor: "DS2(\<^bold>\) \ \ nNORM \" unfolding neg cond conn order by metis lemma lCoP2_DS1: "lCoP2(\<^bold>\) \ \ DS1(\<^bold>\) \" unfolding neg conn order by blast lemma lCoP1_DS2: "lCoP1(\<^bold>\) \ \ DS2(\<^bold>\) \" unfolding neg conn order by blast lemma "CoP2 \ \ DS1(\<^bold>\) \" nitpick oops \\ counterexample \ lemma "CoP1 \ \ DS2(\<^bold>\) \" nitpick oops \\ counterexample \ end diff --git a/thys/Topological_Semantics/sse_boolean_algebra.thy b/thys/Topological_Semantics/sse_boolean_algebra.thy deleted file mode 100644 --- a/thys/Topological_Semantics/sse_boolean_algebra.thy +++ /dev/null @@ -1,155 +0,0 @@ -theory sse_boolean_algebra - imports Main -begin - -declare[[syntax_ambiguity_warning=false]] -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Shallow embedding of a Boolean algebra of propositions\ - -text\\noindent{In this section we present a shallow semantical embedding (SSE, cf. \<^cite>\J41\ and \<^cite>\J23\) -for a family of logics whose semantics is based upon extensions of (complete and atomic) Boolean algebras. -The range of such logics is indeed very wide, including, as we will see, quantified paraconsistent and -paracomplete (e.g. intuitionistic) logics. Aside from illustrating how the SSE approach works in practice -we show how it allows us to effectively harness theorem provers, model finders and `hammers' for reasoning -with quantified non-classical logics. Proof reconstructions have deliberately been avoided. -Most of the proofs (in fact, all one-liners) have been found using Sledgehammer.}\ - -text\\noindent{Two notions play a fundamental role in this work: propositions and propositional functions. -Propositions, qua sentence denotations, are modeled as objects of type @{text "w\bool"} (shortened as @{text "\"}). -Propositional functions, as the name indicates, are basically anything with a (parametric) type @{text "'t\\"}.}\ - -text\\noindent{We introduce a type @{text "w"} for the domain of points (aka. 'worlds', 'states', etc.). -@{text "\"} is a type alias for sets of points (i.e. propositions) encoded as characteristic functions.}\ -typedecl w -type_synonym \ = "w\bool" - -text\\noindent{In the sequel, we introduce the following naming convention for variables: - -(i) Latin letters (A, b, M, P, q, X, y, etc.) denote in general propositions (type @{text "\"}); -however, we reserve letters D and S to denote sets of propositions (aka. domains/spaces) and -the letters u, v and w to denote worlds/points. - -(ii) Greek letters (in particular @{text "\"}) denote propositional functions (type @{text "'t\\"}); -among the latter we may employ the letters @{text "\"}, @{text "\"} and @{text "\"} to explicitly -name those corresponding to unary connectives/operations (type @{text "\\\"}).}\ - -subsection \Encoding Boolean operations\ - -text\\noindent{We start with an ordered algebra,}\ -abbreviation sequ::"\\\\bool" (infixr "\<^bold>\" 45) where "A \<^bold>\ B \ \w. (A w) \ (B w)" -abbreviation subs::"\\\\bool" (infixr "\<^bold>\" 45) where "A \<^bold>\ B \ \w. (A w) \ (B w)" -abbreviation sups::"\\\\bool" (infixr "\<^bold>\" 45) where "B \<^bold>\ A \ A \<^bold>\ B" - -text\\noindent{define meet and join by reusing HOL metalogical conjunction and disjunction,}\ -definition meet::"\\\\\" (infixr "\<^bold>\" 54) where "A \<^bold>\ B \ \w. (A w) \ (B w)" -definition join::"\\\\\" (infixr "\<^bold>\" 53) where "A \<^bold>\ B \ \w. (A w) \ (B w)" - -text\\noindent{and introduce further operations to obtain a Boolean 'algebra of propositions'.}\ -definition top::"\" ("\<^bold>\") where "\<^bold>\ \ \w. True" -definition bottom::"\" ("\<^bold>\") where "\<^bold>\ \ \w. False" -definition impl::"\\\\\" (infixr "\<^bold>\" 51) where "A \<^bold>\ B \ \w. (A w)\(B w)" -definition dimp::"\\\\\" (infixr "\<^bold>\" 51) where "A \<^bold>\ B \ \w. (A w)\(B w)" -definition diff::"\\\\\" (infixr "\<^bold>\" 51) where "A \<^bold>\ B \ \w. (A w) \ \(B w)" -definition compl::"\\\" ("\<^bold>\_" [57]58) where "\<^bold>\A \ \w. \(A w)" - -named_theorems conn (*algebraic connectives*) -declare meet_def[conn] join_def[conn] top_def[conn] bottom_def[conn] - impl_def[conn] dimp_def[conn] diff_def[conn] compl_def[conn] - -text\\noindent{Quite trivially, we can verify that the algebra satisfies some essential lattice properties.}\ -lemma "a \<^bold>\ a \<^bold>\ a" unfolding conn by simp -lemma "a \<^bold>\ a \<^bold>\ a" unfolding conn by simp -lemma "a \<^bold>\ a \<^bold>\ b" unfolding conn by simp -lemma "a \<^bold>\ b \<^bold>\ a" unfolding conn by simp -lemma "(a \<^bold>\ b) \<^bold>\ b \<^bold>\ b" unfolding conn by auto (*absorption 1*) -lemma "a \<^bold>\ (a \<^bold>\ b) \<^bold>\ a" unfolding conn by auto (*absorption 2*) -lemma "a \<^bold>\ c \ b \<^bold>\ c \ a \<^bold>\ b \<^bold>\ c" unfolding conn by simp -lemma "c \<^bold>\ a \ c \<^bold>\ b \ c \<^bold>\ a \<^bold>\ b" unfolding conn by simp -lemma "a \<^bold>\ b \ (a \<^bold>\ b) \<^bold>\ b" unfolding conn by smt -lemma "b \<^bold>\ a \ (a \<^bold>\ b) \<^bold>\ b" unfolding conn by smt -lemma "a \<^bold>\ c \ b \<^bold>\ d \ (a \<^bold>\ b) \<^bold>\ (c \<^bold>\ d)" unfolding conn by simp -lemma "a \<^bold>\ c \ b \<^bold>\ d \ (a \<^bold>\ b) \<^bold>\ (c \<^bold>\ d)" unfolding conn by simp - - -subsection \Second-order operations and fixed-points\ - -text\\noindent{We define equality for propositional functions as follows.}\ -definition equal_op::"('t\\)\('t\\)\bool" (infix "\<^bold>\" 60) where "\ \<^bold>\ \ \ \X. \ X \<^bold>\ \ X" - -text\\noindent{Moreover, we define some useful Boolean (2nd-order) operations on propositional functions,}\ -abbreviation unionOp::"('t\\)\('t\\)\('t\\)" (infixr "\<^bold>\" 61) where "\ \<^bold>\ \ \ \X. \ X \<^bold>\ \ X" -abbreviation interOp::"('t\\)\('t\\)\('t\\)" (infixr "\<^bold>\" 62) where "\ \<^bold>\ \ \ \X. \ X \<^bold>\ \ X" -abbreviation compOp::"('t\\)\('t\\)" ("(_\<^sup>c)") where "\\<^sup>c \ \X. \<^bold>\\ X" -text\\noindent{some of them explicitly targeting operations,}\ -definition dual::"(\\\)\(\\\)" ("(_\<^sup>d)") where "\\<^sup>d \ \X. \<^bold>\(\(\<^bold>\X))" -text\\noindent{and also define an useful operation (for technical purposes).}\ -definition id::"\\\" ("id") where "id A \ A" - -text\\noindent{We now prove some useful lemmas (some of them may help the provers in their hard work).}\ -lemma comp_symm: "\\<^sup>c = \ \ \ = \\<^sup>c" unfolding conn by blast -lemma comp_invol: "\\<^sup>c\<^sup>c = \" unfolding conn by blast -lemma dual_symm: "(\ \ \\<^sup>d) \ (\ \ \\<^sup>d)" unfolding dual_def conn by simp -lemma dual_comp: "\\<^sup>d\<^sup>c = \\<^sup>c\<^sup>d" unfolding dual_def by simp - -lemma "id\<^sup>d \<^bold>\ id" by (simp add: id_def dual_def equal_op_def conn) -lemma "id\<^sup>c \<^bold>\ compl" by (simp add: id_def dual_def equal_op_def conn) -lemma "(A \<^bold>\ B)\<^sup>d \<^bold>\ (A\<^sup>d) \<^bold>\ (B\<^sup>d)" by (simp add: dual_def equal_op_def conn) -lemma "(A \<^bold>\ B)\<^sup>c \<^bold>\ (A\<^sup>c) \<^bold>\ (B\<^sup>c)" by (simp add: equal_op_def conn) -lemma "(A \<^bold>\ B)\<^sup>d \<^bold>\ (A\<^sup>d) \<^bold>\ (B\<^sup>d)" by (simp add: dual_def equal_op_def conn) -lemma "(A \<^bold>\ B)\<^sup>c \<^bold>\ (A\<^sup>c) \<^bold>\ (B\<^sup>c)" by (simp add: equal_op_def conn) - -text\\noindent{The notion of a fixed point is a fundamental one. We speak of propositions being fixed points of -operations. For a given operation we define in the usual way a fixed-point predicate for propositions.}\ -abbreviation fixedpoint::"(\\\)\(\\bool)" ("fp") where "fp \ \ \X. \ X \<^bold>\ X" - -lemma fp_d: "(fp \\<^sup>d) X = (fp \)(\<^bold>\X)" unfolding dual_def conn by auto -lemma fp_c: "(fp \\<^sup>c) X = (X \<^bold>\ \<^bold>\(\ X))" unfolding conn by auto -lemma fp_dc:"(fp \\<^sup>d\<^sup>c) X = (X \<^bold>\ \(\<^bold>\X))" unfolding dual_def conn by auto - -text\\noindent{Indeed, we can 'operationalize' this predicate by defining a fixed-point operator as follows:}\ -abbreviation fixedpoint_op::"(\\\)\(\\\)" ("(_\<^sup>f\<^sup>p)") where "\\<^sup>f\<^sup>p \ \X. (\ X) \<^bold>\ X" - -lemma ofp_c: "(\\<^sup>c)\<^sup>f\<^sup>p \<^bold>\ (\\<^sup>f\<^sup>p)\<^sup>c" unfolding conn equal_op_def by auto -lemma ofp_d: "(\\<^sup>d)\<^sup>f\<^sup>p \<^bold>\ (\\<^sup>f\<^sup>p)\<^sup>d\<^sup>c" unfolding dual_def equal_op_def conn by auto -lemma ofp_dc:"(\\<^sup>d\<^sup>c)\<^sup>f\<^sup>p \<^bold>\ (\\<^sup>f\<^sup>p)\<^sup>d" unfolding dual_def equal_op_def conn by auto -lemma ofp_decomp: "\\<^sup>f\<^sup>p \<^bold>\ (id \<^bold>\ \) \<^bold>\ ((id \<^bold>\ \)\<^sup>c)" unfolding equal_op_def id_def conn by auto -lemma ofp_invol: "(\\<^sup>f\<^sup>p)\<^sup>f\<^sup>p \<^bold>\ \" unfolding conn equal_op_def by auto - -text\\noindent{Fixed-point predicate and fixed-point operator are closely related.}\ -lemma fp_rel: "((fp \) X) = (\\<^sup>f\<^sup>p X \<^bold>\ \<^bold>\)" unfolding conn by auto -lemma fp_d_rel: "((fp \\<^sup>d) X) = (\\<^sup>f\<^sup>p(\<^bold>\X) \<^bold>\ \<^bold>\)" unfolding dual_def conn by auto -lemma fp_c_rel: "((fp \\<^sup>c) X) = (\\<^sup>f\<^sup>p X \<^bold>\ \<^bold>\)" unfolding conn by auto -lemma fp_dc_rel: "((fp \\<^sup>d\<^sup>c) X) = (\\<^sup>f\<^sup>p(\<^bold>\X) \<^bold>\ \<^bold>\)" unfolding dual_def conn by auto - - -subsection \Equality and atomicity\ - -text\\noindent{We prove some facts about equality (which may help improve prover's performance).}\ -lemma eq_ext: "a \<^bold>\ b \ a = b" using ext by blast -lemma eq_ext': "a \<^bold>\ b \ a = b" using ext unfolding equal_op_def by blast -lemma meet_char: "a \<^bold>\ b \ a \<^bold>\ b \<^bold>\ a" unfolding conn by blast -lemma join_char: "a \<^bold>\ b \ a \<^bold>\ b \<^bold>\ b" unfolding conn by blast - -text\\noindent{We can verify indeed that the algebra is atomic (in three different ways) by relying on the -presence of primitive equality in HOL. A more general class of Boolean algebras could in principle -be obtained in systems without primitive equality or by suitably restricting quantification over -propositions (e.g. defining a topology and restricting quantifiers to open/closed sets).}\ -definition "atom a \ \(a \<^bold>\ \<^bold>\) \ (\p. a \<^bold>\ p \ a \<^bold>\ \<^bold>\p)" -lemma atomic1: "\w. \q. q w \ (\p. p w \ q \<^bold>\ p)" using the_sym_eq_trivial by (metis (full_types)) -lemma atomic2: "\w. \q. q w \ atom(q)" using the_sym_eq_trivial by (metis (full_types) atom_def compl_def bottom_def) -lemma atomic3: "\p. \(p \<^bold>\ \<^bold>\) \ (\q. atom(q) \ q \<^bold>\ p)" proof - (*using atom_def unfolding conn by fastforce*) - { fix p - { assume "\(p \<^bold>\ \<^bold>\)" - hence "\v. p v" unfolding conn by simp - then obtain w where 1:"p w" by (rule exE) - let ?q="\v. v = w" (*using HOL primitive equality*) - have 2: "atom ?q" unfolding atom_def unfolding conn by simp - have "\v. ?q v \ p v" using 1 by simp - hence 3: "?q \<^bold>\ p" by simp - from 2 3 have "\q. atom(q) \ q \<^bold>\ p" by blast - } hence "\(p \<^bold>\ \<^bold>\) \ (\q. atom(q) \ q \<^bold>\ p)" by (rule impI) - } thus ?thesis by (rule allI) -qed - -end diff --git a/thys/Topological_Semantics/sse_boolean_algebra_quantification.thy b/thys/Topological_Semantics/sse_boolean_algebra_quantification.thy deleted file mode 100644 --- a/thys/Topological_Semantics/sse_boolean_algebra_quantification.thy +++ /dev/null @@ -1,170 +0,0 @@ -theory sse_boolean_algebra_quantification - imports sse_boolean_algebra -begin -hide_const(open) List.list.Nil no_notation List.list.Nil ("[]") (*We have no use for lists... *) -hide_const(open) Relation.converse no_notation Relation.converse ("(_\)" [1000] 999) (*..nor for relations in this work*) -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - - -subsection \Obtaining a complete Boolean Algebra\ - -text\\noindent{Our aim is to obtain a complete Boolean algebra which we can use to interpret -quantified formulas (in the spirit of Boolean-valued models for set theory).}\ - -text\\noindent{We start by defining infinite meet (infimum) and infinite join (supremum) operations,}\ -definition infimum:: "(\\bool)\\" ("\<^bold>\_") where "\<^bold>\S \ \w. \X. S X \ X w" -definition supremum::"(\\bool)\\" ("\<^bold>\_") where "\<^bold>\S \ \w. \X. S X \ X w" - -text\\noindent{and show that the corresponding lattice is complete.}\ -abbreviation "upper_bound U S \ \X. (S X) \ X \<^bold>\ U" -abbreviation "lower_bound L S \ \X. (S X) \ L \<^bold>\ X" -abbreviation "is_supremum U S \ upper_bound U S \ (\X. upper_bound X S \ U \<^bold>\ X)" -abbreviation "is_infimum L S \ lower_bound L S \ (\X. lower_bound X S \ X \<^bold>\ L)" - -lemma sup_char: "is_supremum \<^bold>\S S" unfolding supremum_def by auto -lemma sup_ext: "\S. \X. is_supremum X S" by (metis supremum_def) -lemma inf_char: "is_infimum \<^bold>\S S" unfolding infimum_def by auto -lemma inf_ext: "\S. \X. is_infimum X S" by (metis infimum_def) - -text\\noindent{We can check that being closed under supremum/infimum entails being closed under join/meet.}\ -abbreviation "meet_closed S \ \X Y. (S X \ S Y) \ S(X \<^bold>\ Y)" -abbreviation "join_closed S \ \X Y. (S X \ S Y) \ S(X \<^bold>\ Y)" - -abbreviation "nonEmpty S \ \x. S x" -abbreviation "contains S D \ \X. D X \ S X" -abbreviation "infimum_closed S \ \D. nonEmpty D \ contains S D \ S(\<^bold>\D)" -abbreviation "supremum_closed S \ \D. nonEmpty D \ contains S D \ S(\<^bold>\D)" - -lemma inf_meet_closed: "\S. infimum_closed S \ meet_closed S" proof - - { fix S - { assume inf_closed: "infimum_closed S" - hence "meet_closed S" proof - - { fix X::"\" and Y::"\" - let ?D="\Z. Z=X \ Z=Y" - { assume "S X \ S Y" - hence "contains S ?D" by simp - moreover have "nonEmpty ?D" by auto - ultimately have "S(\<^bold>\?D)" using inf_closed by simp - hence "S(\w. \Z. (Z=X \ Z=Y) \ Z w)" unfolding infimum_def by simp - moreover have "(\w. \Z. (Z=X \ Z=Y) \ Z w) = (\w. X w \ Y w)" by auto - ultimately have "S(\w. X w \ Y w)" by simp - } hence "(S X \ S Y) \ S(X \<^bold>\ Y)" unfolding conn by (rule impI) - } thus ?thesis by simp qed - } hence "infimum_closed S \ meet_closed S" by simp - } thus ?thesis by (rule allI) -qed -lemma sup_join_closed: "\P. supremum_closed P \ join_closed P" proof - - { fix S - { assume sup_closed: "supremum_closed S" - hence "join_closed S" proof - - { fix X::"\" and Y::"\" - let ?D="\Z. Z=X \ Z=Y" - { assume "S X \ S Y" - hence "contains S ?D" by simp - moreover have "nonEmpty ?D" by auto - ultimately have "S(\<^bold>\?D)" using sup_closed by simp - hence "S(\w. \Z. (Z=X \ Z=Y) \ Z w)" unfolding supremum_def by simp - moreover have "(\w. \Z. (Z=X \ Z=Y) \ Z w) = (\w. X w \ Y w)" by auto - ultimately have "S(\w. X w \ Y w)" by simp - } hence "(S X \ S Y) \ S(X \<^bold>\ Y)" unfolding conn by (rule impI) - } thus ?thesis by simp qed - } hence "supremum_closed S \ join_closed S" by simp - } thus ?thesis by (rule allI) -qed - - -subsection \Adding quantifiers (restricted and unrestricted)\ - -text\\noindent{We can harness HOL to define quantification over individuals of arbitrary type (using polymorphism). -These (unrestricted) quantifiers take a propositional function and give a proposition.}\ -abbreviation mforall::"('t\\)\\" ("\<^bold>\_" [55]56) where "\<^bold>\\ \ \w. \X. (\ X) w" -abbreviation mexists::"('t\\)\\" ("\<^bold>\_" [55]56) where "\<^bold>\\ \ \w. \X. (\ X) w" -text\\noindent{To improve readability, we introduce for them an useful binder notation.}\ -abbreviation mforallB (binder"\<^bold>\"[55]56) where "\<^bold>\X. \ X \ \<^bold>\\" -abbreviation mexistsB (binder"\<^bold>\"[55]56) where "\<^bold>\X. \ X \ \<^bold>\\" - -(*TODO: is it possible to also add binder notation to the ones below?*) -text\\noindent{Moreover, we define restricted quantifiers which take a 'functional domain' as additional parameter. -The latter is a propositional function that maps each element 'e' to the proposition 'e exists'.}\ -abbreviation mforall_restr::"('t\\)\('t\\)\\" ("\<^bold>\\<^sup>R(_)_") where "\<^bold>\\<^sup>R(\)\ \ \w.\X. (\ X) w \ (\ X) w" -abbreviation mexists_restr::"('t\\)\('t\\)\\" ("\<^bold>\\<^sup>R(_)_") where "\<^bold>\\<^sup>R(\)\ \ \w.\X. (\ X) w \ (\ X) w" - - -subsection \Relating quantifiers with further operators\ - -text\\noindent{The following 'type-lifting' function is useful for converting sets into 'rigid' propositional functions.}\ -abbreviation lift_conv::"('t\bool)\('t\\)" ("\_\") where "\S\ \ \X. \w. S X" - -text\\noindent{We introduce an useful operator: the range of a propositional function (resp. restricted over a domain),}\ -definition pfunRange::"('t\\)\(\\bool)" ("Ra(_)") where "Ra(\) \ \Y. \x. (\ x) = Y" -definition pfunRange_restr::"('t\\)\('t\bool)\(\\bool)" ("Ra[_|_]") where "Ra[\|D] \ \Y. \x. (D x) \ (\ x) = Y" - -text\\noindent{and check that taking infinite joins/meets (suprema/infima) over the range of a propositional function -can be equivalently codified by using quantifiers. This is a quite useful simplifying relationship.}\ -lemma Ra_all: "\<^bold>\Ra(\) = \<^bold>\\" by (metis (full_types) infimum_def pfunRange_def) -lemma Ra_ex: "\<^bold>\Ra(\) = \<^bold>\\" by (metis (full_types) pfunRange_def supremum_def) -lemma Ra_restr_all: "\<^bold>\Ra[\|D] = \<^bold>\\<^sup>R\D\\" by (metis (full_types) pfunRange_restr_def infimum_def) -lemma Ra_restr_ex: "\<^bold>\Ra[\|D] = \<^bold>\\<^sup>R\D\\" by (metis pfunRange_restr_def supremum_def) - -text\\noindent{We further introduce the positive (negative) restriction of a propositional function wrt. a domain,}\ -abbreviation pfunRestr_pos::"('t\\)\('t\\)\('t\\)" ("[_|_]\<^sup>P") where "[\|\]\<^sup>P \ \X. \w. (\ X) w \ (\ X) w" -abbreviation pfunRestr_neg::"('t\\)\('t\\)\('t\\)" ("[_|_]\<^sup>N") where "[\|\]\<^sup>N \ \X. \w. (\ X) w \ (\ X) w" - -text\\noindent{and check that some additional simplifying relationships obtain.}\ -lemma all_restr: "\<^bold>\\<^sup>R(\)\ = \<^bold>\[\|\]\<^sup>P" by simp -lemma ex_restr: "\<^bold>\\<^sup>R(\)\ = \<^bold>\[\|\]\<^sup>N" by simp -lemma Ra_all_restr: "\<^bold>\Ra[\|D] = \<^bold>\[\|\D\]\<^sup>P" using Ra_restr_all by blast -lemma Ra_ex_restr: "\<^bold>\Ra[\|D] = \<^bold>\[\|\D\]\<^sup>N" by (simp add: Ra_restr_ex) - -text\\noindent{Observe that using these operators has the advantage of allowing for binder notation,}\ -lemma "\<^bold>\X. [\|\]\<^sup>P X = \<^bold>\[\|\]\<^sup>P" by simp -lemma "\<^bold>\X. [\|\]\<^sup>N X = \<^bold>\[\|\]\<^sup>N" by simp - -text\\noindent{noting that extra care should be taken when working with complements or negations; -always remember to switch P/N (positive/negative restriction) accordingly.}\ -lemma "\<^bold>\\<^sup>R(\)\ = \<^bold>\X. [\|\]\<^sup>P X" by simp -lemma "\<^bold>\\<^sup>R(\)\\<^sup>c = \<^bold>\X. \<^bold>\[\|\]\<^sup>N X" by (simp add: compl_def) -lemma "\<^bold>\\<^sup>R(\)\ = \<^bold>\X. [\|\]\<^sup>N X" by simp -lemma "\<^bold>\\<^sup>R(\)\\<^sup>c = \<^bold>\X. \<^bold>\[\|\]\<^sup>P X" by (simp add: compl_def) - -text\\noindent{The previous definitions allow us to nicely characterize the interaction -between function composition and (restricted) quantification:}\ -lemma Ra_all_comp1: "\<^bold>\(\\\) = \<^bold>\[\|\Ra \\]\<^sup>P" by (metis comp_apply pfunRange_def) -lemma Ra_all_comp2: "\<^bold>\(\\\) = \<^bold>\\<^sup>R\Ra \\ \" by (metis comp_apply pfunRange_def) -lemma Ra_ex_comp1: "\<^bold>\(\\\) = \<^bold>\[\|\Ra \\]\<^sup>N" by (metis comp_apply pfunRange_def) -lemma Ra_ex_comp2: "\<^bold>\(\\\) = \<^bold>\\<^sup>R\Ra \\ \" by (metis comp_apply pfunRange_def) - -text\\noindent{This useful operator returns for a given domain of propositions the domain of their complements:}\ -definition dom_compl::"(\\bool)\(\\bool)" ("(_\)") where "D\ \ \X. \Y. (D Y) \ (X = \<^bold>\Y)" -lemma dom_compl_def2: "D\ = (\X. D(\<^bold>\X))" unfolding dom_compl_def by (metis comp_symm fun_upd_same) -lemma dom_compl_invol: "D = (D\)\" unfolding dom_compl_def by (metis comp_symm fun_upd_same) - -text\\noindent{We can now check an infinite variant of the De Morgan laws,}\ -lemma iDM_a: "\<^bold>\(\<^bold>\S) = \<^bold>\S\" unfolding dom_compl_def2 infimum_def supremum_def using compl_def by force -lemma iDM_b:" \<^bold>\(\<^bold>\S) = \<^bold>\S\" unfolding dom_compl_def2 infimum_def supremum_def using compl_def by force - -text\\noindent{and some useful dualities regarding the range of propositional functions (restricted wrt. a domain).}\ -lemma Ra_compl: "Ra[\\<^sup>c|D] = Ra[\|D]\" unfolding pfunRange_restr_def dom_compl_def by auto -lemma Ra_dual1: "Ra[\\<^sup>d|D] = Ra[\|D\]\" unfolding pfunRange_restr_def dom_compl_def using dual_def by auto -lemma Ra_dual2: "Ra[\\<^sup>d|D] = Ra[\\<^sup>c|D\]" unfolding pfunRange_restr_def dom_compl_def using dual_def by auto -lemma Ra_dual3: "Ra[\\<^sup>d|D]\ = Ra[\|D\]" unfolding pfunRange_restr_def dom_compl_def using dual_def comp_symm by metis -lemma Ra_dual4: "Ra[\\<^sup>d|D\] = Ra[\|D]\" using Ra_dual3 dual_symm by metis - -text\\noindent{Finally, we check some facts concerning duality for quantifiers.}\ -lemma "\<^bold>\\\<^sup>c = \<^bold>\(\<^bold>\\)" using compl_def by auto -lemma "\<^bold>\\\<^sup>c = \<^bold>\(\<^bold>\\)" using compl_def by auto -lemma "\<^bold>\X. \<^bold>\\ X = \<^bold>\(\<^bold>\X. \ X)" using compl_def by auto -lemma "\<^bold>\X. \<^bold>\\ X = \<^bold>\(\<^bold>\X. \ X)" using compl_def by auto - -lemma "\<^bold>\\<^sup>R(\)\\<^sup>c = \<^bold>\(\<^bold>\\<^sup>R(\)\)" using compl_def by auto -lemma "\<^bold>\\<^sup>R(\)\\<^sup>c = \<^bold>\(\<^bold>\\<^sup>R(\)\)" using compl_def by auto -lemma "\<^bold>\X. \<^bold>\[\|\]\<^sup>P X = \<^bold>\(\<^bold>\X. [\|\]\<^sup>P X)" using compl_def by auto -lemma "\<^bold>\X. \<^bold>\[\|\]\<^sup>P X = \<^bold>\(\<^bold>\X. [\|\]\<^sup>P X)" using compl_def by auto -lemma "\<^bold>\X. \<^bold>\[\|\]\<^sup>N X = \<^bold>\(\<^bold>\X. [\|\]\<^sup>N X)" using compl_def by auto -lemma "\<^bold>\X. \<^bold>\[\|\]\<^sup>N X = \<^bold>\(\<^bold>\X. [\|\]\<^sup>N X)" using compl_def by auto - -text\\noindent{Warning: Do not switch P and N when passing to the dual form.}\ -lemma "\<^bold>\X. [\|\]\<^sup>P X = \<^bold>\(\<^bold>\X. \<^bold>\[\|\]\<^sup>N X)" nitpick oops \\ wrong: counterexample \ -lemma "\<^bold>\X. [\|\]\<^sup>P X = \<^bold>\(\<^bold>\X. \<^bold>\[\|\]\<^sup>P X)" using compl_def by auto \\ correct \ - -end diff --git a/thys/Topological_Semantics/sse_operation_negative.thy b/thys/Topological_Semantics/sse_operation_negative.thy deleted file mode 100644 --- a/thys/Topological_Semantics/sse_operation_negative.thy +++ /dev/null @@ -1,391 +0,0 @@ -theory sse_operation_negative - imports sse_boolean_algebra -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Negative semantic conditions for operations\ - -text\\noindent{We define and interrelate some conditions on operations (i.e. propositional functions of type -@{text "\\\"}), this time involving negative-like properties.}\ - -named_theorems Defs - -subsection \Definitions and interrelations (finitary case)\ - -subsubsection \Principles of excluded middle, contradiction and explosion\ - -text\\noindent{TND: tertium non datur, aka. law of excluded middle (resp. strong, weak, minimal).}\ -abbreviation pTND ("TND\<^sup>_ _") where "TND\<^sup>a \ \ \<^bold>\ \<^bold>\ a \<^bold>\ (\ a)" -abbreviation pTNDw ("TNDw\<^sup>_ _") where "TNDw\<^sup>a \ \ \b. (\ b) \<^bold>\ a \<^bold>\ (\ a)" -abbreviation pTNDm ("TNDm\<^sup>_ _") where "TNDm\<^sup>a \ \ (\ \<^bold>\) \<^bold>\ a \<^bold>\ (\ a)" -definition "TND \ \ \\. TND\<^sup>\ \" -definition "TNDw \ \ \\. TNDw\<^sup>\ \" -definition "TNDm \ \ \\. TNDm\<^sup>\ \" -declare TND_def[Defs] TNDw_def[Defs] TNDm_def[Defs] - -text\\noindent{Explore some (non)entailment relations:}\ -lemma "TND \ \ TNDw \" unfolding Defs conn by simp -lemma "TNDw \ \ TND \" nitpick oops -lemma "TNDw \ \ TNDm \" unfolding Defs by simp -lemma "TNDm \ \ TNDw \" nitpick oops - -text\\noindent{ECQ: ex contradictione (sequitur) quodlibet (resp: strong, weak, minimal).}\ -abbreviation pECQ ("ECQ\<^sup>_ _") where "ECQ\<^sup>a \ \ a \<^bold>\ (\ a) \<^bold>\ \<^bold>\" -abbreviation pECQw ("ECQw\<^sup>_ _") where "ECQw\<^sup>a \ \ \b. a \<^bold>\ (\ a) \<^bold>\ (\ b)" -abbreviation pECQm ("ECQm\<^sup>_ _") where "ECQm\<^sup>a \ \ a \<^bold>\ (\ a) \<^bold>\ (\ \<^bold>\)" -definition "ECQ \ \ \a. ECQ\<^sup>a \" -definition "ECQw \ \ \a. ECQw\<^sup>a \" -definition "ECQm \ \ \a. ECQm\<^sup>a \" -declare ECQ_def[Defs] ECQw_def[Defs] ECQm_def[Defs] - -text\\noindent{Explore some (non)entailment relations:}\ -lemma "ECQ \ \ ECQw \" unfolding Defs conn by blast -lemma "ECQw \ \ ECQ \" nitpick oops -lemma "ECQw \ \ ECQm \" unfolding Defs conn by simp -lemma "ECQm \ \ ECQw \" nitpick oops - -text\\noindent{LNC: law of non-contradiction.}\ -abbreviation pLNC ("LNC\<^sup>_ _") where "LNC\<^sup>a \ \ \(a \<^bold>\ \ a) \<^bold>\ \<^bold>\" -definition "LNC \ \ \a. LNC\<^sup>a \" -declare LNC_def[Defs] - -text\\noindent{ECQ and LNC are in general independent.}\ -lemma "ECQ \ \ LNC \" nitpick oops -lemma "LNC \ \ ECQm \" nitpick oops - - -subsubsection \Contraposition rules\ - -text\\noindent{CoP: contraposition (global/rule variants, resp. weak, strong var. 1, strong var. 2, strong var. 3).}\ -abbreviation pCoPw ("CoPw\<^sup>_\<^sup>_ _") where "CoPw\<^sup>a\<^sup>b \ \ a \<^bold>\ b \ (\ b) \<^bold>\ (\ a)" -abbreviation pCoP1 ("CoP1\<^sup>_\<^sup>_ _") where "CoP1\<^sup>a\<^sup>b \ \ a \<^bold>\ (\ b) \ b \<^bold>\ (\ a)" -abbreviation pCoP2 ("CoP2\<^sup>_\<^sup>_ _") where "CoP2\<^sup>a\<^sup>b \ \ (\ a) \<^bold>\ b \ (\ b) \<^bold>\ a" -abbreviation pCoP3 ("CoP3\<^sup>_\<^sup>_ _") where "CoP3\<^sup>a\<^sup>b \ \ (\ a) \<^bold>\ (\ b) \ b \<^bold>\ a" -definition "CoPw \ \ \a b. CoPw\<^sup>a\<^sup>b \" -definition "CoP1 \ \ \a b. CoP1\<^sup>a\<^sup>b \" -definition "CoP1' \ \ \a b. a \<^bold>\ (\ b) \ b \<^bold>\ (\ a)" -definition "CoP2 \ \ \a b. CoP2\<^sup>a\<^sup>b \" -definition "CoP2' \ \ \a b. (\ a) \<^bold>\ b \ (\ b) \<^bold>\ a" -definition "CoP3 \ \ \a b. CoP3\<^sup>a\<^sup>b \" -declare CoPw_def[Defs] CoP1_def[Defs] CoP1'_def[Defs] - CoP2_def[Defs] CoP2'_def[Defs] CoP3_def[Defs] - -lemma CoP1_defs_rel: "CoP1 \ = CoP1' \" unfolding Defs by metis -lemma CoP2_defs_rel: "CoP2 \ = CoP2' \" unfolding Defs by metis - -text\\noindent{Explore some (non)entailment relations:}\ -lemma "CoP1 \ \ CoPw \" unfolding Defs by metis -lemma "CoPw \ \ CoP1 \" nitpick oops -lemma "CoP2 \ \ CoPw \" unfolding Defs by metis -lemma "CoPw \ \ CoP2 \" nitpick oops -lemma "CoP3 \ \ CoPw \" (*nitpick*) oops \\ no countermodel found so far \ -lemma "CoPw \ \ CoP3 \" nitpick oops - -text\\noindent{All three strong variants are pairwise independent. However, CoP3 follows from CoP1 plus CoP2.}\ -lemma CoP123: "CoP1 \ \ CoP2 \ \ CoP3 \" unfolding Defs by smt -text\\noindent{Taking all CoP together still leaves room for a boldly paraconsistent resp. paracomplete logic.}\ -lemma "CoP1 \ \ CoP2 \ \ ECQm \" nitpick oops -lemma "CoP1 \ \ CoP2 \ \ TNDm \" nitpick oops - - -subsubsection \Modus tollens rules\ - -text\\noindent{MT: modus (tollendo) tollens (global/rule variants).}\ -abbreviation pMT0 ("MT0\<^sup>_\<^sup>_ _") where "MT0\<^sup>a\<^sup>b \ \ a \<^bold>\ b \ (\ b) \<^bold>\ \<^bold>\ \ (\ a) \<^bold>\ \<^bold>\" -abbreviation pMT1 ("MT1\<^sup>_\<^sup>_ _") where "MT1\<^sup>a\<^sup>b \ \ a \<^bold>\ (\ b) \ b \<^bold>\ \<^bold>\ \ (\ a) \<^bold>\ \<^bold>\" -abbreviation pMT2 ("MT2\<^sup>_\<^sup>_ _") where "MT2\<^sup>a\<^sup>b \ \ (\ a) \<^bold>\ b \ (\ b) \<^bold>\ \<^bold>\ \ a \<^bold>\ \<^bold>\" -abbreviation pMT3 ("MT3\<^sup>_\<^sup>_ _") where "MT3\<^sup>a\<^sup>b \ \ (\ a) \<^bold>\ (\ b) \ b \<^bold>\ \<^bold>\ \ a \<^bold>\ \<^bold>\" -definition "MT0 \ \ \a b. MT0\<^sup>a\<^sup>b \" -definition "MT1 \ \ \a b. MT1\<^sup>a\<^sup>b \" -definition "MT2 \ \ \a b. MT2\<^sup>a\<^sup>b \" -definition "MT3 \ \ \a b. MT3\<^sup>a\<^sup>b \" -declare MT0_def[Defs] MT1_def[Defs] MT2_def[Defs] MT3_def[Defs] - -text\\noindent{Again, all MT variants are pairwise independent. We explore some (non)entailment relations:}\ -lemma "CoPw \ \ MT0 \" unfolding Defs by (metis top_def) -lemma "CoP1 \ \ MT1 \" unfolding Defs by (metis top_def) -lemma "CoP2 \ \ MT2 \" unfolding Defs by (metis top_def) -lemma "CoP3 \ \ MT3 \" unfolding Defs by (metis top_def) -lemma "MT0 \ \ MT1 \ \ MT2 \ \ MT3 \ \ CoPw \" nitpick oops -lemma "MT0 \ \ MT1 \ \ MT2 \ \ MT3 \ \ ECQm \" nitpick oops -lemma "MT0 \ \ MT1 \ \ MT2 \ \ MT3 \ \ TNDm \" nitpick oops -lemma MT123: "MT1 \ \ MT2 \ \ MT3 \" unfolding Defs by smt - - -subsubsection \Double negation introduction and elimination\ - -text\\noindent{DNI/DNE: double negation introduction/elimination (as axioms).}\ -abbreviation pDNI ("DNI\<^sup>_ _") where "DNI\<^sup>a \ \ a \<^bold>\ \ (\ a)" -abbreviation pDNE ("DNE\<^sup>_ _") where "DNE\<^sup>a \ \ \ (\ a) \<^bold>\ a" -definition "DNI \ \ \a. DNI\<^sup>a \" -definition "DNE \ \ \a. DNE\<^sup>a \" -declare DNI_def[Defs] DNE_def[Defs] - -text\\noindent{CoP1 (resp. CoP2) can alternatively be defined as CoPw plus DNI (resp. DNE).}\ -lemma "DNI \ \ CoP1 \" nitpick oops -lemma CoP1_def2: "CoP1 \ = (CoPw \ \ DNI \)" unfolding Defs by smt -lemma "DNE \ \ CoP2 \" nitpick oops -lemma CoP2_def2: "CoP2 \ = (CoPw \ \ DNE \)" unfolding Defs by smt - -text\\noindent{Explore some non-entailment relations:}\ -lemma "DNI \ \ DNE \ \ CoPw \" nitpick oops -lemma "DNI \ \ DNE \ \ TNDm \" nitpick oops -lemma "DNI \ \ DNE \ \ ECQm \" nitpick oops -lemma "DNI \ \ DNE \ \ MT0 \" nitpick oops -lemma "DNI \ \ DNE \ \ MT1 \" nitpick oops -lemma "DNI \ \ DNE \ \ MT2 \" nitpick oops -lemma "DNI \ \ DNE \ \ MT3 \" nitpick oops - -text\\noindent{DNI/DNE: double negation introduction/elimination (as rules).}\ -abbreviation prDNI ("rDNI\<^sup>_ _") where "rDNI\<^sup>a \ \ a \<^bold>\ \<^bold>\ \ \ (\ a) \<^bold>\ \<^bold>\" -abbreviation prDNE ("rDNE\<^sup>_ _") where "rDNE\<^sup>a \ \ \ (\ a) \<^bold>\ \<^bold>\ \ a \<^bold>\ \<^bold>\" -definition "rDNI \ \ \a. rDNI\<^sup>a \" -definition "rDNE \ \ \a. rDNE\<^sup>a \" -declare rDNI_def[Defs] rDNE_def[Defs] - -text\\noindent{The rule variants are strictly weaker than the axiom variants,}\ -lemma "DNI \ \ rDNI \" by (simp add: DNI_def rDNI_def top_def) -lemma "rDNI \ \ DNI \" nitpick oops -lemma "DNE \ \ rDNE \" by (metis DNE_def rDNE_def top_def) -lemma "rDNE \ \ DNE \" nitpick oops -text\\noindent{and follow already from modus tollens.}\ -lemma MT1_rDNI: "MT1 \ \ rDNI \" unfolding Defs by blast -lemma MT2_rDNE: "MT2 \ \ rDNE \" unfolding Defs by blast - - -subsubsection \Normality and its dual\ - -text\\noindent{n(D)Nor: negative (dual) 'normality'.}\ -definition "nNor \ \ (\ \<^bold>\) \<^bold>\ \<^bold>\" -definition "nDNor \ \ (\ \<^bold>\) \<^bold>\ \<^bold>\" -declare nNor_def[Defs] nDNor_def[Defs] - -text\\noindent{nNor (resp. nDNor) is entailed by CoP1 (resp. CoP2). }\ -lemma CoP1_Nor: "CoP1 \ \ nNor \" unfolding Defs conn by simp -lemma CoP2_DNor: "CoP2 \ \ nDNor \" unfolding Defs conn by fastforce -lemma "DNI \ \ nNor \" nitpick oops -lemma "DNE \ \ nDNor \" nitpick oops -text\\noindent{nNor and nDNor together entail the rule variant of DNI (rDNI).}\ -lemma nDNor_rDNI: "nNor \ \ nDNor \ \ rDNI \" unfolding Defs using nDNor_def nNor_def eq_ext by metis -lemma "nNor \ \ nDNor \ \ rDNE \" nitpick oops - - -subsubsection \De Morgan laws\ - -text\\noindent{DM: De Morgan laws.}\ -abbreviation pDM1 ("DM1\<^sup>_\<^sup>_ _") where "DM1\<^sup>a\<^sup>b \ \ \(a \<^bold>\ b) \<^bold>\ (\ a) \<^bold>\ (\ b)" -abbreviation pDM2 ("DM2\<^sup>_\<^sup>_ _") where "DM2\<^sup>a\<^sup>b \ \ (\ a) \<^bold>\ (\ b) \<^bold>\ \(a \<^bold>\ b)" -abbreviation pDM3 ("DM3\<^sup>_\<^sup>_ _") where "DM3\<^sup>a\<^sup>b \ \ \(a \<^bold>\ b) \<^bold>\ (\ a) \<^bold>\ (\ b)" -abbreviation pDM4 ("DM4\<^sup>_\<^sup>_ _") where "DM4\<^sup>a\<^sup>b \ \ (\ a) \<^bold>\ (\ b) \<^bold>\ \(a \<^bold>\ b)" -definition "DM1 \ \ \a b. DM1\<^sup>a\<^sup>b \" -definition "DM2 \ \ \a b. DM2\<^sup>a\<^sup>b \" -definition "DM3 \ \ \a b. DM3\<^sup>a\<^sup>b \" -definition "DM4 \ \ \a b. DM4\<^sup>a\<^sup>b \" -declare DM1_def[Defs] DM2_def[Defs] DM3_def[Defs] DM4_def[Defs] - -text\\noindent{CoPw, DM1 and DM2 are indeed equivalent.}\ -lemma DM1_CoPw: "DM1 \ = CoPw \" proof - - have LtoR: "DM1 \ \ CoPw \" proof - - assume dm1: "DM1 \" - { fix a b - { assume "a \<^bold>\ b" - hence 1: "a \<^bold>\ b \<^bold>\ b" unfolding conn by simp - have 2: "b \<^bold>\ a \<^bold>\ b" unfolding conn by simp - from 1 2 have "b = a \<^bold>\ b" using eq_ext by blast - hence 3: "\ b \<^bold>\ \ (a \<^bold>\ b)" by auto - from dm1 have "\(a \<^bold>\ b) \<^bold>\ (\ a) \<^bold>\ (\ b)" unfolding Defs by blast - hence 4: "(\ b) \<^bold>\ (\ a) \<^bold>\ (\ b)" using 3 by simp - have 5: "(\ a) \<^bold>\ (\ b) \<^bold>\ (\ a)" unfolding conn by simp - from 4 5 have "(\ b) \<^bold>\ (\ a)" by simp - } hence "a \<^bold>\ b \ (\ b) \<^bold>\ (\ a)" by (rule impI) - } thus ?thesis unfolding Defs by simp - qed - have RtoL: "CoPw \ \ DM1 \" unfolding Defs conn by (metis (no_types, lifting)) - thus ?thesis using LtoR RtoL by blast -qed -lemma DM2_CoPw: "DM2 \ = CoPw \" proof - - have LtoR: "DM2 \ \ CoPw \" proof - - assume dm2: "DM2 \" - { fix a b - { assume "a \<^bold>\ b" - hence 1: "a \<^bold>\ a \<^bold>\ b" unfolding conn by simp - have 2: "a \<^bold>\ b \<^bold>\ a" unfolding conn by simp - from 1 2 have "a = a \<^bold>\ b" using eq_ext by blast - hence 3: "\ a \<^bold>\ \ (a \<^bold>\ b)" by auto - from dm2 have "(\ a) \<^bold>\ (\ b) \<^bold>\ \(a \<^bold>\ b)" unfolding Defs by blast - hence 4: "(\ a) \<^bold>\ (\ b) \<^bold>\ (\ a) " using 3 by simp - have 5: "(\ b) \<^bold>\ (\ a) \<^bold>\ (\ b)" unfolding conn by simp - from 4 5 have "(\ b) \<^bold>\ (\ a)" by simp - } hence "a \<^bold>\ b \ (\ b) \<^bold>\ (\ a)" by (rule impI) - } thus ?thesis unfolding Defs by simp - qed - have RtoL: "CoPw \ \ DM2 \" unfolding Defs conn by (metis (no_types, lifting)) - thus ?thesis using LtoR RtoL by blast -qed -lemma DM12: "DM1 \ = DM2 \" by (simp add: DM1_CoPw DM2_CoPw) - -text\\noindent{DM3 (resp. DM4) are entailed by CoPw together with DNE (resp. DNI).}\ -lemma CoPw_DNE_DM3: "CoPw \ \ DNE \ \ DM3 \" proof - - assume copw: "CoPw \" and dne: "DNE \" - { fix a b - have "\(a) \<^bold>\ (\ a) \<^bold>\ (\ b)" unfolding conn by simp - hence "\(\(a) \<^bold>\ \(b)) \<^bold>\ \((\ a))" using CoPw_def copw by (metis (no_types, lifting)) - hence 1: "\(\(a) \<^bold>\ \(b)) \<^bold>\ a" using DNE_def dne by metis - have "\(b) \<^bold>\ (\ a) \<^bold>\ (\ b)" unfolding conn by simp - hence "\(\(a) \<^bold>\ \(b)) \<^bold>\ \((\ b))" using CoPw_def copw by (metis (no_types, lifting)) - hence 2: "\(\(a) \<^bold>\ \(b)) \<^bold>\ b" using DNE_def dne by metis - from 1 2 have "\(\(a) \<^bold>\ \(b)) \<^bold>\ a \<^bold>\ b" unfolding conn by simp - hence "\(a \<^bold>\ b) \<^bold>\ \(\(\(a) \<^bold>\ \(b)))" using CoPw_def copw by smt - hence "\(a \<^bold>\ b) \<^bold>\ (\ a) \<^bold>\ (\ b)" using DNE_def dne by metis - } thus ?thesis by (simp add: DM3_def) -qed -lemma CoPw_DNI_DM4: "CoPw \ \ DNI \ \ DM4 \" proof - - assume copw: "CoPw \" and dni: "DNI \" - { fix a b - have "(\ a) \<^bold>\ (\ b) \<^bold>\ \(a)" unfolding conn by simp - hence "\((\ a)) \<^bold>\ \(\(a) \<^bold>\ \(b))" using CoPw_def copw by (metis (no_types, lifting)) - hence 1: "a \<^bold>\ \(\(a) \<^bold>\ \(b))" using DNI_def dni by metis - have "(\ a) \<^bold>\ (\ b) \<^bold>\ \(b)" unfolding conn by simp - hence "\((\ b)) \<^bold>\ \(\(a) \<^bold>\ \(b))" using CoPw_def copw by (metis (no_types, lifting)) - hence 2: "b \<^bold>\ \(\(a) \<^bold>\ \(b))" using DNI_def dni by metis - from 1 2 have "a \<^bold>\ b \<^bold>\ \(\(a) \<^bold>\ \(b))" unfolding conn by simp - hence "\(\(\(a) \<^bold>\ \(b))) \<^bold>\ \(a \<^bold>\ b)" using CoPw_def copw by auto - hence "\(a) \<^bold>\ \(b) \<^bold>\ \(a \<^bold>\ b)" using DNI_def dni by simp - } thus ?thesis by (simp add: DM4_def) -qed -text\\noindent{From this follows that DM3 (resp. DM4) is entailed by CoP2 (resp. CoP1).}\ -lemma CoP2_DM3: "CoP2 \ \ DM3 \" by (simp add: CoP2_def2 CoPw_DNE_DM3) -lemma CoP1_DM4: "CoP1 \ \ DM4 \" by (simp add: CoP1_def2 CoPw_DNI_DM4) -text\\noindent{Explore some non-entailment relations:}\ -lemma "CoPw \ \ DM3 \ \ DM4 \ \ nNor \ \ nDNor \ \ DNI \" nitpick oops -lemma "CoPw \ \ DM3 \ \ DM4 \ \ nNor \ \ nDNor \ \ DNE \" nitpick oops -lemma "CoPw \ \ DM3 \ \ DM4 \ \ DNI \ \ DNE \ \ ECQm \" nitpick oops -lemma "CoPw \ \ DM3 \ \ DM4 \ \ DNI \ \ DNE \ \ TNDm \" nitpick oops - - -subsubsection \Contextual (strong) contraposition rule\ - -text\\noindent{XCoP: contextual contraposition (global/rule variant).}\ -abbreviation pXCoP ("XCoP\<^sup>_\<^sup>_ _") where "XCoP\<^sup>a\<^sup>b \ \ \c. c \<^bold>\ a \<^bold>\ b \ c \<^bold>\ (\ b) \<^bold>\ (\ a)" -definition "XCoP \ \ \a b. XCoP\<^sup>a\<^sup>b \" -declare XCoP_def[Defs] - -text\\noindent{XCoP can alternatively be defined as ECQw plus TNDw.}\ -lemma XCoP_def2: "XCoP \ = (ECQw \ \ TNDw \)" proof - - have LtoR1: "XCoP \ \ ECQw \" unfolding XCoP_def ECQw_def conn by auto - have LtoR2: "XCoP \ \ TNDw \" unfolding XCoP_def TNDw_def conn by (smt atom_def atomic2 conn) - have RtoL: "ECQw \ \ TNDw \ \ XCoP \" using XCoP_def ECQw_def TNDw_def unfolding conn by metis - from LtoR1 LtoR2 RtoL show ?thesis by blast -qed -text\\noindent{Explore some (non)entailment relations:}\ -lemma "XCoP \ \ ECQ \" nitpick oops -lemma "XCoP \ \ TND \" nitpick oops -lemma XCoP_CoPw: "XCoP \ \ CoPw \" unfolding Defs conn by metis -lemma "XCoP \ \ CoP1 \" nitpick oops -lemma "XCoP \ \ CoP2 \" nitpick oops -lemma "XCoP \ \ CoP3 \" nitpick oops -lemma "CoP1 \ \ CoP2 \ \ XCoP \" nitpick oops -lemma "XCoP \ \ nNor \" nitpick oops -lemma "XCoP \ \ nDNor \" nitpick oops -lemma "XCoP \ \ rDNI \" nitpick oops -lemma "XCoP \ \ rDNE \" nitpick oops -lemma XCoP_DM3: "XCoP \ \ DM3 \" unfolding DM3_def XCoP_def conn using ECQw_def TNDw_def atom_def atomic2 conn by smt -lemma XCoP_DM4: "XCoP \ \ DM4 \" unfolding DM4_def XCoP_def conn using ECQw_def TNDw_def atom_def atomic2 conn by smt - - -subsubsection \Local contraposition axioms\ -text\\noindent{Observe that the definitions below take implication as an additional parameter: @{text "\"}.}\ - -text\\noindent{lCoP: contraposition (local/axiom variants).}\ -abbreviation plCoPw ("lCoPw\<^sup>_\<^sup>_ _ _") where "lCoPw\<^sup>a\<^sup>b \ \ \ (\ a b::\) \<^bold>\ (\ (\ b) (\ a))" -abbreviation plCoP1 ("lCoP1\<^sup>_\<^sup>_ _ _") where "lCoP1\<^sup>a\<^sup>b \ \ \ (\ a (\ b::\)) \<^bold>\ (\ b (\ a))" -abbreviation plCoP2 ("lCoP2\<^sup>_\<^sup>_ _ _") where "lCoP2\<^sup>a\<^sup>b \ \ \ (\ (\ a) b::\) \<^bold>\ (\ (\ b) a)" -abbreviation plCoP3 ("lCoP3\<^sup>_\<^sup>_ _ _") where "lCoP3\<^sup>a\<^sup>b \ \ \ (\ (\ a) (\ b::\)) \<^bold>\ (\ b a)" -definition "lCoPw \ \ \ \a b. lCoPw\<^sup>a\<^sup>b \ \" -definition "lCoP1 \ \ \ \a b. lCoP1\<^sup>a\<^sup>b \ \" -definition "lCoP1' \ \ \ \a b. (\ a (\ b)) \<^bold>\ (\ b (\ a))" -definition "lCoP2 \ \ \ \a b. lCoP2\<^sup>a\<^sup>b \ \" -definition "lCoP2' \ \ \ \a b. (\ (\ a) b) \<^bold>\ (\ (\ b) a)" -definition "lCoP3 \ \ \ \a b. lCoP3\<^sup>a\<^sup>b \ \" -declare lCoPw_def[Defs] lCoP1_def[Defs] lCoP1'_def[Defs] - lCoP2_def[Defs] lCoP2'_def[Defs] lCoP3_def[Defs] - -lemma lCoP1_defs_rel: "lCoP1 \ \ = lCoP1' \ \" by (metis (full_types) lCoP1'_def lCoP1_def) -lemma lCoP2_defs_rel: "lCoP2 \ \ = lCoP2' \ \" by (metis (full_types) lCoP2'_def lCoP2_def) - -text\\noindent{All local contraposition variants are in general independent from each other. However if we take -classical implication we can verify some relationships.}\ - -lemma lCoP1_def2: "lCoP1(\<^bold>\) \ = (lCoPw(\<^bold>\) \ \ DNI \)" unfolding Defs conn by smt -lemma lCoP2_def2: "lCoP2(\<^bold>\) \ = (lCoPw(\<^bold>\) \ \ DNE \)" unfolding Defs conn by smt - -lemma "lCoP1(\<^bold>\) \ \ lCoPw(\<^bold>\) \" unfolding Defs conn by metis -lemma "lCoPw(\<^bold>\) \ \ lCoP1(\<^bold>\) \" nitpick oops -lemma "lCoP2(\<^bold>\) \ \ lCoPw(\<^bold>\) \" unfolding Defs conn by metis -lemma "lCoPw(\<^bold>\) \ \ lCoP2(\<^bold>\) \" nitpick oops -lemma "lCoP3(\<^bold>\) \ \ lCoPw(\<^bold>\) \" unfolding Defs conn by blast -lemma "lCoPw(\<^bold>\) \ \ lCoP3(\<^bold>\) \" nitpick oops -lemma lCoP123: "lCoP1(\<^bold>\) \ \ lCoP2(\<^bold>\) \ \ lCoP3(\<^bold>\) \" unfolding Defs conn by metis - -text\\noindent{Local variants imply global ones as expected.}\ -lemma "lCoPw(\<^bold>\) \ \ CoPw \" unfolding Defs conn by metis -lemma "lCoP1(\<^bold>\) \ \ CoP1 \" unfolding Defs conn by metis -lemma "lCoP2(\<^bold>\) \ \ CoP2 \" unfolding Defs conn by metis -lemma "lCoP3(\<^bold>\) \ \ CoP3 \" unfolding Defs conn by metis - -text\\noindent{Explore some (non)entailment relations.}\ -lemma lCoPw_XCoP: "lCoPw(\<^bold>\) \ = XCoP \" unfolding Defs conn by (smt XCoP_def XCoP_def2 TNDw_def conn) -lemma lCoP1_TND: "lCoP1(\<^bold>\) \ \ TND \" by (smt XCoP_CoPw XCoP_def2 CoP1_Nor CoP1_def2 nNor_def TND_def TNDw_def lCoP1_def2 lCoPw_XCoP conn) -lemma "TND \ \ lCoP1(\<^bold>\) \" nitpick oops -lemma lCoP2_ECQ: "lCoP2(\<^bold>\) \ \ ECQ \" by (smt XCoP_CoPw XCoP_def2 CoP2_DNor CoP2_def2 nDNor_def ECQ_def ECQw_def lCoP2_def2 lCoPw_XCoP conn) -lemma "ECQ \ \ lCoP2(\<^bold>\) \" nitpick oops - - -subsubsection \Local modus tollens axioms\ - -text\\noindent{lMT: Modus tollens (local/axiom variants).}\ -abbreviation plMT0 ("lMT0\<^sup>_\<^sup>_ _ _") where "lMT0\<^sup>a\<^sup>b \ \ \ (\ a b::\) \<^bold>\ (\ b) \<^bold>\ (\ a)" -abbreviation plMT1 ("lMT1\<^sup>_\<^sup>_ _ _") where "lMT1\<^sup>a\<^sup>b \ \ \ (\ a (\ b::\)) \<^bold>\ b \<^bold>\ (\ a)" -abbreviation plMT2 ("lMT2\<^sup>_\<^sup>_ _ _") where "lMT2\<^sup>a\<^sup>b \ \ \ (\ (\ a) b::\) \<^bold>\ (\ b) \<^bold>\ a" -abbreviation plMT3 ("lMT3\<^sup>_\<^sup>_ _ _") where "lMT3\<^sup>a\<^sup>b \ \ \ (\ (\ a) (\ b::\)) \<^bold>\ b \<^bold>\ a" -definition "lMT0 \ \ \ \a b. lMT0\<^sup>a\<^sup>b \ \" -definition "lMT1 \ \ \ \a b. lMT1\<^sup>a\<^sup>b \ \" -definition "lMT2 \ \ \ \a b. lMT2\<^sup>a\<^sup>b \ \" -definition "lMT3 \ \ \ \a b. lMT3\<^sup>a\<^sup>b \ \" -declare lMT0_def[Defs] lMT1_def[Defs] lMT2_def[Defs] lMT3_def[Defs] - -text\\noindent{All local MT variants are in general independent from each other and also from local CoP instances. -However if we take classical implication we can verify that local MT and CoP are indeed equivalent.}\ -lemma "lMT0(\<^bold>\) \ = lCoPw(\<^bold>\) \" unfolding Defs conn by metis -lemma "lMT1(\<^bold>\) \ = lCoP1(\<^bold>\) \" unfolding Defs conn by metis -lemma "lMT2(\<^bold>\) \ = lCoP2(\<^bold>\) \" unfolding Defs conn by metis -lemma "lMT3(\<^bold>\) \ = lCoP3(\<^bold>\) \" unfolding Defs conn by metis - - -subsubsection \Disjunctive syllogism\ - -text\\noindent{DS: disjunctive syllogism.}\ -abbreviation pDS1 ("DS1\<^sup>_\<^sup>_ _ _") where "DS1\<^sup>a\<^sup>b \ \ \ (a \<^bold>\ b::\) \<^bold>\ (\ (\ a) b)" -abbreviation pDS2 ("DS2\<^sup>_\<^sup>_ _ _") where "DS2\<^sup>a\<^sup>b \ \ \ (\ (\ a) b::\) \<^bold>\ (a \<^bold>\ b)" -abbreviation pDS3 ("DS3\<^sup>_\<^sup>_ _ _") where "DS3\<^sup>a\<^sup>b \ \ \ ((\ a) \<^bold>\ b::\) \<^bold>\ (\ a b)" -abbreviation pDS4 ("DS4\<^sup>_\<^sup>_ _ _") where "DS4\<^sup>a\<^sup>b \ \ \ (\ a b::\) \<^bold>\ ((\ a) \<^bold>\ b)" -definition "DS1 \ \ \ \a b. DS1\<^sup>a\<^sup>b \ \" -definition "DS2 \ \ \ \a b. DS2\<^sup>a\<^sup>b \ \" -definition "DS3 \ \ \ \a b. DS3\<^sup>a\<^sup>b \ \" -definition "DS4 \ \ \ \a b. DS4\<^sup>a\<^sup>b \ \" -declare DS1_def[Defs] DS2_def[Defs] DS3_def[Defs] DS4_def[Defs] - -text\\noindent{All DS variants are in general independent from each other. However if we take classical implication -we can verify that the pairs DS1-DS3 and DS2-DS4 are indeed equivalent. }\ -lemma "DS1(\<^bold>\) \ = DS3(\<^bold>\) \" unfolding Defs by (metis impl_def join_def) -lemma "DS2(\<^bold>\) \ = DS4(\<^bold>\) \" unfolding Defs by (metis impl_def join_def) - -text\\noindent{Explore some (non)entailment relations.}\ -lemma DS1_nDNor: "DS1(\<^bold>\) \ \ nDNor \" unfolding Defs by (metis bottom_def impl_def join_def top_def) -lemma DS2_nNor: "DS2(\<^bold>\) \ \ nNor \" unfolding Defs by (metis bottom_def impl_def join_def top_def) -lemma lCoP2_DS1: "lCoP2(\<^bold>\) \ \ DS1(\<^bold>\) \" unfolding Defs conn by metis -lemma lCoP1_DS2: "lCoP1(\<^bold>\) \ \ DS2(\<^bold>\) \" unfolding Defs by (metis (no_types, lifting) conn) -lemma "CoP2 \ \ DS1(\<^bold>\) \" nitpick oops -lemma "CoP1 \ \ DS2(\<^bold>\) \" nitpick oops - -end diff --git a/thys/Topological_Semantics/sse_operation_negative_quantification.thy b/thys/Topological_Semantics/sse_operation_negative_quantification.thy deleted file mode 100644 --- a/thys/Topological_Semantics/sse_operation_negative_quantification.thy +++ /dev/null @@ -1,114 +0,0 @@ -theory sse_operation_negative_quantification - imports sse_operation_negative sse_boolean_algebra_quantification -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -subsection \Definitions and interrelations (infinitary case)\ - -text\\noindent{We define and interrelate infinitary variants for some previously introduced ('negative') conditions -on operations. We show how they relate to quantifiers as previously defined.}\ - -text\\noindent{iDM: infinitary De Morgan laws.}\ -abbreviation riDM1 ("iDM1\<^sup>_ _") where "iDM1\<^sup>S \ \ \(\<^bold>\S) \<^bold>\ \<^bold>\Ra[\|S]" -abbreviation riDM2 ("iDM2\<^sup>_ _") where "iDM2\<^sup>S \ \ \<^bold>\Ra[\|S] \<^bold>\ \(\<^bold>\S)" -abbreviation riDM3 ("iDM3\<^sup>_ _") where "iDM3\<^sup>S \ \ \(\<^bold>\S) \<^bold>\ \<^bold>\Ra[\|S]" -abbreviation riDM4 ("iDM4\<^sup>_ _") where "iDM4\<^sup>S \ \ \<^bold>\Ra[\|S] \<^bold>\ \(\<^bold>\S)" -definition "iDM1 \ \ \S. iDM1\<^sup>S \" -definition "iDM2 \ \ \S. iDM2\<^sup>S \" -definition "iDM3 \ \ \S. iDM3\<^sup>S \" -definition "iDM4 \ \ \S. iDM4\<^sup>S \" -declare iDM1_def[Defs] iDM2_def[Defs] iDM3_def[Defs] iDM4_def[Defs] - -lemma CoPw_iDM1: "CoPw \ \ iDM1 \" unfolding Defs by (smt Ra_restr_all sup_char) -lemma CoPw_iDM2: "CoPw \ \ iDM2 \" unfolding Defs by (smt Ra_restr_ex inf_char) -lemma CoP2_iDM3: "CoP2 \ \ iDM3 \" by (smt CoP2_def Ra_restr_ex iDM3_def inf_char) -lemma CoP1_iDM4: "CoP1 \ \ iDM4 \" by (smt CoP1_def Ra_restr_all iDM4_def sup_char) -lemma "XCoP \ \ iDM3 \" nitpick oops -lemma "XCoP \ \ iDM4 \" nitpick oops - -text\\noindent{DM1, DM2, iDM1, iDM2 and CoPw are equivalent.}\ -lemma iDM1_rel: "iDM1 \ \ DM1 \" proof - - assume idm1: "iDM1 \" - { fix a::"\" and b::"\" - let ?S="\Z. Z=a \ Z=b" - have "\<^bold>\Ra[\|?S] = \<^bold>\\<^sup>R\?S\ \" using Ra_restr_all by blast - moreover have "\<^bold>\\<^sup>R\?S\ \ = (\ a) \<^bold>\ (\ b)" using meet_def by auto - ultimately have "\<^bold>\Ra[\|?S] = (\ a) \<^bold>\ (\ b)" by simp - moreover have "\<^bold>\?S = a \<^bold>\ b" using supremum_def join_def by auto - moreover from idm1 have "\(\<^bold>\?S) \<^bold>\ \<^bold>\Ra[\|?S]" by (simp add: iDM1_def) - ultimately have "\(a \<^bold>\ b) \<^bold>\ (\ a) \<^bold>\ (\ b)" by simp - } thus ?thesis by (simp add: DM1_def) - qed -lemma iDM2_rel: "iDM2 \ \ DM2 \" proof - - assume idm2: "iDM2 \" - { fix a::"\" and b::"\" - let ?S="\Z. Z=a \ Z=b" - have "\<^bold>\Ra[\|?S] = \<^bold>\\<^sup>R\?S\ \" using Ra_restr_ex by blast - moreover have "\<^bold>\\<^sup>R\?S\ \ = (\ a) \<^bold>\ (\ b)" using join_def by auto - ultimately have "\<^bold>\Ra[\|?S] = (\ a) \<^bold>\ (\ b)" by simp - moreover have "\<^bold>\?S = a \<^bold>\ b" using infimum_def meet_def by auto - moreover from idm2 have "\<^bold>\Ra[\|?S] \<^bold>\ \(\<^bold>\?S)" by (simp add: iDM2_def) - ultimately have "(\ a) \<^bold>\ (\ b) \<^bold>\ \(a \<^bold>\ b)" by auto - } thus ?thesis by (simp add: DM2_def) -qed -lemma "DM1 \ = iDM1 \" using CoPw_iDM1 DM1_CoPw iDM1_rel by blast -lemma "DM2 \ = iDM2 \" using CoPw_iDM2 DM2_CoPw iDM2_rel by blast -lemma "iDM1 \ = iDM2 \" using CoPw_iDM1 CoPw_iDM2 DM1_CoPw DM2_CoPw iDM1_rel iDM2_rel by blast - -text\\noindent{iDM3/4 entail their finitary variants but not the other way round.}\ -lemma iDM3_rel: "iDM3 \ \ DM3 \" proof - - assume idm3: "iDM3 \" - { fix a::"\" and b::"\" - let ?S="\Z. Z=a \ Z=b" - have "\<^bold>\Ra[\|?S] = \<^bold>\\<^sup>R\?S\ \" using Ra_restr_ex by blast - moreover have "\<^bold>\\<^sup>R\?S\ \ = (\ a) \<^bold>\ (\ b)" using join_def by auto - ultimately have "\<^bold>\Ra[\|?S] = (\ a) \<^bold>\ (\ b)" by simp - moreover have "\<^bold>\?S = a \<^bold>\ b" using infimum_def meet_def by auto - moreover from idm3 have "\(\<^bold>\?S) \<^bold>\ \<^bold>\Ra[\|?S]" by (simp add: iDM3_def) - ultimately have "\(a \<^bold>\ b) \<^bold>\ (\ a) \<^bold>\ (\ b)" by auto - } thus ?thesis by (simp add: DM3_def) -qed -lemma iDM4_rel: "iDM4 \ \ DM4 \" proof - - assume idm4: "iDM4 \" - { fix a::"\" and b::"\" - let ?S="\Z. Z=a \ Z=b" - have "\<^bold>\Ra[\|?S] = \<^bold>\\<^sup>R\?S\ \" using Ra_restr_all by blast - moreover have "\<^bold>\\<^sup>R\?S\ \ = (\ a) \<^bold>\ (\ b)" using meet_def by auto - ultimately have "\<^bold>\Ra[\|?S] = (\ a) \<^bold>\ (\ b)" by simp - moreover have "\<^bold>\?S = a \<^bold>\ b" using supremum_def join_def by auto - moreover from idm4 have "\<^bold>\Ra[\|?S] \<^bold>\ \(\<^bold>\?S)" by (simp add: iDM4_def) - ultimately have "(\ a) \<^bold>\ (\ b) \<^bold>\ \(a \<^bold>\ b)" by simp - } thus ?thesis by (simp add: DM4_def) - qed -lemma "DM3 \ \ iDM3 \" nitpick oops -lemma "DM4 \ \ iDM4 \" nitpick oops - -text\\noindent{Indeed the previous characterization of the infinitary De Morgan laws is fairly general and entails -the traditional version employing quantifiers (though not the other way round).}\ -text\\noindent{The first two variants DM1/2 follow easily from DM1/2, iDM1/2 or CoPw (all of them equivalent).}\ -lemma iDM1_trad: "iDM1 \ \ \\. \(\<^bold>\x. \ x) \<^bold>\ (\<^bold>\x. \(\ x))" by (metis (mono_tags, lifting) CoPw_def DM1_CoPw iDM1_rel) -lemma iDM2_trad: "iDM2 \ \ \\. (\<^bold>\x. \(\ x)) \<^bold>\ \(\<^bold>\x. \ x)" by (metis (mono_tags, lifting) CoPw_def DM2_CoPw iDM2_rel) - -text\\noindent{An analogous relationship holds for variants DM3/4, though the proof is less trivial. -To see how let us first consider an intermediate version of the De Morgan laws, obtained as a -particular case of the general variant above, with S as the range of a propositional function.}\ -abbreviation "piDM1 \ \ \ \(\<^bold>\Ra(\)) \<^bold>\ \<^bold>\Ra[\|Ra(\)]" -abbreviation "piDM2 \ \ \ \<^bold>\Ra[\|Ra(\)] \<^bold>\ \(\<^bold>\Ra(\))" -abbreviation "piDM3 \ \ \ \(\<^bold>\Ra(\)) \<^bold>\ \<^bold>\Ra[\|Ra(\)]" -abbreviation "piDM4 \ \ \ \<^bold>\Ra[\|Ra(\)] \<^bold>\ \(\<^bold>\Ra(\))" - -text\\noindent{They are entailed (unidirectionally) by the general De Morgan laws.}\ -lemma "iDM1 \ \ \\. piDM1 \ \" by (simp add: iDM1_def) -lemma "iDM2 \ \ \\. piDM2 \ \" by (simp add: iDM2_def) -lemma "iDM3 \ \ \\. piDM3 \ \" by (simp add: iDM3_def) -lemma "iDM4 \ \ \\. piDM4 \ \" by (simp add: iDM4_def) - -text\\noindent{Drawing upon the relationships shown previously we can rewrite the latter two as:}\ -lemma iDM3_aux: "piDM3 \ \ \ \(\<^bold>\\) \<^bold>\ \<^bold>\[\|\Ra \\]\<^sup>N" unfolding Ra_all Ra_ex_restr by simp -lemma iDM4_aux: "piDM4 \ \ \ \<^bold>\[\|\Ra \\]\<^sup>P \<^bold>\ \(\<^bold>\\)" unfolding Ra_ex Ra_all_restr by simp - -text\\noindent{and thus finally obtain the desired formulas.}\ -lemma iDM3_trad: "iDM3 \ \ \\. \(\<^bold>\x. \ x) \<^bold>\ (\<^bold>\x. \(\ x))" by (metis Ra_ex_comp2 iDM3_def iDM3_aux comp_apply) -lemma iDM4_trad: "iDM4 \ \ \\. (\<^bold>\x. \(\ x)) \<^bold>\ \(\<^bold>\x. \ x)" by (metis Ra_all_comp1 iDM4_def iDM4_aux comp_apply) - -end diff --git a/thys/Topological_Semantics/sse_operation_positive.thy b/thys/Topological_Semantics/sse_operation_positive.thy deleted file mode 100644 --- a/thys/Topological_Semantics/sse_operation_positive.thy +++ /dev/null @@ -1,96 +0,0 @@ -theory sse_operation_positive - imports sse_boolean_algebra -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Positive semantic conditions for operations\ - -text\\noindent{We define and interrelate some useful conditions on propositional functions which do not involve -negative-like properties (hence 'positive'). We focus on propositional functions which correspond to unary -connectives of the algebra (with type @{text "\\\"}). We call such propositional functions 'operations'.}\ - -subsection \Definitions (finitary case)\ - -text\\noindent{Monotonicity (MONO).}\ -definition "MONO \ \ \A B. A \<^bold>\ B \ \ A \<^bold>\ \ B" -lemma MONO_ant: "MONO \ \ \A B C. A \<^bold>\ B \ \(B \<^bold>\ C) \<^bold>\ \(A \<^bold>\ C)" by (smt MONO_def conn) -lemma MONO_cons: "MONO \ \ \A B C. A \<^bold>\ B \ \(C \<^bold>\ A) \<^bold>\ \(C \<^bold>\ B)" by (smt MONO_def conn) -lemma MONO_dual: "MONO \ \ MONO \\<^sup>d" by (smt MONO_def dual_def compl_def) - -text\\noindent{Extensive/expansive (EXP) and its dual (dEXP), aka. 'contractive'.}\ -definition "EXP \ \ \A. A \<^bold>\ \ A" -definition "dEXP \ \ \A. \ A \<^bold>\ A" -lemma EXP_dual1: "EXP \ \ dEXP \\<^sup>d" by (metis EXP_def dEXP_def dual_def compl_def) -lemma EXP_dual2: "dEXP \ \ EXP \\<^sup>d" by (metis EXP_def dEXP_def dual_def compl_def) - -text\\noindent{Idempotence (IDEM).}\ -definition "IDEM \ \ \A. (\ A) \<^bold>\ \(\ A)" -definition "IDEMa \ \ \A. (\ A) \<^bold>\ \(\ A)" -definition "IDEMb \ \ \A. (\ A) \<^bold>\ \(\ A)" -lemma IDEM_dual1: "IDEMa \ \ IDEMb \\<^sup>d" unfolding dual_def IDEMa_def IDEMb_def compl_def by auto -lemma IDEM_dual2: "IDEMb \ \ IDEMa \\<^sup>d" unfolding dual_def IDEMa_def IDEMb_def compl_def by auto -lemma IDEM_dual: "IDEM \ = IDEM \\<^sup>d" by (metis IDEM_def IDEM_dual1 IDEM_dual2 IDEMa_def IDEMb_def dual_symm) - -text\\noindent{Normality (NOR) and its dual (dNOR).}\ -definition "NOR \ \ (\ \<^bold>\) \<^bold>\ \<^bold>\" -definition "dNOR \ \ (\ \<^bold>\) \<^bold>\ \<^bold>\" -lemma NOR_dual1: "NOR \ = dNOR \\<^sup>d" unfolding dual_def NOR_def dNOR_def top_def bottom_def compl_def by simp -lemma NOR_dual2: "dNOR \ = NOR \\<^sup>d" unfolding dual_def NOR_def dNOR_def top_def bottom_def compl_def by simp - -text\\noindent{Distribution over meets or multiplicativity (MULT).}\ -definition "MULT \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" -definition "MULT_a \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" -definition "MULT_b \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" - -text\\noindent{Distribution over joins or additivity (ADDI).}\ -definition "ADDI \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" -definition "ADDI_a \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" -definition "ADDI_b \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" - - -subsection \Relations among conditions (finitary case)\ - -text\\noindent{dEXP and dNOR entail NOR.}\ -lemma "dEXP \ \ dNOR \ \ NOR \" by (meson bottom_def dEXP_def NOR_def) - -text\\noindent{EXP and NOR entail dNOR.}\ -lemma "EXP \ \ NOR \ \ dNOR \" by (simp add: EXP_def dNOR_def top_def) - -text\\noindent{Interestingly, EXP and its dual allow for an alternative characterization of fixed-point operators.}\ -lemma EXP_fp: "EXP \ \ \\<^sup>f\<^sup>p \<^bold>\ (\\<^sup>c \<^bold>\ id)" by (smt id_def EXP_def dual_def dual_symm equal_op_def conn) -lemma dEXP_fp: "dEXP \ \ \\<^sup>f\<^sup>p \<^bold>\ (\ \<^bold>\ compl)" by (smt dEXP_def equal_op_def conn) - -text\\noindent{MONO, MULT-a and ADDI-b are equivalent.}\ -lemma MONO_MULTa: "MONO \ = MULT_a \" proof - - have lr: "MONO \ \ MULT_a \" by (smt MONO_def MULT_a_def meet_def) - have rl: "MULT_a \ \ MONO \" proof- - assume multa: "MULT_a \" - { fix A B - { assume "A \<^bold>\ B" - hence "A \<^bold>\ A \<^bold>\ B" unfolding conn by blast - hence "\ A \<^bold>\ \(A \<^bold>\ B)" unfolding conn by simp - moreover from multa have "\(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" using MULT_a_def by metis - ultimately have "\ A \<^bold>\ (\ A) \<^bold>\ (\ B)" by blast - hence "\ A \<^bold>\ (\ B)" unfolding conn by blast - } hence "A \<^bold>\ B \ \ A \<^bold>\ \ B" by (rule impI) - } thus ?thesis by (simp add: MONO_def) qed - from lr rl show ?thesis by auto -qed -lemma MONO_ADDIb: "MONO \ = ADDI_b \" proof - - have lr: "MONO \ \ ADDI_b \" by (smt ADDI_b_def MONO_def join_def) - have rl: "ADDI_b \ \ MONO \" proof - - assume addib: "ADDI_b \" - { fix A B - { assume "A \<^bold>\ B" - hence "B \<^bold>\ A \<^bold>\ B" unfolding conn by blast - hence "\ B \<^bold>\ \(A \<^bold>\ B)" unfolding conn by simp - moreover from addib have "(\ A) \<^bold>\ (\ B) \<^bold>\ \(A \<^bold>\ B)" using ADDI_b_def by metis - ultimately have "(\ A) \<^bold>\ (\ B) \<^bold>\ \ B" by blast - hence "\ A \<^bold>\ (\ B)" unfolding conn by blast - } hence "A \<^bold>\ B \ \ A \<^bold>\ \ B" by (rule impI) - } thus ?thesis by (simp add: MONO_def) qed - from lr rl show ?thesis by auto -qed -lemma ADDIb_MULTa: "ADDI_b \ = MULT_a \" using MONO_ADDIb MONO_MULTa by auto - -end diff --git a/thys/Topological_Semantics/sse_operation_positive_quantification.thy b/thys/Topological_Semantics/sse_operation_positive_quantification.thy deleted file mode 100644 --- a/thys/Topological_Semantics/sse_operation_positive_quantification.thy +++ /dev/null @@ -1,98 +0,0 @@ -theory sse_operation_positive_quantification - imports sse_operation_positive sse_boolean_algebra_quantification -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - - -subsection \Definitions (infinitary case)\ - -text\\noindent{We define and interrelate infinitary variants for some previously introduced ('positive') conditions -on operations and show how they relate to quantifiers as previously defined.}\ - -text\\noindent{Distribution over infinite meets (infima) or infinite multiplicativity (iMULT).}\ -definition "iMULT \ \ \S. \(\<^bold>\S) \<^bold>\ \<^bold>\Ra[\|S]" -definition "iMULT_a \ \ \S. \(\<^bold>\S) \<^bold>\ \<^bold>\Ra[\|S]" -definition "iMULT_b \ \ \S. \(\<^bold>\S) \<^bold>\ \<^bold>\Ra[\|S]" - -text\\noindent{Distribution over infinite joins (suprema) or infinite additivity (iADDI).}\ -definition "iADDI \ \ \S. \(\<^bold>\S) \<^bold>\ \<^bold>\Ra[\|S]" -definition "iADDI_a \ \ \S. \(\<^bold>\S) \<^bold>\ \<^bold>\Ra[\|S]" -definition "iADDI_b \ \ \S. \(\<^bold>\S) \<^bold>\ \<^bold>\Ra[\|S]" - - -subsection \Relations among conditions (infinitary case)\ - -text\\noindent{We start by noting that there is a duality between iADDI-a and iMULT-b.}\ -lemma iADDI_MULT_dual1: "iADDI_a \ \ iMULT_b \\<^sup>d" unfolding iADDI_a_def iMULT_b_def by (metis compl_def dual_def iDM_a iDM_b Ra_dual1) -lemma iADDI_MULT_dual2: "iMULT_b \ \ iADDI_a \\<^sup>d" unfolding iADDI_a_def iMULT_b_def by (metis compl_def dual_def iDM_b Ra_dual3) - -text\\noindent{MULT-a and iMULT-a are equivalent.}\ -lemma iMULTa_rel: "iMULT_a \ = MULT_a \" proof - - have lr: "iMULT_a \ \ MULT_a \" proof - - assume imulta: "iMULT_a \" - { fix A::"\" and B::"\" - let ?S="\Z. Z=A \ Z=B" - from imulta have "\(\<^bold>\?S) \<^bold>\ \<^bold>\\<^sup>R\?S\ \" by (simp add: iMULT_a_def Ra_restr_all) - moreover have "\<^bold>\?S = A \<^bold>\ B" using infimum_def meet_def by auto - moreover have "\<^bold>\\<^sup>R\?S\ \ = (\ A) \<^bold>\ (\ B)" using meet_def by auto - ultimately have "\(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" by smt - } thus ?thesis by (simp add: MULT_a_def) qed - have rl: "MULT_a \ \ iMULT_a \" by (smt MONO_def MONO_MULTa Ra_restr_all iMULT_a_def inf_char) - from lr rl show ?thesis by auto -qed -text\\noindent{ADDI-b and iADDI-b are equivalent.}\ -lemma iADDIb_rel: "iADDI_b \ = ADDI_b \" proof - - have lr: "iADDI_b \ \ ADDI_b \" proof - - assume iaddib: "iADDI_b \" - { fix A::"\" and B::"\" - let ?S="\Z. Z=A \ Z=B" - from iaddib have "\(\<^bold>\?S) \<^bold>\ \<^bold>\\<^sup>R\?S\(\)" by (simp add: iADDI_b_def Ra_restr_ex) - moreover have "\<^bold>\?S = A \<^bold>\ B" using supremum_def join_def by auto - moreover have "\<^bold>\\<^sup>R\?S\(\) = (\ A) \<^bold>\ (\ B)" using join_def by auto - ultimately have "\(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" by smt - } thus ?thesis by (simp add: ADDI_b_def) qed - have rl: "ADDI_b \ \ iADDI_b \" by (smt MONO_def MONO_ADDIb Ra_restr_ex iADDI_b_def sup_char) - from lr rl show ?thesis by auto -qed - -text\\noindent{Thus we have that MONO, MULT-a/iMULT-a and ADDI-b/iADDI-b are all equivalent.}\ -lemma MONO_iADDIb: "MONO \ = iADDI_b \" using MONO_ADDIb iADDIb_rel by simp -lemma MONO_iMULTa: "MONO \ = iMULT_a \" using MONO_MULTa iMULTa_rel by simp -lemma iADDI_b_iMULTa: "iADDI_b \ = iMULT_a \" using MONO_iADDIb MONO_iMULTa by auto - -lemma PI_imult: "MONO \ \ iMULT_b \ \ iMULT \" using MONO_MULTa iMULT_a_def iMULT_b_def iMULT_def iMULTa_rel by auto -lemma PC_iaddi: "MONO \ \ iADDI_a \ \ iADDI \" using MONO_ADDIb iADDI_a_def iADDI_b_def iADDI_def iADDIb_rel by auto - -text\\noindent{Interestingly, we can show that suitable (infinitary) conditions on an operation can make the set -of its fixed points closed under infinite meets/joins.}\ -lemma fp_inf_closed: "MONO \ \ iMULT_b \ \ infimum_closed (fp \)" by (metis (full_types) PI_imult Ra_restr_all iMULT_def infimum_def) -lemma fp_sup_closed: "MONO \ \ iADDI_a \ \ supremum_closed (fp \)" by (metis (full_types) PC_iaddi Ra_restr_ex iADDI_def supremum_def) - - -subsection \Exploring the Barcan formula and its converse\ - -text\\noindent{The converse Barcan formula follows readily from monotonicity.}\ -lemma CBarcan1: "MONO \ \ \\. \(\<^bold>\x. \ x) \<^bold>\ (\<^bold>\x. \(\ x))" by (metis (mono_tags, lifting) MONO_def) -lemma CBarcan2: "MONO \ \ \\. (\<^bold>\x. \(\ x)) \<^bold>\ \(\<^bold>\x. \ x)" by (metis (mono_tags, lifting) MONO_def) - -text\\noindent{However, the Barcan formula requires a stronger assumption (of an infinitary character).}\ -lemma Barcan1: "iMULT_b \ \ \\. (\<^bold>\x. \(\ x)) \<^bold>\ \(\<^bold>\x. \ x)" proof - - assume imultb: "iMULT_b \" - { fix \::"'a\\" - from imultb have "(\<^bold>\Ra(\\\)) \<^bold>\ \(\<^bold>\Ra(\))" unfolding iMULT_b_def by (smt comp_apply infimum_def pfunRange_def pfunRange_restr_def) - moreover have "\<^bold>\Ra(\) = (\<^bold>\x. \ x)" unfolding Ra_all by simp - moreover have "\<^bold>\Ra(\\\) = (\<^bold>\x. \(\ x))" unfolding Ra_all by simp - ultimately have "(\<^bold>\x. \(\ x)) \<^bold>\ \(\<^bold>\x. \ x)" by simp - } thus ?thesis by simp -qed -lemma Barcan2: "iADDI_a \ \ \\. \(\<^bold>\x. \ x) \<^bold>\ (\<^bold>\x. \(\ x))" proof - - assume iaddia: "iADDI_a \" - { fix \::"'a\\" - from iaddia have "\(\<^bold>\Ra(\)) \<^bold>\ (\<^bold>\Ra(\\\))" unfolding iADDI_a_def Ra_restr_ex by (smt fcomp_comp fcomp_def pfunRange_def sup_char) - moreover have "\<^bold>\Ra(\) = (\<^bold>\x. \ x)" unfolding Ra_ex by simp - moreover have "\<^bold>\Ra(\\\) = (\<^bold>\x. \(\ x))" unfolding Ra_ex by simp - ultimately have "\(\<^bold>\x. \ x) \<^bold>\ (\<^bold>\x. \(\ x))" by simp - } thus ?thesis by simp -qed - -end diff --git a/thys/Topological_Semantics/topo_alexandrov.thy b/thys/Topological_Semantics/topo_alexandrov.thy deleted file mode 100644 --- a/thys/Topological_Semantics/topo_alexandrov.thy +++ /dev/null @@ -1,139 +0,0 @@ -theory topo_alexandrov - imports sse_operation_positive_quantification -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - - -section \Generalized specialization orderings and Alexandrov topologies\ - -text\\noindent{A topology is called 'Alexandrov' (after the Russian mathematician Pavel Alexandrov) if the intersection -(resp. union) of any (finite or infinite) family of open (resp. closed) sets is open (resp. closed); -in algebraic terms, this means that the set of fixed points of the interior (closure) operation is closed -under infinite meets (joins). Another common algebraic formulation requires the closure (interior) operation -to satisfy the infinitary variants of additivity (multiplicativity), i.e. iADDI (iMULT) as introduced before. - -In the literature, the well-known Kuratowski conditions for the closure (resp. interior) operation are assumed, -namely: ADDI, EXP, NOR, IDEM (resp. MULT, dEXP, dNOR, IDEM). This makes both formulations equivalent. -However, this is not the case in general if those conditions become negotiable.}\ - -text\\noindent{Alexandrov topologies have interesting properties relating them to the semantics of modal logic. -Assuming Kuratowski conditions, Alexandrov topological operations defined on subsets of S are in one-to-one -correspondence with preorders on S; in topological terms, Alexandrov topologies are uniquely determined by -their specialization preorders. Since we do not presuppose any Kuratowski conditions to begin with, the -preorders in question are in general not even transitive. Here we just call them 'specialization relations'. -We will still call (generalized) closure/interior-like operations as such (for lack of a better name). -We explore minimal conditions under which some relevant results for the semantics of modal logic obtain.}\ - -subsection \Specialization relations\ - -text\\noindent{Specialization relations (among worlds/points) are particular cases of propositional functions with type @{text "w\\"}.}\ - -text\\noindent{Define some relevant properties of relations: }\ -abbreviation "serial R \ \x. \y. R x y" -abbreviation "reflexive R \ \x. R x x" -abbreviation "transitive R \ \x y z. R x y \ R y z \ R x z" -abbreviation "antisymmetric R \ \x y. R x y \ R y x \ x = y" -abbreviation "symmetric R \ \x y. R x y \ R y x" - -text\\noindent{Closure/interior operations can be derived from an arbitrary relation as operations returning down-/up-sets.}\ -definition Cl_rel::"(w\\)\(\\\)" ("\\<^sub>R") where "\\<^sub>R R \ \A. \w. \v. R w v \ A v" -definition Int_rel::"(w\\)\(\\\)" ("\\<^sub>R") where "\\<^sub>R R \ \A. \w. \v. R w v \ A v" - -text\\noindent{Duality between interior and closure follows directly:}\ -lemma dual_rel1: "\A. (\\<^sub>R R) A \<^bold>\ (\\<^sub>R R)\<^sup>d A" unfolding Cl_rel_def Int_rel_def dual_def conn by simp -lemma dual_rel2: "\A. (\\<^sub>R R) A \<^bold>\ (\\<^sub>R R)\<^sup>d A" unfolding Cl_rel_def Int_rel_def dual_def conn by simp - -text\\noindent{We explore minimal conditions of the specialization relation under which some operation's conditions obtain.}\ -lemma rC1: "ADDI (\\<^sub>R R)" unfolding Cl_rel_def ADDI_def conn by blast -lemma rC1i:"iADDI (\\<^sub>R R)" by (smt Cl_rel_def Ra_restr_ex iADDI_def supremum_def) -lemma rC2: "reflexive R \ EXP (\\<^sub>R R)" unfolding EXP_def Cl_rel_def by auto -lemma rC3: "NOR (\\<^sub>R R)" unfolding Cl_rel_def NOR_def conn by blast -lemma rC4: "reflexive R \ transitive R \ IDEM (\\<^sub>R R)" unfolding Cl_rel_def IDEM_def by smt -lemma rC_Barcan: "\\. (\\<^sub>R R)(\<^bold>\x. \ x) \<^bold>\ (\<^bold>\x. (\\<^sub>R R)(\ x))" unfolding Cl_rel_def by auto - -lemma rI1: "MULT (\\<^sub>R R)" unfolding Int_rel_def MULT_def conn by blast -lemma rI1i:"iMULT (\\<^sub>R R)" by (smt Int_rel_def Ra_restr_all iMULT_def infimum_def) -lemma rI2: "reflexive R \ dEXP (\\<^sub>R R)" unfolding Int_rel_def dEXP_def Int_rel_def by auto -lemma rI3: "dNOR (\\<^sub>R R)" unfolding Int_rel_def dNOR_def conn by simp -lemma rI4: "reflexive R \ transitive R \ IDEM (\\<^sub>R R)" unfolding IDEM_def Int_rel_def by smt -lemma rI_Barcan: "\\. (\<^bold>\x. (\\<^sub>R R)(\ x)) \<^bold>\ (\\<^sub>R R)(\<^bold>\x. \ x)" unfolding Int_rel_def by simp - -text\\noindent{A specialization relation can be derived from a given operation (intended as a closure-like operation).}\ -definition sp_rel::"(\\\)\(w\\)" ("\\<^sup>C") where "\\<^sup>C \ \ \w v. \ (\u. u=v) w" - -text\\noindent{Preorder properties of the specialization relation follow directly from the corresponding operation's conditions.}\ -lemma sp_rel_reflex: "EXP \ \ reflexive (\\<^sup>C \)" by (simp add: EXP_def sp_rel_def) -lemma sp_rel_trans: "MONO \ \ IDEM \ \ transitive (\\<^sup>C \)" by (smt IDEM_def MONO_def sp_rel_def) - -text\\noindent{However, we can obtain finite countermodels for antisymmetry and symmetry given all relevant conditions. -We will revisit this issue later and examine their relation with the topological separation axioms T0 and T1 resp.}\ -lemma "iADDI \ \ EXP \ \ NOR \ \ IDEM \ \ antisymmetric (\\<^sup>C \)" nitpick oops (*counterexample*) -lemma "iADDI \ \ EXP \ \ NOR \ \ IDEM \ \ symmetric (\\<^sup>C \)" nitpick oops (*counterexample*) - - -subsection \Alexandrov topology\ - -text\\noindent{As mentioned previously, Alexandrov closure (and by duality interior) operations correspond to specialization -relations. It is worth mentioning that in Alexandrov topologies every point has a minimal/smallest neighborhood, -namely the set of points related to it by the specialization (aka. accessibility) relation. Alexandrov spaces are -thus also called 'finitely generated'. We examine below minimal conditions under which these relations obtain.}\ - -lemma sp_rel_a: "MONO \ \ \A. (\\<^sub>R (\\<^sup>C \)) A \<^bold>\ \ A" by (smt Cl_rel_def MONO_def sp_rel_def) -lemma sp_rel_b: "iADDI_a \ \ \A. (\\<^sub>R (\\<^sup>C \)) A \<^bold>\ \ A" proof - - assume iaddia: "iADDI_a \" - { fix A - let ?S="\B::\. \w::w. A w \ B=(\u. u=w)" - have "A \<^bold>\ (\<^bold>\?S)" using supremum_def by auto - hence "\(A) \<^bold>\ \(\<^bold>\?S)" by (smt eq_ext) - moreover have "\<^bold>\Ra[\|?S] \<^bold>\ (\\<^sub>R (\\<^sup>C \)) A" by (smt Cl_rel_def Ra_restr_ex sp_rel_def) - moreover from iaddia have "\(\<^bold>\?S) \<^bold>\ \<^bold>\Ra[\|?S]" unfolding iADDI_a_def by simp - ultimately have "\ A \<^bold>\ (\\<^sub>R (\\<^sup>C \)) A" by simp - } thus ?thesis by simp -qed -lemma sp_rel: "iADDI \ \ \A. \ A \<^bold>\ (\\<^sub>R (\\<^sup>C \)) A" by (metis MONO_iADDIb iADDI_a_def iADDI_b_def iADDI_def sp_rel_a sp_rel_b) -text\\noindent{It is instructive to expand the definitions in the above lemma:}\ -lemma "iADDI \ \ \A. \w. (\ A) w \ (\v. A v \ (\ (\u. u=v)) w)" using Cl_rel_def sp_rel by fastforce - - -text\\noindent{We now turn to the more traditional characterization of Alexandrov topologies in terms of closure under -infinite joins/meets.}\ - -text\\noindent{Fixed points of operations satisfying ADDI (MULT) are not in general closed under infinite joins (meets). -For the given conditions countermodels are expected to be infinite. We (sanity) check that nitpick cannot find any.}\ -lemma "ADDI(\) \ supremum_closed (fp \)" (*nitpick*) oops (*cannot find finite countermodels*) -lemma "MULT(\) \ infimum_closed (fp \)" (*nitpick*) oops (*cannot find finite countermodels*) - -text\\noindent{By contrast, we can show that this obtains if assuming the corresponding infinitary variants (iADDI/iMULT).}\ -lemma "iADDI(\) \ supremum_closed (fp \)" by (metis (full_types) Ra_restr_ex iADDI_def supremum_def) -lemma "iMULT(\) \ infimum_closed (fp \)" by (metis (full_types) Ra_restr_all iMULT_def infimum_def) - -text\\noindent{As shown above, closure (interior) operations derived from relations readily satisfy iADDI (iMULT), -being thus closed under infinite joins (meets).}\ -lemma "supremum_closed (fp (\\<^sub>R R))" by (smt Cl_rel_def supremum_def) -lemma "infimum_closed (fp (\\<^sub>R R))" by (smt Int_rel_def infimum_def) - - -subsection \(Anti)symmetry and the separation axioms T0 and T1\ -text\\noindent{We can now revisit the relationship between (anti)symmetry and the separation axioms T1 and T0.}\ - -text\\noindent{T0: any two distinct points in the space can be separated by an open set (i.e. containing one point and not the other).}\ -abbreviation "T0_sep \ \ \w v. w \ v \ (\G. (fp \\<^sup>d)(G) \ (G w \ G v))" -text\\noindent{T1: any two distinct points can be separated by (two not necessarily disjoint) open sets, i.e. all singletons are closed.}\ -abbreviation "T1_sep \ \ \w. (fp \)(\u. u = w)" - -text\\noindent{We can (sanity) check that T1 entails T0 but not viceversa.}\ -lemma "T0_sep \ \ T1_sep \" nitpick oops (*counterexample*) -lemma "T1_sep \ \ T0_sep \" by (smt compl_def dual_def dual_symm) - -text\\noindent{Under appropriate conditions, T0-separation corresponds to antisymmetry of the specialization relation (here an ordering).}\ -lemma "T0_sep \ \ antisymmetric (\\<^sup>C \)" nitpick oops (*counterexample*) -lemma T0_antisymm_a: "MONO \ \ T0_sep \ \ antisymmetric (\\<^sup>C \)" by (smt Cl_rel_def compl_def dual_def sp_rel_a) -lemma T0_antisymm_b: "EXP \ \ IDEM \ \ antisymmetric (\\<^sup>C \) \ T0_sep \" by (metis (full_types) EXP_dual1 IDEM_def IDEM_dual2 IDEMa_def IDEMb_def compl_def dEXP_def dual_def dual_symm sp_rel_def) -lemma T0_antisymm: "MONO \ \ EXP \ \ IDEM \ \ T0_sep \ = antisymmetric (\\<^sup>C \)" by (metis T0_antisymm_a T0_antisymm_b) - -text\\noindent{Also, under the appropriate conditions, T1-separation corresponds to symmetry of the specialization relation.}\ -lemma T1_symm_a: "T1_sep \ \ symmetric (\\<^sup>C \)" using sp_rel_def by auto -lemma T1_symm_b: "MONO \ \ EXP \ \ T0_sep \ \ symmetric (\\<^sup>C \) \ T1_sep \" by (metis T0_antisymm_a sp_rel_def sp_rel_reflex) -lemma T1_symm: "MONO \ \ EXP \ \ T0_sep \ \ symmetric (\\<^sup>C \) = T1_sep \" by (metis T1_symm_a T1_symm_b) - -end diff --git a/thys/Topological_Semantics/topo_border_algebra.thy b/thys/Topological_Semantics/topo_border_algebra.thy deleted file mode 100644 --- a/thys/Topological_Semantics/topo_border_algebra.thy +++ /dev/null @@ -1,54 +0,0 @@ -theory topo_border_algebra - imports topo_operators_basic -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Border algebra\ -text\\noindent{We define a border algebra in an analogous fashion to the well-known closure/interior algebras. -We also verify a few interesting properties.}\ - -text\\noindent{Declares a primitive (unconstrained) border operation and defines others from it.}\ -consts \::"\\\" -abbreviation "\ \ \\<^sub>B \" \\ interior \ -abbreviation "\ \ \\<^sub>B \" \\ closure \ -abbreviation "\ \ \\<^sub>B \" \\ frontier \ - - -subsection \Basic properties\ - -text\\noindent{Verifies minimal conditions under which operators resulting from conversion functions coincide.}\ -lemma ICdual: "\ \<^bold>\ \\<^sup>d" by (simp add: Cl_br_def Int_br_def dual_def equal_op_def conn) -lemma ICdual': "\ \<^bold>\ \\<^sup>d" by (simp add: Cl_br_def Int_br_def dual_def equal_op_def conn) -lemma FI_rel: "Br_1 \ \ \ \<^bold>\ \\<^sub>I \" using Fr_br_def Fr_int_def Int_br_def equal_op_def by (smt Br_5b_def PB5b dual_def conn) -lemma IF_rel: "Br_1 \ \ \ \<^bold>\ \\<^sub>F \" using Br_5b_def Fr_br_def Int_br_def Int_fr_def PB5b unfolding equal_op_def conn by fastforce -lemma FC_rel: "Br_1 \ \ \ \<^bold>\ \\<^sub>C \" using Br_5b_def Cl_br_def Fr_br_def Fr_cl_def PB5b unfolding equal_op_def conn by fastforce -lemma CF_rel: "Br_1 \ \ \ \<^bold>\ \\<^sub>F \" using Br_5b_def Cl_br_def Cl_fr_def Fr_br_def PB5b unfolding equal_op_def conn by fastforce -lemma BI_rel: "Br_1 \ \ \ \<^bold>\ \\<^sub>I \" using Br_5b_def Br_int_def Int_br_def PB5b diff_def equal_op_def by fastforce -lemma BC_rel: "Br_1 \ \ \ \<^bold>\ \\<^sub>C \" using BI_BC_rel BI_rel ICdual' eq_ext' by fastforce -lemma BF_rel: "Br_1 \ \ \ \<^bold>\ \\<^sub>F \" by (smt BI_rel Br_fr_def Br_int_def IF_rel Int_fr_def diff_def equal_op_def meet_def) - -text\\noindent{Fixed-point and other operators are interestingly related.}\ -lemma fp1: "Br_1 \ \ \\<^sup>f\<^sup>p \<^bold>\ \\<^sup>c" using Br_5b_def Int_br_def PB5b unfolding equal_op_def conn by fastforce -lemma fp2: "Br_1 \ \ \\<^sup>f\<^sup>p \<^bold>\ \\<^sup>c" using Br_5b_def Int_br_def PB5b conn equal_op_def by fastforce -lemma fp3: "Br_1 \ \ \\<^sup>f\<^sup>p \<^bold>\ \\<^sup>d" using Br_5c_def Cl_br_def PB5c dual_def unfolding equal_op_def conn by fastforce -lemma fp4: "Br_1 \ \ (\\<^sup>d)\<^sup>f\<^sup>p \<^bold>\ \" by (smt dimp_def equal_op_def fp3) -lemma fp5: "Br_1 \ \ \\<^sup>f\<^sup>p \<^bold>\ \ \<^bold>\ (\\<^sup>c)" by (smt Br_5b_def Cl_br_def Fr_br_def PB5b equal_op_def conn) - -text\\noindent{Define some fixed-point predicates and prove some properties.}\ -abbreviation openset ("Op") where "Op A \ fp \ A" -abbreviation closedset ("Cl") where "Cl A \ fp \ A" -abbreviation borderset ("Br") where "Br A \ fp \ A" -abbreviation frontierset ("Fr") where "Fr A \ fp \ A" - -lemma Int_Open: "Br_1 \ \ Br_3 \ \ \A. Op(\ A)" using IB4 IDEM_def by blast -lemma Cl_Closed: "Br_1 \ \ Br_3 \ \ \A. Cl(\ A)" using CB4 IDEM_def by blast -lemma Br_Border: "Br_1 \ \ \A. Br(\ A)" using IDEM_def PB6 by blast -text\\noindent{In contrast, there is no analogous fixed-point result for frontier:}\ -lemma "\ \ \ \A. Fr(\ A)" nitpick oops (*counterexample even if assuming all border conditions*) - -lemma OpCldual: "\A. Cl A \ Op(\<^bold>\A)" using Cl_br_def Int_br_def conn by auto -lemma ClOpdual: "\A. Op A \ Cl(\<^bold>\A)" using Cl_br_def Int_br_def conn by auto -lemma Fr_ClBr: "Br_1 \ \ \A. Fr(A) = (Cl(A) \ Br(A))" by (metis BF_rel Br_fr_def CF_rel Cl_fr_def eq_ext' join_def meet_def) -lemma Cl_F: "Br_1 \ \ Br_3 \ \ \A. Cl(\ A)" by (metis CF_rel Cl_fr_def FB4 Fr_4_def eq_ext' join_def) - -end diff --git a/thys/Topological_Semantics/topo_closure_algebra.thy b/thys/Topological_Semantics/topo_closure_algebra.thy deleted file mode 100644 --- a/thys/Topological_Semantics/topo_closure_algebra.thy +++ /dev/null @@ -1,53 +0,0 @@ -theory topo_closure_algebra - imports topo_operators_basic -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Closure algebra\ -text\\noindent{We define a topological Boolean algebra with a primitive closure operator and verify a few properties.}\ - -text\\noindent{Declares a primitive (unconstrained) closure operation and defines others from it.}\ -consts \::"\\\" -abbreviation "\ \ \\<^sup>d" \\ interior \ -abbreviation "\ \ \\<^sub>C \" \\ border \ -abbreviation "\ \ \\<^sub>C \" \\ frontier \ - - -subsection \Basic properties\ - -text\\noindent{Verifies minimal conditions under which operators resulting from conversion functions coincide.}\ -lemma ICdual': "\ \<^bold>\ \\<^sup>d" using dual_symm equal_op_def by auto -lemma IB_rel: "Cl_2 \ \ \ \<^bold>\ \\<^sub>B \" using Br_cl_def EXP_dual1 Int_br_def compl_def dEXP_def diff_def dual_def equal_op_def meet_def by fastforce -lemma IF_rel: "Cl_2 \ \ \ \<^bold>\ \\<^sub>F \" by (smt EXP_def Fr_cl_def Int_fr_def compl_def diff_def dual_def equal_op_def meet_def) -lemma CB_rel: "Cl_2 \ \ \ \<^bold>\ \\<^sub>B \" by (smt Cl_br_def EXP_def IB_rel Int_br_def compl_def diff_def dual_def dual_symm eq_ext' equal_op_def join_def) -lemma CF_rel: "Cl_2 \ \ \ \<^bold>\ \\<^sub>F \" by (smt Br_cl_def CB_rel Cl_br_def Cl_fr_def Fr_cl_def equal_op_def join_def meet_def) -lemma BI_rel: "\ \<^bold>\ \\<^sub>I \" using BI_BC_rel dual_symm equal_op_def by metis -lemma BF_rel: "Cl_2 \ \ \ \<^bold>\ \\<^sub>F \" by (smt Br_cl_def Br_fr_def EXP_def Fr_cl_def equal_op_def meet_def) -lemma FI_rel: "\ \<^bold>\ \\<^sub>I \" by (metis FI2 Fr_2_def Fr_cl_def Fr_int_def ICdual' dual_def eq_ext' equal_op_def) -lemma FB_rel: "Cl_2 \ \ \ \<^bold>\ \\<^sub>B \" by (smt BF_rel Br_fr_def CB_rel Cl_br_def Fr_br_def Fr_cl_def equal_op_def join_def meet_def) - -text\\noindent{Fixed-point and other operators are interestingly related.}\ -lemma fp1: "Cl_2 \ \ \\<^sup>f\<^sup>p \<^bold>\ \\<^sup>c" by (smt BI_rel Br_int_def IB_rel Int_br_def compl_def diff_def dimp_def equal_op_def) -lemma fp2: "Cl_2 \ \ \\<^sup>f\<^sup>p \<^bold>\ \\<^sup>c" using fp1 by (smt compl_def dimp_def equal_op_def) -lemma fp3: "Cl_2 \ \ \\<^sup>f\<^sup>p \<^bold>\ \\<^sup>d" by (smt BI_rel Br_int_def CB_rel Cl_br_def compl_def diff_def dimp_def dual_def equal_op_def join_def) -lemma fp4: "Cl_2 \ \ (\\<^sup>d)\<^sup>f\<^sup>p \<^bold>\ \" by (smt dimp_def equal_op_def fp3) -lemma fp5: "Cl_2 \ \ \\<^sup>f\<^sup>p \<^bold>\ \ \<^bold>\ (\\<^sup>c)" by (smt Br_cl_def CF_rel Cl_fr_def FC2 Fr_2_def compl_def dimp_def eq_ext' equal_op_def join_def meet_def) - -text\\noindent{Define some fixed-point predicates and prove some properties.}\ -abbreviation openset ("Op") where "Op A \ fp \ A" -abbreviation closedset ("Cl") where "Cl A \ fp \ A" -abbreviation borderset ("Br") where "Br A \ fp \ A" -abbreviation frontierset ("Fr") where "Fr A \ fp \ A" - -lemma Int_Open: "Cl_4 \ \ \A. Op(\ A)" using IC4 IDEM_def by blast -lemma Cl_Closed: "Cl_4 \ \ \A. Cl(\ A)" by (simp add: IDEM_def) -lemma Br_Border: "Cl_1b \ \ \A. Br(\ A)" by (metis BI_rel Br_cl_def Br_int_def CI1b MULT_a_def diff_def eq_ext' meet_def) -text\\noindent{In contrast, there is no analogous fixed-point result for frontier:}\ -lemma "\ \ \ \A. Fr(\ A)" nitpick oops (*counterexample even if assuming all closure conditions*) - -lemma OpCldual: "\A. Cl A \ Op(\<^bold>\A)" by (simp add: compl_def dual_def) -lemma ClOpdual: "\A. Op A \ Cl(\<^bold>\A)" by (simp add: fp_d) -lemma Fr_ClBr: "Cl_2 \ \ \A. Fr(A) = (Cl(A) \ Br(A))" using BF_rel by (metis Br_fr_def CF_rel Cl_fr_def eq_ext' join_def meet_def) -lemma Cl_F: "Cl_1b \ \ Cl_2 \ \ Cl_4 \ \ \A. Cl(\ A)" using Cl_8_def Fr_cl_def PC8 by auto - -end diff --git a/thys/Topological_Semantics/topo_derivative_algebra.thy b/thys/Topological_Semantics/topo_derivative_algebra.thy deleted file mode 100644 --- a/thys/Topological_Semantics/topo_derivative_algebra.thy +++ /dev/null @@ -1,204 +0,0 @@ -theory topo_derivative_algebra - imports topo_operators_derivative -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Derivative algebra\ -text\\noindent{The closure of a set A (@{text "\(A)"}) can be seen as the set A augmented by (i) its boundary points, or -(ii) its accumulation/limit points. We explore the second variant by drawing on the notion of derivative algebra.}\ - -text\\noindent{Declares a primitive (unconstrained) derivative (aka. derived-set) operation and defines others from it.}\ -consts \::"\\\" -abbreviation "\ \ \\<^sub>D \" \\ interior \ -abbreviation "\ \ \\<^sub>D \" \\ closure \ -abbreviation "\ \ \\<^sub>D \" \\ border \ -abbreviation "\ \ \\<^sub>D \" \\ frontier \ -abbreviation "\ \ \\<^sub>D \" \\ coherence \ - - -subsection \Basic properties\ - -text\\noindent{Verifies minimal conditions under which operators resulting from conversion functions coincide.}\ -lemma ICdual: "\ \<^bold>\ \\<^sup>d" by (simp add: dual_der2 equal_op_def) -lemma ICdual': "\ \<^bold>\ \\<^sup>d" by (simp add: dual_der1 equal_op_def) -lemma BI_rel: "\ \<^bold>\ \\<^sub>I \" using Br_der_def Br_int_def Int_der_def unfolding equal_op_def conn by auto -lemma IB_rel: "\ \<^bold>\ \\<^sub>B \" using Br_der_def Int_br_def Int_der_def unfolding equal_op_def conn by auto -lemma BC_rel: "\ \<^bold>\ \\<^sub>C \" using BI_BC_rel BI_rel dual_der1 by auto -lemma CB_rel: "\ \<^bold>\ \\<^sub>B \" using Br_der_def2 Cl_br_def Int_der_def2 dual_def dual_der1 unfolding equal_op_def conn by auto -lemma FI_rel: "\ \<^bold>\ \\<^sub>I \" by (metis Cl_der_def FI2 Fr_2_def Fr_der_def2 Fr_int_def ICdual' dual_def eq_ext' equal_op_def) -lemma FC_rel: "\ \<^bold>\ \\<^sub>C \" by (simp add: Cl_der_def Fr_cl_def Fr_der_def2 equal_op_def) -lemma FB_rel: "\ \<^bold>\ \\<^sub>B \" by (smt Br_der_def CB_rel Cl_br_def Cl_der_def Fr_br_def Fr_der_def Fr_der_def2 equal_op_def conn) - -text\\noindent{Recall that derivative and coherence operations cannot be obtained from either interior, closure, border -nor frontier. The derivative operation can indeed be seen as being more fundamental than the other ones.}\ - -text\\noindent{Fixed-point and other operators are interestingly related.}\ -lemma fp1: "\\<^sup>f\<^sup>p \<^bold>\ \\<^sup>c" by (smt BI_rel Br_int_def IB_rel Int_br_def equal_op_def conn) -lemma fp2: "\\<^sup>f\<^sup>p \<^bold>\ \\<^sup>c" using Br_der_def Int_der_def unfolding equal_op_def conn by auto -lemma fp3: "\\<^sup>f\<^sup>p \<^bold>\ \\<^sup>d" by (smt BI_rel Br_int_def CB_rel Cl_br_def dual_def equal_op_def conn) -lemma fp4: "(\\<^sup>d)\<^sup>f\<^sup>p \<^bold>\ \" by (smt dimp_def equal_op_def fp3) -lemma fp5: "\\<^sup>f\<^sup>p \<^bold>\ \ \<^bold>\ (\\<^sup>c)" using Br_der_def Cl_der_def Fr_der_def unfolding equal_op_def conn by auto -lemma fp6: "\\<^sup>f\<^sup>p \<^bold>\ \ \<^bold>\ (\\<^sup>c)" using Cl_der_def Kh_der_def equal_op_def conn by fastforce - -text\\noindent{Different inter-relations (some redundant ones are kept to help the provers).}\ -lemma monI: "Der_1b \ \ MONO \" by (simp add: ID1b MONO_MULTa) -lemma monC: "Der_1b \ \ MONO \" by (simp add: CD1b MONO_ADDIb) -lemma pB1: "\A. \ A \<^bold>\ A \<^bold>\ \ A" using BI_rel Br_int_def eq_ext' by fastforce -lemma pB2: "\A. \ A \<^bold>\ A \<^bold>\ \ A" using Br_der_def Fr_der_def conn by auto -lemma pB3: "\A. \(\<^bold>\A) \<^bold>\ \<^bold>\A \<^bold>\ \ A" using FD2 Fr_2_def meet_def pB2 by auto -lemma pB4: "\A. \(\<^bold>\A) \<^bold>\ \<^bold>\A \<^bold>\ \ A" by (simp add: dual_def dual_der1 pB1 conn) -lemma pB5: "Der_1b \ \ \A. \(\ A) \<^bold>\ (\ A) \<^bold>\ \(\<^bold>\A)" using ADDI_b_def Cl_der_def MONO_ADDIb monI pB1 pB4 unfolding conn by auto -lemma pF1: "\A. \ A \<^bold>\ \ A \<^bold>\ \ A" using Cl_der_def Fr_der_def Int_der_def conn by auto -lemma pF2: "\A. \ A \<^bold>\ \ A \<^bold>\ \(\<^bold>\A)" by (simp add: Cl_der_def Fr_der_def2) -lemma pF3: "\A. \ A \<^bold>\ \ A \<^bold>\ \(\<^bold>\A)" by (smt Br_der_def Cl_der_def dual_def dual_der1 dual_der2 pF2 conn) -lemma pF4: "Der_1 \ \ Der_4e \ \ \A. \(\ A) \<^bold>\ \ A" by (metis CD1 CD2 CD4a ICdual ID4 IDEM_def PC1 PC4 PC5 PD8 diff_def eq_ext' pF1) -lemma pF5: "Der_1 \ \ Der_4e \ \ \A. \(\ A) \<^bold>\ \ A" by (metis FD2 Fr_2_def ICdual' dual_def eq_ext' pF4) -lemma pA1: "\A. A \<^bold>\ \ A \<^bold>\ \ A" using Br_der_def2 Int_der_def2 conn by auto -lemma pA2: "\A. A \<^bold>\ \ A \<^bold>\ \(\<^bold>\A)" using Cl_der_def pB4 conn by auto -lemma pC1: "\A. \ A \<^bold>\ A \<^bold>\ \(\<^bold>\A)" using CB_rel Cl_br_def eq_ext' by fastforce -lemma pC2: "\A. \ A \<^bold>\ A \<^bold>\ \ A" using Cl_der_def Fr_der_def2 conn by auto -lemma pI1: "\A. \ A \<^bold>\ A \<^bold>\ \ A" using pA1 pB1 conn by auto -lemma pI2: "\A. \ A \<^bold>\ A \<^bold>\ \ A" using Br_der_def Fr_der_def pI1 conn by auto - -lemma IC_imp: "Der_1 \ \ Der_3 \ \ \A B. \(A \<^bold>\ B) \<^bold>\ \ A \<^bold>\ \ B" proof - - assume der1: "Der_1 \" and der3: "Der_3 \" - { fix a b - have "(a \<^bold>\ b) \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\a = \<^bold>\" unfolding conn by auto - hence "\((a \<^bold>\ b) \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\a) \<^bold>\ \(\<^bold>\)" by simp - hence "\((a \<^bold>\ b) \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\a) \<^bold>\ \<^bold>\" using der3 dNOR_def using ID3 by auto - moreover have "let A=(a \<^bold>\ b) \<^bold>\ \<^bold>\b; B=\<^bold>\a in \(A \<^bold>\ B) \<^bold>\ \(A) \<^bold>\ \(B)" using ID1 Int_7_def PI7 der1 by auto - ultimately have "\((a \<^bold>\ b) \<^bold>\ \<^bold>\b) \<^bold>\ \(\<^bold>\a) \<^bold>\ \<^bold>\" unfolding conn by simp - moreover have "let A=a \<^bold>\ b; B=\<^bold>\b in \(A \<^bold>\ B) \<^bold>\ \(A) \<^bold>\ \(B)" using ID1 MULT_def der1 by auto - ultimately have "\(a \<^bold>\ b) \<^bold>\ \(\<^bold>\b) \<^bold>\ \(\<^bold>\a) \<^bold>\ \<^bold>\" unfolding conn by simp - moreover have "\A. \(\<^bold>\A) \<^bold>\ \<^bold>\\(A)" by (simp add: dual_def dual_der1 conn) - ultimately have "\(a \<^bold>\ b) \<^bold>\ \<^bold>\\(b) \<^bold>\ \<^bold>\\(a) \<^bold>\ \<^bold>\" unfolding conn by simp - hence "\(a \<^bold>\ b) \<^bold>\ \<^bold>\\(b) \<^bold>\ \<^bold>\\(a)" unfolding conn by simp - hence "\(a \<^bold>\ b) \<^bold>\ \ a \<^bold>\ \ b" unfolding conn by metis - } thus ?thesis unfolding conn by simp -qed - -text\\noindent{Define some fixed-point predicates and prove some properties.}\ -abbreviation openset ("Op") where "Op A \ fp \ A" -abbreviation closedset ("Cl") where "Cl A \ fp \ A" -abbreviation borderset ("Br") where "Br A \ fp \ A" -abbreviation frontierset ("Fr") where "Fr A \ fp \ A" - -lemma Int_Open: "Der_1a \ \ Der_4e \ \ \A. Op(\ A)" using ID4a IDEM_def by blast -lemma Cl_Closed: "Der_1a \ \ Der_4e \ \ \A. Cl(\ A)" using CD4a IDEM_def by blast -lemma Br_Border: "Der_1b \ \ \A. Br(\ A)" by (smt Br_der_def CI1b IC1_dual PD1 conn) -text\\noindent{In contrast, there is no analogous fixed-point result for frontier:}\ -lemma "\
\ \ \A. Fr(\ A)" nitpick oops (*counterexample even if assuming all derivative conditions*) - -lemma OpCldual: "\A. Cl A \ Op(\<^bold>\A)" using dual_def dual_der1 conn by auto -lemma ClOpdual: "\A. Op A \ Cl(\<^bold>\A)" by (simp add: dual_def dual_der1 conn) -lemma Fr_ClBr: "\A. Fr(A) = (Cl(A) \ Br(A))" using join_def meet_def pB2 pC2 by auto -lemma Cl_F: "Der_1 \ \ Der_4e \ \ \A. Cl(\ A)" using FD4 Fr_4_def join_def pC2 by auto - - -subsection \Further properties\ - -text\\noindent{The definitions and theorems below are well known in the literature (e.g. \<^cite>\Kuratowski2\). -Here we uncover the minimal conditions under which they hold (taking derivative operation as primitive).}\ -lemma Cl_Bzero: "\A. Cl A \ \(\<^bold>\A) \<^bold>\ \<^bold>\" using pA2 pC1 unfolding conn by metis -lemma Op_Bzero: "\A. Op A \ \ A \<^bold>\ \<^bold>\" using pB1 pI1 unfolding conn by metis -lemma Br_boundary: "\A. Br(A) \ \ A \<^bold>\ \<^bold>\" using Br_der_def2 Int_der_def2 unfolding conn by metis -lemma Fr_nowhereDense: "\A. Fr(A) \ \(\ A) \<^bold>\ \<^bold>\" using Fr_ClBr Br_boundary eq_ext by metis -lemma Cl_FB: "\A. Cl A \ \ A \<^bold>\ \ A" using Br_der_def2 pA2 pF1 pF3 unfolding conn by metis -lemma Op_FB: "\A. Op A \ \ A \<^bold>\ \(\<^bold>\A)" using pA1 pA2 pF3 pI2 unfolding conn by metis -lemma Clopen_Fzero: "\A. Cl A \ Op A \ \ A \<^bold>\ \<^bold>\" using Cl_der_def Int_der_def Fr_der_def unfolding conn by smt - -lemma Int_sup_closed: "Der_1b \ \ supremum_closed (\A. Op A)" by (smt IC1_dual ID1b Int_der_def2 PD1 sup_char diff_def) -lemma Int_meet_closed: "Der_1a \ \ meet_closed (\A. Op A)" by (metis ID1a Int_der_def MULT_b_def meet_def) -lemma Int_inf_closed: "Der_inf \ \ infimum_closed (\A. Op A)" by (simp add: fp_ID_inf_closed) -lemma Cl_inf_closed: "Der_1b \ \ infimum_closed (\A. Cl A)" by (smt Cl_der_def IC1_dual ID1b PD2 dual_der1 inf_char join_def) -lemma Cl_join_closed: "Der_1a \ \ join_closed (\A. Cl A)" using ADDI_a_def Cl_der_def join_def by fastforce -lemma Cl_sup_closed: "Der_inf \ \ supremum_closed (\A. Cl A)" by (simp add: fp_CD_sup_closed) -lemma Br_inf_closed: "Der_1b \ \ infimum_closed (\A. Br A)" by (smt Br_der_def CI1b IC1_dual PD1 inf_char diff_def) -lemma Fr_inf_closed: "Der_1b \ \ infimum_closed (\A. Fr A)" by (metis (full_types) Br_der_def Br_inf_closed Cl_der_def Cl_inf_closed Fr_der_def join_def diff_def) -lemma Br_Fr_join: "Der_1 \ \ Der_4e \ \ \A B. Br A \ Fr B \ Br(A \<^bold>\ B)" proof - - assume der1: "Der_1 \" and der4: "Der_4e \" - { fix A B - { assume bra: "Br A" and frb: "Fr B" - from bra have "\ A \<^bold>\ \<^bold>\" using Br_boundary by auto - hence 1: "\(\<^bold>\A) \<^bold>\ \<^bold>\" by (metis ICdual bottom_def compl_def dual_def eq_ext' top_def) - from frb have "\(\ B) \<^bold>\ \<^bold>\" by (simp add: Fr_nowhereDense) - hence 2: "\(\<^bold>\(\ B)) \<^bold>\ \<^bold>\" by (metis ICdual bottom_def compl_def dual_def eq_ext' top_def) - from der1 have "\(\<^bold>\A) \<^bold>\ \ B \<^bold>\ \((\<^bold>\A) \<^bold>\ B)" by (simp add: CD1 PD4) - hence "\(\<^bold>\A) \<^bold>\ \ B \<^bold>\ \(\<^bold>\(A \<^bold>\ B))" unfolding conn by simp - hence "\<^bold>\ \<^bold>\ \ B \<^bold>\ \(\<^bold>\(A \<^bold>\ B))" using 1 unfolding conn by simp - hence 3: "\<^bold>\(\ B) \<^bold>\ \(\<^bold>\(A \<^bold>\ B))" unfolding conn by simp - from der1 der4 have 4: "let M=\<^bold>\(\ B); N=\<^bold>\(A \<^bold>\ B) in M \<^bold>\ \ N \ \ M \<^bold>\ \ N" by (smt CD1b Cl_Closed PC1 PD1) - from 3 4 have "\(\<^bold>\(\ B)) \<^bold>\ \(\<^bold>\(A \<^bold>\ B))" by simp - hence "\<^bold>\ \<^bold>\ \(\<^bold>\(A \<^bold>\ B))" using 2 unfolding top_def by simp - hence "\<^bold>\ \<^bold>\ \(A \<^bold>\ B)" using ICdual dual_def eq_ext' conn by metis - hence "Br (A \<^bold>\ B)" using Br_boundary by simp - } hence "Br A \ Fr B \ Br (A \<^bold>\ B)" by simp - } hence "\A B. Br A \ Fr B \ Br (A \<^bold>\ B)" by simp - thus ?thesis by simp -qed -lemma Fr_join_closed: "Der_1 \ \ Der_4e \ \ join_closed (\A. Fr A)" by (simp add: Br_Fr_join Cl_join_closed Fr_ClBr PC1) - - -text\\noindent{Introduces a predicate for indicating that two sets are disjoint and proves some properties.}\ -abbreviation "Disj A B \ A \<^bold>\ B \<^bold>\ \<^bold>\" - -lemma Disj_comm: "\A B. Disj A B \ Disj B A" unfolding conn by fastforce -lemma Disj_IF: "\A. Disj (\ A) (\ A)" by (simp add: Cl_der_def Fr_der_def2 dual_def dual_der2 conn) -lemma Disj_B: "\A. Disj (\ A) (\(\<^bold>\A))" by (simp add: Br_der_def2 conn) -lemma Disj_I: "\A. Disj (\ A) (\<^bold>\A)" by (simp add: Int_der_def conn) -lemma Disj_BCI: "\A. Disj (\(\ A)) (\(\<^bold>\A))" by (simp add: Br_der_def2 dual_def dual_der1 conn) -lemma Disj_CBI: "Der_1b \ \ Der_4e \ \ \A. Disj (\(\(\<^bold>\A))) (\(\<^bold>\A))" by (smt Br_der_def2 Der_4e_def Cl_der_def Int_der_def2 PD3 conn) - -text\\noindent{Introduce a predicate for indicating that two sets are separated and proves some properties.}\ -definition "Sep A B \ Disj (\ A) B \ Disj (\ B) A" - -lemma Sep_comm: "\A B. Sep A B \ Sep B A" by (simp add: Sep_def) -lemma Sep_disj: "\A B. Sep A B \ Disj A B" using CD2 EXP_def Sep_def conn by auto -lemma Sep_I: "Der_1 \ \ Der_4e \ \ \A. Sep (\ A) (\ (\<^bold>\A))" unfolding Sep_def by (smt CD2 CD4 IC1 ID1 PC4 PC5 PD8 dual_def dual_der1 dual_der2 conn) - -lemma Sep_sub: "Der_1b \ \ \A B C D. Sep A B \ C \<^bold>\ A \ D \<^bold>\ B \ Sep C D" using MONO_ADDIb PD2 dual_der1 monI unfolding Sep_def conn by metis -lemma Sep_Cl_diff: "Der_1b \ \ \A B. Cl(A) \ Cl(B) \ Sep (A \<^bold>\ B) (B \<^bold>\ A)" unfolding Sep_def using CD1b PD1 bottom_def diff_def meet_def by smt -lemma Sep_Op_diff: "Der_1b \ \ \A B. Op(A) \ Op(B) \ Sep (A \<^bold>\ B) (B \<^bold>\ A)" proof - - assume der1b:"Der_1b \" - { fix A B - from der1b have aux: "let M=\<^bold>\A ; N=\<^bold>\B in (Cl(M) \ Cl(N) \ Sep (M \<^bold>\ N) (N \<^bold>\ M))" using Sep_Cl_diff by simp - { assume "Op(A) \ Op(B)" - hence "Cl(\<^bold>\A) \ Cl(\<^bold>\B)" using der1b ClOpdual by simp - hence "Sep (\<^bold>\A \<^bold>\ \<^bold>\B) (\<^bold>\B \<^bold>\ \<^bold>\A)" using der1b aux unfolding conn by simp - moreover have "(\<^bold>\A \<^bold>\ \<^bold>\B) \<^bold>\ (B \<^bold>\ A)" unfolding conn by auto - moreover have "(\<^bold>\B \<^bold>\ \<^bold>\A) \<^bold>\ (A \<^bold>\ B)" unfolding conn by auto - ultimately have "Sep (B \<^bold>\ A) (A \<^bold>\ B)" unfolding conn by simp - hence "Sep (A \<^bold>\ B) (B \<^bold>\ A)" using Sep_comm by simp - } hence "Op(A) \ Op(B) \ Sep (A \<^bold>\ B) (B \<^bold>\ A)" by (rule impI) - } thus ?thesis by simp -qed -lemma Sep_Cl: "\A B. Cl(A) \ Cl(B) \ Disj A B \ Sep A B" unfolding Sep_def conn by blast -lemma Sep_Op: "Der_1b \ \ \A B. Op(A) \ Op(B) \ Disj A B \ Sep A B" proof - - assume der1b:"Der_1b \" - { fix A B - from der1b have aux: "Op(A) \ Op(B) \ Sep (A \<^bold>\ B) (B \<^bold>\ A)" using Sep_Op_diff by simp - { assume op: "Op(A) \ Op(B)" and disj: "Disj A B" - hence "(A \<^bold>\ B) \<^bold>\ A \ (B \<^bold>\ A) \<^bold>\ B" unfolding conn by blast - hence "Sep A B" using op aux unfolding conn by simp - } hence "Op(A) \ Op(B) \ Disj A B \ Sep A B" by simp - } thus ?thesis by simp -qed -lemma "Der_1a \ \ \A B C. Sep A B \ Sep A C \ Sep A (B \<^bold>\ C)" using ADDI_a_def CD1a unfolding Sep_def conn by metis - - -text\\noindent{Verifies a neighborhood-based definition of interior.}\ -definition "nbhd A p \ \E. E \<^bold>\ A \ Op(E) \ (E p)" -lemma nbhd_def2: "Der_1 \ \ Der_4e \ \ \A p. (nbhd A p) = (\ A p)" unfolding nbhd_def by (smt Int_Open MONO_def PC1 monI pI2 conn) - -lemma I_def'_lr': "\A p. (\ A p) \ (\E. (\ E p) \ E \<^bold>\ A)" by blast -lemma I_def'_rl': "Der_1b \ \ \A p. (\ A p) \ (\E. (\ E p) \ E \<^bold>\ A)" using MONO_def monI by metis -lemma I_def': "Der_1b \ \ \A p. (\ A p) \ (\E. (\ E p) \ E \<^bold>\ A)" using MONO_def monI by metis - - -text\\noindent{Explore the Barcan and converse Barcan formulas.}\ -lemma Barcan_I: "Der_inf \ \ \P. (\<^bold>\x. \(P x)) \<^bold>\ \(\<^bold>\x. P x)" using ID_inf Barcan1 by auto -lemma Barcan_C: "Der_inf \ \ \P. \(\<^bold>\x. P x) \<^bold>\ (\<^bold>\x. \(P x))" using CD_inf Barcan2 by metis -lemma CBarcan_I: "Der_1b \ \ \P. \(\<^bold>\x. P x) \<^bold>\ (\<^bold>\x. \(P x))" by (metis (mono_tags, lifting) MONO_def monI) -lemma CBarcan_C: "Der_1b \ \ \P. (\<^bold>\x. \(P x)) \<^bold>\ \(\<^bold>\x. P x)" by (metis (mono_tags, lifting) MONO_def monC) - -end diff --git a/thys/Topological_Semantics/topo_frontier_algebra.thy b/thys/Topological_Semantics/topo_frontier_algebra.thy deleted file mode 100644 --- a/thys/Topological_Semantics/topo_frontier_algebra.thy +++ /dev/null @@ -1,199 +0,0 @@ -theory topo_frontier_algebra - imports topo_operators_basic -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Frontier Algebra\ - -text\\noindent{The closure of a set A (@{text "\(A)"}) can be seen as the set A augmented by (i) its boundary points, -or (ii) its accumulation/limit points. In this section we explore the first variant by drawing on the notion -of a frontier algebra, defined in an analogous fashion as the well-known closure and interior algebras.}\ - -text\\noindent{Declares a primitive (unconstrained) frontier (aka. boundary) operation and defines others from it.}\ -consts \::"\\\" -abbreviation "\ \ \\<^sub>F \" \\ interior \ -abbreviation "\ \ \\<^sub>F \" \\ closure \ -abbreviation "\ \ \\<^sub>F \" \\ border \ - - -subsection \Basic properties\ - -text\\noindent{Verifies minimal conditions under which operators resulting from conversion functions coincide.}\ -lemma ICdual: "Fr_2 \ \ \ \<^bold>\ \\<^sup>d" by (simp add: Cl_fr_def Fr_2_def Int_fr_def dual_def equal_op_def conn) -lemma ICdual': "Fr_2 \ \ \ \<^bold>\ \\<^sup>d" by (simp add: Cl_fr_def Fr_2_def Int_fr_def dual_def equal_op_def conn) -lemma BI_rel: "\ \<^bold>\ \\<^sub>I \" using Br_fr_def Br_int_def Int_fr_def unfolding equal_op_def conn by auto -lemma IB_rel: "\ \<^bold>\ \\<^sub>B \" using Br_fr_def Int_br_def Int_fr_def unfolding equal_op_def conn by auto -lemma BC_rel: "Fr_2 \ \ \ \<^bold>\ \\<^sub>C \" using BI_BC_rel BI_rel ICdual' eq_ext' by fastforce -lemma CB_rel: "Fr_2 \ \ \ \<^bold>\ \\<^sub>B \" using Br_fr_def Cl_br_def Cl_fr_def Fr_2_def unfolding equal_op_def conn by auto -lemma FI_rel: "Fr_2 \ \ \ \<^bold>\ \\<^sub>I \" by (smt Cl_fr_def Fr_int_def ICdual' Int_fr_def compl_def diff_def equal_op_def join_def meet_def) -lemma FC_rel: "Fr_2 \ \ \ \<^bold>\ \\<^sub>C \" by (metis (mono_tags, lifting) FI_rel Fr_2_def Fr_cl_def Fr_int_def ICdual' dual_def eq_ext' equal_op_def) -lemma FB_rel: "Fr_2 \ \ \ \<^bold>\ \\<^sub>B \" by (smt Br_fr_def CB_rel Cl_br_def FC_rel Fr_br_def Fr_cl_def equal_op_def join_def meet_def) - -text\\noindent{Fixed-point and other operators are interestingly related.}\ -lemma fp1: "\\<^sup>f\<^sup>p \<^bold>\ \\<^sup>c" using Br_fr_def Int_fr_def unfolding equal_op_def conn by auto -lemma fp2: "\\<^sup>f\<^sup>p \<^bold>\ \\<^sup>c" using Br_fr_def Int_fr_def unfolding equal_op_def conn by auto -lemma fp3: "Fr_2 \ \ \\<^sup>f\<^sup>p \<^bold>\ \\<^sup>d" using Br_fr_def Cl_fr_def Fr_2_def dual_def equal_op_def conn by fastforce -lemma fp4: "Fr_2 \ \ (\\<^sup>d)\<^sup>f\<^sup>p \<^bold>\ \" by (smt dimp_def equal_op_def fp3) -lemma fp5: "\\<^sup>f\<^sup>p \<^bold>\ \ \<^bold>\ (\\<^sup>c)" using Br_fr_def Cl_fr_def unfolding equal_op_def conn by auto - -text\\noindent{Different inter-relations (some redundant ones are kept to help the provers).}\ -lemma monI: "Fr_1b \ \ MONO(\)" by (simp add: IF1a MONO_MULTa) -lemma monC: "Fr_6 \ \ MONO(\)" by (simp add: Cl_fr_def Fr_6_def MONO_def conn) -lemma pB1: "\A. \ A \<^bold>\ A \<^bold>\ \ A" using Br_fr_def fp1 unfolding conn equal_op_def by metis -lemma pB2: "\A. \ A \<^bold>\ A \<^bold>\ \ A" by (simp add: Br_fr_def) -lemma pB3: "Fr_2 \ \ \A. \(\<^bold>\A) \<^bold>\ \<^bold>\A \<^bold>\ \ A" by (simp add: Fr_2_def pB2 conn) -lemma pB4: "Fr_2 \ \ \A. \(\<^bold>\A) \<^bold>\ \<^bold>\A \<^bold>\ \ A" using CB_rel Cl_br_def pB3 unfolding conn equal_op_def by metis -lemma pB5: "Fr_1b \ \ Fr_2 \ \ \A. \(\ A) \<^bold>\ (\ A) \<^bold>\ \(\<^bold>\A)" by (smt BC_rel Br_cl_def Cl_fr_def Fr_6_def PF6 equal_op_def conn) -lemma pF1: "\A. \ A \<^bold>\ \ A \<^bold>\ \ A" using Cl_fr_def Int_fr_def conn by auto -lemma pF2: "Fr_2 \ \ \A. \ A \<^bold>\ \ A \<^bold>\ \(\<^bold>\A)" using Cl_fr_def Fr_2_def conn by auto -lemma pF3: "Fr_2 \ \ \A. \ A \<^bold>\ \ A \<^bold>\ \(\<^bold>\A)" using Br_fr_def Fr_2_def conn by auto -lemma pF4: "Fr_1 \ \ Fr_2 \ \ Fr_4(\) \ \A. \(\ A) \<^bold>\ \ A" by (smt IDEMa_def IF2 IF4 Int_fr_def MONO_def PF1 PF6 PI4 diff_def monC pF1) -lemma pF5: "Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ \A. \(\ A) \<^bold>\ \ A" by (metis Br_fr_def CF4 Cl_fr_def IDEM_def PF1 join_def meet_def pB5 pF3) -lemma pA1: "\A. A \<^bold>\ \ A \<^bold>\ \ A" using Br_fr_def Int_fr_def unfolding conn by auto -lemma pA2: "Fr_2 \ \ \A. A \<^bold>\ \ A \<^bold>\ \(\<^bold>\A)" using pB1 dual_def fp3 unfolding equal_op_def conn by smt -lemma pC1: "Fr_2 \ \ \A. \ A \<^bold>\ A \<^bold>\ \(\<^bold>\A)" using Cl_fr_def pB4 conn by auto -lemma pC2: "\A. \ A \<^bold>\ A \<^bold>\ \ A" using Cl_fr_def by auto -lemma pI1: "\A. \ A \<^bold>\ A \<^bold>\ \ A" using pA1 pB1 conn by auto -lemma pI2: "\A. \ A \<^bold>\ A \<^bold>\ \ A" by (simp add: Int_fr_def) - -lemma IC_imp: "Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ \A B. \(A \<^bold>\ B) \<^bold>\ \ A \<^bold>\ \ B" proof - - assume fr1: "Fr_1 \" and fr2: "Fr_2 \" and fr3: "Fr_3 \" - { fix a b - have "(a \<^bold>\ b) \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\a = \<^bold>\" unfolding conn by auto - hence "\((a \<^bold>\ b) \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\a) \<^bold>\ \(\<^bold>\)" by simp - hence "\((a \<^bold>\ b) \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\a) \<^bold>\ \<^bold>\" using fr3 IF3 dNOR_def fr2 by auto - moreover have "let A=(a \<^bold>\ b) \<^bold>\ \<^bold>\b; B=\<^bold>\a in \(A \<^bold>\ B) \<^bold>\ \(A) \<^bold>\ \(B)" using fr1 IF1 PI7 Int_7_def by metis - ultimately have "\((a \<^bold>\ b) \<^bold>\ \<^bold>\b) \<^bold>\ \(\<^bold>\a) \<^bold>\ \<^bold>\" unfolding conn by simp - moreover have "let A=a \<^bold>\ b; B=\<^bold>\b in \(A \<^bold>\ B) \<^bold>\ \(A) \<^bold>\ \(B)" using fr1 IF1 MULT_def by simp - ultimately have "\(a \<^bold>\ b) \<^bold>\ \(\<^bold>\b) \<^bold>\ \(\<^bold>\a) \<^bold>\ \<^bold>\" unfolding conn by simp - moreover have "\A. \(\<^bold>\A) \<^bold>\ \<^bold>\\(A)" using Cl_fr_def Fr_2_def Int_fr_def fr2 unfolding conn by auto - ultimately have "\(a \<^bold>\ b) \<^bold>\ \<^bold>\\(b) \<^bold>\ \<^bold>\\(a) \<^bold>\ \<^bold>\" unfolding conn by simp - hence "\(a \<^bold>\ b) \<^bold>\ \<^bold>\\(b) \<^bold>\ \<^bold>\\(a)" unfolding conn by simp - hence "\(a \<^bold>\ b) \<^bold>\ \ a \<^bold>\ \ b" unfolding conn by metis - } thus ?thesis unfolding conn by simp -qed - -text\\noindent{Defines some fixed-point predicates and prove some properties.}\ -abbreviation openset ("Op") where "Op A \ fp \ A" -abbreviation closedset ("Cl") where "Cl A \ fp \ A" -abbreviation borderset ("Br") where "Br A \ fp \ A" -abbreviation frontierset ("Fr") where "Fr A \ fp \ A" - -lemma Int_Open: "Fr_1a \ \ Fr_2 \ \ Fr_4 \ \ \A. Op(\ A)" using IF4 IDEM_def by blast -lemma Cl_Closed: "Fr_1a \ \ Fr_2 \ \ Fr_4 \ \ \A. Cl(\ A)" using CF4 IDEM_def by blast -lemma Br_Border: "Fr_1b \ \ \A. Br(\ A)" using Br_fr_def Fr_1b_def conn by auto -text\\noindent{In contrast, there is no analogous fixed-point result for frontier:}\ -lemma "\ \ \ \A. Fr(\ A)" nitpick oops (*counterexample even if assuming all frontier conditions*) - -lemma OpCldual: "Fr_2 \ \ \A. Cl A \ Op(\<^bold>\A)" using Cl_fr_def Fr_2_def Int_fr_def conn by auto -lemma ClOpdual: "Fr_2 \ \ \A. Op A \ Cl(\<^bold>\A)" using ICdual dual_def unfolding equal_op_def conn by auto -lemma Fr_ClBr: "\A. Fr(A) = (Cl(A) \ Br(A))" using Br_fr_def Cl_fr_def join_def meet_def by auto -lemma Cl_F: "Fr_4 \ \ \A. Cl(\ A)" using Cl_fr_def Fr_4_def conn by auto - - -subsection \Further properties\ - -text\\noindent{The definitions and theorems below are well known in the literature (e.g. \<^cite>\Kuratowski2\). -Here we uncover the minimal conditions under which they hold (taking frontier operation as primitive).}\ -lemma Cl_Bzero: "Fr_2 \ \ \A. Cl A \ \(\<^bold>\A) \<^bold>\ \<^bold>\" using pA2 pC1 unfolding conn by metis -lemma Op_Bzero: "\A. Op A \ (\ A) \<^bold>\ \<^bold>\" using pB1 pI1 unfolding conn by metis -lemma Br_boundary: "Fr_2 \ \ \A. Br(A) \ \ A \<^bold>\ \<^bold>\" by (metis Br_fr_def Int_fr_def bottom_def diff_def meet_def) -lemma Fr_nowhereDense: "Fr_2 \ \ \A. Fr(A) \ \(\ A) \<^bold>\ \<^bold>\" using Fr_ClBr Br_boundary eq_ext by metis -lemma Cl_FB: "\A. Cl A \ \ A \<^bold>\ \ A" using Br_fr_def Cl_fr_def unfolding conn by auto -lemma Op_FB: "Fr_2 \ \ \A. Op A \ \ A \<^bold>\ \(\<^bold>\A)" using Br_fr_def Fr_2_def Int_fr_def unfolding conn by auto -lemma Clopen_Fzero: "\A. Cl A \ Op A \ \ A \<^bold>\ \<^bold>\" using Cl_fr_def Int_fr_def unfolding conn by auto - -lemma Int_sup_closed: "Fr_1b \ \ supremum_closed (\A. Op A)" by (smt IF1a Int_fr_def MONO_def MONO_MULTa sup_char conn) -lemma Int_meet_closed: "Fr_1a \ \ meet_closed (\A. Op A)" using Fr_1a_def Int_fr_def unfolding conn by metis -lemma Int_inf_closed: "Fr_inf \ \ infimum_closed (\A. Op A)" by (simp add: fp_IF_inf_closed) -lemma Cl_inf_closed: "Fr_6 \ \ infimum_closed (\A. Cl A)" by (smt Cl_fr_def Fr_6_def infimum_def join_def) -lemma Cl_join_closed: "Fr_1a \ \ Fr_2 \ \ join_closed (\A. Cl A)" using ADDI_a_def CF1a CF2 EXP_def unfolding conn by metis -lemma Cl_sup_closed: "Fr_2 \ \ Fr_inf \ \ supremum_closed (\A. Cl A)" by (simp add: fp_CF_sup_closed) -lemma Br_inf_closed: "Fr_1b \ \ infimum_closed (\A. Br A)" by (smt BI_rel Br_int_def IF1a MONO_iMULTa MONO_MULTa Ra_restr_all eq_ext' iMULT_a_def inf_char diff_def) -lemma Fr_inf_closed: "Fr_1b \ \ Fr_2 \ \ infimum_closed (\A. Fr A)" by (metis (full_types) Br_fr_def Br_inf_closed PF6 Cl_fr_def Cl_inf_closed meet_def join_def) -lemma Br_Fr_join: "Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ \A B. Br A \ Fr B \ Br(A \<^bold>\ B)" proof - - assume fr1: "Fr_1 \" and fr2: "Fr_2 \" and fr4: "Fr_4 \" - { fix A B - { assume bra: "Br A" and frb: "Fr B" - from bra have "\ A \<^bold>\ \<^bold>\" using Br_boundary fr2 by auto - hence 1: "\(\<^bold>\A) \<^bold>\ \<^bold>\" by (metis ICdual bottom_def compl_def dual_def eq_ext' fr2 top_def) - from frb have "\(\ B) \<^bold>\ \<^bold>\" by (simp add: Fr_nowhereDense fr2) - hence 2: "\(\<^bold>\(\ B)) \<^bold>\ \<^bold>\" by (metis ICdual bottom_def compl_def dual_def eq_ext' fr2 top_def) - from fr1 fr2 have "\(\<^bold>\A) \<^bold>\ \ B \<^bold>\ \((\<^bold>\A) \<^bold>\ B)" using CF1 Cl_6_def PC6 by metis - hence "\(\<^bold>\A) \<^bold>\ \ B \<^bold>\ \(\<^bold>\(A \<^bold>\ B))" unfolding conn by simp - hence "\<^bold>\ \<^bold>\ \ B \<^bold>\ \(\<^bold>\(A \<^bold>\ B))" using 1 unfolding conn by simp - hence 3: "\<^bold>\(\ B) \<^bold>\ \(\<^bold>\(A \<^bold>\ B))" unfolding conn by simp - from fr1 fr2 fr4 have 4: "let M=\<^bold>\(\ B); N=\<^bold>\(A \<^bold>\ B) in M \<^bold>\ \ N \ \ M \<^bold>\ \ N" using PC9 CF4 Cl_9_def PF1 CF1b by fastforce - from 3 4 have "\(\<^bold>\(\ B)) \<^bold>\ \(\<^bold>\(A \<^bold>\ B))" by simp - hence "\<^bold>\ \<^bold>\ \(\<^bold>\(A \<^bold>\ B))" using 2 unfolding top_def by simp - hence "\<^bold>\ \<^bold>\ \(A \<^bold>\ B)" using fr2 ICdual dual_def eq_ext' conn by metis - hence "Br (A \<^bold>\ B)" using fr2 Br_boundary by simp - } hence "Br A \ Fr B \ Br (A \<^bold>\ B)" by simp - } hence "\A B. Br A \ Fr B \ Br (A \<^bold>\ B)" by simp - thus ?thesis by simp -qed -lemma Fr_join_closed: "Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ join_closed (\A. Fr A)" by (smt Br_Fr_join ADDI_a_def CF1a Cl_fr_def PF1 diff_def join_def meet_def pB2 pF1) - - -text\\noindent{Introduces a predicate for indicating that two sets are disjoint and proves some properties.}\ -abbreviation "Disj A B \ A \<^bold>\ B \<^bold>\ \<^bold>\" - -lemma Disj_comm: "\A B. Disj A B \ Disj B A" unfolding conn by fastforce -lemma Disj_IF: "\A. Disj (\ A) (\ A)" by (simp add: Int_fr_def conn) -lemma Disj_B: "\A. Disj (\ A) (\(\<^bold>\A))" using Br_fr_def unfolding conn by auto -lemma Disj_I: "\A. Disj (\ A) (\<^bold>\A)" by (simp add: Int_fr_def conn) -lemma Disj_BCI: "Fr_2 \ \ \A. Disj (\(\ A)) (\(\<^bold>\A))" using Br_fr_def Cl_fr_def Fr_2_def Int_fr_def conn by metis -lemma Disj_CBI: "Fr_6 \ \ Fr_4 \ \ \A. Disj (\(\(\<^bold>\A))) (\(\<^bold>\A))" by (metis Br_fr_def Cl_F IB_rel Int_br_def MONO_MULTa MULT_a_def eq_ext' monC conn) - -text\\noindent{Introduces a predicate for indicating that two sets are separated and proves some properties.}\ -definition "Sep A B \ Disj (\ A) B \ Disj (\ B) A" - -lemma Sep_comm: "\A B. Sep A B \ Sep B A" by (simp add: Sep_def) -lemma Sep_disj: "\A B. Sep A B \ Disj A B" using Cl_fr_def Sep_def conn by fastforce -lemma Sep_I: "Fr_1(\) \ Fr_2(\) \ Fr_4(\) \ \A. Sep (\ A) (\ (\<^bold>\A))" using Cl_fr_def pF4 Fr_2_def Int_fr_def unfolding Sep_def conn by metis -lemma Sep_sub: "Fr_6 \ \ \A B C D. Sep A B \ C \<^bold>\ A \ D \<^bold>\ B \ Sep C D" using MONO_def monC unfolding Sep_def conn by smt -lemma Sep_Cl_diff: "Fr_6 \ \ \A B. Cl(A) \ Cl(B) \ Sep (A \<^bold>\ B) (B \<^bold>\ A)" using Fr_6_def pC2 unfolding Sep_def conn by smt -lemma Sep_Op_diff: "Fr_1b \ \ Fr_2 \ \ \A B. Op(A) \ Op(B) \ Sep (A \<^bold>\ B) (B \<^bold>\ A)" proof - - assume fr1b:"Fr_1b \" and fr2:"Fr_2 \" - { fix A B - from fr1b fr2 have aux: "let M=\<^bold>\A ; N=\<^bold>\B in (Cl(M) \ Cl(N) \ Sep (M \<^bold>\ N) (N \<^bold>\ M))" using PF6 Sep_Cl_diff by simp - { assume "Op(A) \ Op(B)" - hence "Cl(\<^bold>\A) \ Cl(\<^bold>\B)" using fr2 ClOpdual by simp - hence "Sep (\<^bold>\A \<^bold>\ \<^bold>\B) (\<^bold>\B \<^bold>\ \<^bold>\A)" using fr1b fr2 aux unfolding conn by simp - moreover have "(\<^bold>\A \<^bold>\ \<^bold>\B) \<^bold>\ (B \<^bold>\ A)" unfolding conn by auto - moreover have "(\<^bold>\B \<^bold>\ \<^bold>\A) \<^bold>\ (A \<^bold>\ B)" unfolding conn by auto - ultimately have "Sep (B \<^bold>\ A) (A \<^bold>\ B)" unfolding conn by simp - hence "Sep (A \<^bold>\ B) (B \<^bold>\ A)" using Sep_comm by simp - } hence "Op(A) \ Op(B) \ Sep (A \<^bold>\ B) (B \<^bold>\ A)" by (rule impI) - } thus ?thesis by simp -qed -lemma Sep_Cl: "\A B. Cl(A) \ Cl(B) \ Disj A B \ Sep A B" unfolding Sep_def conn by blast -lemma Sep_Op: "Fr_1b \ \ Fr_2 \ \ \A B. Op(A) \ Op(B) \ Disj A B \ Sep A B" proof - - assume fr1b:"Fr_1b \" and fr2:"Fr_2 \" - { fix A B - from fr1b fr2 have aux: "Op(A) \ Op(B) \ Sep (A \<^bold>\ B) (B \<^bold>\ A)" using Sep_Op_diff by simp - { assume op: "Op(A) \ Op(B)" and disj: "Disj A B" - hence "(A \<^bold>\ B) \<^bold>\ A \ (B \<^bold>\ A) \<^bold>\ B" unfolding conn by blast - hence "Sep A B" using op aux unfolding conn by simp - } hence "Op(A) \ Op(B) \ Disj A B \ Sep A B" by simp - } thus ?thesis by simp -qed -lemma "Fr_1a \ \ Fr_2 \ \ \A B C. Sep A B \ Sep A C \ Sep A (B \<^bold>\ C)" using CF1a ADDI_a_def unfolding Sep_def conn by metis - - -text\\noindent{Verifies a neighborhood-based definition of closure.}\ -definition "nbhd A p \ \E. E \<^bold>\ A \ Op(E) \ (E p)" -lemma nbhd_def2: "Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ \A p. (nbhd A p) = (\ A p)" using pF4 MONO_def PF1 monI pI2 unfolding nbhd_def conn by smt - -lemma C_def_lr: "Fr_1b \ \ Fr_2 \ \ Fr_4 \ \ \A p. (\ A p) \ (\E. (nbhd E p) \ \(Disj E A))" using Cl_fr_def Fr_2_def Fr_6_def PF6 pF1 unfolding nbhd_def conn by smt -lemma C_def_rl: "Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ \A p. (\ A p) \ (\E. (nbhd E p) \ \(Disj E A))" using Cl_fr_def pF5 pA1 pB4 unfolding nbhd_def conn by smt -lemma C_def: "Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ \A p. (\ A p) \ (\E. (nbhd E p) \ \(Disj E A))" using C_def_lr C_def_rl PF1 by blast - - -text\\noindent{Explore the Barcan and converse Barcan formulas.}\ -lemma Barcan_I: "Fr_inf \ \ \P. (\<^bold>\x. \(P x)) \<^bold>\ \(\<^bold>\x. P x)" using IF_inf Barcan1 by auto -lemma Barcan_C: "Fr_2 \ \ Fr_inf \ \ \P. \(\<^bold>\x. P x) \<^bold>\ (\<^bold>\x. \(P x))" using Fr_2_def CF_inf Barcan2 by metis -lemma CBarcan_I: "Fr_1b \ \ \P. \(\<^bold>\x. P x) \<^bold>\ (\<^bold>\x. \(P x))" by (metis (mono_tags, lifting) MONO_def monI) -lemma CBarcan_C: "Fr_6 \ \ \P. (\<^bold>\x. \(P x)) \<^bold>\ \(\<^bold>\x. P x)" by (metis (mono_tags, lifting) MONO_def monC) - -end diff --git a/thys/Topological_Semantics/topo_interior_algebra.thy b/thys/Topological_Semantics/topo_interior_algebra.thy deleted file mode 100644 --- a/thys/Topological_Semantics/topo_interior_algebra.thy +++ /dev/null @@ -1,53 +0,0 @@ -theory topo_interior_algebra - imports topo_operators_basic -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Interior algebra\ -text\\noindent{We define a topological Boolean algebra taking the interior operator as primitive and verify some properties.}\ - -text\\noindent{Declares a primitive (unconstrained) interior operation and defines others from it.}\ -consts \::"\\\" -abbreviation "\ \ \\<^sup>d" \\ closure \ -abbreviation "\ \ \\<^sub>I \" \\ border \ -abbreviation "\ \ \\<^sub>I \" \\ frontier \ - - -subsection \Basic properties\ - -text\\noindent{Verifies minimal conditions under which operators resulting from conversion functions coincide.}\ -lemma ICdual: "\ \<^bold>\ \\<^sup>d" using dual_symm equal_op_def by auto -lemma IB_rel: "Int_2 \ \ \ \<^bold>\ \\<^sub>B \" by (smt Br_int_def Int_br_def dEXP_def diff_def equal_op_def) -lemma IF_rel: "Int_2 \ \ \ \<^bold>\ \\<^sub>F \" using EXP_def EXP_dual2 Fr_int_def IB_rel Int_br_def Int_fr_def compl_def diff_def equal_op_def meet_def by fastforce -lemma CB_rel: "Int_2 \ \ \ \<^bold>\ \\<^sub>B \" using Br_int_def Cl_br_def EXP_def EXP_dual2 compl_def diff_def dual_def equal_op_def join_def by fastforce -lemma CF_rel: "Int_2 \ \ \ \<^bold>\ \\<^sub>F \" by (smt Cl_fr_def EXP_def EXP_dual2 Fr_int_def compl_def dEXP_def equal_op_def join_def meet_def) -lemma BC_rel: "\ \<^bold>\ \\<^sub>C \" by (simp add: BI_BC_rel equal_op_def) -lemma BF_rel: "Int_2 \ \ \ \<^bold>\ \\<^sub>F \" by (smt Br_fr_def Br_int_def IF_rel Int_fr_def diff_def equal_op_def meet_def) -lemma FC_rel: "\ \<^bold>\ \\<^sub>C \" using Fr_cl_def Fr_int_def compl_def dual_def equal_op_def meet_def by fastforce -lemma FB_rel: "Int_2 \ \ \ \<^bold>\ \\<^sub>B \" unfolding Fr_br_def by (smt BC_rel Br_cl_def CB_rel Cl_br_def FC_rel Fr_cl_def equal_op_def join_def meet_def) - -text\\noindent{Fixed-point and other operators are interestingly related.}\ -lemma fp1: "Int_2 \ \ \\<^sup>f\<^sup>p \<^bold>\ \\<^sup>c" by (smt Br_int_def IB_rel Int_br_def compl_def diff_def dimp_def equal_op_def) -lemma fp2: "Int_2 \ \ \\<^sup>f\<^sup>p \<^bold>\ \\<^sup>c" using fp1 unfolding compl_def dimp_def equal_op_def by smt -lemma fp3: "Int_2 \ \ \\<^sup>f\<^sup>p \<^bold>\ \\<^sup>d" by (metis (no_types) dual_comp eq_ext' fp2 ofp_c ofp_d ofp_invol) -lemma fp4: "Int_2 \ \ (\\<^sup>d)\<^sup>f\<^sup>p \<^bold>\ \" by (smt dimp_def equal_op_def fp3) -lemma fp5: "Int_2 \ \ \\<^sup>f\<^sup>p \<^bold>\ \ \<^bold>\ (\\<^sup>c)" by (smt BC_rel Br_cl_def CF_rel Cl_fr_def FI2 Fr_2_def compl_def dimp_def eq_ext' equal_op_def join_def meet_def) - -text\\noindent{Define some fixed-point predicates and prove some properties.}\ -abbreviation openset ("Op") where "Op A \ fp \ A" -abbreviation closedset ("Cl") where "Cl A \ fp \ A" -abbreviation borderset ("Br") where "Br A \ fp \ A" -abbreviation frontierset ("Fr") where "Fr A \ fp \ A" - -lemma Int_Open: "Int_4 \ \ \A. Op(\ A)" by (simp add: IDEM_def) -lemma Cl_Closed: "Int_4 \ \ \A. Cl(\ A)" using IC4 IDEM_def by blast -lemma Br_Border: "Int_1a \ \ \A. Br(\ A)" by (metis Br_int_def MONO_MULTa MONO_def diff_def) -text\\noindent{In contrast, there is no analogous fixed-point result for frontier:}\ -lemma "\ \ \ \A. Fr(\ A)" nitpick oops (*counterexample even if assuming all interior conditions*) - -lemma OpCldual: "\A. Cl A \ Op(\<^bold>\A)" by (simp add: fp_d) -lemma ClOpdual: "\A. Op A \ Cl(\<^bold>\A)" by (simp add: compl_def dual_def) -lemma Fr_ClBr: "Int_2 \ \ \A. Fr(A) = (Cl(A) \ Br(A))" using BF_rel Br_fr_def CF_rel Cl_fr_def eq_ext' join_def meet_def by fastforce -lemma Cl_F: "Int_1a \ \ Int_2 \ \ Int_4 \ \ \A. Cl(\ A)" by (metis CF_rel Cl_fr_def FI4 Fr_4_def eq_ext' join_def) - -end diff --git a/thys/Topological_Semantics/topo_negation_conditions.thy b/thys/Topological_Semantics/topo_negation_conditions.thy deleted file mode 100644 --- a/thys/Topological_Semantics/topo_negation_conditions.thy +++ /dev/null @@ -1,267 +0,0 @@ -theory topo_negation_conditions - imports topo_frontier_algebra sse_operation_negative_quantification -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Frontier-based negation - Semantic conditions\ - -text\\noindent{We define a paracomplete and a paraconsistent negation employing the interior and closure operation resp. -We take the frontier operator as primitive and explore which semantic conditions are minimally required -to obtain some relevant properties of negation.}\ - -definition neg_I::"\\\" ("\<^bold>\\<^sup>I") where "\<^bold>\\<^sup>I A \ \(\<^bold>\A)" -definition neg_C::"\\\" ("\<^bold>\\<^sup>C") where "\<^bold>\\<^sup>C A \ \(\<^bold>\A)" -declare neg_I_def[conn] neg_C_def[conn] - -text\\noindent{(We rename the meta-logical HOL negation to avoid terminological confusion)}\ -abbreviation cneg::"bool\bool" ("\_" [40]40) where "\\ \ \\" - -subsection \'Explosion' (ECQ), non-contradiction (LNC) and excluded middle (TND)\ - -text\\noindent{TND}\ -lemma "\ \ \ TNDm \<^bold>\\<^sup>I" nitpick oops (*minimally weak TND not valid*) -lemma TND_C: "TND \<^bold>\\<^sup>C" by (simp add: pC2 Defs conn) (*TND valid in general*) - -text\\noindent{ECQ}\ -lemma ECQ_I: "ECQ \<^bold>\\<^sup>I" by (simp add: pI2 Defs conn) (*ECQ valid in general*) -lemma "\ \ \ ECQm \<^bold>\\<^sup>C" nitpick oops (*minimally weak ECQ not valid*) - -text\\noindent{LNC}\ -lemma "LNC \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma LNC_I: "Fr_2 \ \ Fr_3 \ \ LNC \<^bold>\\<^sup>I" using ECQ_I ECQ_def IF3 LNC_def dNOR_def unfolding conn by auto -lemma "LNC \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma LNC_C: "Fr_6 \ \ LNC \<^bold>\\<^sup>C" unfolding Defs by (smt Cl_fr_def MONO_def compl_def join_def meet_def monC neg_C_def top_def) - -text\\noindent{Relations between LNC and different ECQ variants (only relevant for paraconsistent negation).}\ -lemma "ECQ \<^bold>\\<^sup>C \ LNC \<^bold>\\<^sup>C" by (simp add: pC2 Defs conn) -lemma ECQw_LNC: "ECQw \<^bold>\\<^sup>C \ LNC \<^bold>\\<^sup>C" by (smt ECQw_def LNC_def pC2 conn) -lemma ECQm_LNC: "Fr_1 \ \ Fr_2 \ \ ECQm \<^bold>\\<^sup>C \ LNC \<^bold>\\<^sup>C" using Cl_fr_def Fr_1_def pF2 unfolding Defs conn by metis -lemma "\ \ \ LNC \<^bold>\\<^sup>C \ ECQm \<^bold>\\<^sup>C" nitpick oops (*countermodel*) - -text\\noindent{Below we show that enforcing all conditions on the frontier operator still leaves room -for both boldly paraconsistent and paracomplete models. We use Nitpick to generate a non-trivial -model (here a set algebra with 8 elements).}\ -lemma "\ \ \ \ECQm \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops (*boldly paraconsistent model found*) -lemma "\ \ \ \TNDm \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops (*boldly paracomplete model found*) - - -subsection \Modus tollens (MT)\ - -text\\noindent{MT-I}\ -lemma MT0_I: "Fr_1b \ \ MT0 \<^bold>\\<^sup>I" unfolding Defs by (smt MONO_def compl_def monI neg_I_def top_def) -lemma MT1_I: "Fr_1b \ \ Fr_2 \ \ Fr_3 \ \ MT1 \<^bold>\\<^sup>I" unfolding Defs by (metis MONO_def monI IF3 Int_fr_def compl_def dNOR_def diff_def neg_I_def top_def) -lemma "\ \ \ MT2 \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ MT2 \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TNDm \<^bold>\\<^sup>I \ Fr_1a \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ MT2 \<^bold>\\<^sup>I" nitpick[satisfy] oops -lemma "\TNDm \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ MT2 \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TNDm \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ MT2 \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\ \ \ MT3 \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TNDm \<^bold>\\<^sup>I \ Fr_1a \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ MT3 \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TNDm \<^bold>\\<^sup>I \ MT0 \<^bold>\\<^sup>I \ MT1 \<^bold>\\<^sup>I \ MT2 \<^bold>\\<^sup>I \ MT3 \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -text\\noindent{MT-C}\ -lemma "Fr_2 \ \ MT0 \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma MT0_C: "Fr_6 \ \ MT0 \<^bold>\\<^sup>C" unfolding Defs by (smt ICdual MONO_def compl_def monC neg_C_def top_def) -lemma MT1_C: "Fr_6 \ \ MT1 \<^bold>\\<^sup>C" unfolding Defs by (smt Cl_fr_def Fr_6_def conn) -lemma "\ \ \ MT2 \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQm \<^bold>\\<^sup>C \ \ \ \ MT2 \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops (*model found*) -lemma MT3_C: "Fr_1b \ \ Fr_2 \ \ Fr_3 \ \ MT3 \<^bold>\\<^sup>C" unfolding Defs using MONO_def monI by (metis ClOpdual IF3 compl_def dNOR_def diff_def neg_C_def pA2 top_def) -lemma "\ECQm \<^bold>\\<^sup>C \ MT0 \<^bold>\\<^sup>C \ MT1 \<^bold>\\<^sup>C \ MT2 \<^bold>\\<^sup>C \ MT3 \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops - - -subsection \Contraposition rules (CoP)\ - -text\\noindent{CoPw-I}\ -lemma "CoPw \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma CoPw_I: "Fr_1b \ \ CoPw \<^bold>\\<^sup>I" unfolding Defs conn by (metis (no_types, lifting) MONO_def monI) -text\\noindent{CoPw-C}\ -lemma "CoPw \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma CoPw_C: "Fr_6 \ \ CoPw \<^bold>\\<^sup>C" by (smt CoPw_def MONO_def monC conn) - -text\\noindent{We can indeed prove that XCoP is entailed by CoP1 (CoP2) in the particular case of a closure- (interior-)based negation.}\ -lemma CoP1_XCoP: "CoP1 \<^bold>\\<^sup>C \ XCoP \<^bold>\\<^sup>C" by (metis XCoP_def2 CoP1_def CoP1_def2 DM2_CoPw DM2_def ECQw_def TND_C TND_def TNDw_def top_def) -lemma CoP2_XCoP: "CoP2 \<^bold>\\<^sup>I \ XCoP \<^bold>\\<^sup>I" by (smt XCoP_def2 CoP2_DM3 CoP2_def2 CoPw_def DM3_def DNE_def ECQ_I ECQ_def ECQw_def TNDw_def bottom_def join_def) - -text\\noindent{CoP1-I}\ -lemma "\ \ \ CoP1 \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TNDm \<^bold>\\<^sup>I \ \ \ \ CoP1 \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -text\\noindent{CoP1-C}\ -lemma "\ \ \ CoP1 \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQ \<^bold>\\<^sup>C \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ CoP1 \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops -lemma "CoP1 \<^bold>\\<^sup>C \ ECQm \<^bold>\\<^sup>C" using XCoP_def2 CoP1_XCoP ECQm_def ECQw_def by blast - -text\\noindent{CoP2-I}\ -lemma "\ \ \ CoP2 \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ CoP2 \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ CoP2 \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "CoP2 \<^bold>\\<^sup>I \ TNDm \<^bold>\\<^sup>I" using XCoP_def2 CoP2_XCoP TNDm_def TNDw_def by auto -text\\noindent{CoP2-C}\ -lemma "\ \ \ CoP2 \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQm \<^bold>\\<^sup>C \ \ \ \ CoP2 \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops - -text\\noindent{CoP3-I}\ -lemma "\ \ \ CoP3 \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TND \<^bold>\\<^sup>I \ CoP3 \<^bold>\\<^sup>I" (*nitpick[satisfy]*) oops (*cannot find (finite) models*) -text\\noindent{CoP3-C}\ -lemma "\ \ \ CoP3 \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQ \<^bold>\\<^sup>C \ CoP3 \<^bold>\\<^sup>C" (*nitpick[satisfy]*) oops (*cannot find (finite) models*) - -text\\noindent{XCoP-I}\ -lemma "\ \ \ XCoP \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ XCoP \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ XCoP \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "XCoP \<^bold>\\<^sup>I \ TNDm \<^bold>\\<^sup>I" by (simp add: XCoP_def2 TNDm_def TNDw_def) -text\\noindent{XCoP-C}\ -lemma "\ \ \ XCoP \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQ \<^bold>\\<^sup>C \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ XCoP \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops -lemma "XCoP \<^bold>\\<^sup>C \ ECQm \<^bold>\\<^sup>C" by (simp add: XCoP_def2 ECQm_def ECQw_def) - - -subsection \Normality (negative) and its dual (nNor/nDNor)\ - -text\\noindent{nNor-I}\ -lemma "nNor \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma nNor_I: "Fr_2 \ \ Fr_3 \ \ nNor \<^bold>\\<^sup>I" using IF3 dNOR_def unfolding Defs conn by auto -text\\noindent{nNor-C}\ -lemma nNor_C: "nNor \<^bold>\\<^sup>C" unfolding Cl_fr_def Defs conn by simp - -text\\noindent{nDNor-I}\ -lemma nDNor_I: "nDNor \<^bold>\\<^sup>I" unfolding Int_fr_def Defs conn by simp -text\\noindent{nDNor-C}\ -lemma "nDNor \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma nDNor_C: "Fr_3 \ \ nDNor \<^bold>\\<^sup>C" using pC2 NOR_def unfolding Defs conn by simp - - -subsection \Double negation introduction/elimination (DNI/DNE)\ - -text\\noindent{DNI-I}\ -lemma "\ \ \ DNI \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TNDm \<^bold>\\<^sup>I \ \ \ \ DNI \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -text\\noindent{DNI-C}\ -lemma "\ \ \ DNI \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQ \<^bold>\\<^sup>C \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ DNI \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops -lemma "\ECQm \<^bold>\\<^sup>C \ Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ DNI \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops -lemma "\ECQm \<^bold>\\<^sup>C \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ DNI \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops - -text\\noindent{DNE-I}\ -lemma "\ \ \ DNE \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ DNE \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ DNE \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TNDm \<^bold>\\<^sup>I \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ DNE \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TND \<^bold>\\<^sup>I \ DNE \<^bold>\\<^sup>I \ DNI \<^bold>\\<^sup>I" (*nitpick[satisfy]*) oops (*cannot find (finite) models*) -text\\noindent{DNE-C}\ -lemma "\ \ \ DNE \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQm \<^bold>\\<^sup>C \ \ \ \ DNE \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops - -lemma "\ECQ \<^bold>\\<^sup>C \ DNE \<^bold>\\<^sup>C \ DNI \<^bold>\\<^sup>C" (*nitpick[satisfy]*) oops (*cannot find (finite) models*) - -text\\noindent{rDNI-I}\ -lemma "Fr_2 \ \ Fr_3 \ \ rDNI \<^bold>\\<^sup>I" using nNor_I nDNor_I nDNor_rDNI by simp -text\\noindent{rDNI-C}\ -lemma "Fr_3 \ \ rDNI \<^bold>\\<^sup>C" using nNor_C nDNor_C nDNor_rDNI by simp - -text\\noindent{rDNE-I}\ -lemma "\ \ \ rDNE \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TNDm \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ rDNE \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TNDm \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ rDNE \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TNDm \<^bold>\\<^sup>I \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ rDNE \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -text\\noindent{rDNE-C}\ -lemma "\ \ \ rDNE \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQm \<^bold>\\<^sup>C \ \ \ \ rDNE \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops - -lemma "\ECQm \<^bold>\\<^sup>C \ rDNE \<^bold>\\<^sup>C \ rDNI \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops - - -subsection \De Morgan laws\ - -text\\noindent{DM1/2 (see CoPw)}\ - -text\\noindent{DM3-I}\ -lemma "\ \ \ DM3 \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ DM3 \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ DM3 \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TNDm \<^bold>\\<^sup>I \ DM3 \<^bold>\\<^sup>I" (*nitpick[satisfy]*) oops (*cannot find (finite) models*) -text\\noindent{DM3-C}\ -lemma "DM3 \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma DM3_C: "Fr_1a \ \ Fr_2 \ \ DM3 \<^bold>\\<^sup>C" using DM3_def Fr_1a_def pA2 pF2 unfolding conn by smt -lemma "\ECQm \<^bold>\\<^sup>C \ \ \ \ DM3 \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops - -text\\noindent{DM4-I}\ -lemma "DM4 \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma DM4_I: "Fr_1a \ \ DM4 \<^bold>\\<^sup>I" using ADDI_a_def Cl_fr_def DM4_def IC1b IF1b dual_def unfolding conn by smt -lemma "\TNDm \<^bold>\\<^sup>I \ \ \ \ DM4 \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -text\\noindent{DM4-C}\ -lemma "\ \ \ DM4 \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQ \<^bold>\\<^sup>C \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ DM4 \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops -lemma "\ECQm \<^bold>\\<^sup>C \ DM4 \<^bold>\\<^sup>C" (*nitpick[satisfy]*) oops (*cannot find (finite) models*) - -text\\noindent{iDM1/2 (see CoPw)}\ - -text\\noindent{iDM3-I}\ -lemma "\ \ \ Fr_inf \ \ iDM3 \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ iDM3 \<^bold>\\<^sup>I" nitpick[satisfy] oops -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ iDM3 \<^bold>\\<^sup>I" nitpick[satisfy] oops -lemma "\TNDm \<^bold>\\<^sup>I \ iDM3 \<^bold>\\<^sup>I" (*nitpick[satisfy]*) oops (*cannot find (finite) models*) -text\\noindent{iDM3-C}\ -lemma "iDM3 \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma iDM3_C: "Fr_2 \ \ Fr_inf \ \ iDM3 \<^bold>\\<^sup>C" unfolding Defs by (metis (full_types) CF_inf Ra_restr_ex dom_compl_def iADDI_a_def iDM_a neg_C_def) -text\\noindent{iDM4-I}\ -lemma "iDM4 \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma iDM4_I: "Fr_inf \ \ iDM4 \<^bold>\\<^sup>I" unfolding Defs by (metis IF_inf Ra_restr_all dom_compl_def iDM_b iMULT_b_def neg_I_def) -text\\noindent{iDM4-C}\ -lemma "\ \ \ Fr_inf \ \ iDM4 \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQ \<^bold>\\<^sup>C \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ iDM4 \<^bold>\\<^sup>C" nitpick[satisfy] oops -lemma "\ECQm \<^bold>\\<^sup>C \ iDM4 \<^bold>\\<^sup>C" (*nitpick[satisfy]*) oops (*cannot find (finite) models*) - - -subsection \Local contraposition axioms (lCoP)\ - -text\\noindent{lCoPw-I}\ -lemma "\ \ \ lCoPw(\<^bold>\) \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ lCoPw(\<^bold>\) \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ lCoPw(\<^bold>\) \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "lCoPw(\<^bold>\) \<^bold>\\<^sup>I \ TNDm \<^bold>\\<^sup>I" by (simp add: XCoP_def2 TNDm_def TNDw_def lCoPw_XCoP) -text\\noindent{lCoPw-C}\ -lemma "\ \ \ lCoPw(\<^bold>\) \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQ \<^bold>\\<^sup>C \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ lCoPw(\<^bold>\) \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops -lemma "lCoPw(\<^bold>\) \<^bold>\\<^sup>C \ ECQm \<^bold>\\<^sup>C" by (simp add: XCoP_def2 ECQm_def ECQw_def lCoPw_XCoP) - -text\\noindent{lCoP1-I}\ -lemma "\ \ \ lCoP1(\<^bold>\) \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "lCoP1(\<^bold>\) \<^bold>\\<^sup>I \ TND \<^bold>\\<^sup>I" by (simp add: lCoP1_TND) -text\\noindent{lCoP1-C}\ -lemma "\ \ \ lCoP1(\<^bold>\) \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "\ECQ \<^bold>\\<^sup>C \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ lCoP1(\<^bold>\) \<^bold>\\<^sup>C" nitpick[satisfy,card w=3] oops -lemma "lCoP1(\<^bold>\) \<^bold>\\<^sup>C \ ECQm \<^bold>\\<^sup>C" by (simp add: XCoP_def2 ECQm_def ECQw_def lCoP1_def2 lCoPw_XCoP) - -text\\noindent{lCoP2-I}\ -lemma "\ \ \ lCoP2(\<^bold>\) \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ lCoP2(\<^bold>\) \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "\TND \<^bold>\\<^sup>I \ Fr_1 \ \ Fr_3 \ \ Fr_4 \ \ lCoP2(\<^bold>\) \<^bold>\\<^sup>I" nitpick[satisfy,card w=3] oops -lemma "lCoP2(\<^bold>\) \<^bold>\\<^sup>I \ TNDm \<^bold>\\<^sup>I" by (simp add: XCoP_def2 TNDm_def TNDw_def lCoP2_def2 lCoPw_XCoP) -text\\noindent{lCoP2-C}\ -lemma "\ \ \ lCoP2(\<^bold>\) \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "lCoP2(\<^bold>\) \<^bold>\\<^sup>C \ ECQ \<^bold>\\<^sup>C" by (simp add: lCoP2_ECQ) - -text\\noindent{lCoP3-I}\ -lemma "\ \ \ lCoP3(\<^bold>\) \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "lCoP3(\<^bold>\) \<^bold>\\<^sup>I \ TND \<^bold>\\<^sup>I" unfolding Defs conn by metis -text\\noindent{lCoP3-C}\ -lemma "\ \ \ lCoP3(\<^bold>\) \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "lCoP3(\<^bold>\) \<^bold>\\<^sup>C \ ECQ \<^bold>\\<^sup>C" unfolding Defs conn by metis - - -subsection \Disjunctive syllogism\ - -text\\noindent{DS1-I}\ -lemma "DS1(\<^bold>\) \<^bold>\\<^sup>I" using DS1_def ECQ_I ECQ_def unfolding conn by auto -text\\noindent{DS1-C}\ -lemma "\ \ \ DS1(\<^bold>\) \<^bold>\\<^sup>C" nitpick oops (*countermodel*) -lemma "DS1(\<^bold>\) \<^bold>\\<^sup>C \ ECQ \<^bold>\\<^sup>C" unfolding Defs conn by metis - -text\\noindent{DS2-I}\ -lemma "\ \ \ DS2(\<^bold>\) \<^bold>\\<^sup>I" nitpick oops (*countermodel*) -lemma "DS2(\<^bold>\) \<^bold>\\<^sup>I \ TND \<^bold>\\<^sup>I" by (simp add: Defs conn) -text\\noindent{DS2-C}\ -lemma "DS2(\<^bold>\) \<^bold>\\<^sup>C" using TND_C unfolding Defs conn by auto - -end diff --git a/thys/Topological_Semantics/topo_negation_fixedpoints.thy b/thys/Topological_Semantics/topo_negation_fixedpoints.thy deleted file mode 100644 --- a/thys/Topological_Semantics/topo_negation_fixedpoints.thy +++ /dev/null @@ -1,156 +0,0 @@ -theory topo_negation_fixedpoints - imports topo_frontier_algebra sse_operation_negative_quantification -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Frontier-based negation - Fixed-points\ - -text\\noindent{We define a paracomplete and a paraconsistent negation employing the interior and the closure operation -respectively. We explore the use of fixed-point predicates to recover some relevant properties of negation, -many of which cannot be readily recovered by just adding semantic conditions. -We take the frontier operator as primitive and explore which semantic conditions are minimally required.}\ - -definition neg_I::"\\\" ("\<^bold>\\<^sup>I") where "\<^bold>\\<^sup>I \ \ \(\<^bold>\\)" -definition neg_C::"\\\" ("\<^bold>\\<^sup>C") where "\<^bold>\\<^sup>C \ \ \(\<^bold>\\)" -declare neg_I_def[conn] neg_C_def[conn] - -text\\noindent{Note that all results obtained for fixed-point predicates extend to their associated operators as follows:}\ -lemma "\A. \\<^sup>f\<^sup>p(A) \<^bold>\ \(A) \<^bold>\ \(A) \ \A. (fp \)(A) \ \(A) \<^bold>\ \(A)" unfolding conn by simp -lemma "\A B. \\<^sup>f\<^sup>p(A) \<^bold>\ \\<^sup>f\<^sup>p(B) \<^bold>\ (\ A B) \<^bold>\ (\ A B) \ \A B. (fp \)(A) \ (fp \)(B) \ (\ A B) \<^bold>\ (\ A B)" unfolding conn by simp - -text\\noindent{Recall from previous results that if we have Fr(A) then we also have both Cl(A) and Br(A). -With this understanding we will tacitly assume the corresponding results for Fr(-) below. -Moreover, we obtained countermodels (using Nitpick) for all formulas featuring other combinations (not shown).}\ - - -subsection \'Explosion' (ECQ) and excluded middle (TND)\ - -text\\noindent{TND-I}\ -lemma "Fr_2 \ \ \A. Cl(A) \ TND\<^sup>A \<^bold>\\<^sup>I" by (simp add: OpCldual conn) -text\\noindent{ECQ-C}\ -lemma "Fr_2 \ \ \A. Op(A) \ ECQ\<^sup>A \<^bold>\\<^sup>C" using Cl_fr_def Int_fr_def pF2 unfolding conn by fastforce - - -subsection \Contraposition rules\ - -text\\noindent{CoPw-I}\ -lemma "\A B. Br(\<^bold>\B) \ CoPw\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using Int_fr_def pB1 conn by auto -lemma "Fr_2 \ \ \A B. Cl(A) \ CoPw\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using Int_fr_def OpCldual conn by auto -text\\noindent{CoPw-C}\ -lemma "Fr_2 \ \ \A B. Br(A) \ CoPw\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using pA2 pB2 pF2 unfolding conn by metis -lemma "Fr_2 \ \ \A B. Op(B) \ CoPw\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using ClOpdual Cl_fr_def unfolding conn by auto - -text\\noindent{CoP1-I}\ -lemma "Fr_2 \ \ \A B. Cl(A) \ CoP1\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using Int_fr_def OpCldual conn by auto -lemma "Fr_1b \ \ \A B. Op(B) \ CoP1\<^sup>A\<^sup>B \<^bold>\\<^sup>I" by (smt IF2 dEXP_def MONO_def monI conn) -lemma CoP1_I_rec: "Fr_2 \ \ Fr_3 \ \ \A B. Br (\<^bold>\B) \ CoP1\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using IF3 dNOR_def Br_boundary unfolding conn by auto -text\\noindent{CoP1-C}\ -lemma "Fr_2 \ \ \A B. Op(B) \ CoP1\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using Int_fr_def pC2 pF2 unfolding conn by metis -lemma "Fr_2 \ \ \A B. Br(A) \ CoP1\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using Br_fr_def Cl_fr_def pF2 unfolding conn by fastforce - -text\\noindent{CoP2-I}\ -lemma "Fr_2 \ \ \A B. Cl(A) \ CoP2\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using Int_fr_def OpCldual unfolding conn by auto -lemma "\A B. Br (\<^bold>\B) \ CoP2\<^sup>A\<^sup>B \<^bold>\\<^sup>I" by (simp add: pI1 conn) -text\\noindent{CoP2-C}\ -lemma "Fr_2 \ \ \A B. Op(B) \ CoP2\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using Cl_fr_def Int_fr_def pF2 unfolding conn by fastforce -lemma "Fr_6 \ \ \A B. Cl(A) \ CoP2\<^sup>A\<^sup>B \<^bold>\\<^sup>C" by (smt Cl_fr_def MONO_def monC conn) -lemma "Fr_2 \ \ Fr_3 \ \ \A B. Br(A) \ CoP2\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using CoP1_I_rec Disj_IF pA2 pF2 pF3 unfolding conn by smt - -text\\noindent{CoP3-I}\ -lemma "Fr_2 \ \ \A B. Cl(A) \ CoP3\<^sup>A\<^sup>B \<^bold>\\<^sup>I" by (metis Disj_I ICdual' compl_def dual_def eq_ext' meet_def neg_I_def) -text\\noindent{CoP3-C}\ -lemma "Fr_2 \ \ \A B. Op(B) \ CoP3\<^sup>A\<^sup>B \<^bold>\\<^sup>C" by (metis Disj_I ICdual compl_def dual_def eq_ext' meet_def neg_C_def) - -text\\noindent{XCoP-I}\ -lemma "Fr_2 \ \ \A B. Cl(A) \ XCoP\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using Fr_2_def pA1 pA2 pF1 unfolding conn by metis -lemma "\A B. Br(\<^bold>\B) \ XCoP\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using IB_rel Int_br_def diff_def eq_ext' conn by fastforce -text\\noindent{XCoP-C}\ -lemma "Fr_2 \ \ \A B. Op(B) \ XCoP\<^sup>A\<^sup>B \<^bold>\\<^sup>C" by (metis ClOpdual compl_def diff_def meet_def neg_C_def pA2) -lemma "Fr_2 \ \ \A B. \A B. Br(A) \ XCoP\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using Cl_fr_def compl_def join_def neg_C_def pF3 by fastforce - - -subsection \Double negation introduction/elimination\ - -text\\noindent{DNI-I}\ -lemma "Fr_1b \ \ \A. Op(A) \ DNI\<^sup>A \<^bold>\\<^sup>I" using MONO_def monI pA1 unfolding conn by smt -lemma "Fr_2 \ \ Fr_3 \ \ \A. Br (\<^bold>\A) \ DNI\<^sup>A \<^bold>\\<^sup>I" using CoP1_I_rec by simp -text\\noindent{DNI-C}\ -lemma "Fr_2 \ \ \A. Op(A) \ DNI\<^sup>A \<^bold>\\<^sup>C" using Cl_fr_def Int_fr_def pF2 unfolding conn by fastforce - -text\\noindent{DNE-I}\ -lemma "Fr_2 \ \ \A. Cl(A) \ DNE\<^sup>A \<^bold>\\<^sup>I" by (simp add: Cl_fr_def Fr_2_def Int_fr_def conn) -text\\noindent{DNE-C}\ -lemma "Fr_6 \ \ \A. Cl(A) \ DNE\<^sup>A \<^bold>\\<^sup>C" by (smt MONO_def monC pC2 conn) -lemma "Fr_2 \ \ Fr_3 \ \ \A. Br(A) \ DNE\<^sup>A \<^bold>\\<^sup>C" using CoP1_I_rec Disj_IF pA2 pF2 pF3 unfolding conn by smt - - -subsection \De Morgan laws\ - -text\\noindent{DM1-I}\ -lemma "Fr_1b \ \ \A B. DM1\<^sup>A\<^sup>B \<^bold>\\<^sup>I" by (smt MONO_def monI conn) -lemma "Fr_2 \ \ \A B. Cl(A) \ Cl(B) \ DM1\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using pF2 pI2 conn by fastforce -text\\noindent{DM1-C}\ -lemma "Fr_6 \ \ \A B. DM1\<^sup>A\<^sup>B \<^bold>\\<^sup>C" by (smt MONO_def monC conn) -lemma "Fr_2 \ \ \A B. Br(A) \ Br(B) \ DM1\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using CF2 EXP_def pF2 pF3 unfolding conn by metis - -text\\noindent{DM2-I}\ -lemma "Fr_1b \ \ \A B. DM2\<^sup>A\<^sup>B \<^bold>\\<^sup>I" by (smt MONO_def monI conn) -lemma "\A B. Br(\<^bold>\A) \ Br(\<^bold>\B) \ DM2\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using Int_fr_def pB1 conn by auto -text\\noindent{DM2-C}\ -lemma "Fr_6 \ \ \A B. DM2\<^sup>A\<^sup>B \<^bold>\\<^sup>C" by (smt MONO_def monC conn) -lemma "Fr_2 \ \ \A B. Op(A) \ Op(B) \ DM2\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using CF2 ClOpdual EXP_def unfolding conn by auto - -text\\noindent{DM3-I}\ -lemma "Fr_2 \ \ \A B. Cl(A) \ Cl(B) \ DM3\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using Int_fr_def pF2 unfolding conn by fastforce -text\\noindent{DM3-C}\ -lemma "Fr_1a \ \ Fr_2 \ \ \A B. DM3\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using Cl_fr_def Fr_1a_def pF2 unfolding conn by metis -lemma "Fr_2 \ \ \A B. Br(A) \ Br(B) \ DM3\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using Cl_fr_def pF3 unfolding conn by fastforce - -text\\noindent{DM4-I}\ -lemma "Fr_1a \ \ Fr_2 \ \ \A B. DM4\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using ADDI_a_def Br_fr_def CF1a Int_fr_def pC1 unfolding conn by (metis (full_types)) -lemma "\A B. Br(\<^bold>\A) \ Br(\<^bold>\B) \ DM4\<^sup>A\<^sup>B \<^bold>\\<^sup>I" using Int_fr_def pB1 conn by auto -text\\noindent{DM4-C}\ -lemma "Fr_2 \ \ \A B. Op(A) \ Op(B) \ DM4\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using Cl_fr_def Int_fr_def pF2 unfolding conn by (metis (full_types)) -lemma "Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ \A B. Fr(A) \ Fr(B) \ DM4\<^sup>A\<^sup>B \<^bold>\\<^sup>C" using Cl_fr_def Fr_join_closed Fr_2_def compl_def join_def neg_C_def by auto - - -subsection \Local contraposition axioms\ - -text\\noindent{lCoPw-I}\ -lemma "Fr_2 \ \ \A B. Cl(A) \ lCoPw\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>I" using Int_fr_def OpCldual unfolding conn by auto -lemma "\A B. Br(\<^bold>\B) \ lCoPw\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>I" by (simp add: pI1 conn) -text\\noindent{lCoPw-C}\ -lemma "Fr_2 \ \ \A B. Op(B) \ lCoPw\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>C" using Cl_fr_def Int_fr_def pF2 unfolding conn by fastforce -lemma "Fr_2 \ \ \A B. Br(A) \ lCoPw\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>C" by (simp add: pC1 conn) - -text\\noindent{lCoP1-I}\ -lemma "Fr_2 \ \ \A B. Cl(A) \ lCoP1\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>I" using Int_fr_def OpCldual unfolding conn by auto -text\\noindent{lCoP1-C}\ -lemma "Fr_2 \ \ \A B. Op(B) \ lCoP1\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>C" using Cl_fr_def Int_fr_def pF2 unfolding conn by fastforce -lemma "Fr_2 \ \ \A B. Br(A) \ lCoP1\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>C" by (simp add: pC1 conn) - -text\\noindent{lCoP2-I}\ -lemma "Fr_2 \ \ \A B. Cl(A) \ lCoP2\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>I" using Int_fr_def OpCldual unfolding conn by auto -lemma "\A B. Br(\<^bold>\B) \ lCoP2\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>I" by (simp add: pI1 conn) -text\\noindent{lCoP2-C}\ -lemma "Fr_2 \ \ \A B. Op(B) \ lCoP2\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>C" using Cl_fr_def Int_fr_def pF2 unfolding conn by fastforce - -text\\noindent{lCoP3-I}\ -lemma "Fr_2 \ \ \A B. Cl(A) \ lCoP3\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>I" using Int_fr_def OpCldual unfolding conn by auto -text\\noindent{lCoP3-C}\ -lemma "Fr_2 \ \ \A B. Op(B) \ lCoP3\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>C" using Cl_fr_def Int_fr_def pF2 unfolding conn by fastforce - - -subsection \Disjunctive syllogism\ - -text\\noindent{DS1-I}\ -lemma "\A B. DS1\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>I" by (simp add: Int_fr_def conn) -text\\noindent{DS1-C}\ -lemma "Fr_2 \ \ \A B. Op(A) \ DS1\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>C" using Cl_fr_def Int_fr_def pF2 unfolding conn by fastforce - -text\\noindent{DS2-I}\ -lemma "Fr_2 \ \ \A B. Cl(A) \ DS2\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>I" using OpCldual unfolding conn by auto -text\\noindent{DS2-C}\ -lemma "\A B. DS2\<^sup>A\<^sup>B(\<^bold>\) \<^bold>\\<^sup>C" using Cl_fr_def unfolding conn by auto - -end diff --git a/thys/Topological_Semantics/topo_operators_basic.thy b/thys/Topological_Semantics/topo_operators_basic.thy deleted file mode 100644 --- a/thys/Topological_Semantics/topo_operators_basic.thy +++ /dev/null @@ -1,416 +0,0 @@ -theory topo_operators_basic - imports sse_operation_positive_quantification -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -abbreviation implies_rl::"bool\bool\bool" (infixl "\" 25) where "\ \ \ \ \ \ \" (*for readability*) - -section \Topological operators\ - -text\\noindent{Below we define some conditions on algebraic operations (aka. operators) with type @{text "\\\"}. -Those operations are aimed at extending a Boolean 'algebra of propositions' towards different -generalizations of topological algebras. -We divide this section into two parts. In the first we define and interrelate the topological operators of -interior, closure, border and frontier. In the second we introduce the (more fundamental) notion of -derivative (aka. derived set) and its related notion of (Cantorian) coherence, defining both as operators. -We follow the naming conventions introduced originally by Kuratowski \<^cite>\Kuratowski1\ -(cf. also \<^cite>\Kuratowski2\) and Zarycki \<^cite>\Zarycki1\.}\ - -subsection \Interior and closure\ -text\\noindent{In this section we examine the traditional notion of topological (closure, resp. interior) algebras -in the spirit of McKinsey \& Tarski \<^cite>\AOT\, but drawing primarily from the works of Zarycki -\<^cite>\Zarycki1\ and Kuratowski \<^cite>\Kuratowski1\. -We also explore the less-known notions of border (cf. 'Rand' \<^cite>\Hausdorff\, 'bord' \<^cite>\Zarycki1\) and -frontier (aka. 'boundary'; cf. 'Grenze' \<^cite>\Hausdorff\, 'fronti\`ere' \<^cite>\Zarycki1\ \<^cite>\Kuratowski2\) -as studied by Zarycki \<^cite>\Zarycki1\ and define corresponding operations for them.}\ - -subsubsection \Interior conditions\ - -abbreviation "Int_1 \ \ MULT \" -abbreviation "Int_1a \ \ MULT_a \" -abbreviation "Int_1b \ \ MULT_b \" -abbreviation "Int_2 \ \ dEXP \" -abbreviation "Int_3 \ \ dNOR \" -abbreviation "Int_4 \ \ IDEM \" -abbreviation "Int_4' \ \ IDEMa \" - -abbreviation "Int_5 \ \ NOR \" -definition "Int_6 \ \ \A B. \(A \<^bold>\ B) \<^bold>\ \(A) \<^bold>\ \(B)" -definition "Int_7 \ \ \A B. \(A \<^bold>\ B) \<^bold>\ \(A) \<^bold>\ \(B)" -definition "Int_8 \ \ \A B. \(\ A \<^bold>\ \ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" -definition "Int_9 \ \ \A B. \ A \<^bold>\ B \ \ A \<^bold>\ \ B" - -text\\noindent{@{text "\"} is an interior operator (@{text "\(\)"}) iff it satisfies conditions 1-4 (cf. \<^cite>\Zarycki1\ -and also \<^cite>\Kuratowski2\). This characterization is shown consistent by generating a non-trivial model.}\ -abbreviation "\ \ \ Int_1 \ \ Int_2 \ \ Int_3 \ \ Int_4 \" -lemma "\ \" nitpick[satisfy, card w=3] oops (*model found: characterization is consistent*) - -text\\noindent{We verify some properties which will become useful later (also to improve provers' performance).}\ -lemma PI1: "Int_1 \ = (Int_1a \ \ Int_1b \)" by (metis MULT_def MULT_a_def MULT_b_def) -lemma PI4: "Int_2 \ \ (Int_4 \ = Int_4' \)" unfolding dEXP_def IDEM_def IDEMa_def by blast -lemma PI5: "Int_2 \ \ Int_5 \" unfolding dEXP_def NOR_def conn by blast -lemma PI6: "Int_1a \ \ Int_2 \ \ Int_6 \" by (smt dEXP_def Int_6_def MONO_MULTa MONO_def conn) -lemma PI7: "Int_1 \ \ Int_7 \" proof - - assume int1: "Int_1 \" - { fix a b - have "a \<^bold>\ b = a \<^bold>\ (a \<^bold>\ b)" unfolding conn by blast - hence "\(a \<^bold>\ b) = \(a \<^bold>\ (a \<^bold>\ b))" by simp - moreover from int1 have "\(a \<^bold>\ b) \<^bold>\ \ a \<^bold>\ \ b" by (simp add: MULT_def) - moreover from int1 have "\(a \<^bold>\ (a \<^bold>\ b)) \<^bold>\ \ a \<^bold>\ \ (a \<^bold>\ b)" by (simp add: MULT_def) - ultimately have "\ a \<^bold>\ \ (a \<^bold>\ b) \<^bold>\ \ a \<^bold>\ \ b" by simp - hence "\ a \<^bold>\ \ (a \<^bold>\ b) \<^bold>\ \ a \<^bold>\ (\ a \<^bold>\ \ b)" unfolding conn by blast - hence "\(a \<^bold>\ b) \<^bold>\ \ a \<^bold>\ \ b" unfolding conn by blast - } thus ?thesis by (simp add: Int_7_def) -qed -lemma PI8: "Int_1a \ \ Int_2 \ \ Int_4 \ \ Int_8 \" using ADDI_b_def IDEM_def Int_8_def MONO_ADDIb MONO_MULTa dEXP_def join_def by auto -lemma PI9: "Int_1a \ \ Int_4 \ \ Int_9 \" using IDEM_def Int_9_def MONO_def MONO_MULTa by simp - - -subsubsection \Closure conditions\ - -abbreviation "Cl_1 \ \ ADDI \" -abbreviation "Cl_1a \ \ ADDI_a \" -abbreviation "Cl_1b \ \ ADDI_b \" -abbreviation "Cl_2 \ \ EXP \" -abbreviation "Cl_3 \ \ NOR \" -abbreviation "Cl_4 \ \ IDEM \" -abbreviation "Cl_4' \ \ IDEMb \" - -abbreviation "Cl_5 \ \ dNOR \" -definition "Cl_6 \ \ \A B. (\ A) \<^bold>\ (\ B) \<^bold>\ \ (A \<^bold>\ B)" -definition "Cl_7 \ \ \A B. (\ A) \<^bold>\ (\ B) \<^bold>\ \ (A \<^bold>\ B)" -definition "Cl_8 \ \ \A B. \(\ A \<^bold>\ \ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" -definition "Cl_9 \ \ \A B. A \<^bold>\ \ B \ \ A \<^bold>\ \ B" - -text\\noindent{@{text "\"} is a closure operator (@{text "\(\)"}) iff it satisfies conditions 1-4 (cf. \<^cite>\Kuratowski1\ -\<^cite>\Kuratowski2\). This characterization is shown consistent by generating a non-trivial model.}\ -abbreviation "\ \ \ Cl_1 \ \ Cl_2 \ \ Cl_3 \ \ Cl_4 \" -lemma "\ \" nitpick[satisfy, card w=3] oops (*model found: characterization is consistent*) - -text\\noindent{We verify some properties that will become useful later.}\ -lemma PC1: "Cl_1 \ = (Cl_1a \ \ Cl_1b \)" unfolding ADDI_def ADDI_a_def ADDI_b_def by blast -lemma PC4: "Cl_2 \ \ (Cl_4 \ = Cl_4' \)" unfolding EXP_def IDEM_def IDEMb_def by smt -lemma PC5: "Cl_2 \ \ Cl_5 \" unfolding EXP_def dNOR_def conn by simp -lemma PC6: "Cl_1 \ \ Cl_6 \" proof - - assume cl1: "Cl_1 \" - { fix a b - have "a \<^bold>\ b = (a \<^bold>\ b) \<^bold>\ b" unfolding conn by blast - hence "\(a \<^bold>\ b) = \((a \<^bold>\ b) \<^bold>\ b)" by simp - moreover from cl1 have "\(a \<^bold>\ b) \<^bold>\ \ a \<^bold>\ \ b" by (simp add: ADDI_def) - moreover from cl1 have "\((a \<^bold>\ b) \<^bold>\ b) \<^bold>\ \ (a \<^bold>\ b) \<^bold>\ (\ b)" by (simp add: ADDI_def) - ultimately have "\ a \<^bold>\ \ b \<^bold>\ \(a \<^bold>\ b) \<^bold>\ \ b" by simp - hence "(\ a \<^bold>\ \ b) \<^bold>\ \ b \<^bold>\ \(a \<^bold>\ b) \<^bold>\ \ b" unfolding conn by blast - hence "\ a \<^bold>\ \ b \<^bold>\ \ (a \<^bold>\ b)" unfolding conn by blast - } thus ?thesis by (simp add: Cl_6_def) -qed -lemma PC7: "Cl_1b \ \ Cl_2 \ \ Cl_7 \" by (smt EXP_def Cl_7_def MONO_def PC1 MONO_ADDIb conn) -lemma PC8: "Cl_1b \ \ Cl_2 \ \ Cl_4 \ \ Cl_8 \" by (smt Cl_8_def EXP_def IDEM_def MONO_ADDIb MONO_MULTa MULT_a_def meet_def) -lemma PC9: "Cl_1b \ \ Cl_4 \ \ Cl_9 \" by (smt IDEM_def Cl_9_def MONO_def MONO_ADDIb) - - -subsubsection \Exploring dualities\ - -lemma IC1_dual: "Int_1a \ = Cl_1b \" using MONO_ADDIb MONO_MULTa by auto -lemma "Int_1b \ = Cl_1a \" nitpick oops - -lemma IC1a: "Int_1a \ \ Cl_1b \\<^sup>d" by (smt MULT_a_def ADDI_b_def MONO_def MONO_MULTa dual_def conn) -lemma IC1b: "Int_1b \ \ Cl_1a \\<^sup>d" unfolding MULT_b_def ADDI_a_def dual_def conn by auto -lemma IC1: "Int_1 \ \ Cl_1 \\<^sup>d" unfolding MULT_def ADDI_def dual_def conn by simp -lemma IC2: "Int_2 \ \ Cl_2 \\<^sup>d" unfolding dEXP_def EXP_def dual_def conn by blast -lemma IC3: "Int_3 \ \ Cl_3 \\<^sup>d" unfolding dNOR_def NOR_def dual_def conn by simp -lemma IC4: "Int_4 \ \ Cl_4 \\<^sup>d" unfolding IDEM_def IDEM_def dual_def conn by simp -lemma IC4': "Int_4' \ \ Cl_4' \\<^sup>d" unfolding IDEMa_def IDEMb_def dual_def conn by metis -lemma IC5: "Int_5 \ \ Cl_5 \\<^sup>d" unfolding NOR_def dNOR_def dual_def conn by simp - -lemma CI1a: "Cl_1a \ \ Int_1b \\<^sup>d" proof - - assume cl1a: "Cl_1a \" - { fix A B - have "\<^bold>\\(\<^bold>\(A \<^bold>\ B)) \<^bold>\ \<^bold>\\(\<^bold>\A \<^bold>\ \<^bold>\B)" unfolding conn by simp - moreover from cl1a have "\<^bold>\(\(\<^bold>\A) \<^bold>\ \(\<^bold>\B)) \<^bold>\ \<^bold>\\(\<^bold>\A \<^bold>\ \<^bold>\B)" using ADDI_a_def conn by metis - ultimately have "\<^bold>\(\(\<^bold>\A)) \<^bold>\ \<^bold>\(\(\<^bold>\B)) \<^bold>\ \<^bold>\\(\<^bold>\(A \<^bold>\ B))" unfolding conn by simp - hence "(\\<^sup>d A) \<^bold>\ (\\<^sup>d B) \<^bold>\ (\\<^sup>d (A \<^bold>\ B))" unfolding dual_def by simp - } thus "Int_1b \\<^sup>d" using MULT_b_def by blast -qed -lemma CI1b: "Cl_1b \ \ Int_1a \\<^sup>d" by (metis IC1a MONO_ADDIb MONO_MULTa) -lemma CI1: "Cl_1 \ \ Int_1 \\<^sup>d" by (simp add: CI1a CI1b PC1 PI1) -lemma CI2: "Cl_2 \ \ Int_2 \\<^sup>d" unfolding EXP_def dEXP_def dual_def conn by metis -lemma CI3: "Cl_3 \ \ Int_3 \\<^sup>d" unfolding NOR_def dNOR_def dual_def conn by simp -lemma CI4: "Cl_4 \ \ Int_4 \\<^sup>d" unfolding IDEM_def IDEM_def dual_def conn by simp -lemma CI4': "Cl_4' \ \ Int_4' \\<^sup>d" unfolding IDEMa_def IDEMb_def dual_def conn by metis -lemma CI5: "Cl_5 \ \ Int_5 \\<^sup>d" unfolding dNOR_def NOR_def dual_def conn by simp - - -subsection \Frontier and border\ - -subsubsection \Frontier conditions\ - -definition "Fr_1a \ \ \A B. (A \<^bold>\ B) \<^bold>\ \(A \<^bold>\ B) \<^bold>\ (A \<^bold>\ B) \<^bold>\ (\ A \<^bold>\ \ B)" -definition "Fr_1b \ \ \A B. (A \<^bold>\ B) \<^bold>\ \(A \<^bold>\ B) \<^bold>\ (A \<^bold>\ B) \<^bold>\ (\ A \<^bold>\ \ B)" -definition "Fr_1 \ \ \A B. (A \<^bold>\ B) \<^bold>\ \(A \<^bold>\ B) \<^bold>\ (A \<^bold>\ B) \<^bold>\ (\ A \<^bold>\ \ B)" -definition "Fr_2 \ \ \A. \ A \<^bold>\ \(\<^bold>\A)" -abbreviation "Fr_3 \ \ NOR \" -definition "Fr_4 \ \ \A. \(\ A) \<^bold>\ (\ A)" - -definition "Fr_5 \ \ \A. \(\(\ A)) \<^bold>\ \(\ A)" -definition "Fr_6 \ \ \A B. A \<^bold>\ B \ (\ A \<^bold>\ B \<^bold>\ \ B)" - -text\\noindent{@{text "\"} is a topological frontier operator (@{text "\(\)"}) iff it satisfies conditions 1-4 -(cf. \<^cite>\Zarycki1\). This is also shown consistent by generating a non-trivial model.}\ -abbreviation "\ \ \ Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \" -lemma "\ \" nitpick[satisfy, card w=3] oops (*model found: characterization is consistent*) - -text\\noindent{We now verify some useful properties of the frontier operator.}\ -lemma PF1: "Fr_1 \ = (Fr_1a \ \ Fr_1b \)" unfolding Fr_1_def Fr_1a_def Fr_1b_def by meson -lemma PF5: "Fr_1 \ \ Fr_4 \ \ Fr_5 \" proof - - assume fr1: "Fr_1 \" and fr4: "Fr_4 \" - { fix A - from fr1 have "(\(\ A) \<^bold>\ (\ A)) \<^bold>\ \(\(\ A) \<^bold>\ (\ A)) \<^bold>\ (\(\ A) \<^bold>\ (\ A)) \<^bold>\ (\(\(\ A)) \<^bold>\ \(\ A))" by (simp add: Fr_1_def) - moreover have r1: "\(\ A) \<^bold>\ (\ A) \<^bold>\ \(\ A)" using meet_char Fr_4_def fr4 by simp - moreover have r2: "\(\(\ A)) \<^bold>\ \(\ A) \<^bold>\ \(\ A)" using join_char Fr_4_def fr4 by simp - ultimately have "(\(\ A) \<^bold>\ (\ A)) \<^bold>\ \(\(\ A)) \<^bold>\ (\(\ A) \<^bold>\ (\ A)) \<^bold>\ \(\ A)" unfolding conn by auto - hence "\(\(\ A)) \<^bold>\ \(\ A)" using r1 r2 conn by auto - } thus ?thesis by (simp add: Fr_5_def) -qed -lemma PF6: "Fr_1b \ \ Fr_2 \ \ Fr_6 \" proof - - assume fr1b: "Fr_1b \" and fr2: "Fr_2 \" - { fix A B - have "\(\<^bold>\(A \<^bold>\ B)) \<^bold>\ \(\<^bold>\A \<^bold>\ \<^bold>\B)" unfolding conn by simp - hence aux: "\(\<^bold>\A \<^bold>\ \<^bold>\B) \<^bold>\ \(A \<^bold>\ B)" by (metis (mono_tags) Fr_2_def fr2) - from fr1b have "(\<^bold>\A \<^bold>\ \<^bold>\B) \<^bold>\ \(\<^bold>\A \<^bold>\ \<^bold>\B) \<^bold>\ (\<^bold>\A \<^bold>\ \<^bold>\B) \<^bold>\ (\(\<^bold>\A) \<^bold>\ \(\<^bold>\B))" using Fr_1b_def by metis - hence "A \<^bold>\ B \<^bold>\ \<^bold>\\(\<^bold>\A \<^bold>\ \<^bold>\B) \<^bold>\ A \<^bold>\ B \<^bold>\ (\<^bold>\(\ A) \<^bold>\ \<^bold>\(\ B))" using Fr_2_def fr2 conn by auto - hence "\<^bold>\A \<^bold>\ \<^bold>\B \<^bold>\ \<^bold>\\(\<^bold>\A \<^bold>\ \<^bold>\B) \<^bold>\ \<^bold>\A \<^bold>\ \<^bold>\B \<^bold>\ \<^bold>\(\ A) \<^bold>\ \<^bold>\(\ B)" unfolding conn by blast - hence "A \<^bold>\ B \<^bold>\ \(A \<^bold>\ B) \<^bold>\ A \<^bold>\ B \<^bold>\ (\ A) \<^bold>\ (\ B)" using aux unfolding conn by blast - hence "A \<^bold>\ B \ B \<^bold>\ \(A \<^bold>\ B) \<^bold>\ B \<^bold>\ (\ A) \<^bold>\ (\ B)" unfolding conn by auto - hence "A \<^bold>\ B \ B \<^bold>\ (\ B) \<^bold>\ B \<^bold>\ (\ A) \<^bold>\ (\ B)" using join_char unfolding conn by simp - hence "A \<^bold>\ B \ (\ A) \<^bold>\ B \<^bold>\ (\ B)" unfolding conn by simp - } thus "Fr_6 \" by (simp add: Fr_6_def) -qed - - -subsubsection \Border conditions\ - -definition "Br_1 \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (A \<^bold>\ \ B) \<^bold>\ (B \<^bold>\ \ A)" -definition "Br_2 \ \ (\ \<^bold>\) \<^bold>\ \<^bold>\" -definition "Br_3 \ \ \A. \(\\<^sup>d A) \<^bold>\ A" - -definition "Br_4 \ \ \A B. A \<^bold>\ B \ A \<^bold>\ (\ B) \<^bold>\ \ A" -definition "Br_5a \ \ \A. \(\\<^sup>d A) \<^bold>\ \ A" -definition "Br_5b \ \ \A. \ A \<^bold>\ A" -definition "Br_5c \ \ \A. A \<^bold>\ \\<^sup>d A" -definition "Br_5d \ \ \A. \\<^sup>d A \<^bold>\ \\<^sup>d(\ A)" -abbreviation "Br_6 \ \ IDEM \" -abbreviation "Br_7 \ \ ADDI_a \" -abbreviation "Br_8 \ \ MULT_b \" -definition "Br_9 \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" -definition "Br_10 \ \ \A. \(\<^bold>\(\ A) \<^bold>\ \\<^sup>d A) \<^bold>\ \<^bold>\" - -text\\noindent{@{text "\"} is a topological border operator (@{text "\(\)"}) iff it satisfies conditions 1-3 -(cf. \<^cite>\Zarycki1\). This is also shown consistent.}\ -abbreviation "\ \ \ Br_1 \ \ Br_2 \ \ Br_3 \" -lemma "\ \" nitpick[satisfy, card w=3] oops (*model found: characterization is consistent*) - -text\\noindent{We now verify some useful properties of the border operator.}\ -lemma PB4: "Br_1 \ \ Br_4 \" proof - - assume br1: "Br_1 \" - { fix A B - have aux: "\(A \<^bold>\ B) \<^bold>\ (A \<^bold>\ \ B) \<^bold>\ (B \<^bold>\ \ A)" using br1 Br_1_def by simp - { assume "A \<^bold>\ B" - hence "\(A \<^bold>\ B) \<^bold>\ \ A" using meet_char unfolding conn by simp - hence "\ A \<^bold>\ (A \<^bold>\ \ B) \<^bold>\ (B \<^bold>\ \ A)" using aux by auto - hence "A \<^bold>\ \ B \<^bold>\ \ A" unfolding conn by auto - } hence "A \<^bold>\ B \ A \<^bold>\ \ B \<^bold>\ \ A" by (rule impI) - } thus "Br_4 \" by (simp add: Br_4_def) -qed -lemma PB5b: "Br_1 \ \ Br_5b \" proof - - assume br1: "Br_1 \" - { fix A - from br1 have aux: "\(A \<^bold>\ A) \<^bold>\ (A \<^bold>\ \ A) \<^bold>\ (A \<^bold>\ \ A)" unfolding Br_1_def by blast - hence "\ A \<^bold>\ (A \<^bold>\ \ A)" unfolding conn by metis - hence "\ A \<^bold>\ A" unfolding conn by blast - } thus "Br_5b \" using Br_5b_def by blast -qed -lemma PB5c: "Br_1 \ \ Br_5c \" using Br_5b_def Br_5c_def PB5b dual_def unfolding conn by force -lemma PB5a: "Br_1 \ \ Br_3 \ \ Br_5a \" proof - - assume br1: "Br_1 \" and br3: "Br_3 \" - hence br5c: "Br_5c \" using PB5c by simp - { fix A - have "A \<^bold>\ \(\\<^sup>d A) \<^bold>\ \ A" by (metis br5c Br_4_def Br_5c_def PB4 br1) - hence "\(\\<^sup>d A) \<^bold>\ \ A" using Br_3_def br3 unfolding conn by metis - } thus "Br_5a \" using Br_5a_def by simp -qed -lemma PB5d: "Br_1 \ \ Br_3 \ \ Br_5d \" proof - - assume br1: "Br_1 \" and br3: "Br_3 \" - hence br5a: "Br_5a \" using PB5a by simp - { fix A - from br5a have "\(\\<^sup>d(\<^bold>\A)) \<^bold>\ \(\<^bold>\A)" unfolding Br_5a_def by simp - hence "\<^bold>\\(\<^bold>\A) \<^bold>\ \<^bold>\\(\\<^sup>d(\<^bold>\A))" unfolding conn by blast - hence "\\<^sup>d A \<^bold>\ \\<^sup>d(\ A)" unfolding dual_def conn by simp - } thus "Br_5d \" using Br_5d_def by simp -qed -lemma PB6: "Br_1 \ \ Br_6 \" by (smt Br_4_def Br_5b_def IDEM_def PB4 PB5b conn) -lemma PB7: "Br_1 \ \ Br_7 \" using Br_4_def Br_5b_def ADDI_a_def PB4 PB5b unfolding conn by smt -lemma PB8: "Br_1 \ \ Br_8 \" using Br_1_def Br_5b_def MULT_b_def PB5b unfolding conn by metis -lemma PB9: "Br_1 \ \ Br_9 \" unfolding Br_1_def Br_9_def conn by simp -lemma PB10: "Br_1 \ \ Br_3 \ \ Br_10 \" proof - \\ proof automagically generated \ - assume a1: "Br_3 \" - assume a2: "Br_1 \" - { fix bb :: "w \ bool" and ww :: w - have "\p pa w pb. ((pb p w \ pb pa w) \ \pb (pa \<^bold>\ p) w) \ \Br_9 pb" - by (metis Br_9_def join_def) (* 20 ms *) - then have "(\ (((\\<^sup>c) \<^bold>\ (\\<^sup>d)) bb) ww) \ (\<^bold>\ ww) \ \(\ (((\\<^sup>c) \<^bold>\ (\\<^sup>d)) bb) ww) \ \(\<^bold>\ ww)" - using a2 a1 by (metis (no_types) Br_5a_def Br_5b_def Br_5d_def PB5a PB5b PB5d PB9 bottom_def compl_def dual_def meet_def) (* 86 ms *) - } then show ?thesis unfolding Br_10_def by blast (* 1 ms *) -qed - - -subsubsection \Relation with closure and interior\ - -text\\noindent{We define and verify some conversion operators useful to derive border and frontier operators from -closure/interior operators and also between each other.}\ - -text\\noindent{Frontier operator as derived from interior.}\ -definition Fr_int::"(\\\)\(\\\)" ("\\<^sub>I") where "\\<^sub>I \ \ \A. \<^bold>\(\ A) \<^bold>\ \\<^sup>d A" -lemma FI1: "Int_1 \ \ Int_2 \ \ Fr_1(\\<^sub>I \)" using EXP_def Fr_1_def Fr_int_def IC2 MULT_def conn by auto -lemma FI2: "Fr_2(\\<^sub>I \)" using Fr_2_def Fr_int_def dual_def dual_symm unfolding conn by smt -lemma FI3: "Int_3 \ \ Fr_3(\\<^sub>I \)" using NOR_def Fr_int_def IC3 unfolding conn by auto -lemma FI4: "Int_1a \ \ Int_2 \ \ Int_4 \ \ Fr_4(\\<^sub>I \)" unfolding Fr_int_def Fr_4_def dual_def conn by (smt Int_9_def MONO_MULTa PI9) - -text\\noindent{Frontier operator as derived from closure.}\ -definition Fr_cl::"(\\\)\(\\\)" ("\\<^sub>C") where "\\<^sub>C \ \ \A. (\ A) \<^bold>\ \(\<^bold>\A)" -lemma FC1: "Cl_1 \ \ Cl_2 \ \ Fr_1(\\<^sub>C \)" using CI1 EXP_def Fr_1_def Fr_cl_def MULT_def dual_def unfolding conn by smt -lemma FC2: "Fr_2(\\<^sub>C \)" using Fr_2_def Fr_cl_def dual_def dual_symm unfolding conn by smt -lemma FC3: "Cl_3 \ \ Fr_3(\\<^sub>C \)" unfolding NOR_def Fr_cl_def conn by simp -lemma FC4: "Cl_1b \ \ Cl_2 \ \ Cl_4 \ \ Fr_4(\\<^sub>C \)" using Cl_8_def MONO_ADDIb PC8 unfolding Fr_cl_def Fr_4_def conn by blast - -text\\noindent{Frontier operator as derived from border.}\ -definition Fr_br::"(\\\)\(\\\)" ("\\<^sub>B") where "\\<^sub>B \ \ \A. \ A \<^bold>\ \(\<^bold>\A)" -lemma FB1: "Br_1 \ \ Fr_1(\\<^sub>B \)" unfolding Fr_br_def Fr_1_def using Br_1_def Br_5b_def PB5b conn by smt -lemma FB2: "Fr_2(\\<^sub>B \)" unfolding Fr_br_def Fr_2_def conn by auto -lemma FB3: "Br_1 \ \ Br_2 \ \ Fr_3(\\<^sub>B \)" using Br_2_def Br_5b_def PB5b unfolding Fr_br_def NOR_def conn by force -lemma FB4: "Br_1 \ \ Br_3 \ \ Fr_4(\\<^sub>B \)" proof - - assume br1: "Br_1 \" and br3: "Br_3 \" - { fix A - have "(\\<^sub>B \) ((\\<^sub>B \) A) = \(\ A \<^bold>\ \(\<^bold>\A)) \<^bold>\ \(\<^bold>\(\ A \<^bold>\ \(\<^bold>\A)))" by (simp add: Fr_br_def conn) - also have "... = \(\ A \<^bold>\ \(\<^bold>\A)) \<^bold>\ \(\<^bold>\(\ A) \<^bold>\ \\<^sup>d A)" by (simp add: dual_def conn) - also from br1 br3 have "... = \(\ A \<^bold>\ \(\<^bold>\A)) \<^bold>\ \<^bold>\" using Br_10_def PB10 by metis - finally have "(\\<^sub>B \) ((\\<^sub>B \) A) \<^bold>\ \(\ A \<^bold>\ \(\<^bold>\A))" unfolding conn by simp - hence "(\\<^sub>B \) ((\\<^sub>B \) A) \<^bold>\ (\\<^sub>B \) A" using Br_5b_def Fr_br_def PB5b br1 by fastforce - } thus "Fr_4(\\<^sub>B \)" using Fr_4_def by blast -qed - -text\\noindent{Border operator as derived from interior.}\ -definition Br_int::"(\\\)\(\\\)" ("\\<^sub>I") where "\\<^sub>I \ \ \A. A \<^bold>\ (\ A)" -lemma BI1: "Int_1 \ \ Br_1 (\\<^sub>I \)" using Br_1_def Br_int_def MULT_def conn by auto -lemma BI2: "Int_3 \ \ Br_2 (\\<^sub>I \)" by (simp add: Br_2_def Br_int_def dNOR_def conn) -lemma BI3: "Int_1a \ \ Int_4 \ \ Br_3 (\\<^sub>I \)" unfolding Br_int_def Br_3_def dual_def by (smt Int_9_def MONO_MULTa PI9 conn) - -text\\noindent{Border operator as derived from closure.}\ -definition Br_cl::"(\\\)\(\\\)" ("\\<^sub>C") where "\\<^sub>C \ \ \A. A \<^bold>\ \(\<^bold>\A)" -lemma BC1: "Cl_1 \ \ Br_1(\\<^sub>C \)" using Br_1_def Br_cl_def CI1 MULT_def dual_def unfolding conn by smt -lemma BC2: "Cl_3 \ \ Br_2(\\<^sub>C \)" using Br_2_def Br_cl_def CI3 dNOR_def dual_def unfolding conn by auto -lemma BC3: "Cl_1b \ \ Cl_4 \ \ Br_3 (\\<^sub>C \)" unfolding Br_cl_def Br_3_def dual_def conn by (metis (mono_tags, lifting) Cl_9_def PC9) - -text\\noindent{Note that the previous two conversion functions are related:}\ -lemma BI_BC_rel: "(\\<^sub>I \) = \\<^sub>C(\\<^sup>d)" by (simp add: Br_cl_def Br_int_def dual_def conn) - -text\\noindent{Border operator as derived from frontier.}\ -definition Br_fr::"(\\\)\(\\\)" ("\\<^sub>F") where "\\<^sub>F \ \ \A. A \<^bold>\ (\ A)" -lemma BF1: "Fr_1 \ \ Br_1(\\<^sub>F \)" using Br_1_def Br_fr_def Fr_1_def conn by auto -lemma BF2: "Fr_2 \ \ Fr_3 \ \ Br_2(\\<^sub>F \)" using Br_2_def Br_fr_def Fr_2_def NOR_def unfolding conn by fastforce -lemma BF3: "Fr_1b \ \ Fr_2 \ \ Fr_4 \ \ Br_3(\\<^sub>F \)" proof - - assume fr1b: "Fr_1b \" and fr2: "Fr_2 \" and fr4: "Fr_4 \" - { fix A - from fr1b fr2 have "let M= \<^bold>\A \<^bold>\ \ A ; N= \ A in M \<^bold>\ N \ (\ M \<^bold>\ N \<^bold>\ \ N)" using PF1 PF6 by (simp add: Fr_6_def) - hence "\(\<^bold>\A \<^bold>\ \ A) \<^bold>\ \ A \<^bold>\ \(\ A)" unfolding conn by meson - hence "\(\<^bold>\A \<^bold>\ \ A) \<^bold>\ \ A" using Fr_4_def fr4 unfolding conn by metis - hence aux: "\(\<^bold>\A \<^bold>\ \ A) \<^bold>\ \<^bold>\(\ A) \<^bold>\ \<^bold>\" unfolding conn by blast - have "(\\<^sub>F \)((\\<^sub>F \)\<^sup>d A) = \<^bold>\(\<^bold>\A \<^bold>\ \(\<^bold>\A)) \<^bold>\ \(\<^bold>\(\<^bold>\A \<^bold>\ \(\<^bold>\A)))" unfolding Br_fr_def dual_def by simp - also have "... = (A \<^bold>\ \<^bold>\\ A) \<^bold>\ \(\<^bold>\A \<^bold>\ \ A)" using Fr_2_def fr2 unfolding conn by force - also have "... = (A \<^bold>\ \(\<^bold>\A \<^bold>\ \ A)) \<^bold>\ (\<^bold>\\ A \<^bold>\ \(\<^bold>\A \<^bold>\ \ A))" unfolding conn by auto - also have "... = (A \<^bold>\ \(\<^bold>\A \<^bold>\ \ A))" using aux unfolding conn by blast - finally have "(\\<^sub>F \)((\\<^sub>F \)\<^sup>d A) = A \<^bold>\ \(\<^bold>\A \<^bold>\ \ A)" unfolding Br_fr_def dual_def conn by simp - } thus "Br_3(\\<^sub>F \)" unfolding Br_3_def Br_fr_def conn by meson -qed - -text\\noindent{Interior operator as derived from border.}\ -definition Int_br::"(\\\)\(\\\)" ("\\<^sub>B") where "\\<^sub>B \ \ \A. A \<^bold>\ (\ A)" -lemma IB1: "Br_1 \ \ Int_1(\\<^sub>B \)" using Br_1_def MULT_def Int_br_def conn by auto -lemma IB2: "Int_2(\\<^sub>B \)" by (simp add: dEXP_def Int_br_def conn) -lemma IB3: "Br_2 \ \ Int_3(\\<^sub>B \)" by (simp add: Br_2_def dNOR_def Int_br_def conn) -lemma IB4: "Br_1 \ \ Br_3 \ \ Int_4(\\<^sub>B \)" unfolding Int_br_def IDEM_def conn - using Br_1_def Br_5c_def Br_5d_def PB5c PB5d dual_def unfolding conn by (metis (full_types)) - -text\\noindent{Interior operator as derived from frontier.}\ -definition Int_fr::"(\\\)\(\\\)" ("\\<^sub>F") where "\\<^sub>F \ \ \A. A \<^bold>\ (\ A)" -lemma IF1a: "Fr_1b \ \ Int_1a(\\<^sub>F \)" using Fr_1b_def MULT_a_def Int_fr_def conn by auto -lemma IF1b: "Fr_1a \ \ Int_1b(\\<^sub>F \)" using Fr_1a_def MULT_b_def Int_fr_def conn by auto -lemma IF1: "Fr_1 \ \ Int_1(\\<^sub>F \)" by (simp add: IF1a IF1b PF1 PI1) -lemma IF2: "Int_2(\\<^sub>F \)" by (simp add: dEXP_def Int_fr_def conn) -lemma IF3: "Fr_2 \ \ Fr_3 \ \ Int_3(\\<^sub>F \)" using BF2 Br_2_def Br_fr_def dNOR_def Int_fr_def unfolding conn by auto -lemma IF4: "Fr_1a \ \ Fr_2 \ \ Fr_4 \ \ Int_4(\\<^sub>F \)" - using Fr_1a_def Fr_2_def Fr_4_def unfolding Int_fr_def IDEM_def conn by metis - -text\\noindent{Closure operator as derived from border.}\ -definition Cl_br::"(\\\)\(\\\)" ("\\<^sub>B") where "\\<^sub>B \ \ \A. A \<^bold>\ \(\<^bold>\A)" -lemma CB1: "Br_1 \ \ Cl_1(\\<^sub>B \)" proof - - assume br1: "Br_1 \" - { fix A B - have "(\\<^sub>B \) (A \<^bold>\ B) = A \<^bold>\ B \<^bold>\ \(\<^bold>\(A \<^bold>\ B))" by (simp add: Cl_br_def conn) - also have "... = A \<^bold>\ B \<^bold>\ \(\<^bold>\A \<^bold>\ \<^bold>\B)" unfolding conn by simp - also have "... = A \<^bold>\ B \<^bold>\ (\<^bold>\A \<^bold>\ \(\<^bold>\B)) \<^bold>\ (\<^bold>\B \<^bold>\ \(\<^bold>\A))" using Br_1_def br1 unfolding conn by auto - also have "... = A \<^bold>\ \(\<^bold>\A) \<^bold>\ B \<^bold>\ \(\<^bold>\B)" unfolding conn by auto - also have "... = (\\<^sub>B \) A \<^bold>\ (\\<^sub>B \) B" by (simp add: Cl_br_def conn) - finally have "(\\<^sub>B \)(A \<^bold>\ B) = (\\<^sub>B \) A \<^bold>\ (\\<^sub>B \) B" unfolding Cl_br_def by blast - } thus "Cl_1(\\<^sub>B \)" unfolding ADDI_def Cl_br_def by simp -qed -lemma CB2: "Cl_2(\\<^sub>B \)" by (simp add: EXP_def Cl_br_def conn) -lemma CB3: "Br_2 \ \ Cl_3(\\<^sub>B \)" using Br_2_def Cl_br_def IC3 dNOR_def dual_def dual_symm unfolding conn by smt -lemma CB4: "Br_1 \ \ Br_3 \ \ Cl_4(\\<^sub>B \)" proof - - assume br1: "Br_1 \" and br3: "Br_3 \" - { fix A - have "(\\<^sub>B \) ((\\<^sub>B \) A) = A \<^bold>\ \(\<^bold>\A) \<^bold>\ \(\<^bold>\(A \<^bold>\ \(\<^bold>\A)))" by (simp add: Cl_br_def conn) - also have "... = A \<^bold>\ \(\<^bold>\A) \<^bold>\ \(\<^bold>\A \<^bold>\ \\<^sup>d A)" unfolding dual_def conn by simp - also have "... = A \<^bold>\ \(\<^bold>\A) \<^bold>\ (\<^bold>\A \<^bold>\ \(\\<^sup>d A)) \<^bold>\ (\<^bold>\A \<^bold>\ \(\<^bold>\A))" unfolding dual_def using Br_1_def br1 conn by auto - also have "... = A \<^bold>\ \(\<^bold>\A)" using Br_3_def br3 unfolding conn by metis - finally have "(\\<^sub>B \) ((\\<^sub>B \) A) = (\\<^sub>B \) A" unfolding Cl_br_def by blast - } thus "Cl_4(\\<^sub>B \)" unfolding IDEM_def Cl_br_def by simp -qed - -text\\noindent{Closure operator as derived from frontier.}\ -definition Cl_fr::"(\\\)\(\\\)" ("\\<^sub>F") where "\\<^sub>F \ \ \A. A \<^bold>\ (\ A)" -lemma CF1b: "Fr_1b \ \ Fr_2 \ \ Cl_1b(\\<^sub>F \)" using Cl_fr_def Fr_6_def MONO_def MONO_ADDIb PF6 unfolding conn by auto -lemma CF1a: "Fr_1a \ \ Fr_2 \ \ Cl_1a(\\<^sub>F \)" proof - - have aux: "Fr_2 \ \ (\\<^sub>F \)\<^sup>d = (\\<^sub>F \)" by (simp add: Cl_fr_def Fr_2_def Int_fr_def dual_def conn) - hence "Fr_1a \ \ Fr_2 \ \ Cl_1a(\\<^sub>F \)\<^sup>d" using IC1b IF1b by blast - thus "Fr_1a \ \ Fr_2 \ \ Cl_1a(\\<^sub>F \)" using ADDI_a_def aux unfolding conn by simp -qed -lemma CF1: "Fr_1 \ \ Fr_2 \ \ Cl_1(\\<^sub>F \)" by (simp add: CF1a CF1b PC1 PF1) -lemma CF2: "Cl_2(\\<^sub>F \)" by (simp add: EXP_def Cl_fr_def conn) -lemma CF3: "Fr_3 \ \ Cl_3(\\<^sub>F \)" by (simp add: Cl_fr_def NOR_def conn) -lemma CF4: "Fr_1a \ \ Fr_2 \ \ Fr_4 \ \ Cl_4(\\<^sub>F \)" proof - - have aux: "Fr_2 \ \ (\\<^sub>F \)\<^sup>d = (\\<^sub>F \)" by (simp add: Cl_fr_def Fr_2_def Int_fr_def dual_def conn) - hence "Fr_1a \ \ Fr_2 \ \ Fr_4 \ \ Cl_4(\\<^sub>F \)\<^sup>d" using IC4 IF4 by blast - thus "Fr_1a \ \ Fr_2 \ \ Fr_4 \ \ Cl_4(\\<^sub>F \)" by (simp add: aux) -qed - - -subsubsection \Infinitary conditions\ - -text\\noindent{We define the essential infinitary conditions for the closure and interior operators (entailing infinite -additivity and multiplicativity resp.). Observe that the other direction is implied by monotonicity (MONO).}\ -abbreviation "Cl_inf \ \ iADDI_a(\)" -abbreviation "Int_inf \ \ iMULT_b(\)" - -text\\noindent{There exists indeed a condition on frontier operators responsible for the infinitary conditions above:}\ -definition "Fr_inf \ \ \S. \<^bold>\S \<^bold>\ \(\<^bold>\S) \<^bold>\ \<^bold>\S \<^bold>\ \<^bold>\Ra[\|S]" - -lemma CF_inf: "Fr_2 \ \ Fr_inf \ \ Cl_inf(\\<^sub>F \)" unfolding iADDI_a_def Fr_inf_def - by (smt Cl_fr_def Fr_2_def Ra_restr_ex compl_def dom_compl_def2 iDM_b join_def meet_def supremum_def) -lemma IF_inf: "Fr_inf \ \ Int_inf(\\<^sub>F \)" unfolding Fr_inf_def iMULT_b_def Int_fr_def Ra_restr_all - by (metis (mono_tags, opaque_lifting) diff_def infimum_def meet_def pfunRange_restr_def supremum_def) - -text\\noindent{This condition is indeed strong enough to entail closure of the fixed-point predicates under infimum/supremum.}\ -lemma fp_IF_inf_closed: "Fr_inf \ \ infimum_closed (fp (\\<^sub>F \))" by (metis (full_types) IF2 IF_inf Ra_restr_all dEXP_def iMULT_b_def infimum_def) -lemma fp_CF_sup_closed: "Fr_inf \ \ Fr_2 \ \ supremum_closed (fp (\\<^sub>F \))" by (metis (full_types) CF2 CF_inf EXP_def Ra_restr_ex iADDI_a_def supremum_def) - -end diff --git a/thys/Topological_Semantics/topo_operators_derivative.thy b/thys/Topological_Semantics/topo_operators_derivative.thy deleted file mode 100644 --- a/thys/Topological_Semantics/topo_operators_derivative.thy +++ /dev/null @@ -1,207 +0,0 @@ -theory topo_operators_derivative - imports topo_operators_basic -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - - -subsection \Derivative and coherence\ - -text\\noindent{In this section we investigate two related operators, namely the `derivative' (or `derived set') -and the (Cantorian) `coherence' of a set. The derivative of a set is the set of its accumulation (aka. limit) -points. The coherence of a set A is the set formed by those limit points of A belonging to A. -For the derivative operator we draw upon the works by Kuratowski \<^cite>\Kuratowski1\ and -(in more detail) by Zarycki \<^cite>\Zarycki3\; cf.~also McKinsey \& Tarski \<^cite>\AOT\. -For the (Cantorian) coherence operator we follow the treatment given by Zarycki in \<^cite>\Zarycki2\.}\ - -subsubsection \Derivative conditions\ - -text\\noindent{The derivative conditions overlap partly with Kuratowski closure conditions \<^cite>\Kuratowski1\. -We try to make both notations coincide when possible.}\ - -abbreviation "Der_1 \ \ Cl_1 \" -abbreviation "Der_1a \ \ Cl_1a \" -abbreviation "Der_1b \ \ Cl_1b \" -abbreviation "Der_2 \ \ Cl_5 \" \\ follows from Cl-2 \ -abbreviation "Der_3 \ \ Cl_3 \" -abbreviation "Der_4 \ \ Cl_4' \" -definition "Der_4e \ \ \A. \(\ A) \<^bold>\ (\ A \<^bold>\ A)" -definition "Der_5 \ \ \A. (\ A \<^bold>\ A) \ (A \<^bold>\ \\<^sup>d A) \ (A \<^bold>\ \<^bold>\ \ A \<^bold>\ \<^bold>\)" - -text\\noindent{Some remarks: -Condition Der-2 basically says (when assuming other derivative axioms) that the space is dense-in-itself, -i.e. that all points are accumulation points (no point is isolated) w.r.t the whole space. -Der-4 is a weakened (left-to-right) variant of Cl-4. -Condition Der-4e corresponds to a (weaker) condition than Der-4 and is used in more recent literature -(in particular in the works of Leo Esakia \<^cite>\Esakia\). -When other derivative axioms are assumed, Der-5 above as used by Zarycki \<^cite>\Zarycki3\ says that -the only clopen sets in the space are the top and bottom elements (empty set and universe, resp.). -We verify some properties:}\ - -lemma Der4e_rel: "Der_4 \ \ Der_4e \" by (simp add: IDEMb_def Der_4e_def conn) -lemma PD1: "Der_1b \ \ \A B. A \<^bold>\ B \ \ A \<^bold>\ \ B" using MONO_def MONO_ADDIb by auto -lemma PD2: "Der_1b \ \ \A B. A \<^bold>\ B \ \\<^sup>d A \<^bold>\ \\<^sup>d B" using CI1b MONO_def MONO_MULTa dual_def by fastforce -lemma PD3: "Der_1b \ \ \A B. \(A \<^bold>\ B) \<^bold>\ \ A \<^bold>\ \ B" unfolding conn by (metis (no_types, lifting) PD1) -lemma PD4: "Der_1 \ \ \A B. (\ A \<^bold>\ \ B) \<^bold>\ \(A \<^bold>\ B)" using Cl_6_def PC6 by metis -lemma PD5: "Der_4 \ \ \A. \(\(\<^bold>\(\ A))) \<^bold>\ \(\<^bold>\(\ A))" unfolding IDEMb_def by blast -text\\noindent{Observe that the lemmas below require Der-2 as premise:}\ -lemma PD6: "Der_1 \ \ Der_2 \ \ \A. \\<^sup>d A \<^bold>\ \ A" unfolding ADDI_def dNOR_def dual_def conn by fastforce -lemma PD7: "Der_1 \ \ Der_2 \ \ \A. \(\\<^sup>d A) \<^bold>\ \(\ A)" by (metis (mono_tags, lifting) PC1 PD1 PD6) -lemma PD8: "Der_1 \ \ Der_2 \ \ Der_4 \ \ \A. \(\\<^sup>d A) \<^bold>\ \ A" by (meson IDEMb_def PD7) -lemma PD9: "Der_1 \ \ Der_2 \ \ Der_4 \ \ \A. \\<^sup>d A \<^bold>\ \\<^sup>d(\ A)" by (metis CI4' IDEMa_def PC1 PD2 PD6) -lemma PD10: "Der_1 \ \ Der_2 \ \ Der_4 \ \ \A. \\<^sup>d A \<^bold>\ \(\\<^sup>d A)" using CI4' IDEMa_def PD6 by metis -lemma PD11: "Der_1 \ \ Der_2 \ \ Der_4 \ \ \A. \<^bold>\(\ A) \<^bold>\ \(\<^bold>\(\ A))" using IDEMb_def PD6 dual_def unfolding conn by metis -lemma PD12: "Der_1 \ \ Der_2 \ \ Der_4 \ \ \A. (\\<^sup>d A) \<^bold>\ (\ A) \<^bold>\ \\<^sup>d(A \<^bold>\ (\ A))" proof - - assume der1: "Der_1 \" and der2: "Der_2 \" and der4: "Der_4 \" - { fix A - from der1 have "let M=\<^bold>\A ; N=\<^bold>\(\ A) in \(M \<^bold>\ N) \<^bold>\ (\ M) \<^bold>\ (\ N)" unfolding ADDI_def by simp - hence aux: "\<^bold>\(\(\<^bold>\A) \<^bold>\ \(\<^bold>\(\ A))) \<^bold>\ \<^bold>\(\ (\<^bold>\A \<^bold>\ \<^bold>\(\ A)))" unfolding dual_def conn by simp - have "(\\<^sup>d A \<^bold>\ \ A) = (\\<^sup>d A \<^bold>\ \\<^sup>d(\ A))" using PD6 PD9 der1 der2 der4 unfolding conn by metis - also have "... = \<^bold>\(\(\<^bold>\A) \<^bold>\ \(\<^bold>\(\ A)))" unfolding dual_def conn by simp - also from aux have "... = \<^bold>\(\ (\<^bold>\A \<^bold>\ \<^bold>\(\ A)))" unfolding dual_def by blast - also have "... = \\<^sup>d(A \<^bold>\ (\ A))" unfolding dual_def conn by simp - finally have "(\\<^sup>d A) \<^bold>\ (\ A) \<^bold>\ \\<^sup>d(A \<^bold>\ (\ A))" by simp - } thus ?thesis by simp -qed - -text\\noindent{The conditions below can serve to axiomatize a derivative operator. Different authors consider different -sets of conditions. We define below some corresponding to Zarycki \<^cite>\Zarycki3\, Kuratowski \<^cite>\Kuratowski1\ -\<^cite>\Zarycki2\, McKinsey \& Tarski \<^cite>\AOT\, and Esakia \<^cite>\Esakia\, respectively.}\ -abbreviation "\
z \ \ Der_1 \ \ Der_2 \ \ Der_3 \ \ Der_4 \ \ Der_5 \" -abbreviation "\
k \ \ Der_1 \ \ Der_2 \ \ Der_3 \ \ Der_4 \" -abbreviation "\
mt \ \ Der_1 \ \ Der_3 \ \ Der_4 \" -abbreviation "\
e \ \ Der_1 \ \ Der_3 \ \ Der_4e \" - -text\\noindent{Our `default' derivative operator will coincide with @{text "\
k"} from Kuratowski (also Zarycki). -However, for proving theorems we will employ the weaker variant Der-4e instead of Der-4 whenever possible. -We start by defining a dual operator and verifying some dualities; we then define conversion operators. -Observe that conditions Der-2 and Der-5 are not used in the rest of this subsection. -Der-2 will be required later when working with the coherence operator.}\ -abbreviation "\
\ \ \
k \" -abbreviation "\ \ \ Int_1 \ \ Int_3 \ \ Int_4' \ \ Int_5 \" \\ 'dual-derivative' operator \ - -lemma SD_dual1: "\(\) \ \
(\\<^sup>d)" using IC1 IC4' IC3 IC5 by simp -lemma SD_dual2: "\(\\<^sup>d) \ \
(\)" using IC1 IC4' IC3 IC5 dual_symm by metis -lemma DS_dual1: "\
(\) \ \(\\<^sup>d)" using CI1 CI4' CI3 CI5 by simp -lemma DS_dual2: "\
(\\<^sup>d) \ \(\)" using CI1 CI4' CI3 CI5 dual_symm by metis -lemma DS_dual: "\
(\) = \(\\<^sup>d)" using SD_dual2 DS_dual1 by smt - -text\\noindent{Closure operator as derived from derivative.}\ -definition Cl_der::"(\\\)\(\\\)" ("\\<^sub>D") where "\\<^sub>D \ \ \A. A \<^bold>\ \(A)" -text\\noindent{Verify properties:}\ -lemma CD1a: "Der_1a \ \ Cl_1a (\\<^sub>D \)" unfolding Cl_der_def ADDI_a_def conn by auto -lemma CD1b: "Der_1b \ \ Cl_1b (\\<^sub>D \)" unfolding Cl_der_def ADDI_b_def conn by auto -lemma CD1 : "Der_1 \ \ Cl_1 (\\<^sub>D \)" unfolding Cl_der_def ADDI_def conn by blast -lemma CD2: "Cl_2 (\\<^sub>D \)" unfolding Cl_der_def EXP_def conn by simp -lemma CD3: "Der_3 \ \ Der_3 (\\<^sub>D \)" unfolding Cl_der_def NOR_def conn by simp -lemma CD4a: "Der_1a \ \ Der_4e \ \ Cl_4 (\\<^sub>D \)" unfolding Cl_der_def by (smt ADDI_a_def Der_4e_def IDEM_def join_def) -lemma "Der_1b \ \ Der_4 \ \ Cl_4 (\\<^sub>D \)" nitpick oops (*countermodel*) -lemma CD4: "Der_1 \ \ Der_4e \ \ Cl_4 (\\<^sub>D \)" unfolding Cl_der_def by (metis (no_types, lifting) CD4a Cl_der_def IDEM_def PC1) - -text\\noindent{Interior operator as derived from (dual) derivative.}\ -definition Int_der::"(\\\)\(\\\)" ("\\<^sub>D") where "\\<^sub>D \ \ \A. A \<^bold>\ \\<^sup>d(A)" -text\\noindent{Verify definition:}\ -lemma Int_der_def2: "\\<^sub>D \ = (\\. \ \<^bold>\ \(\<^bold>\\))" unfolding Int_der_def dual_def conn by simp -lemma dual_der1: "\\<^sub>D \ \ (\\<^sub>D \)\<^sup>d" unfolding Cl_der_def Int_der_def dual_def conn by simp -lemma dual_der2: "\\<^sub>D \ \ (\\<^sub>D \)\<^sup>d" unfolding Cl_der_def Int_der_def dual_def conn by simp -text\\noindent{Verify properties:}\ -lemma ID1: "Der_1 \ \ Int_1 (\\<^sub>D \)" using Int_der_def MULT_def ADDI_def CI1 unfolding conn by auto -lemma ID1a: "Der_1a \ \ Int_1b (\\<^sub>D \)" by (simp add: CI1a dual_der2 CD1a) -lemma ID1b: "Der_1b \ \ Int_1a (\\<^sub>D \)" by (simp add: CI1b dual_der2 CD1b) -lemma ID2: "Int_2 (\\<^sub>D \)" unfolding Int_der_def dEXP_def conn by simp -lemma ID3: "Der_3 \ \ Int_3 (\\<^sub>D \)" by (simp add: CI3 CD3 dual_der2) -lemma ID4: "Der_1 \ \ Der_4e \ \ Int_4 (\\<^sub>D \)" using CI4 CD4 dual_der2 by simp -lemma ID4a: "Der_1a \ \ Der_4e \ \ Int_4 (\\<^sub>D \)" by (simp add: CI4 dual_der2 CD4a) -lemma "Der_1b \ \ Der_4 \ \ Int_4 (\\<^sub>D \)" nitpick oops (*countermodel*) - -text\\noindent{Border operator as derived from (dual) derivative.}\ -definition Br_der::"(\\\)\(\\\)" ("\\<^sub>D") where "\\<^sub>D \ \ \A. A \<^bold>\ \\<^sup>d(A)" -text\\noindent{Verify definition:}\ -lemma Br_der_def2: "\\<^sub>D \ = (\A. A \<^bold>\ \(\<^bold>\A))" unfolding Br_der_def dual_def conn by simp -text\\noindent{Verify properties:}\ -lemma BD1: "Der_1 \ \ Br_1 (\\<^sub>D \)" using Br_der_def Br_1_def CI1 MULT_def conn by auto -lemma BD2: "Der_3 \ \ Br_2 (\\<^sub>D \)" using Br_der_def Br_2_def CI3 dNOR_def unfolding conn by auto -lemma BD3: "Der_1b \ \ Der_4e \ \ Br_3 (\\<^sub>D \)" using dual_def Der_4e_def MONO_ADDIb MONO_def unfolding Br_der_def Br_3_def conn by smt - -text\\noindent{Frontier operator as derived from derivative.}\ -definition Fr_der::"(\\\)\(\\\)" ("\\<^sub>D") where "\\<^sub>D \ \ \A. (A \<^bold>\ \\<^sup>d(A)) \<^bold>\ (\(A) \<^bold>\ A)" -text\\noindent{Verify definition:}\ -lemma Fr_der_def2: "\\<^sub>D \ = (\A. (A \<^bold>\ \(A)) \<^bold>\ (\<^bold>\A \<^bold>\ \(\<^bold>\A)))" unfolding Fr_der_def dual_def conn by auto -text\\noindent{Verify properties:}\ -lemma FD1a: "Der_1a \ \ Fr_1a(\\<^sub>D \)" unfolding Fr_1a_def Fr_der_def using CI1a MULT_b_def conn by auto -lemma FD1b: "Der_1b \ \ Fr_1b(\\<^sub>D \)" unfolding Fr_1b_def Fr_der_def using CI1b MULT_a_def conn by smt -lemma FD1: "Der_1 \ \ Fr_1(\\<^sub>D \)" unfolding Fr_1_def Fr_der_def using CI1 MULT_def conn by auto -lemma FD2: "Fr_2(\\<^sub>D \)" using dual_def dual_symm unfolding Fr_2_def Fr_der_def conn by metis -lemma FD3: "Der_3 \ \ Fr_3(\\<^sub>D \)" by (simp add: NOR_def Fr_der_def conn) -lemma FD4: "Der_1 \ \ Der_4e \ \ Fr_4(\\<^sub>D \)" by (metis CD1b CD2 CD4 Cl_8_def Cl_der_def Fr_4_def Fr_der_def2 PC1 PC8 meet_def) - -text\\noindent{Note that the derivative operation cannot be obtained from interior, closure, border, nor frontier. -In this respect the derivative is a fundamental operation.}\ - -subsubsection \Infinitary conditions\ - -text\\noindent{The corresponding infinitary condition on derivative operators is inherited from the one for closure.}\ -abbreviation "Der_inf \ \ Cl_inf(\)" - -lemma CD_inf: "Der_inf \ \ Cl_inf(\\<^sub>D \)" unfolding iADDI_a_def by (smt Cl_der_def Fr_2_def Ra_restr_ex compl_def dom_compl_def2 iDM_b join_def meet_def supremum_def) -lemma ID_inf: "Der_inf \ \ Int_inf(\\<^sub>D \)" by (simp add: CD_inf dual_der2 iADDI_MULT_dual1) - -text\\noindent{This condition is indeed strong enough as to entail closure of some fixed-point predicates under infimum/supremum.}\ -lemma fp_ID_inf_closed: "Der_inf \ \ infimum_closed (fp (\\<^sub>D \))" by (metis (full_types) ID2 ID_inf Ra_restr_all dEXP_def iMULT_b_def inf_char) -lemma fp_CD_sup_closed: "Der_inf \ \ supremum_closed (fp (\\<^sub>D \))" by (metis (full_types) CD_inf Cl_der_def Ra_restr_ex iADDI_a_def join_def supremum_def) - - -subsubsection \Coherence conditions\ -text\\noindent{We finish this section by introducing the `coherence' operator (Cantor's `Koherenz') as discussed -by Zarycki in \<^cite>\Zarycki2\. As happens with the derivative operator, the coherence operator cannot -be derived from interior, closure, border, nor frontier.}\ - -definition "Kh_1 \ \ ADDI_b \" -definition "Kh_2 \ \ dEXP \" -definition "Kh_3 \ \ \A. \(\\<^sup>d A) \<^bold>\ \\<^sup>d(\ A)" - -lemma PK1: "Kh_1 \ \ MONO \" using ADDI_b_def Kh_1_def MONO_ADDIb by auto -lemma PK2: "Kh_1 \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" using MULT_a_def MONO_MULTa PK1 by auto -lemma PK3: "Kh_2 \ \ \ \<^bold>\ \<^bold>\ \<^bold>\" using Kh_2_def dEXP_def unfolding conn by auto -lemma PK4: "Kh_1 \ \ Kh_3 \ \ \ \<^bold>\ \<^bold>\ \<^bold>\" using MONO_def PK1 dual_def unfolding Kh_3_def conn by metis -lemma PK5: "Kh_2 \ \ \A. \(\<^bold>\A) \<^bold>\ \<^bold>\(\ A)" unfolding Kh_2_def dEXP_def conn by metis -lemma PK6: "Kh_1 \ \ Kh_2 \ \ \A B. \(A \<^bold>\ B) \<^bold>\ (\ A) \<^bold>\ (\ B)" unfolding conn by (metis (no_types, lifting) Kh_2_def dEXP_def MONO_def PK1) -lemma PK7: "Kh_3 \ \ \A. \(\(\<^bold>\(\ A))) \<^bold>\ \(\<^bold>\(\(\\<^sup>d A)))" proof - - assume kh3: "Kh_3 \" - { fix A - from kh3 have "\(\\<^sup>d A) = \\<^sup>d(\ A)" using Kh_3_def by auto - hence "\(\<^bold>\(\(\\<^sup>d A))) \<^bold>\ \(\(\<^bold>\(\ A)))" unfolding dual_def conn by simp - } thus ?thesis by simp -qed -lemma PK8: "Kh_3 \ \ \A. \(\<^bold>\(\(\ A))) \<^bold>\ \\<^sup>d(\(\<^bold>\(\ A)))" proof - - assume kh3: "Kh_3 \" - { fix A - from kh3 have aux: "\\<^sup>d(\ A) \<^bold>\ \(\\<^sup>d A)" using Kh_3_def by simp - have "let A=\<^bold>\(\ A) in (\\<^sup>d(\ A) \<^bold>\ \(\\<^sup>d A))" using Kh_3_def kh3 by auto - hence "\(\<^bold>\(\(\ A))) \<^bold>\ \\<^sup>d(\(\<^bold>\(\ A)))" unfolding dual_def conn by simp - } thus ?thesis by simp -qed - -text\\noindent{Coherence operator as derived from derivative (requires conditions Der-2 and Der-4).}\ -definition Kh_der::"(\\\)\(\\\)" ("\\<^sub>D") where "\\<^sub>D \ \ \A. A \<^bold>\ (\ A)" -text\\noindent{Verify properties:}\ -lemma KD1: "Der_1 \ \ Kh_1 (\\<^sub>D \)" using PC1 PD3 PK2 ADDI_def Kh_der_def unfolding conn by (metis (full_types)) -lemma KD2: "Kh_2 (\\<^sub>D \)" by (simp add: Kh_2_def dEXP_def Kh_der_def conn) -lemma KD3: "Der_1 \ \ Der_2 \ \ Der_4 \ \ Kh_3 (\\<^sub>D \)" proof - - assume der1: "Der_1 \" and der2: "Der_2 \" and der4: "Der_4 \" - { fix A - from der1 have aux1: "let M=A ; N=(\\<^sup>d A) in \(M \<^bold>\ N) \<^bold>\ (\ M \<^bold>\ \ N)" unfolding ADDI_def by simp - from der1 der2 der4 have aux2: "(\\<^sup>d A) \<^bold>\ (\ A) \<^bold>\ \\<^sup>d(A \<^bold>\ \ A)" using PD12 by simp - have "(\\<^sub>D \)((\\<^sub>D \)\<^sup>d A) = (\<^bold>\(\<^bold>\A \<^bold>\ \ (\<^bold>\A)) \<^bold>\ \ (\<^bold>\(\<^bold>\A \<^bold>\ \ (\<^bold>\A))))" unfolding Kh_der_def dual_def by simp - also have "... = (A \<^bold>\ \\<^sup>d A) \<^bold>\ \(A \<^bold>\ \\<^sup>d A)" unfolding dual_def conn by simp - also from aux1 have "... = (A \<^bold>\ \\<^sup>d A) \<^bold>\ (\ A \<^bold>\ \(\\<^sup>d A))" unfolding conn by meson - also have "... = (A \<^bold>\ \ A) \<^bold>\ \\<^sup>d A" using PD6 PD8 der1 der2 der4 unfolding conn by blast - also have "... = (A \<^bold>\ \ A) \<^bold>\ (\\<^sup>d A \<^bold>\ \ A)" using PD6 der1 der2 unfolding conn by blast - also have "... = (A \<^bold>\ \\<^sup>d A) \<^bold>\ (\ A)" unfolding conn by auto - also from aux2 have "... = (A \<^bold>\ \ A) \<^bold>\ \\<^sup>d(A \<^bold>\ \ A)" unfolding dual_def conn by meson - also have "... = \<^bold>\(\<^bold>\(A \<^bold>\ \ A) \<^bold>\ \ (\<^bold>\(A \<^bold>\ \ A)))" unfolding dual_def conn by simp - also have "... = (\\<^sub>D \)\<^sup>d((\\<^sub>D \) A)" unfolding Kh_der_def dual_def by simp - finally have "(\\<^sub>D \)((\\<^sub>D \)\<^sup>d A) \<^bold>\ (\\<^sub>D \)\<^sup>d((\\<^sub>D \) A)" by simp - } thus ?thesis unfolding Kh_3_def by simp -qed - -end diff --git a/thys/Topological_Semantics/topo_strict_implication.thy b/thys/Topological_Semantics/topo_strict_implication.thy deleted file mode 100644 --- a/thys/Topological_Semantics/topo_strict_implication.thy +++ /dev/null @@ -1,115 +0,0 @@ -theory topo_strict_implication - imports topo_frontier_algebra -begin -nitpick_params[assms=true, user_axioms=true, show_all, expect=genuine, format=3] (*default Nitpick settings*) - -section \Strict implication\ - -text\\noindent{We can also employ the closure and interior operations to define different sorts of implications. -In this section we preliminary explore a sort of strict implication and check some relevant properties.}\ - -text\\noindent{A 'strict' implication should entail the classical one. Hence we define it using the interior operator.}\ -definition imp_I::"\\\\\" (infix "\<^bold>\" 51) where "\ \<^bold>\ \ \ \(\ \<^bold>\ \)" -abbreviation imp_I'::"\\\\\" (infix "\<^bold>\" 51) where "\ \<^bold>\ \ \ \ \<^bold>\ \" -declare imp_I_def[conn] - -lemma imp_rel: "\a b. a \<^bold>\ b \<^bold>\ a \<^bold>\ b" by (simp add: Int_fr_def conn) - -text\\noindent{Under appropriate conditions this implication satisfies a weak variant of the deduction theorem (DT),}\ -lemma DTw1: "\a b. a \<^bold>\ b \<^bold>\ \<^bold>\ \ a \<^bold>\ b" by (simp add: Int_fr_def conn) -lemma DTw2: "Fr_2 \ \ Fr_3 \ \ \a b. a \<^bold>\ b \ a \<^bold>\ b \<^bold>\ \<^bold>\" using IF3 dNOR_def unfolding conn by auto -text\\noindent{and a variant of modus ponens and modus tollens resp.}\ -lemma MP: "\a b. a \<^bold>\ (a \<^bold>\ b) \<^bold>\ b" by (simp add: Int_fr_def conn) -lemma MT: "\a b. (a \<^bold>\ b) \<^bold>\ \<^bold>\b \<^bold>\ \<^bold>\a" using MP conn by auto - -text\\noindent{However the full DT (actually right-to-left: implication introduction) is not valid.}\ -lemma DT1: "\a b X. X \<^bold>\ a \<^bold>\ b \ X \<^bold>\ a \<^bold>\ b" by (simp add: Int_fr_def conn) -lemma DT2: "\ \ \ \a b X. X \<^bold>\ a \<^bold>\ b \ X \<^bold>\ a \<^bold>\ b" nitpick oops (*counterexample*) - -text\\noindent{This implication has thus a sort of 'relevant' behaviour. For instance, the following are not valid:}\ -lemma "\ \ \ \a b. (a \<^bold>\ (b \<^bold>\ a)) \<^bold>\ \<^bold>\" nitpick oops (*counterexample*) -lemma "\ \ \ \a b. (a \<^bold>\ ((a \<^bold>\ b) \<^bold>\ b)) \<^bold>\ \<^bold>\" nitpick oops (*counterexample*) -lemma "\ \ \ \a b c. (a \<^bold>\ b) \<^bold>\ (b \<^bold>\ c) \<^bold>\ \<^bold>\" nitpick oops (*counterexample*) -lemma "\ \ \ \a b. ((a \<^bold>\ b) \<^bold>\ a) \<^bold>\ a \<^bold>\ \<^bold>\" nitpick oops (*counterexample*) - -text\\noindent{In contrast the properties below are valid for appropriate conditions.}\ -lemma iprop0: "Fr_2 \ \ Fr_3 \ \ \a. a \<^bold>\ a \<^bold>\ \<^bold>\" using DTw2 pI2 by fastforce -lemma iprop1: "Fr_2 \ \ Fr_3 \ \ \a b. a \<^bold>\ (a \<^bold>\ b) \<^bold>\ b \<^bold>\ \<^bold>\" using DTw2 pI2 unfolding conn by fastforce -lemma iprop2: "Fr_2 \ \ Fr_3 \ \ \a b. a \<^bold>\ (b \<^bold>\ b) \<^bold>\ \<^bold>\" using DTw2 pI2 unfolding conn by fastforce -lemma iprop3: "Fr_2 \ \ Fr_3 \ \ \a b. ((a \<^bold>\ a) \<^bold>\ b) \<^bold>\ b \<^bold>\ \<^bold>\" using DTw2 pI2 unfolding conn by fastforce -lemma iprop4: "Fr_2 \ \ Fr_3 \ \ \a b. (a \<^bold>\ b) \<^bold>\ a \<^bold>\ \<^bold>\" using DTw2 pI2 unfolding conn by fastforce -lemma iprop5: "Fr_2 \ \ Fr_3 \ \ \a b. a \<^bold>\ (a \<^bold>\ b) \<^bold>\ \<^bold>\" using DTw2 pI2 unfolding conn by fastforce -lemma iprop6a: "Fr_2 \ \ Fr_3 \ \ \a b c. (a \<^bold>\ (b \<^bold>\ c)) \<^bold>\ ((a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c)) \<^bold>\ \<^bold>\" using DTw2 pI2 unfolding conn by fastforce -lemma iprop6b: "Fr_2 \ \ Fr_3 \ \ \a b c. (a \<^bold>\ (b \<^bold>\ c)) \<^bold>\ ((a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c)) \<^bold>\ \<^bold>\" using DTw2 unfolding conn by fastforce - -lemma iprop7': "Fr_1 \ \ \a b c. (a \<^bold>\ b) \<^bold>\ (b \<^bold>\ c) \<^bold>\ (a \<^bold>\ c)" proof - - assume fr1: "Fr_1 \" - { fix a b c - have "(a \<^bold>\ b) \<^bold>\ (b \<^bold>\ c) \<^bold>\ (a \<^bold>\ c)" unfolding conn by simp - hence "\((a \<^bold>\ b) \<^bold>\ (b \<^bold>\ c)) \<^bold>\ \(a \<^bold>\ c)" using MONO_def PF1 fr1 monI by simp - moreover from fr1 have "let A=(a \<^bold>\ b); B=(b \<^bold>\ c) in \(A \<^bold>\ B) \<^bold>\ \ A \<^bold>\ \ B" using IF1 MULT_def by simp - ultimately have "\(a \<^bold>\ b) \<^bold>\ \(b \<^bold>\ c) \<^bold>\ \(a \<^bold>\ c)" unfolding conn by simp - } thus ?thesis unfolding conn by blast -qed -lemma iprop7: "Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ \a b c. ((a \<^bold>\ b) \<^bold>\ (b \<^bold>\ c)) \<^bold>\ (a \<^bold>\ c) \<^bold>\ \<^bold>\" by (simp add: DTw2 iprop7') - -lemma iprop8a': "Fr_1 \ \ \a b c. (a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c) \<^bold>\ a \<^bold>\ (b \<^bold>\ c)" proof - - assume fr1: "Fr_1 \" - { fix a b c - have "(a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c) \<^bold>\ (a \<^bold>\ (b \<^bold>\ c))" unfolding conn by simp - hence "\((a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c)) \<^bold>\ \(a \<^bold>\ (b \<^bold>\ c))" using MONO_def PF1 fr1 monI by simp - moreover from fr1 have "let A=(a \<^bold>\ b); B=(a \<^bold>\ c) in \(A \<^bold>\ B) \<^bold>\ \ A \<^bold>\ \ B" using IF1 MULT_def by simp - ultimately have "\(a \<^bold>\ b) \<^bold>\ \(a \<^bold>\ c) \<^bold>\ \(a \<^bold>\ (b \<^bold>\ c))" unfolding conn by simp - } thus ?thesis unfolding conn by simp -qed -lemma iprop8b': "Fr_1b \ \ \a b c. (a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c) \<^bold>\ a \<^bold>\ (b \<^bold>\ c)" by (smt MONO_def monI conn) -lemma iprop8a: "Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ \a b c. ((a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c)) \<^bold>\ (a \<^bold>\ (b \<^bold>\ c)) \<^bold>\ \<^bold>\" by (simp add: DTw2 iprop8a') -lemma iprop8b: "Fr_1b \ \ Fr_2 \ \ Fr_3 \ \ \a b c. ((a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c)) \<^bold>\ (a \<^bold>\ (b \<^bold>\ c)) \<^bold>\ \<^bold>\" by (simp add: DTw2 iprop8b') - -lemma iprop9a': "Fr_1 \ \ \a b c. ((a \<^bold>\ b) \<^bold>\ (c \<^bold>\ b)) \<^bold>\ ((a \<^bold>\ c) \<^bold>\ b)" proof - - assume fr1: "Fr_1 \" - { fix a b c - have "(a \<^bold>\ b) \<^bold>\ (c \<^bold>\ b) \<^bold>\ (a \<^bold>\ c) \<^bold>\ b" unfolding conn by simp - hence "\((a \<^bold>\ b) \<^bold>\ (c \<^bold>\ b)) \<^bold>\ \((a \<^bold>\ c) \<^bold>\ b)" using MONO_def PF1 fr1 monI by simp - moreover from fr1 have "let A=(a \<^bold>\ b); B=(c \<^bold>\ b) in \(A \<^bold>\ B) \<^bold>\ \ A \<^bold>\ \ B" using IF1 MULT_def by simp - ultimately have "\(a \<^bold>\ b) \<^bold>\ \(c \<^bold>\ b) \<^bold>\ \((a \<^bold>\ c) \<^bold>\ b)" unfolding conn by simp - } thus ?thesis unfolding conn by simp -qed -lemma iprop9b': "Fr_1b \ \ \a b c. ((a \<^bold>\ b) \<^bold>\ (c \<^bold>\ b)) \<^bold>\ ((a \<^bold>\ c) \<^bold>\ b)" by (smt MONO_def monI conn) -lemma iprop9a: "Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ \a b c. ((a \<^bold>\ b) \<^bold>\ (c \<^bold>\ b)) \<^bold>\ ((a \<^bold>\ c) \<^bold>\ b) \<^bold>\ \<^bold>\" by (simp add: DTw2 iprop9a') -lemma iprop9b: "Fr_1b \ \ Fr_2 \ \ Fr_3 \ \ \a b c. ((a \<^bold>\ b) \<^bold>\ (c \<^bold>\ b)) \<^bold>\ ((a \<^bold>\ c) \<^bold>\ b) \<^bold>\ \<^bold>\" by (simp add: DTw2 iprop9b') - -lemma iprop10': "Fr_1 \ \ Fr_2 \ \ Fr_4 \ \ \a b c. a \<^bold>\ (b \<^bold>\ c) \<^bold>\ (a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c)" proof - - assume fr1: "Fr_1 \" and fr2: "Fr_2 \" and fr4: "Fr_4 \" - { fix a b c - have "a \<^bold>\ (b \<^bold>\ c) \<^bold>\ (a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c)" unfolding conn by simp - hence "a \<^bold>\ (b \<^bold>\ c) \<^bold>\ (a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c)" using Int_fr_def conn by auto - hence "\(a \<^bold>\ (b \<^bold>\ c)) \<^bold>\ \((a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c))" using MONO_def PF1 fr1 monI by simp - moreover from fr1 have "let A=(a \<^bold>\ b); B=(a \<^bold>\ c) in \(A \<^bold>\ B) \<^bold>\ \ A \<^bold>\ \ B" using IF1 Int_7_def PI7 by auto - ultimately have "\(a \<^bold>\ (b \<^bold>\ c)) \<^bold>\ \(a \<^bold>\ b) \<^bold>\ \(a \<^bold>\ c)" by (metis (full_types)) - hence "\(\(a \<^bold>\ (b \<^bold>\ c))) \<^bold>\ \(\(a \<^bold>\ b) \<^bold>\ \(a \<^bold>\ c))" using MONO_def PF1 fr1 monI by simp - hence "\(a \<^bold>\ (b \<^bold>\ c)) \<^bold>\ \(\(a \<^bold>\ b) \<^bold>\ \(a \<^bold>\ c))" using Int_Open PF1 fr1 fr2 fr4 by blast - hence "(a \<^bold>\ (b \<^bold>\ c)) \<^bold>\ (a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c)" using Int_Open PF1 fr1 fr2 fr4 unfolding conn by blast - } thus ?thesis by blast -qed -lemma iprop10: "Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ \a b c. (a \<^bold>\ (b \<^bold>\ c)) \<^bold>\ ((a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c)) \<^bold>\ \<^bold>\" by (simp add: DTw2 iprop10') -lemma "\ \ \ \a b c. a \<^bold>\ (b \<^bold>\ c) \<^bold>\ (a \<^bold>\ b) \<^bold>\ (a \<^bold>\ c)" nitpick oops (*counterexample*) - -lemma iprop11a': "Fr_1 \ \ \a b. (a \<^bold>\ (a \<^bold>\ b)) \<^bold>\ (a \<^bold>\ b)" by (smt MONO_def PF1 imp_rel monI conn) -lemma iprop11b': "\ \ \ \a b. (a \<^bold>\ (a \<^bold>\ b)) \<^bold>\ (a \<^bold>\ b)" unfolding PF1 by (metis Int_Open MONO_def imp_I_def impl_def monI) -lemma iprop11a: "Fr_1 \ \ Fr_2 \ \ Fr_3 \ \ \a b. (a \<^bold>\ (a \<^bold>\ b)) \<^bold>\ (a \<^bold>\ b) \<^bold>\ \<^bold>\" using DTw2 iprop11a' by blast -lemma iprop11b: "\ \ \ \a b. (a \<^bold>\ (a \<^bold>\ b)) \<^bold>\ (a \<^bold>\ b) \<^bold>\ \<^bold>\" using DTw2 iprop11b' by blast - -text\\noindent{In particular, 'strenghening the antecedent' is valid only under certain conditions:}\ -lemma SA':"Fr_1b \ \ \a b c. a \<^bold>\ c \<^bold>\ (a \<^bold>\ b) \<^bold>\ c" by (smt MONO_def monI conn) -lemma SA: "Fr_1b \ \ Fr_2 \ \ Fr_3 \ \ \a b c. (a \<^bold>\ c) \<^bold>\ ((a \<^bold>\ b) \<^bold>\ c) \<^bold>\ \<^bold>\" using SA' using DTw2 by fastforce -lemma "Fr_1a \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ \a b c. a \<^bold>\ c \<^bold>\ (a \<^bold>\ b) \<^bold>\ c" nitpick oops (*counterexample*) -lemma "Fr_1a \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ \a b c. (a \<^bold>\ c) \<^bold>\ ((a \<^bold>\ b) \<^bold>\ c) \<^bold>\ \<^bold>\" nitpick oops (*counterexample*) - -text\\noindent{Similarly, 'weakening the consequent' is valid only under certain conditions.}\ -lemma WC':"Fr_1b \ \ \a b c. a \<^bold>\ c \<^bold>\ a \<^bold>\ (c \<^bold>\ b)" by (smt MONO_def monI conn) -lemma WC: "Fr_1b \ \ Fr_2 \ \ Fr_3 \ \ \a b c. (a \<^bold>\ c) \<^bold>\ (a \<^bold>\ (c \<^bold>\ b)) \<^bold>\ \<^bold>\" using DTw2 WC' by fastforce -lemma "Fr_1a \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ \a b c. a \<^bold>\ c \<^bold>\ a \<^bold>\ (c \<^bold>\ b)" nitpick oops (*counterexample*) -lemma "Fr_1a \ \ Fr_2 \ \ Fr_3 \ \ Fr_4 \ \ \a b c. (a \<^bold>\ c) \<^bold>\ (a \<^bold>\ (c \<^bold>\ b)) \<^bold>\ \<^bold>\" nitpick oops (*counterexample*) - -end