Nominal/Manual/Term4.thy
changeset 1856 c8e406f64db0
parent 1855 0a306922ace7
child 1862 310b7b768adf
equal deleted inserted replaced
1855:0a306922ace7 1856:c8e406f64db0
    27   done
    27   done
    28 
    28 
    29 thm permute_rtrm4_permute_rtrm4_list.simps
    29 thm permute_rtrm4_permute_rtrm4_list.simps
    30 thm permute_rtrm4_permute_rtrm4_list.simps[simplified repaired]
    30 thm permute_rtrm4_permute_rtrm4_list.simps[simplified repaired]
    31 
    31 
    32 ML define_fv_alpha_export
       
    33 local_setup {* snd o define_fv_alpha_export (Datatype.the_info @{theory} "Term4.rtrm4")
    32 local_setup {* snd o define_fv_alpha_export (Datatype.the_info @{theory} "Term4.rtrm4")
    34   [[[], [], [(NONE, 0, 1, AlphaGen)]], [[], []] ] [] *}
    33   [[[], [], [(NONE, 0, 1, AlphaGen)]], [[], []] ] [] *}
    35 print_theorems
    34 print_theorems
    36 
    35 
    37 lemma fix2: "alpha_rtrm4_list = list_rel alpha_rtrm4"
    36 lemma fix2: "alpha_rtrm4_list = list_rel alpha_rtrm4"
    77 
    76 
    78 local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha4_equivp_no}, []),
    77 local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha4_equivp_no}, []),
    79   (build_equivps [@{term alpha_rtrm4}, @{term alpha_rtrm4_list}] @{thms alpha4_reflp} @{thm alpha_rtrm4_alpha_rtrm4_list.induct} @{thms rtrm4.inject list.inject} @{thms alpha4_inj_no} @{thms rtrm4.distinct list.distinct} @{thms alpha_rtrm4_list.cases alpha_rtrm4.cases} @{thms alpha4_eqvt_no} ctxt)) ctxt)) *}
    78   (build_equivps [@{term alpha_rtrm4}, @{term alpha_rtrm4_list}] @{thms alpha4_reflp} @{thm alpha_rtrm4_alpha_rtrm4_list.induct} @{thms rtrm4.inject list.inject} @{thms alpha4_inj_no} @{thms rtrm4.distinct list.distinct} @{thms alpha_rtrm4_list.cases alpha_rtrm4.cases} @{thms alpha4_eqvt_no} ctxt)) ctxt)) *}
    80 lemmas alpha4_equivp = alpha4_equivp_no[simplified fix2]
    79 lemmas alpha4_equivp = alpha4_equivp_no[simplified fix2]
    81 
    80 
    82 (*lemma fv_rtrm4_rsp:
       
    83   "xa \<approx>4 ya \<Longrightarrow> fv_rtrm4 xa = fv_rtrm4 ya"
       
    84   "x \<approx>4l y \<Longrightarrow> fv_rtrm4_list x = fv_rtrm4_list y"
       
    85   apply (induct rule: alpha_rtrm4_alpha_rtrm4_list.inducts)
       
    86   apply (simp_all add: alpha_gen)
       
    87 done*)
       
    88 
       
    89 
       
    90 quotient_type 
    81 quotient_type 
    91   trm4 = rtrm4 / alpha_rtrm4
    82   trm4 = rtrm4 / alpha_rtrm4
    92 (*and
    83 (*and
    93   trm4list = "rtrm4 list" / alpha_rtrm4_list*)
    84   trm4list = "rtrm4 list" / alpha_rtrm4_list*)
    94   by (simp_all add: alpha4_equivp)
    85   by (simp_all add: alpha4_equivp)
    95 
    86 
    96 local_setup {*
    87 local_setup {*
    97 (fn ctxt => ctxt
    88 (fn ctxt => ctxt
    98  |> snd o (Quotient_Def.quotient_lift_const [] ("Vr4", @{term rVr4}))
    89  |> snd o (Quotient_Def.quotient_lift_const [] ("Vr4", @{term rVr4}))
    99  |> snd o (Quotient_Def.quotient_lift_const [] ("Ap4", @{term rAp4}))
    90  |> snd o (Quotient_Def.quotient_lift_const [@{typ "trm4"}] ("Ap4", @{term rAp4}))
   100  |> snd o (Quotient_Def.quotient_lift_const [] ("Lm4", @{term rLm4})))
    91  |> snd o (Quotient_Def.quotient_lift_const [] ("Lm4", @{term rLm4}))
       
    92  |> snd o (Quotient_Def.quotient_lift_const [] ("fv_trm4", @{term fv_rtrm4})))
   101 *}
    93 *}
   102 print_theorems
    94 print_theorems
   103 
    95 
   104 local_setup {* snd o prove_const_rsp @{binding fv_rtrm4_rsp} [@{term fv_rtrm4}]
    96 
   105   (fn _ => fvbv_rsp_tac @{thm alpha_rtrm4_alpha_rtrm4_list.inducts(1)} @{thms fv_rtrm4_fv_rtrm4_list.simps} 1) *}
    97 
       
    98 lemma fv_rtrm4_rsp:
       
    99   "xa \<approx>4 ya \<Longrightarrow> fv_rtrm4 xa = fv_rtrm4 ya"
       
   100   "x \<approx>4l y \<Longrightarrow> fv_rtrm4_list x = fv_rtrm4_list y"
       
   101   apply (induct rule: alpha_rtrm4_alpha_rtrm4_list.inducts)
       
   102   apply (simp_all add: alpha_gen)
       
   103 done
       
   104 
       
   105 local_setup {* snd o prove_const_rsp [] @{binding fv_rtrm4_rsp'} [@{term fv_rtrm4}]
       
   106   (fn _ => asm_full_simp_tac (@{simpset} addsimps @{thms fv_rtrm4_rsp}) 1) *}
   106 print_theorems
   107 print_theorems
   107 
   108 
   108 local_setup {* snd o prove_const_rsp @{binding rVr4_rsp} [@{term rVr4}]
   109 ML constr_rsp_tac
   109   (fn _ => constr_rsp_tac @{thms alpha4_inj} @{thms fv_rtrm4_rsp} @{thms alpha4_equivp} 1) *}
       
   110 lemma "(alpha_rtrm4 ===> list_rel alpha_rtrm4 ===> alpha_rtrm4) rAp4 rAp4"
       
   111 apply simp
       
   112 apply clarify
       
   113 apply (simp add: alpha4_inj)
       
   114 
   110 
       
   111 local_setup {* snd o prove_const_rsp [] @{binding rVr4_rsp} [@{term rVr4}]
       
   112   (fn _ => constr_rsp_tac @{thms alpha4_inj} @{thms fv_rtrm4_rsp alpha4_equivp} 1) *}
       
   113 local_setup {* snd o prove_const_rsp [] @{binding rLm4_rsp} [@{term rLm4}]
       
   114   (fn _ => constr_rsp_tac @{thms alpha4_inj} @{thms fv_rtrm4_rsp alpha4_equivp} 1) *}
   115 
   115 
   116 local_setup {* snd o prove_const_rsp @{binding rLm4_rsp} [@{term rLm4}]
   116 lemma [quot_respect]:
   117   (fn _ => constr_rsp_tac @{thms alpha4_inj} @{thms fv_rtrm4_rsp} @{thms alpha4_equivp} 1) *}
   117   "(alpha_rtrm4 ===> list_rel alpha_rtrm4 ===> alpha_rtrm4) rAp4 rAp4"
   118 local_setup {* snd o prove_const_rsp @{binding permute_rtrm4_rsp}
   118   by (simp add: alpha4_inj)
   119   [@{term "permute :: perm \<Rightarrow> rtrm4 \<Rightarrow> rtrm4"}, @{term "permute :: perm \<Rightarrow> rtrm4 list \<Rightarrow> rtrm4 list"}] 
   119 
       
   120 (* Maybe also need: @{term "permute :: perm \<Rightarrow> rtrm4 list \<Rightarrow> rtrm4 list"} *)
       
   121 local_setup {* snd o prove_const_rsp [] @{binding permute_rtrm4_rsp}
       
   122   [@{term "permute :: perm \<Rightarrow> rtrm4 \<Rightarrow> rtrm4"}]
   120   (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha4_eqvt}) 1) *}
   123   (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha4_eqvt}) 1) *}
       
   124 print_theorems
   121 
   125 
   122 thm rtrm4.induct
   126 lemma list_rel_rsp:
   123 lemmas trm1_bp_induct = rtrm4.induct[quot_lifted]
   127   "\<lbrakk>\<forall>x y. R x y \<longrightarrow> (\<forall>a b. R a b \<longrightarrow> S x a = T y b); list_rel R x y; list_rel R a b\<rbrakk>
       
   128     \<Longrightarrow> list_rel S x a = list_rel T y b"
       
   129   sorry
       
   130 
       
   131 lemma[quot_respect]:
       
   132   "((R ===> R ===> op =) ===> list_rel R ===> list_rel R ===> op =) list_rel list_rel"
       
   133   by (simp add: list_rel_rsp)
       
   134 
       
   135 lemma[quot_preserve]:
       
   136   assumes a: "Quotient R abs1 rep1"
       
   137   shows "((abs1 ---> abs1 ---> id) ---> map rep1 ---> map rep1 ---> id) list_rel = list_rel"
       
   138   apply (simp add: expand_fun_eq)
       
   139   apply clarify
       
   140   apply (induct_tac xa xb rule: list_induct2')
       
   141   apply (simp_all add: Quotient_abs_rep[OF a])
       
   142   done
       
   143 
       
   144 lemma[quot_preserve]:
       
   145   assumes a: "Quotient R abs1 rep1"
       
   146   shows "(list_rel ((rep1 ---> rep1 ---> id) R) l m) = (l = m)"
       
   147   by (induct l m rule: list_induct2') (simp_all add: Quotient_rel_rep[OF a])
       
   148 
       
   149 lemma bla: "(Ap4 trm4 list = Ap4 trm4a lista) =
       
   150   (trm4 = trm4a \<and> list_rel (op =) list lista)"
       
   151   by (lifting alpha4_inj(2))
       
   152 
       
   153 thm bla[simplified list_rel_eq]
       
   154 
       
   155 lemma " (Lm4 name rtrm4 = Lm4 namea rtrm4a) =
       
   156  (\<exists>pi\<Colon>perm.
       
   157      fv_trm4 rtrm4 - {atom name} = fv_trm4 rtrm4a - {atom namea} \<and>
       
   158      (fv_trm4 rtrm4 - {atom name}) \<sharp>* pi \<and>
       
   159      pi \<bullet> rtrm4 = rtrm4a \<and> pi \<bullet> {atom name} = {atom namea})"
       
   160 
       
   161 ML {* lift_thm [@{typ trm4}] @{context} @{thm alpha4_inj(1)} *}
       
   162 ML {* lift_thm [@{typ trm4}] @{context} @{thm alpha4_inj(2)} *}
       
   163 ML {* lift_thm [@{typ trm4}] @{context} @{thm alpha4_inj(3)[unfolded alpha_gen]} *}
       
   164 ML {* lift_thm [@{typ trm4}] @{context} @{thm rtrm4.induct} *}
       
   165 .
       
   166 
       
   167 (*lemmas trm1_bp_induct = rtrm4.induct[quot_lifted]*)
   124 
   168 
   125 end
   169 end