| author | haftmann | 
| Mon, 03 Feb 2014 08:23:21 +0100 | |
| changeset 55293 | 42cf5802d36a | 
| parent 55075 | b3d0a02a756d | 
| child 57207 | df0f8ad7cc30 | 
| permissions | -rw-r--r-- | 
| 55075 | 1 | (* Title: HOL/BNF_Examples/TreeFI.thy | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 2 | Author: Dmitriy Traytel, TU Muenchen | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 3 | Author: Andrei Popescu, TU Muenchen | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 4 | Copyright 2012 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 5 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 6 | Finitely branching possibly infinite trees. | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 7 | *) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 8 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 9 | header {* Finitely Branching Possibly Infinite Trees *}
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 10 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 11 | theory TreeFI | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 12 | imports ListF | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 13 | begin | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 14 | |
| 51804 
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
 blanchet parents: 
51766diff
changeset | 15 | codatatype 'a treeFI = Tree (lab: 'a) (sub: "'a treeFI listF") | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 16 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 17 | (* Tree reverse:*) | 
| 54027 
e5853a648b59
use new coinduction method and primcorec in examples
 traytel parents: 
53470diff
changeset | 18 | primcorec trev where | 
| 
e5853a648b59
use new coinduction method and primcorec in examples
 traytel parents: 
53470diff
changeset | 19 | "lab (trev t) = lab t" | 
| 
e5853a648b59
use new coinduction method and primcorec in examples
 traytel parents: 
53470diff
changeset | 20 | | "sub (trev t) = mapF trev (lrev (sub t))" | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 21 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 22 | lemma treeFI_coinduct: | 
| 53470 | 23 | assumes *: "phi x y" | 
| 24 | and step: "\<And>a b. phi a b \<Longrightarrow> | |
| 25 | lab a = lab b \<and> | |
| 26 | lengthh (sub a) = lengthh (sub b) \<and> | |
| 27 | (\<forall>i < lengthh (sub a). phi (nthh (sub a) i) (nthh (sub b) i))" | |
| 28 | shows "x = y" | |
| 54027 
e5853a648b59
use new coinduction method and primcorec in examples
 traytel parents: 
53470diff
changeset | 29 | using * proof (coinduction arbitrary: x y) | 
| 
e5853a648b59
use new coinduction method and primcorec in examples
 traytel parents: 
53470diff
changeset | 30 | case (Eq_treeFI t1 t2) | 
| 
e5853a648b59
use new coinduction method and primcorec in examples
 traytel parents: 
53470diff
changeset | 31 | from conjunct1[OF conjunct2[OF step[OF Eq_treeFI]]] conjunct2[OF conjunct2[OF step[OF Eq_treeFI]]] | 
| 
e5853a648b59
use new coinduction method and primcorec in examples
 traytel parents: 
53470diff
changeset | 32 | have "relF phi (sub t1) (sub t2)" | 
| 53355 | 33 | proof (induction "sub t1" "sub t2" arbitrary: t1 t2 rule: listF_induct2) | 
| 34 | case (Conss x xs y ys) | |
| 35 | note sub = Conss(2,3)[symmetric] and phi = mp[OF spec[OF Conss(4)], unfolded sub] | |
| 36 | and IH = Conss(1)[of "Tree (lab t1) (tlF (sub t1))" "Tree (lab t2) (tlF (sub t2))", | |
| 37 | unfolded sub, simplified] | |
| 38 | from phi[of 0] show ?case unfolding sub by (auto intro!: IH dest: phi[simplified, OF Suc_mono]) | |
| 39 | qed simp | |
| 54027 
e5853a648b59
use new coinduction method and primcorec in examples
 traytel parents: 
53470diff
changeset | 40 | with conjunct1[OF step[OF Eq_treeFI]] show ?case by simp | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 41 | qed | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 42 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 43 | lemma trev_trev: "trev (trev tr) = tr" | 
| 54027 
e5853a648b59
use new coinduction method and primcorec in examples
 traytel parents: 
53470diff
changeset | 44 | by (coinduction arbitrary: tr rule: treeFI_coinduct) auto | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 45 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 46 | end |