Thu, 25 Feb 2010 12:32:15 +0100 Cezary Kaliszyk merge
Thu, 25 Feb 2010 12:30:50 +0100 Cezary Kaliszyk Move the eqvt code out of Terms and fixed induction for single-rule examples.
Thu, 25 Feb 2010 12:24:37 +0100 Christian Urban merged
Thu, 25 Feb 2010 11:51:34 +0100 Christian Urban a few simplifications
Thu, 25 Feb 2010 11:30:00 +0100 Christian Urban first attempt to make sense out of the core-haskell definition
Thu, 25 Feb 2010 12:15:11 +0100 Cezary Kaliszyk Code for proving eqvt, still in Terms.
Thu, 25 Feb 2010 09:41:14 +0100 Cezary Kaliszyk Use eqvt infrastructure.
Thu, 25 Feb 2010 09:22:29 +0100 Cezary Kaliszyk Simple function eqvt code.
Thu, 25 Feb 2010 08:40:52 +0100 Christian Urban added IsaMakefile...but so far included only a test for the parser
Thu, 25 Feb 2010 07:57:17 +0100 Christian Urban moved Quot package to Attic (still compiles there with "isabelle make")
Thu, 25 Feb 2010 07:48:57 +0100 Christian Urban merged
Thu, 25 Feb 2010 07:48:33 +0100 Christian Urban moved Nominal to "toplevel"
Thu, 25 Feb 2010 07:05:52 +0100 Cezary Kaliszyk Export perm_frees.
Wed, 24 Feb 2010 23:25:30 +0100 Cezary Kaliszyk Restructuring the code in Perm
Wed, 24 Feb 2010 18:38:49 +0100 Cezary Kaliszyk Simplified and finised eqvt proofs for t1 and t5
Wed, 24 Feb 2010 17:42:52 +0100 Cezary Kaliszyk merge
Wed, 24 Feb 2010 17:42:37 +0100 Cezary Kaliszyk Define lifted perms.
Wed, 24 Feb 2010 17:32:43 +0100 Christian Urban merged
Wed, 24 Feb 2010 17:32:22 +0100 Christian Urban parsing and definition of raw datatype and bv-function work (not very beautiful)
Wed, 24 Feb 2010 15:39:17 +0100 Cezary Kaliszyk With permute_rsp we can lift the instance proofs :).
Wed, 24 Feb 2010 15:36:07 +0100 Cezary Kaliszyk Note the instance proofs, since they can be easily lifted.
Wed, 24 Feb 2010 15:13:22 +0100 Cezary Kaliszyk More refactoring and removed references to the global simpset in Perm.
Wed, 24 Feb 2010 14:55:09 +0100 Cezary Kaliszyk Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
Wed, 24 Feb 2010 14:37:51 +0100 Cezary Kaliszyk Regularize finite support proof for trm1
Wed, 24 Feb 2010 14:09:34 +0100 Cezary Kaliszyk Made the fv-supp proof much more straightforward.
Wed, 24 Feb 2010 12:06:55 +0100 Cezary Kaliszyk Regularize the proofs about finite support.
Wed, 24 Feb 2010 11:28:34 +0100 Cezary Kaliszyk Respects of permute and constructors.
Wed, 24 Feb 2010 11:03:30 +0100 Cezary Kaliszyk Generate fv_rsp automatically.
Wed, 24 Feb 2010 10:59:31 +0100 Cezary Kaliszyk Define the constants automatically.
Wed, 24 Feb 2010 10:47:41 +0100 Cezary Kaliszyk Rename also the lifted types to non-capital.
Wed, 24 Feb 2010 10:44:38 +0100 Cezary Kaliszyk Use the infrastructure in LF. Much shorter :).
Wed, 24 Feb 2010 10:38:45 +0100 Cezary Kaliszyk Final synchronization of names.
Wed, 24 Feb 2010 10:25:59 +0100 Cezary Kaliszyk LF renaming part 3 (proper names of alpha equvalences)
Wed, 24 Feb 2010 10:08:54 +0100 Cezary Kaliszyk LF renaming part 2 (proper fv functions)
Wed, 24 Feb 2010 09:58:44 +0100 Cezary Kaliszyk merge
Wed, 24 Feb 2010 09:58:12 +0100 Cezary Kaliszyk LF renaming part1.
Wed, 24 Feb 2010 09:56:32 +0100 Christian Urban merged
Wed, 24 Feb 2010 09:56:12 +0100 Christian Urban parsing of function definitions almost works now; still an error with undefined constants
Tue, 23 Feb 2010 18:28:48 +0100 Cezary Kaliszyk merge
Tue, 23 Feb 2010 18:27:32 +0100 Cezary Kaliszyk rsp for bv; the only issue is that it requires an appropriate induction principle.
Tue, 23 Feb 2010 16:32:04 +0100 Christian Urban merged
Tue, 23 Feb 2010 16:31:40 +0100 Christian Urban declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
Tue, 23 Feb 2010 16:12:30 +0100 Cezary Kaliszyk rsp infrastructure.
Tue, 23 Feb 2010 14:20:42 +0100 Cezary Kaliszyk merge
Tue, 23 Feb 2010 14:19:44 +0100 Cezary Kaliszyk Progress towards automatic rsp of constants and fv.
Tue, 23 Feb 2010 13:33:01 +0100 Christian Urban merged
Tue, 23 Feb 2010 13:32:35 +0100 Christian Urban "raw"-ified the term-constructors and types given in the specification
Tue, 23 Feb 2010 12:49:45 +0100 Cezary Kaliszyk Looking at proving the rsp rules automatically.
Tue, 23 Feb 2010 11:56:47 +0100 Cezary Kaliszyk Minor beutification.
Tue, 23 Feb 2010 11:22:06 +0100 Cezary Kaliszyk Define the quotient from ML
Tue, 23 Feb 2010 10:47:14 +0100 Cezary Kaliszyk All works in LF but will require renaming.
Tue, 23 Feb 2010 09:34:41 +0100 Cezary Kaliszyk Reordering in LF.
Tue, 23 Feb 2010 09:31:59 +0100 Cezary Kaliszyk Fixes for auxiliary datatypes.
Mon, 22 Feb 2010 18:09:44 +0100 Cezary Kaliszyk Fixed pseudo_injectivity for trm4
Mon, 22 Feb 2010 17:19:28 +0100 Cezary Kaliszyk Testing auto equivp code.
Mon, 22 Feb 2010 16:44:58 +0100 Cezary Kaliszyk A tactic for final equivp
Mon, 22 Feb 2010 16:16:04 +0100 Cezary Kaliszyk More equivp infrastructure.
Mon, 22 Feb 2010 15:41:30 +0100 Cezary Kaliszyk tactify transp
Mon, 22 Feb 2010 15:09:53 +0100 Cezary Kaliszyk export the reflp and symp tacs.
Mon, 22 Feb 2010 15:03:48 +0100 Cezary Kaliszyk Generalize atom_trans and atom_sym.
Mon, 22 Feb 2010 14:50:53 +0100 Cezary Kaliszyk Some progress about transp
Mon, 22 Feb 2010 13:41:13 +0100 Cezary Kaliszyk alpha-symmetric addons.
Mon, 22 Feb 2010 12:12:32 +0100 Cezary Kaliszyk alpha reflexivity
Mon, 22 Feb 2010 10:57:39 +0100 Cezary Kaliszyk Renaming.
Mon, 22 Feb 2010 10:39:05 +0100 Cezary Kaliszyk Added missing description.
Mon, 22 Feb 2010 10:16:13 +0100 Cezary Kaliszyk Added Brian's suggestion.
Mon, 22 Feb 2010 09:55:43 +0100 Cezary Kaliszyk Update TODO
Sun, 21 Feb 2010 22:39:11 +0100 Cezary Kaliszyk Removed bindings 'in itself' where possible.
Sat, 20 Feb 2010 06:31:03 +0100 Cezary Kaliszyk Some adaptation
Fri, 19 Feb 2010 17:50:43 +0100 Cezary Kaliszyk proof cleaning and standardizing.
Fri, 19 Feb 2010 16:45:24 +0100 Cezary Kaliszyk Automatic production and proving of pseudo-injectivity.
Fri, 19 Feb 2010 12:05:58 +0100 Cezary Kaliszyk Experiments for the pseudo-injectivity tactic.
Fri, 19 Feb 2010 10:26:38 +0100 Cezary Kaliszyk merge
Fri, 19 Feb 2010 10:17:35 +0100 Cezary Kaliszyk Constructing alpha_inj goal.
Thu, 18 Feb 2010 23:07:52 +0100 Christian Urban merged
Thu, 18 Feb 2010 23:07:28 +0100 Christian Urban start work with the parser
Thu, 18 Feb 2010 18:33:53 +0100 Cezary Kaliszyk Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Thu, 18 Feb 2010 15:03:09 +0100 Cezary Kaliszyk First (non-working) version of alpha-equivalence
Thu, 18 Feb 2010 13:36:38 +0100 Cezary Kaliszyk Description of the fv procedure.
Thu, 18 Feb 2010 12:06:59 +0100 Cezary Kaliszyk Testing auto constant lifting.
Thu, 18 Feb 2010 11:28:20 +0100 Cezary Kaliszyk Fix for new Isabelle (primrec)
Thu, 18 Feb 2010 11:19:16 +0100 Cezary Kaliszyk Automatic lifting of constants.
Thu, 18 Feb 2010 10:01:48 +0100 Cezary Kaliszyk Changed back to original version of trm5
Thu, 18 Feb 2010 10:00:58 +0100 Cezary Kaliszyk The alternate version of trm5 with additional binding. All proofs work the same.
Thu, 18 Feb 2010 09:46:38 +0100 Cezary Kaliszyk Code for handling atom sets.
Thu, 18 Feb 2010 08:43:13 +0100 Cezary Kaliszyk Replace Terms by Terms2.
Thu, 18 Feb 2010 08:37:45 +0100 Cezary Kaliszyk Fixed proofs in Terms2 and found a mistake in Terms.
Wed, 17 Feb 2010 17:51:35 +0100 Cezary Kaliszyk Terms2 with bindings for binders synchronized with bindings they are used in.
Wed, 17 Feb 2010 17:29:26 +0100 Cezary Kaliszyk Cleaning of proofs in Terms.
Wed, 17 Feb 2010 16:22:16 +0100 Cezary Kaliszyk Testing Fv
Wed, 17 Feb 2010 15:52:08 +0100 Cezary Kaliszyk Fix the strong induction principle.
Wed, 17 Feb 2010 15:45:03 +0100 Cezary Kaliszyk Reorder
Wed, 17 Feb 2010 15:28:50 +0100 Cezary Kaliszyk Add bindings of recursive types by free_variables.
Wed, 17 Feb 2010 15:20:22 +0100 Cezary Kaliszyk Bindings adapted to multiple defined datatypes.
Wed, 17 Feb 2010 15:00:04 +0100 Cezary Kaliszyk Reorganization
Wed, 17 Feb 2010 14:44:32 +0100 Cezary Kaliszyk Now should work.
Wed, 17 Feb 2010 14:35:06 +0100 Cezary Kaliszyk Some optimizations and fixes.
Wed, 17 Feb 2010 14:17:02 +0100 Cezary Kaliszyk Simplified format of bindings.
Wed, 17 Feb 2010 13:56:31 +0100 Cezary Kaliszyk Tested the Perm code; works everywhere in Terms.
Wed, 17 Feb 2010 13:54:35 +0100 Cezary Kaliszyk Wrapped the permutation code.
Wed, 17 Feb 2010 10:20:26 +0100 Cezary Kaliszyk Description of intended bindings.
Wed, 17 Feb 2010 10:12:01 +0100 Cezary Kaliszyk Code for generating the fv function, no bindings yet.
Wed, 17 Feb 2010 09:27:02 +0100 Cezary Kaliszyk merge
Wed, 17 Feb 2010 09:26:49 +0100 Cezary Kaliszyk indent
Wed, 17 Feb 2010 09:26:38 +0100 Cezary Kaliszyk merge
Wed, 17 Feb 2010 09:26:10 +0100 Cezary Kaliszyk Simplifying perm_eq
Tue, 16 Feb 2010 15:13:14 +0100 Cezary Kaliszyk merge
Tue, 16 Feb 2010 15:12:31 +0100 Cezary Kaliszyk indenting
Tue, 16 Feb 2010 15:12:49 +0100 Cezary Kaliszyk Minor
Tue, 16 Feb 2010 14:57:39 +0100 Cezary Kaliszyk Merge
Tue, 16 Feb 2010 14:57:22 +0100 Cezary Kaliszyk Ported Stefan's permutation code, still needs some localizing.
Mon, 15 Feb 2010 16:54:09 +0100 Cezary Kaliszyk merge
Mon, 15 Feb 2010 16:53:51 +0100 Cezary Kaliszyk Removed varifyT.
Mon, 15 Feb 2010 17:02:46 +0100 Christian Urban merged
Mon, 15 Feb 2010 17:02:26 +0100 Christian Urban 2-spaces rule (where it makes sense)
Mon, 15 Feb 2010 16:52:32 +0100 Cezary Kaliszyk merge
Mon, 15 Feb 2010 16:51:30 +0100 Cezary Kaliszyk Fixed the definition of less and finished the missing proof.
Mon, 15 Feb 2010 16:50:11 +0100 Christian Urban further tuning
Mon, 15 Feb 2010 16:37:48 +0100 Christian Urban small tuning
Mon, 15 Feb 2010 16:28:07 +0100 Christian Urban tuned the parsing and testing code in quotient_def.ML; cleaned out old stuff in AbsRepTest.thy
Mon, 15 Feb 2010 14:58:03 +0100 Cezary Kaliszyk der_bname -> derived_bname
Mon, 15 Feb 2010 14:51:17 +0100 Cezary Kaliszyk Names of files.
Mon, 15 Feb 2010 14:28:03 +0100 Cezary Kaliszyk Finished introducing the binding.
Mon, 15 Feb 2010 13:40:03 +0100 Cezary Kaliszyk Synchronize the commands.
Mon, 15 Feb 2010 12:23:02 +0100 Cezary Kaliszyk Passing the binding to quotient_def
Mon, 15 Feb 2010 12:15:14 +0100 Cezary Kaliszyk Added a binding to the parser.
Mon, 15 Feb 2010 10:25:17 +0100 Cezary Kaliszyk Second inline
Mon, 15 Feb 2010 10:11:26 +0100 Cezary Kaliszyk remove one-line wrapper.
Fri, 12 Feb 2010 16:27:25 +0100 Cezary Kaliszyk Undid the read_terms change; now compiles.
Fri, 12 Feb 2010 16:06:09 +0100 Cezary Kaliszyk merge
Fri, 12 Feb 2010 16:04:10 +0100 Cezary Kaliszyk renamed 'as' to 'is' everywhere.
Fri, 12 Feb 2010 15:50:43 +0100 Cezary Kaliszyk "is" defined as the keyword
Fri, 12 Feb 2010 15:06:20 +0100 Christian Urban moved "strange" lemma to quotient_tacs; marked a number of lemmas as unused; tuned
Fri, 12 Feb 2010 12:06:09 +0100 Cezary Kaliszyk The lattice instantiations are gone from Isabelle/Main, so
Thu, 11 Feb 2010 17:58:06 +0100 Cezary Kaliszyk the lam/bla example.
Thu, 11 Feb 2010 16:54:04 +0100 Cezary Kaliszyk Finished a working foo/bar.
Thu, 11 Feb 2010 16:05:15 +0100 Cezary Kaliszyk fv_foo is not regular.
Thu, 11 Feb 2010 15:08:45 +0100 Cezary Kaliszyk Testing foo/bar
Thu, 11 Feb 2010 14:23:26 +0100 Cezary Kaliszyk Even when bv = fv it still doesn't lift.
Thu, 11 Feb 2010 14:02:34 +0100 Cezary Kaliszyk Added the missing syntax file
Thu, 11 Feb 2010 14:00:00 +0100 Cezary Kaliszyk Notation available locally
Thu, 11 Feb 2010 10:06:02 +0100 Cezary Kaliszyk Main renaming + fixes for new Isabelle in IntEx2.
Thu, 11 Feb 2010 09:23:59 +0100 Cezary Kaliszyk Merging QuotBase into QuotMain.
Wed, 10 Feb 2010 21:39:40 +0100 Christian Urban removed dead code
Wed, 10 Feb 2010 20:35:54 +0100 Christian Urban cleaned a bit
Wed, 10 Feb 2010 17:22:18 +0100 Cezary Kaliszyk lowercase locale
Wed, 10 Feb 2010 17:10:52 +0100 Cezary Kaliszyk hg-added the added file.
Wed, 10 Feb 2010 17:02:29 +0100 Cezary Kaliszyk Changes from Makarius's code review + some noticed fixes.
Wed, 10 Feb 2010 12:30:26 +0100 Cezary Kaliszyk example with a respectful bn function defined over the type itself
Wed, 10 Feb 2010 11:53:15 +0100 Cezary Kaliszyk Finishe the renaming.
Wed, 10 Feb 2010 11:39:22 +0100 Cezary Kaliszyk Another mistake found with OTT.
Wed, 10 Feb 2010 11:31:53 +0100 Cezary Kaliszyk merge
Wed, 10 Feb 2010 11:31:43 +0100 Cezary Kaliszyk Fixed rbv6, when translating to OTT.
Wed, 10 Feb 2010 11:27:49 +0100 Cezary Kaliszyk Some cleaning of proofs.
Wed, 10 Feb 2010 11:11:06 +0100 Christian Urban merged again
Wed, 10 Feb 2010 11:10:44 +0100 Christian Urban merged
Wed, 10 Feb 2010 11:09:30 +0100 Cezary Kaliszyk more minor space and bracket modifications.
Wed, 10 Feb 2010 10:55:14 +0100 Cezary Kaliszyk More changes according to the standards.
Wed, 10 Feb 2010 10:36:47 +0100 Cezary Kaliszyk A concrete example, with a proof that rbv is not regular and
Tue, 09 Feb 2010 19:08:08 +0100 Christian Urban proper declaration of types and terms during parsing (removes the varifyT when storing data)
Tue, 09 Feb 2010 17:26:28 +0100 Christian Urban merged
Tue, 09 Feb 2010 17:26:08 +0100 Christian Urban slight correction
Tue, 09 Feb 2010 17:26:00 +0100 Cezary Kaliszyk merge
Tue, 09 Feb 2010 17:24:08 +0100 Cezary Kaliszyk More about trm6
Tue, 09 Feb 2010 17:17:06 +0100 Christian Urban merged
Tue, 09 Feb 2010 17:05:07 +0100 Cezary Kaliszyk the specifications of the respects.
Tue, 09 Feb 2010 16:44:06 +0100 Cezary Kaliszyk trm6 with the 'Foo' constructor.
Tue, 09 Feb 2010 16:10:08 +0100 Cezary Kaliszyk removing unnecessary brackets
Tue, 09 Feb 2010 15:55:58 +0100 Cezary Kaliszyk More indentation cleaning.
Tue, 09 Feb 2010 15:43:39 +0100 Cezary Kaliszyk 'exc' -> 'exn' and more name and space cleaning.
Tue, 09 Feb 2010 15:36:23 +0100 Cezary Kaliszyk Fully qualified exception names.
Tue, 09 Feb 2010 15:28:30 +0100 Cezary Kaliszyk merge
Tue, 09 Feb 2010 15:28:15 +0100 Cezary Kaliszyk More indentation, names and todo cleaning in the quotient package
Tue, 09 Feb 2010 15:20:52 +0100 Christian Urban merged
Tue, 09 Feb 2010 15:20:40 +0100 Christian Urban a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Tue, 09 Feb 2010 11:40:32 +0100 Christian Urban minor tuning
Tue, 09 Feb 2010 14:32:37 +0100 Cezary Kaliszyk Explicitly marked what is bound.
Tue, 09 Feb 2010 12:22:00 +0100 Cezary Kaliszyk Cleaning and updating in Terms.
Tue, 09 Feb 2010 11:22:34 +0100 Cezary Kaliszyk Looking at the trm2 example
Tue, 09 Feb 2010 10:48:42 +0100 Cezary Kaliszyk Fixed pattern matching, now the test in Abs works correctly.
Mon, 08 Feb 2010 13:50:52 +0100 Christian Urban added a test case
Mon, 08 Feb 2010 13:13:20 +0100 Christian Urban merged
Mon, 08 Feb 2010 13:12:55 +0100 Christian Urban moved some lemmas to Nominal; updated all files
Mon, 08 Feb 2010 13:04:29 +0100 Cezary Kaliszyk merge
Mon, 08 Feb 2010 13:04:13 +0100 Cezary Kaliszyk Comments.
Mon, 08 Feb 2010 11:56:22 +0100 Christian Urban slightly tuned
Mon, 08 Feb 2010 11:41:25 +0100 Cezary Kaliszyk Proper context fixes lifting inside instantiations.
Mon, 08 Feb 2010 10:47:19 +0100 Cezary Kaliszyk Fixed the context import/export and simplified LFex.
Mon, 08 Feb 2010 06:27:20 +0100 Christian Urban added 2 papers about core haskell
Sun, 07 Feb 2010 10:20:29 +0100 Christian Urban fixed lemma name
Sun, 07 Feb 2010 10:16:21 +0100 Christian Urban updated to latest Nominal2
Sat, 06 Feb 2010 12:58:56 +0100 Christian Urban minor
Sat, 06 Feb 2010 10:04:56 +0100 Christian Urban some tuning
Fri, 05 Feb 2010 15:17:21 +0100 Cezary Kaliszyk merge
Fri, 05 Feb 2010 14:52:27 +0100 Cezary Kaliszyk merge
Fri, 05 Feb 2010 10:32:21 +0100 Cezary Kaliszyk Fixes for Bex1 removal.
Fri, 05 Feb 2010 15:09:49 +0100 Cezary Kaliszyk Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
Fri, 05 Feb 2010 11:37:18 +0100 Cezary Kaliszyk A procedure that properly instantiates the types too.
Fri, 05 Feb 2010 11:28:49 +0100 Cezary Kaliszyk More code abstracted away
Fri, 05 Feb 2010 11:19:21 +0100 Cezary Kaliszyk A bit more intelligent and cleaner code.
Fri, 05 Feb 2010 11:09:43 +0100 Cezary Kaliszyk merge
Fri, 05 Feb 2010 10:45:49 +0100 Cezary Kaliszyk A proper version of the attribute
Fri, 05 Feb 2010 09:06:49 +0100 Christian Urban merged
Fri, 05 Feb 2010 09:06:27 +0100 Christian Urban eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Thu, 04 Feb 2010 18:09:20 +0100 Cezary Kaliszyk The automatic lifting translation function, still with dummy types,
Thu, 04 Feb 2010 17:58:23 +0100 Cezary Kaliszyk Quotdata_dest needed for lifting theorem translation.
Thu, 04 Feb 2010 17:39:04 +0100 Christian Urban fixed (permute_eqvt in eqvts makes this simpset always looping)
Thu, 04 Feb 2010 15:19:24 +0100 Christian Urban rollback of the test
Thu, 04 Feb 2010 15:16:34 +0100 Christian Urban linked versions - instead of copies
Thu, 04 Feb 2010 14:55:52 +0100 Christian Urban merged
Thu, 04 Feb 2010 14:55:21 +0100 Christian Urban restored the old behaviour of having an eqvts list; the transformed theorems are stored in eqvts_raw
Wed, 03 Feb 2010 18:28:50 +0100 Cezary Kaliszyk More let-rec experiments
Wed, 03 Feb 2010 17:36:25 +0100 Christian Urban proposal for an alpha equivalence
Wed, 03 Feb 2010 15:17:29 +0100 Cezary Kaliszyk Lets different.
Wed, 03 Feb 2010 14:39:19 +0100 Cezary Kaliszyk Simplified the proof.
Wed, 03 Feb 2010 14:36:48 +0100 Christian Urban merged
Wed, 03 Feb 2010 14:36:22 +0100 Christian Urban proved that bv for lists respects alpha for terms
Wed, 03 Feb 2010 14:28:00 +0100 Cezary Kaliszyk Finished remains on the let proof.
Wed, 03 Feb 2010 14:22:25 +0100 Cezary Kaliszyk merge
Wed, 03 Feb 2010 14:19:53 +0100 Cezary Kaliszyk Lets are ok.
Wed, 03 Feb 2010 14:15:07 +0100 Christian Urban merged
Wed, 03 Feb 2010 14:12:50 +0100 Christian Urban added type-scheme example
Wed, 03 Feb 2010 13:00:37 +0100 Cezary Kaliszyk merge
Wed, 03 Feb 2010 13:00:07 +0100 Cezary Kaliszyk Definitions for trm5
Wed, 03 Feb 2010 12:58:02 +0100 Christian Urban another adaptation for the eqvt-change
Wed, 03 Feb 2010 12:45:06 +0100 Christian Urban merged
Wed, 03 Feb 2010 12:44:29 +0100 Christian Urban fixed proofs that broke because of eqvt
Wed, 03 Feb 2010 12:34:53 +0100 Cezary Kaliszyk Minor fix.
Wed, 03 Feb 2010 12:34:01 +0100 Cezary Kaliszyk merge
Wed, 03 Feb 2010 12:29:45 +0100 Cezary Kaliszyk alpha5 pseudo-injective
Wed, 03 Feb 2010 12:31:58 +0100 Christian Urban fixed proofs in Abs.thy
Wed, 03 Feb 2010 12:13:22 +0100 Christian Urban merged
Wed, 03 Feb 2010 12:06:10 +0100 Christian Urban added a first eqvt_tac which pushes permutations inside terms
Wed, 03 Feb 2010 12:11:23 +0100 Cezary Kaliszyk The alpha-equivalence relation for let-rec. Not sure if correct...
Wed, 03 Feb 2010 11:47:37 +0100 Cezary Kaliszyk Starting with a let-rec example.
Wed, 03 Feb 2010 11:21:34 +0100 Cezary Kaliszyk Minor
Wed, 03 Feb 2010 10:50:24 +0100 Cezary Kaliszyk Some cleaning and eqvt proof
Wed, 03 Feb 2010 09:25:21 +0100 Cezary Kaliszyk The trm1_support lemma explicitly and stated a strong induction principle.
Wed, 03 Feb 2010 08:32:24 +0100 Cezary Kaliszyk More ingredients in Terms.
Tue, 02 Feb 2010 17:10:42 +0100 Cezary Kaliszyk Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
Tue, 02 Feb 2010 16:51:00 +0100 Cezary Kaliszyk More in Terms
Tue, 02 Feb 2010 14:55:07 +0100 Cezary Kaliszyk First experiments in Terms.
Tue, 02 Feb 2010 13:10:46 +0100 Cezary Kaliszyk LF ported to alpha_gen, equivp solved and one of the missing proofs in support<-> fv solved. Still some supp properties left.
Tue, 02 Feb 2010 12:48:12 +0100 Cezary Kaliszyk Disambiguating the syntax.
Tue, 02 Feb 2010 12:36:01 +0100 Cezary Kaliszyk Minor uncommited changes from LamEx2.
Tue, 02 Feb 2010 11:56:37 +0100 Cezary Kaliszyk Some equivariance machinery that comes useful in LF.
Tue, 02 Feb 2010 11:23:17 +0100 Cezary Kaliszyk Generalized the eqvt proof for single binders.
Tue, 02 Feb 2010 10:43:48 +0100 Cezary Kaliszyk With induct instead of induct_tac, just one induction is sufficient.
Tue, 02 Feb 2010 10:20:54 +0100 Cezary Kaliszyk General alpha_gen_trans for one-variable abstraction.
Tue, 02 Feb 2010 09:51:39 +0100 Cezary Kaliszyk With unfolding Rep/Abs_eqvt no longer needed.
Tue, 02 Feb 2010 08:16:34 +0100 Cezary Kaliszyk Lam2 finished apart from Rep_eqvt.
Mon, 01 Feb 2010 20:02:44 +0100 Cezary Kaliszyk merge
Mon, 01 Feb 2010 16:05:59 +0100 Cezary Kaliszyk All should be ok now.
Mon, 01 Feb 2010 18:57:39 +0100 Christian Urban repaired according to changes in Abs.thy
Mon, 01 Feb 2010 18:57:20 +0100 Christian Urban added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Mon, 01 Feb 2010 16:46:07 +0100 Christian Urban cleaned
Mon, 01 Feb 2010 16:23:47 +0100 Christian Urban updated from huffman
Mon, 01 Feb 2010 16:13:24 +0100 Christian Urban updated from nominal-huffman
Mon, 01 Feb 2010 15:57:37 +0100 Cezary Kaliszyk Fixed wrong rename.
Mon, 01 Feb 2010 15:46:25 +0100 Cezary Kaliszyk merge
Mon, 01 Feb 2010 15:45:40 +0100 Cezary Kaliszyk Lambda based on alpha_gen, under construction.
Mon, 01 Feb 2010 15:32:20 +0100 Christian Urban updated from huffman - repo
Mon, 01 Feb 2010 13:00:01 +0100 Christian Urban renamed Abst/abst to Abs/abs
Mon, 01 Feb 2010 12:48:18 +0100 Christian Urban got rid of RAbst type - is now just pairs
Mon, 01 Feb 2010 12:06:46 +0100 Cezary Kaliszyk Monotonicity of ~~gen, needed for using it in inductive definitions.
Mon, 01 Feb 2010 11:39:59 +0100 Cezary Kaliszyk The current state of fv vs supp proofs in LF.
Mon, 01 Feb 2010 11:16:31 +0100 Cezary Kaliszyk merge
Mon, 01 Feb 2010 11:16:13 +0100 Cezary Kaliszyk More proofs in the LF example.
Mon, 01 Feb 2010 11:00:51 +0100 Christian Urban merged
Mon, 01 Feb 2010 10:00:03 +0100 Christian Urban slight tuning
Mon, 01 Feb 2010 09:47:46 +0100 Christian Urban renamed function according to the name of the constant
Mon, 01 Feb 2010 09:04:22 +0100 Christian Urban fixed problem with Bex1_rel renaming
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
Wed, 16 Dec 2009 14:28:48 +0100 Christian Urban complete fix for IsaMakefile
Wed, 16 Dec 2009 14:26:14 +0100 Christian Urban first fix
Wed, 16 Dec 2009 14:09:03 +0100 Christian Urban merged
Wed, 16 Dec 2009 14:08:42 +0100 Christian Urban added a paper for possible notes
Wed, 16 Dec 2009 12:15:41 +0100 Cezary Kaliszyk Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Tue, 15 Dec 2009 16:40:00 +0100 Cezary Kaliszyk lambda_prs & solve_quotient_assum cleaned.
Tue, 15 Dec 2009 15:38:17 +0100 Christian Urban some commenting
Mon, 14 Dec 2009 14:24:08 +0100 Cezary Kaliszyk Fixed previous commit.
Mon, 14 Dec 2009 13:59:08 +0100 Cezary Kaliszyk merge
Mon, 14 Dec 2009 13:58:51 +0100 Cezary Kaliszyk Moved DETERM inside Repeat & added SOLVE around quotient_tac.
Mon, 14 Dec 2009 13:57:39 +0100 Cezary Kaliszyk merge.
Mon, 14 Dec 2009 13:56:24 +0100 Cezary Kaliszyk FIXME/TODO.
Mon, 14 Dec 2009 10:19:27 +0100 Cezary Kaliszyk reply to question in code
Mon, 14 Dec 2009 10:12:23 +0100 Cezary Kaliszyk Reply in code.
Mon, 14 Dec 2009 10:09:49 +0100 Cezary Kaliszyk Replies to questions from the weekend: Uncommenting the renamed theorem commented out in 734.
Sun, 13 Dec 2009 02:47:47 +0100 Christian Urban a few code annotations
Sun, 13 Dec 2009 02:35:34 +0100 Christian Urban another pass on apply_rsp
Sun, 13 Dec 2009 01:56:19 +0100 Christian Urban managed to simplify apply_rsp
Sat, 12 Dec 2009 18:43:42 +0100 Christian Urban tried to simplify apply_rsp_tac; failed at the moment; added some questions
Sat, 12 Dec 2009 18:01:22 +0100 Christian Urban some trivial changes
Sat, 12 Dec 2009 16:40:29 +0100 Christian Urban trivial cleaning of make_inst
Sat, 12 Dec 2009 15:23:58 +0100 Christian Urban tried to improve test; but fails
Sat, 12 Dec 2009 15:08:25 +0100 Christian Urban merged
Sat, 12 Dec 2009 15:07:59 +0100 Christian Urban annotated some questions to the code; some simple changes
Sat, 12 Dec 2009 14:57:34 +0100 Cezary Kaliszyk Answering the question in code.
Sat, 12 Dec 2009 13:54:01 +0100 Christian Urban merged
Sat, 12 Dec 2009 13:53:46 +0100 Christian Urban trivial
Sat, 12 Dec 2009 02:01:33 +0100 Christian Urban tuned code
Sat, 12 Dec 2009 09:27:06 +0100 Cezary Kaliszyk Minor
Sat, 12 Dec 2009 05:12:50 +0100 Cezary Kaliszyk Some proofs.
Sat, 12 Dec 2009 04:48:43 +0100 Cezary Kaliszyk Proof of finite_set_storng_cases_raw.
Sat, 12 Dec 2009 04:25:47 +0100 Cezary Kaliszyk A bracket was missing; with it proved the 'definitely false' lemma.
Sat, 12 Dec 2009 01:44:56 +0100 Christian Urban renamed quotient.ML to quotient_typ.ML
Fri, 11 Dec 2009 19:22:30 +0100 Christian Urban merged
Fri, 11 Dec 2009 19:19:50 +0100 Christian Urban tuned
Fri, 11 Dec 2009 19:19:24 +0100 Christian Urban started to have a look at it; redefined the relation
Fri, 11 Dec 2009 17:59:29 +0100 Cezary Kaliszyk More name and indentation cleaning.
Fri, 11 Dec 2009 17:22:26 +0100 Cezary Kaliszyk Merge + Added LarryInt & Fset3 to tests.
Fri, 11 Dec 2009 17:19:38 +0100 Cezary Kaliszyk Renaming
Fri, 11 Dec 2009 17:03:52 +0100 Christian Urban merged
Fri, 11 Dec 2009 17:03:34 +0100 Christian Urban deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Fri, 11 Dec 2009 16:32:40 +0100 Cezary Kaliszyk FSet3 minor fixes + cases
Fri, 11 Dec 2009 15:58:15 +0100 Christian Urban added Int example from Larry
Fri, 11 Dec 2009 15:49:15 +0100 Cezary Kaliszyk Added FSet3 with a formalisation of finite sets based on Michael's one.
Fri, 11 Dec 2009 13:51:08 +0100 Cezary Kaliszyk Updated TODO list together.
Fri, 11 Dec 2009 11:32:29 +0100 Cezary Kaliszyk Merge
Fri, 11 Dec 2009 11:30:00 +0100 Cezary Kaliszyk More theorem renaming.
Fri, 11 Dec 2009 11:25:52 +0100 Cezary Kaliszyk Renamed theorems in IntEx2 to conform to names in Int.
Fri, 11 Dec 2009 11:19:41 +0100 Cezary Kaliszyk Updated comments.
Fri, 11 Dec 2009 11:14:05 +0100 Christian Urban merged
Fri, 11 Dec 2009 11:12:53 +0100 Christian Urban tuned
Fri, 11 Dec 2009 10:57:46 +0100 Christian Urban renamed Larrys example
Fri, 11 Dec 2009 11:08:58 +0100 Cezary Kaliszyk New syntax for definitions.
Fri, 11 Dec 2009 08:28:41 +0100 Christian Urban changed error message
Fri, 11 Dec 2009 06:58:31 +0100 Christian Urban reformulated the lemma lifting_procedure as ML value; gave better warning message for injection case
Thu, 10 Dec 2009 19:05:56 +0100 Christian Urban slightly tuned
Thu, 10 Dec 2009 18:28:41 +0100 Christian Urban merged
Thu, 10 Dec 2009 18:28:30 +0100 Christian Urban added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Thu, 10 Dec 2009 16:56:03 +0100 Christian Urban added maps-printout and tuned some comments
Thu, 10 Dec 2009 14:35:06 +0100 Cezary Kaliszyk Option and Sum quotients.
Thu, 10 Dec 2009 12:25:12 +0100 Cezary Kaliszyk Regularized the hard lemma.
Thu, 10 Dec 2009 11:19:34 +0100 Cezary Kaliszyk Simplification of Babses for regularize; will probably become injection
Thu, 10 Dec 2009 10:54:45 +0100 Cezary Kaliszyk Found the problem with ttt3.
Thu, 10 Dec 2009 10:36:05 +0100 Cezary Kaliszyk minor
Thu, 10 Dec 2009 10:21:51 +0100 Cezary Kaliszyk Moved Unused part of locale to Unused QuotMain.
Thu, 10 Dec 2009 08:55:30 +0100 Cezary Kaliszyk Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
Thu, 10 Dec 2009 08:44:01 +0100 Cezary Kaliszyk Removed 'Presburger' as it introduces int & other minor cleaning in Int2.
Thu, 10 Dec 2009 05:11:53 +0100 Christian Urban more tuning
Thu, 10 Dec 2009 05:02:34 +0100 Christian Urban tuned
Thu, 10 Dec 2009 04:53:48 +0100 Christian Urban simplified the instantiation of QUOT_TRUE in procedure_tac
Thu, 10 Dec 2009 04:35:08 +0100 Christian Urban completed previous commit
Thu, 10 Dec 2009 04:34:24 +0100 Christian Urban deleted DT/NDT diagnostic code
Thu, 10 Dec 2009 04:23:13 +0100 Christian Urban moved the interpretation code into Unused.thy
Thu, 10 Dec 2009 03:48:39 +0100 Christian Urban added an attempt to get a finite set theory
Thu, 10 Dec 2009 03:47:10 +0100 Christian Urban removed memb and used standard mem (member from List.thy)
Thu, 10 Dec 2009 03:25:42 +0100 Christian Urban merged
Thu, 10 Dec 2009 03:11:19 +0100 Christian Urban simplified proofs
Thu, 10 Dec 2009 02:46:08 +0100 Christian Urban removed quot_respect attribute of a non-standard lemma
Thu, 10 Dec 2009 02:42:09 +0100 Cezary Kaliszyk With int_of_nat as a quotient_def, lemmas about it can be easily lifted.
Thu, 10 Dec 2009 01:48:39 +0100 Christian Urban merged
Thu, 10 Dec 2009 01:47:55 +0100 Christian Urban naming in this file cannot be made to agree to the original (PROBLEM?)
Thu, 10 Dec 2009 01:39:47 +0100 Cezary Kaliszyk Lifted some kind of induction.
Wed, 09 Dec 2009 23:32:16 +0100 Christian Urban more proofs in IntEx2
Wed, 09 Dec 2009 22:43:11 +0100 Cezary Kaliszyk Finished one proof in IntEx2.
Wed, 09 Dec 2009 22:05:11 +0100 Christian Urban slightly more on IntEx2
Wed, 09 Dec 2009 20:35:52 +0100 Christian Urban proved (with a lot of pain) that times_raw is respectful
Wed, 09 Dec 2009 17:31:19 +0100 Christian Urban merged
Wed, 09 Dec 2009 17:31:01 +0100 Christian Urban fixed minor stupidity
Wed, 09 Dec 2009 17:16:39 +0100 Cezary Kaliszyk Exception handling.
Wed, 09 Dec 2009 17:05:33 +0100 Cezary Kaliszyk Code cleaning.
Wed, 09 Dec 2009 16:44:34 +0100 Cezary Kaliszyk merge
Wed, 09 Dec 2009 16:43:12 +0100 Cezary Kaliszyk foldr_rsp.
Wed, 09 Dec 2009 16:09:25 +0100 Christian Urban tuned
Wed, 09 Dec 2009 15:59:02 +0100 Cezary Kaliszyk merge
Wed, 09 Dec 2009 15:57:47 +0100 Cezary Kaliszyk Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Wed, 09 Dec 2009 15:35:21 +0100 Christian Urban deleted make_inst3
Wed, 09 Dec 2009 15:29:36 +0100 Christian Urban tuned
Wed, 09 Dec 2009 15:28:01 +0100 Christian Urban tuned
Wed, 09 Dec 2009 15:24:11 +0100 Christian Urban moved function and tuned comment
Wed, 09 Dec 2009 15:11:49 +0100 Christian Urban improved fun_map_conv
Wed, 09 Dec 2009 06:21:09 +0100 Cezary Kaliszyk Arbitrary number of fun_map_tacs.
Wed, 09 Dec 2009 05:59:49 +0100 Cezary Kaliszyk Temporarily repeated fun_map_tac 4 times. Cleaning for all examples work.
Wed, 09 Dec 2009 00:54:46 +0100 Christian Urban tuned code
Wed, 09 Dec 2009 00:03:18 +0100 Christian Urban tuned the examples and flagged the problematic cleaning lemmas in FSet
Tue, 08 Dec 2009 23:32:54 +0100 Christian Urban merged
Tue, 08 Dec 2009 23:30:47 +0100 Christian Urban implemented cleaning strategy with fun_map.simps on non-bounded variables; still a few rough edges
Tue, 08 Dec 2009 23:04:40 +0100 Cezary Kaliszyk merge
Tue, 08 Dec 2009 23:04:25 +0100 Cezary Kaliszyk manually cleaned the hard lemma.
Tue, 08 Dec 2009 22:26:01 +0100 Christian Urban merged
Tue, 08 Dec 2009 22:24:24 +0100 Christian Urban decoupled QuotProd from QuotMain and also started new cleaning strategy
Tue, 08 Dec 2009 22:05:01 +0100 Cezary Kaliszyk merge
Tue, 08 Dec 2009 22:03:34 +0100 Cezary Kaliszyk An example which is hard to lift because of the interplay between lambda_prs and unfolding.
Tue, 08 Dec 2009 22:02:14 +0100 Christian Urban proper formulation of all preservation theorems
Tue, 08 Dec 2009 20:55:55 +0100 Christian Urban started to reformulate preserve lemmas
Tue, 08 Dec 2009 20:34:00 +0100 Christian Urban properly set up the prs_rules
Tue, 08 Dec 2009 17:43:32 +0100 Christian Urban merged
Tue, 08 Dec 2009 17:40:58 +0100 Christian Urban added preserve rules to the cleaning_tac
Tue, 08 Dec 2009 17:39:34 +0100 Cezary Kaliszyk merge
Tue, 08 Dec 2009 17:35:04 +0100 Cezary Kaliszyk cleaning.
Tue, 08 Dec 2009 17:34:10 +0100 Christian Urban merged
Tue, 08 Dec 2009 17:33:51 +0100 Christian Urban chnaged syntax to "lifting theorem"
Tue, 08 Dec 2009 17:30:00 +0100 Christian Urban changed names of attributes
Tue, 08 Dec 2009 16:56:51 +0100 Cezary Kaliszyk merge
Tue, 08 Dec 2009 16:56:37 +0100 Cezary Kaliszyk Manual regularization of a goal in FSet.
Tue, 08 Dec 2009 16:36:01 +0100 Christian Urban merged
Tue, 08 Dec 2009 16:35:40 +0100 Christian Urban added methods for the lifting_tac and the other tacs
Tue, 08 Dec 2009 15:42:29 +0100 Cezary Kaliszyk make_inst3
Tue, 08 Dec 2009 15:12:55 +0100 Cezary Kaliszyk Merge
Tue, 08 Dec 2009 15:12:36 +0100 Cezary Kaliszyk trans2 replaced with equals_rsp_tac
Tue, 08 Dec 2009 14:00:48 +0100 Christian Urban corrected name of FSet in ROOT.ML
Tue, 08 Dec 2009 13:09:21 +0100 Cezary Kaliszyk Made fset work again to test all.
Tue, 08 Dec 2009 13:08:56 +0100 Cezary Kaliszyk Finished the proof of ttt2 and found bug in regularize when trying ttt3.
Tue, 08 Dec 2009 13:01:23 +0100 Cezary Kaliszyk Another lambda example theorem proved. Seems it starts working properly.
Tue, 08 Dec 2009 13:00:36 +0100 Cezary Kaliszyk Removed pattern from quot_rel_rsp, since list_rel and all used introduced ones cannot be patterned
Tue, 08 Dec 2009 12:59:38 +0100 Cezary Kaliszyk Proper checked map_rsp.
Tue, 08 Dec 2009 12:36:28 +0100 Cezary Kaliszyk Nitpick found a counterexample for one lemma.
Tue, 08 Dec 2009 11:59:16 +0100 Cezary Kaliszyk Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Tue, 08 Dec 2009 11:38:58 +0100 Cezary Kaliszyk It also regularizes.
Tue, 08 Dec 2009 11:28:04 +0100 Cezary Kaliszyk inj_repabs also works.
Tue, 08 Dec 2009 11:20:01 +0100 Cezary Kaliszyk merge
Tue, 08 Dec 2009 11:17:56 +0100 Cezary Kaliszyk An example of working cleaning for lambda lifting. Still not sure why Babs helps.
Tue, 08 Dec 2009 04:21:14 +0100 Christian Urban tuned
Tue, 08 Dec 2009 04:14:02 +0100 Christian Urban the lift_tac produces a warning message if one of the three automatic proofs fails
Tue, 08 Dec 2009 01:25:43 +0100 Christian Urban added a thm list for ids
Tue, 08 Dec 2009 01:00:21 +0100 Christian Urban removed a fixme: map_info is now checked
Mon, 07 Dec 2009 23:45:51 +0100 Christian Urban tuning of the code
Mon, 07 Dec 2009 21:54:14 +0100 Christian Urban merged
Mon, 07 Dec 2009 21:53:50 +0100 Christian Urban removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Mon, 07 Dec 2009 21:25:49 +0100 Cezary Kaliszyk 3 lambda examples in FSet. In the last one regularize_term fails.
Mon, 07 Dec 2009 21:21:57 +0100 Cezary Kaliszyk Handling of errors in lambda_prs_conv.
Mon, 07 Dec 2009 21:21:23 +0100 Cezary Kaliszyk babs_prs
Mon, 07 Dec 2009 18:49:14 +0100 Christian Urban clarified the function examples
Mon, 07 Dec 2009 17:57:33 +0100 Christian Urban first attempt to deal with Babs in regularise and cleaning (not yet working)
Mon, 07 Dec 2009 15:21:51 +0100 Christian Urban isabelle make tests all examples
Mon, 07 Dec 2009 15:18:44 +0100 Cezary Kaliszyk merge
Mon, 07 Dec 2009 15:18:00 +0100 Cezary Kaliszyk make_inst for lambda_prs where the second quotient is not identity.
Mon, 07 Dec 2009 14:37:10 +0100 Christian Urban added "end" to each example theory
Mon, 07 Dec 2009 14:35:45 +0100 Cezary Kaliszyk List moved after QuotMain
Mon, 07 Dec 2009 14:14:07 +0100 Christian Urban cleaning
Mon, 07 Dec 2009 14:12:29 +0100 Christian Urban final move
Mon, 07 Dec 2009 14:09:50 +0100 Christian Urban directory re-arrangement
Mon, 07 Dec 2009 14:00:36 +0100 Cezary Kaliszyk inj_repabs_tac handles Babs now.
Mon, 07 Dec 2009 12:14:25 +0100 Cezary Kaliszyk Fix of regularize for babs and proof of babs_rsp.
Mon, 07 Dec 2009 11:14:21 +0100 Cezary Kaliszyk Using pair_prs; debugging the error in regularize of a lambda.
Mon, 07 Dec 2009 08:45:04 +0100 Cezary Kaliszyk QuotProd with product_quotient and a 3 respects and preserves lemmas.
Mon, 07 Dec 2009 04:41:42 +0100 Cezary Kaliszyk merge
Mon, 07 Dec 2009 04:39:42 +0100 Cezary Kaliszyk 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Mon, 07 Dec 2009 02:34:24 +0100 Christian Urban simplified the regularize simproc
Mon, 07 Dec 2009 01:28:10 +0100 Christian Urban now simpler regularize_tac with added solver works
Mon, 07 Dec 2009 01:22:20 +0100 Christian Urban removed usage of HOL_basic_ss by using a slighly extended version of empty_ss
Mon, 07 Dec 2009 00:13:36 +0100 Christian Urban merged
Mon, 07 Dec 2009 00:07:23 +0100 Christian Urban fixed examples
Mon, 07 Dec 2009 00:03:12 +0100 Cezary Kaliszyk Fix IntEx2 for equiv_list
Sun, 06 Dec 2009 23:35:02 +0100 Christian Urban merged
Sun, 06 Dec 2009 23:32:27 +0100 Christian Urban working state again
Sun, 06 Dec 2009 13:41:42 +0100 Christian Urban added a theorem list for equivalence theorems
Sun, 06 Dec 2009 22:58:03 +0100 Cezary Kaliszyk Merge
Sun, 06 Dec 2009 22:57:44 +0100 Cezary Kaliszyk Name changes.
Sun, 06 Dec 2009 22:57:03 +0100 Cezary Kaliszyk Solved all quotient goals.
Sun, 06 Dec 2009 11:39:34 +0100 Christian Urban updated Isabelle and deleted mono rules
Sun, 06 Dec 2009 11:21:29 +0100 Christian Urban more tuning of the code
Sun, 06 Dec 2009 11:09:51 +0100 Christian Urban puting code in separate sections
Sun, 06 Dec 2009 06:58:24 +0100 Cezary Kaliszyk Handle 'find_qt_asm' exception. Now all inj_repabs_goals should be solved automatically.
Sun, 06 Dec 2009 06:41:52 +0100 Cezary Kaliszyk merge
Sun, 06 Dec 2009 06:39:32 +0100 Cezary Kaliszyk Simpler definition code that works with any type maps.
Sun, 06 Dec 2009 04:03:08 +0100 Christian Urban working on lambda_prs with examples; polished code of clean_tac
Sun, 06 Dec 2009 02:41:35 +0100 Christian Urban renamed lambda_allex_prs
Sun, 06 Dec 2009 01:43:46 +0100 Christian Urban added more to IntEx2
Sun, 06 Dec 2009 00:19:45 +0100 Christian Urban merged
Sun, 06 Dec 2009 00:13:35 +0100 Christian Urban added new example for Ints; regularise does not work in all instances
Sun, 06 Dec 2009 00:00:47 +0100 Cezary Kaliszyk Definitions folded first.
Sat, 05 Dec 2009 23:35:09 +0100 Cezary Kaliszyk Used symmetric definitions. Moved quotient_rsp to QuotMain.
Sat, 05 Dec 2009 23:26:08 +0100 Cezary Kaliszyk Proved foldl_rsp and ho_map_rsp
Sat, 05 Dec 2009 22:38:42 +0100 Christian Urban moved all_prs and ex_prs out from the conversion into the simplifier
Sat, 05 Dec 2009 22:16:17 +0100 Christian Urban further cleaning
Sat, 05 Dec 2009 22:07:46 +0100 Cezary Kaliszyk Merge
Sat, 05 Dec 2009 22:05:09 +0100 Cezary Kaliszyk Added nil_rsp and cons_rsp to quotient_rsp; simplified IntEx.
Sat, 05 Dec 2009 22:02:32 +0100 Christian Urban simplified inj_repabs_trm
Sat, 05 Dec 2009 21:50:31 +0100 Christian Urban merged
Sat, 05 Dec 2009 21:47:48 +0100 Christian Urban merged
Sat, 05 Dec 2009 21:45:56 +0100 Christian Urban merged
Sat, 05 Dec 2009 21:44:01 +0100 Christian Urban simpler version of clean_tac
Sat, 05 Dec 2009 21:45:39 +0100 Cezary Kaliszyk Handling of respects in the fast inj_repabs_tac; includes respects with quotient assumptions.
Sat, 05 Dec 2009 21:28:24 +0100 Cezary Kaliszyk Solutions to IntEx tests.
Sat, 05 Dec 2009 16:26:18 +0100 Christian Urban made some slight simplifications to the examples
Sat, 05 Dec 2009 13:49:35 +0100 Christian Urban added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Sat, 05 Dec 2009 00:06:27 +0100 Christian Urban tuned code
Fri, 04 Dec 2009 21:43:29 +0100 Christian Urban merged
Fri, 04 Dec 2009 21:42:55 +0100 Christian Urban not yet quite functional treatment of constants
Fri, 04 Dec 2009 19:47:58 +0100 Cezary Kaliszyk Changed FOCUS to SUBGOAL in rep_abs_rsp_tac; another 20% speedup :)
Fri, 04 Dec 2009 19:06:53 +0100 Cezary Kaliszyk Changing FOCUS to CSUBGOAL (part 1)
Fri, 04 Dec 2009 18:32:19 +0100 Cezary Kaliszyk abs_rep as ==
Fri, 04 Dec 2009 17:57:03 +0100 Cezary Kaliszyk Cleaning the Quotients file
Fri, 04 Dec 2009 17:50:02 +0100 Cezary Kaliszyk Minor renames and moving
Fri, 04 Dec 2009 17:36:45 +0100 Cezary Kaliszyk Cleaning/review of QuotScript.
Fri, 04 Dec 2009 17:15:55 +0100 Cezary Kaliszyk More cleaning
Fri, 04 Dec 2009 16:53:11 +0100 Cezary Kaliszyk more name cleaning and removing
Fri, 04 Dec 2009 16:40:23 +0100 Cezary Kaliszyk More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
Fri, 04 Dec 2009 16:12:40 +0100 Cezary Kaliszyk Cleaning & Renaming coming from QuotList
Fri, 04 Dec 2009 16:01:23 +0100 Cezary Kaliszyk Cleaning in Quotients
Fri, 04 Dec 2009 15:50:57 +0100 Cezary Kaliszyk Even more name changes and cleaning
Fri, 04 Dec 2009 15:41:09 +0100 Cezary Kaliszyk More code cleaning and name changes
Fri, 04 Dec 2009 15:25:51 +0100 Cezary Kaliszyk merge
Fri, 04 Dec 2009 15:25:26 +0100 Cezary Kaliszyk merged
Fri, 04 Dec 2009 15:23:10 +0100 Christian Urban smaller theory footprint
Fri, 04 Dec 2009 15:20:06 +0100 Christian Urban merged
Fri, 04 Dec 2009 15:19:39 +0100 Christian Urban merge
Fri, 04 Dec 2009 15:18:37 +0100 Christian Urban smaller theory footprint
Fri, 04 Dec 2009 15:18:33 +0100 Cezary Kaliszyk More name changes
Fri, 04 Dec 2009 15:04:05 +0100 Cezary Kaliszyk Naming changes
Fri, 04 Dec 2009 14:35:36 +0100 Cezary Kaliszyk code cleaning and renaming
Fri, 04 Dec 2009 14:11:20 +0100 Cezary Kaliszyk merge
Fri, 04 Dec 2009 14:11:03 +0100 Cezary Kaliszyk Removed previous inj_repabs_tac
Fri, 04 Dec 2009 13:58:23 +0100 Christian Urban some tuning
Fri, 04 Dec 2009 12:21:15 +0100 Cezary Kaliszyk merge
Fri, 04 Dec 2009 12:20:49 +0100 Cezary Kaliszyk rep_abs_rsp_tac to replace the last use of instantiate_tac with matching and unification.
Fri, 04 Dec 2009 11:34:49 +0100 Christian Urban merged
Fri, 04 Dec 2009 11:34:21 +0100 Christian Urban merged
Fri, 04 Dec 2009 11:33:58 +0100 Cezary Kaliszyk Change equiv_trans2 to EQUALS_RSP, since we can prove it for any quotient type, not only for eqv relations.
Fri, 04 Dec 2009 11:06:43 +0100 Cezary Kaliszyk compose_tac instead of rtac to avoid unification; some code cleaning.
Fri, 04 Dec 2009 10:36:35 +0100 Cezary Kaliszyk Got to about 5 seconds for the longest proof. APPLY_RSP_TAC' solves the quotient internally without instantiation resulting in a faster proof.
Fri, 04 Dec 2009 10:12:17 +0100 Cezary Kaliszyk Using APPLY_RSP1; again a little bit faster.
Fri, 04 Dec 2009 09:33:32 +0100 Cezary Kaliszyk Fixes after big merge.
Fri, 04 Dec 2009 09:25:27 +0100 Cezary Kaliszyk The big merge; probably non-functional.
Fri, 04 Dec 2009 09:08:51 +0100 Cezary Kaliszyk Testing the new tactic everywhere
Fri, 04 Dec 2009 09:01:13 +0100 Cezary Kaliszyk First version of the deterministic rep-abs-inj-tac.
Fri, 04 Dec 2009 09:18:46 +0100 Cezary Kaliszyk Changing = to \<equiv> in case if we want to use simp.
Fri, 04 Dec 2009 09:10:31 +0100 Cezary Kaliszyk Even better: Completely got rid of the simps in both quotient_tac and inj_repabs_tac
Fri, 04 Dec 2009 08:19:56 +0100 Cezary Kaliszyk merge
Fri, 04 Dec 2009 08:18:38 +0100 Cezary Kaliszyk Made your version work again with LIST_REL_EQ.
Thu, 03 Dec 2009 19:06:14 +0100 Christian Urban the error occurs in APPLY_RSP_TAC where the wrong quotient (for LIST_REL) is applied
Thu, 03 Dec 2009 15:03:31 +0100 Christian Urban removed quot argument...not all examples work anymore
Thu, 03 Dec 2009 14:02:05 +0100 Christian Urban merged
Thu, 03 Dec 2009 14:00:43 +0100 Christian Urban merged
Thu, 03 Dec 2009 13:59:53 +0100 Christian Urban first version of internalised quotient theorems; added FIXME-TODO
Thu, 03 Dec 2009 11:40:10 +0100 Christian Urban further dead code
Thu, 03 Dec 2009 13:56:59 +0100 Cezary Kaliszyk Reintroduced varifyT; we still need it for permutation definition.
Thu, 03 Dec 2009 13:45:52 +0100 Cezary Kaliszyk Updated the examples
Thu, 03 Dec 2009 12:33:05 +0100 Cezary Kaliszyk Fixed previous mistake
Thu, 03 Dec 2009 12:31:05 +0100 Cezary Kaliszyk defs used automatically by clean_tac
Thu, 03 Dec 2009 12:22:19 +0100 Cezary Kaliszyk Added qoutient_consts dest for getting all the constant definitions in the cleaning step.
Thu, 03 Dec 2009 12:17:23 +0100 Cezary Kaliszyk Added the definition to quotient constant data.
Thu, 03 Dec 2009 11:58:46 +0100 Cezary Kaliszyk removing unused code
Thu, 03 Dec 2009 11:34:34 +0100 Christian Urban merged
Thu, 03 Dec 2009 11:33:24 +0100 Christian Urban deleted some dead code
Thu, 03 Dec 2009 11:28:19 +0100 Cezary Kaliszyk Included all_prs and ex_prs in the lambda_prs conversion.
Thu, 03 Dec 2009 02:53:54 +0100 Christian Urban further simplification
Thu, 03 Dec 2009 02:18:21 +0100 Christian Urban simplified lambda_prs_conv
Wed, 02 Dec 2009 23:31:30 +0100 Christian Urban deleted now obsolete argument rty everywhere
Wed, 02 Dec 2009 23:11:50 +0100 Christian Urban deleted tests at the beginning of QuotMain
Wed, 02 Dec 2009 20:54:59 +0100 Cezary Kaliszyk Experiments with OPTION_map
Wed, 02 Dec 2009 17:16:44 +0100 Cezary Kaliszyk merge
Wed, 02 Dec 2009 17:15:36 +0100 Cezary Kaliszyk More experiments with higher order quotients and theorems with non-lifted constants.
Wed, 02 Dec 2009 16:12:41 +0100 Christian Urban merged
Wed, 02 Dec 2009 14:37:21 +0100 Cezary Kaliszyk Lifting to 2 different types :)
Wed, 02 Dec 2009 14:11:46 +0100 Cezary Kaliszyk New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Wed, 02 Dec 2009 12:07:54 +0100 Cezary Kaliszyk Fixed unlam for non-abstractions and updated list_induct_part proof.
Wed, 02 Dec 2009 11:30:40 +0100 Cezary Kaliszyk Removed the use of 'rty' from APPLY_RSP, finally LF proofs go automatically.
Wed, 02 Dec 2009 10:56:35 +0100 Cezary Kaliszyk The conversion approach works.
Wed, 02 Dec 2009 10:30:20 +0100 Cezary Kaliszyk Trying a conversion based approach.
Wed, 02 Dec 2009 09:23:48 +0100 Cezary Kaliszyk A bit of progress; but the object-logic vs meta-logic distinction is troublesome.
Wed, 02 Dec 2009 07:21:10 +0100 Cezary Kaliszyk Added tactic for dealing with QUOT_TRUE and introducing QUOT_TRUE.
Tue, 01 Dec 2009 19:18:43 +0100 Cezary Kaliszyk back in working state
Tue, 01 Dec 2009 19:01:51 +0100 Cezary Kaliszyk clean
Tue, 01 Dec 2009 18:51:14 +0100 Christian Urban fixed previous commit
Tue, 01 Dec 2009 18:48:52 +0100 Christian Urban fixed problems with FOCUS
Tue, 01 Dec 2009 18:41:01 +0100 Christian Urban added a make_inst test
Tue, 01 Dec 2009 18:22:48 +0100 Cezary Kaliszyk Transformation of QUOT_TRUE assumption by any given function
Tue, 01 Dec 2009 16:27:42 +0100 Christian Urban QUOT_TRUE joke
Tue, 01 Dec 2009 14:08:56 +0100 Cezary Kaliszyk Removed last HOL_ss
Tue, 01 Dec 2009 14:04:00 +0100 Cezary Kaliszyk more cleaning
Tue, 01 Dec 2009 12:16:45 +0100 Cezary Kaliszyk Cleaning 'aps'.
Mon, 30 Nov 2009 15:40:22 +0100 Cezary Kaliszyk merge
Mon, 30 Nov 2009 15:36:49 +0100 Christian Urban cleaned inj_regabs_trm
Mon, 30 Nov 2009 15:33:06 +0100 Cezary Kaliszyk merge
Mon, 30 Nov 2009 15:32:14 +0100 Cezary Kaliszyk clean_tac rewrites the definitions the other way
Mon, 30 Nov 2009 13:58:05 +0100 Christian Urban merged
Mon, 30 Nov 2009 12:26:08 +0100 Christian Urban added facilities to get all stored quotient data (equiv thms etc)
Mon, 30 Nov 2009 12:14:20 +0100 Cezary Kaliszyk More code cleaning
Mon, 30 Nov 2009 11:53:20 +0100 Cezary Kaliszyk Code cleaning.
Mon, 30 Nov 2009 10:16:10 +0100 Cezary Kaliszyk Commented clean-tac
Mon, 30 Nov 2009 02:05:22 +0100 Cezary Kaliszyk Added another induction to LFex
Sun, 29 Nov 2009 23:59:37 +0100 Christian Urban tried to improve the inj_repabs_trm function but left the new part commented out
Sun, 29 Nov 2009 19:48:55 +0100 Christian Urban added a new version of QuotMain to experiment with qids
Sun, 29 Nov 2009 17:47:37 +0100 Christian Urban started functions for qid-insertion and fixed a bug in regularise
Sun, 29 Nov 2009 09:38:07 +0100 Cezary Kaliszyk Removed unnecessary HOL_ss which proved one of the subgoals.
Sun, 29 Nov 2009 08:48:06 +0100 Cezary Kaliszyk Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Sun, 29 Nov 2009 03:59:18 +0100 Christian Urban introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Sun, 29 Nov 2009 02:51:42 +0100 Christian Urban tuned
Sat, 28 Nov 2009 19:14:12 +0100 Christian Urban improved pattern matching inside the inj_repabs_tacs
Sat, 28 Nov 2009 18:49:39 +0100 Christian Urban selective debugging of the inj_repabs_tac (at the moment for step 3 and 4 debugging information is printed)
Sat, 28 Nov 2009 14:45:22 +0100 Christian Urban removed old inj_repabs_tac; kept only the one with (selective) debugging information
Sat, 28 Nov 2009 14:33:04 +0100 Christian Urban renamed r_mk_comb_tac to inj_repabs_tac
Sat, 28 Nov 2009 14:15:05 +0100 Christian Urban tuning
Sat, 28 Nov 2009 14:03:01 +0100 Christian Urban tuned comments
Sat, 28 Nov 2009 13:54:48 +0100 Christian Urban renamed LAMBDA_RES_TAC and WEAK_LAMBDA_RES_TAC to lower case names
Sat, 28 Nov 2009 08:46:24 +0100 Cezary Kaliszyk Manually finished LF induction.
Sat, 28 Nov 2009 08:04:23 +0100 Cezary Kaliszyk Moved fast instantiation to QuotMain
Sat, 28 Nov 2009 07:44:17 +0100 Cezary Kaliszyk LFex proof a bit further.
Sat, 28 Nov 2009 06:15:06 +0100 Cezary Kaliszyk merge
Sat, 28 Nov 2009 06:14:50 +0100 Cezary Kaliszyk Looking at repabs proof in LF.
Sat, 28 Nov 2009 05:53:31 +0100 Christian Urban further proper merge
Sat, 28 Nov 2009 05:49:16 +0100 Christian Urban merged
Sat, 28 Nov 2009 05:47:13 +0100 Christian Urban more simplification
Sat, 28 Nov 2009 05:43:18 +0100 Cezary Kaliszyk Merged and tested that all works.
Sat, 28 Nov 2009 05:29:30 +0100 Cezary Kaliszyk Finished and tested the new regularize
Sat, 28 Nov 2009 05:09:22 +0100 Christian Urban more tuning of the repabs-tactics
Sat, 28 Nov 2009 04:46:03 +0100 Christian Urban fixed examples in IntEx and FSet
Sat, 28 Nov 2009 04:37:30 +0100 Christian Urban merged
Sat, 28 Nov 2009 04:37:04 +0100 Christian Urban fixed previous commit
Sat, 28 Nov 2009 04:02:54 +0100 Cezary Kaliszyk Cleaned all lemmas about regularisation of Ball and Bex and moved in one place. Second Ball simprox.
Sat, 28 Nov 2009 03:17:22 +0100 Cezary Kaliszyk Merged comment
Sat, 28 Nov 2009 03:07:38 +0100 Cezary Kaliszyk Integrated Stefan's tactic and changed substs to simps with empty context.
Sat, 28 Nov 2009 03:06:22 +0100 Christian Urban some slight tuning of the apply-tactic
Sat, 28 Nov 2009 02:54:24 +0100 Christian Urban annotated a proof with all steps and simplified LAMBDA_RES_TAC
Fri, 27 Nov 2009 18:38:44 +0100 Cezary Kaliszyk Merge
Fri, 27 Nov 2009 18:38:09 +0100 Cezary Kaliszyk The magical code from Stefan, will need to be integrated in the Simproc.
Fri, 27 Nov 2009 13:59:52 +0100 Christian Urban replaced FIRST' (map rtac list) with resolve_tac list
Fri, 27 Nov 2009 10:04:49 +0100 Cezary Kaliszyk Simplifying arguments; got rid of trans2_thm.
Fri, 27 Nov 2009 09:16:32 +0100 Cezary Kaliszyk Cleaning of LFex. Lambda_prs fails to unify in 2 places.
Fri, 27 Nov 2009 08:22:46 +0100 Cezary Kaliszyk Recommit
Fri, 27 Nov 2009 08:15:23 +0100 Cezary Kaliszyk Removing arguments of tactics: absrep, rel_refl, reps_same are computed.
Fri, 27 Nov 2009 07:16:16 +0100 Cezary Kaliszyk More cleaning in QuotMain, identity handling.
Fri, 27 Nov 2009 07:00:14 +0100 Cezary Kaliszyk Minor cleaning
Fri, 27 Nov 2009 04:02:20 +0100 Christian Urban tuned
Fri, 27 Nov 2009 03:56:18 +0100 Christian Urban some tuning
Fri, 27 Nov 2009 03:33:30 +0100 Christian Urban simplified gen_frees_tac and properly named abstracted variables
Fri, 27 Nov 2009 02:58:28 +0100 Christian Urban removed CHANGED'
Fri, 27 Nov 2009 02:55:56 +0100 Christian Urban introduced a separate lemma for id_simps
Fri, 27 Nov 2009 02:45:54 +0100 Christian Urban renamed inj_REPABS to inj_repabs_trm
Fri, 27 Nov 2009 02:44:11 +0100 Christian Urban tuned comments and moved slightly some code
Fri, 27 Nov 2009 02:35:50 +0100 Christian Urban deleted obsolete qenv code
Fri, 27 Nov 2009 02:23:49 +0100 Christian Urban renamed REGULARIZE to be regularize
Thu, 26 Nov 2009 21:16:59 +0100 Christian Urban more tuning
Thu, 26 Nov 2009 21:04:17 +0100 Christian Urban deleted get_fun_old and stuff
Thu, 26 Nov 2009 21:01:53 +0100 Christian Urban recommited changes of comments
Thu, 26 Nov 2009 20:32:56 +0100 Cezary Kaliszyk Merge Again
Thu, 26 Nov 2009 20:32:33 +0100 Cezary Kaliszyk Merged
Thu, 26 Nov 2009 20:18:36 +0100 Christian Urban tuned comments
Thu, 26 Nov 2009 19:51:31 +0100 Christian Urban some diagnostic code for r_mk_comb
Thu, 26 Nov 2009 16:23:24 +0100 Christian Urban introduced a new property for Ball and ===> on the left
Thu, 26 Nov 2009 13:52:46 +0100 Christian Urban fixed QuotList
Thu, 26 Nov 2009 13:46:00 +0100 Christian Urban changed left-res
Thu, 26 Nov 2009 12:21:47 +0100 Cezary Kaliszyk Manually regularized akind_aty_atrm.induct
Thu, 26 Nov 2009 10:52:24 +0100 Cezary Kaliszyk Playing with Monos in LFex.
Thu, 26 Nov 2009 10:32:31 +0100 Cezary Kaliszyk Fixed FSet after merge.
Thu, 26 Nov 2009 03:18:38 +0100 Christian Urban merged
Thu, 26 Nov 2009 02:31:00 +0100 Christian Urban test with monos
Wed, 25 Nov 2009 21:48:32 +0100 Cezary Kaliszyk applic_prs
Wed, 25 Nov 2009 15:43:12 +0100 Christian Urban polishing
Wed, 25 Nov 2009 15:20:10 +0100 Christian Urban reordered the code
Wed, 25 Nov 2009 14:25:29 +0100 Cezary Kaliszyk Moved exception handling to QuotMain and cleaned FSet.
Wed, 25 Nov 2009 14:16:33 +0100 Cezary Kaliszyk Merge
Wed, 25 Nov 2009 14:15:34 +0100 Cezary Kaliszyk Finished manual lifting of list_induct_part :)
Wed, 25 Nov 2009 12:27:28 +0100 Christian Urban comments tuning and slight reordering
Wed, 25 Nov 2009 11:59:49 +0100 Cezary Kaliszyk Merge
Wed, 25 Nov 2009 11:58:56 +0100 Cezary Kaliszyk More moving from QuotMain to UnusedQuotMain
Wed, 25 Nov 2009 11:46:59 +0100 Christian Urban deleted some obsolete diagnostic code
Wed, 25 Nov 2009 11:41:42 +0100 Cezary Kaliszyk Removed unused things from QuotMain.
Wed, 25 Nov 2009 10:52:21 +0100 Cezary Kaliszyk All examples work again.
Wed, 25 Nov 2009 10:39:53 +0100 Cezary Kaliszyk cleaning in MyInt
Wed, 25 Nov 2009 10:34:03 +0100 Cezary Kaliszyk lambda_prs and cleaning the existing examples.
Wed, 25 Nov 2009 03:47:07 +0100 Christian Urban merged
Wed, 25 Nov 2009 03:45:44 +0100 Christian Urban fixed the problem with generalising variables; at the moment it is quite a hack
Tue, 24 Nov 2009 20:13:16 +0100 Cezary Kaliszyk Ho-matching failures...
Tue, 24 Nov 2009 19:09:29 +0100 Christian Urban changed unification to matching
Tue, 24 Nov 2009 18:13:57 +0100 Christian Urban unification
Tue, 24 Nov 2009 18:13:18 +0100 Cezary Kaliszyk Lambda & SOLVED' for new quotient_tac
Tue, 24 Nov 2009 17:35:31 +0100 Christian Urban merged
Tue, 24 Nov 2009 17:00:53 +0100 Cezary Kaliszyk Conversion
Tue, 24 Nov 2009 16:20:34 +0100 Cezary Kaliszyk The non-working procedure_tac.
Tue, 24 Nov 2009 16:10:31 +0100 Christian Urban merged
Tue, 24 Nov 2009 15:31:29 +0100 Christian Urban use error instead of raising our own exception
Tue, 24 Nov 2009 15:18:00 +0100 Cezary Kaliszyk Fixes to the tactic after quotient_tac changed.
Tue, 24 Nov 2009 15:15:33 +0100 Christian Urban merged
Tue, 24 Nov 2009 15:15:10 +0100 Christian Urban added a prepare_tac
Tue, 24 Nov 2009 14:41:47 +0100 Cezary Kaliszyk TRY' for clean_tac
Tue, 24 Nov 2009 14:19:54 +0100 Cezary Kaliszyk Moved cleaning to QuotMain
Tue, 24 Nov 2009 14:16:57 +0100 Cezary Kaliszyk New cleaning tactic
Tue, 24 Nov 2009 13:46:36 +0100 Christian Urban explicit phases for the cleaning
Tue, 24 Nov 2009 12:25:04 +0100 Cezary Kaliszyk Separate regularize_tac
Tue, 24 Nov 2009 08:36:28 +0100 Cezary Kaliszyk Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Tue, 24 Nov 2009 08:35:04 +0100 Cezary Kaliszyk More fixes for inj_REPABS
Tue, 24 Nov 2009 01:36:50 +0100 Christian Urban addded a tactic, which sets up the three goals of the `algorithm'
Mon, 23 Nov 2009 23:09:03 +0100 Christian Urban fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Mon, 23 Nov 2009 22:00:54 +0100 Christian Urban merged
Mon, 23 Nov 2009 21:59:57 +0100 Christian Urban tuned some comments
Mon, 23 Nov 2009 21:56:30 +0100 Cezary Kaliszyk Another not-typechecking regularized term.
Mon, 23 Nov 2009 21:48:44 +0100 Cezary Kaliszyk domain_type in regularizing equality
Mon, 23 Nov 2009 21:12:16 +0100 Cezary Kaliszyk More theorems lifted in the goal-directed way.
Mon, 23 Nov 2009 20:10:39 +0100 Cezary Kaliszyk Finished temporary goal-directed lift_theorem wrapper.
Mon, 23 Nov 2009 16:13:19 +0100 Christian Urban merged
Mon, 23 Nov 2009 16:12:50 +0100 Christian Urban a version of inj_REPABS (needs to be looked at again later)
Mon, 23 Nov 2009 15:47:14 +0100 Cezary Kaliszyk Fixes for atomize
Mon, 23 Nov 2009 15:08:25 +0100 Cezary Kaliszyk merge
Mon, 23 Nov 2009 15:08:09 +0100 Cezary Kaliszyk lift_thm with a goal.
Mon, 23 Nov 2009 15:02:48 +0100 Christian Urban slight change in code layout
Mon, 23 Nov 2009 14:40:53 +0100 Cezary Kaliszyk Fixes for new code
Mon, 23 Nov 2009 14:32:11 +0100 Cezary Kaliszyk Removing dead code
Mon, 23 Nov 2009 14:16:41 +0100 Cezary Kaliszyk Move atomize_goal to QuotMain
Mon, 23 Nov 2009 14:05:02 +0100 Cezary Kaliszyk Removed second implementation of Regularize/Inject from FSet.
Mon, 23 Nov 2009 13:55:31 +0100 Cezary Kaliszyk Moved new repabs_inj code to QuotMain
Mon, 23 Nov 2009 13:46:14 +0100 Cezary Kaliszyk New repabs behaves the same way as old one.
Mon, 23 Nov 2009 13:24:12 +0100 Christian Urban code review with Cezary
Mon, 23 Nov 2009 10:26:59 +0100 Cezary Kaliszyk The other branch does not seem to work...
Mon, 23 Nov 2009 10:04:35 +0100 Cezary Kaliszyk Fixes for recent changes.
Sun, 22 Nov 2009 15:30:23 +0100 Christian Urban updated to Isabelle 22nd November
Sun, 22 Nov 2009 00:01:06 +0100 Christian Urban a little tuning of comments
Sat, 21 Nov 2009 23:23:01 +0100 Christian Urban slight tuning
Sat, 21 Nov 2009 14:45:25 +0100 Christian Urban some debugging code, but cannot find the place where the cprems_of exception is raised
Sat, 21 Nov 2009 14:18:31 +0100 Christian Urban tried to prove the repabs_inj lemma, but failed for the moment
Sat, 21 Nov 2009 13:14:35 +0100 Christian Urban my first version of repabs injection
Sat, 21 Nov 2009 11:16:48 +0100 Christian Urban tuned
Sat, 21 Nov 2009 10:58:08 +0100 Christian Urban tunded
Sat, 21 Nov 2009 03:12:50 +0100 Christian Urban tuned
Sat, 21 Nov 2009 02:53:23 +0100 Christian Urban flagged qenv-stuff as obsolete
Sat, 21 Nov 2009 02:49:39 +0100 Christian Urban simplified get_fun so that it uses directly rty and qty, instead of qenv
Fri, 20 Nov 2009 13:03:01 +0100 Christian Urban started regularize of rtrm/qtrm version; looks quite promising
Thu, 19 Nov 2009 14:17:10 +0100 Christian Urban updated to new Isabelle
Wed, 18 Nov 2009 23:52:48 +0100 Christian Urban fixed the storage of qconst definitions
Fri, 13 Nov 2009 19:32:12 +0100 Cezary Kaliszyk Still don't know how to do the proof automatically.
Fri, 13 Nov 2009 16:44:36 +0100 Christian Urban added some tracing information to all phases of lifting to the function lift_thm
Thu, 12 Nov 2009 13:57:20 +0100 Cezary Kaliszyk merge of the merge?
Thu, 12 Nov 2009 13:56:07 +0100 Cezary Kaliszyk merged
Thu, 12 Nov 2009 12:15:41 +0100 Christian Urban added a FIXME commment
Thu, 12 Nov 2009 12:07:33 +0100 Christian Urban looking up data in quot_info works now (needs qualified string)
Thu, 12 Nov 2009 02:54:40 +0100 Christian Urban changed the quotdata to be a symtab table (needs fixing)
Thu, 12 Nov 2009 02:18:36 +0100 Christian Urban added a container for quotient constants (does not work yet though)
Wed, 11 Nov 2009 22:30:43 +0100 Cezary Kaliszyk Lifting towards goal and manually finished the proof.
Wed, 11 Nov 2009 18:51:59 +0100 Cezary Kaliszyk merge
Wed, 11 Nov 2009 18:49:46 +0100 Cezary Kaliszyk Modifications while preparing the goal-directed version.
Wed, 11 Nov 2009 11:59:22 +0100 Christian Urban updated to new Theory_Data and to new Isabelle
Wed, 11 Nov 2009 10:22:47 +0100 Cezary Kaliszyk Removed 'Toplevel.program' for polyml 5.3
Tue, 10 Nov 2009 17:43:05 +0100 Cezary Kaliszyk Atomizing a "goal" theorems.
Tue, 10 Nov 2009 09:32:16 +0100 Cezary Kaliszyk More code cleaning and commenting
Mon, 09 Nov 2009 15:40:43 +0100 Cezary Kaliszyk Minor cleaning and removing of some 'handle _'.
Mon, 09 Nov 2009 15:23:33 +0100 Cezary Kaliszyk Cleaning and commenting
Mon, 09 Nov 2009 13:47:46 +0100 Cezary Kaliszyk Fixes for the other get_fun implementation.
Fri, 06 Nov 2009 19:43:09 +0100 Christian Urban permutation lifting works now also
Fri, 06 Nov 2009 19:26:32 +0100 Christian Urban merged
Fri, 06 Nov 2009 19:26:08 +0100 Christian Urban updated to new Isabelle version and added a new example file
Fri, 06 Nov 2009 17:42:20 +0100 Cezary Kaliszyk Minor changes
Fri, 06 Nov 2009 11:02:11 +0100 Cezary Kaliszyk merge
Fri, 06 Nov 2009 11:01:22 +0100 Cezary Kaliszyk fold_rsp
Fri, 06 Nov 2009 09:48:37 +0100 Christian Urban tuned the code in quotient and quotient_def
Thu, 05 Nov 2009 16:43:57 +0100 Cezary Kaliszyk More functionality for lifting list.cases and list.recs.
Thu, 05 Nov 2009 13:47:41 +0100 Christian Urban merged
Thu, 05 Nov 2009 13:47:04 +0100 Christian Urban removed typing information from get_fun in quotient_def; *potentially* dangerous
Thu, 05 Nov 2009 13:36:46 +0100 Cezary Kaliszyk Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Thu, 05 Nov 2009 10:46:54 +0100 Christian Urban removed Simplifier.context
Thu, 05 Nov 2009 10:23:27 +0100 Christian Urban replaced check_term o parse_term by read_term
Thu, 05 Nov 2009 09:55:21 +0100 Christian Urban merged
Thu, 05 Nov 2009 09:38:34 +0100 Cezary Kaliszyk Infrastructure for polymorphic types
Wed, 04 Nov 2009 16:10:39 +0100 Cezary Kaliszyk Two new tests for get_fun. Second one fails.
Wed, 04 Nov 2009 15:27:32 +0100 Cezary Kaliszyk Type instantiation in regularize
Wed, 04 Nov 2009 14:03:46 +0100 Cezary Kaliszyk Description of regularize
Wed, 04 Nov 2009 13:33:13 +0100 Cezary Kaliszyk Experiments with lifting partially applied constants.
Wed, 04 Nov 2009 12:19:04 +0100 Christian Urban more tuning
Wed, 04 Nov 2009 12:07:22 +0100 Christian Urban slightly tuned
Wed, 04 Nov 2009 11:59:48 +0100 Christian Urban merged
Wed, 04 Nov 2009 11:59:15 +0100 Christian Urban separated the quotient_def into a separate file
Wed, 04 Nov 2009 11:08:05 +0100 Cezary Kaliszyk Experiments in Int
Wed, 04 Nov 2009 10:43:33 +0100 Christian Urban fixed definition of PLUS
Wed, 04 Nov 2009 10:31:20 +0100 Christian Urban simplified the quotient_def code
Wed, 04 Nov 2009 09:52:31 +0100 Cezary Kaliszyk Lifting 'fold1.simps(2)' and some cleaning.
Tue, 03 Nov 2009 18:09:59 +0100 Cezary Kaliszyk Playing with alpha_refl.
Tue, 03 Nov 2009 17:51:10 +0100 Cezary Kaliszyk Alpha.induct now lifts automatically.
Tue, 03 Nov 2009 17:30:43 +0100 Cezary Kaliszyk merge
Tue, 03 Nov 2009 17:30:27 +0100 Cezary Kaliszyk applic_prs
Tue, 03 Nov 2009 16:51:33 +0100 Christian Urban simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Tue, 03 Nov 2009 16:17:19 +0100 Cezary Kaliszyk Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Tue, 03 Nov 2009 14:04:45 +0100 Cezary Kaliszyk merge
Tue, 03 Nov 2009 14:04:21 +0100 Cezary Kaliszyk Preparing infrastructure for general FORALL_PRS
Mon, 02 Nov 2009 18:26:55 +0100 Christian Urban split quotient.ML into two files
Mon, 02 Nov 2009 18:16:19 +0100 Christian Urban slightly saner way of parsing the quotient_def
Mon, 02 Nov 2009 15:39:25 +0100 Christian Urban merged
Mon, 02 Nov 2009 15:38:49 +0100 Christian Urban changed Type.typ_match to Sign.typ_match
Mon, 02 Nov 2009 15:38:03 +0100 Cezary Kaliszyk Fixes after optimization and preparing for a general FORALL_PRS
Mon, 02 Nov 2009 14:57:56 +0100 Cezary Kaliszyk Optimization
Mon, 02 Nov 2009 11:51:50 +0100 Cezary Kaliszyk Map does not fully work yet.
Mon, 02 Nov 2009 11:15:26 +0100 Cezary Kaliszyk Fixed quotdata_lookup.
Mon, 02 Nov 2009 09:47:49 +0100 Christian Urban fixed problem with maps_update
Mon, 02 Nov 2009 09:39:29 +0100 Christian Urban merged
Mon, 02 Nov 2009 09:33:48 +0100 Christian Urban fixed the problem with types in map
Sat, 31 Oct 2009 11:20:55 +0100 Cezary Kaliszyk Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Fri, 30 Oct 2009 19:03:53 +0100 Cezary Kaliszyk Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Fri, 30 Oct 2009 18:31:06 +0100 Cezary Kaliszyk Regularization
Fri, 30 Oct 2009 16:37:09 +0100 Christian Urban merged
Fri, 30 Oct 2009 16:35:43 +0100 Christian Urban added some facts about fresh and support of lam
Fri, 30 Oct 2009 16:24:07 +0100 Cezary Kaliszyk Finally merged the code of the versions of regularize and tested examples.
Fri, 30 Oct 2009 15:52:47 +0100 Cezary Kaliszyk Lemmas about fv.
Fri, 30 Oct 2009 15:35:42 +0100 Christian Urban changed the order of rfv and reformulated a3 with rfv
Fri, 30 Oct 2009 15:32:04 +0100 Christian Urban merged
Fri, 30 Oct 2009 15:30:33 +0100 Christian Urban not sure what changed here
Fri, 30 Oct 2009 15:28:44 +0100 Christian Urban added fv-function
Fri, 30 Oct 2009 15:22:59 +0100 Cezary Kaliszyk The proper real_alpha
Fri, 30 Oct 2009 14:25:37 +0100 Cezary Kaliszyk Finding applications and duplicates filtered out in abstractions
Fri, 30 Oct 2009 12:22:03 +0100 Cezary Kaliszyk Cleaning also in Lam
Fri, 30 Oct 2009 11:25:29 +0100 Cezary Kaliszyk Cleaning of the interface to lift.
Thu, 29 Oct 2009 17:35:03 +0100 Cezary Kaliszyk Tried manually lifting real_alpha
Thu, 29 Oct 2009 13:30:11 +0100 Cezary Kaliszyk More tests in Lam
Thu, 29 Oct 2009 13:29:03 +0100 Cezary Kaliszyk Cleaning of 'map id' and 'prod_fun id id' in lower_defs.
Thu, 29 Oct 2009 12:09:31 +0100 Cezary Kaliszyk Using subst for identity definition.
Thu, 29 Oct 2009 08:46:34 +0100 Cezary Kaliszyk Lifting of the 3 lemmas in LamEx
Thu, 29 Oct 2009 08:06:49 +0100 Cezary Kaliszyk Fixed wrong CARD definition and removed the "Does not work anymore" comment.
Thu, 29 Oct 2009 07:29:12 +0100 Christian Urban merged
Wed, 28 Oct 2009 20:01:20 +0100 Christian Urban updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Wed, 28 Oct 2009 19:46:15 +0100 Christian Urban ported all constant definitions to new scheme
Wed, 28 Oct 2009 19:36:52 +0100 Christian Urban fixed the definition of alpha; this *breaks* some of the experiments
Wed, 28 Oct 2009 18:08:38 +0100 Cezary Kaliszyk disambiguate ===> syntax
Wed, 28 Oct 2009 17:38:37 +0100 Cezary Kaliszyk More cleaning in Lam code
Wed, 28 Oct 2009 17:17:21 +0100 Cezary Kaliszyk cleaned FSet
Wed, 28 Oct 2009 16:48:57 +0100 Cezary Kaliszyk Some cleaning
Wed, 28 Oct 2009 16:16:38 +0100 Cezary Kaliszyk Cleaning the unnecessary theorems in 'IntEx'.
Wed, 28 Oct 2009 16:11:28 +0100 Cezary Kaliszyk Fix also in the general procedure.
Wed, 28 Oct 2009 16:06:19 +0100 Cezary Kaliszyk merge
Wed, 28 Oct 2009 16:05:59 +0100 Cezary Kaliszyk Fixes
Wed, 28 Oct 2009 15:48:38 +0100 Christian Urban updated all definitions
Wed, 28 Oct 2009 15:25:36 +0100 Christian Urban merged
Wed, 28 Oct 2009 15:25:11 +0100 Christian Urban added infrastructure for defining lifted constants
Wed, 28 Oct 2009 14:59:24 +0100 Cezary Kaliszyk First experiments with Lambda
Wed, 28 Oct 2009 12:22:06 +0100 Cezary Kaliszyk Fixed mistake in const generation, will postpone this.
Wed, 28 Oct 2009 10:29:00 +0100 Cezary Kaliszyk More finshed proofs and cleaning
Wed, 28 Oct 2009 10:17:07 +0100 Cezary Kaliszyk Proof of append_rsp
Wed, 28 Oct 2009 01:49:31 +0100 Christian Urban merged
Wed, 28 Oct 2009 01:48:45 +0100 Christian Urban added a function for matching types
Tue, 27 Oct 2009 18:05:45 +0100 Cezary Kaliszyk Manual conversion of equality to equivalence allows lifting append_assoc.
Tue, 27 Oct 2009 18:02:35 +0100 Cezary Kaliszyk Simplfied interface to repabs_injection.
Tue, 27 Oct 2009 17:08:47 +0100 Cezary Kaliszyk map_append lifted automatically.
Tue, 27 Oct 2009 16:15:56 +0100 Cezary Kaliszyk Manually lifted Map_Append.
Tue, 27 Oct 2009 15:00:15 +0100 Cezary Kaliszyk Merged
Tue, 27 Oct 2009 14:59:00 +0100 Cezary Kaliszyk Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
Tue, 27 Oct 2009 14:46:38 +0100 Christian Urban tuned
Tue, 27 Oct 2009 14:15:40 +0100 Christian Urban merged
Tue, 27 Oct 2009 14:14:30 +0100 Christian Urban added equiv-thm to the quot_info
Tue, 27 Oct 2009 12:20:57 +0100 Cezary Kaliszyk Simplifying FSet with new functions.
Tue, 27 Oct 2009 11:43:02 +0100 Christian Urban added an example about lambda-terms
Tue, 27 Oct 2009 11:27:53 +0100 Christian Urban made quotients compatiple with Nominal; updated keyword file
Tue, 27 Oct 2009 11:03:38 +0100 Christian Urban merged
Tue, 27 Oct 2009 09:01:12 +0100 Cezary Kaliszyk Completely cleaned Int.
Tue, 27 Oct 2009 07:46:52 +0100 Cezary Kaliszyk Further reordering in Int code.
Mon, 26 Oct 2009 19:35:30 +0100 Cezary Kaliszyk Simplifying Int.
Mon, 26 Oct 2009 15:32:17 +0100 Cezary Kaliszyk Merge
Mon, 26 Oct 2009 15:31:53 +0100 Cezary Kaliszyk Simplifying Int and Working on map
Mon, 26 Oct 2009 14:18:26 +0100 Christian Urban merged
Mon, 26 Oct 2009 14:16:32 +0100 Cezary Kaliszyk Simplifying code in int
Mon, 26 Oct 2009 13:33:28 +0100 Cezary Kaliszyk Symmetry of integer addition
Mon, 26 Oct 2009 11:55:36 +0100 Cezary Kaliszyk Finished the code for adding lower defs, and more things moved to QuotMain
Mon, 26 Oct 2009 11:34:02 +0100 Cezary Kaliszyk Making all the definitions from the original ones
Mon, 26 Oct 2009 10:20:20 +0100 Cezary Kaliszyk Finished COND_PRS proof.
Mon, 26 Oct 2009 10:02:50 +0100 Cezary Kaliszyk Cleaning and fixing.
Mon, 26 Oct 2009 02:06:01 +0100 Christian Urban updated with quotient_def
Sun, 25 Oct 2009 23:44:41 +0100 Christian Urban added code for declaring map-functions
Sun, 25 Oct 2009 01:31:04 +0200 Christian Urban added "print_quotients" command to th ekeyword file
Sun, 25 Oct 2009 01:15:03 +0200 Christian Urban proved the two lemmas in QuotScript (reformulated them without leading forall)
Sun, 25 Oct 2009 00:14:40 +0200 Christian Urban added data-storage about the quotients
Sat, 24 Oct 2009 22:52:23 +0200 Christian Urban added another example file about integers (see HOL/Int.thy)
Sat, 24 Oct 2009 18:17:38 +0200 Christian Urban changed the definitions of liftet constants to use fun_maps
Sat, 24 Oct 2009 17:29:20 +0200 cek Merge
Sat, 24 Oct 2009 17:29:27 +0200 Cezary Kaliszyk Finally lifted induction, with some manually added simplification lemmas.
Sat, 24 Oct 2009 16:31:07 +0200 Christian Urban changed encoding from utf8 to ISO8 (needed to work with xemacs)
Sat, 24 Oct 2009 16:15:33 +0200 cek Merge
Sat, 24 Oct 2009 16:15:58 +0200 Cezary Kaliszyk Preparing infrastructire for LAMBDA_PRS
Sat, 24 Oct 2009 16:09:05 +0200 Christian Urban moved the map_funs setup into QuotMain
Sat, 24 Oct 2009 14:00:18 +0200 Cezary Kaliszyk Finally completely lift the previously lifted theorems + clean some old stuff
Sat, 24 Oct 2009 13:00:54 +0200 Cezary Kaliszyk More infrastructure for automatic lifting of theorems lifted before
Sat, 24 Oct 2009 10:16:53 +0200 Cezary Kaliszyk More infrastructure for automatic lifting of theorems lifted before
Sat, 24 Oct 2009 08:34:14 +0200 cek Undid wrong merge
Sat, 24 Oct 2009 08:29:11 +0200 cek Tried rolling back
Sat, 24 Oct 2009 08:24:26 +0200 cek Cleaning the mess
Sat, 24 Oct 2009 08:09:40 +0200 cek Merge
Sat, 24 Oct 2009 08:09:09 +0200 cek Better tactic and simplified the proof further
Sat, 24 Oct 2009 01:33:29 +0200 Christian Urban fixed problem with incorrect ABS/REP name
Fri, 23 Oct 2009 18:20:06 +0200 Cezary Kaliszyk Stronger tactic, simpler proof.
Fri, 23 Oct 2009 16:34:20 +0200 Cezary Kaliszyk Split Finite Set example into separate file
Fri, 23 Oct 2009 16:01:13 +0200 Cezary Kaliszyk eqsubst_tac
Fri, 23 Oct 2009 11:24:43 +0200 Cezary Kaliszyk Trying to get a simpler lemma with the whole infrastructure
Fri, 23 Oct 2009 09:21:45 +0200 Cezary Kaliszyk Using RANGE tactical allows getting rid of the quotients immediately.
Thu, 22 Oct 2009 17:35:40 +0200 Cezary Kaliszyk Further developing the tactic and simplifying the proof
Thu, 22 Oct 2009 16:24:02 +0200 Cezary Kaliszyk res_forall_rsp_tac further simplifies the proof
Thu, 22 Oct 2009 16:10:06 +0200 Cezary Kaliszyk Working on the proof and the tactic.
Thu, 22 Oct 2009 15:45:05 +0200 Cezary Kaliszyk The proof gets simplified
Thu, 22 Oct 2009 15:44:16 +0200 Cezary Kaliszyk Removed an assumption
Thu, 22 Oct 2009 15:02:01 +0200 Cezary Kaliszyk The proof now including manually unfolded higher-order RES_FORALL_RSP.
Thu, 22 Oct 2009 13:45:48 +0200 Cezary Kaliszyk The problems with 'abs' term.
Thu, 22 Oct 2009 11:43:12 +0200 Cezary Kaliszyk Simplified the proof with some tactic... Still hangs sometimes.
Thu, 22 Oct 2009 10:45:33 +0200 Cezary Kaliszyk More proof
Thu, 22 Oct 2009 10:38:00 +0200 Cezary Kaliszyk Got rid of instantiations in the proof
Thu, 22 Oct 2009 06:51:27 +0200 Cezary Kaliszyk Removed some debugging messages
Thu, 22 Oct 2009 01:59:17 +0200 Christian Urban tuned and attempted to store data about the quotients (does not work yet)
Thu, 22 Oct 2009 01:16:42 +0200 Christian Urban tuned
Thu, 22 Oct 2009 01:15:01 +0200 Christian Urban slight tuning
Wed, 21 Oct 2009 18:30:42 +0200 Christian Urban fixed my_reg
Wed, 21 Oct 2009 16:13:39 +0200 Cezary Kaliszyk Reorganization of the construction part
Wed, 21 Oct 2009 15:01:50 +0200 Cezary Kaliszyk Simplified proof more
Wed, 21 Oct 2009 14:30:29 +0200 Cezary Kaliszyk Cleaning the code
Wed, 21 Oct 2009 14:15:22 +0200 Cezary Kaliszyk Further reorganization
Wed, 21 Oct 2009 14:09:06 +0200 Cezary Kaliszyk Further reorganizing the file
Wed, 21 Oct 2009 13:47:39 +0200 Cezary Kaliszyk Reordering
Wed, 21 Oct 2009 11:50:53 +0200 Cezary Kaliszyk cterm_instantiate also fails for some strange reason...
Wed, 21 Oct 2009 10:55:32 +0200 Cezary Kaliszyk preparing arguments for res_inst_tac
Wed, 21 Oct 2009 10:30:29 +0200 Cezary Kaliszyk Trying res_inst_tac
Tue, 20 Oct 2009 19:46:22 +0200 Christian Urban started to write code for storing data about the quotients
Tue, 20 Oct 2009 09:37:22 +0200 Christian Urban some minor tuning
Tue, 20 Oct 2009 09:31:34 +0200 Christian Urban tuned and fixed the earlier fix
Tue, 20 Oct 2009 09:21:18 +0200 Christian Urban fixed the abs case in my_reg and added an app case
Tue, 20 Oct 2009 01:17:22 +0200 Christian Urban my version of regularise (still needs to be completed)
Tue, 20 Oct 2009 00:12:05 +0200 Christian Urban moved the map-info and fun-info section to quotient.ML
Sun, 18 Oct 2009 10:34:53 +0200 Cezary Kaliszyk Test if we can already do sth with the transformed theorem.
Sun, 18 Oct 2009 08:44:16 +0200 Christian Urban slight fix and tuning
Sun, 18 Oct 2009 00:52:10 +0200 Christian Urban the command "quotient" can now define more than one quotient at the same time; quotients need to be separated by and
Sat, 17 Oct 2009 16:06:54 +0200 Cezary Kaliszyk Partial simplification of the proof
Sat, 17 Oct 2009 15:42:57 +0200 Cezary Kaliszyk Some QUOTIENTS
Sat, 17 Oct 2009 14:16:57 +0200 Cezary Kaliszyk Only QUOTIENSs are left to fnish proof
Sat, 17 Oct 2009 12:44:58 +0200 Cezary Kaliszyk More higher order unification problems
Sat, 17 Oct 2009 12:31:48 +0200 Cezary Kaliszyk Merged
Sat, 17 Oct 2009 12:31:36 +0200 Cezary Kaliszyk Simplified
Sat, 17 Oct 2009 12:20:56 +0200 Cezary Kaliszyk Further in the proof
Sat, 17 Oct 2009 11:54:50 +0200 Cezary Kaliszyk compose_tac works with the full instantiation.
Sat, 17 Oct 2009 12:19:06 +0200 Christian Urban slightly simplified get_fun function
Sat, 17 Oct 2009 10:40:54 +0200 Cezary Kaliszyk The instantiated version is the same modulo beta
Sat, 17 Oct 2009 10:07:52 +0200 Cezary Kaliszyk Fully manually instantiated. Still fails...
Sat, 17 Oct 2009 09:04:24 +0200 Cezary Kaliszyk Little progress with match/instantiate
Fri, 16 Oct 2009 19:21:05 +0200 Cezary Kaliszyk Fighting with the instantiation
Fri, 16 Oct 2009 17:05:52 +0200 Cezary Kaliszyk Symmetric version of REP_ABS_RSP
Fri, 16 Oct 2009 16:51:01 +0200 Cezary Kaliszyk Progressing with the proof
Fri, 16 Oct 2009 10:54:31 +0200 Cezary Kaliszyk Finally fix get_fun.
Fri, 16 Oct 2009 08:48:56 +0200 Cezary Kaliszyk A fix for one fun_map; doesn't work for more.
Fri, 16 Oct 2009 03:26:54 +0200 Christian Urban fixed the problem with function types; but only type_of works; cterm_of does not work
Thu, 15 Oct 2009 16:51:24 +0200 Cezary Kaliszyk Description of the problem with get_fun.
Thu, 15 Oct 2009 16:36:20 +0200 Cezary Kaliszyk A proper build_goal_term function.
Thu, 15 Oct 2009 12:06:00 +0200 Cezary Kaliszyk Cleaning the code
Thu, 15 Oct 2009 11:57:33 +0200 Cezary Kaliszyk Merged
Thu, 15 Oct 2009 11:56:30 +0200 Cezary Kaliszyk Cleaning the proofs
Thu, 15 Oct 2009 11:55:52 +0200 Cezary Kaliszyk Cleaning the code, part 4
Thu, 15 Oct 2009 11:53:11 +0200 Christian Urban slightly improved tyRel
Thu, 15 Oct 2009 11:42:14 +0200 Cezary Kaliszyk Reordering the code, part 3
Thu, 15 Oct 2009 11:29:38 +0200 Cezary Kaliszyk Reordering the code, part 2.
Thu, 15 Oct 2009 11:25:25 +0200 Cezary Kaliszyk Reordering the code, part 1.
Thu, 15 Oct 2009 11:20:50 +0200 Cezary Kaliszyk Minor cleaning.
Thu, 15 Oct 2009 11:17:27 +0200 Cezary Kaliszyk The definition of Fold1
Thu, 15 Oct 2009 10:25:23 +0200 Cezary Kaliszyk A number of lemmas for REGULARIZE_TAC and regularizing card1.
Wed, 14 Oct 2009 18:13:16 +0200 Cezary Kaliszyk Proving the proper RepAbs version
Wed, 14 Oct 2009 16:23:49 +0200 Cezary Kaliszyk Forgot to save, second part of the commit
Wed, 14 Oct 2009 16:23:32 +0200 Cezary Kaliszyk Manually regularized list_induct2
Wed, 14 Oct 2009 12:05:39 +0200 Cezary Kaliszyk Fixed bug in regularise types. Now we can regularise list.induct.
Wed, 14 Oct 2009 11:23:55 +0200 Cezary Kaliszyk Function for checking recursively if lifting is needed
Wed, 14 Oct 2009 09:50:13 +0200 Cezary Kaliszyk Merge
Wed, 14 Oct 2009 09:47:16 +0200 Cezary Kaliszyk Proper handling of non-lifted quantifiers, testing type freezing.
Tue, 13 Oct 2009 22:22:15 +0200 Christian Urban slight simplification of atomize_thm
Tue, 13 Oct 2009 18:01:54 +0200 Cezary Kaliszyk atomize_thm and meta_quantify.
Tue, 13 Oct 2009 13:37:07 +0200 Cezary Kaliszyk Regularizing HOL all.
Tue, 13 Oct 2009 11:38:35 +0200 Cezary Kaliszyk ":" is used for being in a set, "IN" means something else...
Tue, 13 Oct 2009 11:03:55 +0200 Cezary Kaliszyk First (untested) version of regularize for abstractions.
Tue, 13 Oct 2009 09:26:19 +0200 Christian Urban restored old version
Tue, 13 Oct 2009 00:02:22 +0200 Christian Urban tuned
Mon, 12 Oct 2009 23:39:14 +0200 Christian Urban slightly modified the parser
Mon, 12 Oct 2009 23:16:20 +0200 Christian Urban deleted diagnostic code
Mon, 12 Oct 2009 23:06:14 +0200 Christian Urban added quotient command (you need to update isar-keywords-prove.el)
Mon, 12 Oct 2009 22:44:16 +0200 Christian Urban added new keyword
Mon, 12 Oct 2009 16:31:29 +0200 Cezary Kaliszyk Bounded quantifier
Mon, 12 Oct 2009 15:47:27 +0200 Cezary Kaliszyk The tyREL function.
Mon, 12 Oct 2009 14:30:50 +0200 Christian Urban started some strange functions
Mon, 12 Oct 2009 13:58:31 +0200 Cezary Kaliszyk Further with the manual proof
Fri, 09 Oct 2009 17:05:45 +0200 Cezary Kaliszyk Further experiments with proving induction manually
Fri, 09 Oct 2009 15:03:43 +0200 Cezary Kaliszyk Testing if I can prove the regularized version of induction manually
Thu, 08 Oct 2009 14:27:50 +0200 Christian Urban exported parts of QuotMain into a separate ML-file
Tue, 06 Oct 2009 15:11:30 +0200 Christian Urban consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Tue, 06 Oct 2009 11:56:23 +0200 Christian Urban simplified typedef_quot_type_tac (using MetaSimplifier.rewrite_rule instead of the simplifier)
Tue, 06 Oct 2009 11:41:35 +0200 Christian Urban renamed unlam_def to unabs_def (matching the function abs_def in drule.ML)
Tue, 06 Oct 2009 11:36:08 +0200 Christian Urban tuned; nothing serious
Tue, 06 Oct 2009 09:28:59 +0200 Christian Urban another improvement to unlam_def
Tue, 06 Oct 2009 02:02:51 +0200 Christian Urban one further improvement to unlam_def
Tue, 06 Oct 2009 01:50:13 +0200 Christian Urban simplified the unlam_def function
Mon, 05 Oct 2009 11:54:02 +0200 Christian Urban added an explicit syntax-argument to the function make_def (is needed if the user gives an syntax annotation for quotient types)
Mon, 05 Oct 2009 11:24:32 +0200 Christian Urban used prop_of to get the term of a theorem (replaces crep_thm)
Fri, 02 Oct 2009 11:10:21 +0200 Cezary Kaliszyk Merged
Fri, 02 Oct 2009 11:09:33 +0200 Cezary Kaliszyk First theorem with quantifiers. Learned how to use sledgehammer.
Thu, 01 Oct 2009 16:10:14 +0200 Christian Urban simplified the storage of the map-functions by using TheoryDataFun
Wed, 30 Sep 2009 16:57:09 +0200 Cezary Kaliszyk Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Tue, 29 Sep 2009 22:35:48 +0200 Christian Urban used new cong_tac
Tue, 29 Sep 2009 17:46:18 +0200 Cezary Kaliszyk First version of handling of the universal quantifier
Tue, 29 Sep 2009 15:58:14 +0200 Cezary Kaliszyk Handling abstraction correctly.
Tue, 29 Sep 2009 13:38:27 +0200 Cezary Kaliszyk second test
Tue, 29 Sep 2009 13:24:57 +0200 Cezary Kaliszyk Test line
Tue, 29 Sep 2009 11:55:37 +0200 Cezary Kaliszyk Partially fix lifting of card_suc. The quantified variable still needs to be changed.
Tue, 29 Sep 2009 10:49:31 +0200 Cezary Kaliszyk Tested new build_goal and removed old one, eq_reflection is included in atomize, card_suc tests.
Tue, 29 Sep 2009 09:58:02 +0200 Cezary Kaliszyk Checked again with the version on my hard-drive backedup yesterday and fixed small whitespace issues.
Tue, 29 Sep 2009 09:26:22 +0200 Cezary Kaliszyk Merged
Mon, 28 Sep 2009 23:17:29 +0200 Christian Urban added name to prove
Mon, 28 Sep 2009 19:24:55 +0200 Christian Urban consistent state
Mon, 28 Sep 2009 19:22:28 +0200 Christian Urban some tuning of my code
Mon, 28 Sep 2009 19:15:19 +0200 Cezary Kaliszyk Cleanup
Mon, 28 Sep 2009 19:10:27 +0200 Cezary Kaliszyk Merged
Mon, 28 Sep 2009 18:59:04 +0200 Cezary Kaliszyk Cleaning the code with Makarius
Mon, 28 Sep 2009 15:37:38 +0200 Cezary Kaliszyk Instnatiation with a schematic variable
Mon, 28 Sep 2009 02:39:44 +0200 Christian Urban added ctxt as explicit argument to build_goal; tuned
Fri, 25 Sep 2009 19:26:18 +0200 Christian Urban slightly tuned the comment for unlam
Fri, 25 Sep 2009 17:08:50 +0200 Christian Urban fixed a bug in my code: function typedef_term constructs now now the correct term when the relation is called x
Fri, 25 Sep 2009 16:50:11 +0200 Christian Urban tuned slightly one proof
Fri, 25 Sep 2009 14:50:35 +0200 Cezary Kaliszyk A version of the tactic that exports variables correctly.
Fri, 25 Sep 2009 09:38:16 +0200 Cezary Kaliszyk Minor cleaning: whitespace, commas etc.
Thu, 24 Sep 2009 17:43:26 +0200 Cezary Kaliszyk Correctly handling abstractions while building goals
Thu, 24 Sep 2009 13:32:52 +0200 Cezary Kaliszyk Minor cleanup
Thu, 24 Sep 2009 11:38:20 +0200 Cezary Kaliszyk Found the source for the exception and added a handle for it.
Thu, 24 Sep 2009 09:09:46 +0200 Cezary Kaliszyk Make the tactic use Trueprop_cong and atomize.
Wed, 23 Sep 2009 16:57:56 +0200 Cezary Kaliszyk Proper code for getting the prop out of the theorem.
Wed, 23 Sep 2009 16:44:56 +0200 Cezary Kaliszyk Using "atomize" the versions with arbitrary Trueprops can be proven.
Tue, 22 Sep 2009 17:39:52 +0200 Cezary Kaliszyk Proper definition of CARD and some properties of it.
Tue, 22 Sep 2009 09:42:27 +0200 Christian Urban some comments
Mon, 21 Sep 2009 18:18:29 +0200 Christian Urban merged
Mon, 21 Sep 2009 18:18:01 +0200 Christian Urban slight tuning
Mon, 21 Sep 2009 16:45:44 +0200 Cezary Kaliszyk The tactic with REPEAT, CHANGED and a proper simpset.
Mon, 21 Sep 2009 10:53:01 +0200 Cezary Kaliszyk Merged with my changes from the morning:
Sun, 20 Sep 2009 05:48:49 +0100 Ning some more beautification
Sun, 20 Sep 2009 05:18:36 +0100 Ning beautification of some proofs
Fri, 18 Sep 2009 17:30:33 +0200 Cezary Kaliszyk Added more useful quotient facts.
Fri, 18 Sep 2009 13:44:06 +0200 Cezary Kaliszyk Testing the tactic further.
Thu, 17 Sep 2009 16:55:11 +0200 Cezary Kaliszyk The tactic still only for fset
Thu, 17 Sep 2009 12:06:06 +0200 Cezary Kaliszyk Infrastructure for the tactic
Wed, 16 Sep 2009 11:50:41 +0200 Cezary Kaliszyk More naming/binding suggestions from Makarius
Tue, 15 Sep 2009 11:31:35 +0200 Cezary Kaliszyk Code cleanup
Tue, 15 Sep 2009 10:00:34 +0200 Cezary Kaliszyk Cleaning the code
Tue, 15 Sep 2009 09:59:56 +0200 Cezary Kaliszyk Generalized interpretation, works for all examples.
Fri, 28 Aug 2009 17:12:19 +0200 Cezary Kaliszyk Fixes after suggestions from Makarius:
Fri, 28 Aug 2009 10:01:25 +0200 Cezary Kaliszyk More examples.
Thu, 27 Aug 2009 09:04:39 +0200 Cezary Kaliszyk Use metavariables in the 'non-lambda' definitions
Wed, 26 Aug 2009 11:31:55 +0200 Cezary Kaliszyk Make both kinds of definitions.
Tue, 25 Aug 2009 17:37:50 +0200 Cezary Kaliszyk Changed to the use of "modern interface"
Tue, 25 Aug 2009 14:37:11 +0200 Cezary Kaliszyk Initial version of the function that builds goals.
Tue, 25 Aug 2009 10:35:28 +0200 Cezary Kaliszyk - Build an interpretation for fset from ML level and use it
Tue, 25 Aug 2009 00:30:23 +0200 Christian Urban added the prove command
Fri, 21 Aug 2009 13:56:20 +0200 Cezary Kaliszyk Fixed
Fri, 21 Aug 2009 13:36:58 +0200 Christian Urban tuned
Thu, 20 Aug 2009 14:44:29 +0200 cek UNION - Append theorem
Thu, 20 Aug 2009 10:31:44 +0200 Christian Urban update
Tue, 18 Aug 2009 14:04:51 +0200 Christian Urban updated slightly
Tue, 11 Aug 2009 12:29:15 +0200 Christian Urban initial commit
(0) +1920 tip