theory IntEx
imports QuotMain
begin
fun
intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50)
where
"intrel (x, y) (u, v) = (x + v = u + y)"
quotient my_int = "nat \<times> nat" / intrel
apply(unfold EQUIV_def)
apply(auto simp add: mem_def expand_fun_eq)
done
thm my_int_equiv
print_theorems
print_quotients
quotient_def
ZERO::"my_int"
where
"ZERO \<equiv> (0::nat, 0::nat)"
ML {* print_qconstinfo @{context} *}
term ZERO
thm ZERO_def
ML {* prop_of @{thm ZERO_def} *}
ML {* separate *}
quotient_def
ONE::"my_int"
where
"ONE \<equiv> (1::nat, 0::nat)"
ML {* print_qconstinfo @{context} *}
term ONE
thm ONE_def
fun
my_plus :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
where
"my_plus (x, y) (u, v) = (x + u, y + v)"
quotient_def
PLUS::"my_int \<Rightarrow> my_int \<Rightarrow> my_int"
where
"PLUS \<equiv> my_plus"
term my_plus
term PLUS
thm PLUS_def
fun
my_neg :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
where
"my_neg (x, y) = (y, x)"
quotient_def
NEG::"my_int \<Rightarrow> my_int"
where
"NEG \<equiv> my_neg"
term NEG
thm NEG_def
definition
MINUS :: "my_int \<Rightarrow> my_int \<Rightarrow> my_int"
where
"MINUS z w = PLUS z (NEG w)"
fun
my_mult :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
where
"my_mult (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
quotient_def
MULT::"my_int \<Rightarrow> my_int \<Rightarrow> my_int"
where
"MULT \<equiv> my_mult"
term MULT
thm MULT_def
(* NOT SURE WETHER THIS DEFINITION IS CORRECT *)
fun
my_le :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"
where
"my_le (x, y) (u, v) = (x+v \<le> u+y)"
quotient_def
LE :: "my_int \<Rightarrow> my_int \<Rightarrow> bool"
where
"LE \<equiv> my_le"
term LE
thm LE_def
definition
LESS :: "my_int \<Rightarrow> my_int \<Rightarrow> bool"
where
"LESS z w = (LE z w \<and> z \<noteq> w)"
term LESS
thm LESS_def
definition
ABS :: "my_int \<Rightarrow> my_int"
where
"ABS i = (if (LESS i ZERO) then (NEG i) else i)"
definition
SIGN :: "my_int \<Rightarrow> my_int"
where
"SIGN i = (if i = ZERO then ZERO else if (LESS ZERO i) then ONE else (NEG ONE))"
ML {* print_qconstinfo @{context} *}
lemma plus_sym_pre:
shows "my_plus a b \<approx> my_plus b a"
apply(cases a)
apply(cases b)
apply(auto)
done
lemma ho_plus_rsp:
"(intrel ===> intrel ===> intrel) my_plus my_plus"
by (simp)
ML {* val qty = @{typ "my_int"} *}
ML {* val ty_name = "my_int" *}
ML {* val rsp_thms = @{thms ho_plus_rsp} @ @{thms ho_all_prs ho_ex_prs} *}
ML {* val defs = @{thms PLUS_def} *}
ML {*
fun lift_thm_my_int lthy t =
lift_thm lthy qty ty_name rsp_thms defs t
*}
print_quotients
ML {* quotdata_lookup @{context} "IntEx.my_int" *}
ML {* quotdata_lookup @{context} "my_int" *}
ML {*
val test = lift_thm_my_int @{context} @{thm plus_sym_pre}
*}
ML {*
val aqtrm = (prop_of (atomize_thm test))
val artrm = (prop_of (atomize_thm @{thm plus_sym_pre}))
val reg_artm = mk_REGULARIZE_goal @{context} artrm aqtrm
*}
prove {* reg_artm *}
apply(tactic {* ObjectLogic.full_atomize_tac 1 *})
apply(tactic {* REGULARIZE_tac' @{context} @{thm "TrueI"} @{thm my_int_equiv} *})
done
lemma plus_assoc_pre:
shows "my_plus (my_plus i j) k \<approx> my_plus i (my_plus j k)"
apply (cases i)
apply (cases j)
apply (cases k)
apply (simp)
done
ML {* val test2 = lift_thm_my_int @{context} @{thm plus_assoc_pre} *}
ML {*
val arthm = atomize_thm @{thm plus_assoc_pre}
val aqthm = atomize_thm test2
val aqtrm = prop_of aqthm
val artrm = prop_of arthm
*}
prove testtest: {* mk_REGULARIZE_goal @{context} artrm aqtrm *}
apply(tactic {* ObjectLogic.full_atomize_tac 1 *})
apply(tactic {* REGULARIZE_tac' @{context} @{thm "TrueI"} @{thm my_int_equiv} *})
done
ML {* @{thm testtest} OF [arthm] *}
ML {* val reg_atrm = prop_of (@{thm testtest} OF [arthm]) *}
(*
does not work.
ML {*
fun r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms =
(REPEAT1 o FIRST1)
[(K (print_tac "start")) THEN' (K no_tac),
DT ctxt "1" (rtac trans_thm),
DT ctxt "2" (LAMBDA_RES_TAC ctxt),
DT ctxt "3" (ball_rsp_tac ctxt),
DT ctxt "4" (bex_rsp_tac ctxt),
DT ctxt "5" (FIRST' (map rtac rsp_thms)),
DT ctxt "6" (instantiate_tac @{thm REP_ABS_RSP(1)} ctxt THEN' (RANGE [quotient_tac quot_thm])),
DT ctxt "7" (rtac refl),
DT ctxt "8" (APPLY_RSP_TAC rty ctxt THEN' (RANGE [quotient_tac quot_thm, quotient_tac quot_thm])),
DT ctxt "9" (Cong_Tac.cong_tac @{thm cong}),
DT ctxt "A" (rtac @{thm ext}),
DT ctxt "B" (rtac reflex_thm),
DT ctxt "C" (atac),
DT ctxt "D" (SOLVES' (simp_tac (HOL_ss addsimps @{thms FUN_REL.simps}))),
DT ctxt "E" (WEAK_LAMBDA_RES_TAC ctxt),
DT ctxt "F" (CHANGED' (asm_full_simp_tac (HOL_ss addsimps @{thms FUN_REL_EQ})))]
*}
*)
ML {*
val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} @{typ "my_int"}
val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "my_int"
val consts = lookup_quot_consts defs
*}
ML {* cprem_of *}
ML {*
mk_inj_REPABS_goal @{context} (reg_atrm, aqtrm)
|> Syntax.check_term @{context}
*}
ML {* val my_goal = cterm_of @{theory} (mk_inj_REPABS_goal @{context} (reg_atrm, aqtrm)) *}
ML {* val yr_goal = cterm_of @{theory} (build_repabs_goal @{context} (@{thm testtest} OF [arthm]) consts rty qty) *}
prove {* mk_inj_REPABS_goal @{context} (reg_atrm, aqtrm) *}
apply(tactic {* ObjectLogic.full_atomize_tac 1 *})
apply(tactic {* REPEAT_ALL_NEW (r_mk_comb_tac @{context} rty quot rel_refl trans2 rsp_thms) 1 *})
done
ML {*
inj_REPABS @{context} (reg_atrm, aqtrm)
|> Syntax.string_of_term @{context}
|> writeln
*}
lemma ho_tst: "foldl my_plus x [] = x"
apply simp
done
text {* Below is the construction site code used if things do not work *}
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "my_int" *}
(* ML {* val consts = [@{const_name my_plus}] *}*)
ML {* val consts = lookup_quot_consts defs *}
ML {* val t_a = atomize_thm @{thm ho_tst} *}
(*
prove t_r: {* build_regularize_goal t_a rty rel @{context} *}
ML_prf {* fun tac ctxt =
(ObjectLogic.full_atomize_tac) THEN'
REPEAT_ALL_NEW (FIRST' [
rtac rel_refl,
atac,
rtac @{thm universal_twice},
(rtac @{thm impI} THEN' atac),
(*rtac @{thm equality_twice},*)
EqSubst.eqsubst_tac ctxt [0]
[(@{thm equiv_res_forall} OF [rel_eqv]),
(@{thm equiv_res_exists} OF [rel_eqv])],
(rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
(rtac @{thm RIGHT_RES_FORALL_REGULAR})
]);*}
apply (atomize(full))
apply (tactic {* tac @{context} 1 *})
apply (auto)
done
ML {* val t_r = @{thm t_r} OF [t_a] *}*)
ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
ML {*
val rt = build_repabs_term @{context} t_r consts rty qty
val rg = Logic.mk_equals ((Thm.prop_of t_r), rt);
*}
lemma foldl_rsp:
"((IntEx.intrel ===> IntEx.intrel ===> IntEx.intrel) ===>
IntEx.intrel ===> op = ===> IntEx.intrel)
foldl foldl"
apply (simp only:FUN_REL.simps)
apply (rule allI)
apply (rule allI)
apply (rule impI)
apply (rule allI)
apply (rule allI)
apply (rule impI)
apply (rule allI)
apply (rule allI)
apply (rule impI)
apply (simp only:)
apply (rule list.induct)
apply (simp)
apply (simp only: foldl.simps)
sorry
ML {* val rsp_thms = @{thm foldl_rsp} :: rsp_thms *}
prove t_t: rg
apply(atomize(full))
ML_prf {* fun r_mk_comb_tac_int lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms *}
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_int @{context}) 1 *})
done
ML {* val t_t = @{thm Pure.equal_elim_rule1} OF [@{thm t_t},t_r] *}
ML {* val abs = findabs rty (prop_of t_a) *}
ML {* val aps = findaps rty (prop_of t_a); *}
ML {* val simp_lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
(*ML {* val t_t = repabs @{context} @{thm t_r} consts rty qty quot rel_refl trans2 rsp_thms *}*)
ML findallex
ML {* val (alls, exs) = findallex @{context} rty qty (prop_of t_a) *}
ML {* val allthms = map (make_allex_prs_thm @{context} quot @{thm FORALL_PRS}) alls *}
ML {* val t_a = MetaSimplifier.rewrite_rule (allthms) t_t *}
ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_a *}
ML {* val defs_sym = flat (map (add_lower_defs @{context}) defs) *}
ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_l *}
ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *}
ML {* ObjectLogic.rulify t_r *}