diff -r ca37f4b6457c -r 8dfae5d97387 Quot/Examples/LamEx.thy --- a/Quot/Examples/LamEx.thy Tue Dec 08 11:17:56 2009 +0100 +++ b/Quot/Examples/LamEx.thy Tue Dec 08 11:20:01 2009 +0100 @@ -169,54 +169,48 @@ apply (simp_all add: rlam.inject alpha_refl) done -ML {* val qty = @{typ "lam"} *} -ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} *} - -ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *} -ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *} -ML {* fun lift_tac_lam lthy t = lift_tac lthy t *} lemma pi_var: "(pi\('x \ 'x) list) \ Var a = Var (pi \ a)" -apply (tactic {* lift_tac_lam @{context} @{thm pi_var_com} 1 *}) +apply (tactic {* lift_tac @{context} @{thm pi_var_com} 1 *}) done lemma pi_app: "(pi\('x \ 'x) list) \ App (x\lam) (xa\lam) = App (pi \ x) (pi \ xa)" -apply (tactic {* lift_tac_lam @{context} @{thm pi_app_com} 1 *}) +apply (tactic {* lift_tac @{context} @{thm pi_app_com} 1 *}) done lemma pi_lam: "(pi\('x \ 'x) list) \ Lam (a\name) (x\lam) = Lam (pi \ a) (pi \ x)" -apply (tactic {* lift_tac_lam @{context} @{thm pi_lam_com} 1 *}) +apply (tactic {* lift_tac @{context} @{thm pi_lam_com} 1 *}) done lemma fv_var: "fv (Var (a\name)) = {a}" -apply (tactic {* lift_tac_lam @{context} @{thm rfv_var} 1 *}) +apply (tactic {* lift_tac @{context} @{thm rfv_var} 1 *}) done lemma fv_app: "fv (App (x\lam) (xa\lam)) = fv x \ fv xa" -apply (tactic {* lift_tac_lam @{context} @{thm rfv_app} 1 *}) +apply (tactic {* lift_tac @{context} @{thm rfv_app} 1 *}) done lemma fv_lam: "fv (Lam (a\name) (x\lam)) = fv x - {a}" -apply (tactic {* lift_tac_lam @{context} @{thm rfv_lam} 1 *}) +apply (tactic {* lift_tac @{context} @{thm rfv_lam} 1 *}) done lemma a1: "(a\name) = (b\name) \ Var a = Var b" -apply (tactic {* lift_tac_lam @{context} @{thm a1} 1 *}) +apply (tactic {* lift_tac @{context} @{thm a1} 1 *}) done lemma a2: "\(x\lam) = (xa\lam); (xb\lam) = (xc\lam)\ \ App x xb = App xa xc" -apply (tactic {* lift_tac_lam @{context} @{thm a2} 1 *}) +apply (tactic {* lift_tac @{context} @{thm a2} 1 *}) done lemma a3: "\(x\lam) = [(a\name, b\name)] \ (xa\lam); a \ fv (Lam b x)\ \ Lam a x = Lam b xa" -apply (tactic {* lift_tac_lam @{context} @{thm a3} 1 *}) +apply (tactic {* lift_tac @{context} @{thm a3} 1 *}) done lemma alpha_cases: "\a1 = a2; \a b. \a1 = Var a; a2 = Var b; a = b\ \ P; \x xa xb xc. \a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\ \ P; \x a b xa. \a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \ xa; a \ fv (Lam b x)\ \ P\ \ P" -apply (tactic {* lift_tac_lam @{context} @{thm alpha.cases} 1 *}) +apply (tactic {* lift_tac @{context} @{thm alpha.cases} 1 *}) done lemma alpha_induct: "\(qx\lam) = (qxa\lam); \(a\name) b\name. a = b \ (qxb\lam \ lam \ bool) (Var a) (Var b); @@ -224,16 +218,16 @@ \(x\lam) (a\name) (b\name) xa\lam. \x = [(a, b)] \ xa; qxb x ([(a, b)] \ xa); a \ fv (Lam b x)\ \ qxb (Lam a x) (Lam b xa)\ \ qxb qx qxa" -apply (tactic {* lift_tac_lam @{context} @{thm alpha.induct} 1 *}) +apply (tactic {* lift_tac @{context} @{thm alpha.induct} 1 *}) done lemma var_inject: "(Var a = Var b) = (a = b)" -apply (tactic {* lift_tac_lam @{context} @{thm rvar_inject} 1 *}) +apply (tactic {* lift_tac @{context} @{thm rvar_inject} 1 *}) done lemma lam_induct:" \\name. P (Var name); \lam1 lam2. \P lam1; P lam2\ \ P (App lam1 lam2); \name lam. P lam \ P (Lam name lam)\ \ P lam" -apply (tactic {* lift_tac_lam @{context} @{thm rlam.induct} 1 *}) +apply (tactic {* lift_tac @{context} @{thm rlam.induct} 1 *}) done lemma var_supp: