LamEx.thy
changeset 217 9cc87d02190a
parent 203 7384115df9fd
child 219 329111e1c4ba
equal deleted inserted replaced
216:8934d2153469 217:9cc87d02190a
    30 
    30 
    31 lemma real_alpha:
    31 lemma real_alpha:
    32   assumes "t = ([(a,b)]\<bullet>s)" "a\<sharp>s"
    32   assumes "t = ([(a,b)]\<bullet>s)" "a\<sharp>s"
    33   shows "Lam a t = Lam b s"
    33   shows "Lam a t = Lam b s"
    34 sorry
    34 sorry
       
    35 
       
    36 
       
    37 
       
    38 
       
    39 
       
    40 (* Construction Site code *)
       
    41 
       
    42 lemma perm_rsp: "op = ===> alpha ===> alpha op \<bullet> op \<bullet>" sorry
       
    43 lemma fresh_rsp: "op = ===> (alpha ===> op =) fresh fresh" sorry
       
    44 lemma rLam_rsp: "op = ===> (alpha ===> alpha) rLam rLam" sorry
       
    45 
       
    46 ML {* val defs = @{thms Var_def App_def Lam_def} *}
       
    47 ML {* val consts = [@{const_name "rVar"}, @{const_name "rApp"}, @{const_name "rLam"}]; *}
       
    48 
       
    49 ML {* val rty = @{typ "rlam"} *}
       
    50 ML {* val qty = @{typ "lam"} *}
       
    51 ML {* val rel = @{term "alpha"} *}
       
    52 ML {* val rel_eqv = (#equiv_thm o hd) (quotdata_lookup @{theory}) *}
       
    53 ML {* val rel_refl = @{thm EQUIV_REFL} OF [rel_eqv] *}
       
    54 ML {* val quot = @{thm QUOTIENT_lam} *}
       
    55 ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rLam_rsp} @ @{thms ho_all_prs ho_ex_prs} *}
       
    56 ML {* val trans2 = @{thm QUOT_TYPE_I_lam.R_trans2} *}
       
    57 ML {* val reps_same = @{thm QUOT_TYPE_I_lam.REPS_same} *}
       
    58 
       
    59 thm a3
       
    60 ML {* val t_a = atomize_thm @{thm a3} *}
       
    61 ML {* val t_r = regularize t_a rty rel rel_eqv @{context} *}
       
    62 ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
       
    63 ML {* val abs = findabs rty (prop_of t_a) *}
       
    64 ML {* val simp_lam_prs_thms = map (make_simp_lam_prs_thm @{context} quot) abs *}
       
    65 ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *}
       
    66 ML {* val t_c = simp_allex_prs @{context} quot t_l *}
       
    67 ML {* val t_defs_sym = add_lower_defs @{context} defs *}
       
    68 ML {* val t_d = repeat_eqsubst_thm @{context} t_defs_sym t_c *}
       
    69 ML {* val t_b = MetaSimplifier.rewrite_rule [reps_same] t_d *}
       
    70 ML {* ObjectLogic.rulify t_b *}
       
    71 
       
    72 local_setup {*
       
    73   make_const_def @{binding lfresh} @{term "fresh :: 'a \<Rightarrow> rlam \<Rightarrow> bool"} NoSyn @{typ "rlam"} @{typ "lam"} #> snd #>
       
    74   make_const_def @{binding lperm} @{term "perm :: ('a \<times> 'a) list \<Rightarrow> rlam \<Rightarrow> rlam"} NoSyn @{typ "rlam"} @{typ "lam"} #> snd
       
    75 *}
       
    76 
       
    77 ML {* val consts = @{const_name fresh} :: @{const_name perm} :: consts *}
       
    78 ML {* val defs = @{thms lfresh_def lperm_def} @ defs *}
       
    79 ML {* val t_t = Toplevel.program (fn () => repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms) *}
       
    80 ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *}
       
    81 ML {* val t_c = simp_allex_prs @{context} quot t_l *}
       
    82 ML {* val t_defs_sym = add_lower_defs @{context} defs *}
       
    83 ML {* val t_d = repeat_eqsubst_thm @{context} t_defs_sym t_c *}
       
    84 ML {* val t_b = MetaSimplifier.rewrite_rule [reps_same] t_d *}
       
    85 ML {* ObjectLogic.rulify t_b *}
       
    86 ML {* val rr =  (add_lower_defs @{context} @{thms lperm_def}) *}
       
    87 ML {* val rrr = @{thm eq_reflection} OF [rr] *}
       
    88 ML {* MetaSimplifier.rewrite_rule [rrr] t_b *}
       
    89 
       
    90 
       
    91 
       
    92 
       
    93 
       
    94 ML {*
       
    95 fun lift_thm_lam lthy t =
       
    96   lift_thm lthy consts rty qty rel rel_eqv rel_refl quot rsp_thms trans2 reps_same defs t
       
    97 *}
       
    98 
       
    99 ML {* Toplevel.program (fn () => lift_thm_lam @{context} @{thm a3}) *}
       
   100