| author | wenzelm | 
| Mon, 04 Oct 1999 21:43:05 +0200 | |
| changeset 7699 | 09d8fd81cc1f | 
| parent 7681 | 4434adbe24a1 | 
| child 7721 | cb353d802ade | 
| permissions | -rw-r--r-- | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 1 | (* Title: HOL/ex/Acc | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 4 | Copyright 1994 University of Cambridge | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 5 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 6 | Inductive definition of acc(r) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 7 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 8 | See Ch. Paulin-Mohring, Inductive Definitions in the System Coq. | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 9 | Research Report 92-49, LIP, ENS Lyon. Dec 1992. | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 10 | *) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 11 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 12 | (*The intended introduction rule*) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 13 | val prems = goal Acc.thy | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 14 | "[| !!b. (b,a):r ==> b: acc(r) |] ==> a: acc(r)"; | 
| 5618 | 15 | by (fast_tac (claset() addIs prems @ | 
| 16 | map (rewrite_rule [pred_def]) acc.intrs) 1); | |
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 17 | qed "accI"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 18 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 19 | Goal "[| b: acc(r); (a,b): r |] ==> a: acc(r)"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 20 | by (etac acc.elim 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 21 | by (rewtac pred_def); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 22 | by (Fast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 23 | qed "acc_downward"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 24 | |
| 5163 | 25 | Goal "(b,a) : r^* ==> a : acc r --> b : acc r"; | 
| 5168 | 26 | by (etac rtrancl_induct 1); | 
| 27 | by (Blast_tac 1); | |
| 28 | by (blast_tac (claset() addDs [acc_downward]) 1); | |
| 7681 | 29 | no_qed(); | 
| 5163 | 30 | val lemma = result(); | 
| 7453 | 31 | |
| 5163 | 32 | Goal "[| a : acc r; (b,a) : r^* |] ==> b : acc r"; | 
| 5168 | 33 | by (blast_tac (claset() addDs [lemma]) 1); | 
| 5163 | 34 | qed "acc_downwards"; | 
| 35 | ||
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 36 | val [major,indhyp] = goal Acc.thy | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 37 | "[| a : acc(r); \ | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 38 | \ !!x. [| x: acc(r); ALL y. (y,x):r --> P(y) |] ==> P(x) \ | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 39 | \ |] ==> P(a)"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 40 | by (rtac (major RS acc.induct) 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 41 | by (rtac indhyp 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 42 | by (resolve_tac acc.intrs 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 43 | by (rewtac pred_def); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 44 | by (Fast_tac 2); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 45 | by (etac (Int_lower1 RS Pow_mono RS subsetD) 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 46 | qed "acc_induct"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 47 | |
| 5580 | 48 | Goal "!x. x : acc(r) ==> wf(r)"; | 
| 49 | by (rtac wfUNIVI 1); | |
| 6301 | 50 | by (deepen_tac (claset() addEs [acc_induct]) 1 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 51 | qed "acc_wfI"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 52 | |
| 5580 | 53 | Goal "wf(r) ==> x : acc(r)"; | 
| 5223 
4cb05273f764
Removal of obsolete "open" commands from heads of .ML files
 paulson parents: 
5168diff
changeset | 54 | by (etac wf_induct 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 55 | by (rtac accI 1); | 
| 5580 | 56 | by (Blast_tac 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 57 | qed "acc_wfD"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 58 | |
| 5580 | 59 | Goal "wf(r) = (!x. x : acc(r))"; | 
| 60 | by (blast_tac (claset() addIs [acc_wfI] addDs [acc_wfD]) 1); | |
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 61 | qed "wf_acc_iff"; |