IntEx.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 25 Nov 2009 14:25:29 +0100
changeset 386 4fcbbb5b3b58
parent 379 57bde65f6eb2
child 389 d67240113f68
permissions -rw-r--r--
Moved exception handling to QuotMain and cleaned FSet.

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 {* 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 = lookup_quot_consts defs *}

ML {*
fun lift_tac_fset lthy t =
  lift_tac lthy t rel_eqv rel_refl rty quot trans2 rsp_thms reps_same defs
*}

lemma "PLUS a b = PLUS b a"
by (tactic {* lift_tac_fset @{context} @{thm plus_sym_pre} 1 *})

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

lemma plus_assoc: "PLUS (PLUS x xa) xb = PLUS x (PLUS xa xb)"
by (tactic {* lift_tac_fset @{context} @{thm plus_assoc_pre} 1 *})


lemma ho_tst: "foldl my_plus x [] = x"
apply simp
done

lemma "foldl PLUS x [] = x"
apply (tactic {* lift_tac_fset @{context} @{thm ho_tst} 1 *})
prefer 3
apply(tactic {* clean_tac @{context} quot @{thms PLUS_def} reps_same 1 *})
sorry

(*
  FIXME: All below is your construction code; mostly commented out as it
  does not work.
*)

ML {*
  mk_REGULARIZE_goal @{context} 
    @{term "\<forall>i j k. my_plus (my_plus i j) k \<approx> my_plus i (my_plus j k)"}
    @{term "\<forall>i j k. PLUS (PLUS i j) k = PLUS i (PLUS j k)"}
  |> Syntax.string_of_term @{context}
  |> writeln
*}

lemma "PLUS (PLUS i j) k = PLUS i (PLUS j k)"
apply(tactic {* procedure_tac @{thm plus_assoc_pre} @{context} 1 *})
apply(tactic {* regularize_tac @{context} rel_eqv rel_refl 1 *})
apply(tactic {* REPEAT_ALL_NEW (r_mk_comb_tac @{context} rty quot rel_refl trans2 rsp_thms) 1 *})
apply(tactic {* clean_tac @{context} quot @{thms PLUS_def} reps_same 1 *})
done 


(*

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 {* 
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
*}

*)