# HG changeset patch # User Cezary Kaliszyk # Date 1256738364 -3600 # Node ID 9cc87d02190a6aac114200c7ba0b47b6ccba75a2 # Parent 8934d21534693f4a0bd1a40bef98a023c660c378 First experiments with Lambda diff -r 8934d2153469 -r 9cc87d02190a LamEx.thy --- a/LamEx.thy Wed Oct 28 12:22:06 2009 +0100 +++ b/LamEx.thy Wed Oct 28 14:59:24 2009 +0100 @@ -31,4 +31,70 @@ lemma real_alpha: assumes "t = ([(a,b)]\s)" "a\s" shows "Lam a t = Lam b s" -sorry \ No newline at end of file +sorry + + + + + +(* Construction Site code *) + +lemma perm_rsp: "op = ===> alpha ===> alpha op \ op \" sorry +lemma fresh_rsp: "op = ===> (alpha ===> op =) fresh fresh" sorry +lemma rLam_rsp: "op = ===> (alpha ===> alpha) rLam rLam" sorry + +ML {* val defs = @{thms Var_def App_def Lam_def} *} +ML {* val consts = [@{const_name "rVar"}, @{const_name "rApp"}, @{const_name "rLam"}]; *} + +ML {* val rty = @{typ "rlam"} *} +ML {* val qty = @{typ "lam"} *} +ML {* val rel = @{term "alpha"} *} +ML {* val rel_eqv = (#equiv_thm o hd) (quotdata_lookup @{theory}) *} +ML {* val rel_refl = @{thm EQUIV_REFL} OF [rel_eqv] *} +ML {* val quot = @{thm QUOTIENT_lam} *} +ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rLam_rsp} @ @{thms ho_all_prs ho_ex_prs} *} +ML {* val trans2 = @{thm QUOT_TYPE_I_lam.R_trans2} *} +ML {* val reps_same = @{thm QUOT_TYPE_I_lam.REPS_same} *} + +thm a3 +ML {* val t_a = atomize_thm @{thm a3} *} +ML {* val t_r = regularize t_a rty rel rel_eqv @{context} *} +ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *} +ML {* val abs = findabs rty (prop_of t_a) *} +ML {* val simp_lam_prs_thms = map (make_simp_lam_prs_thm @{context} quot) abs *} +ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *} +ML {* val t_c = simp_allex_prs @{context} quot t_l *} +ML {* val t_defs_sym = add_lower_defs @{context} defs *} +ML {* val t_d = repeat_eqsubst_thm @{context} t_defs_sym t_c *} +ML {* val t_b = MetaSimplifier.rewrite_rule [reps_same] t_d *} +ML {* ObjectLogic.rulify t_b *} + +local_setup {* + make_const_def @{binding lfresh} @{term "fresh :: 'a \ rlam \ bool"} NoSyn @{typ "rlam"} @{typ "lam"} #> snd #> + make_const_def @{binding lperm} @{term "perm :: ('a \ 'a) list \ rlam \ rlam"} NoSyn @{typ "rlam"} @{typ "lam"} #> snd +*} + +ML {* val consts = @{const_name fresh} :: @{const_name perm} :: consts *} +ML {* val defs = @{thms lfresh_def lperm_def} @ defs *} +ML {* val t_t = Toplevel.program (fn () => repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms) *} +ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *} +ML {* val t_c = simp_allex_prs @{context} quot t_l *} +ML {* val t_defs_sym = add_lower_defs @{context} defs *} +ML {* val t_d = repeat_eqsubst_thm @{context} t_defs_sym t_c *} +ML {* val t_b = MetaSimplifier.rewrite_rule [reps_same] t_d *} +ML {* ObjectLogic.rulify t_b *} +ML {* val rr = (add_lower_defs @{context} @{thms lperm_def}) *} +ML {* val rrr = @{thm eq_reflection} OF [rr] *} +ML {* MetaSimplifier.rewrite_rule [rrr] t_b *} + + + + + +ML {* +fun lift_thm_lam lthy t = + lift_thm lthy consts rty qty rel rel_eqv rel_refl quot rsp_thms trans2 reps_same defs t +*} + +ML {* Toplevel.program (fn () => lift_thm_lam @{context} @{thm a3}) *} + diff -r 8934d2153469 -r 9cc87d02190a QuotMain.thy --- a/QuotMain.thy Wed Oct 28 12:22:06 2009 +0100 +++ b/QuotMain.thy Wed Oct 28 14:59:24 2009 +0100 @@ -355,7 +355,7 @@ val thm' = forall_intr_vars thm val thm'' = ObjectLogic.atomize (cprop_of thm') in - Thm.freezeT (Simplifier.rewrite_rule [thm''] thm') + Thm.freezeT @{thm Pure.equal_elim_rule1} OF [thm'', thm'] end *} diff -r 8934d2153469 -r 9cc87d02190a QuotScript.thy --- a/QuotScript.thy Wed Oct 28 12:22:06 2009 +0100 +++ b/QuotScript.thy Wed Oct 28 14:59:24 2009 +0100 @@ -19,6 +19,10 @@ unfolding EQUIV_def REFL_def SYM_def TRANS_def expand_fun_eq by (blast) +lemma EQUIV_REFL: + shows "EQUIV E ==> REFL E" + by (simp add: EQUIV_REFL_SYM_TRANS) + definition "PART_EQUIV E \ (\x. E x x) \ (\x y. E x y = (E x x \ E y y \ (E x = E y)))"