src/HOL/BNF_Examples/TreeFI.thy
author haftmann
Thu, 01 May 2014 09:30:36 +0200
changeset 56812 baef1c110f12
parent 55075 b3d0a02a756d
child 57207 df0f8ad7cc30
permissions -rw-r--r--
centralized upper/lowercase name mangling
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55075
b3d0a02a756d dissolved BNF session
blanchet
parents: 55071
diff changeset
     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: 51766
diff 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: 53470
diff changeset
    18
primcorec trev where
e5853a648b59 use new coinduction method and primcorec in examples
traytel
parents: 53470
diff changeset
    19
  "lab (trev t) = lab t"
e5853a648b59 use new coinduction method and primcorec in examples
traytel
parents: 53470
diff 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
fe80dd7cd543 tuned whitespace
traytel
parents: 53355
diff changeset
    23
  assumes *: "phi x y"
fe80dd7cd543 tuned whitespace
traytel
parents: 53355
diff changeset
    24
  and step: "\<And>a b. phi a b \<Longrightarrow>
fe80dd7cd543 tuned whitespace
traytel
parents: 53355
diff changeset
    25
     lab a = lab b \<and>
fe80dd7cd543 tuned whitespace
traytel
parents: 53355
diff changeset
    26
     lengthh (sub a) = lengthh (sub b) \<and>
fe80dd7cd543 tuned whitespace
traytel
parents: 53355
diff changeset
    27
     (\<forall>i < lengthh (sub a). phi (nthh (sub a) i) (nthh (sub b) i))"
fe80dd7cd543 tuned whitespace
traytel
parents: 53355
diff changeset
    28
  shows "x = y"
54027
e5853a648b59 use new coinduction method and primcorec in examples
traytel
parents: 53470
diff changeset
    29
using * proof (coinduction arbitrary: x y)
e5853a648b59 use new coinduction method and primcorec in examples
traytel
parents: 53470
diff changeset
    30
  case (Eq_treeFI t1 t2)
e5853a648b59 use new coinduction method and primcorec in examples
traytel
parents: 53470
diff 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: 53470
diff changeset
    32
  have "relF phi (sub t1) (sub t2)"
53355
603e6e97c391 modernized more examples
traytel
parents: 53353
diff changeset
    33
  proof (induction "sub t1" "sub t2" arbitrary: t1 t2 rule: listF_induct2)
603e6e97c391 modernized more examples
traytel
parents: 53353
diff changeset
    34
    case (Conss x xs y ys)
603e6e97c391 modernized more examples
traytel
parents: 53353
diff changeset
    35
    note sub = Conss(2,3)[symmetric] and phi = mp[OF spec[OF Conss(4)], unfolded sub]
603e6e97c391 modernized more examples
traytel
parents: 53353
diff changeset
    36
      and IH = Conss(1)[of "Tree (lab t1) (tlF (sub t1))" "Tree (lab t2) (tlF (sub t2))",
603e6e97c391 modernized more examples
traytel
parents: 53353
diff changeset
    37
        unfolded sub, simplified]
603e6e97c391 modernized more examples
traytel
parents: 53353
diff changeset
    38
    from phi[of 0] show ?case unfolding sub by (auto intro!: IH dest: phi[simplified, OF Suc_mono])
603e6e97c391 modernized more examples
traytel
parents: 53353
diff changeset
    39
  qed simp
54027
e5853a648b59 use new coinduction method and primcorec in examples
traytel
parents: 53470
diff 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: 53470
diff 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