Tue, 28 Jun 2011 14:45:30 +0900 trying new fcb in let/subst
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 28 Jun 2011 14:45:30 +0900] rev 2921
trying new fcb in let/subst
Tue, 28 Jun 2011 14:30:30 +0900 Leftover only inj and eqvt
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 28 Jun 2011 14:30:30 +0900] rev 2920
Leftover only inj and eqvt
Tue, 28 Jun 2011 14:18:26 +0900 eapply fcb ok
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 28 Jun 2011 14:18:26 +0900] rev 2919
eapply fcb ok
Tue, 28 Jun 2011 14:01:15 +0900 Removed Inl and Inr
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 28 Jun 2011 14:01:15 +0900] rev 2918
Removed Inl and Inr
Tue, 28 Jun 2011 14:49:48 +0100 relaxed type in fcb
Christian Urban <urbanc@in.tum.de> [Tue, 28 Jun 2011 14:49:48 +0100] rev 2917
relaxed type in fcb
Tue, 28 Jun 2011 14:34:07 +0100 fcb with explicit bn function
Christian Urban <urbanc@in.tum.de> [Tue, 28 Jun 2011 14:34:07 +0100] rev 2916
fcb with explicit bn function
Tue, 28 Jun 2011 14:01:52 +0100 added let-rec example
Christian Urban <urbanc@in.tum.de> [Tue, 28 Jun 2011 14:01:52 +0100] rev 2915
added let-rec example
Tue, 28 Jun 2011 12:36:34 +0900 Experiments with res
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 28 Jun 2011 12:36:34 +0900] rev 2914
Experiments with res
Tue, 28 Jun 2011 00:48:57 +0100 proved the fcb also for sets (no restriction yet)
Christian Urban <urbanc@in.tum.de> [Tue, 28 Jun 2011 00:48:57 +0100] rev 2913
proved the fcb also for sets (no restriction yet)
Tue, 28 Jun 2011 00:30:30 +0100 copied all work to Lambda.thy; had to derive a special version of fcb1 for concrete atom
Christian Urban <urbanc@in.tum.de> [Tue, 28 Jun 2011 00:30:30 +0100] rev 2912
copied all work to Lambda.thy; had to derive a special version of fcb1 for concrete atom
Mon, 27 Jun 2011 22:51:42 +0100 streamlined the fcb-proof and made fcb1 a special case of fcb
Christian Urban <urbanc@in.tum.de> [Mon, 27 Jun 2011 22:51:42 +0100] rev 2911
streamlined the fcb-proof and made fcb1 a special case of fcb
Mon, 27 Jun 2011 19:22:10 +0100 completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de> [Mon, 27 Jun 2011 19:22:10 +0100] rev 2910
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Mon, 27 Jun 2011 19:15:18 +0100 fcb for multible (list) binders; at the moment all of them have to have the same sort (at-class); this should also work for set binders, but not yet for restriction.
Christian Urban <urbanc@in.tum.de> [Mon, 27 Jun 2011 19:15:18 +0100] rev 2909
fcb for multible (list) binders; at the moment all of them have to have the same sort (at-class); this should also work for set binders, but not yet for restriction.
Mon, 27 Jun 2011 19:13:55 +0100 renamed ds to dset (disagreement set)
Christian Urban <urbanc@in.tum.de> [Mon, 27 Jun 2011 19:13:55 +0100] rev 2908
renamed ds to dset (disagreement set)
Mon, 27 Jun 2011 12:15:21 +0100 added small lemma about disagreement set
Christian Urban <urbanc@in.tum.de> [Mon, 27 Jun 2011 12:15:21 +0100] rev 2907
added small lemma about disagreement set
Mon, 27 Jun 2011 08:42:02 +0900 merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 27 Jun 2011 08:42:02 +0900] rev 2906
merge
Mon, 27 Jun 2011 08:38:54 +0900 New-style fcb for multiple binders.
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 27 Jun 2011 08:38:54 +0900] rev 2905
New-style fcb for multiple binders.
Mon, 27 Jun 2011 04:01:55 +0900 equality of lst_binder and a few helper lemmas
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 27 Jun 2011 04:01:55 +0900] rev 2904
equality of lst_binder and a few helper lemmas [l]lst. T = [l]lst. S <-> T = S
Sun, 26 Jun 2011 21:42:07 +0100 only one of the fcb precondistions are needed (probably the same with the perm-conditions)
Christian Urban <urbanc@in.tum.de> [Sun, 26 Jun 2011 21:42:07 +0100] rev 2903
only one of the fcb precondistions are needed (probably the same with the perm-conditions)
Sun, 26 Jun 2011 17:55:22 +0100 another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
Christian Urban <urbanc@in.tum.de> [Sun, 26 Jun 2011 17:55:22 +0100] rev 2902
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
Sat, 25 Jun 2011 22:51:43 +0100 did all cases, except the multiple binder case
Christian Urban <urbanc@in.tum.de> [Sat, 25 Jun 2011 22:51:43 +0100] rev 2901
did all cases, except the multiple binder case
Sat, 25 Jun 2011 21:28:24 +0100 an alternative FCB for Abs_lst1; seems simpler but not as simple as I thought; not sure whether it generalises to multiple binders.
Christian Urban <urbanc@in.tum.de> [Sat, 25 Jun 2011 21:28:24 +0100] rev 2900
an alternative FCB for Abs_lst1; seems simpler but not as simple as I thought; not sure whether it generalises to multiple binders.
Fri, 24 Jun 2011 09:42:44 +0100 except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de> [Fri, 24 Jun 2011 09:42:44 +0100] rev 2899
except for the interated binder case, finished definition in Calssical.thy
Fri, 24 Jun 2011 11:18:18 +0900 Make examples work with non-precompiled image
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 24 Jun 2011 11:18:18 +0900] rev 2898
Make examples work with non-precompiled image
Fri, 24 Jun 2011 11:15:22 +0900 Remove Lambda_add.thy
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 24 Jun 2011 11:15:22 +0900] rev 2897
Remove Lambda_add.thy
Fri, 24 Jun 2011 11:14:58 +0900 The examples in Lambda_add can be defined by nominal_function directly
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 24 Jun 2011 11:14:58 +0900] rev 2896
The examples in Lambda_add can be defined by nominal_function directly
Fri, 24 Jun 2011 11:03:53 +0900 Theory name changes for JEdit
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 24 Jun 2011 11:03:53 +0900] rev 2895
Theory name changes for JEdit
Fri, 24 Jun 2011 10:54:31 +0900 More usual names for substitution properties
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 24 Jun 2011 10:54:31 +0900] rev 2894
More usual names for substitution properties
Fri, 24 Jun 2011 10:30:06 +0900 Second Fixed Point Theorem
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 24 Jun 2011 10:30:06 +0900] rev 2893
Second Fixed Point Theorem
Fri, 24 Jun 2011 10:12:47 +0900 Speed-up the completeness proof.
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 24 Jun 2011 10:12:47 +0900] rev 2892
Speed-up the completeness proof.
Thu, 23 Jun 2011 22:21:43 +0100 the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de> [Thu, 23 Jun 2011 22:21:43 +0100] rev 2891
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Thu, 23 Jun 2011 13:09:17 +0100 added file
Christian Urban <urbanc@in.tum.de> [Thu, 23 Jun 2011 13:09:17 +0100] rev 2890
added file
Thu, 23 Jun 2011 12:28:25 +0100 expanded the example
Christian Urban <urbanc@in.tum.de> [Thu, 23 Jun 2011 12:28:25 +0100] rev 2889
expanded the example
Thu, 23 Jun 2011 11:30:39 +0100 fixed nasty bug with type variables in nominal_datatypes; this included to be careful with the output of the inductive and function package
Christian Urban <urbanc@in.tum.de> [Thu, 23 Jun 2011 11:30:39 +0100] rev 2888
fixed nasty bug with type variables in nominal_datatypes; this included to be careful with the output of the inductive and function package
Wed, 22 Jun 2011 14:14:54 +0100 tuned
Christian Urban <urbanc@in.tum.de> [Wed, 22 Jun 2011 14:14:54 +0100] rev 2887
tuned
Wed, 22 Jun 2011 13:40:25 +0100 deleted some dead code
Christian Urban <urbanc@in.tum.de> [Wed, 22 Jun 2011 13:40:25 +0100] rev 2886
deleted some dead code
Wed, 22 Jun 2011 12:18:22 +0100 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de> [Wed, 22 Jun 2011 12:18:22 +0100] rev 2885
some rudimentary infrastructure for storing data about nominal datatypes
Wed, 22 Jun 2011 17:57:15 +0900 constants with the same names
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 22 Jun 2011 17:57:15 +0900] rev 2884
constants with the same names
Wed, 22 Jun 2011 04:49:56 +0900 Quotients/TODO addtion
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 22 Jun 2011 04:49:56 +0900] rev 2883
Quotients/TODO addtion
Tue, 21 Jun 2011 23:59:36 +0900 Minor
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 21 Jun 2011 23:59:36 +0900] rev 2882
Minor
Tue, 21 Jun 2011 10:39:25 +0900 merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 21 Jun 2011 10:39:25 +0900] rev 2881
merge
Tue, 21 Jun 2011 10:37:43 +0900 spelling
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 21 Jun 2011 10:37:43 +0900] rev 2880
spelling
Mon, 20 Jun 2011 20:09:51 +0900 merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 20 Jun 2011 20:09:51 +0900] rev 2879
merge
Mon, 20 Jun 2011 20:08:16 +0900 Abs_set_fcb
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 20 Jun 2011 20:08:16 +0900] rev 2878
Abs_set_fcb
Mon, 20 Jun 2011 20:09:30 +0900 function for let-rec
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 20 Jun 2011 20:09:30 +0900] rev 2877
function for let-rec
Mon, 20 Jun 2011 10:16:12 +0900 TODO/minor
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 20 Jun 2011 10:16:12 +0900] rev 2876
TODO/minor
Mon, 20 Jun 2011 09:59:18 +0900 Move lst_fcb to Nominal2_Abs
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 20 Jun 2011 09:59:18 +0900] rev 2875
Move lst_fcb to Nominal2_Abs
Mon, 20 Jun 2011 09:38:57 +0900 More minor TODOs
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 20 Jun 2011 09:38:57 +0900] rev 2874
More minor TODOs
Mon, 20 Jun 2011 09:36:16 +0900 Update TODO
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 20 Jun 2011 09:36:16 +0900] rev 2873
Update TODO
Mon, 20 Jun 2011 09:29:42 +0900 Let/minor
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 20 Jun 2011 09:29:42 +0900] rev 2872
Let/minor
Mon, 20 Jun 2011 08:50:13 +0900 Update Quotient/TODO and remove some attic code
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 20 Jun 2011 08:50:13 +0900] rev 2871
Update Quotient/TODO and remove some attic code
Sun, 19 Jun 2011 13:14:37 +0900 merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Sun, 19 Jun 2011 13:14:37 +0900] rev 2870
merge
Sun, 19 Jun 2011 13:10:15 +0900 little on cps2
Cezary Kaliszyk <kaliszyk@in.tum.de> [Sun, 19 Jun 2011 13:10:15 +0900] rev 2869
little on cps2
Thu, 16 Jun 2011 20:07:03 +0100 got rid of the boolean flag in the raw_equivariance function
Christian Urban <urbanc@in.tum.de> [Thu, 16 Jun 2011 20:07:03 +0100] rev 2868
got rid of the boolean flag in the raw_equivariance function
Thu, 16 Jun 2011 23:11:50 +0900 Fix
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 16 Jun 2011 23:11:50 +0900] rev 2867
Fix
Thu, 16 Jun 2011 22:00:52 +0900 merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 16 Jun 2011 22:00:52 +0900] rev 2866
merge
Thu, 16 Jun 2011 21:23:38 +0900 CPS3 can be defined with eqvt_rhs.
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 16 Jun 2011 21:23:38 +0900] rev 2865
CPS3 can be defined with eqvt_rhs.
Thu, 16 Jun 2011 13:32:36 +0100 moved for the moment CPS translations into the example directory
Christian Urban <urbanc@in.tum.de> [Thu, 16 Jun 2011 13:32:36 +0100] rev 2864
moved for the moment CPS translations into the example directory
Thu, 16 Jun 2011 13:14:53 +0100 merged
Christian Urban <urbanc@in.tum.de> [Thu, 16 Jun 2011 13:14:53 +0100] rev 2863
merged
Thu, 16 Jun 2011 13:14:16 +0100 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de> [Thu, 16 Jun 2011 13:14:16 +0100] rev 2862
added eqvt_at and invariant for boths sides of the equations
Thu, 16 Jun 2011 20:56:30 +0900 Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 16 Jun 2011 20:56:30 +0900] rev 2861
Added the CPS translation experiments. CPS1 comes with all the proofs, CPS2,3 just have the function and need eqvt_rhs to finish the obligations.
Thu, 16 Jun 2011 12:12:25 +0100 added a test that every function must be of pt-sort
Christian Urban <urbanc@in.tum.de> [Thu, 16 Jun 2011 12:12:25 +0100] rev 2860
added a test that every function must be of pt-sort
Thu, 16 Jun 2011 11:03:01 +0100 all tests work again
Christian Urban <urbanc@in.tum.de> [Thu, 16 Jun 2011 11:03:01 +0100] rev 2859
all tests work again
Wed, 15 Jun 2011 22:36:59 +0100 added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
Christian Urban <urbanc@in.tum.de> [Wed, 15 Jun 2011 22:36:59 +0100] rev 2858
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
Wed, 15 Jun 2011 12:52:48 +0100 added an abstract
Christian Urban <urbanc@in.tum.de> [Wed, 15 Jun 2011 12:52:48 +0100] rev 2857
added an abstract
Wed, 15 Jun 2011 12:32:40 +0100 added a stub for function paper; "isabelle make fnpaper"
Christian Urban <urbanc@in.tum.de> [Wed, 15 Jun 2011 12:32:40 +0100] rev 2856
added a stub for function paper; "isabelle make fnpaper"
Wed, 15 Jun 2011 11:06:48 +0900 merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 15 Jun 2011 11:06:48 +0900] rev 2855
merge
Wed, 15 Jun 2011 11:06:31 +0900 one TODO and one Problem?
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 15 Jun 2011 11:06:31 +0900] rev 2854
one TODO and one Problem?
Wed, 15 Jun 2011 09:51:26 +0900 merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 15 Jun 2011 09:51:26 +0900] rev 2853
merge
Wed, 15 Jun 2011 09:50:53 +0900 Some TODOs
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 15 Jun 2011 09:50:53 +0900] rev 2852
Some TODOs
Wed, 15 Jun 2011 09:31:59 +0900 merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 15 Jun 2011 09:31:59 +0900] rev 2851
merge
Wed, 15 Jun 2011 09:25:36 +0900 TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 15 Jun 2011 09:25:36 +0900] rev 2850
TypeSchemes work with 'default'.
Tue, 14 Jun 2011 19:11:44 +0100 tuned some proofs
Christian Urban <urbanc@in.tum.de> [Tue, 14 Jun 2011 19:11:44 +0100] rev 2849
tuned some proofs
Tue, 14 Jun 2011 14:07:07 +0100 fixed the problem when giving a complex default-term; the fundef lemmas in Nominal_Base were not general enough
Christian Urban <urbanc@in.tum.de> [Tue, 14 Jun 2011 14:07:07 +0100] rev 2848
fixed the problem when giving a complex default-term; the fundef lemmas in Nominal_Base were not general enough
Tue, 14 Jun 2011 11:43:06 +0100 tuned
Christian Urban <urbanc@in.tum.de> [Tue, 14 Jun 2011 11:43:06 +0100] rev 2847
tuned
Fri, 10 Jun 2011 15:52:17 +0900 Move working examples before non-working ones
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 10 Jun 2011 15:52:17 +0900] rev 2846
Move working examples before non-working ones
Fri, 10 Jun 2011 15:45:49 +0900 Optimized proofs and removed some garbage.
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 10 Jun 2011 15:45:49 +0900] rev 2845
Optimized proofs and removed some garbage.
Fri, 10 Jun 2011 15:31:14 +0900 merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 10 Jun 2011 15:31:14 +0900] rev 2844
merge
Fri, 10 Jun 2011 15:30:09 +0900 Slightly modify fcb for list1 and put in common place.
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 10 Jun 2011 15:30:09 +0900] rev 2843
Slightly modify fcb for list1 and put in common place.
Fri, 10 Jun 2011 09:00:24 +0900 Experiments with Let
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 10 Jun 2011 09:00:24 +0900] rev 2842
Experiments with Let
Thu, 09 Jun 2011 15:34:51 +0900 Eval can be defined with additional freshness
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 09 Jun 2011 15:34:51 +0900] rev 2841
Eval can be defined with additional freshness
Thu, 09 Jun 2011 15:03:58 +0900 Minor simplification
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 09 Jun 2011 15:03:58 +0900] rev 2840
Minor simplification
Thu, 09 Jun 2011 11:10:41 +0900 abs_res_fcb will be enough to finish the multiple-recursive proof, if we have a working 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 09 Jun 2011 11:10:41 +0900] rev 2839
abs_res_fcb will be enough to finish the multiple-recursive proof, if we have a working 'default'.
Thu, 09 Jun 2011 09:44:51 +0900 More experiments with 'default'
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 09 Jun 2011 09:44:51 +0900] rev 2838
More experiments with 'default'
Wed, 08 Jun 2011 21:44:03 +0900 Finished the proof with the invariant
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 08 Jun 2011 21:44:03 +0900] rev 2837
Finished the proof with the invariant
Wed, 08 Jun 2011 21:32:35 +0900 Issue with 'default'
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 08 Jun 2011 21:32:35 +0900] rev 2836
Issue with 'default'
Wed, 08 Jun 2011 12:30:56 +0100 merged
Christian Urban <urbanc@in.tum.de> [Wed, 08 Jun 2011 12:30:56 +0100] rev 2835
merged
Wed, 08 Jun 2011 12:30:46 +0100 merged
Christian Urban <urbanc@in.tum.de> [Wed, 08 Jun 2011 12:30:46 +0100] rev 2834
merged
Wed, 08 Jun 2011 08:44:01 +0100 using the option "default" the function package allows one to give an explicit default value
Christian Urban <urbanc@in.tum.de> [Wed, 08 Jun 2011 08:44:01 +0100] rev 2833
using the option "default" the function package allows one to give an explicit default value
Wed, 08 Jun 2011 17:52:06 +0900 Simpler proof of TypeSchemes/substs
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 08 Jun 2011 17:52:06 +0900] rev 2832
Simpler proof of TypeSchemes/substs
Wed, 08 Jun 2011 17:25:54 +0900 Simplify fcb_res
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 08 Jun 2011 17:25:54 +0900] rev 2831
Simplify fcb_res
Wed, 08 Jun 2011 09:56:39 +0900 FCB for res binding and simplified proof of subst for type schemes
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 08 Jun 2011 09:56:39 +0900] rev 2830
FCB for res binding and simplified proof of subst for type schemes
Wed, 08 Jun 2011 07:06:20 +0900 Simplify ln-trans proof
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 08 Jun 2011 07:06:20 +0900] rev 2829
Simplify ln-trans proof
Wed, 08 Jun 2011 07:02:52 +0900 cbvs can be easily defined without an invariant
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 08 Jun 2011 07:02:52 +0900] rev 2828
cbvs can be easily defined without an invariant
Tue, 07 Jun 2011 20:58:00 +0100 defined the "count-bound-variables-occurences" function which has an accumulator like trans
Christian Urban <urbanc@in.tum.de> [Tue, 07 Jun 2011 20:58:00 +0100] rev 2827
defined the "count-bound-variables-occurences" function which has an accumulator like trans
Tue, 07 Jun 2011 17:45:38 +0100 merged
Christian Urban <urbanc@in.tum.de> [Tue, 07 Jun 2011 17:45:38 +0100] rev 2826
merged
Tue, 07 Jun 2011 23:42:12 +0900 remove garbage (proofs that assumes the invariant outside function)
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 07 Jun 2011 23:42:12 +0900] rev 2825
remove garbage (proofs that assumes the invariant outside function)
Tue, 07 Jun 2011 23:38:39 +0900 Proof of trans with invariant
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 07 Jun 2011 23:38:39 +0900] rev 2824
Proof of trans with invariant
Tue, 07 Jun 2011 23:22:58 +0900 Testing invariant in Lambda_F_T
Cezary Kaliszyk <kaliszyk@in.tum.de> [Tue, 07 Jun 2011 23:22:58 +0900] rev 2823
Testing invariant in Lambda_F_T
Tue, 07 Jun 2011 10:40:06 +0100 cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
Christian Urban <urbanc@in.tum.de> [Tue, 07 Jun 2011 10:40:06 +0100] rev 2822
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
Tue, 07 Jun 2011 08:52:59 +0100 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de> [Tue, 07 Jun 2011 08:52:59 +0100] rev 2821
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Mon, 06 Jun 2011 13:11:04 +0100 slightly stronger property in fundef_ex_prop
Christian Urban <urbanc@in.tum.de> [Mon, 06 Jun 2011 13:11:04 +0100] rev 2820
slightly stronger property in fundef_ex_prop
Sun, 05 Jun 2011 21:14:23 +0100 added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de> [Sun, 05 Jun 2011 21:14:23 +0100] rev 2819
added an option for an invariant (at the moment only a stub)
Sun, 05 Jun 2011 16:58:18 +0100 added a more general lemma fro fundef_ex1
Christian Urban <urbanc@in.tum.de> [Sun, 05 Jun 2011 16:58:18 +0100] rev 2818
added a more general lemma fro fundef_ex1
Sat, 04 Jun 2011 14:50:57 +0900 Trying the induction on the graph
Cezary Kaliszyk <kaliszyk@in.tum.de> [Sat, 04 Jun 2011 14:50:57 +0900] rev 2817
Trying the induction on the graph
Sat, 04 Jun 2011 09:07:50 +0900 Finish and test the locale approach
Cezary Kaliszyk <kaliszyk@in.tum.de> [Sat, 04 Jun 2011 09:07:50 +0900] rev 2816
Finish and test the locale approach
Fri, 03 Jun 2011 22:31:44 +0900 FiniteSupp precondition in the function is enough to get rid of completeness obligationss
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 03 Jun 2011 22:31:44 +0900] rev 2815
FiniteSupp precondition in the function is enough to get rid of completeness obligationss
Fri, 03 Jun 2011 12:46:23 +0100 recursion combinator inside a locale
Christian Urban <urbanc@in.tum.de> [Fri, 03 Jun 2011 12:46:23 +0100] rev 2814
recursion combinator inside a locale
Fri, 03 Jun 2011 18:33:11 +0900 merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 03 Jun 2011 18:33:11 +0900] rev 2813
merge
Fri, 03 Jun 2011 18:32:22 +0900 F for lambda used to define translation to locally nameless
Cezary Kaliszyk <kaliszyk@in.tum.de> [Fri, 03 Jun 2011 18:32:22 +0900] rev 2812
F for lambda used to define translation to locally nameless
Thu, 02 Jun 2011 16:15:18 +0100 typo
Christian Urban <urbanc@in.tum.de> [Thu, 02 Jun 2011 16:15:18 +0100] rev 2811
typo
Thu, 02 Jun 2011 15:35:33 +0100 removed dead code
Christian Urban <urbanc@in.tum.de> [Thu, 02 Jun 2011 15:35:33 +0100] rev 2810
removed dead code
Thu, 02 Jun 2011 22:24:33 +0900 finished the missing obligations
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 02 Jun 2011 22:24:33 +0900] rev 2809
finished the missing obligations
Thu, 02 Jun 2011 12:14:03 +0100 merged
Christian Urban <urbanc@in.tum.de> [Thu, 02 Jun 2011 12:14:03 +0100] rev 2808
merged
Thu, 02 Jun 2011 12:09:31 +0100 a test with a recursion combinator defined on top of nominal_primrec
Christian Urban <urbanc@in.tum.de> [Thu, 02 Jun 2011 12:09:31 +0100] rev 2807
a test with a recursion combinator defined on top of nominal_primrec
Thu, 02 Jun 2011 16:41:09 +0900 Use FCB to simplify proof
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 02 Jun 2011 16:41:09 +0900] rev 2806
Use FCB to simplify proof
Thu, 02 Jun 2011 10:11:50 +0900 merge
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 02 Jun 2011 10:11:50 +0900] rev 2805
merge
Thu, 02 Jun 2011 10:09:23 +0900 Remove SMT
Cezary Kaliszyk <kaliszyk@in.tum.de> [Thu, 02 Jun 2011 10:09:23 +0900] rev 2804
Remove SMT
Wed, 01 Jun 2011 22:55:14 +0100 hopefully final fix for ho-functions
Christian Urban <urbanc@in.tum.de> [Wed, 01 Jun 2011 22:55:14 +0100] rev 2803
hopefully final fix for ho-functions
Wed, 01 Jun 2011 21:03:30 +0100 first test to fix the problem with free variables
Christian Urban <urbanc@in.tum.de> [Wed, 01 Jun 2011 21:03:30 +0100] rev 2802
first test to fix the problem with free variables
(0) -1000 -120 +120 tip