142 (* the auxiliary data for the quotient types *) |
142 (* the auxiliary data for the quotient types *) |
143 use "quotient_info.ML" |
143 use "quotient_info.ML" |
144 |
144 |
145 declare [[map * = (prod_fun, prod_rel)]] |
145 declare [[map * = (prod_fun, prod_rel)]] |
146 declare [[map "fun" = (fun_map, fun_rel)]] |
146 declare [[map "fun" = (fun_map, fun_rel)]] |
147 (* FIXME: This should throw an exception: |
147 |
148 declare [[map "option" = (bla, blu)]] |
148 (* FIXME: This should throw an exception: *) |
149 *) |
149 (* declare [[map "option" = (bla, blu)]] *) |
|
150 |
150 |
151 |
151 (* identity quotient is not here as it has to be applied first *) |
152 (* identity quotient is not here as it has to be applied first *) |
152 lemmas [quotient_thm] = |
153 lemmas [quotient_thm] = |
153 fun_quotient prod_quotient |
154 fun_quotient prod_quotient |
154 |
155 |
292 | (Abs (_, T, t), Abs (_, T', t')) => matches_typ (T, T') andalso matches_term (t, t') |
293 | (Abs (_, T, t), Abs (_, T', t')) => matches_typ (T, T') andalso matches_term (t, t') |
293 | (t $ s, t' $ s') => matches_term (t, t') andalso matches_term (s, s') |
294 | (t $ s, t' $ s') => matches_term (t, t') andalso matches_term (s, s') |
294 | _ => false |
295 | _ => false |
295 *} |
296 *} |
296 |
297 |
297 section {* Regularization *} |
298 section {* Computation of the Regularize Goal *} |
298 |
299 |
299 (* |
300 (* |
300 Regularizing an rtrm means: |
301 Regularizing an rtrm means: |
301 - quantifiers over a type that needs lifting are replaced by |
302 - quantifiers over a type that needs lifting are replaced by |
302 bounded quantifiers, for example: |
303 bounded quantifiers, for example: |
530 |
533 |
531 lemma eq_imp_rel: |
534 lemma eq_imp_rel: |
532 shows "equivp R \<Longrightarrow> a = b \<longrightarrow> R a b" |
535 shows "equivp R \<Longrightarrow> a = b \<longrightarrow> R a b" |
533 by (simp add: equivp_reflp) |
536 by (simp add: equivp_reflp) |
534 |
537 |
535 (* FIXME/TODO: How does regularizing work? *) |
538 |
536 (* FIXME/TODO: needs to be adapted |
539 (* Regularize Tactic *) |
537 |
540 |
538 To prove that the raw theorem implies the regularised one, |
541 (* 0. preliminary simplification step according to *) |
539 we try in order: |
542 thm ball_reg_eqv bex_reg_eqv babs_reg_eqv (* the latter of no use *) |
540 |
543 ball_reg_eqv_range bex_reg_eqv_range |
541 - Reflexivity of the relation |
544 (* 1. eliminating simple Ball/Bex instances*) |
542 - Assumption |
545 thm ball_reg_right bex_reg_left |
543 - Elimnating quantifiers on both sides of toplevel implication |
546 (* 2. monos *) |
544 - Simplifying implications on both sides of toplevel implication |
547 (* 3. commutation rules for ball and bex *) |
545 - Ball (Respects ?E) ?P = All ?P |
548 thm ball_all_comm bex_ex_comm |
546 - (\<And>x. ?R x \<Longrightarrow> ?P x \<longrightarrow> ?Q x) \<Longrightarrow> All ?P \<longrightarrow> Ball ?R ?Q |
549 (* 4. then rel-equality (which need to be instantiated to avoid loops *) |
547 |
550 thm eq_imp_rel |
548 *) |
551 (* 5. then simplification like 0 *) |
|
552 (* finally jump back to 1 *) |
|
553 |
|
554 |
549 ML {* |
555 ML {* |
550 fun regularize_tac ctxt = |
556 fun regularize_tac ctxt = |
551 let |
557 let |
552 val thy = ProofContext.theory_of ctxt |
558 val thy = ProofContext.theory_of ctxt |
553 val pat_ball = @{term "Ball (Respects (R1 ===> R2)) P"} |
559 val pat_ball = @{term "Ball (Respects (R1 ===> R2)) P"} |
556 val simpset = (mk_minimal_ss ctxt) |
562 val simpset = (mk_minimal_ss ctxt) |
557 addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv} |
563 addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv} |
558 addsimprocs [simproc] addSolver equiv_solver |
564 addsimprocs [simproc] addSolver equiv_solver |
559 (* TODO: Make sure that there are no list_rel, pair_rel etc involved *) |
565 (* TODO: Make sure that there are no list_rel, pair_rel etc involved *) |
560 (* can this cause loops in equiv_tac ? *) |
566 (* can this cause loops in equiv_tac ? *) |
561 val eq_eqvs = map (fn x => @{thm eq_imp_rel} OF [x]) (equiv_rules_get ctxt) |
567 val eq_eqvs = map (OF1 @{thm eq_imp_rel}) (equiv_rules_get ctxt) |
562 in |
568 in |
563 simp_tac simpset THEN' |
569 simp_tac simpset THEN' |
564 REPEAT_ALL_NEW (CHANGED o FIRST' [ |
570 REPEAT_ALL_NEW (CHANGED o FIRST' [ |
565 rtac @{thm ball_reg_right}, |
571 resolve_tac @{thms ball_reg_right bex_reg_left}, |
566 rtac @{thm bex_reg_left}, |
|
567 resolve_tac (Inductive.get_monos ctxt), |
572 resolve_tac (Inductive.get_monos ctxt), |
568 rtac @{thm ball_all_comm}, |
573 resolve_tac @{thms ball_all_comm bex_ex_comm}, |
569 rtac @{thm bex_ex_comm}, |
|
570 resolve_tac eq_eqvs, |
574 resolve_tac eq_eqvs, |
571 (* should be equivalent to the above, but causes loops: *) |
|
572 (* rtac @{thm eq_imp_rel} THEN' SOLVES' (equiv_tac ctxt), *) |
|
573 (* the culprit is aslread rtac @{thm eq_imp_rel} *) |
|
574 simp_tac simpset]) |
575 simp_tac simpset]) |
575 end |
576 end |
576 *} |
577 *} |
577 |
578 |
578 section {* Injections of rep and abses *} |
579 section {* Calculation of the Injected Goal *} |
579 |
580 |
580 (* |
581 (* |
581 Injecting repabs means: |
582 Injecting repabs means: |
582 |
583 |
583 For abstractions: |
584 For abstractions: |
1041 More_Conv.top_conv lambda_prs_simple_conv |
1042 More_Conv.top_conv lambda_prs_simple_conv |
1042 |
1043 |
1043 fun lambda_prs_tac ctxt = CONVERSION (lambda_prs_conv ctxt) |
1044 fun lambda_prs_tac ctxt = CONVERSION (lambda_prs_conv ctxt) |
1044 *} |
1045 *} |
1045 |
1046 |
1046 (* |
1047 (* 1. conversion (is not a pattern) *) |
1047 Cleaning the theorem consists of three rewriting steps. |
1048 thm lambda_prs |
1048 The first two need to be done before fun_map is unfolded |
1049 (* 2. folding of definitions: (rep ---> abs) oldConst == newconst *) |
1049 |
1050 (* and simplification with *) |
1050 1) lambda_prs: |
1051 thm all_prs ex_prs |
1051 (Rep1 ---> Abs2) (\<lambda>x. Rep2 (f (Abs1 x))) ----> f |
1052 (* 3. simplification with *) |
1052 |
1053 thm fun_map.simps id_simps Quotient_abs_rep Quotient_rel_rep |
1053 Implemented as conversion since it is not a pattern. |
1054 (* 4. Test for refl *) |
1054 |
|
1055 2) all_prs (the same for exists): |
|
1056 Ball (Respects R) ((abs ---> id) f) ----> All f |
|
1057 |
|
1058 Rewriting with definitions from the argument defs |
|
1059 (rep ---> abs) oldConst ----> newconst |
|
1060 |
|
1061 3) Quotient_rel_rep: |
|
1062 Rel (Rep x) (Rep y) ----> x = y |
|
1063 |
|
1064 Quotient_abs_rep: |
|
1065 Abs (Rep x) ----> x |
|
1066 |
|
1067 id_simps; fun_map.simps |
|
1068 *) |
|
1069 |
1055 |
1070 ML {* |
1056 ML {* |
1071 fun clean_tac lthy = |
1057 fun clean_tac lthy = |
1072 let |
1058 let |
1073 val thy = ProofContext.theory_of lthy; |
1059 val thy = ProofContext.theory_of lthy; |
1154 val rtrm' = HOLogic.dest_Trueprop rtrm |
1140 val rtrm' = HOLogic.dest_Trueprop rtrm |
1155 val qtrm' = HOLogic.dest_Trueprop qtrm |
1141 val qtrm' = HOLogic.dest_Trueprop qtrm |
1156 val reg_goal = |
1142 val reg_goal = |
1157 Syntax.check_term ctxt (regularize_trm ctxt rtrm' qtrm') |
1143 Syntax.check_term ctxt (regularize_trm ctxt rtrm' qtrm') |
1158 handle (LIFT_MATCH s) => lift_match_error ctxt s rtrm qtrm |
1144 handle (LIFT_MATCH s) => lift_match_error ctxt s rtrm qtrm |
1159 val _ = warning "Regularization done." |
|
1160 val inj_goal = |
1145 val inj_goal = |
1161 Syntax.check_term ctxt (inj_repabs_trm ctxt (reg_goal, qtrm')) |
1146 Syntax.check_term ctxt (inj_repabs_trm ctxt (reg_goal, qtrm')) |
1162 handle (LIFT_MATCH s) => lift_match_error ctxt s rtrm qtrm |
1147 handle (LIFT_MATCH s) => lift_match_error ctxt s rtrm qtrm |
1163 val _ = warning "RepAbs Injection done." |
|
1164 in |
1148 in |
1165 Drule.instantiate' [] |
1149 Drule.instantiate' [] |
1166 [SOME (cterm_of thy rtrm'), |
1150 [SOME (cterm_of thy rtrm'), |
1167 SOME (cterm_of thy reg_goal), |
1151 SOME (cterm_of thy reg_goal), |
1168 SOME (cterm_of thy inj_goal)] @{thm lifting_procedure} |
1152 SOME (cterm_of thy inj_goal)] @{thm lifting_procedure} |
1169 end |
1153 end |
1170 *} |
1154 *} |
1171 |
1155 |
1172 (* Left for debugging *) |
1156 ML {* |
1173 ML {* |
1157 (* leaves three subgoales to be proved *) |
1174 fun procedure_tac ctxt rthm = |
1158 fun procedure_tac ctxt rthm = |
1175 ObjectLogic.full_atomize_tac |
|
1176 THEN' gen_frees_tac ctxt |
|
1177 THEN' CSUBGOAL (fn (gl, i) => |
|
1178 let |
|
1179 val rthm' = atomize_thm rthm |
|
1180 val rule = procedure_inst ctxt (prop_of rthm') (term_of gl) |
|
1181 val thm = Drule.instantiate' [] [SOME (snd (Thm.dest_comb gl))] @{thm QUOT_TRUE_i} |
|
1182 in |
|
1183 (rtac rule THEN' RANGE [rtac rthm', (fn _ => all_tac), rtac thm]) i |
|
1184 end) |
|
1185 *} |
|
1186 |
|
1187 ML {* |
|
1188 fun lift_tac ctxt rthm = |
|
1189 ObjectLogic.full_atomize_tac |
1159 ObjectLogic.full_atomize_tac |
1190 THEN' gen_frees_tac ctxt |
1160 THEN' gen_frees_tac ctxt |
1191 THEN' CSUBGOAL (fn (goal, i) => |
1161 THEN' CSUBGOAL (fn (goal, i) => |
1192 let |
1162 let |
1193 val rthm' = atomize_thm rthm |
1163 val rthm' = atomize_thm rthm |
1194 val rule = procedure_inst ctxt (prop_of rthm') (term_of goal) |
1164 val rule = procedure_inst ctxt (prop_of rthm') (term_of goal) |
1195 val thm = Drule.instantiate' [] [SOME (snd (Thm.dest_comb goal))] @{thm QUOT_TRUE_i} |
1165 val bare_goal = snd (Thm.dest_comb goal) |
|
1166 val quot_weak = Drule.instantiate' [] [SOME bare_goal] @{thm QUOT_TRUE_i} |
1196 in |
1167 in |
1197 (rtac rule THEN' |
1168 (rtac rule THEN' RANGE [rtac rthm', K all_tac, rtac quot_weak]) i |
1198 RANGE [rtac rthm', |
|
1199 regularize_tac ctxt, |
|
1200 rtac thm THEN' all_inj_repabs_tac ctxt, |
|
1201 clean_tac ctxt]) i |
|
1202 end) |
1169 end) |
1203 *} |
1170 *} |
1204 |
1171 |
1205 end |
1172 ML {* |
1206 |
1173 (* automated tactics *) |
|
1174 fun lift_tac ctxt rthm = |
|
1175 procedure_tac ctxt rthm |
|
1176 THEN' RANGE [regularize_tac ctxt, |
|
1177 all_inj_repabs_tac ctxt, |
|
1178 clean_tac ctxt] |
|
1179 *} |
|
1180 |
|
1181 end |
|
1182 |