Mon, 01 Feb 2010 09:56:32 +0100 Cezary Kaliszyk Ported LF to the generic lambda and solved the simpler _supp cases.
Sat, 30 Jan 2010 12:12:52 +0100 Christian Urban merged
Sat, 30 Jan 2010 11:44:25 +0100 Christian Urban introduced a generic alpha (but not sure whether it is helpful)
Fri, 29 Jan 2010 19:42:07 +0100 Cezary Kaliszyk More in the LF example in the new nominal way, all is clear until support.
Fri, 29 Jan 2010 13:47:05 +0100 Cezary Kaliszyk Fixed the induction problem + some more proofs.
Fri, 29 Jan 2010 12:16:08 +0100 Cezary Kaliszyk equivariance of rfv and alpha.
Fri, 29 Jan 2010 10:13:07 +0100 Cezary Kaliszyk Added the experiments with fun and function.
Fri, 29 Jan 2010 07:09:52 +0100 Christian Urban now also final step is proved - the supp of lambdas is now completely characterised
Fri, 29 Jan 2010 00:22:00 +0100 Christian Urban the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
Thu, 28 Jan 2010 23:47:02 +0100 Christian Urban improved the proof slightly by defining alpha as a function and completely characterised the equality between two abstractions
Thu, 28 Jan 2010 23:36:58 +0100 Christian Urban merged
Thu, 28 Jan 2010 23:36:38 +0100 Christian Urban general abstraction operator and complete characterisation of its support and freshness
Thu, 28 Jan 2010 19:23:55 +0100 Cezary Kaliszyk Ported existing part of LF to new permutations and alphas.
Thu, 28 Jan 2010 15:47:35 +0100 Christian Urban attempt of a general abstraction operator
Thu, 28 Jan 2010 14:20:26 +0100 Christian Urban attempt to prove equivalence between alpha definitions
Thu, 28 Jan 2010 12:28:50 +0100 Cezary Kaliszyk End of renaming.
Thu, 28 Jan 2010 12:25:38 +0100 Cezary Kaliszyk Minor when looking at lam.distinct and lam.inject
Thu, 28 Jan 2010 12:24:49 +0100 Cezary Kaliszyk Renamed Bexeq to Bex1_rel
Thu, 28 Jan 2010 10:52:10 +0100 Cezary Kaliszyk Substracting bounds from free variables.
Thu, 28 Jan 2010 10:26:36 +0100 Cezary Kaliszyk Improper interface for datatype and function packages and proper interface lateron.
Thu, 28 Jan 2010 09:28:20 +0100 Christian Urban merged
Thu, 28 Jan 2010 09:28:06 +0100 Christian Urban minor
Thu, 28 Jan 2010 01:24:09 +0100 Christian Urban test about supp/freshness for lam (old proofs work in principle - for single binders)
Thu, 28 Jan 2010 08:13:39 +0100 Cezary Kaliszyk Recommited the changes for nitpick
Wed, 27 Jan 2010 18:26:01 +0100 Cezary Kaliszyk Correct types which fixes the printing.
Wed, 27 Jan 2010 18:06:14 +0100 Cezary Kaliszyk fv for subterms
Wed, 27 Jan 2010 17:39:13 +0100 Cezary Kaliszyk Fix the problem with later examples. Maybe need to go back to textual specifications.
Wed, 27 Jan 2010 17:18:30 +0100 Cezary Kaliszyk Some processing of variables in constructors to get free variables.
Wed, 27 Jan 2010 16:40:16 +0100 Cezary Kaliszyk Parsing of the input as terms and types, and passing them as such to the function package.
Wed, 27 Jan 2010 16:07:49 +0100 Cezary Kaliszyk Undid the parsing, as it is not possible with thy->lthy interaction.
Wed, 27 Jan 2010 14:57:11 +0100 Cezary Kaliszyk merge
Wed, 27 Jan 2010 14:56:58 +0100 Cezary Kaliszyk Some cleaning of thy vs lthy vs context.
Wed, 27 Jan 2010 14:06:34 +0100 Christian Urban merged
Wed, 27 Jan 2010 14:06:17 +0100 Christian Urban tuned comment
Wed, 27 Jan 2010 14:05:42 +0100 Christian Urban completely ported
Wed, 27 Jan 2010 13:44:05 +0100 Cezary Kaliszyk Another string in the specification.
Wed, 27 Jan 2010 13:32:28 +0100 Cezary Kaliszyk Variable takes a 'name'.
Wed, 27 Jan 2010 12:21:40 +0100 Cezary Kaliszyk merge
Wed, 27 Jan 2010 12:19:58 +0100 Cezary Kaliszyk When commenting discovered a missing case of Babs->Abs regularization.
Wed, 27 Jan 2010 12:19:21 +0100 Christian Urban merged
Wed, 27 Jan 2010 12:19:00 +0100 Christian Urban mostly ported Terms.thy to new Nominal
Wed, 27 Jan 2010 12:06:43 +0100 Cezary Kaliszyk merge
Wed, 27 Jan 2010 12:06:24 +0100 Cezary Kaliszyk Commenting regularize
Wed, 27 Jan 2010 11:48:04 +0100 Christian Urban very rough example file for how nominal2 specification can be parsed
Wed, 27 Jan 2010 11:31:16 +0100 Christian Urban reordered cases in regularize (will be merged into two cases)
Wed, 27 Jan 2010 08:41:42 +0100 Christian Urban use of equiv_relation_chk in quotient_term
Wed, 27 Jan 2010 08:20:31 +0100 Christian Urban some slight tuning
Wed, 27 Jan 2010 07:49:43 +0100 Christian Urban added Terms to Nominal - Instantiation of two types does not work (ask Florian)
Wed, 27 Jan 2010 07:45:01 +0100 Christian Urban added another example with indirect recursion over lists
Tue, 26 Jan 2010 20:12:41 +0100 Christian Urban just moved obsolete material into Attic
Tue, 26 Jan 2010 20:07:50 +0100 Christian Urban added an LamEx example together with the new nominal infrastructure
Tue, 26 Jan 2010 16:30:51 +0100 Cezary Kaliszyk Bex1_Bexeq_regular.
Tue, 26 Jan 2010 15:59:04 +0100 Cezary Kaliszyk Hom Theorem with exists unique
Tue, 26 Jan 2010 14:48:25 +0100 Cezary Kaliszyk 2 cases for regularize with split, lemmas with split now lift.
Tue, 26 Jan 2010 14:08:47 +0100 Cezary Kaliszyk Simpler statement that has the problem.
Tue, 26 Jan 2010 13:58:28 +0100 Cezary Kaliszyk Found a term that does not regularize.
Tue, 26 Jan 2010 13:53:56 +0100 Cezary Kaliszyk A triple is still ok.
Tue, 26 Jan 2010 13:38:42 +0100 Cezary Kaliszyk Combined the simpsets in clean_tac and updated the comment. Now cleaning of splits does work.
Tue, 26 Jan 2010 12:24:23 +0100 Cezary Kaliszyk Changed the lambda_prs_simple_conv to use id_apply, now last eq_reflection can be removed from id_simps.
Tue, 26 Jan 2010 12:06:47 +0100 Cezary Kaliszyk Sigma cleaning works with split_prs (still manual proof).
Tue, 26 Jan 2010 11:13:08 +0100 Christian Urban tuned
Tue, 26 Jan 2010 10:53:44 +0100 Christian Urban merged
Tue, 26 Jan 2010 01:42:46 +0100 Christian Urban cleaning of QuotProd; a little cleaning of QuotList
Tue, 26 Jan 2010 01:00:35 +0100 Christian Urban added prs and rsp lemmas for Inl and Inr
Tue, 26 Jan 2010 00:47:40 +0100 Christian Urban used split_option_all lemma
Tue, 26 Jan 2010 00:18:48 +0100 Christian Urban used the internal Option.map instead of custom option_map
Tue, 26 Jan 2010 09:54:43 +0100 Cezary Kaliszyk Generalized split_prs and split_rsp
Tue, 26 Jan 2010 09:28:32 +0100 Cezary Kaliszyk All eq_reflections apart from the one of 'id_apply' can be removed.
Tue, 26 Jan 2010 08:55:55 +0100 Cezary Kaliszyk continued
Tue, 26 Jan 2010 08:09:22 +0100 Cezary Kaliszyk More eqreflection/equiv cleaning.
Tue, 26 Jan 2010 07:42:52 +0100 Cezary Kaliszyk more eq_reflection & other cleaning.
Tue, 26 Jan 2010 07:14:10 +0100 Cezary Kaliszyk Removing more eq_reflections.
Mon, 25 Jan 2010 20:47:20 +0100 Christian Urban ids *cannot* be object equalities
Mon, 25 Jan 2010 20:35:42 +0100 Christian Urban re-inserted lemma in QuotList
Mon, 25 Jan 2010 19:52:53 +0100 Christian Urban added prs and rsp lemmas for Some and None
Mon, 25 Jan 2010 19:14:46 +0100 Christian Urban tuned proofs (mainly in QuotProd)
Mon, 25 Jan 2010 18:52:22 +0100 Christian Urban properly commented out the "unused lemmas section" and moved actually used lemmas elsewhere; added two minor items to the TODO list
Mon, 25 Jan 2010 18:13:44 +0100 Christian Urban renamed QuotScript to QuotBase
Mon, 25 Jan 2010 17:53:08 +0100 Christian Urban cleaned some theorems
Sun, 24 Jan 2010 23:41:27 +0100 Christian Urban test with splits
Sat, 23 Jan 2010 17:25:18 +0100 Cezary Kaliszyk The alpha equivalence relations for structures in 'Terms'
Sat, 23 Jan 2010 15:41:54 +0100 Cezary Kaliszyk More experiments with defining the homomorphism directly, lifting of 'distinct' and of 'exhaust'.
Sat, 23 Jan 2010 07:22:27 +0100 Cezary Kaliszyk Trying to define hom for the lifted type directly.
Fri, 22 Jan 2010 17:44:46 +0100 Cezary Kaliszyk Proper alpha equivalence for Sigma calculus.
Thu, 21 Jan 2010 19:52:46 +0100 Cezary Kaliszyk Changed fun_map and rel_map to definitions.
Thu, 21 Jan 2010 12:50:43 +0100 Cezary Kaliszyk Lifted Peter's Sigma lemma with Ex1.
Thu, 21 Jan 2010 12:03:47 +0100 Cezary Kaliszyk Automatic injection of Bexeq
Thu, 21 Jan 2010 11:11:22 +0100 Cezary Kaliszyk Automatic cleaning of Bexeq<->Ex1 theorems.
Thu, 21 Jan 2010 10:55:09 +0100 Cezary Kaliszyk Using Bexeq_rsp, and manually lifted lemma with Ex1.
Thu, 21 Jan 2010 09:55:05 +0100 Cezary Kaliszyk Bexeq definition, Ex1_prs lemma, Bex1_rsp lemma, compiles.
Thu, 21 Jan 2010 09:02:04 +0100 Cezary Kaliszyk The missing rule.
Thu, 21 Jan 2010 07:38:34 +0100 Cezary Kaliszyk Ex1 -> Bex1 Regularization, Preparing Exeq.
Wed, 20 Jan 2010 16:50:31 +0100 Cezary Kaliszyk Added the Sigma Calculus example
Wed, 20 Jan 2010 16:44:31 +0100 Cezary Kaliszyk Better error messages for non matching quantifiers.
Wed, 20 Jan 2010 12:33:19 +0100 Cezary Kaliszyk Statement of term1_hom_rsp
Wed, 20 Jan 2010 12:20:18 +0100 Christian Urban proved that the function is a function
Wed, 20 Jan 2010 11:30:32 +0100 Cezary Kaliszyk term1_hom as a function
Tue, 19 Jan 2010 18:17:42 +0100 Cezary Kaliszyk A version of hom with quantifiers.
Sun, 17 Jan 2010 02:24:15 +0100 Christian Urban added permutation functions for the raw calculi
Sat, 16 Jan 2010 04:23:27 +0100 Christian Urban fixed broken (partial) proof
Sat, 16 Jan 2010 03:56:00 +0100 Christian Urban used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
Sat, 16 Jan 2010 02:09:38 +0100 Christian Urban liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Fri, 15 Jan 2010 17:09:36 +0100 Christian Urban added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Fri, 15 Jan 2010 16:13:49 +0100 Christian Urban tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Fri, 15 Jan 2010 15:56:25 +0100 Christian Urban merged
Fri, 15 Jan 2010 15:56:06 +0100 Christian Urban added free_variable function (do not know about the algorithm yet)
Fri, 15 Jan 2010 15:51:25 +0100 Cezary Kaliszyk hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Fri, 15 Jan 2010 12:17:30 +0100 Christian Urban slight tuning of relation_error
Fri, 15 Jan 2010 11:04:21 +0100 Cezary Kaliszyk Appropriate respects and a statement of the lifted hom lemma
Fri, 15 Jan 2010 10:48:49 +0100 Christian Urban recursion-hom for lambda
Fri, 15 Jan 2010 10:36:48 +0100 Cezary Kaliszyk Incorrect version of the homomorphism lemma
Thu, 14 Jan 2010 23:51:17 +0100 Christian Urban trivial
Thu, 14 Jan 2010 23:48:31 +0100 Christian Urban tuned quotient_typ.ML
Thu, 14 Jan 2010 23:17:21 +0100 Christian Urban tuned quotient_def.ML and cleaned somewhat LamEx.thy
Thu, 14 Jan 2010 19:03:08 +0100 Christian Urban a few more lemmas...except supp of lambda-abstractions
Thu, 14 Jan 2010 18:41:50 +0100 Christian Urban removed one sorry
Thu, 14 Jan 2010 18:35:38 +0100 Christian Urban nearly all of the proof
Thu, 14 Jan 2010 17:57:20 +0100 Christian Urban right generalisation
Thu, 14 Jan 2010 17:53:23 +0100 Cezary Kaliszyk First subgoal.
Thu, 14 Jan 2010 17:13:11 +0100 Christian Urban setup for strong induction
Thu, 14 Jan 2010 16:41:17 +0100 Cezary Kaliszyk exported absrep_const for nitpick.
Thu, 14 Jan 2010 15:36:29 +0100 Cezary Kaliszyk minor
Thu, 14 Jan 2010 15:25:24 +0100 Cezary Kaliszyk Simplified matches_typ.
Thu, 14 Jan 2010 12:23:59 +0100 Christian Urban added bound-variable functions to terms
Thu, 14 Jan 2010 12:17:39 +0100 Christian Urban merged
Thu, 14 Jan 2010 12:14:35 +0100 Christian Urban added 3 calculi with interesting binding structure
Thu, 14 Jan 2010 10:51:03 +0100 Cezary Kaliszyk produce defs with lthy, like prs and ids
Thu, 14 Jan 2010 10:47:19 +0100 Cezary Kaliszyk Remove SOLVED from quotient_tac. Move atomize_eqv to 'Unused'.
Thu, 14 Jan 2010 10:06:29 +0100 Cezary Kaliszyk Finished organising an efficient datastructure for qconst_info.
Thu, 14 Jan 2010 08:02:20 +0100 Cezary Kaliszyk Undid changes from symtab to termtab, since we need to lookup specialized types.
Wed, 13 Jan 2010 16:46:25 +0100 Cezary Kaliszyk Moved the matches_typ function outside a?d simplified it.
Wed, 13 Jan 2010 16:39:20 +0100 Christian Urban one more item in the list of Markus
Wed, 13 Jan 2010 16:23:32 +0100 Cezary Kaliszyk Put relation_error as a separate function.
Wed, 13 Jan 2010 16:14:02 +0100 Cezary Kaliszyk Better error message for definition failure.
Wed, 13 Jan 2010 15:17:52 +0100 Cezary Kaliszyk merge
Wed, 13 Jan 2010 15:17:36 +0100 Cezary Kaliszyk Stored Termtab for constant information.
Wed, 13 Jan 2010 13:40:47 +0100 Christian Urban merged
Wed, 13 Jan 2010 13:40:23 +0100 Christian Urban deleted SOLVED'
Wed, 13 Jan 2010 13:12:04 +0100 Cezary Kaliszyk Removed the 'oops' in IntEx.
Wed, 13 Jan 2010 09:41:57 +0100 Christian Urban tuned
Wed, 13 Jan 2010 09:30:59 +0100 Christian Urban added SOLVED' which is now part of Isabelle....must be removed eventually
Wed, 13 Jan 2010 09:19:20 +0100 Christian Urban merged
Wed, 13 Jan 2010 00:46:31 +0100 Christian Urban tuned
Wed, 13 Jan 2010 00:45:28 +0100 Christian Urban absrep_fun and equiv_relation do not produce anymore spurious maps; two problems arose in IntEx, which are marked with "INJECTION PROBLEM"
Tue, 12 Jan 2010 17:46:35 +0100 Cezary Kaliszyk More indenting, bracket removing and comment restructuring.
Tue, 12 Jan 2010 16:44:33 +0100 Cezary Kaliszyk Finished replacing OO by OOO
Tue, 12 Jan 2010 16:28:53 +0100 Cezary Kaliszyk Change OO to OOO in FSet3.
Tue, 12 Jan 2010 16:21:42 +0100 Cezary Kaliszyk minor comment editing
Tue, 12 Jan 2010 16:12:54 +0100 Cezary Kaliszyk modifying comments/indentation in quotient_term.ml
Tue, 12 Jan 2010 16:03:51 +0100 Cezary Kaliszyk Cleaning comments, indentation etc in quotient_tacs.
Tue, 12 Jan 2010 15:48:46 +0100 Cezary Kaliszyk No more exception handling in rep_abs_rsp_tac
Tue, 12 Jan 2010 12:14:33 +0100 Cezary Kaliszyk handle all is no longer necessary in lambda_prs.
Tue, 12 Jan 2010 12:04:47 +0100 Cezary Kaliszyk removed 3 hacks.
Tue, 12 Jan 2010 11:25:38 +0100 Cezary Kaliszyk Updated some comments.
Tue, 12 Jan 2010 10:59:51 +0100 Cezary Kaliszyk merge
Tue, 12 Jan 2010 10:59:38 +0100 Cezary Kaliszyk Removed exception handling from equals_rsp_tac.
Mon, 11 Jan 2010 22:36:21 +0100 Christian Urban added an abbreviation for OOO
Mon, 11 Jan 2010 20:04:19 +0100 Cezary Kaliszyk merge
Mon, 11 Jan 2010 20:03:43 +0100 Cezary Kaliszyk Undid the non-working part.
Mon, 11 Jan 2010 16:33:00 +0100 Christian Urban started to adhere to Wenzel-Standard
Mon, 11 Jan 2010 15:58:38 +0100 Cezary Kaliszyk Changing exceptions to 'try', part 1.
Mon, 11 Jan 2010 15:13:09 +0100 Cezary Kaliszyk removed quotdata_lookup_type
Mon, 11 Jan 2010 11:51:19 +0100 Cezary Kaliszyk Fix for testing matching constants in regularize.
Mon, 11 Jan 2010 01:03:34 +0100 Christian Urban tuned previous commit further
Mon, 11 Jan 2010 00:31:29 +0100 Christian Urban the chk-functions in quotient_term also simplify the result according to the id_simps; had to remove id_def from this theorem list though; this caused in FSet3 that relied on this rule; the problem is marked with "ID PROBLEM"
Sat, 09 Jan 2010 09:38:34 +0100 Christian Urban introduced separate match function
Sat, 09 Jan 2010 08:52:06 +0100 Christian Urban removed obsolete equiv_relation and rnamed new_equiv_relation
Fri, 08 Jan 2010 19:46:22 +0100 Cezary Kaliszyk New_relations, all works again including concat examples.
Fri, 08 Jan 2010 15:02:12 +0100 Cezary Kaliszyk map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Fri, 08 Jan 2010 14:43:30 +0100 Cezary Kaliszyk id_simps needs to be taken out not used directly, otherwise the new lemmas are not there.
Fri, 08 Jan 2010 11:20:12 +0100 Cezary Kaliszyk Experimients with fconcat_insert
Fri, 08 Jan 2010 10:44:30 +0100 Cezary Kaliszyk Modifications for new_equiv_rel, part2
Fri, 08 Jan 2010 10:39:08 +0100 Cezary Kaliszyk Modifictaions for new_relation.
Fri, 08 Jan 2010 10:08:01 +0100 Cezary Kaliszyk Proved concat_empty.
Thu, 07 Jan 2010 16:51:38 +0100 Cezary Kaliszyk Replacing equivp by reflp in the assumptions leads to non-provable subgoals in the gen_pre lemmas.
Thu, 07 Jan 2010 16:06:13 +0100 Cezary Kaliszyk some cleaning.
Thu, 07 Jan 2010 15:50:22 +0100 Cezary Kaliszyk First generalization.
Thu, 07 Jan 2010 14:14:17 +0100 Cezary Kaliszyk The working proof of the special case.
Thu, 07 Jan 2010 10:55:20 +0100 Cezary Kaliszyk Reduced the proof to two simple but not obvious to prove facts.
Thu, 07 Jan 2010 10:13:15 +0100 Cezary Kaliszyk More cleaning and commenting AbsRepTest. Now tests work; just slow.
Thu, 07 Jan 2010 09:55:42 +0100 Cezary Kaliszyk cleaning in AbsRepTest.
Wed, 06 Jan 2010 16:24:21 +0100 Cezary Kaliszyk Further in the proof
Wed, 06 Jan 2010 09:19:23 +0100 Cezary Kaliszyk Tried to prove the lemma manually; only left with quotient proofs.
Wed, 06 Jan 2010 08:24:37 +0100 Cezary Kaliszyk Sledgehammer bug.
Tue, 05 Jan 2010 18:10:20 +0100 Cezary Kaliszyk merge
Tue, 05 Jan 2010 18:09:03 +0100 Cezary Kaliszyk Trying the proof
Tue, 05 Jan 2010 17:12:35 +0100 Christian Urban merged
Tue, 05 Jan 2010 17:06:51 +0100 Cezary Kaliszyk Struggling with composition
Tue, 05 Jan 2010 15:25:31 +0100 Cezary Kaliszyk Trying to state composition quotient.
Tue, 05 Jan 2010 14:23:45 +0100 Christian Urban proper handling of error messages (code copy - maybe this can be avoided)
Tue, 05 Jan 2010 14:09:04 +0100 Christian Urban added a new version of equiv_relation (is not yet used anywhere except in AbsRepTest)
Tue, 05 Jan 2010 10:41:20 +0100 Cezary Kaliszyk Readded 'regularize_to_injection' which I believe will be needed.
Sat, 02 Jan 2010 23:15:15 +0100 Christian Urban added a warning to the quotient_type definition, if a map function is missing
Fri, 01 Jan 2010 23:59:32 +0100 Christian Urban tuned
Fri, 01 Jan 2010 11:30:00 +0100 Christian Urban a slight change to abs/rep generation
Fri, 01 Jan 2010 04:39:43 +0100 Christian Urban tuned
Fri, 01 Jan 2010 01:10:38 +0100 Christian Urban fixed comment errors
Fri, 01 Jan 2010 01:08:19 +0100 Christian Urban some slight tuning
Thu, 31 Dec 2009 23:53:10 +0100 Christian Urban renamed transfer to transform (Markus)
Wed, 30 Dec 2009 12:10:57 +0000 cu some small changes
Sun, 27 Dec 2009 23:33:10 +0100 Christian Urban added a functor that allows checking what is added to the theorem lists
Sat, 26 Dec 2009 23:20:46 +0100 Christian Urban corrected wrong [quot_respect] attribute; tuned
Sat, 26 Dec 2009 21:36:20 +0100 Christian Urban renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Sat, 26 Dec 2009 20:45:37 +0100 Christian Urban added an item about when the same quotient constant is defined twice (throws a bind exception in quoteint_def.ML)
Sat, 26 Dec 2009 20:24:53 +0100 Christian Urban as expected problems occure when lifting concat lemmas
Sat, 26 Dec 2009 09:03:35 +0100 Christian Urban tuned
Sat, 26 Dec 2009 08:06:45 +0100 Christian Urban commeted the absrep function
Sat, 26 Dec 2009 07:15:30 +0100 Christian Urban generalised absrep function; needs consolidation
Fri, 25 Dec 2009 00:58:06 +0100 Christian Urban tuned
Fri, 25 Dec 2009 00:17:55 +0100 Christian Urban added sanity checks for quotient_type
Thu, 24 Dec 2009 22:28:19 +0100 Christian Urban made the quotient_type definition more like typedef; now type variables need to be explicitly given
Thu, 24 Dec 2009 00:58:50 +0100 Christian Urban used Local_Theory.declaration for storing quotdata
Wed, 23 Dec 2009 23:53:03 +0100 Christian Urban tuning
Wed, 23 Dec 2009 23:22:02 +0100 Christian Urban renamed some fields in the info records
Wed, 23 Dec 2009 22:42:30 +0100 Christian Urban modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Wed, 23 Dec 2009 21:30:23 +0100 Christian Urban cleaed a bit function mk_typedef_main
Wed, 23 Dec 2009 13:45:42 +0100 Christian Urban renamed QUOT_TYPE to Quot_Type
Wed, 23 Dec 2009 13:23:33 +0100 Christian Urban explicit handling of mem_def, avoiding the use of the simplifier; this fixes some quotient_type definitions
Wed, 23 Dec 2009 10:31:54 +0100 Christian Urban corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Tue, 22 Dec 2009 22:10:48 +0100 Christian Urban added "Highest Priority" category; and tuned slightly code
Tue, 22 Dec 2009 21:44:50 +0100 Christian Urban added a print_maps command; updated the keyword file accordingly
Tue, 22 Dec 2009 21:31:44 +0100 Christian Urban renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Tue, 22 Dec 2009 21:16:11 +0100 Christian Urban tuned
Tue, 22 Dec 2009 21:06:46 +0100 Christian Urban moved get_fun into quotient_term; this simplifies the overall including structure of the package
Tue, 22 Dec 2009 20:51:37 +0100 Christian Urban tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Tue, 22 Dec 2009 07:42:16 +0100 Christian Urban on the hunt for what condition raises which exception in the CLEVER CODE of calculate_inst
Tue, 22 Dec 2009 07:28:09 +0100 Christian Urban simplified calculate_instance; worked around some clever code; clever code is unfortunately still there...needs to be removed
Mon, 21 Dec 2009 23:13:40 +0100 Christian Urban used eq_reflection not with OF, but directly in @{thm ...}
Mon, 21 Dec 2009 23:01:58 +0100 Christian Urban cleaned a bit calculate_inst a bit; eta-contraction seems to be not necessary? (all examples go through)
Mon, 21 Dec 2009 22:36:31 +0100 Christian Urban get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Sun, 20 Dec 2009 00:53:35 +0100 Christian Urban on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Sun, 20 Dec 2009 00:26:53 +0100 Christian Urban renamed "quotient" command to "quotient_type"; needs new keyword file to be installed
Sun, 20 Dec 2009 00:15:40 +0100 Christian Urban this file is now obsolete; replaced by isar-keywords-quot.el
Sun, 20 Dec 2009 00:14:46 +0100 Christian Urban with "isabelle make keywords" you can create automatically a "quot" keywordfile, provided all Logics are in place
Sat, 19 Dec 2009 22:42:31 +0100 Christian Urban added a very old paper about Quotients in Isabelle (related work)
Sat, 19 Dec 2009 22:21:51 +0100 Christian Urban avoided global "open"s - replaced by local "open"s
Sat, 19 Dec 2009 22:09:57 +0100 Christian Urban small tuning
Sat, 19 Dec 2009 22:04:34 +0100 Christian Urban various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Thu, 17 Dec 2009 17:59:12 +0100 Christian Urban minor cleaning
Thu, 17 Dec 2009 14:58:33 +0100 Christian Urban moved the QuotMain code into two ML-files
(0) -240 +240 +1000 tip