diff -r 934d4aed8497 -r 2510e6f7b11c src/CTT/ex/Equality.thy --- a/src/CTT/ex/Equality.thy Wed Oct 26 17:22:12 2022 +0100 +++ b/src/CTT/ex/Equality.thy Wed Oct 26 18:08:44 2022 +0100 @@ -6,63 +6,62 @@ section "Equality reasoning by rewriting" theory Equality -imports "../CTT" + 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 + 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 + apply (rule EqE) + apply (rule elim_rls, assumption) + apply rew + done -(*in the "rec" formulation of addition, 0+n=n *) +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 + apply (rule EqE) + apply (rule elim_rls, assumption) + apply rew + done -(*the harder version, n+0=n: recursive, uses induction hypothesis*) +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 + apply (rule EqE) + apply (rule elim_rls, assumption) + apply hyp_rew + done -(*Associativity of addition*) +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 + apply (NE a) + apply hyp_rew + done -(*Martin-Löf (1984) page 62: pairing is surjective*) +text \Martin-Löf (1984) page 62: pairing is surjective\ lemma "p : Sum(A,B) \ 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 + apply (rule EqE) + apply (rule elim_rls, assumption) + apply (tactic \DEPTH_SOLVE_1 (rew_tac \<^context> [])\) (*!!!!!!!*) + done lemma "\a : A; b : B\ \ (\<^bold>\u. split(u, \v w.)) ` = : \x:B. A" -apply rew -done + by rew -(*a contrived, complicated simplication, requires sum-elimination also*) +text \a contrived, complicated simplication, requires sum-elimination also\ lemma "(\<^bold>\f. \<^bold>\x. f`(f`x)) ` (\<^bold>\u. split(u, \v w.)) = \<^bold>\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 + 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