# HG changeset patch # User wenzelm # Date 809958386 -7200 # Node ID 706cfddca75c59c5b0116133376a713dd6eac583 # Parent 934183dfc7868a86e615f941dd9938398f58f45f the NatClass demo of the axclass tutorial; diff -r 934183dfc786 -r 706cfddca75c src/FOL/ex/NatClass.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/NatClass.ML Fri Sep 01 14:26:26 1995 +0200 @@ -0,0 +1,62 @@ +(* Title: FOL/ex/NatClass.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +This is Nat.ML trivially modified to make it work with NatClass.thy. +*) + +open NatClass; + +goal NatClass.thy "Suc(k) ~= (k::'a::nat)"; +by (res_inst_tac [("n","k")] induct 1); +by (resolve_tac [notI] 1); +by (eresolve_tac [Suc_neq_0] 1); +by (resolve_tac [notI] 1); +by (eresolve_tac [notE] 1); +by (eresolve_tac [Suc_inject] 1); +qed "Suc_n_not_n"; + + +goal NatClass.thy "(k+m)+n = k+(m+n)"; +prths ([induct] RL [topthm()]); (*prints all 14 next states!*) +by (resolve_tac [induct] 1); +back(); +back(); +back(); +back(); +back(); +back(); + +goalw NatClass.thy [add_def] "0+n = n"; +by (resolve_tac [rec_0] 1); +qed "add_0"; + +goalw NatClass.thy [add_def] "Suc(m)+n = Suc(m+n)"; +by (resolve_tac [rec_Suc] 1); +qed "add_Suc"; + +val add_ss = FOL_ss addsimps [add_0, add_Suc]; + +goal NatClass.thy "(k+m)+n = k+(m+n)"; +by (res_inst_tac [("n","k")] induct 1); +by (simp_tac add_ss 1); +by (asm_simp_tac add_ss 1); +qed "add_assoc"; + +goal NatClass.thy "m+0 = m"; +by (res_inst_tac [("n","m")] induct 1); +by (simp_tac add_ss 1); +by (asm_simp_tac add_ss 1); +qed "add_0_right"; + +goal NatClass.thy "m+Suc(n) = Suc(m+n)"; +by (res_inst_tac [("n","m")] induct 1); +by (ALLGOALS (asm_simp_tac add_ss)); +qed "add_Suc_right"; + +val [prem] = goal NatClass.thy "(!!n. f(Suc(n)) = Suc(f(n))) ==> f(i+j) = i+f(j)"; +by (res_inst_tac [("n","i")] induct 1); +by (simp_tac add_ss 1); +by (asm_simp_tac (add_ss addsimps [prem]) 1); +result(); diff -r 934183dfc786 -r 706cfddca75c src/FOL/ex/NatClass.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/NatClass.thy Fri Sep 01 14:26:26 1995 +0200 @@ -0,0 +1,34 @@ +(* Title: FOL/ex/NatClass.thy + ID: $Id$ + Author: Markus Wenzel, TU Muenchen + +This is an abstract version of Nat.thy. Instead of axiomatizing a +single type "nat" we define the class of all these types (up to +isomorphism). + +Note: The "rec" operator had to be made 'monomorphic', because class +axioms may not contain more than one type variable. +*) + +NatClass = FOL + + +consts + "0" :: "'a" ("0") + Suc :: "'a => 'a" + rec :: "['a, 'a, ['a, 'a] => 'a] => 'a" + +axclass + nat < term + induct "[| P(0); !!x. P(x) ==> P(Suc(x)) |] ==> P(n)" + Suc_inject "Suc(m) = Suc(n) ==> m = n" + Suc_neq_0 "Suc(m) = 0 ==> R" + rec_0 "rec(0, a, f) = a" + rec_Suc "rec(Suc(m), a, f) = f(m, rec(m, a, f))" + +consts + "+" :: "['a::nat, 'a] => 'a" (infixl 60) + +defs + add_def "m + n == rec(m, n, %x y. Suc(y))" + +end