Theory Equality

(*  Title:      CTT/ex/Equality.thy
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1991  University of Cambridge
*)

section "Equality reasoning by rewriting"

theory Equality
  imports "../CTT"
begin

lemma split_eq: "p : Sum(A,B)  split(p,pair) = p : Sum(A,B)"
  apply (rule EqE)
  apply (rule elim_rls, assumption)
  apply rew
  done

lemma when_eq: "A type; B type; p : A+B  when(p,inl,inr) = p : A + B"
  apply (rule EqE)
  apply (rule elim_rls, assumption)
   apply rew
  done

text ‹in the "rec" formulation of addition, $0+n=n$›
lemma "p:N  rec(p,0, λy z. succ(y)) = p : N"
  apply (rule EqE)
  apply (rule elim_rls, assumption)
   apply rew
  done

text ‹the harder version, $n+0=n$: recursive, uses induction hypothesis›
lemma "p:N  rec(p,0, λy z. succ(z)) = p : N"
  apply (rule EqE)
  apply (rule elim_rls, assumption)
   apply hyp_rew
  done

text ‹Associativity of addition›
lemma "a:N; b:N; c:N
   rec(rec(a, b, λx y. succ(y)), c, λx y. succ(y)) =
    rec(a, rec(b, c, λx y. succ(y)), λx y. succ(y)) : N"
  apply (NE a)
    apply hyp_rew
  done

text ‹Martin-Löf (1984) page 62: pairing is surjective›
lemma "p : Sum(A,B)  <split(p,λx y. x), split(p,λx y. y)> = p : Sum(A,B)"
  apply (rule EqE)
  apply (rule elim_rls, assumption)
  apply (tactic DEPTH_SOLVE_1 (rew_tac context [])) (*!!!!!!!*)
  done

lemma "a : A; b : B  (λu. split(u, λv w.<w,v>)) ` <a,b> = <b,a> : x:B. A"
  by rew

text ‹a contrived, complicated simplication, requires sum-elimination also›
lemma "(λf. λx. f`(f`x)) ` (λu. split(u, λv w.<w,v>)) =
      λx. x  :  x:(y:N. N). (y:N. N)"
  apply (rule reduction_rls)
    apply (rule_tac [3] intrL_rls)
     apply (rule_tac [4] EqE)
     apply (erule_tac [4] SumE)
    (*order of unifiers is essential here*)
     apply rew
  done

end