src/HOL/BNF/Examples/TreeFI.thy
author blanchet
Mon, 29 Apr 2013 09:10:49 +0200
changeset 51804 be6e703908f4
parent 51766 f19a4d0ab1bf
child 52990 6b6c4ec42024
permissions -rw-r--r--
renamed BNF "(co)data" commands to names that are closer to their final names
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49508
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
49220
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    17
lemma pre_treeFI_listF_set[simp]: "pre_treeFI_set2 (i, xs) = listF_set xs"
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    18
unfolding pre_treeFI_set2_def collect_def[abs_def] prod_set_defs
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 50516
diff changeset
    19
by (auto simp add: listF.set_map')
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
49508
1e205327f059 adapted examples to renamings
blanchet
parents: 49463
diff changeset
    21
lemma dtor[simp]: "treeFI_dtor tr = (lab tr, sub tr)"
49606
afc7f88370a8 partly ported "TreeFI" example to new syntax
blanchet
parents: 49594
diff changeset
    22
unfolding lab_def sub_def treeFI_case_def
afc7f88370a8 partly ported "TreeFI" example to new syntax
blanchet
parents: 49594
diff changeset
    23
by (metis fst_def pair_collapse snd_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    25
definition pair_fun (infixr "\<odot>" 50) where
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
  "f \<odot> g \<equiv> \<lambda>x. (f x, g x)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    27
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
(* Tree reverse:*)
49606
afc7f88370a8 partly ported "TreeFI" example to new syntax
blanchet
parents: 49594
diff changeset
    29
definition "trev \<equiv> treeFI_unfold lab (lrev o sub)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    30
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
lemma trev_simps1[simp]: "lab (trev t) = lab t"
49606
afc7f88370a8 partly ported "TreeFI" example to new syntax
blanchet
parents: 49594
diff changeset
    32
unfolding trev_def by simp
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    33
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
lemma trev_simps2[simp]: "sub (trev t) = listF_map trev (lrev (sub t))"
49606
afc7f88370a8 partly ported "TreeFI" example to new syntax
blanchet
parents: 49594
diff changeset
    35
unfolding trev_def by simp
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    36
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    37
lemma treeFI_coinduct:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    38
assumes *: "phi x y"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    39
and step: "\<And>a b. phi a b \<Longrightarrow>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
   lab a = lab b \<and>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
   lengthh (sub a) = lengthh (sub b) \<and>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    42
   (\<forall>i < lengthh (sub a). phi (nthh (sub a) i) (nthh (sub b) i))"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    43
shows "x = y"
49588
9b72d207617b export "dtor_map_coinduct" theorems, since they're used in one example
blanchet
parents: 49510
diff changeset
    44
proof (rule mp[OF treeFI.dtor_map_coinduct, of phi, OF _ *])
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
  fix a b :: "'a treeFI"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
  let ?zs = "zipp (sub a) (sub b)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    47
  let ?z = "(lab a, ?zs)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
  assume "phi a b"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    49
  with step have step': "lab a = lab b" "lengthh (sub a) = lengthh (sub b)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
    "\<forall>i < lengthh (sub a). phi (nthh (sub a) i) (nthh (sub b) i)" by auto
49508
1e205327f059 adapted examples to renamings
blanchet
parents: 49463
diff changeset
    51
  hence "pre_treeFI_map id fst ?z = treeFI_dtor a" "pre_treeFI_map id snd ?z = treeFI_dtor b"
49220
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    52
    unfolding pre_treeFI_map_def by auto
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    53
  moreover have "\<forall>(x, y) \<in> pre_treeFI_set2 ?z. phi x y"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    54
  proof safe
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    55
    fix z1 z2
49220
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    56
    assume "(z1, z2) \<in> pre_treeFI_set2 ?z"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    57
    hence "(z1, z2) \<in> listF_set ?zs" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    58
    hence "\<exists>i < lengthh ?zs. nthh ?zs i = (z1, z2)" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    59
    with step'(2) obtain i where "i < lengthh (sub a)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    60
      "nthh (sub a) i = z1" "nthh (sub b) i = z2" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
    with step'(3) show "phi z1 z2" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    63
  ultimately show "\<exists>z.
49508
1e205327f059 adapted examples to renamings
blanchet
parents: 49463
diff changeset
    64
    (pre_treeFI_map id fst z = treeFI_dtor a \<and>
1e205327f059 adapted examples to renamings
blanchet
parents: 49463
diff changeset
    65
    pre_treeFI_map id snd z = treeFI_dtor b) \<and>
49220
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    66
    (\<forall>x y. (x, y) \<in> pre_treeFI_set2 z \<longrightarrow> phi x y)" by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    67
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    68
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    69
lemma trev_trev: "trev (trev tr) = tr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    70
by (rule treeFI_coinduct[of "%a b. trev (trev b) = a"]) auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    71
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    72
end