IntEx.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Mon, 26 Oct 2009 19:35:30 +0100
changeset 196 9163c0f9830d
parent 195 72165b83b9d6
child 197 c0f2db9a243b
permissions -rw-r--r--
Simplifying Int.

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

print_quotients

typ my_int

local_setup {*
  make_const_def @{binding "ZERO"} @{term "(0::nat, 0::nat)"} NoSyn @{typ "nat \<times> nat"} @{typ "my_int"} #> snd
*}

term ZERO
thm ZERO_def

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

local_setup {*
  make_const_def @{binding ONE} @{term "(1::nat, 0::nat)"} NoSyn @{typ "nat \<times> nat"} @{typ "my_int"} #> snd
*}

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)"

local_setup {*
  make_const_def @{binding PLUS} @{term "my_plus"} NoSyn @{typ "nat \<times> nat"} @{typ "my_int"} #> snd
*}

term PLUS
thm PLUS_def

fun
  my_neg :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
where
  "my_neg (x, y) = (y, x)"

local_setup {*
  make_const_def @{binding NEG} @{term "my_neg"} NoSyn @{typ "nat \<times> nat"} @{typ "my_int"} #> snd
*}

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)"

local_setup {*
  make_const_def @{binding MULT} @{term "my_mult"} NoSyn @{typ "nat \<times> nat"} @{typ "my_int"} #> snd
*}

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)"

local_setup {*
  make_const_def @{binding LE} @{term "my_le"} NoSyn @{typ "nat \<times> nat"} @{typ "my_int"} #> snd
*}

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)"
  sorry

lemma equiv_intrel: "EQUIV intrel"
  sorry

lemma intrel_refl: "intrel a a"
  sorry

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 = @{thm equiv_intrel} *}
ML {* val rel_refl = @{thm intrel_refl} *}
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 t_a = atomize_thm @{thm plus_sym_pre} *}
ML {* val t_r = regularize t_a rty rel rel_eqv @{context} *}

ML {* val (g, thm, othm) =
  Toplevel.program (fn () =>
  repabs_eq @{context} t_r consts rty qty
   quot rel_refl trans2
   rsp_thms
  )
*}
ML {*
  val t_t2 =
  Toplevel.program (fn () =>
    repabs_eq2 @{context} (g, thm, othm)
  )
*}

ML {*
fun make_simp_lam_prs_thm lthy quot_thm typ =
  let
    val (_, [lty, rty]) = dest_Type typ;
    val thy = ProofContext.theory_of lthy;
    val (lcty, rcty) = (ctyp_of thy lty, ctyp_of thy rty)
    val inst = [SOME lcty, NONE, SOME rcty];
    val lpi = Drule.instantiate' inst [] @{thm LAMBDA_PRS};
    val tac =
      (compose_tac (false, @{thm LAMBDA_PRS}, 2)) THEN_ALL_NEW
      (quotient_tac quot_thm);
    val t = Goal.prove lthy [] [] (concl_of lpi) (fn _ => tac 1);
    val ts = @{thm HOL.sym} OF [t]
  in
    MetaSimplifier.rewrite_rule [@{thm eq_reflection} OF @{thms id_apply}] ts
  end
*}


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 {*
  fun simp_lam_prs lthy thm =
    simp_lam_prs lthy (eqsubst_thm lthy simp_lam_prs_thms thm)
    handle _ => thm
*}
ML {* t_t2 *}
ML {* val t_l = simp_lam_prs @{context} t_t2 *}

ML {*
  fun simp_allex_prs lthy quot thm =
    let
      val rwf = @{thm FORALL_PRS} OF [quot];
      val rwfs = @{thm "HOL.sym"} OF [rwf];
      val rwe = @{thm EXISTS_PRS} OF [quot];
      val rwes = @{thm "HOL.sym"} OF [rwe]
    in
      (simp_allex_prs lthy quot (eqsubst_thm lthy [rwfs, rwes] thm))
    end
    handle _ => thm
*}

ML {* val t_a = simp_allex_prs @{context} quot t_l *}

ML {* val t_defs = @{thms PLUS_def} *}
ML {* val t_defs_sym = add_lower_defs @{context} t_defs *}
ML {* val t_d = MetaSimplifier.rewrite_rule t_defs_sym t_a *}
ML {* val t_r = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_my_int.REPS_same} t_d *}
ML {* ObjectLogic.rulify t_r *}

lemma 
  fixes i j k::"my_int"
  shows "(PLUS (PLUS i j) k) = (PLUS i (PLUS j k))"
  apply(unfold PLUS_def)
  apply(simp add: expand_fun_eq)
  sorry