Quot/Examples/LamEx.thy
changeset 610 2bee5ca44ef5
parent 601 81f40b8bde7b
child 636 520a4084d064
equal deleted inserted replaced
606:38aa6b67a80b 610:2bee5ca44ef5
   167 apply (auto)
   167 apply (auto)
   168 apply (erule alpha.cases)
   168 apply (erule alpha.cases)
   169 apply (simp_all add: rlam.inject alpha_refl)
   169 apply (simp_all add: rlam.inject alpha_refl)
   170 done
   170 done
   171 
   171 
   172 ML {* val qty = @{typ "lam"} *}
       
   173 ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} *}
       
   174 
       
   175 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
       
   176 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *}
       
   177 ML {* fun lift_tac_lam lthy t = lift_tac lthy t *}
       
   178 
   172 
   179 lemma pi_var: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Var a = Var (pi \<bullet> a)"
   173 lemma pi_var: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Var a = Var (pi \<bullet> a)"
   180 apply (tactic {* lift_tac_lam @{context} @{thm pi_var_com} 1 *})
   174 apply (tactic {* lift_tac @{context} @{thm pi_var_com} 1 *})
   181 done
   175 done
   182 
   176 
   183 lemma pi_app: "(pi\<Colon>('x \<times> 'x) list) \<bullet> App (x\<Colon>lam) (xa\<Colon>lam) = App (pi \<bullet> x) (pi \<bullet> xa)"
   177 lemma pi_app: "(pi\<Colon>('x \<times> 'x) list) \<bullet> App (x\<Colon>lam) (xa\<Colon>lam) = App (pi \<bullet> x) (pi \<bullet> xa)"
   184 apply (tactic {* lift_tac_lam @{context} @{thm pi_app_com} 1 *})
   178 apply (tactic {* lift_tac @{context} @{thm pi_app_com} 1 *})
   185 done
   179 done
   186 
   180 
   187 lemma pi_lam: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Lam (a\<Colon>name) (x\<Colon>lam) = Lam (pi \<bullet> a) (pi \<bullet> x)"
   181 lemma pi_lam: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Lam (a\<Colon>name) (x\<Colon>lam) = Lam (pi \<bullet> a) (pi \<bullet> x)"
   188 apply (tactic {* lift_tac_lam @{context} @{thm pi_lam_com} 1 *})
   182 apply (tactic {* lift_tac @{context} @{thm pi_lam_com} 1 *})
   189 done
   183 done
   190 
   184 
   191 lemma fv_var: "fv (Var (a\<Colon>name)) = {a}"
   185 lemma fv_var: "fv (Var (a\<Colon>name)) = {a}"
   192 apply (tactic {* lift_tac_lam @{context} @{thm rfv_var} 1 *})
   186 apply (tactic {* lift_tac @{context} @{thm rfv_var} 1 *})
   193 done
   187 done
   194 
   188 
   195 lemma fv_app: "fv (App (x\<Colon>lam) (xa\<Colon>lam)) = fv x \<union> fv xa"
   189 lemma fv_app: "fv (App (x\<Colon>lam) (xa\<Colon>lam)) = fv x \<union> fv xa"
   196 apply (tactic {* lift_tac_lam @{context} @{thm rfv_app} 1 *})
   190 apply (tactic {* lift_tac @{context} @{thm rfv_app} 1 *})
   197 done
   191 done
   198 
   192 
   199 lemma fv_lam: "fv (Lam (a\<Colon>name) (x\<Colon>lam)) = fv x - {a}"
   193 lemma fv_lam: "fv (Lam (a\<Colon>name) (x\<Colon>lam)) = fv x - {a}"
   200 apply (tactic {* lift_tac_lam @{context} @{thm rfv_lam} 1 *})
   194 apply (tactic {* lift_tac @{context} @{thm rfv_lam} 1 *})
   201 done
   195 done
   202 
   196 
   203 lemma a1: "(a\<Colon>name) = (b\<Colon>name) \<Longrightarrow> Var a = Var b"
   197 lemma a1: "(a\<Colon>name) = (b\<Colon>name) \<Longrightarrow> Var a = Var b"
   204 apply (tactic {* lift_tac_lam @{context} @{thm a1} 1 *})
   198 apply (tactic {* lift_tac @{context} @{thm a1} 1 *})
   205 done
   199 done
   206 
   200 
   207 lemma a2: "\<lbrakk>(x\<Colon>lam) = (xa\<Colon>lam); (xb\<Colon>lam) = (xc\<Colon>lam)\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
   201 lemma a2: "\<lbrakk>(x\<Colon>lam) = (xa\<Colon>lam); (xb\<Colon>lam) = (xc\<Colon>lam)\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
   208 apply (tactic {* lift_tac_lam @{context} @{thm a2} 1 *})
   202 apply (tactic {* lift_tac @{context} @{thm a2} 1 *})
   209 done
   203 done
   210 
   204 
   211 lemma a3: "\<lbrakk>(x\<Colon>lam) = [(a\<Colon>name, b\<Colon>name)] \<bullet> (xa\<Colon>lam); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> Lam a x = Lam b xa"
   205 lemma a3: "\<lbrakk>(x\<Colon>lam) = [(a\<Colon>name, b\<Colon>name)] \<bullet> (xa\<Colon>lam); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> Lam a x = Lam b xa"
   212 apply (tactic {* lift_tac_lam @{context} @{thm a3} 1 *})
   206 apply (tactic {* lift_tac @{context} @{thm a3} 1 *})
   213 done
   207 done
   214 
   208 
   215 lemma alpha_cases: "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
   209 lemma alpha_cases: "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
   216      \<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P;
   210      \<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P;
   217      \<And>x a b xa. \<lbrakk>a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> P\<rbrakk>
   211      \<And>x a b xa. \<lbrakk>a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> P\<rbrakk>
   218     \<Longrightarrow> P"
   212     \<Longrightarrow> P"
   219 apply (tactic {* lift_tac_lam @{context} @{thm alpha.cases} 1 *})
   213 apply (tactic {* lift_tac @{context} @{thm alpha.cases} 1 *})
   220 done
   214 done
   221 
   215 
   222 lemma alpha_induct: "\<lbrakk>(qx\<Colon>lam) = (qxa\<Colon>lam); \<And>(a\<Colon>name) b\<Colon>name. a = b \<Longrightarrow> (qxb\<Colon>lam \<Rightarrow> lam \<Rightarrow> bool) (Var a) (Var b);
   216 lemma alpha_induct: "\<lbrakk>(qx\<Colon>lam) = (qxa\<Colon>lam); \<And>(a\<Colon>name) b\<Colon>name. a = b \<Longrightarrow> (qxb\<Colon>lam \<Rightarrow> lam \<Rightarrow> bool) (Var a) (Var b);
   223      \<And>(x\<Colon>lam) (xa\<Colon>lam) (xb\<Colon>lam) xc\<Colon>lam. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc);
   217      \<And>(x\<Colon>lam) (xa\<Colon>lam) (xb\<Colon>lam) xc\<Colon>lam. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc);
   224      \<And>(x\<Colon>lam) (a\<Colon>name) (b\<Colon>name) xa\<Colon>lam.
   218      \<And>(x\<Colon>lam) (a\<Colon>name) (b\<Colon>name) xa\<Colon>lam.
   225         \<lbrakk>x = [(a, b)] \<bullet> xa; qxb x ([(a, b)] \<bullet> xa); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> qxb (Lam a x) (Lam b xa)\<rbrakk>
   219         \<lbrakk>x = [(a, b)] \<bullet> xa; qxb x ([(a, b)] \<bullet> xa); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> qxb (Lam a x) (Lam b xa)\<rbrakk>
   226     \<Longrightarrow> qxb qx qxa"
   220     \<Longrightarrow> qxb qx qxa"
   227 apply (tactic {* lift_tac_lam @{context} @{thm alpha.induct} 1 *})
   221 apply (tactic {* lift_tac @{context} @{thm alpha.induct} 1 *})
   228 done
   222 done
   229 
   223 
   230 lemma var_inject: "(Var a = Var b) = (a = b)"
   224 lemma var_inject: "(Var a = Var b) = (a = b)"
   231 apply (tactic {* lift_tac_lam @{context} @{thm rvar_inject} 1 *})
   225 apply (tactic {* lift_tac @{context} @{thm rvar_inject} 1 *})
   232 done
   226 done
   233 
   227 
   234 lemma lam_induct:" \<lbrakk>\<And>name. P (Var name); \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2);
   228 lemma lam_induct:" \<lbrakk>\<And>name. P (Var name); \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2);
   235               \<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> \<Longrightarrow> P lam"
   229               \<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> \<Longrightarrow> P lam"
   236 apply (tactic {* lift_tac_lam @{context} @{thm rlam.induct} 1 *})
   230 apply (tactic {* lift_tac @{context} @{thm rlam.induct} 1 *})
   237 done
   231 done
   238 
   232 
   239 lemma var_supp:
   233 lemma var_supp:
   240   shows "supp (Var a) = ((supp a)::name set)"
   234   shows "supp (Var a) = ((supp a)::name set)"
   241   apply(simp add: supp_def)
   235   apply(simp add: supp_def)