IntEx.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 28 Oct 2009 16:16:38 +0100
changeset 224 f9a25fe22037
parent 222 37b29231265b
child 228 268a727b0f10
permissions -rw-r--r--
Cleaning the unnecessary theorems in 'IntEx'.

theory IntEx
imports QuotMain
begin

fun
  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"
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

quotient_def (for my_int)
  ZERO::"my_int"
where
  "ZERO \<equiv> (0::nat, 0::nat)"

thm ZERO_def

quotient_def (for my_int)
  ONE::"my_int"
where
  "ONE \<equiv> (1::nat, 0::nat)"

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 (for my_int)
  PLUS::"my_int \<Rightarrow> my_int \<Rightarrow> my_int"
where
  "PLUS \<equiv> 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 (for my_int)
  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 (for my_int)
  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 (for my_int)
  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))"

lemma plus_sym_pre:
  shows "intrel (my_plus a b) (my_plus b a)"
  apply(cases a)
  apply(cases b)
  apply(auto)
  done

lemma ho_plus_rsp:
  "IntEx.intrel ===> IntEx.intrel ===> IntEx.intrel my_plus my_plus"
  by (simp)

ML {* val consts = [@{const_name "my_plus"}] *}
ML {* val rty = @{typ "nat \<times> nat"} *}
ML {* val qty = @{typ "my_int"} *}
ML {* val rel = @{term "intrel"} *}
ML {* val rel_eqv = (#equiv_thm o hd) (quotdata_lookup @{context}) *}
ML {* val rel_refl = @{thm EQUIV_REFL} OF [rel_eqv] *}
ML {* val quot = @{thm QUOTIENT_my_int} *}
ML {* val rsp_thms = @{thms ho_plus_rsp} @ @{thms ho_all_prs ho_ex_prs} *}
ML {* val trans2 = @{thm QUOT_TYPE_I_my_int.R_trans2} *}
ML {* val reps_same = @{thm QUOT_TYPE_I_my_int.REPS_same} *}
ML {* val t_defs = @{thms PLUS_def} *}

ML {*
fun lift_thm_my_int lthy t =
  lift_thm lthy consts rty qty rel rel_eqv rel_refl quot rsp_thms trans2 reps_same t_defs t
*}

ML {* lift_thm_my_int @{context} @{thm plus_sym_pre} *}

lemma plus_assoc_pre:
  shows "intrel (my_plus (my_plus i j) k) (my_plus i (my_plus j k))"
  apply (cases i)
  apply (cases j)
  apply (cases k)
  apply (simp)
  done

ML {* lift_thm_my_int @{context} @{thm plus_assoc_pre} *}








text {* Below is the construction site code used if things do not work *}


ML {* val abs = findabs rty (prop_of (atomize_thm @{thm plus_sym_pre})) *}
ML {* val simp_lam_prs_thms = map (make_simp_lam_prs_thm @{context} quot) abs *}
ML {* val t_defs_sym = add_lower_defs @{context} t_defs *}


ML {* val t_a = atomize_thm @{thm plus_sym_pre} *}
ML {* val t_r = regularize t_a rty rel rel_eqv @{context} *}
ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *}
ML {* val t_a = simp_allex_prs @{context} quot t_l *}
ML {* val t_d = repeat_eqsubst_thm @{context} t_defs_sym t_a *}
ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *}
ML {* ObjectLogic.rulify t_r *}