62858
|
1 |
(* Title: HOL/Corec_Examples/Tests/Merge_C.thy
|
62696
|
2 |
Author: Aymeric Bouzy, Ecole polytechnique
|
|
3 |
Author: Jasmin Blanchette, Inria, LORIA, MPII
|
|
4 |
Copyright 2015, 2016
|
|
5 |
|
|
6 |
Tests theory merges.
|
|
7 |
*)
|
|
8 |
|
62726
|
9 |
section \<open>Tests Theory Merges\<close>
|
62696
|
10 |
|
|
11 |
theory Merge_C
|
|
12 |
imports Merge_A
|
|
13 |
begin
|
|
14 |
|
|
15 |
consts fc :: "'a ta \<Rightarrow> 'a ta"
|
|
16 |
consts gc :: "'a ta \<Rightarrow> 'a ta"
|
|
17 |
|
|
18 |
friend_of_corec fc :: "'a ta \<Rightarrow> 'a ta" where
|
|
19 |
"fc t = Ca (sa1 t) (sa2 t)"
|
|
20 |
sorry
|
|
21 |
|
|
22 |
friend_of_corec gb :: "'a ta \<Rightarrow> 'a ta" where
|
|
23 |
"gb t = Ca (sa1 t) (sa2 t)"
|
|
24 |
sorry
|
|
25 |
|
|
26 |
friend_of_corec gc :: "'a ta \<Rightarrow> 'a ta" where
|
|
27 |
"gc t = Ca (sa1 t) (sa2 t)"
|
|
28 |
sorry
|
|
29 |
|
|
30 |
end
|