# HG changeset patch # User Christian Urban # Date 1259681262 -3600 # Node ID 6d077eac6ad77ce65ec0af636d2302779ed8fcfb # Parent 10d75457792f4a19905fff8a60d4b06b3b5bd333 QUOT_TRUE joke diff -r 10d75457792f -r 6d077eac6ad7 FSet.thy --- a/FSet.thy Tue Dec 01 14:08:56 2009 +0100 +++ b/FSet.thy Tue Dec 01 16:27:42 2009 +0100 @@ -12,6 +12,12 @@ | "xs \ ys \ a#xs \ a#ys" | "\xs1 \ xs2; xs2 \ xs3\ \ xs1 \ xs3" +lemma + "(list_eq ===> (list_eq ===> (op =))) (list_eq) (list_eq)" +apply(simp add: FUN_REL.simps) +apply(rule allI | rule impI)+ +sorry + lemma list_eq_refl: shows "xs \ xs" by (induct xs) (auto intro: list_eq.intros) diff -r 10d75457792f -r 6d077eac6ad7 QuotMain.thy --- a/QuotMain.thy Tue Dec 01 14:08:56 2009 +0100 +++ b/QuotMain.thy Tue Dec 01 16:27:42 2009 +0100 @@ -836,12 +836,24 @@ *) +definition + "QUOT_TRUE x \ True" + +lemma quot_true_dests: + shows QT_all: "QUOT_TRUE (All P) \ QUOT_TRUE P" + and QT_appL: "QUOT_TRUE (A B) \ QUOT_TRUE A" + and QT_appR: "QUOT_TRUE (A B) \ QUOT_TRUE B" + and QT_lam: "QUOT_TRUE (\x. P x) \ (\x. QUOT_TRUE (P x))" +apply(simp_all add: QUOT_TRUE_def) +done + + ML {* fun inj_repabs_tac ctxt rty quot_thms rel_refl trans2 = (FIRST' [ (* "cong" rule of the of the relation / transitivity*) (* (op =) (R a b) (R c d) ----> \R a c; R b d\ *) - NDT ctxt "1" (resolve_tac trans2), + DT ctxt "1" (resolve_tac trans2), (* (R1 ===> R2) (\x\) (\y\) ----> \R1 x y\ \ R2 (\x) (\y) *) NDT ctxt "2" (lambda_res_tac ctxt), @@ -887,12 +899,6 @@ (* resolving with R x y assumptions *) NDT ctxt "E" (atac), - (* seems not necessay:: NDT ctxt "F" (SOLVES' (simp_tac (HOL_ss addsimps @{thms FUN_REL.simps}))),*) - - (* (R1 ===> R2) (\) (\y\) ----> \R1 x y\ \ R2 (\x) (\y) *) - (* (R1 ===> R2) (\x\) (\) ----> \R1 x y\ \ R2 (\x) (\y) *) - (*NDT ctxt "G" (weak_lambda_res_tac ctxt),*) - (* (op =) ===> (op =) \ (op =), needed in order to apply respectfulness theorems *) (* global simplification *) NDT ctxt "H" (CHANGED o (asm_full_simp_tac ((Simplifier.context ctxt empty_ss) addsimps @{thms eq_reflection[OF FUN_REL_EQ]})))]) @@ -929,7 +935,7 @@ if incr_boundvars i u aconv t then Bound i else (case t of t1 $ t2 => mk_abs i t1 $ mk_abs i t2 - | Abs (s, T, t') => Abs (s, T, mk_abs (i+1) t') + | Abs (s, T, t') => Abs (s, T, mk_abs (i + 1) t') | Bound j => if i = j then error "make_inst" else t | _ => t); in (f, Abs ("x", T, mk_abs 0 g)) end; @@ -986,21 +992,30 @@ The first two need to be done before fun_map is unfolded - LAMBDA_PRS: - (Rep1 ---> Abs2) (\x. Rep2 (f (Abs1 x))) \ f + (Rep1 ---> Abs2) (\x. Rep2 (f (Abs1 x))) ----> f + - FORALL_PRS (and the same for exists: EXISTS_PRS) - \x\Respects R. (abs ---> id) ?f \ \x. ?f + \x\Respects R. (abs ---> id) f ----> \x. f + - Rewriting with definitions from the argument defs - NewConst \ (rep ---> abs) oldConst + NewConst ----> (rep ---> abs) oldConst + - QUOTIENT_REL_REP: - Rel (Rep x) (Rep y) \ x = y + Rel (Rep x) (Rep y) ----> x = y + - ABS_REP - Abs (Rep x) \ x + Abs (Rep x) ----> x + - id_simps; fun_map.simps The first one is implemented as a conversion (fast). The second one is an EqSubst (slow). The rest are a simp_tac and are fast. *) + +thm all_prs ex_prs + + ML {* fun clean_tac lthy quot defs = let