X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Ftopology%2Figft.ma;h=6b2b73f881f5aa4ef29989ee24a763a212c459db;hb=3cb42e0873c101c6c5a8b9967d765b5135882685;hp=3c97b9c25000735971c09018c1a1f4c42a6e86a3;hpb=a8aaa095aad443c8eca8f64e3f22f54615e8dd9b;p=helm.git diff --git a/helm/software/matita/nlibrary/topology/igft.ma b/helm/software/matita/nlibrary/topology/igft.ma index 3c97b9c25..6b2b73f88 100644 --- a/helm/software/matita/nlibrary/topology/igft.ma +++ b/helm/software/matita/nlibrary/topology/igft.ma @@ -1,4 +1,4 @@ -(*DOCBEGIN +(*D Matita Tutorial: inductively generated formal topologies ======================================================== @@ -12,12 +12,27 @@ the formalization of the paper by S.Berardi and S. Valentini. The tutorial is by Enrico Tassi. +The tutorial spends a considerable amount of effort in defining +notations that resemble the ones used in the original paper. We believe +this a important part of every formalization, not only for the estetic +point of view, but also from the practical point of view. Being +consistent allows to follow the paper in a pedantic way, and hopefully +to make the formalization (at least the definitions and proved +statements) readable to the author of the paper. + Orientering ----------- + ? : A +apply (f : A -> B): -------------------- + (f ? ) : B + + f: A1 -> ... -> An -> B ?1: A1 ... ?n: An +apply (f : A -> B): ------------------------------------------------ + apply f == f \ldots == f ? ... ? : B TODO -buttons, PG-interaction-model, sequent window, script window +buttons, PG-interaction-model, sequent window, script window, ncheck The library, inclusion of `sets/sets.ma`, notation defined: Ω^A. Symbols (see menu: View ▹ TeX/UTF-8 Table): @@ -40,13 +55,14 @@ some notation attached to them: - A ∪ B `A \cup B` - A ∩ B `A \cap B` +- A ≬ B `A \between B` - x ∈ A `x \in A` - Ω^A, that is the type of the subsets of A, `\Omega ^ A` The `include` command tells Matita to load a part of the library, in particular the part that we will use can be loaded as follows: -DOCEND*) +D*) include "sets/sets.ma". @@ -65,7 +81,7 @@ nlemma subseteq_intersection_r: ∀A.∀U,V,W:Ω^A.W ⊆ U → W ⊆ V → W ⊆ nqed. (*UNHIDE*) -(*DOCBEGIN +(*D Some basic results that we will use are also part of the sets library: @@ -77,7 +93,7 @@ Defining Axiom set records, projections, types of projections.. -DOCEND*) +D*) nrecord Ax : Type[1] ≝ { S :> setoid; @@ -85,7 +101,7 @@ nrecord Ax : Type[1] ≝ { C : ∀a:S. I a → Ω ^ S }. -(*DOCBEGIN +(*D Note that the field `S` was declared with `:>` instead of a simple `:`. This declares the `S` projection to be a coercion. A coercion is @@ -121,12 +137,12 @@ Something that is not still satisfactory, in that the dependent type of `I` and `C` are abstracted over the Axiom set. To obtain the precise type of a term, you can use the `ncheck` command as follows. -DOCEND*) +D*) (* ncheck I. *) (* ncheck C. *) -(*DOCBEGIN +(*D One would like to write `I a` and not `I A a` under a context where `A` is an axiom set and `a` has type `S A` (or thanks to the coercion @@ -136,11 +152,11 @@ infer. Matita performs some sort of type inference, thus writing `I ? a` is enough: since the second argument of `I` is typed by the first one, the first one can be inferred just computing the type of `a`. -DOCEND*) +D*) (* ncheck (∀A:Ax.∀a:A.I ? a). *) -(*DOCBEGIN +(*D This is still not completely satisfactory, since you have always type `?`; to fix this minor issue we have to introduce the notational @@ -163,12 +179,12 @@ keyboard and what is displayed in the sequent window) and the content level is defined with the `notation` command. When followed by `>`, it defines an input (only) notation. -DOCEND*) +D*) notation > "𝐈 term 90 a" non associative with precedence 70 for @{ 'I $a }. notation > "𝐂 term 90 a term 90 i" non associative with precedence 70 for @{ 'C $a $i }. -(*DOCBEGIN +(*D The forst notation defines the writing `𝐈 a` where `a` is a generic term of precedence 90, the maximum one. This high precedence forces @@ -188,12 +204,12 @@ new content element to which a term `$a` is passed. Content elements have to be interpreted, and possibly multiple, incompatible, interpretations can be defined. -DOCEND*) +D*) interpretation "I" 'I a = (I ? a). interpretation "C" 'C a i = (C ? a i). -(*DOCBEGIN +(*D The `interpretation` command allows to define the mapping between the content level and the terms level. Here we associate the `I` and @@ -204,12 +220,12 @@ Interpretation are bi-directional, thus when displaying a term like `C _ a i`, the system looks for a presentation for the content element `'C a i`. -DOCEND*) +D*) notation < "𝐈 \sub( ❨a❩ )" non associative with precedence 70 for @{ 'I $a }. notation < "𝐂 \sub( ❨a,\emsp i❩ )" non associative with precedence 70 for @{ 'C $a $i }. -(*DOCBEGIN +(*D For output purposes we can define more complex notations, for example we can put bold parenteses around the arguments of `𝐈` and `𝐂`, decreasing @@ -219,38 +235,183 @@ as subscript), separating them with a comma followed by a little space. The first (technical) definition -------------------------------- +Before defining the cover relation as an inductive predicate, one +has to notice that the infinity rule uses, in its hypotheses, the +cover relation between two subsets, while the inductive predicate +we are going to define relates an element and a subset. + +An option would be to unfold the definition of cover between subsets, +but we prefer to define the abstract notion of cover between subsets +(so that we can attach a (ambiguous) notation to it). + +Anyway, to ease the understaing of the definition of the cover relation +between subsets, we first define the inductive predicate unfolding the +definition, and we later refine it with. + +D*) + +ninductive xcover (A : Ax) (U : Ω^A) : A → CProp[0] ≝ +| xcreflexivity : ∀a:A. a ∈ U → xcover A U a +| xcinfinity : ∀a:A.∀i:𝐈 a. (∀y.y ∈ 𝐂 a i → xcover A U y) → xcover A U a. + +(*D + +We defined the xcover (x will be removed in the final version of the +definition) as an inductive predicate. The arity of the inductive +predicate has to be carefully analyzed: + +> (A : Ax) (U : Ω^A) : A → CProp[0] + +The syntax separates with `:` abstractions that are fixed for every +constructor (introduction rule) and abstractions that can change. In that +case the parameter `U` is abstracted once and forall in front of every +constructor, and every occurrence of the inductive predicate is applied to +`U` in a consistent way. Arguments abstracted on the right of `:` are not +constant, for example the xcinfinity constructor introduces `a ◃ U`, +but under the assumption that (for every y) `y ◃ U`. In that rule, the left +had side of the predicate changes, thus it has to be abstrated (in the arity +of the inductive predicate) on the right of `:`. + +D*) + +(* ncheck xcreflexivity. *) + +(*D + +We want now to abstract out `(∀y.y ∈ 𝐂 a i → xcover A U y)` and define +a notion `cover_set` to which a notation `𝐂 a i ◃ U` can be attached. +This notion has to be abstracted over the cover relation (whose +type is the arity of the inductive `xcover` predicate just defined). -DOCEND*) +Then it has to be abstracted over the arguments of that cover relation, +i.e. the axiom set and the set U, and the subset (in that case `𝐂 a i`) +sitting on the left hand side of `◃`. + +D*) ndefinition cover_set : - ∀c: ∀A:Ax.Ω^A → A → CProp[0]. ∀A:Ax.∀C,U:Ω^A. CProp[0] + ∀cover: ∀A:Ax.Ω^A → A → CProp[0]. ∀A:Ax.∀C,U:Ω^A. CProp[0] ≝ - λc: ∀A:Ax.Ω^A → A → CProp[0]. λA,C,U.∀y.y ∈ C → c A U y. + λcover. λA, C,U. ∀y.y ∈ C → cover A U y. -ndefinition cover_set_interactive : - ∀c: ∀A:Ax.Ω^A → A → CProp[0]. ∀A:Ax.∀C,U:Ω^A. CProp[0]. -#cover; #A; #C; #U; napply (∀y:A.y ∈ C → ?); napply cover; -##[ napply A; -##| napply U; -##| napply y; -##] -nqed. +(*D + +The `ndefinition` command takes a name, a type and body (of that type). +The type can be omitted, and in that case it is inferred by the system. +If the type is given, the system uses it to infer implicit arguments +of the body. In that case all types are left implicit in the body. + +We now define the notation `a ◃ b`. Here the keywork `hvbox` +and `break` tell the system how to wrap text when it does not +fit the screen (they can be safely ignore for the scope of +this tutorial). we also add an interpretation for that notation, +where the (abstracted) cover relation is implicit. The system +will not be able to infer it from the other arguments `C` and `U` +and will thus prompt the user for it. This is also why we named this +interpretation `covers set temp`: we will later define another +interpretation in which the cover relation is the one we are going to +define. + +D*) -(* a \ltri b *) notation "hvbox(a break ◃ b)" non associative with precedence 45 for @{ 'covers $a $b }. interpretation "covers set temp" 'covers C U = (cover_set ?? C U). +(*D + +We can now define the cover relation using the `◃` notation for +the premise of infinity. + +D*) + ninductive cover (A : Ax) (U : Ω^A) : A → CProp[0] ≝ -| creflexivity : ∀a:A. a ∈ U → cover ? U a -| cinfinity : ∀a:A.∀i:𝐈 a. 𝐂 a i ◃ U → cover ? U a. +| creflexivity : ∀a. a ∈ U → cover ? U a +| cinfinity : ∀a. ∀i. 𝐂 a i ◃ U → cover ? U a. +(** screenshot "cover". *) napply cover; nqed. +(*D + +Note that the system accepts the definition +but prompts the user for the relation the `cover_set` notion is +abstracted on. + + + +The orizontal line separates the hypotheses from the conclusion. +The `napply cover` command tells the system that the relation +it is looking for is exactly our first context entry (i.e. the inductive +predicate we are defining, up to α-conversion); while the `nqed` command +ends a definition or proof. + +We can now define the interpretation for the cover relation between an +element and a subset fist, then between two subsets (but this time +we fixed the relation `cover_set` is abstracted on). + +D*) + interpretation "covers" 'covers a U = (cover ? U a). -(* interpretation "covers set" 'covers a U = (cover_set cover ? a U). *) +interpretation "covers set" 'covers a U = (cover_set cover ? a U). + +(*D + +We will proceed similarly for the fish relation, but before going +on it is better to give a short introduction to the proof mode of Matita. +We define again the `cover_set` term, but this time we will build +its body interactively. In λ-calculus Matita is based on, CIC, proofs +and terms share the same syntax, and it thus possible to use the +commands devoted to build proof term to build regular definitions. + +D*) + +ndefinition xcover_set : + ∀c: ∀A:Ax.Ω^A → A → CProp[0]. ∀A:Ax.∀C,U:Ω^A. CProp[0]. + (** screenshot "xcover-set-1". *) +#cover; #A; #C; #U; (** screenshot "xcover-set-2". *) +napply (∀y:A.y ∈ C → ?); (** screenshot "xcover-set-3". *) +napply cover; (** screenshot "xcover-set-4". *) +##[ napply A; +##| napply U; +##| napply y; +##] +nqed. + +(*D[xcover-set-1] +The system asks for a proof of the full statement, in an empty context. + +The `#` command is the ∀-introduction rule, it gives a name to an +assumption putting it in the context, and generates a λ-abstraction +in the proof term. + +D[xcover-set-2] +We have now to provide a proposition, and we exhibit it. We left +a part of it implicit; since the system cannot infer it it will +ask it later. Note that the type of `∀y:A.y ∈ C → ?` is a proposition +whenever `?` is. + +D[xcover-set-3] +The proposition we want to provide is an application of the +cover relation we have abstracted in the context. The command +`napply`, if the given term has not the expected type (in that +case it is a product versus a proposition) it applies it to as many +implicit arguments as necessary (in that case `? ? ?`). + +D[xcover-set-4] +The system will now ask in turn the three implicit arguments +passed to cover. The syntax `##[` allows to start a branching +to tackle every sub proof individually, otherwise every command +is applied to every subrpoof. The command `##|` switches to the next +subproof and `##]` ends the branching. +D*) + +(*D +The definition of fish works exactly the same way as for cover, except +that it is defined as a coinductive proposition. +D*) ndefinition fish_set ≝ λf:∀A:Ax.Ω^A → A → CProp[0]. λA,U,V. @@ -270,22 +431,119 @@ nqed. interpretation "fish set" 'fish A U = (fish_set fish ? U A). interpretation "fish" 'fish a U = (fish ? U a). +(*D + +Matita is able to generate elimination rules for inductive types, +but not introduction rules for the coinductive case. + +D*) + +(* ncheck cover_rect_CProp0. *) + +(*D + +We thus have to define the introduction rule for fish by corecursion. +Here we again use the proof mode of Matita to exhibit the body of the +corecursive function. + +D*) + nlet corec fish_rec (A:Ax) (U: Ω^A) (P: Ω^A) (H1: P ⊆ U) - (H2: ∀a:A. a ∈ P → ∀j: 𝐈 a. 𝐂 a j ≬ P): - ∀a:A. ∀p: a ∈ P. a ⋉ U ≝ ?. -#a; #p; napply cfish; (** screenshot "def-fish-rec". *) -##[ napply H1; napply p; -##| #i; ncases (H2 a p i); #x; *; #xC; #xP; @; ##[napply x] - @; ##[ napply xC ] napply (fish_rec ? U P); nassumption; + (H2: ∀a:A. a ∈ P → ∀j: 𝐈 a. 𝐂 a j ≬ P): ∀a:A. ∀p: a ∈ P. a ⋉ U ≝ ?. + (** screenshot "def-fish-rec-1". *) +#a; #p; napply cfish; (** screenshot "def-fish-rec-2". *) +##[ nchange in H1 with (∀b.b∈P → b∈U); (** screenshot "def-fish-rec-2-1". *) + napply H1; (** screenshot "def-fish-rec-3". *) + nassumption; +##| #i; ncases (H2 a p i); (** screenshot "def-fish-rec-5". *) + #x; *; #xC; #xP; (** screenshot "def-fish-rec-5-1". *) + @; (** screenshot "def-fish-rec-6". *) + ##[ napply x + ##| @; (** screenshot "def-fish-rec-7". *) + ##[ napply xC; + ##| napply (fish_rec ? U P); (** screenshot "def-fish-rec-9". *) + nassumption; + ##] + ##] ##] nqed. -notation "◃U" non associative with precedence 55 -for @{ 'coverage $U }. +(*D +D[def-fish-rec-1] +Note the first item of the context, it is the corecursive function we are +defining. This item allows to perform the recursive call, but we will be +allowed to do such call only after having generated a constructor of +the fish coinductive type. + +We introduce `a` and `p`, and then return the fish constructor `cfish`. +Since the constructor accepts two arguments, the system asks for them. + +D[def-fish-rec-2] +The first one is a proof that `a ∈ U`. This can be proved using `H1` and `p`. +With the `nchange` tactic we change `H1` into an equivalent form (this step +can be skipped, since the systeem would be able to unfold the definition +of inclusion by itself) + +D[def-fish-rec-2-1] +It is now clear that `H1` can be applied. Again `napply` adds two +implicit arguments to `H1 ? ?`, obtaining a proof of `? ∈ U` given a proof +that `? ∈ P`. Thanks to unification, the system understands that `?` is actually +`a`, and it asks a proof that `a ∈ P`. + +D[def-fish-rec-3] +The `nassumption` tactic looks for the required proof in the context, and in +that cases finds it in the last context position. + +We move now to the second branch of the proof, corresponding to the second +argument of the `cfish` constructor. + +We introduce `i` and then we destruct `H2 a p i`, that being a proof +of an overlap predicate, give as an element and a proof that it is +both in `𝐂 a i` and `P`. + +D[def-fish-rec-5] +We then introduce `x`, break the conjunction (the `*;` command is the +equivalent of `ncases` but operates on the first hypothesis that can +be introduced. We then introduce the two sides of the conjuction. + +D[def-fish-rec-5-1] +The goal is now the existence of an a point in `𝐂 a i` fished by `U`. +We thus need to use the introduction rulle for the existential quantifier. +In CIC it is a defined notion, that is an inductive type with just +one constructor (one introduction rule) holding the witness and the proof +that the witness satisfies a proposition. + +> ncheck Ex. + +Instead of trying to remember the name of the constructor, that should +be used as the argument of `napply`, we can ask the system to find by +itself the constructor name and apply it with the `@` tactic. +Note that some inductive predicates, like the disjunction, have multiple +introduction rules, and thus `@` can be followed by a number identifying +the constructor. + +D[def-fish-rec-6] +After choosing `x` as the witness, we have to prove a conjunction, +and we again apply the introduction rule for the inductively defined +predicate `∧`. + +D[def-fish-rec-7] +The left hand side of the conjunction is trivial to prove, since it +is already in the context. The right hand side needs to perform +the co-recursive call. + +D[def-fish-rec-9] +The co-recursive call needs some arguments, but all of them live +in the context. Instead of explicitly mention them, we use the +`nassumption` tactic, that simply tries to apply every context item. + +D*) ndefinition coverage : ∀A:Ax.∀U:Ω^A.Ω^A ≝ λA,U.{ a | a ◃ U }. +notation "◃U" non associative with precedence 55 for @{ 'coverage $U }. + interpretation "coverage cover" 'coverage U = (coverage ? U). ndefinition cover_equation : ∀A:Ax.∀U,X:Ω^A.CProp[0] ≝ λA,U,X. @@ -426,13 +684,12 @@ naxiom AC : ∀A,a,i,U.(∀j:𝐃 a i.∃x:Ord A.𝐝 a i j ∈ U⎽x) → (Σf. naxiom setoidification : ∀A:nAx.∀a,b:A.∀U.a=b → b ∈ U → a ∈ U. -(*DOCBEGIN +(*D Bla Bla, -
foo
-DOCEND*) +D*) alias symbol "covers" = "new covers". alias symbol "covers" = "new covers set". @@ -519,8 +776,8 @@ nelim o; ##] nqed. -(*DOCBEGIN +(*D [1]: http://upsilon.cc/~zack/research/publications/notation.pdf -*) \ No newline at end of file +*)