# HG changeset patch # User Christian Urban # Date 1274978410 -7200 # Node ID c785fff02a8fb780c86d674f2ddad6ffdbed2647 # Parent c6db12ddb60c3afde65c46d7ce5a8b59c7f3df67# Parent 4d8d9b8af76fd8489bb52fd40511a353b56d5c81 merged diff -r c6db12ddb60c -r c785fff02a8f Nominal-General/Nominal2_Eqvt.thy --- a/Nominal-General/Nominal2_Eqvt.thy Thu May 27 18:37:52 2010 +0200 +++ b/Nominal-General/Nominal2_Eqvt.thy Thu May 27 18:40:10 2010 +0200 @@ -383,7 +383,7 @@ declare [[trace_eqvt = false]] -text {* Problem: there is no raw eqvt-rule for The *} +text {* there is no raw eqvt-rule for The *} lemma "p \ (THE x. P x) = foo" apply(perm_strict_simp exclude: The) apply(perm_simp exclude: The) diff -r c6db12ddb60c -r c785fff02a8f Nominal-General/nominal_thmdecls.ML --- a/Nominal-General/nominal_thmdecls.ML Thu May 27 18:37:52 2010 +0200 +++ b/Nominal-General/nominal_thmdecls.ML Thu May 27 18:40:10 2010 +0200 @@ -47,13 +47,13 @@ val merge = Item_Net.merge); structure EqvtRawData = Generic_Data -( type T = thm Symtab.table; - val empty = Symtab.empty; +( type T = thm Termtab.table; + val empty = Termtab.empty; val extend = I; - val merge = Symtab.merge (K true)); + val merge = Termtab.merge (K true)); val eqvts = Item_Net.content o EqvtData.get; -val eqvts_raw = map snd o Symtab.dest o EqvtRawData.get; +val eqvts_raw = map snd o Termtab.dest o EqvtRawData.get; val get_eqvts_thms = eqvts o Context.Proof; val get_eqvts_raw_thms = eqvts_raw o Context.Proof; @@ -63,17 +63,17 @@ fun add_raw_thm thm = case prop_of thm of - Const ("==", _) $ _ $ Const (c, _) => EqvtRawData.map (Symtab.update (c, thm)) + Const ("==", _) $ _ $ (c as Const _) => EqvtRawData.map (Termtab.update (c, thm)) | _ => raise THM ("Theorem must be a meta-equality where the right-hand side is a constant.", 0, [thm]) fun del_raw_thm thm = case prop_of thm of - Const ("==", _) $ _ $ Const (c, _) => EqvtRawData.map (Symtab.delete c) + Const ("==", _) $ _ $ (c as Const _) => EqvtRawData.map (Termtab.delete c) | _ => raise THM ("Theorem must be a meta-equality where the right-hand side is a constant.", 0, [thm]) fun is_eqvt ctxt trm = case trm of - Const (c, _) => Symtab.defined (EqvtRawData.get (Context.Proof ctxt)) c + (c as Const _) => Termtab.defined (EqvtRawData.get (Context.Proof ctxt)) c | _ => raise TERM ("Term must be a constsnt.", [trm]) diff -r c6db12ddb60c -r c785fff02a8f Nominal/Ex/SingleLet.thy --- a/Nominal/Ex/SingleLet.thy Thu May 27 18:37:52 2010 +0200 +++ b/Nominal/Ex/SingleLet.thy Thu May 27 18:40:10 2010 +0200 @@ -74,8 +74,6 @@ thm trm_assg.fv[simplified trm_assg.supp(1-2)] *) - - end diff -r c6db12ddb60c -r c785fff02a8f Nominal/FSet.thy --- a/Nominal/FSet.thy Thu May 27 18:37:52 2010 +0200 +++ b/Nominal/FSet.thy Thu May 27 18:40:10 2010 +0200 @@ -160,10 +160,14 @@ text {* Respectfullness *} -lemma [quot_respect]: +lemma append_rsp[quot_respect]: shows "(op \ ===> op \ ===> op \) op @ op @" by auto +lemma append_rsp_unfolded: + "\x \ y; v \ w\ \ x @ v \ y @ w" + by auto + lemma [quot_respect]: shows "(op \ ===> op \ ===> op =) sub_list sub_list" by (auto simp add: sub_list_def) @@ -350,7 +354,7 @@ then show ?thesis using f i by auto qed -lemma [quot_respect]: +lemma concat_rsp[quot_respect]: shows "(list_rel op \ OOO op \ ===> op \) concat concat" proof (rule fun_relI, elim pred_compE) fix a b ba bb @@ -373,6 +377,31 @@ then show "concat a \ concat b" by simp qed + + +lemma concat_rsp_unfolded: + "\list_rel op \ a ba; ba \ bb; list_rel op \ bb b\ \ concat a \ concat b" +proof - + fix a b ba bb + assume a: "list_rel op \ a ba" + assume b: "ba \ bb" + assume c: "list_rel op \ bb b" + have "\x. (\xa\set a. x \ set xa) = (\xa\set b. x \ set xa)" proof + fix x + show "(\xa\set a. x \ set xa) = (\xa\set b. x \ set xa)" proof + assume d: "\xa\set a. x \ set xa" + show "\xa\set b. x \ set xa" by (rule concat_rsp_pre[OF a b c d]) + next + assume e: "\xa\set b. x \ set xa" + have a': "list_rel op \ ba a" by (rule list_rel_symp[OF list_eq_equivp, OF a]) + have b': "bb \ ba" by (rule equivp_symp[OF list_eq_equivp, OF b]) + have c': "list_rel op \ b bb" by (rule list_rel_symp[OF list_eq_equivp, OF c]) + show "\xa\set a. x \ set xa" by (rule concat_rsp_pre[OF c' b' a' e]) + qed + qed + then show "concat a \ concat b" by simp +qed + lemma [quot_respect]: "((op =) ===> op \ ===> op \) filter filter" by auto @@ -597,6 +626,11 @@ apply auto done +lemma insert_preserve2: + shows "((rep_fset ---> (map rep_fset \ rep_fset) ---> (abs_fset \ map abs_fset)) op #) = + (id ---> rep_fset ---> abs_fset) op #" + by (simp add: expand_fun_eq abs_o_rep[OF Quotient_fset] map_id Quotient_abs_rep[OF Quotient_fset]) + lemma [quot_preserve]: "(rep_fset ---> (map rep_fset \ rep_fset) ---> (abs_fset \ map abs_fset)) op # = finsert" by (simp add: expand_fun_eq Quotient_abs_rep[OF Quotient_fset] diff -r c6db12ddb60c -r c785fff02a8f Nominal/NewAlpha.thy diff -r c6db12ddb60c -r c785fff02a8f Paper/Paper.thy --- a/Paper/Paper.thy Thu May 27 18:37:52 2010 +0200 +++ b/Paper/Paper.thy Thu May 27 18:40:10 2010 +0200 @@ -1,6 +1,6 @@ (*<*) theory Paper -imports "../Nominal/Test" "../Nominal/FSet" "LaTeXsugar" +imports "../Nominal/Test" "LaTeXsugar" begin consts diff -r c6db12ddb60c -r c785fff02a8f Paper/ROOT.ML --- a/Paper/ROOT.ML Thu May 27 18:37:52 2010 +0200 +++ b/Paper/ROOT.ML Thu May 27 18:40:10 2010 +0200 @@ -1,3 +1,3 @@ quick_and_dirty := true; -no_document use_thys ["LaTeXsugar", "../Nominal/FSet"]; +no_document use_thys ["LaTeXsugar"]; use_thys ["Paper"]; \ No newline at end of file diff -r c6db12ddb60c -r c785fff02a8f Quotient-Paper/Paper.thy --- a/Quotient-Paper/Paper.thy Thu May 27 18:37:52 2010 +0200 +++ b/Quotient-Paper/Paper.thy Thu May 27 18:40:10 2010 +0200 @@ -1,7 +1,9 @@ + (*<*) theory Paper imports "Quotient" "LaTeXsugar" + "../Nominal/FSet" begin notation (latex output) @@ -10,6 +12,8 @@ fun_map ("_ ---> _" [51, 51] 50) and fun_rel ("_ ===> _" [51, 51] 50) +and + list_eq (infix "\" 50) (* Not sure if we want this notation...? *) ML {* fun nth_conj n (_, r) = nth (HOLogic.dest_conj r) n; @@ -75,18 +79,19 @@ surprising, none of them can deal compositions of quotients, for example with lifting theorems about @{text "concat"}: - @{thm concat.simps(1)}\\ - @{thm concat.simps(2)[no_vars]} + @{thm [display] concat.simps(1)} + @{thm [display] concat.simps(2)[no_vars]} \noindent - One would like to lift this definition to the operation - - @{text [display] "union definition"} + One would like to lift this definition to the operation: + + @{thm [display] fconcat_empty[no_vars]} + @{thm [display] fconcat_insert[no_vars]} \noindent What is special about this operation is that we have as input lists of lists which after lifting turn into finite sets of finite - sets. + sets. *} subsection {* Contributions *} @@ -101,7 +106,7 @@ \item We define quotient composition, function map composition and relation map composition. This lets lifting polymorphic types with subtypes quotiented as well. We extend the notions of - respectfullness and preservation to cope with quotient + respectfulness and preservation to cope with quotient composition. \item We allow lifting only some occurrences of quotiented @@ -140,8 +145,8 @@ \end{definition} - \begin{definition}[Relation map and function map] - @{thm fun_rel_def[no_vars]}\\ + \begin{definition}[Relation map and function map]\\ + @{thm fun_rel_def[of "R1" "R2", no_vars]}\\ @{thm fun_map_def[no_vars]} \end{definition} @@ -153,75 +158,303 @@ *} -section {* Constants *} - -(* Describe what containers are? *) +subsection {* Higher Order Logic *} text {* - \begin{definition}[Composition of Relations] - @{abbrev "rel_conj R1 R2"} - \end{definition} + + Types: + \begin{eqnarray}\nonumber + @{text "\ ::="} & @{text "\"} & \textrm{(type variable)} \\ \nonumber + @{text "|"} & @{text "(\,\,\)\"} & \textrm{(type construction)} + \end{eqnarray} + + Terms: + \begin{eqnarray}\nonumber + @{text "t ::="} & @{text "x\<^isup>\"} & \textrm{(variable)} \\ \nonumber + @{text "|"} & @{text "c\<^isup>\"} & \textrm{(constant)} \\ \nonumber + @{text "|"} & @{text "t t"} & \textrm{(application)} \\ \nonumber + @{text "|"} & @{text "\x\<^isup>\. t"} & \textrm{(abstraction)} \\ \nonumber + \end{eqnarray} + +*} - The first operation that we describe is the generation of - aggregate Abs or Rep function for two given compound types. - This operation will be used for the constant defnitions - and for the translation of theorems statements. It relies on - knowing map functions and relation functions for container types. - It follows the following algorithm: +section {* Constants *} + +(* Say more about containers? *) + +text {* + + To define a constant on the lifted type, an aggregate abstraction + function is applied to the raw constant. Below we describe the operation + that generates + an aggregate @{term "Abs"} or @{term "Rep"} function given the + compound raw type and the compound quotient type. + This operation will also be used in translations of theorem statements + and in the lifting procedure. + + The operation is additionally able to descend into types for which + maps are known. Such maps for most common types (list, pair, sum, + option, \ldots) are described in Homeier, and we assume that @{text "map"} + is the function that returns a map for a given type. Then REP/ABS is defined + as follows: \begin{itemize} - \item For equal types or free type variables return identity. - - \item For function types recurse, change the Rep/Abs flag to - the opposite one for the domain type and compose the - results with @{term "fun_map"}. - - \item For equal type constructors use the appropriate map function - applied to the results for the arguments. - - \item For unequal type constructors are unequal, we look in the - quotients information for a raw type quotient type pair that - matches the given types. We apply the environment to the arguments - and recurse composing it with the aggregate map function. + \item @{text "ABS(\\<^isub>1, \\<^isub>2)"} = @{text "id"} + \item @{text "REP(\\<^isub>1, \\<^isub>2)"} = @{text "id"} + \item @{text "ABS(\, \)"} = @{text "id"} + \item @{text "REP(\, \)"} = @{text "id"} + \item @{text "ABS(\\<^isub>1\\\<^isub>2,\\<^isub>1\\\<^isub>2)"} = @{text "REP(\\<^isub>1,\\<^isub>1) ---> ABS(\\<^isub>2,\\<^isub>2)"} + \item @{text "REP(\\<^isub>1\\\<^isub>2,\\<^isub>1\\\<^isub>2)"} = @{text "ABS(\\<^isub>1,\\<^isub>1) ---> REP(\\<^isub>2,\\<^isub>2)"} + \item @{text "ABS((\\<^isub>1,\,\\<^isub>n))\, (\\<^isub>1,\,\\<^isub>n))\)"} = @{text "(map \) (ABS(\\<^isub>1,\\<^isub>1)) \ (ABS(\\<^isub>n,\\<^isub>n))"} + \item @{text "REP((\\<^isub>1,\,\\<^isub>n))\, (\\<^isub>1,\,\\<^isub>n))\)"} = @{text "(map \) (REP(\\<^isub>1,\\<^isub>1)) \ (REP(\\<^isub>n,\\<^isub>n))"} + \item @{text "ABS((\\<^isub>1,\,\\<^isub>n))\\<^isub>1, (\\<^isub>1,\,\\<^isub>m))\\<^isub>2)"} = @{text "Abs_\\<^isub>2 \ (map \\<^isub>1) (ABS(\\<^isub>1,\\<^isub>1) \ (ABS(\\<^isub>p,\\<^isub>p)"} provided @{text "\ \\<^isub>2 = (\\<^isub>1\\\<^isub>p)\\<^isub>1 \ \s. s(\s\\<^isub>1)=\s\\<^isub>1 \ s(\s\\<^isub>2)=\s\\<^isub>2"} + \item @{text "REP((\\<^isub>1,\,\\<^isub>n))\\<^isub>1, (\\<^isub>1,\,\\<^isub>m))\\<^isub>2)"} = @{text "(map \\<^isub>1) (REP(\\<^isub>1,\\<^isub>1) \ (REP(\\<^isub>p,\\<^isub>p) \ Rep_\\<^isub>2"} provided @{text "\ \\<^isub>2 = (\\<^isub>1\\\<^isub>p)\\<^isub>1 \ \s. s(\s\\<^isub>1)=\s\\<^isub>1 \ s(\s\\<^isub>2)=\s\\<^isub>2"} \end{itemize} + Apart from the last 2 points the definition is same as the one implemented in + in Homeier's HOL package, below is the definition of @{term fconcat} + that shows the last points: + @{thm fconcat_def[no_vars]} + + The aggregate @{term Abs} function takes a finite set of finite sets + and applies @{term "map rep_fset"} composed with @{term rep_fset} to + its input, obtaining a list of lists, passes the result to @{term concat} + obtaining a list and applies @{term abs_fset} obtaining the composed + finite set. +*} + +subsection {* Respectfulness *} + +text {* + + A respectfulness lemma for a constant states that the equivalence + class returned by this constant depends only on the equivalence + classes of the arguments applied to the constant. This can be + expressed in terms of an aggregate relation between the constant + and itself, for example the respectfullness for @{term "append"} + can be stated as: + + @{thm [display] append_rsp[no_vars]} + + \noindent + Which is equivalent to: + + @{thm [display] append_rsp_unfolded[no_vars]} + + Below we show the algorithm for finding the aggregate relation. + This algorithm uses + the relation composition which we define as: + + \begin{definition}[Composition of Relations] + @{abbrev "rel_conj R1 R2"} where @{text OO} is the predicate + composition @{thm pred_compI[no_vars]} + \end{definition} + + Given an aggregate raw type and quotient type: + + \begin{itemize} + \item For equal types or free type variables return equality + + \item For equal type constructors use the appropriate rel + function applied to the results for the argument pairs + + \item For unequal type constructors, look in the quotients information + for a quotient type that matches the type constructor, and instantiate + the type appropriately getting back an instantiation environment. We + apply the environment to the arguments and recurse composing it with + the aggregate relation function. + + \end{itemize} + + Again, the the behaviour of our algorithm in the last situation is + novel, so lets look at the example of respectfullness for @{term concat}. + The statement as computed by the algorithm above is: - Rsp and Prs + @{thm [display] concat_rsp[no_vars]} + + \noindent + By unfolding the definition of relation composition and relation map + we can see the equivalent statement just using the primitive list + equivalence relation: + + @{thm [display] concat_rsp_unfolded[of "a" "a'" "b'" "b", no_vars]} + + The statement reads that, for any lists of lists @{term a} and @{term b} + if there exist intermediate lists of lists @{term "a'"} and @{term "b'"} + such that each element of @{term a} is in the relation with an appropriate + element of @{term a'}, @{term a'} is in relation with @{term b'} and each + element of @{term b'} is in relation with the appropriate element of + @{term b}. + *} +subsection {* Preservation *} + +text {* + To be able to lift theorems that talk about constants that are not + lifted but whose type changes when lifting is performed additionally + preservation theorems are needed. + + To lift theorems that talk about insertion in lists of lifted types + we need to know that for any quotient type with the abstraction and + representation functions @{text "Abs"} and @{text Rep} we have: + + @{thm [display] (concl) cons_prs[no_vars]} + + This is not enough to lift theorems that talk about quotient compositions. + For some constants (for example empty list) it is possible to show a + general compositional theorem, but for @{term "op #"} it is necessary + to show that it respects the particular quotient type: + + @{thm [display] insert_preserve2[no_vars]} +*} + +subsection {* Composition of Quotient theorems *} + +text {* + Given two quotients, one of which quotients a container, and the + other quotients the type in the container, we can write the + composition of those quotients. To compose two quotient theorems + we compose the relations with relation composition + and the abstraction and relation functions with function composition. + The @{term "Rep"} and @{term "Abs"} functions that we obtain are + the same as the ones created by in the aggregate functions and the + relation is the same as the one given by aggregate relations. + This becomes especially interesting + when we compose the quotient with itself, as there is no simple + intermediate step. + + Lets take again the example of @{term concat}. To be able to lift + theorems that talk about it we will first prove the composition + quotient theorems, which then lets us perform the lifting procedure + in an unchanged way: + + @{thm [display] quotient_compose_list[no_vars]} +*} + + section {* Lifting Theorems *} -text {* TBD *} +text {* + The core of the quotient package takes an original theorem that + talks about the raw types, and the statement of the theorem that + it is supposed to produce. This is different from other existing + quotient packages, where only the raw theorems was necessary. + We notice that in some cases only some occurrences of the raw + types need to be lifted. This is for example the case in the + new Nominal package, where a raw datatype that talks about + pairs of natural numbers or strings (being lists of characters) + should not be changed to a quotient datatype with constructors + taking integers or finite sets of characters. To simplify the + use of the quotient package we additionally provide an automated + statement translation mechanism that replaces occurrences of + types that match given quotients by appropriate lifted types. -text {* Why providing a statement to prove is necessary is some cases *} + Lifting the theorems is performed in three steps. In the following + we call these steps \emph{regularization}, \emph{injection} and + \emph{cleaning} following the names used in Homeier's HOL + implementation. -subsection {* Regularization *} + We first define the statement of the regularized theorem based + on the original theorem and the goal theorem. Then we define + the statement of the injected theorem, based on the regularized + theorem and the goal. We then show the 3 proofs, and all three + can be performed independently from each other. + +*} + +subsection {* Regularization and Injection statements *} text {* -Transformation of the theorem statement: -\begin{itemize} -\item Quantifiers and abstractions involving raw types replaced by bounded ones. -\item Equalities involving raw types replaced by bounded ones. -\end{itemize} + + The function that gives the statement of the regularized theorem + takes the statement of the raw theorem (a term) and the statement + of the lifted theorem. The intuition behind the procedure is that + it replaces quantifiers and abstractions involving raw types + by bounded ones, and equalities involving raw types are replaced + by appropriate aggregate relations. It is defined as follows: -The procedure. + \begin{itemize} + \item @{text "REG (\x : \. t, \x : \. s) = \x : \. REG (t, s)"} + \item @{text "REG (\x : \. t, \x : \. s) = \x : \ \ Res (REL (\, \)). REG (t, s)"} + \item @{text "REG (\x : \. t, \x : \. s) = \x : \. REG (t, s)"} + \item @{text "REG (\x : \. t, \x : \. s) = \x : \ \ Res (REL (\, \)). REG (t, s)"} + \item @{text "REG ((op =) : \, (op =) : \) = (op =) : \"} + \item @{text "REG ((op =) : \, (op =) : \) = REL (\, \) : \"} + \item @{text "REG (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2) = REG (t\<^isub>1, s\<^isub>1) REG (t\<^isub>2, s\<^isub>2)"} + \item @{text "REG (v\<^isub>1, v\<^isub>2) = v\<^isub>1"} + \item @{text "REG (c\<^isub>1, c\<^isub>2) = c\<^isub>1"} + \end{itemize} + + Existential quantifiers and unique existential quantifiers are defined + similarily to the universal one. + + The function that gives the statment of the injected theorem + takes the statement of the regularized theorems and the statement + of the lifted theorem both as terms. -Example of non-regularizable theorem ($0 = 1$). + \begin{itemize} + \item @{text "INJ ((\x. t) : \, (\x. s) : \) = \x. (INJ (t, s)"} + \item @{text "INJ ((\x. t) : \, (\x. s) : \) = REP(\,\) (ABS (\,\) (\x. (INJ (t, s))))"} + \item @{text "INJ ((\x \ R. t) : \, (\x. s) : \) = REP(\,\) (ABS (\,\) (\x \ R. (INJ (t, s))))"} + \item @{text "INJ (\ t, \ s) = \ (INJ (t, s)"} + \item @{text "INJ (\ t \ R, \ s) = \ (INJ (t, s) \ R"} + \item @{text "INJ (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2) = INJ (t\<^isub>1, s\<^isub>1) INJ (t\<^isub>2, s\<^isub>2)"} + \item @{text "INJ (v\<^isub>1 : \, v\<^isub>2 : \) = v\<^isub>1"} + \item @{text "INJ (v\<^isub>1 : \, v\<^isub>2 : \) = REP(\,\) (ABS (\,\) (v\<^isub>1))"} + \item @{text "INJ (c\<^isub>1 : \, c\<^isub>2 : \) = c\<^isub>1"} + \item @{text "INJ (c\<^isub>1 : \, c\<^isub>2 : \) = REP(\,\) (ABS (\,\) (c\<^isub>1))"} + \end{itemize} + + For existential quantifiers and unique existential quantifiers it is + defined similarily to the universal one. + +*} + +subsection {* Proof of Regularization *} -Separtion of regularization from injection thanks to the following 2 lemmas: -\begin{lemma} -If @{term R2} is an equivalence relation, then: -\begin{eqnarray} -@{thm (rhs) ball_reg_eqv_range[no_vars]} & = & @{thm (lhs) ball_reg_eqv_range[no_vars]}\\ -@{thm (rhs) bex_reg_eqv_range[no_vars]} & = & @{thm (lhs) bex_reg_eqv_range[no_vars]} -\end{eqnarray} -\end{lemma} +text {* + Example of non-regularizable theorem ($0 = 1$). + + + Separtion of regularization from injection thanks to the following 2 lemmas: + \begin{lemma} + If @{term R2} is an equivalence relation, then: + \begin{eqnarray} + @{thm (rhs) ball_reg_eqv_range[no_vars]} & = & @{thm (lhs) ball_reg_eqv_range[no_vars]}\\ + @{thm (rhs) bex_reg_eqv_range[no_vars]} & = & @{thm (lhs) bex_reg_eqv_range[no_vars]} + \end{eqnarray} + \end{lemma} + + Other lemmas used in regularization: + @{thm [display] ball_reg_eqv[no_vars]} + @{thm [display] bex_reg_eqv[no_vars]} + @{thm [display] babs_reg_eqv[no_vars]} + @{thm [display] babs_simp[no_vars]} + + @{thm [display] ball_reg_right[no_vars]} + @{thm [display] bex_reg_left[no_vars]} + @{thm [display] bex1_bexeq_reg[no_vars]} *} subsection {* Injection *} +text {* + + The 2 key lemmas are: + + @{thm [display] apply_rsp[no_vars]} + @{thm [display] rep_abs_rsp[no_vars]} + + + +*} + + + + subsection {* Cleaning *} text {* Preservation of quantifiers, abstractions, relations, quotient-constants diff -r c6db12ddb60c -r c785fff02a8f Quotient-Paper/ROOT.ML --- a/Quotient-Paper/ROOT.ML Thu May 27 18:37:52 2010 +0200 +++ b/Quotient-Paper/ROOT.ML Thu May 27 18:40:10 2010 +0200 @@ -1,4 +1,5 @@ no_document use_thys ["Quotient", - "LaTeXsugar"]; + "LaTeXsugar", + "../Nominal/FSet" ]; use_thys ["Paper"]; \ No newline at end of file